make assertion labels unique already when loading a verification condition,
keep abstract view on verification conditions and provide various splitting operations on verification conditions,
allow to discharge only parts of a verification condition,
extended the command "boogie_vc" with options to consider only some assertions or to split a verification condition into its paths,
added a narrowing option to "boogie_status" (a divide-and-conquer approach for identifying the "hard" subset of assertions of a verification conditions),
added tactics "boogie", "boogie_all" and "boogie_cases",
dropped tactic "split_vc",
split example Boogie_Max into Boogie_Max (proof based on SMT) and Boogie_Max_Stepwise (proof based on metis and auto with documentation of the available Boogie commands),
dropped (mostly unused) abbreviations
--- a/src/HOL/Boogie/Boogie.thy Fri Dec 11 15:06:12 2009 +0100
+++ b/src/HOL/Boogie/Boogie.thy Fri Dec 11 15:35:29 2009 +0100
@@ -10,7 +10,7 @@
("Tools/boogie_vcs.ML")
("Tools/boogie_loader.ML")
("Tools/boogie_commands.ML")
- ("Tools/boogie_split.ML")
+ ("Tools/boogie_tactics.ML")
begin
text {*
@@ -29,6 +29,7 @@
*}
+
section {* Built-in types and functions of Boogie *}
subsection {* Labels *}
@@ -45,16 +46,6 @@
lemmas labels = assert_at_def block_at_def
-subsection {* Arrays *}
-
-abbreviation (input) boogie_select :: "('i \<Rightarrow> 'v) \<Rightarrow> 'i \<Rightarrow> 'v"
-where "boogie_select \<equiv> (\<lambda>f x. f x)"
-
-abbreviation (input) boogie_store ::
- "('i \<Rightarrow> 'v) \<Rightarrow> 'i \<Rightarrow> 'v \<Rightarrow> 'i \<Rightarrow> 'v"
-where "boogie_store \<equiv> fun_upd"
-
-
subsection {* Integer arithmetics *}
text {*
@@ -69,104 +60,6 @@
boogie_mod :: "int \<Rightarrow> int \<Rightarrow> int" (infixl "mod'_b" 70)
-subsection {* Bitvectors *}
-
-text {*
-Boogie's and Z3's built-in bitvector functions are modelled with
-functions of the HOL-Word library and the SMT theory of bitvectors.
-Every of the following bitvector functions is supported by the SMT
-binding.
-*}
-
-abbreviation (input) boogie_bv_concat ::
- "'a::len0 word \<Rightarrow> 'b::len0 word \<Rightarrow> 'c::len0 word"
-where "boogie_bv_concat \<equiv> (\<lambda>w1 w2. word_cat w1 w2)"
-
-abbreviation (input) boogie_bv_extract ::
- "nat \<Rightarrow> nat \<Rightarrow> 'a::len0 word \<Rightarrow> 'b::len0 word"
-where "boogie_bv_extract \<equiv> (\<lambda>mb lb w. slice lb w)"
-
-abbreviation (input) z3_bvnot :: "'a::len0 word \<Rightarrow> 'a word"
-where "z3_bvnot \<equiv> (\<lambda>w. NOT w)"
-
-abbreviation (input) z3_bvand :: "'a::len0 word \<Rightarrow> 'a word \<Rightarrow> 'a word"
-where "z3_bvand \<equiv> (\<lambda>w1 w2. w1 AND w2)"
-
-abbreviation (input) z3_bvor :: "'a::len0 word \<Rightarrow> 'a word \<Rightarrow> 'a word"
-where "z3_bvor \<equiv> (\<lambda>w1 w2. w1 OR w2)"
-
-abbreviation (input) z3_bvxor :: "'a::len0 word \<Rightarrow> 'a word \<Rightarrow> 'a word"
-where "z3_bvxor \<equiv> (\<lambda>w1 w2. w1 XOR w2)"
-
-abbreviation (input) z3_bvneg :: "'a::len0 word \<Rightarrow> 'a word"
-where "z3_bvneg \<equiv> (\<lambda>w. - w)"
-
-abbreviation (input) z3_bvadd :: "'a::len0 word \<Rightarrow> 'a word \<Rightarrow> 'a word"
-where "z3_bvadd \<equiv> (\<lambda>w1 w2. w1 + w2)"
-
-abbreviation (input) z3_bvsub :: "'a::len0 word \<Rightarrow> 'a word \<Rightarrow> 'a word"
-where "z3_bvsub \<equiv> (\<lambda>w1 w2. w1 - w2)"
-
-abbreviation (input) z3_bvmul :: "'a::len0 word \<Rightarrow> 'a word \<Rightarrow> 'a word"
-where "z3_bvmul \<equiv> (\<lambda>w1 w2. w1 * w2)"
-
-abbreviation (input) z3_bvudiv :: "'a::len0 word \<Rightarrow> 'a word \<Rightarrow> 'a word"
-where "z3_bvudiv \<equiv> (\<lambda>w1 w2. w1 div w2)"
-
-abbreviation (input) z3_bvurem :: "'a::len0 word \<Rightarrow> 'a word \<Rightarrow> 'a word"
-where "z3_bvurem \<equiv> (\<lambda>w1 w2. w1 mod w2)"
-
-abbreviation (input) z3_bvsdiv :: "'a::len word \<Rightarrow> 'a word \<Rightarrow> 'a word"
-where "z3_bvsdiv \<equiv> (\<lambda>w1 w2. w1 sdiv w2)"
-
-abbreviation (input) z3_bvsrem :: "'a::len word \<Rightarrow> 'a word \<Rightarrow> 'a word"
-where "z3_bvsrem \<equiv> (\<lambda>w1 w2. w1 srem w2)"
-
-abbreviation (input) z3_bvshl :: "'a::len0 word \<Rightarrow> 'a word \<Rightarrow> 'a word"
-where "z3_bvshl \<equiv> (\<lambda>w1 w2. bv_shl w1 w2)"
-
-abbreviation (input) z3_bvlshr :: "'a::len0 word \<Rightarrow> 'a word \<Rightarrow> 'a word"
-where "z3_bvlshr \<equiv> (\<lambda>w1 w2. bv_lshr w1 w2)"
-
-abbreviation (input) z3_bvashr :: "'a::len word \<Rightarrow> 'a word \<Rightarrow> 'a word"
-where "z3_bvashr \<equiv> (\<lambda>w1 w2. bv_ashr w1 w2)"
-
-abbreviation (input) z3_sign_extend :: "'a::len word \<Rightarrow> 'b::len word"
-where "z3_sign_extend \<equiv> (\<lambda>w. scast w)"
-
-abbreviation (input) z3_zero_extend :: "'a::len0 word \<Rightarrow> 'b::len0 word"
-where "z3_zero_extend \<equiv> (\<lambda>w. ucast w)"
-
-abbreviation (input) z3_rotate_left :: "nat \<Rightarrow> 'a::len0 word \<Rightarrow> 'a word"
-where "z3_rotate_left \<equiv> (\<lambda>n w. word_rotl n w)"
-
-abbreviation (input) z3_rotate_right :: "nat \<Rightarrow> 'a::len0 word \<Rightarrow> 'a word"
-where "z3_rotate_right \<equiv> (\<lambda>n w. word_rotr n w)"
-
-abbreviation (input) z3_bvult :: "'a::len0 word \<Rightarrow> 'a word \<Rightarrow> bool"
-where "z3_bvult \<equiv> (\<lambda>w1 w2. w1 < w2)"
-
-abbreviation (input) z3_bvule :: "'a::len0 word \<Rightarrow> 'a word \<Rightarrow> bool"
-where "z3_bvule \<equiv> (\<lambda>w1 w2. w1 \<le> w2)"
-
-abbreviation (input) z3_bvugt :: "'a::len0 word \<Rightarrow> 'a word \<Rightarrow> bool"
-where "z3_bvugt \<equiv> (\<lambda>w1 w2. w1 > w2)"
-
-abbreviation (input) z3_bvuge :: "'a::len0 word \<Rightarrow> 'a word \<Rightarrow> bool"
-where "z3_bvuge \<equiv> (\<lambda>w1 w2. w1 \<ge> w2)"
-
-abbreviation (input) z3_bvslt :: "'a::len word \<Rightarrow> 'a word \<Rightarrow> bool"
-where "z3_bvslt \<equiv> (\<lambda>w1 w2. w1 <s w2)"
-
-abbreviation (input) z3_bvsle :: "'a::len word \<Rightarrow> 'a word \<Rightarrow> bool"
-where "z3_bvsle \<equiv> (\<lambda>w1 w2. w1 <=s w2)"
-
-abbreviation (input) z3_bvsgt :: "'a::len word \<Rightarrow> 'a word \<Rightarrow> bool"
-where "z3_bvsgt \<equiv> (\<lambda>w1 w2. w2 <s w1)"
-
-abbreviation (input) z3_bvsge :: "'a::len word \<Rightarrow> 'a word \<Rightarrow> bool"
-where "z3_bvsge \<equiv> (\<lambda>w1 w2. w2 <=s w1)"
-
section {* Boogie environment *}
@@ -187,6 +80,7 @@
*}
+
section {* Setup *}
ML {*
@@ -199,26 +93,12 @@
*}
setup Boogie_Axioms.setup
-text {*
-Opening a Boogie environment causes the following list of theorems to be
-enhanced by all theorems found in Boogie_Axioms.
-*}
-ML {*
-structure Split_VC_SMT_Rules = Named_Thms
-(
- val name = "split_vc_smt"
- val description =
- "theorems given to the SMT sub-tactic of the split_vc method"
-)
-*}
-setup Split_VC_SMT_Rules.setup
-
use "Tools/boogie_vcs.ML"
use "Tools/boogie_loader.ML"
use "Tools/boogie_commands.ML"
setup Boogie_Commands.setup
-use "Tools/boogie_split.ML"
-setup Boogie_Split.setup
+use "Tools/boogie_tactics.ML"
+setup Boogie_Tactics.setup
end
--- a/src/HOL/Boogie/Examples/Boogie_Dijkstra.thy Fri Dec 11 15:06:12 2009 +0100
+++ b/src/HOL/Boogie/Examples/Boogie_Dijkstra.thy Fri Dec 11 15:35:29 2009 +0100
@@ -82,10 +82,9 @@
boogie_open "~~/src/HOL/Boogie/Examples/Boogie_Dijkstra"
boogie_vc Dijkstra
- unfolding labels
+ using [[z3_proofs=true]]
using [[smt_cert="~~/src/HOL/Boogie/Examples/cert/Boogie_Dijkstra"]]
- using [[z3_proofs=true]]
- by (smt boogie)
+ by boogie
boogie_end
--- a/src/HOL/Boogie/Examples/Boogie_Max.thy Fri Dec 11 15:06:12 2009 +0100
+++ b/src/HOL/Boogie/Examples/Boogie_Max.thy Fri Dec 11 15:35:29 2009 +0100
@@ -1,4 +1,4 @@
-(* Title: HOL/Boogie/Examples/Boogie_Dijkstra.thy
+(* Title: HOL/Boogie/Examples/Boogie_Max.thy
Author: Sascha Boehme, TU Muenchen
*)
@@ -38,40 +38,10 @@
boogie_open "~~/src/HOL/Boogie/Examples/Boogie_Max"
-text {*
-Approach 1: Discharge the verification condition fully automatically by SMT:
-*}
boogie_vc max
- unfolding labels
- using [[smt_cert="~~/src/HOL/Boogie/Examples/cert/Boogie_max"]]
using [[z3_proofs=true]]
- by (smt boogie)
-
-text {*
-Approach 2: Split the verification condition, try to solve the splinters by
-a selection of automated proof tools, and show the remaining subgoals by an
-explicit proof. This approach is most useful in an interactive debug-and-fix
-mode.
-*}
-boogie_vc max
-proof (split_vc (verbose) try: fast simp)
- print_cases
- case L_14_5c
- thus ?case by (metis less_le_not_le zle_add1_eq_le zless_linear)
-next
- case L_14_5b
- thus ?case by (metis less_le_not_le xt1(10) zle_linear zless_add1_eq)
-next
- case L_14_5a
- note max0 = `max0 = array 0`
- {
- fix i :: int
- assume "0 \<le> i \<and> i < 1"
- hence "i = 0" by simp
- with max0 have "array i \<le> max0" by simp
- }
- thus ?case by simp
-qed
+ using [[smt_cert="~~/src/HOL/Boogie/Examples/cert/Boogie_max"]]
+ by boogie
boogie_end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Boogie/Examples/Boogie_Max_Stepwise.thy Fri Dec 11 15:35:29 2009 +0100
@@ -0,0 +1,98 @@
+(* Title: HOL/Boogie/Examples/Boogie_Max_Stepwise.thy
+ Author: Sascha Boehme, TU Muenchen
+*)
+
+header {* Boogie example: get the greatest element of a list *}
+
+theory Boogie_Max_Stepwise
+imports Boogie
+begin
+
+text {*
+We prove correct the verification condition generated from the following
+Boogie code:
+
+\begin{verbatim}
+procedure max(array : [int]int, length : int)
+ returns (max : int);
+ requires (0 < length);
+ ensures (forall i: int :: 0 <= i && i < length ==> array[i] <= max);
+ ensures (exists i: int :: 0 <= i && i < length ==> array[i] == max);
+
+implementation max(array : [int]int, length : int) returns (max : int)
+{
+ var p : int, k : int;
+ max := array[0];
+ k := 0;
+ p := 1;
+ while (p < length)
+ invariant (forall i: int :: 0 <= i && i < p ==> array[i] <= max);
+ invariant (array[k] == max);
+ {
+ if (array[p] > max) { max := array[p]; k := p; }
+ p := p + 1;
+ }
+}
+\end{verbatim}
+*}
+
+boogie_open "~~/src/HOL/Boogie/Examples/Boogie_Max"
+
+boogie_status -- {* gives an overview of all verification conditions *}
+
+boogie_status max -- {* shows the names of all unproved assertions
+ of the verification condition @{text max} *}
+
+boogie_status (full) max -- {* shows the state of all assertions
+ of the verification condition @{text max} *}
+
+
+text {* Let's find out which assertions of @{text max} are hard to prove: *}
+
+boogie_status (narrow timeout: 4) max
+ (try: (simp add: labels, (fast | blast)?))
+ -- {* The argument @{text timeout} is optional, restricting the runtime of
+ each narrowing step by the given number of seconds. *}
+ -- {* The tag @{text try} expects a method to be applied at each narrowing
+ step. Note that different methods may be composed to one method by
+ a language similar to regular expressions. *}
+
+
+text {* Now, let's prove the three hard assertions of @{text max}: *}
+
+boogie_vc (only: L_14_5c L_14_5b L_14_5a) max
+proof boogie_cases
+ case L_14_5c
+ thus ?case by (metis less_le_not_le zle_add1_eq_le zless_linear)
+next
+ case L_14_5b
+ thus ?case by (metis less_le_not_le xt1(10) zle_linear zless_add1_eq)
+next
+ case L_14_5a
+ note max0 = `max0 = array 0`
+ {
+ fix i :: int
+ assume "0 \<le> i \<and> i < 1"
+ hence "i = 0" by simp
+ with max0 have "array i \<le> max0" by simp
+ }
+ thus ?case by simp
+qed
+
+
+boogie_status max -- {* the above proved assertions are not shown anymore *}
+
+boogie_status (full) max -- {* states the above proved assertions as proved
+ and all other assertions as not proved *}
+
+
+text {* Now we prove the remaining assertions all at once: *}
+
+boogie_vc max by (auto simp add: labels)
+
+
+boogie_status -- {* @{text max} has been completely proved *}
+
+boogie_end
+
+end
--- a/src/HOL/Boogie/Examples/ROOT.ML Fri Dec 11 15:06:12 2009 +0100
+++ b/src/HOL/Boogie/Examples/ROOT.ML Fri Dec 11 15:35:29 2009 +0100
@@ -1,1 +1,1 @@
-use_thys ["Boogie_Max", "Boogie_Dijkstra", "VCC_Max"];
+use_thys ["Boogie_Max_Stepwise", "Boogie_Max", "Boogie_Dijkstra", "VCC_Max"];
--- a/src/HOL/Boogie/Examples/VCC_Max.thy Fri Dec 11 15:06:12 2009 +0100
+++ b/src/HOL/Boogie/Examples/VCC_Max.thy Fri Dec 11 15:35:29 2009 +0100
@@ -49,10 +49,9 @@
boogie_status
boogie_vc maximum
- unfolding labels
+ using [[z3_proofs=true]]
using [[smt_cert="~~/src/HOL/Boogie/Examples/cert/VCC_maximum"]]
- using [[z3_proofs=true]]
- by (smt boogie)
+ by boogie
boogie_end
--- a/src/HOL/Boogie/Examples/cert/Boogie_Dijkstra Fri Dec 11 15:06:12 2009 +0100
+++ b/src/HOL/Boogie/Examples/cert/Boogie_Dijkstra Fri Dec 11 15:35:29 2009 +0100
@@ -38,6 +38,6 @@
:assumption (< 0 uf_9)
:assumption (forall (?x20 T2) (?x21 T2) (implies (= ?x20 ?x21) (= (uf_10 (uf_1 ?x20 ?x21)) 0)))
:assumption (forall (?x22 T2) (?x23 T2) (implies (not (= ?x22 ?x23)) (< 0 (uf_10 (uf_1 ?x22 ?x23)))))
-:assumption (not (implies true (implies true (implies (forall (?x24 T2) (implies (= ?x24 uf_11) (= (uf_12 ?x24) 0))) (implies (forall (?x25 T2) (implies (not (= ?x25 uf_11)) (= (uf_12 ?x25) uf_9))) (implies (forall (?x26 T2) (not (up_13 ?x26))) (implies true (and (= (uf_12 uf_11) 0) (implies (= (uf_12 uf_11) 0) (and (forall (?x27 T2) (<= 0 (uf_12 ?x27))) (implies (forall (?x28 T2) (<= 0 (uf_12 ?x28))) (and (forall (?x29 T2) (?x30 T2) (implies (and (not (up_13 ?x30)) (up_13 ?x29)) (<= (uf_12 ?x29) (uf_12 ?x30)))) (implies (forall (?x31 T2) (?x32 T2) (implies (and (not (up_13 ?x32)) (up_13 ?x31)) (<= (uf_12 ?x31) (uf_12 ?x32)))) (and (forall (?x33 T2) (?x34 T2) (implies (and (up_13 ?x34) (< (uf_10 (uf_1 ?x34 ?x33)) uf_9)) (<= (uf_12 ?x33) (+ (uf_12 ?x34) (uf_10 (uf_1 ?x34 ?x33)))))) (implies (forall (?x35 T2) (?x36 T2) (implies (and (up_13 ?x36) (< (uf_10 (uf_1 ?x36 ?x35)) uf_9)) (<= (uf_12 ?x35) (+ (uf_12 ?x36) (uf_10 (uf_1 ?x36 ?x35)))))) (and (forall (?x37 T2) (implies (and (not (= ?x37 uf_11)) (< (uf_12 ?x37) uf_9)) (exists (?x38 T2) (and (< (uf_12 ?x38) (uf_12 ?x37)) (and (up_13 ?x38) (= (uf_12 ?x37) (+ (uf_12 ?x38) (uf_10 (uf_1 ?x38 ?x37))))))))) (implies (forall (?x39 T2) (implies (and (not (= ?x39 uf_11)) (< (uf_12 ?x39) uf_9)) (exists (?x40 T2) (and (< (uf_12 ?x40) (uf_12 ?x39)) (and (up_13 ?x40) (= (uf_12 ?x39) (+ (uf_12 ?x40) (uf_10 (uf_1 ?x40 ?x39))))))))) (implies true (implies true (implies (= (uf_4 uf_14 uf_11) 0) (implies (forall (?x41 T2) (<= 0 (uf_4 uf_14 ?x41))) (implies (forall (?x42 T2) (?x43 T2) (implies (and (not (up_6 uf_15 ?x43)) (up_6 uf_15 ?x42)) (<= (uf_4 uf_14 ?x42) (uf_4 uf_14 ?x43)))) (implies (forall (?x44 T2) (?x45 T2) (implies (and (up_6 uf_15 ?x45) (< (uf_10 (uf_1 ?x45 ?x44)) uf_9)) (<= (uf_4 uf_14 ?x44) (+ (uf_4 uf_14 ?x45) (uf_10 (uf_1 ?x45 ?x44)))))) (implies (forall (?x46 T2) (implies (and (not (= ?x46 uf_11)) (< (uf_4 uf_14 ?x46) uf_9)) (exists (?x47 T2) (and (< (uf_4 uf_14 ?x47) (uf_4 uf_14 ?x46)) (and (up_6 uf_15 ?x47) (= (uf_4 uf_14 ?x46) (+ (uf_4 uf_14 ?x47) (uf_10 (uf_1 ?x47 ?x46))))))))) (implies true (and (implies true (implies true (implies (not (exists (?x48 T2) (and (not (up_6 uf_15 ?x48)) (< (uf_4 uf_14 ?x48) uf_9)))) (implies true (implies true (implies (= uf_16 uf_15) (implies (= uf_17 uf_18) (implies (= uf_19 uf_14) (implies (= uf_20 uf_21) (implies true (and (forall (?x49 T2) (implies (and (not (= ?x49 uf_11)) (< (uf_4 uf_19 ?x49) uf_9)) (exists (?x50 T2) (and (< (uf_4 uf_19 ?x50) (uf_4 uf_19 ?x49)) (= (uf_4 uf_19 ?x49) (+ (uf_4 uf_19 ?x50) (uf_10 (uf_1 ?x50 ?x49)))))))) (implies (forall (?x51 T2) (implies (and (not (= ?x51 uf_11)) (< (uf_4 uf_19 ?x51) uf_9)) (exists (?x52 T2) (and (< (uf_4 uf_19 ?x52) (uf_4 uf_19 ?x51)) (= (uf_4 uf_19 ?x51) (+ (uf_4 uf_19 ?x52) (uf_10 (uf_1 ?x52 ?x51)))))))) (and (forall (?x53 T2) (?x54 T2) (implies (and (< (uf_4 uf_19 ?x54) uf_9) (< (uf_10 (uf_1 ?x54 ?x53)) uf_9)) (<= (uf_4 uf_19 ?x53) (+ (uf_4 uf_19 ?x54) (uf_10 (uf_1 ?x54 ?x53)))))) (implies (forall (?x55 T2) (?x56 T2) (implies (and (< (uf_4 uf_19 ?x56) uf_9) (< (uf_10 (uf_1 ?x56 ?x55)) uf_9)) (<= (uf_4 uf_19 ?x55) (+ (uf_4 uf_19 ?x56) (uf_10 (uf_1 ?x56 ?x55)))))) (and (= (uf_4 uf_19 uf_11) 0) (implies (= (uf_4 uf_19 uf_11) 0) true)))))))))))))))) (implies true (implies true (implies (exists (?x57 T2) (and (not (up_6 uf_15 ?x57)) (< (uf_4 uf_14 ?x57) uf_9))) (implies (not (up_6 uf_15 uf_22)) (implies (< (uf_4 uf_14 uf_22) uf_9) (implies (forall (?x58 T2) (implies (not (up_6 uf_15 ?x58)) (<= (uf_4 uf_14 uf_22) (uf_4 uf_14 ?x58)))) (implies (= uf_23 (uf_7 uf_15 uf_22 uf_8)) (implies (forall (?x59 T2) (implies (and (< (uf_10 (uf_1 uf_22 ?x59)) uf_9) (< (+ (uf_4 uf_14 uf_22) (uf_10 (uf_1 uf_22 ?x59))) (uf_4 uf_14 ?x59))) (= (uf_24 ?x59) (+ (uf_4 uf_14 uf_22) (uf_10 (uf_1 uf_22 ?x59)))))) (implies (forall (?x60 T2) (implies (not (and (< (uf_10 (uf_1 uf_22 ?x60)) uf_9) (< (+ (uf_4 uf_14 uf_22) (uf_10 (uf_1 uf_22 ?x60))) (uf_4 uf_14 ?x60)))) (= (uf_24 ?x60) (uf_4 uf_14 ?x60)))) (and (forall (?x61 T2) (<= (uf_24 ?x61) (uf_4 uf_14 ?x61))) (implies (forall (?x62 T2) (<= (uf_24 ?x62) (uf_4 uf_14 ?x62))) (and (forall (?x63 T2) (implies (up_6 uf_23 ?x63) (= (uf_24 ?x63) (uf_4 uf_14 ?x63)))) (implies (forall (?x64 T2) (implies (up_6 uf_23 ?x64) (= (uf_24 ?x64) (uf_4 uf_14 ?x64)))) (implies true (implies true (and (= (uf_24 uf_11) 0) (implies (= (uf_24 uf_11) 0) (and (forall (?x65 T2) (<= 0 (uf_24 ?x65))) (implies (forall (?x66 T2) (<= 0 (uf_24 ?x66))) (and (forall (?x67 T2) (?x68 T2) (implies (and (not (up_6 uf_23 ?x68)) (up_6 uf_23 ?x67)) (<= (uf_24 ?x67) (uf_24 ?x68)))) (implies (forall (?x69 T2) (?x70 T2) (implies (and (not (up_6 uf_23 ?x70)) (up_6 uf_23 ?x69)) (<= (uf_24 ?x69) (uf_24 ?x70)))) (and (forall (?x71 T2) (?x72 T2) (implies (and (up_6 uf_23 ?x72) (< (uf_10 (uf_1 ?x72 ?x71)) uf_9)) (<= (uf_24 ?x71) (+ (uf_24 ?x72) (uf_10 (uf_1 ?x72 ?x71)))))) (implies (forall (?x73 T2) (?x74 T2) (implies (and (up_6 uf_23 ?x74) (< (uf_10 (uf_1 ?x74 ?x73)) uf_9)) (<= (uf_24 ?x73) (+ (uf_24 ?x74) (uf_10 (uf_1 ?x74 ?x73)))))) (and (forall (?x75 T2) (implies (and (not (= ?x75 uf_11)) (< (uf_24 ?x75) uf_9)) (exists (?x76 T2) (and (< (uf_24 ?x76) (uf_24 ?x75)) (and (up_6 uf_23 ?x76) (= (uf_24 ?x75) (+ (uf_24 ?x76) (uf_10 (uf_1 ?x76 ?x75))))))))) (implies (forall (?x77 T2) (implies (and (not (= ?x77 uf_11)) (< (uf_24 ?x77) uf_9)) (exists (?x78 T2) (and (< (uf_24 ?x78) (uf_24 ?x77)) (and (up_6 uf_23 ?x78) (= (uf_24 ?x77) (+ (uf_24 ?x78) (uf_10 (uf_1 ?x78 ?x77))))))))) (implies false true))))))))))))))))))))))))))))))))))))))))))))))))))))
+:assumption (not (implies (and true (and (forall (?x24 T2) (implies (= ?x24 uf_11) (= (uf_12 ?x24) 0))) (and (forall (?x25 T2) (implies (not (= ?x25 uf_11)) (= (uf_12 ?x25) uf_9))) (forall (?x26 T2) (not (up_13 ?x26)))))) (and (= (uf_12 uf_11) 0) (implies (= (uf_12 uf_11) 0) (and (forall (?x27 T2) (<= 0 (uf_12 ?x27))) (implies (forall (?x28 T2) (<= 0 (uf_12 ?x28))) (and (forall (?x29 T2) (?x30 T2) (implies (and (not (up_13 ?x30)) (up_13 ?x29)) (<= (uf_12 ?x29) (uf_12 ?x30)))) (implies (forall (?x31 T2) (?x32 T2) (implies (and (not (up_13 ?x32)) (up_13 ?x31)) (<= (uf_12 ?x31) (uf_12 ?x32)))) (and (forall (?x33 T2) (?x34 T2) (implies (and (up_13 ?x34) (< (uf_10 (uf_1 ?x34 ?x33)) uf_9)) (<= (uf_12 ?x33) (+ (uf_12 ?x34) (uf_10 (uf_1 ?x34 ?x33)))))) (implies (forall (?x35 T2) (?x36 T2) (implies (and (up_13 ?x36) (< (uf_10 (uf_1 ?x36 ?x35)) uf_9)) (<= (uf_12 ?x35) (+ (uf_12 ?x36) (uf_10 (uf_1 ?x36 ?x35)))))) (and (forall (?x37 T2) (implies (and (not (= ?x37 uf_11)) (< (uf_12 ?x37) uf_9)) (exists (?x38 T2) (and (< (uf_12 ?x38) (uf_12 ?x37)) (and (up_13 ?x38) (= (uf_12 ?x37) (+ (uf_12 ?x38) (uf_10 (uf_1 ?x38 ?x37))))))))) (implies (and (forall (?x39 T2) (implies (and (not (= ?x39 uf_11)) (< (uf_12 ?x39) uf_9)) (exists (?x40 T2) (and (< (uf_12 ?x40) (uf_12 ?x39)) (and (up_13 ?x40) (= (uf_12 ?x39) (+ (uf_12 ?x40) (uf_10 (uf_1 ?x40 ?x39))))))))) (and true (and (= (uf_4 uf_14 uf_11) 0) (and (forall (?x41 T2) (<= 0 (uf_4 uf_14 ?x41))) (and (forall (?x42 T2) (?x43 T2) (implies (and (not (up_6 uf_15 ?x43)) (up_6 uf_15 ?x42)) (<= (uf_4 uf_14 ?x42) (uf_4 uf_14 ?x43)))) (and (forall (?x44 T2) (?x45 T2) (implies (and (up_6 uf_15 ?x45) (< (uf_10 (uf_1 ?x45 ?x44)) uf_9)) (<= (uf_4 uf_14 ?x44) (+ (uf_4 uf_14 ?x45) (uf_10 (uf_1 ?x45 ?x44)))))) (forall (?x46 T2) (implies (and (not (= ?x46 uf_11)) (< (uf_4 uf_14 ?x46) uf_9)) (exists (?x47 T2) (and (< (uf_4 uf_14 ?x47) (uf_4 uf_14 ?x46)) (and (up_6 uf_15 ?x47) (= (uf_4 uf_14 ?x46) (+ (uf_4 uf_14 ?x47) (uf_10 (uf_1 ?x47 ?x46))))))))))))))) (and (implies (and true (and (not (exists (?x48 T2) (and (not (up_6 uf_15 ?x48)) (< (uf_4 uf_14 ?x48) uf_9)))) (and (= uf_16 uf_15) (and (= uf_17 uf_18) (and (= uf_19 uf_14) (= uf_20 uf_21)))))) (and (forall (?x49 T2) (implies (and (not (= ?x49 uf_11)) (< (uf_4 uf_19 ?x49) uf_9)) (exists (?x50 T2) (and (< (uf_4 uf_19 ?x50) (uf_4 uf_19 ?x49)) (= (uf_4 uf_19 ?x49) (+ (uf_4 uf_19 ?x50) (uf_10 (uf_1 ?x50 ?x49)))))))) (implies (forall (?x51 T2) (implies (and (not (= ?x51 uf_11)) (< (uf_4 uf_19 ?x51) uf_9)) (exists (?x52 T2) (and (< (uf_4 uf_19 ?x52) (uf_4 uf_19 ?x51)) (= (uf_4 uf_19 ?x51) (+ (uf_4 uf_19 ?x52) (uf_10 (uf_1 ?x52 ?x51)))))))) (and (forall (?x53 T2) (?x54 T2) (implies (and (< (uf_4 uf_19 ?x54) uf_9) (< (uf_10 (uf_1 ?x54 ?x53)) uf_9)) (<= (uf_4 uf_19 ?x53) (+ (uf_4 uf_19 ?x54) (uf_10 (uf_1 ?x54 ?x53)))))) (implies (forall (?x55 T2) (?x56 T2) (implies (and (< (uf_4 uf_19 ?x56) uf_9) (< (uf_10 (uf_1 ?x56 ?x55)) uf_9)) (<= (uf_4 uf_19 ?x55) (+ (uf_4 uf_19 ?x56) (uf_10 (uf_1 ?x56 ?x55)))))) (= (uf_4 uf_19 uf_11) 0)))))) (implies (and true (and (exists (?x57 T2) (and (not (up_6 uf_15 ?x57)) (< (uf_4 uf_14 ?x57) uf_9))) (and (not (up_6 uf_15 uf_22)) (and (< (uf_4 uf_14 uf_22) uf_9) (and (forall (?x58 T2) (implies (not (up_6 uf_15 ?x58)) (<= (uf_4 uf_14 uf_22) (uf_4 uf_14 ?x58)))) (and (= uf_23 (uf_7 uf_15 uf_22 uf_8)) (and (forall (?x59 T2) (implies (and (< (uf_10 (uf_1 uf_22 ?x59)) uf_9) (< (+ (uf_4 uf_14 uf_22) (uf_10 (uf_1 uf_22 ?x59))) (uf_4 uf_14 ?x59))) (= (uf_24 ?x59) (+ (uf_4 uf_14 uf_22) (uf_10 (uf_1 uf_22 ?x59)))))) (forall (?x60 T2) (implies (not (and (< (uf_10 (uf_1 uf_22 ?x60)) uf_9) (< (+ (uf_4 uf_14 uf_22) (uf_10 (uf_1 uf_22 ?x60))) (uf_4 uf_14 ?x60)))) (= (uf_24 ?x60) (uf_4 uf_14 ?x60))))))))))) (and (forall (?x61 T2) (<= (uf_24 ?x61) (uf_4 uf_14 ?x61))) (implies (forall (?x62 T2) (<= (uf_24 ?x62) (uf_4 uf_14 ?x62))) (and (forall (?x63 T2) (implies (up_6 uf_23 ?x63) (= (uf_24 ?x63) (uf_4 uf_14 ?x63)))) (implies (forall (?x64 T2) (implies (up_6 uf_23 ?x64) (= (uf_24 ?x64) (uf_4 uf_14 ?x64)))) (and (= (uf_24 uf_11) 0) (implies (= (uf_24 uf_11) 0) (and (forall (?x65 T2) (<= 0 (uf_24 ?x65))) (implies (forall (?x66 T2) (<= 0 (uf_24 ?x66))) (and (forall (?x67 T2) (?x68 T2) (implies (and (not (up_6 uf_23 ?x68)) (up_6 uf_23 ?x67)) (<= (uf_24 ?x67) (uf_24 ?x68)))) (implies (forall (?x69 T2) (?x70 T2) (implies (and (not (up_6 uf_23 ?x70)) (up_6 uf_23 ?x69)) (<= (uf_24 ?x69) (uf_24 ?x70)))) (and (forall (?x71 T2) (?x72 T2) (implies (and (up_6 uf_23 ?x72) (< (uf_10 (uf_1 ?x72 ?x71)) uf_9)) (<= (uf_24 ?x71) (+ (uf_24 ?x72) (uf_10 (uf_1 ?x72 ?x71)))))) (implies (forall (?x73 T2) (?x74 T2) (implies (and (up_6 uf_23 ?x74) (< (uf_10 (uf_1 ?x74 ?x73)) uf_9)) (<= (uf_24 ?x73) (+ (uf_24 ?x74) (uf_10 (uf_1 ?x74 ?x73)))))) (forall (?x75 T2) (implies (and (not (= ?x75 uf_11)) (< (uf_24 ?x75) uf_9)) (exists (?x76 T2) (and (< (uf_24 ?x76) (uf_24 ?x75)) (and (up_6 uf_23 ?x76) (= (uf_24 ?x75) (+ (uf_24 ?x76) (uf_10 (uf_1 ?x76 ?x75)))))))))))))))))))))))))))))))))))
:formula true
)
--- a/src/HOL/Boogie/Examples/cert/Boogie_Dijkstra.proof Fri Dec 11 15:06:12 2009 +0100
+++ b/src/HOL/Boogie/Examples/cert/Boogie_Dijkstra.proof Fri Dec 11 15:35:29 2009 +0100
@@ -1,6631 +1,6631 @@
#2 := false
-decl up_6 :: (-> T4 T2 bool)
-decl ?x47!7 :: (-> T2 T2)
-decl ?x75!20 :: T2
-#2235 := ?x75!20
-#5912 := (?x47!7 ?x75!20)
-decl uf_23 :: T4
-#187 := uf_23
-#16889 := (up_6 uf_23 #5912)
-decl uf_2 :: (-> T1 T2)
+#52 := 0::int
+decl uf_10 :: (-> T1 int)
decl uf_1 :: (-> T2 T2 T1)
decl uf_3 :: (-> T1 T2)
+decl ?x71!19 :: T2
+#2002 := ?x71!19
+decl ?x72!18 :: T2
+#2001 := ?x72!18
+#2007 := (uf_1 ?x72!18 ?x71!19)
+#11538 := (uf_3 #2007)
decl uf_22 :: T2
-#179 := uf_22
-#4603 := (uf_1 uf_22 uf_22)
-#10571 := (uf_3 #4603)
-#15148 := (uf_1 #10571 ?x75!20)
-#15931 := (uf_3 #15148)
-#16881 := (uf_1 #5912 #15931)
-#19932 := (uf_2 #16881)
+#182 := uf_22
+#18065 := (uf_1 uf_22 #11538)
+#18066 := (uf_10 #18065)
+decl uf_24 :: (-> T2 int)
+#18030 := (uf_24 #11538)
+#952 := -1::int
+#18049 := (* -1::int #18030)
+#18091 := (+ #18049 #18066)
+decl uf_4 :: (-> T3 T2 int)
+decl uf_14 :: T3
+#107 := uf_14
+#185 := (uf_4 uf_14 uf_22)
+#18092 := (+ #185 #18091)
+#18113 := (>= #18092 0::int)
+#18070 := (* -1::int #18066)
+decl uf_9 :: int
+#53 := uf_9
+#18074 := (+ uf_9 #18070)
+#18075 := (<= #18074 0::int)
+#17499 := (not #18075)
+#2008 := (uf_10 #2007)
+#2011 := (* -1::int #2008)
+#2012 := (+ uf_9 #2011)
+#2013 := (<= #2012 0::int)
+#2014 := (not #2013)
+#2005 := (uf_24 ?x72!18)
+#2543 := (* -1::int #2005)
+#2544 := (+ #2543 #2011)
+#2003 := (uf_24 ?x71!19)
+#2545 := (+ #2003 #2544)
+#2546 := (<= #2545 0::int)
+decl up_6 :: (-> T4 T2 bool)
+decl uf_23 :: T4
+#190 := uf_23
+#2015 := (up_6 uf_23 ?x72!18)
+#3106 := (not #2015)
+#3121 := (or #2013 #3106 #2546)
+#3126 := (not #3121)
+decl ?x75!20 :: T2
+#2031 := ?x75!20
+#11 := (:var 0 T2)
+#2035 := (uf_1 #11 ?x75!20)
+#4121 := (pattern #2035)
+#199 := (uf_24 #11)
+#4062 := (pattern #199)
+#216 := (up_6 uf_23 #11)
+#4087 := (pattern #216)
+#2036 := (uf_10 #2035)
+#2032 := (uf_24 ?x75!20)
+#2033 := (* -1::int #2032)
+#2581 := (+ #2033 #2036)
+#2582 := (+ #199 #2581)
+#2585 := (= #2582 0::int)
+#3151 := (not #2585)
+#2034 := (+ #199 #2033)
+#2039 := (>= #2034 0::int)
+#223 := (not #216)
+#3152 := (or #223 #2039 #3151)
+#4122 := (forall (vars (?x76 T2)) (:pat #4087 #4062 #4121) #3152)
+#4127 := (not #4122)
+#10 := (:var 1 T2)
+#90 := (uf_1 #11 #10)
+#3916 := (pattern #90)
+#226 := (uf_24 #10)
+#1337 := (* -1::int #226)
+#1338 := (+ #199 #1337)
+#91 := (uf_10 #90)
+#1363 := (+ #91 #1338)
+#1361 := (>= #1363 0::int)
+#967 := (* -1::int #91)
+#968 := (+ uf_9 #967)
+#969 := (<= #968 0::int)
+#3143 := (or #223 #969 #1361)
+#4113 := (forall (vars (?x71 T2) (?x72 T2)) (:pat #3916) #3143)
+#4118 := (not #4113)
+decl uf_11 :: T2
+#64 := uf_11
+#2557 := (= uf_11 ?x75!20)
+#2043 := (+ uf_9 #2033)
+#2044 := (<= #2043 0::int)
+#4130 := (or #2044 #2557 #4118 #4127)
+decl ?x47!7 :: (-> T2 T2)
+#11097 := (?x47!7 ?x75!20)
+decl uf_15 :: T4
+#113 := uf_15
+#11136 := (up_6 uf_15 #11097)
+#11143 := (uf_4 uf_14 #11097)
+#11144 := (* -1::int #11143)
+#10301 := (uf_4 uf_14 ?x75!20)
+#11145 := (+ #10301 #11144)
+#11146 := (<= #11145 0::int)
+#11137 := (not #11136)
+#11147 := (uf_1 #11097 ?x75!20)
+#11148 := (uf_10 #11147)
+#11149 := (* -1::int #11148)
+#11150 := (+ #11144 #11149)
+#11133 := (+ #10301 #11150)
+#11134 := (= #11133 0::int)
+#11135 := (not #11134)
+#11142 := (or #11135 #11137 #11146)
+#11168 := (not #11142)
+#10403 := (* -1::int #10301)
+#11139 := (+ uf_9 #10403)
+#11140 := (<= #11139 0::int)
+#33430 := (not #11140)
+#2045 := (not #2044)
+#4133 := (not #4130)
+#25286 := [hypothesis]: #4133
+#3757 := (or #4130 #2045)
+#3755 := [def-axiom]: #3757
+#25348 := [unit-resolution #3755 #25286]: #2045
+#10404 := (+ #2032 #10403)
+#7777 := (>= #10404 0::int)
+#10311 := (= #2032 #10301)
+#3601 := (or #4130 #4122)
+#3749 := [def-axiom]: #3601
+#25346 := [unit-resolution #3749 #25286]: #4122
+#30592 := (or #10311 #4127)
+#5832 := (uf_1 uf_22 uf_22)
+#6904 := (uf_3 #5832)
+#8027 := (uf_24 #6904)
+#7883 := (* -1::int #8027)
+#10078 := (+ #2032 #7883)
+#10697 := (<= #10078 0::int)
+#30559 := (not #10697)
+#10700 := (uf_1 #6904 ?x75!20)
+#10701 := (uf_10 #10700)
+#30273 := (<= #10701 0::int)
+#30274 := (not #30273)
+#13994 := (= ?x75!20 #6904)
+#30199 := (not #13994)
+#11031 := (up_6 uf_15 ?x75!20)
+#14001 := (or #11031 #13994)
+#30202 := (not #14001)
decl uf_7 :: (-> T4 T2 T5 T4)
decl uf_8 :: T5
#33 := uf_8
-decl uf_15 :: T4
-#110 := uf_15
-#11533 := (uf_7 uf_15 #10571 uf_8)
-#27162 := (up_6 #11533 #19932)
-#27198 := (not #27162)
-#16890 := (not #16889)
-#27320 := (iff #16890 #27198)
-#27318 := (iff #16889 #27162)
-#27316 := (iff #27162 #16889)
-#27304 := (= #19932 #5912)
-#20977 := (= #5912 #19932)
-#11 := (:var 0 T2)
-#10 := (:var 1 T2)
-#12 := (uf_1 #10 #11)
-#4070 := (pattern #12)
-#16 := (uf_2 #12)
-#317 := (= #10 #16)
-#4077 := (forall (vars (?x4 T2) (?x5 T2)) (:pat #4070) #317)
-#321 := (forall (vars (?x4 T2) (?x5 T2)) #317)
-#4080 := (iff #321 #4077)
-#4078 := (iff #317 #317)
-#4079 := [refl]: #4078
-#4081 := [quant-intro #4079]: #4080
-#1731 := (~ #321 #321)
-#1765 := (~ #317 #317)
-#1766 := [refl]: #1765
-#1732 := [nnf-pos #1766]: #1731
-#17 := (= #16 #10)
-#18 := (forall (vars (?x4 T2) (?x5 T2)) #17)
-#322 := (iff #18 #321)
-#319 := (iff #17 #317)
-#320 := [rewrite]: #319
-#323 := [quant-intro #320]: #322
-#316 := [asserted]: #18
-#326 := [mp #316 #323]: #321
-#1767 := [mp~ #326 #1732]: #321
-#4082 := [mp #1767 #4081]: #4077
-#8504 := (not #4077)
-#20954 := (or #8504 #20977)
-#20968 := [quant-inst]: #20954
-#27303 := [unit-resolution #20968 #4082]: #20977
-#27305 := [symm #27303]: #27304
-#13612 := (= #11533 uf_23)
-#188 := (uf_7 uf_15 uf_22 uf_8)
-#7202 := (= #188 uf_23)
-#189 := (= uf_23 #188)
-#2239 := (uf_1 #11 ?x75!20)
-#4360 := (pattern #2239)
-decl uf_24 :: (-> T2 int)
-#196 := (uf_24 #11)
-#4300 := (pattern #196)
-#206 := (up_6 uf_23 #11)
-#4326 := (pattern #206)
-#52 := 0::int
-decl uf_10 :: (-> T1 int)
-#2240 := (uf_10 #2239)
-#2236 := (uf_24 ?x75!20)
-#1127 := -1::int
-#2237 := (* -1::int #2236)
-#2834 := (+ #2237 #2240)
-#2835 := (+ #196 #2834)
-#2838 := (= #2835 0::int)
-#3400 := (not #2838)
-#2238 := (+ #196 #2237)
-#2243 := (>= #2238 0::int)
-#213 := (not #206)
-#3401 := (or #213 #2243 #3400)
-#4361 := (forall (vars (?x76 T2)) (:pat #4326 #4300 #4360) #3401)
-#4366 := (not #4361)
-#87 := (uf_1 #11 #10)
-#4155 := (pattern #87)
-#216 := (uf_24 #10)
-#1407 := (* -1::int #216)
-#1408 := (+ #196 #1407)
-#88 := (uf_10 #87)
-#1433 := (+ #88 #1408)
-#1431 := (>= #1433 0::int)
-#1142 := (* -1::int #88)
-decl uf_9 :: int
-#53 := uf_9
-#1143 := (+ uf_9 #1142)
-#1144 := (<= #1143 0::int)
-#3392 := (or #213 #1144 #1431)
-#4352 := (forall (vars (?x71 T2) (?x72 T2)) (:pat #4155) #3392)
-#4357 := (not #4352)
-decl uf_11 :: T2
-#64 := uf_11
-#2810 := (= uf_11 ?x75!20)
-#2247 := (+ uf_9 #2237)
-#2248 := (<= #2247 0::int)
-#4369 := (or #2248 #2810 #4357 #4366)
-#4372 := (not #4369)
-decl ?x71!19 :: T2
-#2206 := ?x71!19
-decl ?x72!18 :: T2
-#2205 := ?x72!18
-#2211 := (uf_1 ?x72!18 ?x71!19)
-#2212 := (uf_10 #2211)
-#2215 := (* -1::int #2212)
-#2209 := (uf_24 ?x72!18)
-#2796 := (* -1::int #2209)
-#2797 := (+ #2796 #2215)
-#2207 := (uf_24 ?x71!19)
-#2798 := (+ #2207 #2797)
-#2799 := (<= #2798 0::int)
-#2219 := (up_6 uf_23 ?x72!18)
-#3355 := (not #2219)
-#2216 := (+ uf_9 #2215)
-#2217 := (<= #2216 0::int)
-#3370 := (or #2217 #3355 #2799)
-#3375 := (not #3370)
-#4375 := (or #3375 #4372)
-#4378 := (not #4375)
-#4343 := (pattern #196 #216)
-#1406 := (>= #1408 0::int)
-#214 := (up_6 uf_23 #10)
-#3332 := (not #214)
-#3347 := (or #206 #3332 #1406)
-#4344 := (forall (vars (?x67 T2) (?x68 T2)) (:pat #4343) #3347)
-#4349 := (not #4344)
-#4381 := (or #4349 #4378)
-#4384 := (not #4381)
+#8048 := (uf_7 uf_15 #6904 uf_8)
+#13996 := (up_6 #8048 ?x75!20)
+#14006 := (iff #13996 #14001)
+#30 := (:var 1 T5)
+#20 := (:var 2 T2)
+#29 := (:var 3 T4)
+#31 := (uf_7 #29 #20 #30)
+#32 := (up_6 #31 #11)
+#3845 := (pattern #32)
+#35 := (up_6 #29 #11)
+#34 := (= #30 uf_8)
+#24 := (= #11 #20)
+#36 := (ite #24 #34 #35)
+#37 := (iff #32 #36)
+#3846 := (forall (vars (?x10 T4) (?x11 T2) (?x12 T5) (?x13 T2)) (:pat #3845) #37)
+#38 := (forall (vars (?x10 T4) (?x11 T2) (?x12 T5) (?x13 T2)) #37)
+#3849 := (iff #38 #3846)
+#3847 := (iff #37 #37)
+#3848 := [refl]: #3847
+#3850 := [quant-intro #3848]: #3849
+#1505 := (~ #38 #38)
+#1541 := (~ #37 #37)
+#1542 := [refl]: #1541
+#1506 := [nnf-pos #1542]: #1505
+#309 := [asserted]: #38
+#1543 := [mp~ #309 #1506]: #38
+#3851 := [mp #1543 #3850]: #3846
+#6010 := (not #3846)
+#30184 := (or #6010 #14006)
+#5970 := (= uf_8 uf_8)
+#13995 := (ite #13994 #5970 #11031)
+#13997 := (iff #13996 #13995)
+#30185 := (or #6010 #13997)
+#30191 := (iff #30185 #30184)
+#30193 := (iff #30184 #30184)
+#30194 := [rewrite]: #30193
+#14007 := (iff #13997 #14006)
+#14004 := (iff #13995 #14001)
+#1 := true
+#13998 := (ite #13994 true #11031)
+#14002 := (iff #13998 #14001)
+#14003 := [rewrite]: #14002
+#13999 := (iff #13995 #13998)
+#5975 := (iff #5970 true)
+#5976 := [rewrite]: #5975
+#14000 := [monotonicity #5976]: #13999
+#14005 := [trans #14000 #14003]: #14004
+#14008 := [monotonicity #14005]: #14007
+#30192 := [monotonicity #14008]: #30191
+#30195 := [trans #30192 #30194]: #30191
+#30190 := [quant-inst]: #30185
+#30196 := [mp #30190 #30195]: #30184
+#30517 := [unit-resolution #30196 #3851]: #14006
+#30208 := (not #13996)
+#9424 := (up_6 uf_23 ?x75!20)
+#9640 := (not #9424)
+#30526 := (iff #9640 #30208)
+#30524 := (iff #9424 #13996)
+#30522 := (iff #13996 #9424)
+#8591 := (= #8048 uf_23)
+#191 := (uf_7 uf_15 uf_22 uf_8)
+#7418 := (= #191 uf_23)
+#192 := (= uf_23 #191)
+#4136 := (or #3126 #4133)
+#4139 := (not #4136)
+#4104 := (pattern #199 #226)
+#1336 := (>= #1338 0::int)
+#224 := (up_6 uf_23 #10)
+#3083 := (not #224)
+#3098 := (or #216 #3083 #1336)
+#4105 := (forall (vars (?x67 T2) (?x68 T2)) (:pat #4104) #3098)
+#4110 := (not #4105)
+#4142 := (or #4110 #4139)
+#4145 := (not #4142)
decl ?x68!16 :: T2
-#2180 := ?x68!16
-#2184 := (uf_24 ?x68!16)
-#2770 := (* -1::int #2184)
+#1976 := ?x68!16
+#1980 := (uf_24 ?x68!16)
+#2517 := (* -1::int #1980)
decl ?x67!17 :: T2
-#2181 := ?x67!17
-#2182 := (uf_24 ?x67!17)
-#2771 := (+ #2182 #2770)
-#2772 := (<= #2771 0::int)
-#2188 := (up_6 uf_23 ?x68!16)
-#2187 := (up_6 uf_23 ?x67!17)
-#3309 := (not #2187)
-#3324 := (or #3309 #2188 #2772)
-#3329 := (not #3324)
-#4387 := (or #3329 #4384)
-#4390 := (not #4387)
-#1397 := (>= #196 0::int)
-#4335 := (forall (vars (?x65 T2)) (:pat #4300) #1397)
-#4340 := (not #4335)
-#4393 := (or #4340 #4390)
-#4396 := (not #4393)
+#1977 := ?x67!17
+#1978 := (uf_24 ?x67!17)
+#2518 := (+ #1978 #2517)
+#2519 := (<= #2518 0::int)
+#1984 := (up_6 uf_23 ?x68!16)
+#1983 := (up_6 uf_23 ?x67!17)
+#3060 := (not #1983)
+#3075 := (or #3060 #1984 #2519)
+#3080 := (not #3075)
+#4148 := (or #3080 #4145)
+#4151 := (not #4148)
+#1327 := (>= #199 0::int)
+#4096 := (forall (vars (?x65 T2)) (:pat #4062) #1327)
+#4101 := (not #4096)
+#4154 := (or #4101 #4151)
+#4157 := (not #4154)
decl ?x65!15 :: T2
-#2165 := ?x65!15
-#2166 := (uf_24 ?x65!15)
-#2167 := (>= #2166 0::int)
-#2168 := (not #2167)
-#4399 := (or #2168 #4396)
-#4402 := (not #4399)
-#209 := (uf_24 uf_11)
-#210 := (= #209 0::int)
-#1394 := (not #210)
-#4405 := (or #1394 #4402)
-#4408 := (not #4405)
-#4411 := (or #1394 #4408)
-#4414 := (not #4411)
-decl uf_4 :: (-> T3 T2 int)
-decl uf_14 :: T3
-#104 := uf_14
-#107 := (uf_4 uf_14 #11)
-#4179 := (pattern #107)
-#688 := (= #107 #196)
-#705 := (or #213 #688)
-#4327 := (forall (vars (?x63 T2)) (:pat #4326 #4179 #4300) #705)
-#4332 := (not #4327)
-#4417 := (or #4332 #4414)
-#4420 := (not #4417)
+#1961 := ?x65!15
+#1962 := (uf_24 ?x65!15)
+#1963 := (>= #1962 0::int)
+#1964 := (not #1963)
+#4160 := (or #1964 #4157)
+#4163 := (not #4160)
+#219 := (uf_24 uf_11)
+#220 := (= #219 0::int)
+#1324 := (not #220)
+#4166 := (or #1324 #4163)
+#4169 := (not #4166)
+#4172 := (or #1324 #4169)
+#4175 := (not #4172)
+#110 := (uf_4 uf_14 #11)
+#3940 := (pattern #110)
+#658 := (= #110 #199)
+#700 := (or #223 #658)
+#4088 := (forall (vars (?x63 T2)) (:pat #4087 #3940 #4062) #700)
+#4093 := (not #4088)
+#4178 := (or #4093 #4175)
+#4181 := (not #4178)
decl ?x63!14 :: T2
-#2141 := ?x63!14
-#2143 := (uf_4 uf_14 ?x63!14)
-#2142 := (uf_24 ?x63!14)
-#2747 := (= #2142 #2143)
-#2145 := (up_6 uf_23 ?x63!14)
-#2146 := (not #2145)
-#2750 := (or #2146 #2747)
-#2753 := (not #2750)
-#4423 := (or #2753 #4420)
-#4426 := (not #4423)
-#1382 := (* -1::int #196)
-#1383 := (+ #107 #1382)
-#1381 := (>= #1383 0::int)
-#4318 := (forall (vars (?x61 T2)) (:pat #4179 #4300) #1381)
-#4323 := (not #4318)
-#4429 := (or #4323 #4426)
-#4432 := (not #4429)
+#1937 := ?x63!14
+#1939 := (uf_4 uf_14 ?x63!14)
+#1938 := (uf_24 ?x63!14)
+#2494 := (= #1938 #1939)
+#1941 := (up_6 uf_23 ?x63!14)
+#1942 := (not #1941)
+#2497 := (or #1942 #2494)
+#2500 := (not #2497)
+#4184 := (or #2500 #4181)
+#4187 := (not #4184)
+#1248 := (* -1::int #199)
+#1313 := (+ #110 #1248)
+#1312 := (>= #1313 0::int)
+#4079 := (forall (vars (?x61 T2)) (:pat #3940 #4062) #1312)
+#4084 := (not #4079)
+#4190 := (or #4084 #4187)
+#4193 := (not #4190)
decl ?x61!13 :: T2
-#2123 := ?x61!13
-#2126 := (uf_4 uf_14 ?x61!13)
-#2737 := (* -1::int #2126)
-#2124 := (uf_24 ?x61!13)
-#2738 := (+ #2124 #2737)
-#2739 := (<= #2738 0::int)
-#2744 := (not #2739)
-#4435 := (or #2744 #4432)
-#4438 := (not #4435)
-#190 := (uf_1 uf_22 #11)
-#4301 := (pattern #190)
-#191 := (uf_10 #190)
-#1520 := (+ #191 #1382)
-#182 := (uf_4 uf_14 uf_22)
-#1521 := (+ #182 #1520)
-#1522 := (= #1521 0::int)
-#1351 := (* -1::int #191)
-#1357 := (* -1::int #182)
-#1358 := (+ #1357 #1351)
-#1359 := (+ #107 #1358)
-#1360 := (<= #1359 0::int)
-#1352 := (+ uf_9 #1351)
-#1353 := (<= #1352 0::int)
-#3301 := (or #1353 #1360 #1522)
-#4310 := (forall (vars (?x59 T2)) (:pat #4301 #4179 #4300) #3301)
-#4315 := (not #4310)
-#3281 := (or #1353 #1360)
-#3282 := (not #3281)
-#3285 := (or #688 #3282)
-#4302 := (forall (vars (?x60 T2)) (:pat #4179 #4300 #4301) #3285)
-#4307 := (not #4302)
+#1919 := ?x61!13
+#1922 := (uf_4 uf_14 ?x61!13)
+#2484 := (* -1::int #1922)
+#1920 := (uf_24 ?x61!13)
+#2485 := (+ #1920 #2484)
+#2486 := (<= #2485 0::int)
+#2491 := (not #2486)
+#4196 := (or #2491 #4193)
+#4199 := (not #4196)
+#193 := (uf_1 uf_22 #11)
+#4061 := (pattern #193)
+#194 := (uf_10 #193)
+#1229 := (* -1::int #194)
+#1235 := (* -1::int #185)
+#1236 := (+ #1235 #1229)
+#1237 := (+ #110 #1236)
+#1238 := (<= #1237 0::int)
+#1230 := (+ uf_9 #1229)
+#1231 := (<= #1230 0::int)
+#3032 := (or #1231 #1238)
+#3033 := (not #3032)
+#3054 := (or #658 #3033)
+#4071 := (forall (vars (?x60 T2)) (:pat #3940 #4062 #4061) #3054)
+#4076 := (not #4071)
+#1249 := (+ #194 #1248)
+#1250 := (+ #185 #1249)
+#1251 := (= #1250 0::int)
+#3046 := (or #1231 #1238 #1251)
+#4063 := (forall (vars (?x59 T2)) (:pat #4061 #3940 #4062) #3046)
+#4068 := (not #4063)
decl ?x48!12 :: T2
-#2100 := ?x48!12
-#2106 := (up_6 uf_15 ?x48!12)
-#2101 := (uf_4 uf_14 ?x48!12)
-#2102 := (* -1::int #2101)
-#2103 := (+ uf_9 #2102)
-#2104 := (<= #2103 0::int)
-#1547 := (+ uf_9 #1357)
-#1548 := (<= #1547 0::int)
-#111 := (up_6 uf_15 #11)
-#4221 := (pattern #111)
-#1535 := (+ #107 #1357)
-#1534 := (>= #1535 0::int)
-#1538 := (or #111 #1534)
-#4292 := (forall (vars (?x58 T2)) (:pat #4221 #4179) #1538)
-#4297 := (not #4292)
-#888 := (not #189)
-#180 := (up_6 uf_15 uf_22)
-#4441 := (or #180 #888 #4297 #1548 #2104 #2106 #4307 #4315 #4438)
-#4444 := (not #4441)
+#1886 := ?x48!12
+#1892 := (up_6 uf_15 ?x48!12)
+#1887 := (uf_4 uf_14 ?x48!12)
+#1888 := (* -1::int #1887)
+#1889 := (+ uf_9 #1888)
+#1890 := (<= #1889 0::int)
+#1281 := (+ uf_9 #1235)
+#1282 := (<= #1281 0::int)
+#114 := (up_6 uf_15 #11)
+#3982 := (pattern #114)
+#1272 := (+ #110 #1235)
+#1271 := (>= #1272 0::int)
+#1275 := (or #114 #1271)
+#4053 := (forall (vars (?x58 T2)) (:pat #3982 #3940) #1275)
+#4058 := (not #4053)
+#3249 := (not #192)
+#183 := (up_6 uf_15 uf_22)
+#4202 := (or #183 #3249 #4058 #1282 #1890 #1892 #4068 #4076 #4199)
+#4205 := (not #4202)
decl ?x53!11 :: T2
-#2034 := ?x53!11
+#1827 := ?x53!11
decl ?x54!10 :: T2
-#2033 := ?x54!10
-#2039 := (uf_1 ?x54!10 ?x53!11)
-#2040 := (uf_10 #2039)
-#2047 := (* -1::int #2040)
+#1826 := ?x54!10
+#1832 := (uf_1 ?x54!10 ?x53!11)
+#1833 := (uf_10 #1832)
+#1840 := (* -1::int #1833)
decl uf_19 :: T3
-#141 := uf_19
-#2037 := (uf_4 uf_19 ?x54!10)
-#2043 := (* -1::int #2037)
-#2694 := (+ #2043 #2047)
-#2035 := (uf_4 uf_19 ?x53!11)
-#2695 := (+ #2035 #2694)
-#2696 := (<= #2695 0::int)
-#2048 := (+ uf_9 #2047)
-#2049 := (<= #2048 0::int)
-#2044 := (+ uf_9 #2043)
-#2045 := (<= #2044 0::int)
-#3245 := (or #2045 #2049 #2696)
-#3250 := (not #3245)
-#149 := (uf_4 uf_19 #10)
-#1264 := (* -1::int #149)
-#146 := (uf_4 uf_19 #11)
-#1265 := (+ #146 #1264)
-#1271 := (+ #88 #1265)
-#1294 := (>= #1271 0::int)
-#1251 := (* -1::int #146)
-#1252 := (+ uf_9 #1251)
-#1253 := (<= #1252 0::int)
-#3213 := (or #1144 #1253 #1294)
-#4254 := (forall (vars (?x53 T2) (?x54 T2)) (:pat #4155) #3213)
-#4259 := (not #4254)
-#161 := (uf_4 uf_19 uf_11)
-#162 := (= #161 0::int)
-#4262 := (or #162 #4259)
-#4265 := (not #4262)
-#4268 := (or #4265 #3250)
-#4271 := (not #4268)
-#4230 := (pattern #146)
+#150 := uf_19
+#1830 := (uf_4 uf_19 ?x54!10)
+#1836 := (* -1::int #1830)
+#2443 := (+ #1836 #1840)
+#1828 := (uf_4 uf_19 ?x53!11)
+#2444 := (+ #1828 #2443)
+#2445 := (<= #2444 0::int)
+#1841 := (+ uf_9 #1840)
+#1842 := (<= #1841 0::int)
+#1837 := (+ uf_9 #1836)
+#1838 := (<= #1837 0::int)
+#2992 := (or #1838 #1842 #2445)
+#2997 := (not #2992)
+#163 := (uf_4 uf_19 #10)
+#1165 := (* -1::int #163)
+#160 := (uf_4 uf_19 #11)
+#1166 := (+ #160 #1165)
+#1172 := (+ #91 #1166)
+#1195 := (>= #1172 0::int)
+#1152 := (* -1::int #160)
+#1153 := (+ uf_9 #1152)
+#1154 := (<= #1153 0::int)
+#2960 := (or #969 #1154 #1195)
+#4015 := (forall (vars (?x53 T2) (?x54 T2)) (:pat #3916) #2960)
+#4020 := (not #4015)
+#175 := (uf_4 uf_19 uf_11)
+#176 := (= #175 0::int)
+#4023 := (or #176 #4020)
+#4026 := (not #4023)
+#4029 := (or #4026 #2997)
+#4032 := (not #4029)
+#3991 := (pattern #160)
decl ?x50!9 :: (-> T2 T2)
-#2010 := (?x50!9 #11)
-#2013 := (uf_1 #2010 #11)
-#2014 := (uf_10 #2013)
-#2664 := (* -1::int #2014)
-#2011 := (uf_4 uf_19 #2010)
-#2647 := (* -1::int #2011)
-#2665 := (+ #2647 #2664)
-#2666 := (+ #146 #2665)
-#2667 := (= #2666 0::int)
-#3183 := (not #2667)
-#2648 := (+ #146 #2647)
-#2649 := (<= #2648 0::int)
-#3184 := (or #2649 #3183)
-#3185 := (not #3184)
+#1803 := (?x50!9 #11)
+#1806 := (uf_1 #1803 #11)
+#1807 := (uf_10 #1806)
+#2413 := (* -1::int #1807)
+#1804 := (uf_4 uf_19 #1803)
+#2396 := (* -1::int #1804)
+#2414 := (+ #2396 #2413)
+#2415 := (+ #160 #2414)
+#2416 := (= #2415 0::int)
+#2930 := (not #2416)
+#2397 := (+ #160 #2396)
+#2398 := (<= #2397 0::int)
+#2931 := (or #2398 #2930)
+#2932 := (not #2931)
#65 := (= #11 uf_11)
-#3191 := (or #65 #1253 #3185)
-#4246 := (forall (vars (?x49 T2)) (:pat #4230) #3191)
-#4251 := (not #4246)
-#4274 := (or #4251 #4271)
-#4277 := (not #4274)
+#2938 := (or #65 #1154 #2932)
+#4007 := (forall (vars (?x49 T2)) (:pat #3991) #2938)
+#4012 := (not #4007)
+#4035 := (or #4012 #4032)
+#4038 := (not #4035)
decl ?x49!8 :: T2
-#1970 := ?x49!8
-#1974 := (uf_1 #11 ?x49!8)
-#4231 := (pattern #1974)
-#1975 := (uf_10 #1974)
-#1971 := (uf_4 uf_19 ?x49!8)
-#1972 := (* -1::int #1971)
-#2617 := (+ #1972 #1975)
-#2618 := (+ #146 #2617)
-#2621 := (= #2618 0::int)
-#3147 := (not #2621)
-#1973 := (+ #146 #1972)
-#1978 := (>= #1973 0::int)
-#3148 := (or #1978 #3147)
-#4232 := (forall (vars (?x50 T2)) (:pat #4230 #4231) #3148)
-#4237 := (not #4232)
-#2593 := (= uf_11 ?x49!8)
-#1982 := (+ uf_9 #1972)
-#1983 := (<= #1982 0::int)
-#4240 := (or #1983 #2593 #4237)
-#4243 := (not #4240)
-#4280 := (or #4243 #4277)
-#4283 := (not #4280)
-#1206 := (* -1::int #107)
-#1207 := (+ uf_9 #1206)
-#1208 := (<= #1207 0::int)
-#3133 := (or #111 #1208)
-#4222 := (forall (vars (?x48 T2)) (:pat #4221 #4179) #3133)
-#4227 := (not #4222)
-#509 := (= uf_14 uf_19)
-#614 := (not #509)
+#1763 := ?x49!8
+#1767 := (uf_1 #11 ?x49!8)
+#3992 := (pattern #1767)
+#1768 := (uf_10 #1767)
+#1764 := (uf_4 uf_19 ?x49!8)
+#1765 := (* -1::int #1764)
+#2366 := (+ #1765 #1768)
+#2367 := (+ #160 #2366)
+#2370 := (= #2367 0::int)
+#2894 := (not #2370)
+#1766 := (+ #160 #1765)
+#1771 := (>= #1766 0::int)
+#2895 := (or #1771 #2894)
+#3993 := (forall (vars (?x50 T2)) (:pat #3991 #3992) #2895)
+#3998 := (not #3993)
+#2342 := (= uf_11 ?x49!8)
+#1775 := (+ uf_9 #1765)
+#1776 := (<= #1775 0::int)
+#4001 := (or #1776 #2342 #3998)
+#4004 := (not #4001)
+#4041 := (or #4004 #4038)
+#4044 := (not #4041)
+#1054 := (* -1::int #110)
+#1055 := (+ uf_9 #1054)
+#1056 := (<= #1055 0::int)
+#2880 := (or #114 #1056)
+#3983 := (forall (vars (?x48 T2)) (:pat #3982 #3940) #2880)
+#3988 := (not #3983)
+#533 := (= uf_14 uf_19)
+#3023 := (not #533)
decl uf_16 :: T4
-#136 := uf_16
-#506 := (= uf_15 uf_16)
-#632 := (not #506)
+#145 := uf_16
+#530 := (= uf_15 uf_16)
+#3022 := (not #530)
decl uf_21 :: T3
-#144 := uf_21
+#153 := uf_21
decl uf_20 :: T3
-#143 := uf_20
-#145 := (= uf_20 uf_21)
-#605 := (not #145)
+#152 := uf_20
+#154 := (= uf_20 uf_21)
+#3021 := (not #154)
decl uf_18 :: T2
-#139 := uf_18
+#148 := uf_18
decl uf_17 :: T2
-#138 := uf_17
-#140 := (= uf_17 uf_18)
-#623 := (not #140)
-#4286 := (or #623 #605 #632 #614 #4227 #4283)
-#4289 := (not #4286)
-#4447 := (or #4289 #4444)
-#4450 := (not #4447)
-#1934 := (?x47!7 #11)
-#1935 := (uf_4 uf_14 #1934)
-#2552 := (* -1::int #1935)
-#2567 := (+ #107 #2552)
-#2568 := (<= #2567 0::int)
-#1939 := (uf_1 #1934 #11)
-#1940 := (uf_10 #1939)
-#2553 := (* -1::int #1940)
-#2554 := (+ #2552 #2553)
-#2555 := (+ #107 #2554)
-#2556 := (= #2555 0::int)
-#3117 := (not #2556)
-#1943 := (up_6 uf_15 #1934)
-#3116 := (not #1943)
-#3118 := (or #3116 #3117 #2568)
-#3119 := (not #3118)
-#3125 := (or #65 #1208 #3119)
-#4213 := (forall (vars (?x46 T2)) (:pat #4179) #3125)
-#4218 := (not #4213)
+#147 := uf_17
+#149 := (= uf_17 uf_18)
+#3020 := (not #149)
+#4047 := (or #3020 #3021 #3022 #3023 #3988 #4044)
+#4050 := (not #4047)
+#4208 := (or #4050 #4205)
+#4211 := (not #4208)
+#1708 := (?x47!7 #11)
+#1709 := (uf_4 uf_14 #1708)
+#2309 := (* -1::int #1709)
+#2324 := (+ #110 #2309)
+#2325 := (<= #2324 0::int)
+#1713 := (uf_1 #1708 #11)
+#1714 := (uf_10 #1713)
+#2310 := (* -1::int #1714)
+#2311 := (+ #2309 #2310)
+#2312 := (+ #110 #2311)
+#2313 := (= #2312 0::int)
+#2864 := (not #2313)
+#1717 := (up_6 uf_15 #1708)
+#2863 := (not #1717)
+#2865 := (or #2863 #2864 #2325)
+#2866 := (not #2865)
+#2872 := (or #65 #1056 #2866)
+#3974 := (forall (vars (?x46 T2)) (:pat #3940) #2872)
+#3979 := (not #3974)
decl uf_12 :: (-> T2 int)
#66 := (uf_12 #11)
-#4131 := (pattern #66)
+#3886 := (pattern #66)
decl ?x38!6 :: (-> T2 T2)
-#1907 := (?x38!6 #11)
-#1911 := (uf_12 #1907)
-#2511 := (* -1::int #1911)
-#1908 := (uf_1 #1907 #11)
-#1909 := (uf_10 #1908)
-#2528 := (* -1::int #1909)
-#2529 := (+ #2528 #2511)
-#2530 := (+ #66 #2529)
-#2531 := (= #2530 0::int)
-#3089 := (not #2531)
-#2512 := (+ #66 #2511)
-#2513 := (<= #2512 0::int)
+#1680 := (?x38!6 #11)
+#1684 := (uf_12 #1680)
+#2268 := (* -1::int #1684)
+#1681 := (uf_1 #1680 #11)
+#1682 := (uf_10 #1681)
+#2285 := (* -1::int #1682)
+#2286 := (+ #2285 #2268)
+#2287 := (+ #66 #2286)
+#2288 := (= #2287 0::int)
+#2836 := (not #2288)
+#2269 := (+ #66 #2268)
+#2270 := (<= #2269 0::int)
decl up_13 :: (-> T2 bool)
-#1917 := (up_13 #1907)
-#3088 := (not #1917)
-#3090 := (or #3088 #2513 #3089)
-#3091 := (not #3090)
-#1168 := (* -1::int #66)
-#1169 := (+ uf_9 #1168)
-#1170 := (<= #1169 0::int)
-#3097 := (or #65 #1170 #3091)
-#4205 := (forall (vars (?x37 T2)) (:pat #4131) #3097)
-#4210 := (not #4205)
-#113 := (up_6 uf_15 #10)
-#4196 := (pattern #111 #113)
-#115 := (uf_4 uf_14 #10)
-#1220 := (* -1::int #115)
-#1221 := (+ #107 #1220)
-#1224 := (>= #1221 0::int)
-#3054 := (not #113)
-#3069 := (or #111 #3054 #1224)
-#4197 := (forall (vars (?x42 T2) (?x43 T2)) (:pat #4196) #3069)
-#4202 := (not #4197)
-#1222 := (+ #88 #1221)
-#1602 := (>= #1222 0::int)
-#112 := (not #111)
-#3046 := (or #112 #1144 #1602)
-#4188 := (forall (vars (?x44 T2) (?x45 T2)) (:pat #4155) #3046)
-#4193 := (not #4188)
-#1625 := (>= #107 0::int)
-#4180 := (forall (vars (?x41 T2)) (:pat #4179) #1625)
-#4185 := (not #4180)
-#105 := (uf_4 uf_14 uf_11)
-#106 := (= #105 0::int)
-#1636 := (not #106)
-#4453 := (or #1636 #4185 #4193 #4202 #4210 #4218 #4450)
-#4456 := (not #4453)
+#1690 := (up_13 #1680)
+#2835 := (not #1690)
+#2837 := (or #2835 #2270 #2836)
+#2838 := (not #2837)
+#993 := (* -1::int #66)
+#994 := (+ uf_9 #993)
+#995 := (<= #994 0::int)
+#2844 := (or #65 #995 #2838)
+#3966 := (forall (vars (?x37 T2)) (:pat #3886) #2844)
+#3971 := (not #3966)
+#116 := (up_6 uf_15 #10)
+#3957 := (pattern #114 #116)
+#118 := (uf_4 uf_14 #10)
+#1045 := (* -1::int #118)
+#1046 := (+ #110 #1045)
+#1069 := (>= #1046 0::int)
+#2801 := (not #116)
+#2816 := (or #114 #2801 #1069)
+#3958 := (forall (vars (?x42 T2) (?x43 T2)) (:pat #3957) #2816)
+#3963 := (not #3958)
+#1047 := (+ #91 #1046)
+#1043 := (>= #1047 0::int)
+#115 := (not #114)
+#2793 := (or #115 #969 #1043)
+#3949 := (forall (vars (?x44 T2) (?x45 T2)) (:pat #3916) #2793)
+#3954 := (not #3949)
+#1101 := (>= #110 0::int)
+#3941 := (forall (vars (?x41 T2)) (:pat #3940) #1101)
+#3946 := (not #3941)
+#108 := (uf_4 uf_14 uf_11)
+#109 := (= #108 0::int)
+#3266 := (not #109)
+#4214 := (or #3266 #3946 #3954 #3963 #3971 #3979 #4211)
+#4217 := (not #4214)
decl ?x37!5 :: T2
-#1863 := ?x37!5
-#1864 := (uf_1 #11 ?x37!5)
-#4164 := (pattern #1864)
+#1636 := ?x37!5
+#1637 := (uf_1 #11 ?x37!5)
+#3925 := (pattern #1637)
#74 := (up_13 #11)
-#4124 := (pattern #74)
-#1866 := (uf_12 ?x37!5)
-#1867 := (* -1::int #1866)
-#1865 := (uf_10 #1864)
-#2479 := (+ #1865 #1867)
-#2480 := (+ #66 #2479)
-#2483 := (= #2480 0::int)
-#3007 := (not #2483)
-#1871 := (+ #66 #1867)
-#1872 := (>= #1871 0::int)
+#3879 := (pattern #74)
+#1639 := (uf_12 ?x37!5)
+#1640 := (* -1::int #1639)
+#1638 := (uf_10 #1637)
+#2238 := (+ #1638 #1640)
+#2239 := (+ #66 #2238)
+#2242 := (= #2239 0::int)
+#2754 := (not #2242)
+#1644 := (+ #66 #1640)
+#1645 := (>= #1644 0::int)
#75 := (not #74)
-#3008 := (or #75 #1872 #3007)
-#4165 := (forall (vars (?x38 T2)) (:pat #4124 #4131 #4164) #3008)
-#4170 := (not #4165)
-#2455 := (= uf_11 ?x37!5)
-#1876 := (+ uf_9 #1867)
-#1877 := (<= #1876 0::int)
-#4173 := (or #1877 #2455 #4170)
-#4176 := (not #4173)
-#4459 := (or #4176 #4456)
-#4462 := (not #4459)
-#83 := (uf_12 #10)
-#1130 := (* -1::int #83)
-#1157 := (+ #1130 #88)
-#1158 := (+ #66 #1157)
-#1155 := (>= #1158 0::int)
-#2999 := (or #75 #1144 #1155)
-#4156 := (forall (vars (?x33 T2) (?x34 T2)) (:pat #4155) #2999)
-#4161 := (not #4156)
-#4465 := (or #4161 #4462)
-#4468 := (not #4465)
+#2755 := (or #75 #1645 #2754)
+#3926 := (forall (vars (?x38 T2)) (:pat #3879 #3886 #3925) #2755)
+#3931 := (not #3926)
+#2214 := (= uf_11 ?x37!5)
+#1649 := (+ uf_9 #1640)
+#1650 := (<= #1649 0::int)
+#3934 := (or #1650 #2214 #3931)
+#3937 := (not #3934)
+#4220 := (or #3937 #4217)
+#4223 := (not #4220)
+#86 := (uf_12 #10)
+#955 := (* -1::int #86)
+#982 := (+ #955 #91)
+#983 := (+ #66 #982)
+#980 := (>= #983 0::int)
+#2746 := (or #75 #969 #980)
+#3917 := (forall (vars (?x33 T2) (?x34 T2)) (:pat #3916) #2746)
+#3922 := (not #3917)
+#4226 := (or #3922 #4223)
+#4229 := (not #4226)
decl ?x34!3 :: T2
-#1833 := ?x34!3
-#1840 := (uf_12 ?x34!3)
+#1606 := ?x34!3
+#1613 := (uf_12 ?x34!3)
decl ?x33!4 :: T2
-#1834 := ?x33!4
-#1837 := (uf_12 ?x33!4)
-#1838 := (* -1::int #1837)
-#2442 := (+ #1838 #1840)
-#1835 := (uf_1 ?x34!3 ?x33!4)
-#1836 := (uf_10 #1835)
-#2443 := (+ #1836 #2442)
-#2446 := (>= #2443 0::int)
-#1847 := (up_13 ?x34!3)
-#2962 := (not #1847)
-#1843 := (* -1::int #1836)
-#1844 := (+ uf_9 #1843)
-#1845 := (<= #1844 0::int)
-#2977 := (or #1845 #2962 #2446)
-#8241 := [hypothesis]: #1847
-#4125 := (forall (vars (?x26 T2)) (:pat #4124) #75)
-#76 := (forall (vars (?x26 T2)) #75)
-#4128 := (iff #76 #4125)
-#4126 := (iff #75 #75)
-#4127 := [refl]: #4126
-#4129 := [quant-intro #4127]: #4128
-#1747 := (~ #76 #76)
-#1784 := (~ #75 #75)
-#1785 := [refl]: #1784
-#1748 := [nnf-pos #1785]: #1747
+#1607 := ?x33!4
+#1610 := (uf_12 ?x33!4)
+#1611 := (* -1::int #1610)
+#2201 := (+ #1611 #1613)
+#1608 := (uf_1 ?x34!3 ?x33!4)
+#1609 := (uf_10 #1608)
+#2202 := (+ #1609 #2201)
+#2205 := (>= #2202 0::int)
+#1620 := (up_13 ?x34!3)
+#2709 := (not #1620)
+#1616 := (* -1::int #1609)
+#1617 := (+ uf_9 #1616)
+#1618 := (<= #1617 0::int)
+#2724 := (or #1618 #2709 #2205)
+#2729 := (not #2724)
+#4232 := (or #2729 #4229)
+#4235 := (not #4232)
+#84 := (up_13 #10)
+#3907 := (pattern #74 #84)
+#956 := (+ #66 #955)
+#954 := (>= #956 0::int)
+#1898 := (not #84)
+#2701 := (or #74 #1898 #954)
+#3908 := (forall (vars (?x29 T2) (?x30 T2)) (:pat #3907) #2701)
+#3913 := (not #3908)
+#4238 := (or #3913 #4235)
+#4241 := (not #4238)
+decl ?x30!1 :: T2
+#1581 := ?x30!1
+#1585 := (uf_12 ?x30!1)
+#2182 := (* -1::int #1585)
+decl ?x29!2 :: T2
+#1582 := ?x29!2
+#1583 := (uf_12 ?x29!2)
+#2183 := (+ #1583 #2182)
+#2184 := (<= #2183 0::int)
+#1589 := (up_13 ?x30!1)
+#1588 := (up_13 ?x29!2)
+#1597 := (not #1588)
+#1948 := (or #1597 #1589 #2184)
+#2022 := (not #1948)
+#4244 := (or #2022 #4241)
+#4247 := (not #4244)
+#945 := (>= #66 0::int)
+#3899 := (forall (vars (?x27 T2)) (:pat #3886) #945)
+#3904 := (not #3899)
+#4250 := (or #3904 #4247)
+#4253 := (not #4250)
+decl ?x27!0 :: T2
+#1566 := ?x27!0
+#1567 := (uf_12 ?x27!0)
+#1568 := (>= #1567 0::int)
+#1569 := (not #1568)
+#4256 := (or #1569 #4253)
+#4259 := (not #4256)
+#80 := (uf_12 uf_11)
+#81 := (= #80 0::int)
+#940 := (not #81)
+#4262 := (or #940 #4259)
+#4265 := (not #4262)
+#4992 := [hypothesis]: #940
#67 := (= #66 0::int)
#70 := (not #65)
-#1694 := (or #70 #67)
-#1697 := (forall (vars (?x24 T2)) #1694)
-#1700 := (not #1697)
-#1628 := (forall (vars (?x41 T2)) #1625)
-#1631 := (not #1628)
-#114 := (and #112 #113)
-#454 := (not #114)
-#1616 := (or #454 #1224)
-#1619 := (forall (vars (?x42 T2) (?x43 T2)) #1616)
-#1622 := (not #1619)
-#1145 := (not #1144)
-#1594 := (and #111 #1145)
-#1599 := (not #1594)
-#1605 := (or #1599 #1602)
-#1608 := (forall (vars (?x44 T2) (?x45 T2)) #1605)
-#1611 := (not #1608)
-#1541 := (forall (vars (?x58 T2)) #1538)
-#1544 := (not #1541)
-#1361 := (not #1360)
-#1354 := (not #1353)
-#1364 := (and #1354 #1361)
-#1517 := (not #1364)
-#1525 := (or #1517 #1522)
-#1528 := (forall (vars (?x59 T2)) #1525)
-#1531 := (not #1528)
-#1455 := (= #1433 0::int)
-#1458 := (not #1406)
-#1467 := (and #206 #1458 #1455)
-#1472 := (exists (vars (?x76 T2)) #1467)
-#1444 := (+ uf_9 #1382)
-#1445 := (<= #1444 0::int)
-#1446 := (not #1445)
-#1449 := (and #70 #1446)
-#1452 := (not #1449)
-#1475 := (or #1452 #1472)
-#1478 := (forall (vars (?x75 T2)) #1475)
-#1423 := (and #206 #1145)
-#1428 := (not #1423)
-#1435 := (or #1428 #1431)
-#1438 := (forall (vars (?x71 T2) (?x72 T2)) #1435)
-#1441 := (not #1438)
-#1481 := (or #1441 #1478)
-#1484 := (and #1438 #1481)
-#215 := (and #213 #214)
-#716 := (not #215)
-#1411 := (or #716 #1406)
-#1414 := (forall (vars (?x67 T2) (?x68 T2)) #1411)
-#1417 := (not #1414)
-#1487 := (or #1417 #1484)
-#1490 := (and #1414 #1487)
-#1400 := (forall (vars (?x65 T2)) #1397)
-#1403 := (not #1400)
-#1493 := (or #1403 #1490)
-#1496 := (and #1400 #1493)
-#1499 := (or #1394 #1496)
-#1502 := (and #210 #1499)
-#710 := (forall (vars (?x63 T2)) #705)
-#846 := (not #710)
-#1505 := (or #846 #1502)
-#1508 := (and #710 #1505)
-#1386 := (forall (vars (?x61 T2)) #1381)
-#1389 := (not #1386)
-#1511 := (or #1389 #1508)
-#1514 := (and #1386 #1511)
-#1370 := (or #688 #1364)
-#1375 := (forall (vars (?x60 T2)) #1370)
-#1378 := (not #1375)
-#1209 := (not #1208)
-#1325 := (and #112 #1209)
-#1328 := (exists (vars (?x48 T2)) #1325)
-#1559 := (not #1328)
-#1583 := (or #180 #888 #1559 #1378 #1514 #1531 #1544 #1548)
-#1254 := (not #1253)
-#1288 := (and #1145 #1254)
-#1291 := (not #1288)
-#1297 := (or #1291 #1294)
-#1300 := (forall (vars (?x53 T2) (?x54 T2)) #1297)
-#1303 := (not #1300)
-#1311 := (or #162 #1303)
-#1316 := (and #1300 #1311)
-#1269 := (= #1271 0::int)
-#1263 := (>= #1265 0::int)
-#1266 := (not #1263)
-#1273 := (and #1266 #1269)
-#1276 := (exists (vars (?x50 T2)) #1273)
-#1257 := (and #70 #1254)
-#1260 := (not #1257)
-#1279 := (or #1260 #1276)
-#1282 := (forall (vars (?x49 T2)) #1279)
-#1285 := (not #1282)
-#1319 := (or #1285 #1316)
-#1322 := (and #1282 #1319)
-#1346 := (or #623 #605 #632 #614 #1322 #1328)
-#1588 := (and #1346 #1583)
-#1225 := (not #1224)
-#1218 := (= #1222 0::int)
-#1234 := (and #111 #1218 #1225)
-#1239 := (exists (vars (?x47 T2)) #1234)
-#1212 := (and #70 #1209)
-#1215 := (not #1212)
-#1242 := (or #1215 #1239)
-#1245 := (forall (vars (?x46 T2)) #1242)
-#1248 := (not #1245)
-#1180 := (= #1158 0::int)
-#1131 := (+ #66 #1130)
-#1129 := (>= #1131 0::int)
-#1183 := (not #1129)
-#1192 := (and #74 #1183 #1180)
-#1197 := (exists (vars (?x38 T2)) #1192)
-#1171 := (not #1170)
-#1174 := (and #70 #1171)
-#1177 := (not #1174)
-#1200 := (or #1177 #1197)
-#1203 := (forall (vars (?x37 T2)) #1200)
-#1639 := (not #1203)
-#1660 := (or #1636 #1639 #1248 #1588 #1611 #1622 #1631)
-#1665 := (and #1203 #1660)
-#1148 := (and #74 #1145)
-#1151 := (not #1148)
-#1159 := (or #1151 #1155)
-#1162 := (forall (vars (?x33 T2) (?x34 T2)) #1159)
-#1165 := (not #1162)
-#1668 := (or #1165 #1665)
-#1671 := (and #1162 #1668)
-#81 := (up_13 #10)
-#82 := (and #75 #81)
-#430 := (not #82)
-#1133 := (or #430 #1129)
-#1136 := (forall (vars (?x29 T2) (?x30 T2)) #1133)
-#1139 := (not #1136)
-#1674 := (or #1139 #1671)
-#1677 := (and #1136 #1674)
-#1120 := (>= #66 0::int)
-#1121 := (forall (vars (?x27 T2)) #1120)
-#1124 := (not #1121)
-#1680 := (or #1124 #1677)
-#1683 := (and #1121 #1680)
-#77 := (uf_12 uf_11)
-#78 := (= #77 0::int)
-#1115 := (not #78)
-#1686 := (or #1115 #1683)
-#1689 := (and #78 #1686)
-#413 := (= uf_9 #66)
-#419 := (or #65 #413)
-#424 := (forall (vars (?x25 T2)) #419)
-#1084 := (not #424)
-#1075 := (not #76)
-#1712 := (or #1075 #1084 #1689 #1700)
-#1717 := (not #1712)
-#1 := true
-#234 := (implies false true)
-#221 := (+ #196 #88)
-#228 := (= #216 #221)
-#229 := (and #206 #228)
-#227 := (< #196 #216)
-#230 := (and #227 #229)
-#231 := (exists (vars (?x76 T2)) #230)
-#225 := (< #196 uf_9)
-#226 := (and #70 #225)
-#232 := (implies #226 #231)
-#233 := (forall (vars (?x75 T2)) #232)
-#235 := (implies #233 #234)
-#236 := (and #233 #235)
-#222 := (<= #216 #221)
-#89 := (< #88 uf_9)
-#220 := (and #206 #89)
-#223 := (implies #220 #222)
-#224 := (forall (vars (?x71 T2) (?x72 T2)) #223)
-#237 := (implies #224 #236)
-#238 := (and #224 #237)
-#217 := (<= #216 #196)
-#218 := (implies #215 #217)
-#219 := (forall (vars (?x67 T2) (?x68 T2)) #218)
-#239 := (implies #219 #238)
-#240 := (and #219 #239)
-#211 := (<= 0::int #196)
-#212 := (forall (vars (?x65 T2)) #211)
-#241 := (implies #212 #240)
-#242 := (and #212 #241)
-#243 := (implies #210 #242)
-#244 := (and #210 #243)
-#245 := (implies true #244)
-#246 := (implies true #245)
-#201 := (= #196 #107)
-#207 := (implies #206 #201)
-#208 := (forall (vars (?x63 T2)) #207)
-#247 := (implies #208 #246)
-#248 := (and #208 #247)
-#204 := (<= #196 #107)
-#205 := (forall (vars (?x61 T2)) #204)
-#249 := (implies #205 #248)
-#250 := (and #205 #249)
-#193 := (+ #182 #191)
-#194 := (< #193 #107)
-#192 := (< #191 uf_9)
-#195 := (and #192 #194)
-#200 := (not #195)
-#202 := (implies #200 #201)
-#203 := (forall (vars (?x60 T2)) #202)
-#251 := (implies #203 #250)
-#197 := (= #196 #193)
-#198 := (implies #195 #197)
-#199 := (forall (vars (?x59 T2)) #198)
-#252 := (implies #199 #251)
-#253 := (implies #189 #252)
-#184 := (<= #182 #107)
-#185 := (implies #112 #184)
-#186 := (forall (vars (?x58 T2)) #185)
-#254 := (implies #186 #253)
-#183 := (< #182 uf_9)
-#255 := (implies #183 #254)
-#181 := (not #180)
-#256 := (implies #181 #255)
-#124 := (< #107 uf_9)
-#133 := (and #112 #124)
-#134 := (exists (vars (?x48 T2)) #133)
-#257 := (implies #134 #256)
-#258 := (implies true #257)
-#259 := (implies true #258)
-#163 := (implies #162 true)
-#164 := (and #162 #163)
-#151 := (+ #146 #88)
-#158 := (<= #149 #151)
-#147 := (< #146 uf_9)
-#157 := (and #147 #89)
-#159 := (implies #157 #158)
-#160 := (forall (vars (?x53 T2) (?x54 T2)) #159)
-#165 := (implies #160 #164)
-#166 := (and #160 #165)
-#152 := (= #149 #151)
-#150 := (< #146 #149)
-#153 := (and #150 #152)
-#154 := (exists (vars (?x50 T2)) #153)
-#148 := (and #70 #147)
-#155 := (implies #148 #154)
-#156 := (forall (vars (?x49 T2)) #155)
-#167 := (implies #156 #166)
-#168 := (and #156 #167)
-#169 := (implies true #168)
-#170 := (implies #145 #169)
-#142 := (= uf_19 uf_14)
-#171 := (implies #142 #170)
-#172 := (implies #140 #171)
-#137 := (= uf_16 uf_15)
-#173 := (implies #137 #172)
-#174 := (implies true #173)
-#175 := (implies true #174)
-#135 := (not #134)
-#176 := (implies #135 #175)
-#177 := (implies true #176)
-#178 := (implies true #177)
-#260 := (and #178 #259)
-#261 := (implies true #260)
-#120 := (+ #107 #88)
-#127 := (= #115 #120)
-#128 := (and #111 #127)
-#126 := (< #107 #115)
-#129 := (and #126 #128)
-#130 := (exists (vars (?x47 T2)) #129)
-#125 := (and #70 #124)
-#131 := (implies #125 #130)
-#132 := (forall (vars (?x46 T2)) #131)
-#262 := (implies #132 #261)
-#121 := (<= #115 #120)
-#119 := (and #111 #89)
-#122 := (implies #119 #121)
-#123 := (forall (vars (?x44 T2) (?x45 T2)) #122)
-#263 := (implies #123 #262)
-#116 := (<= #115 #107)
-#117 := (implies #114 #116)
-#118 := (forall (vars (?x42 T2) (?x43 T2)) #117)
-#264 := (implies #118 #263)
-#108 := (<= 0::int #107)
-#109 := (forall (vars (?x41 T2)) #108)
-#265 := (implies #109 #264)
-#266 := (implies #106 #265)
-#267 := (implies true #266)
-#268 := (implies true #267)
-#91 := (+ #66 #88)
-#98 := (= #83 #91)
-#99 := (and #74 #98)
-#97 := (< #66 #83)
-#100 := (and #97 #99)
-#101 := (exists (vars (?x38 T2)) #100)
-#95 := (< #66 uf_9)
-#96 := (and #70 #95)
-#102 := (implies #96 #101)
-#103 := (forall (vars (?x37 T2)) #102)
-#269 := (implies #103 #268)
-#270 := (and #103 #269)
-#92 := (<= #83 #91)
-#90 := (and #74 #89)
-#93 := (implies #90 #92)
-#94 := (forall (vars (?x33 T2) (?x34 T2)) #93)
-#271 := (implies #94 #270)
-#272 := (and #94 #271)
-#84 := (<= #83 #66)
-#85 := (implies #82 #84)
-#86 := (forall (vars (?x29 T2) (?x30 T2)) #85)
-#273 := (implies #86 #272)
-#274 := (and #86 #273)
-#79 := (<= 0::int #66)
-#80 := (forall (vars (?x27 T2)) #79)
-#275 := (implies #80 #274)
-#276 := (and #80 #275)
-#277 := (implies #78 #276)
-#278 := (and #78 #277)
-#279 := (implies true #278)
-#280 := (implies #76 #279)
+#921 := (or #70 #67)
+#3893 := (forall (vars (?x24 T2)) (:pat #3886) #921)
+#924 := (forall (vars (?x24 T2)) #921)
+#3896 := (iff #924 #3893)
+#3894 := (iff #921 #921)
+#3895 := [refl]: #3894
+#3897 := [quant-intro #3895]: #3896
+#1521 := (~ #924 #924)
+#1560 := (~ #921 #921)
+#1561 := [refl]: #1560
+#1522 := [nnf-pos #1561]: #1521
+#397 := (= uf_9 #66)
+#403 := (or #65 #397)
+#408 := (forall (vars (?x25 T2)) #403)
+#76 := (forall (vars (?x26 T2)) #75)
+#930 := (and #76 #408 #924)
+#1385 := (= #1363 0::int)
+#1388 := (not #1336)
+#1397 := (and #216 #1388 #1385)
+#1402 := (exists (vars (?x76 T2)) #1397)
+#1374 := (+ uf_9 #1248)
+#1375 := (<= #1374 0::int)
+#1376 := (not #1375)
+#1379 := (and #70 #1376)
+#1382 := (not #1379)
+#1405 := (or #1382 #1402)
+#1408 := (forall (vars (?x75 T2)) #1405)
+#970 := (not #969)
+#1353 := (and #216 #970)
+#1358 := (not #1353)
+#1365 := (or #1358 #1361)
+#1368 := (forall (vars (?x71 T2) (?x72 T2)) #1365)
+#1371 := (not #1368)
+#1411 := (or #1371 #1408)
+#1414 := (and #1368 #1411)
+#225 := (and #223 #224)
+#711 := (not #225)
+#1341 := (or #711 #1336)
+#1344 := (forall (vars (?x67 T2) (?x68 T2)) #1341)
+#1347 := (not #1344)
+#1417 := (or #1347 #1414)
+#1420 := (and #1344 #1417)
+#1330 := (forall (vars (?x65 T2)) #1327)
+#1333 := (not #1330)
+#1423 := (or #1333 #1420)
+#1426 := (and #1330 #1423)
+#1429 := (or #1324 #1426)
+#1432 := (and #220 #1429)
+#705 := (forall (vars (?x63 T2)) #700)
+#814 := (not #705)
+#1435 := (or #814 #1432)
+#1438 := (and #705 #1435)
+#1316 := (forall (vars (?x61 T2)) #1312)
+#1319 := (not #1316)
+#1441 := (or #1319 #1438)
+#1444 := (and #1316 #1441)
+#1283 := (not #1282)
+#1278 := (forall (vars (?x58 T2)) #1275)
+#1239 := (not #1238)
+#1232 := (not #1231)
+#1242 := (and #1232 #1239)
+#1263 := (or #658 #1242)
+#1268 := (forall (vars (?x60 T2)) #1263)
+#1245 := (not #1242)
+#1254 := (or #1245 #1251)
+#1257 := (forall (vars (?x59 T2)) #1254)
+#1057 := (not #1056)
+#1132 := (and #115 #1057)
+#1135 := (exists (vars (?x48 T2)) #1132)
+#184 := (not #183)
+#1304 := (and #184 #192 #1135 #1257 #1268 #1278 #1283)
+#1309 := (not #1304)
+#1447 := (or #1309 #1444)
+#1155 := (not #1154)
+#1189 := (and #970 #1155)
+#1192 := (not #1189)
+#1198 := (or #1192 #1195)
+#1201 := (forall (vars (?x53 T2) (?x54 T2)) #1198)
+#1204 := (not #1201)
+#1212 := (or #176 #1204)
+#1217 := (and #1201 #1212)
+#1170 := (= #1172 0::int)
+#1164 := (>= #1166 0::int)
+#1167 := (not #1164)
+#1174 := (and #1167 #1170)
+#1177 := (exists (vars (?x50 T2)) #1174)
+#1158 := (and #70 #1155)
+#1161 := (not #1158)
+#1180 := (or #1161 #1177)
+#1183 := (forall (vars (?x49 T2)) #1180)
+#1186 := (not #1183)
+#1220 := (or #1186 #1217)
+#1223 := (and #1183 #1220)
+#1138 := (not #1135)
+#1144 := (and #149 #154 #530 #533 #1138)
+#1149 := (not #1144)
+#1226 := (or #1149 #1223)
+#1450 := (and #1226 #1447)
+#1104 := (forall (vars (?x41 T2)) #1101)
+#117 := (and #115 #116)
+#456 := (not #117)
+#1095 := (or #456 #1069)
+#1098 := (forall (vars (?x42 T2) (?x43 T2)) #1095)
+#1070 := (not #1069)
+#1066 := (= #1047 0::int)
+#1079 := (and #114 #1066 #1070)
+#1084 := (exists (vars (?x47 T2)) #1079)
+#1060 := (and #70 #1057)
+#1063 := (not #1060)
+#1087 := (or #1063 #1084)
+#1090 := (forall (vars (?x46 T2)) #1087)
+#1034 := (and #114 #970)
+#1039 := (not #1034)
+#1048 := (or #1039 #1043)
+#1051 := (forall (vars (?x44 T2) (?x45 T2)) #1048)
+#1005 := (= #983 0::int)
+#1008 := (not #954)
+#1017 := (and #74 #1008 #1005)
+#1022 := (exists (vars (?x38 T2)) #1017)
+#996 := (not #995)
+#999 := (and #70 #996)
+#1002 := (not #999)
+#1025 := (or #1002 #1022)
+#1028 := (forall (vars (?x37 T2)) #1025)
+#1124 := (and #109 #1028 #1051 #1090 #1098 #1104)
+#1129 := (not #1124)
+#1453 := (or #1129 #1450)
+#1456 := (and #1028 #1453)
+#973 := (and #74 #970)
+#976 := (not #973)
+#984 := (or #976 #980)
+#987 := (forall (vars (?x33 T2) (?x34 T2)) #984)
+#990 := (not #987)
+#1459 := (or #990 #1456)
+#1462 := (and #987 #1459)
+#85 := (and #75 #84)
+#432 := (not #85)
+#958 := (or #432 #954)
+#961 := (forall (vars (?x29 T2) (?x30 T2)) #958)
+#964 := (not #961)
+#1465 := (or #964 #1462)
+#1468 := (and #961 #1465)
+#946 := (forall (vars (?x27 T2)) #945)
+#949 := (not #946)
+#1471 := (or #949 #1468)
+#1474 := (and #946 #1471)
+#1477 := (or #940 #1474)
+#1480 := (and #81 #1477)
+#935 := (not #930)
+#1483 := (or #935 #1480)
+#1486 := (not #1483)
+#231 := (+ #199 #91)
+#238 := (= #226 #231)
+#239 := (and #216 #238)
+#237 := (< #199 #226)
+#240 := (and #237 #239)
+#241 := (exists (vars (?x76 T2)) #240)
+#235 := (< #199 uf_9)
+#236 := (and #70 #235)
+#242 := (implies #236 #241)
+#243 := (forall (vars (?x75 T2)) #242)
+#232 := (<= #226 #231)
+#92 := (< #91 uf_9)
+#230 := (and #216 #92)
+#233 := (implies #230 #232)
+#234 := (forall (vars (?x71 T2) (?x72 T2)) #233)
+#244 := (implies #234 #243)
+#245 := (and #234 #244)
+#227 := (<= #226 #199)
+#228 := (implies #225 #227)
+#229 := (forall (vars (?x67 T2) (?x68 T2)) #228)
+#246 := (implies #229 #245)
+#247 := (and #229 #246)
+#221 := (<= 0::int #199)
+#222 := (forall (vars (?x65 T2)) #221)
+#248 := (implies #222 #247)
+#249 := (and #222 #248)
+#250 := (implies #220 #249)
+#251 := (and #220 #250)
+#204 := (= #199 #110)
+#217 := (implies #216 #204)
+#218 := (forall (vars (?x63 T2)) #217)
+#252 := (implies #218 #251)
+#253 := (and #218 #252)
+#214 := (<= #199 #110)
+#215 := (forall (vars (?x61 T2)) #214)
+#254 := (implies #215 #253)
+#255 := (and #215 #254)
+#196 := (+ #185 #194)
+#197 := (< #196 #110)
+#195 := (< #194 uf_9)
+#198 := (and #195 #197)
+#203 := (not #198)
+#205 := (implies #203 #204)
+#206 := (forall (vars (?x60 T2)) #205)
+#200 := (= #199 #196)
+#201 := (implies #198 #200)
+#202 := (forall (vars (?x59 T2)) #201)
+#207 := (and #202 #206)
+#208 := (and #192 #207)
+#187 := (<= #185 #110)
+#188 := (implies #115 #187)
+#189 := (forall (vars (?x58 T2)) #188)
+#209 := (and #189 #208)
+#186 := (< #185 uf_9)
+#210 := (and #186 #209)
+#211 := (and #184 #210)
+#127 := (< #110 uf_9)
+#142 := (and #115 #127)
+#143 := (exists (vars (?x48 T2)) #142)
+#212 := (and #143 #211)
+#213 := (and true #212)
+#256 := (implies #213 #255)
+#165 := (+ #160 #91)
+#172 := (<= #163 #165)
+#161 := (< #160 uf_9)
+#171 := (and #161 #92)
+#173 := (implies #171 #172)
+#174 := (forall (vars (?x53 T2) (?x54 T2)) #173)
+#177 := (implies #174 #176)
+#178 := (and #174 #177)
+#166 := (= #163 #165)
+#164 := (< #160 #163)
+#167 := (and #164 #166)
+#168 := (exists (vars (?x50 T2)) #167)
+#162 := (and #70 #161)
+#169 := (implies #162 #168)
+#170 := (forall (vars (?x49 T2)) #169)
+#179 := (implies #170 #178)
+#180 := (and #170 #179)
+#151 := (= uf_19 uf_14)
+#155 := (and #151 #154)
+#156 := (and #149 #155)
+#146 := (= uf_16 uf_15)
+#157 := (and #146 #156)
+#144 := (not #143)
+#158 := (and #144 #157)
+#159 := (and true #158)
+#181 := (implies #159 #180)
+#257 := (and #181 #256)
+#123 := (+ #110 #91)
+#130 := (= #118 #123)
+#131 := (and #114 #130)
+#129 := (< #110 #118)
+#132 := (and #129 #131)
+#133 := (exists (vars (?x47 T2)) #132)
+#128 := (and #70 #127)
+#134 := (implies #128 #133)
+#135 := (forall (vars (?x46 T2)) #134)
+#124 := (<= #118 #123)
+#122 := (and #114 #92)
+#125 := (implies #122 #124)
+#126 := (forall (vars (?x44 T2) (?x45 T2)) #125)
+#136 := (and #126 #135)
+#119 := (<= #118 #110)
+#120 := (implies #117 #119)
+#121 := (forall (vars (?x42 T2) (?x43 T2)) #120)
+#137 := (and #121 #136)
+#111 := (<= 0::int #110)
+#112 := (forall (vars (?x41 T2)) #111)
+#138 := (and #112 #137)
+#139 := (and #109 #138)
+#140 := (and true #139)
+#94 := (+ #66 #91)
+#101 := (= #86 #94)
+#102 := (and #74 #101)
+#100 := (< #66 #86)
+#103 := (and #100 #102)
+#104 := (exists (vars (?x38 T2)) #103)
+#98 := (< #66 uf_9)
+#99 := (and #70 #98)
+#105 := (implies #99 #104)
+#106 := (forall (vars (?x37 T2)) #105)
+#141 := (and #106 #140)
+#258 := (implies #141 #257)
+#259 := (and #106 #258)
+#95 := (<= #86 #94)
+#93 := (and #74 #92)
+#96 := (implies #93 #95)
+#97 := (forall (vars (?x33 T2) (?x34 T2)) #96)
+#260 := (implies #97 #259)
+#261 := (and #97 #260)
+#87 := (<= #86 #66)
+#88 := (implies #85 #87)
+#89 := (forall (vars (?x29 T2) (?x30 T2)) #88)
+#262 := (implies #89 #261)
+#263 := (and #89 #262)
+#82 := (<= 0::int #66)
+#83 := (forall (vars (?x27 T2)) #82)
+#264 := (implies #83 #263)
+#265 := (and #83 #264)
+#266 := (implies #81 #265)
+#267 := (and #81 #266)
#71 := (= #66 uf_9)
#72 := (implies #70 #71)
#73 := (forall (vars (?x25 T2)) #72)
-#281 := (implies #73 #280)
+#77 := (and #73 #76)
#68 := (implies #65 #67)
#69 := (forall (vars (?x24 T2)) #68)
-#282 := (implies #69 #281)
-#283 := (implies true #282)
-#284 := (implies true #283)
-#285 := (not #284)
-#1720 := (iff #285 #1717)
-#726 := (+ #88 #196)
-#744 := (= #216 #726)
-#747 := (and #206 #744)
-#750 := (and #227 #747)
-#753 := (exists (vars (?x76 T2)) #750)
-#759 := (not #226)
-#760 := (or #759 #753)
-#765 := (forall (vars (?x75 T2)) #760)
-#729 := (<= #216 #726)
-#723 := (and #89 #206)
-#735 := (not #723)
-#736 := (or #735 #729)
-#741 := (forall (vars (?x71 T2) (?x72 T2)) #736)
-#787 := (not #741)
-#788 := (or #787 #765)
-#793 := (and #741 #788)
-#717 := (or #716 #217)
-#720 := (forall (vars (?x67 T2) (?x68 T2)) #717)
-#799 := (not #720)
-#800 := (or #799 #793)
-#805 := (and #720 #800)
-#811 := (not #212)
-#812 := (or #811 #805)
-#817 := (and #212 #812)
-#713 := (= 0::int #209)
-#823 := (not #713)
-#824 := (or #823 #817)
-#829 := (and #713 #824)
-#847 := (or #846 #829)
-#852 := (and #710 #847)
-#858 := (not #205)
-#859 := (or #858 #852)
-#864 := (and #205 #859)
-#694 := (or #195 #688)
-#699 := (forall (vars (?x60 T2)) #694)
-#870 := (not #699)
-#871 := (or #870 #864)
-#674 := (= #193 #196)
-#680 := (or #200 #674)
-#685 := (forall (vars (?x59 T2)) #680)
-#879 := (not #685)
-#880 := (or #879 #871)
-#889 := (or #888 #880)
-#668 := (or #111 #184)
-#671 := (forall (vars (?x58 T2)) #668)
-#897 := (not #671)
-#898 := (or #897 #889)
-#906 := (not #183)
-#907 := (or #906 #898)
-#915 := (or #180 #907)
-#923 := (or #135 #915)
-#554 := (= 0::int #161)
-#512 := (+ #88 #146)
-#539 := (<= #149 #512)
-#536 := (and #89 #147)
-#545 := (not #536)
-#546 := (or #545 #539)
-#551 := (forall (vars (?x53 T2) (?x54 T2)) #546)
-#574 := (not #551)
-#575 := (or #574 #554)
-#580 := (and #551 #575)
-#515 := (= #149 #512)
-#518 := (and #150 #515)
-#521 := (exists (vars (?x50 T2)) #518)
-#527 := (not #148)
-#528 := (or #527 #521)
-#533 := (forall (vars (?x49 T2)) #528)
-#586 := (not #533)
-#587 := (or #586 #580)
-#592 := (and #533 #587)
-#606 := (or #605 #592)
-#615 := (or #614 #606)
-#624 := (or #623 #615)
-#633 := (or #632 #624)
-#652 := (or #134 #633)
-#939 := (and #652 #923)
-#464 := (+ #88 #107)
-#482 := (= #115 #464)
-#485 := (and #111 #482)
-#488 := (and #126 #485)
-#491 := (exists (vars (?x47 T2)) #488)
-#497 := (not #125)
-#498 := (or #497 #491)
-#503 := (forall (vars (?x46 T2)) #498)
-#952 := (not #503)
-#953 := (or #952 #939)
-#467 := (<= #115 #464)
-#461 := (and #89 #111)
-#473 := (not #461)
-#474 := (or #473 #467)
-#479 := (forall (vars (?x44 T2) (?x45 T2)) #474)
-#961 := (not #479)
-#962 := (or #961 #953)
-#455 := (or #454 #116)
-#458 := (forall (vars (?x42 T2) (?x43 T2)) #455)
-#970 := (not #458)
-#971 := (or #970 #962)
-#979 := (not #109)
-#980 := (or #979 #971)
-#451 := (= 0::int #105)
-#988 := (not #451)
-#989 := (or #988 #980)
-#444 := (not #96)
-#445 := (or #444 #101)
-#448 := (forall (vars (?x37 T2)) #445)
-#1008 := (not #448)
-#1009 := (or #1008 #989)
-#1014 := (and #448 #1009)
-#437 := (not #90)
-#438 := (or #437 #92)
-#441 := (forall (vars (?x33 T2) (?x34 T2)) #438)
-#1020 := (not #441)
-#1021 := (or #1020 #1014)
-#1026 := (and #441 #1021)
-#431 := (or #430 #84)
-#434 := (forall (vars (?x29 T2) (?x30 T2)) #431)
-#1032 := (not #434)
-#1033 := (or #1032 #1026)
-#1038 := (and #434 #1033)
-#1044 := (not #80)
-#1045 := (or #1044 #1038)
-#1050 := (and #80 #1045)
-#427 := (= 0::int #77)
-#1056 := (not #427)
-#1057 := (or #1056 #1050)
-#1062 := (and #427 #1057)
-#1076 := (or #1075 #1062)
-#1085 := (or #1084 #1076)
-#399 := (= 0::int #66)
-#405 := (or #70 #399)
-#410 := (forall (vars (?x24 T2)) #405)
-#1093 := (not #410)
-#1094 := (or #1093 #1085)
-#1110 := (not #1094)
-#1718 := (iff #1110 #1717)
-#1715 := (iff #1094 #1712)
-#1703 := (or #1075 #1689)
-#1706 := (or #1084 #1703)
-#1709 := (or #1700 #1706)
-#1713 := (iff #1709 #1712)
-#1714 := [rewrite]: #1713
-#1710 := (iff #1094 #1709)
-#1707 := (iff #1085 #1706)
-#1704 := (iff #1076 #1703)
-#1690 := (iff #1062 #1689)
-#1687 := (iff #1057 #1686)
-#1684 := (iff #1050 #1683)
-#1681 := (iff #1045 #1680)
-#1678 := (iff #1038 #1677)
-#1675 := (iff #1033 #1674)
-#1672 := (iff #1026 #1671)
-#1669 := (iff #1021 #1668)
-#1666 := (iff #1014 #1665)
-#1663 := (iff #1009 #1660)
-#1642 := (or #1248 #1588)
-#1645 := (or #1611 #1642)
-#1648 := (or #1622 #1645)
-#1651 := (or #1631 #1648)
-#1654 := (or #1636 #1651)
-#1657 := (or #1639 #1654)
-#1661 := (iff #1657 #1660)
-#1662 := [rewrite]: #1661
-#1658 := (iff #1009 #1657)
-#1655 := (iff #989 #1654)
-#1652 := (iff #980 #1651)
-#1649 := (iff #971 #1648)
-#1646 := (iff #962 #1645)
-#1643 := (iff #953 #1642)
-#1589 := (iff #939 #1588)
-#1586 := (iff #923 #1583)
-#1562 := (or #1378 #1514)
-#1565 := (or #1531 #1562)
-#1568 := (or #888 #1565)
-#1571 := (or #1544 #1568)
-#1574 := (or #1548 #1571)
-#1577 := (or #180 #1574)
-#1580 := (or #1559 #1577)
-#1584 := (iff #1580 #1583)
-#1585 := [rewrite]: #1584
-#1581 := (iff #923 #1580)
-#1578 := (iff #915 #1577)
-#1575 := (iff #907 #1574)
-#1572 := (iff #898 #1571)
-#1569 := (iff #889 #1568)
-#1566 := (iff #880 #1565)
-#1563 := (iff #871 #1562)
-#1515 := (iff #864 #1514)
-#1512 := (iff #859 #1511)
-#1509 := (iff #852 #1508)
-#1506 := (iff #847 #1505)
-#1503 := (iff #829 #1502)
-#1500 := (iff #824 #1499)
-#1497 := (iff #817 #1496)
-#1494 := (iff #812 #1493)
-#1491 := (iff #805 #1490)
-#1488 := (iff #800 #1487)
-#1485 := (iff #793 #1484)
-#1482 := (iff #788 #1481)
-#1479 := (iff #765 #1478)
-#1476 := (iff #760 #1475)
-#1473 := (iff #753 #1472)
-#1470 := (iff #750 #1467)
-#1461 := (and #206 #1455)
-#1464 := (and #1458 #1461)
-#1468 := (iff #1464 #1467)
-#1469 := [rewrite]: #1468
-#1465 := (iff #750 #1464)
-#1462 := (iff #747 #1461)
-#1456 := (iff #744 #1455)
-#1457 := [rewrite]: #1456
-#1463 := [monotonicity #1457]: #1462
-#1459 := (iff #227 #1458)
-#1460 := [rewrite]: #1459
-#1466 := [monotonicity #1460 #1463]: #1465
-#1471 := [trans #1466 #1469]: #1470
-#1474 := [quant-intro #1471]: #1473
-#1453 := (iff #759 #1452)
-#1450 := (iff #226 #1449)
-#1447 := (iff #225 #1446)
-#1448 := [rewrite]: #1447
-#1451 := [monotonicity #1448]: #1450
-#1454 := [monotonicity #1451]: #1453
-#1477 := [monotonicity #1454 #1474]: #1476
-#1480 := [quant-intro #1477]: #1479
-#1442 := (iff #787 #1441)
-#1439 := (iff #741 #1438)
-#1436 := (iff #736 #1435)
-#1432 := (iff #729 #1431)
-#1434 := [rewrite]: #1432
-#1429 := (iff #735 #1428)
-#1426 := (iff #723 #1423)
-#1420 := (and #1145 #206)
-#1424 := (iff #1420 #1423)
-#1425 := [rewrite]: #1424
-#1421 := (iff #723 #1420)
-#1146 := (iff #89 #1145)
-#1147 := [rewrite]: #1146
-#1422 := [monotonicity #1147]: #1421
-#1427 := [trans #1422 #1425]: #1426
-#1430 := [monotonicity #1427]: #1429
-#1437 := [monotonicity #1430 #1434]: #1436
-#1440 := [quant-intro #1437]: #1439
-#1443 := [monotonicity #1440]: #1442
-#1483 := [monotonicity #1443 #1480]: #1482
-#1486 := [monotonicity #1440 #1483]: #1485
-#1418 := (iff #799 #1417)
-#1415 := (iff #720 #1414)
-#1412 := (iff #717 #1411)
-#1409 := (iff #217 #1406)
-#1410 := [rewrite]: #1409
-#1413 := [monotonicity #1410]: #1412
-#1416 := [quant-intro #1413]: #1415
-#1419 := [monotonicity #1416]: #1418
-#1489 := [monotonicity #1419 #1486]: #1488
-#1492 := [monotonicity #1416 #1489]: #1491
-#1404 := (iff #811 #1403)
-#1401 := (iff #212 #1400)
-#1398 := (iff #211 #1397)
+#78 := (and #69 #77)
+#79 := (and true #78)
+#268 := (implies #79 #267)
+#269 := (not #268)
+#1489 := (iff #269 #1486)
+#721 := (+ #91 #199)
+#739 := (= #226 #721)
+#742 := (and #216 #739)
+#745 := (and #237 #742)
+#748 := (exists (vars (?x76 T2)) #745)
+#754 := (not #236)
+#755 := (or #754 #748)
+#760 := (forall (vars (?x75 T2)) #755)
+#724 := (<= #226 #721)
+#718 := (and #92 #216)
+#730 := (not #718)
+#731 := (or #730 #724)
+#736 := (forall (vars (?x71 T2) (?x72 T2)) #731)
+#766 := (not #736)
+#767 := (or #766 #760)
+#772 := (and #736 #767)
+#712 := (or #711 #227)
+#715 := (forall (vars (?x67 T2) (?x68 T2)) #712)
+#778 := (not #715)
+#779 := (or #778 #772)
+#784 := (and #715 #779)
+#790 := (not #222)
+#791 := (or #790 #784)
+#796 := (and #222 #791)
+#708 := (= 0::int #219)
+#802 := (not #708)
+#803 := (or #802 #796)
+#808 := (and #708 #803)
+#815 := (or #814 #808)
+#820 := (and #705 #815)
+#826 := (not #215)
+#827 := (or #826 #820)
+#832 := (and #215 #827)
+#664 := (or #198 #658)
+#669 := (forall (vars (?x60 T2)) #664)
+#644 := (= #196 #199)
+#650 := (or #203 #644)
+#655 := (forall (vars (?x59 T2)) #650)
+#672 := (and #655 #669)
+#675 := (and #192 #672)
+#638 := (or #114 #187)
+#641 := (forall (vars (?x58 T2)) #638)
+#678 := (and #641 #675)
+#681 := (and #186 #678)
+#684 := (and #184 #681)
+#687 := (and #143 #684)
+#838 := (not #687)
+#839 := (or #838 #832)
+#602 := (= 0::int #175)
+#560 := (+ #91 #160)
+#587 := (<= #163 #560)
+#584 := (and #92 #161)
+#593 := (not #584)
+#594 := (or #593 #587)
+#599 := (forall (vars (?x53 T2) (?x54 T2)) #594)
+#608 := (not #599)
+#609 := (or #608 #602)
+#614 := (and #599 #609)
+#563 := (= #163 #560)
+#566 := (and #164 #563)
+#569 := (exists (vars (?x50 T2)) #566)
+#575 := (not #162)
+#576 := (or #575 #569)
+#581 := (forall (vars (?x49 T2)) #576)
+#620 := (not #581)
+#621 := (or #620 #614)
+#626 := (and #581 #621)
+#539 := (and #154 #533)
+#544 := (and #149 #539)
+#547 := (and #530 #544)
+#550 := (and #144 #547)
+#632 := (not #550)
+#633 := (or #632 #626)
+#844 := (and #633 #839)
+#466 := (+ #91 #110)
+#484 := (= #118 #466)
+#487 := (and #114 #484)
+#490 := (and #129 #487)
+#493 := (exists (vars (?x47 T2)) #490)
+#499 := (not #128)
+#500 := (or #499 #493)
+#505 := (forall (vars (?x46 T2)) #500)
+#469 := (<= #118 #466)
+#463 := (and #92 #114)
+#475 := (not #463)
+#476 := (or #475 #469)
+#481 := (forall (vars (?x44 T2) (?x45 T2)) #476)
+#508 := (and #481 #505)
+#457 := (or #456 #119)
+#460 := (forall (vars (?x42 T2) (?x43 T2)) #457)
+#511 := (and #460 #508)
+#514 := (and #112 #511)
+#453 := (= 0::int #108)
+#517 := (and #453 #514)
+#446 := (not #99)
+#447 := (or #446 #104)
+#450 := (forall (vars (?x37 T2)) #447)
+#527 := (and #450 #517)
+#850 := (not #527)
+#851 := (or #850 #844)
+#856 := (and #450 #851)
+#439 := (not #93)
+#440 := (or #439 #95)
+#443 := (forall (vars (?x33 T2) (?x34 T2)) #440)
+#862 := (not #443)
+#863 := (or #862 #856)
+#868 := (and #443 #863)
+#433 := (or #432 #87)
+#436 := (forall (vars (?x29 T2) (?x30 T2)) #433)
+#874 := (not #436)
+#875 := (or #874 #868)
+#880 := (and #436 #875)
+#886 := (not #83)
+#887 := (or #886 #880)
+#892 := (and #83 #887)
+#429 := (= 0::int #80)
+#898 := (not #429)
+#899 := (or #898 #892)
+#904 := (and #429 #899)
+#414 := (and #76 #408)
+#383 := (= 0::int #66)
+#389 := (or #70 #383)
+#394 := (forall (vars (?x24 T2)) #389)
+#419 := (and #394 #414)
+#910 := (not #419)
+#911 := (or #910 #904)
+#916 := (not #911)
+#1487 := (iff #916 #1486)
+#1484 := (iff #911 #1483)
+#1481 := (iff #904 #1480)
+#1478 := (iff #899 #1477)
+#1475 := (iff #892 #1474)
+#1472 := (iff #887 #1471)
+#1469 := (iff #880 #1468)
+#1466 := (iff #875 #1465)
+#1463 := (iff #868 #1462)
+#1460 := (iff #863 #1459)
+#1457 := (iff #856 #1456)
+#1454 := (iff #851 #1453)
+#1451 := (iff #844 #1450)
+#1448 := (iff #839 #1447)
+#1445 := (iff #832 #1444)
+#1442 := (iff #827 #1441)
+#1439 := (iff #820 #1438)
+#1436 := (iff #815 #1435)
+#1433 := (iff #808 #1432)
+#1430 := (iff #803 #1429)
+#1427 := (iff #796 #1426)
+#1424 := (iff #791 #1423)
+#1421 := (iff #784 #1420)
+#1418 := (iff #779 #1417)
+#1415 := (iff #772 #1414)
+#1412 := (iff #767 #1411)
+#1409 := (iff #760 #1408)
+#1406 := (iff #755 #1405)
+#1403 := (iff #748 #1402)
+#1400 := (iff #745 #1397)
+#1391 := (and #216 #1385)
+#1394 := (and #1388 #1391)
+#1398 := (iff #1394 #1397)
#1399 := [rewrite]: #1398
-#1402 := [quant-intro #1399]: #1401
-#1405 := [monotonicity #1402]: #1404
-#1495 := [monotonicity #1405 #1492]: #1494
-#1498 := [monotonicity #1402 #1495]: #1497
-#1395 := (iff #823 #1394)
-#1392 := (iff #713 #210)
-#1393 := [rewrite]: #1392
-#1396 := [monotonicity #1393]: #1395
-#1501 := [monotonicity #1396 #1498]: #1500
-#1504 := [monotonicity #1393 #1501]: #1503
-#1507 := [monotonicity #1504]: #1506
-#1510 := [monotonicity #1507]: #1509
-#1390 := (iff #858 #1389)
-#1387 := (iff #205 #1386)
-#1384 := (iff #204 #1381)
-#1385 := [rewrite]: #1384
-#1388 := [quant-intro #1385]: #1387
-#1391 := [monotonicity #1388]: #1390
-#1513 := [monotonicity #1391 #1510]: #1512
-#1516 := [monotonicity #1388 #1513]: #1515
-#1379 := (iff #870 #1378)
-#1376 := (iff #699 #1375)
-#1373 := (iff #694 #1370)
-#1367 := (or #1364 #688)
-#1371 := (iff #1367 #1370)
-#1372 := [rewrite]: #1371
-#1368 := (iff #694 #1367)
-#1365 := (iff #195 #1364)
-#1362 := (iff #194 #1361)
-#1363 := [rewrite]: #1362
-#1355 := (iff #192 #1354)
-#1356 := [rewrite]: #1355
-#1366 := [monotonicity #1356 #1363]: #1365
-#1369 := [monotonicity #1366]: #1368
-#1374 := [trans #1369 #1372]: #1373
-#1377 := [quant-intro #1374]: #1376
-#1380 := [monotonicity #1377]: #1379
-#1564 := [monotonicity #1380 #1516]: #1563
-#1532 := (iff #879 #1531)
-#1529 := (iff #685 #1528)
-#1526 := (iff #680 #1525)
-#1523 := (iff #674 #1522)
-#1524 := [rewrite]: #1523
-#1518 := (iff #200 #1517)
-#1519 := [monotonicity #1366]: #1518
-#1527 := [monotonicity #1519 #1524]: #1526
-#1530 := [quant-intro #1527]: #1529
-#1533 := [monotonicity #1530]: #1532
-#1567 := [monotonicity #1533 #1564]: #1566
-#1570 := [monotonicity #1567]: #1569
-#1545 := (iff #897 #1544)
-#1542 := (iff #671 #1541)
-#1539 := (iff #668 #1538)
-#1536 := (iff #184 #1534)
-#1537 := [rewrite]: #1536
-#1540 := [monotonicity #1537]: #1539
-#1543 := [quant-intro #1540]: #1542
-#1546 := [monotonicity #1543]: #1545
-#1573 := [monotonicity #1546 #1570]: #1572
-#1557 := (iff #906 #1548)
-#1549 := (not #1548)
-#1552 := (not #1549)
-#1555 := (iff #1552 #1548)
-#1556 := [rewrite]: #1555
-#1553 := (iff #906 #1552)
-#1550 := (iff #183 #1549)
-#1551 := [rewrite]: #1550
-#1554 := [monotonicity #1551]: #1553
-#1558 := [trans #1554 #1556]: #1557
-#1576 := [monotonicity #1558 #1573]: #1575
-#1579 := [monotonicity #1576]: #1578
-#1560 := (iff #135 #1559)
-#1329 := (iff #134 #1328)
-#1326 := (iff #133 #1325)
-#1210 := (iff #124 #1209)
-#1211 := [rewrite]: #1210
-#1327 := [monotonicity #1211]: #1326
-#1330 := [quant-intro #1327]: #1329
-#1561 := [monotonicity #1330]: #1560
-#1582 := [monotonicity #1561 #1579]: #1581
-#1587 := [trans #1582 #1585]: #1586
-#1349 := (iff #652 #1346)
-#1331 := (or #605 #1322)
-#1334 := (or #614 #1331)
-#1337 := (or #623 #1334)
-#1340 := (or #632 #1337)
-#1343 := (or #1328 #1340)
-#1347 := (iff #1343 #1346)
-#1348 := [rewrite]: #1347
-#1344 := (iff #652 #1343)
-#1341 := (iff #633 #1340)
-#1338 := (iff #624 #1337)
-#1335 := (iff #615 #1334)
-#1332 := (iff #606 #1331)
-#1323 := (iff #592 #1322)
-#1320 := (iff #587 #1319)
-#1317 := (iff #580 #1316)
-#1314 := (iff #575 #1311)
-#1308 := (or #1303 #162)
-#1312 := (iff #1308 #1311)
-#1313 := [rewrite]: #1312
-#1309 := (iff #575 #1308)
-#1306 := (iff #554 #162)
-#1307 := [rewrite]: #1306
-#1304 := (iff #574 #1303)
-#1301 := (iff #551 #1300)
-#1298 := (iff #546 #1297)
-#1295 := (iff #539 #1294)
-#1296 := [rewrite]: #1295
-#1292 := (iff #545 #1291)
-#1289 := (iff #536 #1288)
-#1255 := (iff #147 #1254)
-#1256 := [rewrite]: #1255
-#1290 := [monotonicity #1147 #1256]: #1289
-#1293 := [monotonicity #1290]: #1292
-#1299 := [monotonicity #1293 #1296]: #1298
-#1302 := [quant-intro #1299]: #1301
-#1305 := [monotonicity #1302]: #1304
-#1310 := [monotonicity #1305 #1307]: #1309
-#1315 := [trans #1310 #1313]: #1314
-#1318 := [monotonicity #1302 #1315]: #1317
-#1286 := (iff #586 #1285)
-#1283 := (iff #533 #1282)
-#1280 := (iff #528 #1279)
-#1277 := (iff #521 #1276)
-#1274 := (iff #518 #1273)
-#1270 := (iff #515 #1269)
-#1272 := [rewrite]: #1270
-#1267 := (iff #150 #1266)
-#1268 := [rewrite]: #1267
-#1275 := [monotonicity #1268 #1272]: #1274
-#1278 := [quant-intro #1275]: #1277
-#1261 := (iff #527 #1260)
-#1258 := (iff #148 #1257)
-#1259 := [monotonicity #1256]: #1258
-#1262 := [monotonicity #1259]: #1261
-#1281 := [monotonicity #1262 #1278]: #1280
-#1284 := [quant-intro #1281]: #1283
-#1287 := [monotonicity #1284]: #1286
-#1321 := [monotonicity #1287 #1318]: #1320
-#1324 := [monotonicity #1284 #1321]: #1323
-#1333 := [monotonicity #1324]: #1332
-#1336 := [monotonicity #1333]: #1335
-#1339 := [monotonicity #1336]: #1338
-#1342 := [monotonicity #1339]: #1341
-#1345 := [monotonicity #1330 #1342]: #1344
-#1350 := [trans #1345 #1348]: #1349
-#1590 := [monotonicity #1350 #1587]: #1589
-#1249 := (iff #952 #1248)
-#1246 := (iff #503 #1245)
-#1243 := (iff #498 #1242)
-#1240 := (iff #491 #1239)
-#1237 := (iff #488 #1234)
-#1228 := (and #111 #1218)
-#1231 := (and #1225 #1228)
-#1235 := (iff #1231 #1234)
-#1236 := [rewrite]: #1235
-#1232 := (iff #488 #1231)
-#1229 := (iff #485 #1228)
-#1219 := (iff #482 #1218)
-#1223 := [rewrite]: #1219
-#1230 := [monotonicity #1223]: #1229
-#1226 := (iff #126 #1225)
-#1227 := [rewrite]: #1226
-#1233 := [monotonicity #1227 #1230]: #1232
-#1238 := [trans #1233 #1236]: #1237
-#1241 := [quant-intro #1238]: #1240
-#1216 := (iff #497 #1215)
-#1213 := (iff #125 #1212)
-#1214 := [monotonicity #1211]: #1213
-#1217 := [monotonicity #1214]: #1216
-#1244 := [monotonicity #1217 #1241]: #1243
-#1247 := [quant-intro #1244]: #1246
-#1250 := [monotonicity #1247]: #1249
-#1644 := [monotonicity #1250 #1590]: #1643
-#1612 := (iff #961 #1611)
-#1609 := (iff #479 #1608)
-#1606 := (iff #474 #1605)
-#1603 := (iff #467 #1602)
-#1604 := [rewrite]: #1603
-#1600 := (iff #473 #1599)
-#1597 := (iff #461 #1594)
-#1591 := (and #1145 #111)
-#1595 := (iff #1591 #1594)
-#1596 := [rewrite]: #1595
-#1592 := (iff #461 #1591)
-#1593 := [monotonicity #1147]: #1592
-#1598 := [trans #1593 #1596]: #1597
-#1601 := [monotonicity #1598]: #1600
-#1607 := [monotonicity #1601 #1604]: #1606
-#1610 := [quant-intro #1607]: #1609
-#1613 := [monotonicity #1610]: #1612
-#1647 := [monotonicity #1613 #1644]: #1646
-#1623 := (iff #970 #1622)
-#1620 := (iff #458 #1619)
-#1617 := (iff #455 #1616)
-#1614 := (iff #116 #1224)
-#1615 := [rewrite]: #1614
-#1618 := [monotonicity #1615]: #1617
-#1621 := [quant-intro #1618]: #1620
-#1624 := [monotonicity #1621]: #1623
-#1650 := [monotonicity #1624 #1647]: #1649
-#1632 := (iff #979 #1631)
-#1629 := (iff #109 #1628)
-#1626 := (iff #108 #1625)
-#1627 := [rewrite]: #1626
-#1630 := [quant-intro #1627]: #1629
-#1633 := [monotonicity #1630]: #1632
-#1653 := [monotonicity #1633 #1650]: #1652
-#1637 := (iff #988 #1636)
-#1634 := (iff #451 #106)
-#1635 := [rewrite]: #1634
-#1638 := [monotonicity #1635]: #1637
-#1656 := [monotonicity #1638 #1653]: #1655
-#1640 := (iff #1008 #1639)
-#1204 := (iff #448 #1203)
-#1201 := (iff #445 #1200)
-#1198 := (iff #101 #1197)
-#1195 := (iff #100 #1192)
-#1186 := (and #74 #1180)
-#1189 := (and #1183 #1186)
-#1193 := (iff #1189 #1192)
-#1194 := [rewrite]: #1193
-#1190 := (iff #100 #1189)
-#1187 := (iff #99 #1186)
-#1181 := (iff #98 #1180)
-#1182 := [rewrite]: #1181
-#1188 := [monotonicity #1182]: #1187
-#1184 := (iff #97 #1183)
-#1185 := [rewrite]: #1184
-#1191 := [monotonicity #1185 #1188]: #1190
-#1196 := [trans #1191 #1194]: #1195
-#1199 := [quant-intro #1196]: #1198
-#1178 := (iff #444 #1177)
-#1175 := (iff #96 #1174)
-#1172 := (iff #95 #1171)
-#1173 := [rewrite]: #1172
-#1176 := [monotonicity #1173]: #1175
-#1179 := [monotonicity #1176]: #1178
-#1202 := [monotonicity #1179 #1199]: #1201
-#1205 := [quant-intro #1202]: #1204
-#1641 := [monotonicity #1205]: #1640
-#1659 := [monotonicity #1641 #1656]: #1658
-#1664 := [trans #1659 #1662]: #1663
-#1667 := [monotonicity #1205 #1664]: #1666
-#1166 := (iff #1020 #1165)
-#1163 := (iff #441 #1162)
-#1160 := (iff #438 #1159)
-#1154 := (iff #92 #1155)
-#1156 := [rewrite]: #1154
-#1152 := (iff #437 #1151)
-#1149 := (iff #90 #1148)
-#1150 := [monotonicity #1147]: #1149
-#1153 := [monotonicity #1150]: #1152
-#1161 := [monotonicity #1153 #1156]: #1160
-#1164 := [quant-intro #1161]: #1163
-#1167 := [monotonicity #1164]: #1166
-#1670 := [monotonicity #1167 #1667]: #1669
-#1673 := [monotonicity #1164 #1670]: #1672
-#1140 := (iff #1032 #1139)
-#1137 := (iff #434 #1136)
-#1134 := (iff #431 #1133)
-#1128 := (iff #84 #1129)
-#1132 := [rewrite]: #1128
-#1135 := [monotonicity #1132]: #1134
-#1138 := [quant-intro #1135]: #1137
-#1141 := [monotonicity #1138]: #1140
-#1676 := [monotonicity #1141 #1673]: #1675
-#1679 := [monotonicity #1138 #1676]: #1678
-#1125 := (iff #1044 #1124)
-#1122 := (iff #80 #1121)
-#1118 := (iff #79 #1120)
-#1119 := [rewrite]: #1118
-#1123 := [quant-intro #1119]: #1122
-#1126 := [monotonicity #1123]: #1125
-#1682 := [monotonicity #1126 #1679]: #1681
-#1685 := [monotonicity #1123 #1682]: #1684
-#1116 := (iff #1056 #1115)
-#1113 := (iff #427 #78)
-#1114 := [rewrite]: #1113
-#1117 := [monotonicity #1114]: #1116
-#1688 := [monotonicity #1117 #1685]: #1687
-#1691 := [monotonicity #1114 #1688]: #1690
-#1705 := [monotonicity #1691]: #1704
-#1708 := [monotonicity #1705]: #1707
-#1701 := (iff #1093 #1700)
-#1698 := (iff #410 #1697)
-#1695 := (iff #405 #1694)
-#1692 := (iff #399 #67)
-#1693 := [rewrite]: #1692
-#1696 := [monotonicity #1693]: #1695
-#1699 := [quant-intro #1696]: #1698
-#1702 := [monotonicity #1699]: #1701
-#1711 := [monotonicity #1702 #1708]: #1710
-#1716 := [trans #1711 #1714]: #1715
-#1719 := [monotonicity #1716]: #1718
-#1111 := (iff #285 #1110)
-#1108 := (iff #284 #1094)
-#1099 := (implies true #1094)
-#1102 := (iff #1099 #1094)
+#1395 := (iff #745 #1394)
+#1392 := (iff #742 #1391)
+#1386 := (iff #739 #1385)
+#1387 := [rewrite]: #1386
+#1393 := [monotonicity #1387]: #1392
+#1389 := (iff #237 #1388)
+#1390 := [rewrite]: #1389
+#1396 := [monotonicity #1390 #1393]: #1395
+#1401 := [trans #1396 #1399]: #1400
+#1404 := [quant-intro #1401]: #1403
+#1383 := (iff #754 #1382)
+#1380 := (iff #236 #1379)
+#1377 := (iff #235 #1376)
+#1378 := [rewrite]: #1377
+#1381 := [monotonicity #1378]: #1380
+#1384 := [monotonicity #1381]: #1383
+#1407 := [monotonicity #1384 #1404]: #1406
+#1410 := [quant-intro #1407]: #1409
+#1372 := (iff #766 #1371)
+#1369 := (iff #736 #1368)
+#1366 := (iff #731 #1365)
+#1362 := (iff #724 #1361)
+#1364 := [rewrite]: #1362
+#1359 := (iff #730 #1358)
+#1356 := (iff #718 #1353)
+#1350 := (and #970 #216)
+#1354 := (iff #1350 #1353)
+#1355 := [rewrite]: #1354
+#1351 := (iff #718 #1350)
+#971 := (iff #92 #970)
+#972 := [rewrite]: #971
+#1352 := [monotonicity #972]: #1351
+#1357 := [trans #1352 #1355]: #1356
+#1360 := [monotonicity #1357]: #1359
+#1367 := [monotonicity #1360 #1364]: #1366
+#1370 := [quant-intro #1367]: #1369
+#1373 := [monotonicity #1370]: #1372
+#1413 := [monotonicity #1373 #1410]: #1412
+#1416 := [monotonicity #1370 #1413]: #1415
+#1348 := (iff #778 #1347)
+#1345 := (iff #715 #1344)
+#1342 := (iff #712 #1341)
+#1339 := (iff #227 #1336)
+#1340 := [rewrite]: #1339
+#1343 := [monotonicity #1340]: #1342
+#1346 := [quant-intro #1343]: #1345
+#1349 := [monotonicity #1346]: #1348
+#1419 := [monotonicity #1349 #1416]: #1418
+#1422 := [monotonicity #1346 #1419]: #1421
+#1334 := (iff #790 #1333)
+#1331 := (iff #222 #1330)
+#1328 := (iff #221 #1327)
+#1329 := [rewrite]: #1328
+#1332 := [quant-intro #1329]: #1331
+#1335 := [monotonicity #1332]: #1334
+#1425 := [monotonicity #1335 #1422]: #1424
+#1428 := [monotonicity #1332 #1425]: #1427
+#1325 := (iff #802 #1324)
+#1322 := (iff #708 #220)
+#1323 := [rewrite]: #1322
+#1326 := [monotonicity #1323]: #1325
+#1431 := [monotonicity #1326 #1428]: #1430
+#1434 := [monotonicity #1323 #1431]: #1433
+#1437 := [monotonicity #1434]: #1436
+#1440 := [monotonicity #1437]: #1439
+#1320 := (iff #826 #1319)
+#1317 := (iff #215 #1316)
+#1314 := (iff #214 #1312)
+#1315 := [rewrite]: #1314
+#1318 := [quant-intro #1315]: #1317
+#1321 := [monotonicity #1318]: #1320
+#1443 := [monotonicity #1321 #1440]: #1442
+#1446 := [monotonicity #1318 #1443]: #1445
+#1310 := (iff #838 #1309)
+#1307 := (iff #687 #1304)
+#1286 := (and #1257 #1268)
+#1289 := (and #192 #1286)
+#1292 := (and #1278 #1289)
+#1295 := (and #1283 #1292)
+#1298 := (and #184 #1295)
+#1301 := (and #1135 #1298)
+#1305 := (iff #1301 #1304)
+#1306 := [rewrite]: #1305
+#1302 := (iff #687 #1301)
+#1299 := (iff #684 #1298)
+#1296 := (iff #681 #1295)
+#1293 := (iff #678 #1292)
+#1290 := (iff #675 #1289)
+#1287 := (iff #672 #1286)
+#1269 := (iff #669 #1268)
+#1266 := (iff #664 #1263)
+#1260 := (or #1242 #658)
+#1264 := (iff #1260 #1263)
+#1265 := [rewrite]: #1264
+#1261 := (iff #664 #1260)
+#1243 := (iff #198 #1242)
+#1240 := (iff #197 #1239)
+#1241 := [rewrite]: #1240
+#1233 := (iff #195 #1232)
+#1234 := [rewrite]: #1233
+#1244 := [monotonicity #1234 #1241]: #1243
+#1262 := [monotonicity #1244]: #1261
+#1267 := [trans #1262 #1265]: #1266
+#1270 := [quant-intro #1267]: #1269
+#1258 := (iff #655 #1257)
+#1255 := (iff #650 #1254)
+#1252 := (iff #644 #1251)
+#1253 := [rewrite]: #1252
+#1246 := (iff #203 #1245)
+#1247 := [monotonicity #1244]: #1246
+#1256 := [monotonicity #1247 #1253]: #1255
+#1259 := [quant-intro #1256]: #1258
+#1288 := [monotonicity #1259 #1270]: #1287
+#1291 := [monotonicity #1288]: #1290
+#1279 := (iff #641 #1278)
+#1276 := (iff #638 #1275)
+#1273 := (iff #187 #1271)
+#1274 := [rewrite]: #1273
+#1277 := [monotonicity #1274]: #1276
+#1280 := [quant-intro #1277]: #1279
+#1294 := [monotonicity #1280 #1291]: #1293
+#1284 := (iff #186 #1283)
+#1285 := [rewrite]: #1284
+#1297 := [monotonicity #1285 #1294]: #1296
+#1300 := [monotonicity #1297]: #1299
+#1136 := (iff #143 #1135)
+#1133 := (iff #142 #1132)
+#1058 := (iff #127 #1057)
+#1059 := [rewrite]: #1058
+#1134 := [monotonicity #1059]: #1133
+#1137 := [quant-intro #1134]: #1136
+#1303 := [monotonicity #1137 #1300]: #1302
+#1308 := [trans #1303 #1306]: #1307
+#1311 := [monotonicity #1308]: #1310
+#1449 := [monotonicity #1311 #1446]: #1448
+#1227 := (iff #633 #1226)
+#1224 := (iff #626 #1223)
+#1221 := (iff #621 #1220)
+#1218 := (iff #614 #1217)
+#1215 := (iff #609 #1212)
+#1209 := (or #1204 #176)
+#1213 := (iff #1209 #1212)
+#1214 := [rewrite]: #1213
+#1210 := (iff #609 #1209)
+#1207 := (iff #602 #176)
+#1208 := [rewrite]: #1207
+#1205 := (iff #608 #1204)
+#1202 := (iff #599 #1201)
+#1199 := (iff #594 #1198)
+#1196 := (iff #587 #1195)
+#1197 := [rewrite]: #1196
+#1193 := (iff #593 #1192)
+#1190 := (iff #584 #1189)
+#1156 := (iff #161 #1155)
+#1157 := [rewrite]: #1156
+#1191 := [monotonicity #972 #1157]: #1190
+#1194 := [monotonicity #1191]: #1193
+#1200 := [monotonicity #1194 #1197]: #1199
+#1203 := [quant-intro #1200]: #1202
+#1206 := [monotonicity #1203]: #1205
+#1211 := [monotonicity #1206 #1208]: #1210
+#1216 := [trans #1211 #1214]: #1215
+#1219 := [monotonicity #1203 #1216]: #1218
+#1187 := (iff #620 #1186)
+#1184 := (iff #581 #1183)
+#1181 := (iff #576 #1180)
+#1178 := (iff #569 #1177)
+#1175 := (iff #566 #1174)
+#1171 := (iff #563 #1170)
+#1173 := [rewrite]: #1171
+#1168 := (iff #164 #1167)
+#1169 := [rewrite]: #1168
+#1176 := [monotonicity #1169 #1173]: #1175
+#1179 := [quant-intro #1176]: #1178
+#1162 := (iff #575 #1161)
+#1159 := (iff #162 #1158)
+#1160 := [monotonicity #1157]: #1159
+#1163 := [monotonicity #1160]: #1162
+#1182 := [monotonicity #1163 #1179]: #1181
+#1185 := [quant-intro #1182]: #1184
+#1188 := [monotonicity #1185]: #1187
+#1222 := [monotonicity #1188 #1219]: #1221
+#1225 := [monotonicity #1185 #1222]: #1224
+#1150 := (iff #632 #1149)
+#1147 := (iff #550 #1144)
+#1141 := (and #1138 #547)
+#1145 := (iff #1141 #1144)
+#1146 := [rewrite]: #1145
+#1142 := (iff #550 #1141)
+#1139 := (iff #144 #1138)
+#1140 := [monotonicity #1137]: #1139
+#1143 := [monotonicity #1140]: #1142
+#1148 := [trans #1143 #1146]: #1147
+#1151 := [monotonicity #1148]: #1150
+#1228 := [monotonicity #1151 #1225]: #1227
+#1452 := [monotonicity #1228 #1449]: #1451
+#1130 := (iff #850 #1129)
+#1127 := (iff #527 #1124)
+#1109 := (and #1051 #1090)
+#1112 := (and #1098 #1109)
+#1115 := (and #1104 #1112)
+#1118 := (and #109 #1115)
+#1121 := (and #1028 #1118)
+#1125 := (iff #1121 #1124)
+#1126 := [rewrite]: #1125
+#1122 := (iff #527 #1121)
+#1119 := (iff #517 #1118)
+#1116 := (iff #514 #1115)
+#1113 := (iff #511 #1112)
+#1110 := (iff #508 #1109)
+#1091 := (iff #505 #1090)
+#1088 := (iff #500 #1087)
+#1085 := (iff #493 #1084)
+#1082 := (iff #490 #1079)
+#1073 := (and #114 #1066)
+#1076 := (and #1070 #1073)
+#1080 := (iff #1076 #1079)
+#1081 := [rewrite]: #1080
+#1077 := (iff #490 #1076)
+#1074 := (iff #487 #1073)
+#1067 := (iff #484 #1066)
+#1068 := [rewrite]: #1067
+#1075 := [monotonicity #1068]: #1074
+#1071 := (iff #129 #1070)
+#1072 := [rewrite]: #1071
+#1078 := [monotonicity #1072 #1075]: #1077
+#1083 := [trans #1078 #1081]: #1082
+#1086 := [quant-intro #1083]: #1085
+#1064 := (iff #499 #1063)
+#1061 := (iff #128 #1060)
+#1062 := [monotonicity #1059]: #1061
+#1065 := [monotonicity #1062]: #1064
+#1089 := [monotonicity #1065 #1086]: #1088
+#1092 := [quant-intro #1089]: #1091
+#1052 := (iff #481 #1051)
+#1049 := (iff #476 #1048)
+#1042 := (iff #469 #1043)
+#1044 := [rewrite]: #1042
+#1040 := (iff #475 #1039)
+#1037 := (iff #463 #1034)
+#1031 := (and #970 #114)
+#1035 := (iff #1031 #1034)
+#1036 := [rewrite]: #1035
+#1032 := (iff #463 #1031)
+#1033 := [monotonicity #972]: #1032
+#1038 := [trans #1033 #1036]: #1037
+#1041 := [monotonicity #1038]: #1040
+#1050 := [monotonicity #1041 #1044]: #1049
+#1053 := [quant-intro #1050]: #1052
+#1111 := [monotonicity #1053 #1092]: #1110
+#1099 := (iff #460 #1098)
+#1096 := (iff #457 #1095)
+#1093 := (iff #119 #1069)
+#1094 := [rewrite]: #1093
+#1097 := [monotonicity #1094]: #1096
+#1100 := [quant-intro #1097]: #1099
+#1114 := [monotonicity #1100 #1111]: #1113
+#1105 := (iff #112 #1104)
+#1102 := (iff #111 #1101)
#1103 := [rewrite]: #1102
-#1106 := (iff #284 #1099)
-#1104 := (iff #283 #1094)
-#1100 := (iff #283 #1099)
-#1097 := (iff #282 #1094)
-#1090 := (implies #410 #1085)
-#1095 := (iff #1090 #1094)
-#1096 := [rewrite]: #1095
-#1091 := (iff #282 #1090)
-#1088 := (iff #281 #1085)
-#1081 := (implies #424 #1076)
-#1086 := (iff #1081 #1085)
-#1087 := [rewrite]: #1086
-#1082 := (iff #281 #1081)
-#1079 := (iff #280 #1076)
-#1072 := (implies #76 #1062)
-#1077 := (iff #1072 #1076)
-#1078 := [rewrite]: #1077
-#1073 := (iff #280 #1072)
-#1070 := (iff #279 #1062)
-#1065 := (implies true #1062)
-#1068 := (iff #1065 #1062)
-#1069 := [rewrite]: #1068
-#1066 := (iff #279 #1065)
-#1063 := (iff #278 #1062)
-#1060 := (iff #277 #1057)
-#1053 := (implies #427 #1050)
-#1058 := (iff #1053 #1057)
-#1059 := [rewrite]: #1058
-#1054 := (iff #277 #1053)
-#1051 := (iff #276 #1050)
-#1048 := (iff #275 #1045)
-#1041 := (implies #80 #1038)
-#1046 := (iff #1041 #1045)
-#1047 := [rewrite]: #1046
-#1042 := (iff #275 #1041)
-#1039 := (iff #274 #1038)
-#1036 := (iff #273 #1033)
-#1029 := (implies #434 #1026)
-#1034 := (iff #1029 #1033)
-#1035 := [rewrite]: #1034
-#1030 := (iff #273 #1029)
-#1027 := (iff #272 #1026)
-#1024 := (iff #271 #1021)
-#1017 := (implies #441 #1014)
-#1022 := (iff #1017 #1021)
-#1023 := [rewrite]: #1022
-#1018 := (iff #271 #1017)
-#1015 := (iff #270 #1014)
-#1012 := (iff #269 #1009)
-#1005 := (implies #448 #989)
-#1010 := (iff #1005 #1009)
-#1011 := [rewrite]: #1010
-#1006 := (iff #269 #1005)
-#1003 := (iff #268 #989)
-#994 := (implies true #989)
-#997 := (iff #994 #989)
+#1106 := [quant-intro #1103]: #1105
+#1117 := [monotonicity #1106 #1114]: #1116
+#1107 := (iff #453 #109)
+#1108 := [rewrite]: #1107
+#1120 := [monotonicity #1108 #1117]: #1119
+#1029 := (iff #450 #1028)
+#1026 := (iff #447 #1025)
+#1023 := (iff #104 #1022)
+#1020 := (iff #103 #1017)
+#1011 := (and #74 #1005)
+#1014 := (and #1008 #1011)
+#1018 := (iff #1014 #1017)
+#1019 := [rewrite]: #1018
+#1015 := (iff #103 #1014)
+#1012 := (iff #102 #1011)
+#1006 := (iff #101 #1005)
+#1007 := [rewrite]: #1006
+#1013 := [monotonicity #1007]: #1012
+#1009 := (iff #100 #1008)
+#1010 := [rewrite]: #1009
+#1016 := [monotonicity #1010 #1013]: #1015
+#1021 := [trans #1016 #1019]: #1020
+#1024 := [quant-intro #1021]: #1023
+#1003 := (iff #446 #1002)
+#1000 := (iff #99 #999)
+#997 := (iff #98 #996)
#998 := [rewrite]: #997
-#1001 := (iff #268 #994)
-#999 := (iff #267 #989)
-#995 := (iff #267 #994)
-#992 := (iff #266 #989)
-#985 := (implies #451 #980)
-#990 := (iff #985 #989)
-#991 := [rewrite]: #990
-#986 := (iff #266 #985)
-#983 := (iff #265 #980)
-#976 := (implies #109 #971)
-#981 := (iff #976 #980)
-#982 := [rewrite]: #981
-#977 := (iff #265 #976)
-#974 := (iff #264 #971)
-#967 := (implies #458 #962)
-#972 := (iff #967 #971)
-#973 := [rewrite]: #972
-#968 := (iff #264 #967)
-#965 := (iff #263 #962)
-#958 := (implies #479 #953)
-#963 := (iff #958 #962)
-#964 := [rewrite]: #963
-#959 := (iff #263 #958)
-#956 := (iff #262 #953)
-#949 := (implies #503 #939)
-#954 := (iff #949 #953)
-#955 := [rewrite]: #954
-#950 := (iff #262 #949)
-#947 := (iff #261 #939)
-#942 := (implies true #939)
-#945 := (iff #942 #939)
-#946 := [rewrite]: #945
-#943 := (iff #261 #942)
-#940 := (iff #260 #939)
-#937 := (iff #259 #923)
-#928 := (implies true #923)
-#931 := (iff #928 #923)
+#1001 := [monotonicity #998]: #1000
+#1004 := [monotonicity #1001]: #1003
+#1027 := [monotonicity #1004 #1024]: #1026
+#1030 := [quant-intro #1027]: #1029
+#1123 := [monotonicity #1030 #1120]: #1122
+#1128 := [trans #1123 #1126]: #1127
+#1131 := [monotonicity #1128]: #1130
+#1455 := [monotonicity #1131 #1452]: #1454
+#1458 := [monotonicity #1030 #1455]: #1457
+#991 := (iff #862 #990)
+#988 := (iff #443 #987)
+#985 := (iff #440 #984)
+#979 := (iff #95 #980)
+#981 := [rewrite]: #979
+#977 := (iff #439 #976)
+#974 := (iff #93 #973)
+#975 := [monotonicity #972]: #974
+#978 := [monotonicity #975]: #977
+#986 := [monotonicity #978 #981]: #985
+#989 := [quant-intro #986]: #988
+#992 := [monotonicity #989]: #991
+#1461 := [monotonicity #992 #1458]: #1460
+#1464 := [monotonicity #989 #1461]: #1463
+#965 := (iff #874 #964)
+#962 := (iff #436 #961)
+#959 := (iff #433 #958)
+#953 := (iff #87 #954)
+#957 := [rewrite]: #953
+#960 := [monotonicity #957]: #959
+#963 := [quant-intro #960]: #962
+#966 := [monotonicity #963]: #965
+#1467 := [monotonicity #966 #1464]: #1466
+#1470 := [monotonicity #963 #1467]: #1469
+#950 := (iff #886 #949)
+#947 := (iff #83 #946)
+#943 := (iff #82 #945)
+#944 := [rewrite]: #943
+#948 := [quant-intro #944]: #947
+#951 := [monotonicity #948]: #950
+#1473 := [monotonicity #951 #1470]: #1472
+#1476 := [monotonicity #948 #1473]: #1475
+#941 := (iff #898 #940)
+#938 := (iff #429 #81)
+#939 := [rewrite]: #938
+#942 := [monotonicity #939]: #941
+#1479 := [monotonicity #942 #1476]: #1478
+#1482 := [monotonicity #939 #1479]: #1481
+#936 := (iff #910 #935)
+#933 := (iff #419 #930)
+#927 := (and #924 #414)
+#931 := (iff #927 #930)
#932 := [rewrite]: #931
-#935 := (iff #259 #928)
-#933 := (iff #258 #923)
-#929 := (iff #258 #928)
-#926 := (iff #257 #923)
-#920 := (implies #134 #915)
-#924 := (iff #920 #923)
-#925 := [rewrite]: #924
-#921 := (iff #257 #920)
-#918 := (iff #256 #915)
-#912 := (implies #181 #907)
-#916 := (iff #912 #915)
-#917 := [rewrite]: #916
-#913 := (iff #256 #912)
-#910 := (iff #255 #907)
-#903 := (implies #183 #898)
-#908 := (iff #903 #907)
-#909 := [rewrite]: #908
-#904 := (iff #255 #903)
-#901 := (iff #254 #898)
-#894 := (implies #671 #889)
-#899 := (iff #894 #898)
-#900 := [rewrite]: #899
-#895 := (iff #254 #894)
-#892 := (iff #253 #889)
-#885 := (implies #189 #880)
-#890 := (iff #885 #889)
-#891 := [rewrite]: #890
-#886 := (iff #253 #885)
-#883 := (iff #252 #880)
-#876 := (implies #685 #871)
-#881 := (iff #876 #880)
-#882 := [rewrite]: #881
-#877 := (iff #252 #876)
-#874 := (iff #251 #871)
-#867 := (implies #699 #864)
-#872 := (iff #867 #871)
-#873 := [rewrite]: #872
-#868 := (iff #251 #867)
-#865 := (iff #250 #864)
-#862 := (iff #249 #859)
-#855 := (implies #205 #852)
-#860 := (iff #855 #859)
-#861 := [rewrite]: #860
-#856 := (iff #249 #855)
-#853 := (iff #248 #852)
-#850 := (iff #247 #847)
-#843 := (implies #710 #829)
-#848 := (iff #843 #847)
-#849 := [rewrite]: #848
-#844 := (iff #247 #843)
-#841 := (iff #246 #829)
-#832 := (implies true #829)
-#835 := (iff #832 #829)
-#836 := [rewrite]: #835
-#839 := (iff #246 #832)
-#837 := (iff #245 #829)
-#833 := (iff #245 #832)
-#830 := (iff #244 #829)
-#827 := (iff #243 #824)
-#820 := (implies #713 #817)
-#825 := (iff #820 #824)
-#826 := [rewrite]: #825
-#821 := (iff #243 #820)
-#818 := (iff #242 #817)
-#815 := (iff #241 #812)
-#808 := (implies #212 #805)
-#813 := (iff #808 #812)
-#814 := [rewrite]: #813
-#809 := (iff #241 #808)
-#806 := (iff #240 #805)
-#803 := (iff #239 #800)
-#796 := (implies #720 #793)
-#801 := (iff #796 #800)
-#802 := [rewrite]: #801
-#797 := (iff #239 #796)
-#794 := (iff #238 #793)
-#791 := (iff #237 #788)
-#784 := (implies #741 #765)
-#789 := (iff #784 #788)
-#790 := [rewrite]: #789
-#785 := (iff #237 #784)
-#782 := (iff #236 #765)
-#777 := (and #765 true)
-#780 := (iff #777 #765)
+#928 := (iff #419 #927)
+#925 := (iff #394 #924)
+#922 := (iff #389 #921)
+#919 := (iff #383 #67)
+#920 := [rewrite]: #919
+#923 := [monotonicity #920]: #922
+#926 := [quant-intro #923]: #925
+#929 := [monotonicity #926]: #928
+#934 := [trans #929 #932]: #933
+#937 := [monotonicity #934]: #936
+#1485 := [monotonicity #937 #1482]: #1484
+#1488 := [monotonicity #1485]: #1487
+#917 := (iff #269 #916)
+#914 := (iff #268 #911)
+#907 := (implies #419 #904)
+#912 := (iff #907 #911)
+#913 := [rewrite]: #912
+#908 := (iff #268 #907)
+#905 := (iff #267 #904)
+#902 := (iff #266 #899)
+#895 := (implies #429 #892)
+#900 := (iff #895 #899)
+#901 := [rewrite]: #900
+#896 := (iff #266 #895)
+#893 := (iff #265 #892)
+#890 := (iff #264 #887)
+#883 := (implies #83 #880)
+#888 := (iff #883 #887)
+#889 := [rewrite]: #888
+#884 := (iff #264 #883)
+#881 := (iff #263 #880)
+#878 := (iff #262 #875)
+#871 := (implies #436 #868)
+#876 := (iff #871 #875)
+#877 := [rewrite]: #876
+#872 := (iff #262 #871)
+#869 := (iff #261 #868)
+#866 := (iff #260 #863)
+#859 := (implies #443 #856)
+#864 := (iff #859 #863)
+#865 := [rewrite]: #864
+#860 := (iff #260 #859)
+#857 := (iff #259 #856)
+#854 := (iff #258 #851)
+#847 := (implies #527 #844)
+#852 := (iff #847 #851)
+#853 := [rewrite]: #852
+#848 := (iff #258 #847)
+#845 := (iff #257 #844)
+#842 := (iff #256 #839)
+#835 := (implies #687 #832)
+#840 := (iff #835 #839)
+#841 := [rewrite]: #840
+#836 := (iff #256 #835)
+#833 := (iff #255 #832)
+#830 := (iff #254 #827)
+#823 := (implies #215 #820)
+#828 := (iff #823 #827)
+#829 := [rewrite]: #828
+#824 := (iff #254 #823)
+#821 := (iff #253 #820)
+#818 := (iff #252 #815)
+#811 := (implies #705 #808)
+#816 := (iff #811 #815)
+#817 := [rewrite]: #816
+#812 := (iff #252 #811)
+#809 := (iff #251 #808)
+#806 := (iff #250 #803)
+#799 := (implies #708 #796)
+#804 := (iff #799 #803)
+#805 := [rewrite]: #804
+#800 := (iff #250 #799)
+#797 := (iff #249 #796)
+#794 := (iff #248 #791)
+#787 := (implies #222 #784)
+#792 := (iff #787 #791)
+#793 := [rewrite]: #792
+#788 := (iff #248 #787)
+#785 := (iff #247 #784)
+#782 := (iff #246 #779)
+#775 := (implies #715 #772)
+#780 := (iff #775 #779)
#781 := [rewrite]: #780
-#778 := (iff #236 #777)
-#775 := (iff #235 true)
-#770 := (implies #765 true)
-#773 := (iff #770 true)
-#774 := [rewrite]: #773
-#771 := (iff #235 #770)
-#768 := (iff #234 true)
+#776 := (iff #246 #775)
+#773 := (iff #245 #772)
+#770 := (iff #244 #767)
+#763 := (implies #736 #760)
+#768 := (iff #763 #767)
#769 := [rewrite]: #768
-#766 := (iff #233 #765)
-#763 := (iff #232 #760)
-#756 := (implies #226 #753)
-#761 := (iff #756 #760)
-#762 := [rewrite]: #761
-#757 := (iff #232 #756)
-#754 := (iff #231 #753)
-#751 := (iff #230 #750)
-#748 := (iff #229 #747)
-#745 := (iff #228 #744)
-#727 := (= #221 #726)
-#728 := [rewrite]: #727
-#746 := [monotonicity #728]: #745
-#749 := [monotonicity #746]: #748
-#752 := [monotonicity #749]: #751
-#755 := [quant-intro #752]: #754
-#758 := [monotonicity #755]: #757
-#764 := [trans #758 #762]: #763
-#767 := [quant-intro #764]: #766
-#772 := [monotonicity #767 #769]: #771
-#776 := [trans #772 #774]: #775
-#779 := [monotonicity #767 #776]: #778
-#783 := [trans #779 #781]: #782
-#742 := (iff #224 #741)
-#739 := (iff #223 #736)
-#732 := (implies #723 #729)
-#737 := (iff #732 #736)
-#738 := [rewrite]: #737
-#733 := (iff #223 #732)
-#730 := (iff #222 #729)
-#731 := [monotonicity #728]: #730
-#724 := (iff #220 #723)
-#725 := [rewrite]: #724
-#734 := [monotonicity #725 #731]: #733
-#740 := [trans #734 #738]: #739
-#743 := [quant-intro #740]: #742
-#786 := [monotonicity #743 #783]: #785
-#792 := [trans #786 #790]: #791
-#795 := [monotonicity #743 #792]: #794
-#721 := (iff #219 #720)
-#718 := (iff #218 #717)
-#719 := [rewrite]: #718
-#722 := [quant-intro #719]: #721
-#798 := [monotonicity #722 #795]: #797
-#804 := [trans #798 #802]: #803
-#807 := [monotonicity #722 #804]: #806
-#810 := [monotonicity #807]: #809
-#816 := [trans #810 #814]: #815
-#819 := [monotonicity #816]: #818
-#714 := (iff #210 #713)
-#715 := [rewrite]: #714
-#822 := [monotonicity #715 #819]: #821
-#828 := [trans #822 #826]: #827
-#831 := [monotonicity #715 #828]: #830
+#764 := (iff #244 #763)
+#761 := (iff #243 #760)
+#758 := (iff #242 #755)
+#751 := (implies #236 #748)
+#756 := (iff #751 #755)
+#757 := [rewrite]: #756
+#752 := (iff #242 #751)
+#749 := (iff #241 #748)
+#746 := (iff #240 #745)
+#743 := (iff #239 #742)
+#740 := (iff #238 #739)
+#722 := (= #231 #721)
+#723 := [rewrite]: #722
+#741 := [monotonicity #723]: #740
+#744 := [monotonicity #741]: #743
+#747 := [monotonicity #744]: #746
+#750 := [quant-intro #747]: #749
+#753 := [monotonicity #750]: #752
+#759 := [trans #753 #757]: #758
+#762 := [quant-intro #759]: #761
+#737 := (iff #234 #736)
+#734 := (iff #233 #731)
+#727 := (implies #718 #724)
+#732 := (iff #727 #731)
+#733 := [rewrite]: #732
+#728 := (iff #233 #727)
+#725 := (iff #232 #724)
+#726 := [monotonicity #723]: #725
+#719 := (iff #230 #718)
+#720 := [rewrite]: #719
+#729 := [monotonicity #720 #726]: #728
+#735 := [trans #729 #733]: #734
+#738 := [quant-intro #735]: #737
+#765 := [monotonicity #738 #762]: #764
+#771 := [trans #765 #769]: #770
+#774 := [monotonicity #738 #771]: #773
+#716 := (iff #229 #715)
+#713 := (iff #228 #712)
+#714 := [rewrite]: #713
+#717 := [quant-intro #714]: #716
+#777 := [monotonicity #717 #774]: #776
+#783 := [trans #777 #781]: #782
+#786 := [monotonicity #717 #783]: #785
+#789 := [monotonicity #786]: #788
+#795 := [trans #789 #793]: #794
+#798 := [monotonicity #795]: #797
+#709 := (iff #220 #708)
+#710 := [rewrite]: #709
+#801 := [monotonicity #710 #798]: #800
+#807 := [trans #801 #805]: #806
+#810 := [monotonicity #710 #807]: #809
+#706 := (iff #218 #705)
+#703 := (iff #217 #700)
+#697 := (implies #216 #658)
+#701 := (iff #697 #700)
+#702 := [rewrite]: #701
+#698 := (iff #217 #697)
+#659 := (iff #204 #658)
+#660 := [rewrite]: #659
+#699 := [monotonicity #660]: #698
+#704 := [trans #699 #702]: #703
+#707 := [quant-intro #704]: #706
+#813 := [monotonicity #707 #810]: #812
+#819 := [trans #813 #817]: #818
+#822 := [monotonicity #707 #819]: #821
+#825 := [monotonicity #822]: #824
+#831 := [trans #825 #829]: #830
#834 := [monotonicity #831]: #833
-#838 := [trans #834 #836]: #837
-#840 := [monotonicity #838]: #839
-#842 := [trans #840 #836]: #841
-#711 := (iff #208 #710)
-#708 := (iff #207 #705)
-#702 := (implies #206 #688)
-#706 := (iff #702 #705)
-#707 := [rewrite]: #706
-#703 := (iff #207 #702)
-#689 := (iff #201 #688)
-#690 := [rewrite]: #689
-#704 := [monotonicity #690]: #703
-#709 := [trans #704 #707]: #708
-#712 := [quant-intro #709]: #711
-#845 := [monotonicity #712 #842]: #844
-#851 := [trans #845 #849]: #850
-#854 := [monotonicity #712 #851]: #853
-#857 := [monotonicity #854]: #856
-#863 := [trans #857 #861]: #862
-#866 := [monotonicity #863]: #865
-#700 := (iff #203 #699)
-#697 := (iff #202 #694)
-#691 := (implies #200 #688)
-#695 := (iff #691 #694)
-#696 := [rewrite]: #695
-#692 := (iff #202 #691)
-#693 := [monotonicity #690]: #692
-#698 := [trans #693 #696]: #697
-#701 := [quant-intro #698]: #700
-#869 := [monotonicity #701 #866]: #868
-#875 := [trans #869 #873]: #874
-#686 := (iff #199 #685)
-#683 := (iff #198 #680)
-#677 := (implies #195 #674)
-#681 := (iff #677 #680)
-#682 := [rewrite]: #681
-#678 := (iff #198 #677)
-#675 := (iff #197 #674)
-#676 := [rewrite]: #675
-#679 := [monotonicity #676]: #678
-#684 := [trans #679 #682]: #683
-#687 := [quant-intro #684]: #686
-#878 := [monotonicity #687 #875]: #877
-#884 := [trans #878 #882]: #883
-#887 := [monotonicity #884]: #886
-#893 := [trans #887 #891]: #892
-#672 := (iff #186 #671)
-#669 := (iff #185 #668)
-#670 := [rewrite]: #669
-#673 := [quant-intro #670]: #672
-#896 := [monotonicity #673 #893]: #895
-#902 := [trans #896 #900]: #901
-#905 := [monotonicity #902]: #904
-#911 := [trans #905 #909]: #910
-#914 := [monotonicity #911]: #913
-#919 := [trans #914 #917]: #918
-#922 := [monotonicity #919]: #921
-#927 := [trans #922 #925]: #926
-#930 := [monotonicity #927]: #929
-#934 := [trans #930 #932]: #933
-#936 := [monotonicity #934]: #935
-#938 := [trans #936 #932]: #937
-#666 := (iff #178 #652)
-#657 := (implies true #652)
-#660 := (iff #657 #652)
-#661 := [rewrite]: #660
-#664 := (iff #178 #657)
-#662 := (iff #177 #652)
-#658 := (iff #177 #657)
-#655 := (iff #176 #652)
-#649 := (implies #135 #633)
-#653 := (iff #649 #652)
-#654 := [rewrite]: #653
-#650 := (iff #176 #649)
-#647 := (iff #175 #633)
-#638 := (implies true #633)
-#641 := (iff #638 #633)
-#642 := [rewrite]: #641
-#645 := (iff #175 #638)
-#643 := (iff #174 #633)
-#639 := (iff #174 #638)
-#636 := (iff #173 #633)
-#629 := (implies #506 #624)
+#695 := (iff #213 #687)
+#690 := (and true #687)
+#693 := (iff #690 #687)
+#694 := [rewrite]: #693
+#691 := (iff #213 #690)
+#688 := (iff #212 #687)
+#685 := (iff #211 #684)
+#682 := (iff #210 #681)
+#679 := (iff #209 #678)
+#676 := (iff #208 #675)
+#673 := (iff #207 #672)
+#670 := (iff #206 #669)
+#667 := (iff #205 #664)
+#661 := (implies #203 #658)
+#665 := (iff #661 #664)
+#666 := [rewrite]: #665
+#662 := (iff #205 #661)
+#663 := [monotonicity #660]: #662
+#668 := [trans #663 #666]: #667
+#671 := [quant-intro #668]: #670
+#656 := (iff #202 #655)
+#653 := (iff #201 #650)
+#647 := (implies #198 #644)
+#651 := (iff #647 #650)
+#652 := [rewrite]: #651
+#648 := (iff #201 #647)
+#645 := (iff #200 #644)
+#646 := [rewrite]: #645
+#649 := [monotonicity #646]: #648
+#654 := [trans #649 #652]: #653
+#657 := [quant-intro #654]: #656
+#674 := [monotonicity #657 #671]: #673
+#677 := [monotonicity #674]: #676
+#642 := (iff #189 #641)
+#639 := (iff #188 #638)
+#640 := [rewrite]: #639
+#643 := [quant-intro #640]: #642
+#680 := [monotonicity #643 #677]: #679
+#683 := [monotonicity #680]: #682
+#686 := [monotonicity #683]: #685
+#689 := [monotonicity #686]: #688
+#692 := [monotonicity #689]: #691
+#696 := [trans #692 #694]: #695
+#837 := [monotonicity #696 #834]: #836
+#843 := [trans #837 #841]: #842
+#636 := (iff #181 #633)
+#629 := (implies #550 #626)
#634 := (iff #629 #633)
#635 := [rewrite]: #634
-#630 := (iff #173 #629)
-#627 := (iff #172 #624)
-#620 := (implies #140 #615)
-#625 := (iff #620 #624)
-#626 := [rewrite]: #625
-#621 := (iff #172 #620)
-#618 := (iff #171 #615)
-#611 := (implies #509 #606)
-#616 := (iff #611 #615)
-#617 := [rewrite]: #616
-#612 := (iff #171 #611)
-#609 := (iff #170 #606)
-#602 := (implies #145 #592)
-#607 := (iff #602 #606)
-#608 := [rewrite]: #607
-#603 := (iff #170 #602)
-#600 := (iff #169 #592)
-#595 := (implies true #592)
-#598 := (iff #595 #592)
-#599 := [rewrite]: #598
-#596 := (iff #169 #595)
-#593 := (iff #168 #592)
-#590 := (iff #167 #587)
-#583 := (implies #533 #580)
-#588 := (iff #583 #587)
-#589 := [rewrite]: #588
-#584 := (iff #167 #583)
-#581 := (iff #166 #580)
-#578 := (iff #165 #575)
-#571 := (implies #551 #554)
-#576 := (iff #571 #575)
-#577 := [rewrite]: #576
-#572 := (iff #165 #571)
-#569 := (iff #164 #554)
-#564 := (and #554 true)
-#567 := (iff #564 #554)
-#568 := [rewrite]: #567
-#565 := (iff #164 #564)
-#562 := (iff #163 true)
-#557 := (implies #554 true)
-#560 := (iff #557 true)
-#561 := [rewrite]: #560
-#558 := (iff #163 #557)
-#555 := (iff #162 #554)
-#556 := [rewrite]: #555
-#559 := [monotonicity #556]: #558
-#563 := [trans #559 #561]: #562
-#566 := [monotonicity #556 #563]: #565
-#570 := [trans #566 #568]: #569
-#552 := (iff #160 #551)
-#549 := (iff #159 #546)
-#542 := (implies #536 #539)
-#547 := (iff #542 #546)
-#548 := [rewrite]: #547
-#543 := (iff #159 #542)
-#540 := (iff #158 #539)
-#513 := (= #151 #512)
-#514 := [rewrite]: #513
-#541 := [monotonicity #514]: #540
-#537 := (iff #157 #536)
-#538 := [rewrite]: #537
-#544 := [monotonicity #538 #541]: #543
-#550 := [trans #544 #548]: #549
-#553 := [quant-intro #550]: #552
-#573 := [monotonicity #553 #570]: #572
-#579 := [trans #573 #577]: #578
-#582 := [monotonicity #553 #579]: #581
-#534 := (iff #156 #533)
-#531 := (iff #155 #528)
-#524 := (implies #148 #521)
-#529 := (iff #524 #528)
-#530 := [rewrite]: #529
-#525 := (iff #155 #524)
-#522 := (iff #154 #521)
-#519 := (iff #153 #518)
-#516 := (iff #152 #515)
-#517 := [monotonicity #514]: #516
-#520 := [monotonicity #517]: #519
-#523 := [quant-intro #520]: #522
-#526 := [monotonicity #523]: #525
-#532 := [trans #526 #530]: #531
-#535 := [quant-intro #532]: #534
-#585 := [monotonicity #535 #582]: #584
-#591 := [trans #585 #589]: #590
-#594 := [monotonicity #535 #591]: #593
-#597 := [monotonicity #594]: #596
-#601 := [trans #597 #599]: #600
-#604 := [monotonicity #601]: #603
-#610 := [trans #604 #608]: #609
-#510 := (iff #142 #509)
-#511 := [rewrite]: #510
-#613 := [monotonicity #511 #610]: #612
-#619 := [trans #613 #617]: #618
-#622 := [monotonicity #619]: #621
-#628 := [trans #622 #626]: #627
-#507 := (iff #137 #506)
-#508 := [rewrite]: #507
-#631 := [monotonicity #508 #628]: #630
+#630 := (iff #181 #629)
+#627 := (iff #180 #626)
+#624 := (iff #179 #621)
+#617 := (implies #581 #614)
+#622 := (iff #617 #621)
+#623 := [rewrite]: #622
+#618 := (iff #179 #617)
+#615 := (iff #178 #614)
+#612 := (iff #177 #609)
+#605 := (implies #599 #602)
+#610 := (iff #605 #609)
+#611 := [rewrite]: #610
+#606 := (iff #177 #605)
+#603 := (iff #176 #602)
+#604 := [rewrite]: #603
+#600 := (iff #174 #599)
+#597 := (iff #173 #594)
+#590 := (implies #584 #587)
+#595 := (iff #590 #594)
+#596 := [rewrite]: #595
+#591 := (iff #173 #590)
+#588 := (iff #172 #587)
+#561 := (= #165 #560)
+#562 := [rewrite]: #561
+#589 := [monotonicity #562]: #588
+#585 := (iff #171 #584)
+#586 := [rewrite]: #585
+#592 := [monotonicity #586 #589]: #591
+#598 := [trans #592 #596]: #597
+#601 := [quant-intro #598]: #600
+#607 := [monotonicity #601 #604]: #606
+#613 := [trans #607 #611]: #612
+#616 := [monotonicity #601 #613]: #615
+#582 := (iff #170 #581)
+#579 := (iff #169 #576)
+#572 := (implies #162 #569)
+#577 := (iff #572 #576)
+#578 := [rewrite]: #577
+#573 := (iff #169 #572)
+#570 := (iff #168 #569)
+#567 := (iff #167 #566)
+#564 := (iff #166 #563)
+#565 := [monotonicity #562]: #564
+#568 := [monotonicity #565]: #567
+#571 := [quant-intro #568]: #570
+#574 := [monotonicity #571]: #573
+#580 := [trans #574 #578]: #579
+#583 := [quant-intro #580]: #582
+#619 := [monotonicity #583 #616]: #618
+#625 := [trans #619 #623]: #624
+#628 := [monotonicity #583 #625]: #627
+#558 := (iff #159 #550)
+#553 := (and true #550)
+#556 := (iff #553 #550)
+#557 := [rewrite]: #556
+#554 := (iff #159 #553)
+#551 := (iff #158 #550)
+#548 := (iff #157 #547)
+#545 := (iff #156 #544)
+#542 := (iff #155 #539)
+#536 := (and #533 #154)
+#540 := (iff #536 #539)
+#541 := [rewrite]: #540
+#537 := (iff #155 #536)
+#534 := (iff #151 #533)
+#535 := [rewrite]: #534
+#538 := [monotonicity #535]: #537
+#543 := [trans #538 #541]: #542
+#546 := [monotonicity #543]: #545
+#531 := (iff #146 #530)
+#532 := [rewrite]: #531
+#549 := [monotonicity #532 #546]: #548
+#552 := [monotonicity #549]: #551
+#555 := [monotonicity #552]: #554
+#559 := [trans #555 #557]: #558
+#631 := [monotonicity #559 #628]: #630
#637 := [trans #631 #635]: #636
-#640 := [monotonicity #637]: #639
-#644 := [trans #640 #642]: #643
-#646 := [monotonicity #644]: #645
-#648 := [trans #646 #642]: #647
-#651 := [monotonicity #648]: #650
-#656 := [trans #651 #654]: #655
-#659 := [monotonicity #656]: #658
-#663 := [trans #659 #661]: #662
-#665 := [monotonicity #663]: #664
-#667 := [trans #665 #661]: #666
-#941 := [monotonicity #667 #938]: #940
-#944 := [monotonicity #941]: #943
-#948 := [trans #944 #946]: #947
-#504 := (iff #132 #503)
-#501 := (iff #131 #498)
-#494 := (implies #125 #491)
-#499 := (iff #494 #498)
-#500 := [rewrite]: #499
-#495 := (iff #131 #494)
-#492 := (iff #130 #491)
-#489 := (iff #129 #488)
-#486 := (iff #128 #485)
-#483 := (iff #127 #482)
-#465 := (= #120 #464)
-#466 := [rewrite]: #465
-#484 := [monotonicity #466]: #483
-#487 := [monotonicity #484]: #486
-#490 := [monotonicity #487]: #489
-#493 := [quant-intro #490]: #492
-#496 := [monotonicity #493]: #495
-#502 := [trans #496 #500]: #501
-#505 := [quant-intro #502]: #504
-#951 := [monotonicity #505 #948]: #950
-#957 := [trans #951 #955]: #956
-#480 := (iff #123 #479)
-#477 := (iff #122 #474)
-#470 := (implies #461 #467)
-#475 := (iff #470 #474)
-#476 := [rewrite]: #475
-#471 := (iff #122 #470)
-#468 := (iff #121 #467)
-#469 := [monotonicity #466]: #468
-#462 := (iff #119 #461)
-#463 := [rewrite]: #462
-#472 := [monotonicity #463 #469]: #471
-#478 := [trans #472 #476]: #477
-#481 := [quant-intro #478]: #480
-#960 := [monotonicity #481 #957]: #959
-#966 := [trans #960 #964]: #965
-#459 := (iff #118 #458)
-#456 := (iff #117 #455)
-#457 := [rewrite]: #456
-#460 := [quant-intro #457]: #459
-#969 := [monotonicity #460 #966]: #968
-#975 := [trans #969 #973]: #974
-#978 := [monotonicity #975]: #977
-#984 := [trans #978 #982]: #983
-#452 := (iff #106 #451)
-#453 := [rewrite]: #452
-#987 := [monotonicity #453 #984]: #986
-#993 := [trans #987 #991]: #992
-#996 := [monotonicity #993]: #995
-#1000 := [trans #996 #998]: #999
-#1002 := [monotonicity #1000]: #1001
-#1004 := [trans #1002 #998]: #1003
-#449 := (iff #103 #448)
-#446 := (iff #102 #445)
-#447 := [rewrite]: #446
-#450 := [quant-intro #447]: #449
-#1007 := [monotonicity #450 #1004]: #1006
-#1013 := [trans #1007 #1011]: #1012
-#1016 := [monotonicity #450 #1013]: #1015
-#442 := (iff #94 #441)
-#439 := (iff #93 #438)
-#440 := [rewrite]: #439
-#443 := [quant-intro #440]: #442
-#1019 := [monotonicity #443 #1016]: #1018
-#1025 := [trans #1019 #1023]: #1024
-#1028 := [monotonicity #443 #1025]: #1027
-#435 := (iff #86 #434)
-#432 := (iff #85 #431)
-#433 := [rewrite]: #432
-#436 := [quant-intro #433]: #435
-#1031 := [monotonicity #436 #1028]: #1030
-#1037 := [trans #1031 #1035]: #1036
-#1040 := [monotonicity #436 #1037]: #1039
-#1043 := [monotonicity #1040]: #1042
-#1049 := [trans #1043 #1047]: #1048
-#1052 := [monotonicity #1049]: #1051
-#428 := (iff #78 #427)
-#429 := [rewrite]: #428
-#1055 := [monotonicity #429 #1052]: #1054
-#1061 := [trans #1055 #1059]: #1060
-#1064 := [monotonicity #429 #1061]: #1063
-#1067 := [monotonicity #1064]: #1066
-#1071 := [trans #1067 #1069]: #1070
-#1074 := [monotonicity #1071]: #1073
-#1080 := [trans #1074 #1078]: #1079
-#425 := (iff #73 #424)
-#422 := (iff #72 #419)
-#416 := (implies #70 #413)
-#420 := (iff #416 #419)
-#421 := [rewrite]: #420
-#417 := (iff #72 #416)
-#414 := (iff #71 #413)
-#415 := [rewrite]: #414
-#418 := [monotonicity #415]: #417
-#423 := [trans #418 #421]: #422
-#426 := [quant-intro #423]: #425
-#1083 := [monotonicity #426 #1080]: #1082
-#1089 := [trans #1083 #1087]: #1088
-#411 := (iff #69 #410)
-#408 := (iff #68 #405)
-#402 := (implies #65 #399)
-#406 := (iff #402 #405)
-#407 := [rewrite]: #406
-#403 := (iff #68 #402)
-#400 := (iff #67 #399)
-#401 := [rewrite]: #400
-#404 := [monotonicity #401]: #403
-#409 := [trans #404 #407]: #408
-#412 := [quant-intro #409]: #411
-#1092 := [monotonicity #412 #1089]: #1091
-#1098 := [trans #1092 #1096]: #1097
-#1101 := [monotonicity #1098]: #1100
-#1105 := [trans #1101 #1103]: #1104
-#1107 := [monotonicity #1105]: #1106
-#1109 := [trans #1107 #1103]: #1108
-#1112 := [monotonicity #1109]: #1111
-#1721 := [trans #1112 #1719]: #1720
-#398 := [asserted]: #285
-#1722 := [mp #398 #1721]: #1717
-#1723 := [not-or-elim #1722]: #76
-#1786 := [mp~ #1723 #1748]: #76
-#4130 := [mp #1786 #4129]: #4125
-#5292 := (not #4125)
-#5293 := (or #5292 #2962)
-#5294 := [quant-inst]: #5293
-#8242 := [unit-resolution #5294 #4130 #8241]: false
-#8245 := [lemma #8242]: #2962
-#3719 := (or #2977 #1847)
-#4054 := [def-axiom]: #3719
-#10109 := [unit-resolution #4054 #8245]: #2977
-#2982 := (not #2977)
-#4471 := (or #2982 #4468)
-#4474 := (not #4471)
-#4146 := (pattern #74 #81)
-#2408 := (not #81)
-#2954 := (or #74 #2408 #1129)
-#4147 := (forall (vars (?x29 T2) (?x30 T2)) (:pat #4146) #2954)
-#4152 := (not #4147)
-#4477 := (or #4152 #4474)
-#4480 := (not #4477)
-decl ?x30!1 :: T2
-#1808 := ?x30!1
-#1812 := (uf_12 ?x30!1)
-#2423 := (* -1::int #1812)
-decl ?x29!2 :: T2
-#1809 := ?x29!2
-#1810 := (uf_12 ?x29!2)
-#2424 := (+ #1810 #2423)
-#2425 := (<= #2424 0::int)
-#1816 := (up_13 ?x30!1)
-#1815 := (up_13 ?x29!2)
-#2058 := (not #1815)
-#2132 := (or #2058 #1816 #2425)
-#8816 := [hypothesis]: #1815
-#5238 := (or #5292 #2058)
-#5267 := [quant-inst]: #5238
-#8817 := [unit-resolution #5267 #4130 #8816]: false
-#8818 := [lemma #8817]: #2058
-#3648 := (or #2132 #1815)
-#3733 := [def-axiom]: #3648
-#10110 := [unit-resolution #3733 #8818]: #2132
-#1948 := (not #2132)
-#4483 := (or #1948 #4480)
-#4486 := (not #4483)
-#4138 := (forall (vars (?x27 T2)) (:pat #4131) #1120)
-#4143 := (not #4138)
-#4489 := (or #4143 #4486)
-#4492 := (not #4489)
-decl ?x27!0 :: T2
-#1793 := ?x27!0
-#1794 := (uf_12 ?x27!0)
-#1795 := (>= #1794 0::int)
-#1796 := (not #1795)
-#4495 := (or #1796 #4492)
-#4498 := (not #4495)
-#4501 := (or #1115 #4498)
-#4504 := (not #4501)
-#4511 := (forall (vars (?x24 T2)) (:pat #4131) #1694)
-#4514 := (iff #1697 #4511)
-#4512 := (iff #1694 #1694)
-#4513 := [refl]: #4512
-#4515 := [quant-intro #4513]: #4514
-#2226 := (~ #1697 #1697)
-#2022 := (~ #1694 #1694)
-#2023 := [refl]: #2022
-#2227 := [nnf-pos #2023]: #2226
-#1727 := [not-or-elim #1722]: #1697
-#2057 := [mp~ #1727 #2227]: #1697
-#4516 := [mp #2057 #4515]: #4511
-#5053 := [hypothesis]: #1115
-#3659 := (not #4511)
-#5075 := (or #3659 #78)
-#4998 := (= uf_11 uf_11)
-#4996 := (not #4998)
-#4988 := (or #4996 #78)
-#5076 := (or #3659 #4988)
-#5078 := (iff #5076 #5075)
-#5069 := (iff #5075 #5075)
-#5103 := [rewrite]: #5069
-#5070 := (iff #4988 #78)
-#5059 := (or false #78)
-#5063 := (iff #5059 #78)
-#5064 := [rewrite]: #5063
-#5062 := (iff #4988 #5059)
-#5012 := (iff #4996 false)
-#8701 := (not true)
-#8736 := (iff #8701 false)
-#8737 := [rewrite]: #8736
-#5010 := (iff #4996 #8701)
-#5008 := (iff #4998 true)
-#5009 := [rewrite]: #5008
-#5011 := [monotonicity #5009]: #5010
-#5061 := [trans #5011 #8737]: #5012
-#5052 := [monotonicity #5061]: #5062
-#5071 := [trans #5052 #5064]: #5070
-#5079 := [monotonicity #5071]: #5078
-#5104 := [trans #5079 #5103]: #5078
-#5077 := [quant-inst]: #5076
-#5105 := [mp #5077 #5104]: #5075
-#5060 := [unit-resolution #5105 #5053 #4516]: false
-#5109 := [lemma #5060]: #78
-#4507 := (or #1115 #4504)
-#3412 := (forall (vars (?x76 T2)) #3401)
-#3419 := (not #3412)
-#3397 := (forall (vars (?x71 T2) (?x72 T2)) #3392)
-#3418 := (not #3397)
-#3420 := (or #2248 #2810 #3418 #3419)
-#3421 := (not #3420)
-#3426 := (or #3375 #3421)
-#3433 := (not #3426)
-#3352 := (forall (vars (?x67 T2) (?x68 T2)) #3347)
-#3432 := (not #3352)
-#3434 := (or #3432 #3433)
-#3435 := (not #3434)
-#3440 := (or #3329 #3435)
-#3446 := (not #3440)
-#3447 := (or #1403 #3446)
-#3448 := (not #3447)
-#3453 := (or #2168 #3448)
-#3459 := (not #3453)
-#3460 := (or #1394 #3459)
-#3461 := (not #3460)
-#3466 := (or #1394 #3461)
-#3472 := (not #3466)
-#3473 := (or #846 #3472)
-#3474 := (not #3473)
-#3479 := (or #2753 #3474)
-#3485 := (not #3479)
-#3486 := (or #1389 #3485)
-#3487 := (not #3486)
-#3492 := (or #2744 #3487)
-#3500 := (not #3492)
-#3306 := (forall (vars (?x59 T2)) #3301)
-#3499 := (not #3306)
-#3288 := (forall (vars (?x60 T2)) #3285)
-#3498 := (not #3288)
-#3501 := (or #180 #888 #1544 #1548 #2104 #2106 #3498 #3499 #3500)
-#3502 := (not #3501)
-#3218 := (forall (vars (?x53 T2) (?x54 T2)) #3213)
-#3224 := (not #3218)
-#3225 := (or #162 #3224)
-#3226 := (not #3225)
-#3253 := (or #3226 #3250)
-#3260 := (not #3253)
-#3196 := (forall (vars (?x49 T2)) #3191)
-#3259 := (not #3196)
-#3261 := (or #3259 #3260)
-#3262 := (not #3261)
-#3159 := (forall (vars (?x50 T2)) #3148)
-#3165 := (not #3159)
-#3166 := (or #1983 #2593 #3165)
-#3167 := (not #3166)
-#3267 := (or #3167 #3262)
-#3274 := (not #3267)
-#3144 := (forall (vars (?x48 T2)) #3133)
-#3273 := (not #3144)
-#3275 := (or #623 #605 #632 #614 #3273 #3274)
-#3276 := (not #3275)
-#3507 := (or #3276 #3502)
-#3517 := (not #3507)
-#3130 := (forall (vars (?x46 T2)) #3125)
-#3516 := (not #3130)
-#3102 := (forall (vars (?x37 T2)) #3097)
-#3515 := (not #3102)
-#3074 := (forall (vars (?x42 T2) (?x43 T2)) #3069)
-#3514 := (not #3074)
-#3051 := (forall (vars (?x44 T2) (?x45 T2)) #3046)
-#3513 := (not #3051)
-#3518 := (or #1636 #1631 #3513 #3514 #3515 #3516 #3517)
-#3519 := (not #3518)
-#3019 := (forall (vars (?x38 T2)) #3008)
-#3025 := (not #3019)
-#3026 := (or #1877 #2455 #3025)
+#846 := [monotonicity #637 #843]: #845
+#528 := (iff #141 #527)
+#525 := (iff #140 #517)
+#520 := (and true #517)
+#523 := (iff #520 #517)
+#524 := [rewrite]: #523
+#521 := (iff #140 #520)
+#518 := (iff #139 #517)
+#515 := (iff #138 #514)
+#512 := (iff #137 #511)
+#509 := (iff #136 #508)
+#506 := (iff #135 #505)
+#503 := (iff #134 #500)
+#496 := (implies #128 #493)
+#501 := (iff #496 #500)
+#502 := [rewrite]: #501
+#497 := (iff #134 #496)
+#494 := (iff #133 #493)
+#491 := (iff #132 #490)
+#488 := (iff #131 #487)
+#485 := (iff #130 #484)
+#467 := (= #123 #466)
+#468 := [rewrite]: #467
+#486 := [monotonicity #468]: #485
+#489 := [monotonicity #486]: #488
+#492 := [monotonicity #489]: #491
+#495 := [quant-intro #492]: #494
+#498 := [monotonicity #495]: #497
+#504 := [trans #498 #502]: #503
+#507 := [quant-intro #504]: #506
+#482 := (iff #126 #481)
+#479 := (iff #125 #476)
+#472 := (implies #463 #469)
+#477 := (iff #472 #476)
+#478 := [rewrite]: #477
+#473 := (iff #125 #472)
+#470 := (iff #124 #469)
+#471 := [monotonicity #468]: #470
+#464 := (iff #122 #463)
+#465 := [rewrite]: #464
+#474 := [monotonicity #465 #471]: #473
+#480 := [trans #474 #478]: #479
+#483 := [quant-intro #480]: #482
+#510 := [monotonicity #483 #507]: #509
+#461 := (iff #121 #460)
+#458 := (iff #120 #457)
+#459 := [rewrite]: #458
+#462 := [quant-intro #459]: #461
+#513 := [monotonicity #462 #510]: #512
+#516 := [monotonicity #513]: #515
+#454 := (iff #109 #453)
+#455 := [rewrite]: #454
+#519 := [monotonicity #455 #516]: #518
+#522 := [monotonicity #519]: #521
+#526 := [trans #522 #524]: #525
+#451 := (iff #106 #450)
+#448 := (iff #105 #447)
+#449 := [rewrite]: #448
+#452 := [quant-intro #449]: #451
+#529 := [monotonicity #452 #526]: #528
+#849 := [monotonicity #529 #846]: #848
+#855 := [trans #849 #853]: #854
+#858 := [monotonicity #452 #855]: #857
+#444 := (iff #97 #443)
+#441 := (iff #96 #440)
+#442 := [rewrite]: #441
+#445 := [quant-intro #442]: #444
+#861 := [monotonicity #445 #858]: #860
+#867 := [trans #861 #865]: #866
+#870 := [monotonicity #445 #867]: #869
+#437 := (iff #89 #436)
+#434 := (iff #88 #433)
+#435 := [rewrite]: #434
+#438 := [quant-intro #435]: #437
+#873 := [monotonicity #438 #870]: #872
+#879 := [trans #873 #877]: #878
+#882 := [monotonicity #438 #879]: #881
+#885 := [monotonicity #882]: #884
+#891 := [trans #885 #889]: #890
+#894 := [monotonicity #891]: #893
+#430 := (iff #81 #429)
+#431 := [rewrite]: #430
+#897 := [monotonicity #431 #894]: #896
+#903 := [trans #897 #901]: #902
+#906 := [monotonicity #431 #903]: #905
+#427 := (iff #79 #419)
+#422 := (and true #419)
+#425 := (iff #422 #419)
+#426 := [rewrite]: #425
+#423 := (iff #79 #422)
+#420 := (iff #78 #419)
+#417 := (iff #77 #414)
+#411 := (and #408 #76)
+#415 := (iff #411 #414)
+#416 := [rewrite]: #415
+#412 := (iff #77 #411)
+#409 := (iff #73 #408)
+#406 := (iff #72 #403)
+#400 := (implies #70 #397)
+#404 := (iff #400 #403)
+#405 := [rewrite]: #404
+#401 := (iff #72 #400)
+#398 := (iff #71 #397)
+#399 := [rewrite]: #398
+#402 := [monotonicity #399]: #401
+#407 := [trans #402 #405]: #406
+#410 := [quant-intro #407]: #409
+#413 := [monotonicity #410]: #412
+#418 := [trans #413 #416]: #417
+#395 := (iff #69 #394)
+#392 := (iff #68 #389)
+#386 := (implies #65 #383)
+#390 := (iff #386 #389)
+#391 := [rewrite]: #390
+#387 := (iff #68 #386)
+#384 := (iff #67 #383)
+#385 := [rewrite]: #384
+#388 := [monotonicity #385]: #387
+#393 := [trans #388 #391]: #392
+#396 := [quant-intro #393]: #395
+#421 := [monotonicity #396 #418]: #420
+#424 := [monotonicity #421]: #423
+#428 := [trans #424 #426]: #427
+#909 := [monotonicity #428 #906]: #908
+#915 := [trans #909 #913]: #914
+#918 := [monotonicity #915]: #917
+#1490 := [trans #918 #1488]: #1489
+#382 := [asserted]: #269
+#1491 := [mp #382 #1490]: #1486
+#1492 := [not-or-elim #1491]: #930
+#1495 := [and-elim #1492]: #924
+#1562 := [mp~ #1495 #1522]: #924
+#3898 := [mp #1562 #3897]: #3893
+#4872 := (not #3893)
+#4974 := (or #4872 #81)
+#3593 := (= uf_11 uf_11)
+#4532 := (not #3593)
+#4917 := (or #4532 #81)
+#4975 := (or #4872 #4917)
+#4976 := (iff #4975 #4974)
+#4978 := (iff #4974 #4974)
+#4979 := [rewrite]: #4978
+#4972 := (iff #4917 #81)
+#4951 := (or false #81)
+#4969 := (iff #4951 #81)
+#4971 := [rewrite]: #4969
+#4952 := (iff #4917 #4951)
+#4949 := (iff #4532 false)
+#4945 := (not true)
+#4948 := (iff #4945 false)
+#4943 := [rewrite]: #4948
+#4946 := (iff #4532 #4945)
+#4918 := (iff #3593 true)
+#4944 := [rewrite]: #4918
+#4947 := [monotonicity #4944]: #4946
+#4950 := [trans #4947 #4943]: #4949
+#4953 := [monotonicity #4950]: #4952
+#4973 := [trans #4953 #4971]: #4972
+#4977 := [monotonicity #4973]: #4976
+#4980 := [trans #4977 #4979]: #4976
+#4970 := [quant-inst]: #4975
+#4990 := [mp #4970 #4980]: #4974
+#4993 := [unit-resolution #4990 #3898 #4992]: false
+#4994 := [lemma #4993]: #81
+#4268 := (or #940 #4265)
+#3163 := (forall (vars (?x76 T2)) #3152)
+#3170 := (not #3163)
+#3148 := (forall (vars (?x71 T2) (?x72 T2)) #3143)
+#3169 := (not #3148)
+#3171 := (or #2044 #2557 #3169 #3170)
+#3172 := (not #3171)
+#3177 := (or #3126 #3172)
+#3184 := (not #3177)
+#3103 := (forall (vars (?x67 T2) (?x68 T2)) #3098)
+#3183 := (not #3103)
+#3185 := (or #3183 #3184)
+#3186 := (not #3185)
+#3191 := (or #3080 #3186)
+#3197 := (not #3191)
+#3198 := (or #1333 #3197)
+#3199 := (not #3198)
+#3204 := (or #1964 #3199)
+#3210 := (not #3204)
+#3211 := (or #1324 #3210)
+#3212 := (not #3211)
+#3217 := (or #1324 #3212)
+#3223 := (not #3217)
+#3224 := (or #814 #3223)
+#3225 := (not #3224)
+#3230 := (or #2500 #3225)
+#3236 := (not #3230)
+#3237 := (or #1319 #3236)
+#3238 := (not #3237)
+#3243 := (or #2491 #3238)
+#3253 := (not #3243)
+#3057 := (forall (vars (?x60 T2)) #3054)
+#3251 := (not #3057)
+#3051 := (forall (vars (?x59 T2)) #3046)
+#3250 := (not #3051)
+#3252 := (not #1278)
+#3254 := (or #183 #3249 #3252 #1282 #1890 #1892 #3250 #3251 #3253)
+#3255 := (not #3254)
+#2965 := (forall (vars (?x53 T2) (?x54 T2)) #2960)
+#2971 := (not #2965)
+#2972 := (or #176 #2971)
+#2973 := (not #2972)
+#3000 := (or #2973 #2997)
+#3007 := (not #3000)
+#2943 := (forall (vars (?x49 T2)) #2938)
+#3006 := (not #2943)
+#3008 := (or #3006 #3007)
+#3009 := (not #3008)
+#2906 := (forall (vars (?x50 T2)) #2895)
+#2912 := (not #2906)
+#2913 := (or #1776 #2342 #2912)
+#2914 := (not #2913)
+#3014 := (or #2914 #3009)
+#3025 := (not #3014)
+#2891 := (forall (vars (?x48 T2)) #2880)
+#3024 := (not #2891)
+#3026 := (or #3020 #3021 #3022 #3023 #3024 #3025)
#3027 := (not #3026)
-#3524 := (or #3027 #3519)
-#3531 := (not #3524)
-#3004 := (forall (vars (?x33 T2) (?x34 T2)) #2999)
-#3530 := (not #3004)
-#3532 := (or #3530 #3531)
-#3533 := (not #3532)
-#3538 := (or #2982 #3533)
-#3545 := (not #3538)
-#2959 := (forall (vars (?x29 T2) (?x30 T2)) #2954)
-#3544 := (not #2959)
-#3546 := (or #3544 #3545)
-#3547 := (not #3546)
-#3552 := (or #1948 #3547)
-#3558 := (not #3552)
-#3559 := (or #1124 #3558)
-#3560 := (not #3559)
-#3565 := (or #1796 #3560)
-#3571 := (not #3565)
-#3572 := (or #1115 #3571)
-#3573 := (not #3572)
-#3578 := (or #1115 #3573)
-#4508 := (iff #3578 #4507)
-#4505 := (iff #3573 #4504)
-#4502 := (iff #3572 #4501)
-#4499 := (iff #3571 #4498)
-#4496 := (iff #3565 #4495)
-#4493 := (iff #3560 #4492)
-#4490 := (iff #3559 #4489)
-#4487 := (iff #3558 #4486)
-#4484 := (iff #3552 #4483)
-#4481 := (iff #3547 #4480)
-#4478 := (iff #3546 #4477)
-#4475 := (iff #3545 #4474)
-#4472 := (iff #3538 #4471)
-#4469 := (iff #3533 #4468)
-#4466 := (iff #3532 #4465)
-#4463 := (iff #3531 #4462)
-#4460 := (iff #3524 #4459)
-#4457 := (iff #3519 #4456)
-#4454 := (iff #3518 #4453)
-#4451 := (iff #3517 #4450)
-#4448 := (iff #3507 #4447)
-#4445 := (iff #3502 #4444)
-#4442 := (iff #3501 #4441)
-#4439 := (iff #3500 #4438)
-#4436 := (iff #3492 #4435)
-#4433 := (iff #3487 #4432)
-#4430 := (iff #3486 #4429)
-#4427 := (iff #3485 #4426)
-#4424 := (iff #3479 #4423)
-#4421 := (iff #3474 #4420)
-#4418 := (iff #3473 #4417)
-#4415 := (iff #3472 #4414)
-#4412 := (iff #3466 #4411)
-#4409 := (iff #3461 #4408)
-#4406 := (iff #3460 #4405)
-#4403 := (iff #3459 #4402)
-#4400 := (iff #3453 #4399)
-#4397 := (iff #3448 #4396)
-#4394 := (iff #3447 #4393)
-#4391 := (iff #3446 #4390)
-#4388 := (iff #3440 #4387)
-#4385 := (iff #3435 #4384)
-#4382 := (iff #3434 #4381)
-#4379 := (iff #3433 #4378)
-#4376 := (iff #3426 #4375)
-#4373 := (iff #3421 #4372)
-#4370 := (iff #3420 #4369)
-#4367 := (iff #3419 #4366)
-#4364 := (iff #3412 #4361)
-#4362 := (iff #3401 #3401)
-#4363 := [refl]: #4362
-#4365 := [quant-intro #4363]: #4364
-#4368 := [monotonicity #4365]: #4367
-#4358 := (iff #3418 #4357)
-#4355 := (iff #3397 #4352)
-#4353 := (iff #3392 #3392)
-#4354 := [refl]: #4353
-#4356 := [quant-intro #4354]: #4355
-#4359 := [monotonicity #4356]: #4358
-#4371 := [monotonicity #4359 #4368]: #4370
-#4374 := [monotonicity #4371]: #4373
-#4377 := [monotonicity #4374]: #4376
-#4380 := [monotonicity #4377]: #4379
-#4350 := (iff #3432 #4349)
-#4347 := (iff #3352 #4344)
-#4345 := (iff #3347 #3347)
-#4346 := [refl]: #4345
-#4348 := [quant-intro #4346]: #4347
-#4351 := [monotonicity #4348]: #4350
-#4383 := [monotonicity #4351 #4380]: #4382
-#4386 := [monotonicity #4383]: #4385
-#4389 := [monotonicity #4386]: #4388
-#4392 := [monotonicity #4389]: #4391
-#4341 := (iff #1403 #4340)
-#4338 := (iff #1400 #4335)
-#4336 := (iff #1397 #1397)
-#4337 := [refl]: #4336
-#4339 := [quant-intro #4337]: #4338
-#4342 := [monotonicity #4339]: #4341
-#4395 := [monotonicity #4342 #4392]: #4394
-#4398 := [monotonicity #4395]: #4397
-#4401 := [monotonicity #4398]: #4400
-#4404 := [monotonicity #4401]: #4403
-#4407 := [monotonicity #4404]: #4406
-#4410 := [monotonicity #4407]: #4409
-#4413 := [monotonicity #4410]: #4412
-#4416 := [monotonicity #4413]: #4415
-#4333 := (iff #846 #4332)
-#4330 := (iff #710 #4327)
-#4328 := (iff #705 #705)
-#4329 := [refl]: #4328
-#4331 := [quant-intro #4329]: #4330
-#4334 := [monotonicity #4331]: #4333
-#4419 := [monotonicity #4334 #4416]: #4418
-#4422 := [monotonicity #4419]: #4421
-#4425 := [monotonicity #4422]: #4424
-#4428 := [monotonicity #4425]: #4427
-#4324 := (iff #1389 #4323)
-#4321 := (iff #1386 #4318)
-#4319 := (iff #1381 #1381)
-#4320 := [refl]: #4319
-#4322 := [quant-intro #4320]: #4321
-#4325 := [monotonicity #4322]: #4324
-#4431 := [monotonicity #4325 #4428]: #4430
-#4434 := [monotonicity #4431]: #4433
-#4437 := [monotonicity #4434]: #4436
-#4440 := [monotonicity #4437]: #4439
-#4316 := (iff #3499 #4315)
-#4313 := (iff #3306 #4310)
-#4311 := (iff #3301 #3301)
-#4312 := [refl]: #4311
-#4314 := [quant-intro #4312]: #4313
-#4317 := [monotonicity #4314]: #4316
-#4308 := (iff #3498 #4307)
-#4305 := (iff #3288 #4302)
-#4303 := (iff #3285 #3285)
-#4304 := [refl]: #4303
-#4306 := [quant-intro #4304]: #4305
-#4309 := [monotonicity #4306]: #4308
-#4298 := (iff #1544 #4297)
-#4295 := (iff #1541 #4292)
-#4293 := (iff #1538 #1538)
-#4294 := [refl]: #4293
-#4296 := [quant-intro #4294]: #4295
-#4299 := [monotonicity #4296]: #4298
-#4443 := [monotonicity #4299 #4309 #4317 #4440]: #4442
-#4446 := [monotonicity #4443]: #4445
-#4290 := (iff #3276 #4289)
-#4287 := (iff #3275 #4286)
-#4284 := (iff #3274 #4283)
-#4281 := (iff #3267 #4280)
-#4278 := (iff #3262 #4277)
-#4275 := (iff #3261 #4274)
-#4272 := (iff #3260 #4271)
-#4269 := (iff #3253 #4268)
-#4266 := (iff #3226 #4265)
-#4263 := (iff #3225 #4262)
-#4260 := (iff #3224 #4259)
-#4257 := (iff #3218 #4254)
-#4255 := (iff #3213 #3213)
-#4256 := [refl]: #4255
-#4258 := [quant-intro #4256]: #4257
+#3260 := (or #3027 #3255)
+#3272 := (not #3260)
+#2877 := (forall (vars (?x46 T2)) #2872)
+#3271 := (not #2877)
+#2849 := (forall (vars (?x37 T2)) #2844)
+#3270 := (not #2849)
+#2821 := (forall (vars (?x42 T2) (?x43 T2)) #2816)
+#3268 := (not #2821)
+#2798 := (forall (vars (?x44 T2) (?x45 T2)) #2793)
+#3267 := (not #2798)
+#3269 := (not #1104)
+#3273 := (or #3266 #3269 #3267 #3268 #3270 #3271 #3272)
+#3274 := (not #3273)
+#2766 := (forall (vars (?x38 T2)) #2755)
+#2772 := (not #2766)
+#2773 := (or #1650 #2214 #2772)
+#2774 := (not #2773)
+#3279 := (or #2774 #3274)
+#3286 := (not #3279)
+#2751 := (forall (vars (?x33 T2) (?x34 T2)) #2746)
+#3285 := (not #2751)
+#3287 := (or #3285 #3286)
+#3288 := (not #3287)
+#3293 := (or #2729 #3288)
+#3300 := (not #3293)
+#2706 := (forall (vars (?x29 T2) (?x30 T2)) #2701)
+#3299 := (not #2706)
+#3301 := (or #3299 #3300)
+#3302 := (not #3301)
+#3307 := (or #2022 #3302)
+#3313 := (not #3307)
+#3314 := (or #949 #3313)
+#3315 := (not #3314)
+#3320 := (or #1569 #3315)
+#3326 := (not #3320)
+#3327 := (or #940 #3326)
+#3328 := (not #3327)
+#3333 := (or #940 #3328)
+#4269 := (iff #3333 #4268)
+#4266 := (iff #3328 #4265)
+#4263 := (iff #3327 #4262)
+#4260 := (iff #3326 #4259)
+#4257 := (iff #3320 #4256)
+#4254 := (iff #3315 #4253)
+#4251 := (iff #3314 #4250)
+#4248 := (iff #3313 #4247)
+#4245 := (iff #3307 #4244)
+#4242 := (iff #3302 #4241)
+#4239 := (iff #3301 #4238)
+#4236 := (iff #3300 #4235)
+#4233 := (iff #3293 #4232)
+#4230 := (iff #3288 #4229)
+#4227 := (iff #3287 #4226)
+#4224 := (iff #3286 #4223)
+#4221 := (iff #3279 #4220)
+#4218 := (iff #3274 #4217)
+#4215 := (iff #3273 #4214)
+#4212 := (iff #3272 #4211)
+#4209 := (iff #3260 #4208)
+#4206 := (iff #3255 #4205)
+#4203 := (iff #3254 #4202)
+#4200 := (iff #3253 #4199)
+#4197 := (iff #3243 #4196)
+#4194 := (iff #3238 #4193)
+#4191 := (iff #3237 #4190)
+#4188 := (iff #3236 #4187)
+#4185 := (iff #3230 #4184)
+#4182 := (iff #3225 #4181)
+#4179 := (iff #3224 #4178)
+#4176 := (iff #3223 #4175)
+#4173 := (iff #3217 #4172)
+#4170 := (iff #3212 #4169)
+#4167 := (iff #3211 #4166)
+#4164 := (iff #3210 #4163)
+#4161 := (iff #3204 #4160)
+#4158 := (iff #3199 #4157)
+#4155 := (iff #3198 #4154)
+#4152 := (iff #3197 #4151)
+#4149 := (iff #3191 #4148)
+#4146 := (iff #3186 #4145)
+#4143 := (iff #3185 #4142)
+#4140 := (iff #3184 #4139)
+#4137 := (iff #3177 #4136)
+#4134 := (iff #3172 #4133)
+#4131 := (iff #3171 #4130)
+#4128 := (iff #3170 #4127)
+#4125 := (iff #3163 #4122)
+#4123 := (iff #3152 #3152)
+#4124 := [refl]: #4123
+#4126 := [quant-intro #4124]: #4125
+#4129 := [monotonicity #4126]: #4128
+#4119 := (iff #3169 #4118)
+#4116 := (iff #3148 #4113)
+#4114 := (iff #3143 #3143)
+#4115 := [refl]: #4114
+#4117 := [quant-intro #4115]: #4116
+#4120 := [monotonicity #4117]: #4119
+#4132 := [monotonicity #4120 #4129]: #4131
+#4135 := [monotonicity #4132]: #4134
+#4138 := [monotonicity #4135]: #4137
+#4141 := [monotonicity #4138]: #4140
+#4111 := (iff #3183 #4110)
+#4108 := (iff #3103 #4105)
+#4106 := (iff #3098 #3098)
+#4107 := [refl]: #4106
+#4109 := [quant-intro #4107]: #4108
+#4112 := [monotonicity #4109]: #4111
+#4144 := [monotonicity #4112 #4141]: #4143
+#4147 := [monotonicity #4144]: #4146
+#4150 := [monotonicity #4147]: #4149
+#4153 := [monotonicity #4150]: #4152
+#4102 := (iff #1333 #4101)
+#4099 := (iff #1330 #4096)
+#4097 := (iff #1327 #1327)
+#4098 := [refl]: #4097
+#4100 := [quant-intro #4098]: #4099
+#4103 := [monotonicity #4100]: #4102
+#4156 := [monotonicity #4103 #4153]: #4155
+#4159 := [monotonicity #4156]: #4158
+#4162 := [monotonicity #4159]: #4161
+#4165 := [monotonicity #4162]: #4164
+#4168 := [monotonicity #4165]: #4167
+#4171 := [monotonicity #4168]: #4170
+#4174 := [monotonicity #4171]: #4173
+#4177 := [monotonicity #4174]: #4176
+#4094 := (iff #814 #4093)
+#4091 := (iff #705 #4088)
+#4089 := (iff #700 #700)
+#4090 := [refl]: #4089
+#4092 := [quant-intro #4090]: #4091
+#4095 := [monotonicity #4092]: #4094
+#4180 := [monotonicity #4095 #4177]: #4179
+#4183 := [monotonicity #4180]: #4182
+#4186 := [monotonicity #4183]: #4185
+#4189 := [monotonicity #4186]: #4188
+#4085 := (iff #1319 #4084)
+#4082 := (iff #1316 #4079)
+#4080 := (iff #1312 #1312)
+#4081 := [refl]: #4080
+#4083 := [quant-intro #4081]: #4082
+#4086 := [monotonicity #4083]: #4085
+#4192 := [monotonicity #4086 #4189]: #4191
+#4195 := [monotonicity #4192]: #4194
+#4198 := [monotonicity #4195]: #4197
+#4201 := [monotonicity #4198]: #4200
+#4077 := (iff #3251 #4076)
+#4074 := (iff #3057 #4071)
+#4072 := (iff #3054 #3054)
+#4073 := [refl]: #4072
+#4075 := [quant-intro #4073]: #4074
+#4078 := [monotonicity #4075]: #4077
+#4069 := (iff #3250 #4068)
+#4066 := (iff #3051 #4063)
+#4064 := (iff #3046 #3046)
+#4065 := [refl]: #4064
+#4067 := [quant-intro #4065]: #4066
+#4070 := [monotonicity #4067]: #4069
+#4059 := (iff #3252 #4058)
+#4056 := (iff #1278 #4053)
+#4054 := (iff #1275 #1275)
+#4055 := [refl]: #4054
+#4057 := [quant-intro #4055]: #4056
+#4060 := [monotonicity #4057]: #4059
+#4204 := [monotonicity #4060 #4070 #4078 #4201]: #4203
+#4207 := [monotonicity #4204]: #4206
+#4051 := (iff #3027 #4050)
+#4048 := (iff #3026 #4047)
+#4045 := (iff #3025 #4044)
+#4042 := (iff #3014 #4041)
+#4039 := (iff #3009 #4038)
+#4036 := (iff #3008 #4035)
+#4033 := (iff #3007 #4032)
+#4030 := (iff #3000 #4029)
+#4027 := (iff #2973 #4026)
+#4024 := (iff #2972 #4023)
+#4021 := (iff #2971 #4020)
+#4018 := (iff #2965 #4015)
+#4016 := (iff #2960 #2960)
+#4017 := [refl]: #4016
+#4019 := [quant-intro #4017]: #4018
+#4022 := [monotonicity #4019]: #4021
+#4025 := [monotonicity #4022]: #4024
+#4028 := [monotonicity #4025]: #4027
+#4031 := [monotonicity #4028]: #4030
+#4034 := [monotonicity #4031]: #4033
+#4013 := (iff #3006 #4012)
+#4010 := (iff #2943 #4007)
+#4008 := (iff #2938 #2938)
+#4009 := [refl]: #4008
+#4011 := [quant-intro #4009]: #4010
+#4014 := [monotonicity #4011]: #4013
+#4037 := [monotonicity #4014 #4034]: #4036
+#4040 := [monotonicity #4037]: #4039
+#4005 := (iff #2914 #4004)
+#4002 := (iff #2913 #4001)
+#3999 := (iff #2912 #3998)
+#3996 := (iff #2906 #3993)
+#3994 := (iff #2895 #2895)
+#3995 := [refl]: #3994
+#3997 := [quant-intro #3995]: #3996
+#4000 := [monotonicity #3997]: #3999
+#4003 := [monotonicity #4000]: #4002
+#4006 := [monotonicity #4003]: #4005
+#4043 := [monotonicity #4006 #4040]: #4042
+#4046 := [monotonicity #4043]: #4045
+#3989 := (iff #3024 #3988)
+#3986 := (iff #2891 #3983)
+#3984 := (iff #2880 #2880)
+#3985 := [refl]: #3984
+#3987 := [quant-intro #3985]: #3986
+#3990 := [monotonicity #3987]: #3989
+#4049 := [monotonicity #3990 #4046]: #4048
+#4052 := [monotonicity #4049]: #4051
+#4210 := [monotonicity #4052 #4207]: #4209
+#4213 := [monotonicity #4210]: #4212
+#3980 := (iff #3271 #3979)
+#3977 := (iff #2877 #3974)
+#3975 := (iff #2872 #2872)
+#3976 := [refl]: #3975
+#3978 := [quant-intro #3976]: #3977
+#3981 := [monotonicity #3978]: #3980
+#3972 := (iff #3270 #3971)
+#3969 := (iff #2849 #3966)
+#3967 := (iff #2844 #2844)
+#3968 := [refl]: #3967
+#3970 := [quant-intro #3968]: #3969
+#3973 := [monotonicity #3970]: #3972
+#3964 := (iff #3268 #3963)
+#3961 := (iff #2821 #3958)
+#3959 := (iff #2816 #2816)
+#3960 := [refl]: #3959
+#3962 := [quant-intro #3960]: #3961
+#3965 := [monotonicity #3962]: #3964
+#3955 := (iff #3267 #3954)
+#3952 := (iff #2798 #3949)
+#3950 := (iff #2793 #2793)
+#3951 := [refl]: #3950
+#3953 := [quant-intro #3951]: #3952
+#3956 := [monotonicity #3953]: #3955
+#3947 := (iff #3269 #3946)
+#3944 := (iff #1104 #3941)
+#3942 := (iff #1101 #1101)
+#3943 := [refl]: #3942
+#3945 := [quant-intro #3943]: #3944
+#3948 := [monotonicity #3945]: #3947
+#4216 := [monotonicity #3948 #3956 #3965 #3973 #3981 #4213]: #4215
+#4219 := [monotonicity #4216]: #4218
+#3938 := (iff #2774 #3937)
+#3935 := (iff #2773 #3934)
+#3932 := (iff #2772 #3931)
+#3929 := (iff #2766 #3926)
+#3927 := (iff #2755 #2755)
+#3928 := [refl]: #3927
+#3930 := [quant-intro #3928]: #3929
+#3933 := [monotonicity #3930]: #3932
+#3936 := [monotonicity #3933]: #3935
+#3939 := [monotonicity #3936]: #3938
+#4222 := [monotonicity #3939 #4219]: #4221
+#4225 := [monotonicity #4222]: #4224
+#3923 := (iff #3285 #3922)
+#3920 := (iff #2751 #3917)
+#3918 := (iff #2746 #2746)
+#3919 := [refl]: #3918
+#3921 := [quant-intro #3919]: #3920
+#3924 := [monotonicity #3921]: #3923
+#4228 := [monotonicity #3924 #4225]: #4227
+#4231 := [monotonicity #4228]: #4230
+#4234 := [monotonicity #4231]: #4233
+#4237 := [monotonicity #4234]: #4236
+#3914 := (iff #3299 #3913)
+#3911 := (iff #2706 #3908)
+#3909 := (iff #2701 #2701)
+#3910 := [refl]: #3909
+#3912 := [quant-intro #3910]: #3911
+#3915 := [monotonicity #3912]: #3914
+#4240 := [monotonicity #3915 #4237]: #4239
+#4243 := [monotonicity #4240]: #4242
+#4246 := [monotonicity #4243]: #4245
+#4249 := [monotonicity #4246]: #4248
+#3905 := (iff #949 #3904)
+#3902 := (iff #946 #3899)
+#3900 := (iff #945 #945)
+#3901 := [refl]: #3900
+#3903 := [quant-intro #3901]: #3902
+#3906 := [monotonicity #3903]: #3905
+#4252 := [monotonicity #3906 #4249]: #4251
+#4255 := [monotonicity #4252]: #4254
+#4258 := [monotonicity #4255]: #4257
#4261 := [monotonicity #4258]: #4260
#4264 := [monotonicity #4261]: #4263
#4267 := [monotonicity #4264]: #4266
#4270 := [monotonicity #4267]: #4269
-#4273 := [monotonicity #4270]: #4272
-#4252 := (iff #3259 #4251)
-#4249 := (iff #3196 #4246)
-#4247 := (iff #3191 #3191)
-#4248 := [refl]: #4247
-#4250 := [quant-intro #4248]: #4249
-#4253 := [monotonicity #4250]: #4252
-#4276 := [monotonicity #4253 #4273]: #4275
-#4279 := [monotonicity #4276]: #4278
-#4244 := (iff #3167 #4243)
-#4241 := (iff #3166 #4240)
-#4238 := (iff #3165 #4237)
-#4235 := (iff #3159 #4232)
-#4233 := (iff #3148 #3148)
-#4234 := [refl]: #4233
-#4236 := [quant-intro #4234]: #4235
-#4239 := [monotonicity #4236]: #4238
-#4242 := [monotonicity #4239]: #4241
-#4245 := [monotonicity #4242]: #4244
-#4282 := [monotonicity #4245 #4279]: #4281
-#4285 := [monotonicity #4282]: #4284
-#4228 := (iff #3273 #4227)
-#4225 := (iff #3144 #4222)
-#4223 := (iff #3133 #3133)
-#4224 := [refl]: #4223
-#4226 := [quant-intro #4224]: #4225
-#4229 := [monotonicity #4226]: #4228
-#4288 := [monotonicity #4229 #4285]: #4287
-#4291 := [monotonicity #4288]: #4290
-#4449 := [monotonicity #4291 #4446]: #4448
-#4452 := [monotonicity #4449]: #4451
-#4219 := (iff #3516 #4218)
-#4216 := (iff #3130 #4213)
-#4214 := (iff #3125 #3125)
-#4215 := [refl]: #4214
-#4217 := [quant-intro #4215]: #4216
-#4220 := [monotonicity #4217]: #4219
-#4211 := (iff #3515 #4210)
-#4208 := (iff #3102 #4205)
-#4206 := (iff #3097 #3097)
-#4207 := [refl]: #4206
-#4209 := [quant-intro #4207]: #4208
-#4212 := [monotonicity #4209]: #4211
-#4203 := (iff #3514 #4202)
-#4200 := (iff #3074 #4197)
-#4198 := (iff #3069 #3069)
-#4199 := [refl]: #4198
-#4201 := [quant-intro #4199]: #4200
-#4204 := [monotonicity #4201]: #4203
-#4194 := (iff #3513 #4193)
-#4191 := (iff #3051 #4188)
-#4189 := (iff #3046 #3046)
-#4190 := [refl]: #4189
-#4192 := [quant-intro #4190]: #4191
-#4195 := [monotonicity #4192]: #4194
-#4186 := (iff #1631 #4185)
-#4183 := (iff #1628 #4180)
-#4181 := (iff #1625 #1625)
-#4182 := [refl]: #4181
-#4184 := [quant-intro #4182]: #4183
-#4187 := [monotonicity #4184]: #4186
-#4455 := [monotonicity #4187 #4195 #4204 #4212 #4220 #4452]: #4454
-#4458 := [monotonicity #4455]: #4457
-#4177 := (iff #3027 #4176)
-#4174 := (iff #3026 #4173)
-#4171 := (iff #3025 #4170)
-#4168 := (iff #3019 #4165)
-#4166 := (iff #3008 #3008)
-#4167 := [refl]: #4166
-#4169 := [quant-intro #4167]: #4168
-#4172 := [monotonicity #4169]: #4171
-#4175 := [monotonicity #4172]: #4174
-#4178 := [monotonicity #4175]: #4177
-#4461 := [monotonicity #4178 #4458]: #4460
-#4464 := [monotonicity #4461]: #4463
-#4162 := (iff #3530 #4161)
-#4159 := (iff #3004 #4156)
-#4157 := (iff #2999 #2999)
-#4158 := [refl]: #4157
-#4160 := [quant-intro #4158]: #4159
-#4163 := [monotonicity #4160]: #4162
-#4467 := [monotonicity #4163 #4464]: #4466
-#4470 := [monotonicity #4467]: #4469
-#4473 := [monotonicity #4470]: #4472
-#4476 := [monotonicity #4473]: #4475
-#4153 := (iff #3544 #4152)
-#4150 := (iff #2959 #4147)
-#4148 := (iff #2954 #2954)
-#4149 := [refl]: #4148
-#4151 := [quant-intro #4149]: #4150
-#4154 := [monotonicity #4151]: #4153
-#4479 := [monotonicity #4154 #4476]: #4478
-#4482 := [monotonicity #4479]: #4481
-#4485 := [monotonicity #4482]: #4484
-#4488 := [monotonicity #4485]: #4487
-#4144 := (iff #1124 #4143)
-#4141 := (iff #1121 #4138)
-#4139 := (iff #1120 #1120)
-#4140 := [refl]: #4139
-#4142 := [quant-intro #4140]: #4141
-#4145 := [monotonicity #4142]: #4144
-#4491 := [monotonicity #4145 #4488]: #4490
-#4494 := [monotonicity #4491]: #4493
-#4497 := [monotonicity #4494]: #4496
-#4500 := [monotonicity #4497]: #4499
-#4503 := [monotonicity #4500]: #4502
-#4506 := [monotonicity #4503]: #4505
-#4509 := [monotonicity #4506]: #4508
-#2244 := (not #2243)
-#2841 := (and #206 #2244 #2838)
-#2844 := (not #2841)
-#2847 := (forall (vars (?x76 T2)) #2844)
-#2813 := (not #2810)
-#2249 := (not #2248)
-#2856 := (and #1438 #2249 #2813 #2847)
-#2218 := (not #2217)
-#2783 := (and #2218 #2219)
-#2786 := (not #2783)
-#2804 := (or #2786 #2799)
-#2807 := (not #2804)
-#2861 := (or #2807 #2856)
-#2864 := (and #1414 #2861)
-#2189 := (not #2188)
-#2758 := (and #2187 #2189)
-#2761 := (not #2758)
-#2777 := (or #2761 #2772)
-#2780 := (not #2777)
-#2867 := (or #2780 #2864)
-#2870 := (and #1400 #2867)
-#2873 := (or #2168 #2870)
-#2876 := (and #210 #2873)
-#2879 := (or #1394 #2876)
-#2882 := (and #710 #2879)
-#2885 := (or #2753 #2882)
-#2888 := (and #1386 #2885)
-#2891 := (or #2744 #2888)
-#2107 := (not #2106)
-#2105 := (not #2104)
-#2897 := (and #181 #189 #1375 #1528 #1541 #1549 #2105 #2107 #2891)
-#2050 := (not #2049)
-#2046 := (not #2045)
-#2681 := (and #2046 #2050)
-#2684 := (not #2681)
-#2701 := (or #2684 #2696)
-#2704 := (not #2701)
-#2059 := (not #162)
-#2069 := (and #2059 #1300)
-#2710 := (or #2069 #2704)
-#2654 := (not #2649)
-#2672 := (and #2654 #2667)
-#2675 := (or #1260 #2672)
-#2678 := (forall (vars (?x49 T2)) #2675)
-#2715 := (and #2678 #2710)
-#1979 := (not #1978)
-#2624 := (and #1979 #2621)
-#2627 := (not #2624)
-#2630 := (forall (vars (?x50 T2)) #2627)
-#2596 := (not #2593)
-#1984 := (not #1983)
-#2636 := (and #1984 #2596 #2630)
-#2718 := (or #2636 #2715)
-#2085 := (not #1325)
-#2088 := (forall (vars (?x48 T2)) #2085)
-#2724 := (and #140 #145 #506 #509 #2088 #2718)
-#2902 := (or #2724 #2897)
-#2573 := (not #2568)
-#2576 := (and #1943 #2556 #2573)
-#2579 := (or #1215 #2576)
-#2582 := (forall (vars (?x46 T2)) #2579)
-#2518 := (not #2513)
-#2536 := (and #1917 #2518 #2531)
-#2539 := (or #1177 #2536)
-#2542 := (forall (vars (?x37 T2)) #2539)
-#2908 := (and #106 #1608 #1619 #1628 #2542 #2582 #2902)
-#1873 := (not #1872)
-#2486 := (and #74 #1873 #2483)
-#2489 := (not #2486)
-#2492 := (forall (vars (?x38 T2)) #2489)
-#2458 := (not #2455)
-#1878 := (not #1877)
-#2498 := (and #1878 #2458 #2492)
-#2913 := (or #2498 #2908)
-#2916 := (and #1162 #2913)
-#1846 := (not #1845)
-#2436 := (and #1846 #1847)
-#2439 := (not #2436)
-#2449 := (or #2439 #2446)
-#2452 := (not #2449)
-#2919 := (or #2452 #2916)
-#2922 := (and #1136 #2919)
-#1817 := (not #1816)
-#2411 := (and #1815 #1817)
-#2414 := (not #2411)
-#2430 := (or #2414 #2425)
+#2040 := (not #2039)
+#2588 := (and #216 #2040 #2585)
+#2591 := (not #2588)
+#2594 := (forall (vars (?x76 T2)) #2591)
+#2560 := (not #2557)
+#2603 := (and #1368 #2045 #2560 #2594)
+#2530 := (and #2014 #2015)
+#2533 := (not #2530)
+#2551 := (or #2533 #2546)
+#2554 := (not #2551)
+#2608 := (or #2554 #2603)
+#2611 := (and #1344 #2608)
+#1985 := (not #1984)
+#2505 := (and #1983 #1985)
+#2508 := (not #2505)
+#2524 := (or #2508 #2519)
+#2527 := (not #2524)
+#2614 := (or #2527 #2611)
+#2617 := (and #1330 #2614)
+#2620 := (or #1964 #2617)
+#2623 := (and #220 #2620)
+#2626 := (or #1324 #2623)
+#2629 := (and #705 #2626)
+#2632 := (or #2500 #2629)
+#2635 := (and #1316 #2632)
+#2638 := (or #2491 #2635)
+#1893 := (not #1892)
+#1891 := (not #1890)
+#2644 := (and #184 #192 #1257 #1268 #1278 #1283 #1891 #1893 #2638)
+#1843 := (not #1842)
+#1839 := (not #1838)
+#2430 := (and #1839 #1843)
#2433 := (not #2430)
-#2925 := (or #2433 #2922)
-#2928 := (and #1121 #2925)
-#2931 := (or #1796 #2928)
-#2934 := (and #78 #2931)
-#2937 := (or #1115 #2934)
-#3579 := (iff #2937 #3578)
-#3576 := (iff #2934 #3573)
-#3568 := (and #78 #3565)
-#3574 := (iff #3568 #3573)
-#3575 := [rewrite]: #3574
-#3569 := (iff #2934 #3568)
-#3566 := (iff #2931 #3565)
-#3563 := (iff #2928 #3560)
-#3555 := (and #1121 #3552)
-#3561 := (iff #3555 #3560)
-#3562 := [rewrite]: #3561
-#3556 := (iff #2928 #3555)
-#3553 := (iff #2925 #3552)
-#3550 := (iff #2922 #3547)
-#3541 := (and #2959 #3538)
-#3548 := (iff #3541 #3547)
-#3549 := [rewrite]: #3548
-#3542 := (iff #2922 #3541)
-#3539 := (iff #2919 #3538)
-#3536 := (iff #2916 #3533)
-#3527 := (and #3004 #3524)
-#3534 := (iff #3527 #3533)
-#3535 := [rewrite]: #3534
-#3528 := (iff #2916 #3527)
-#3525 := (iff #2913 #3524)
-#3522 := (iff #2908 #3519)
-#3510 := (and #106 #3051 #3074 #1628 #3102 #3130 #3507)
-#3520 := (iff #3510 #3519)
-#3521 := [rewrite]: #3520
-#3511 := (iff #2908 #3510)
-#3508 := (iff #2902 #3507)
-#3505 := (iff #2897 #3502)
-#3495 := (and #181 #189 #3288 #3306 #1541 #1549 #2105 #2107 #3492)
-#3503 := (iff #3495 #3502)
-#3504 := [rewrite]: #3503
-#3496 := (iff #2897 #3495)
-#3493 := (iff #2891 #3492)
-#3490 := (iff #2888 #3487)
-#3482 := (and #1386 #3479)
-#3488 := (iff #3482 #3487)
-#3489 := [rewrite]: #3488
-#3483 := (iff #2888 #3482)
-#3480 := (iff #2885 #3479)
-#3477 := (iff #2882 #3474)
-#3469 := (and #710 #3466)
-#3475 := (iff #3469 #3474)
-#3476 := [rewrite]: #3475
-#3470 := (iff #2882 #3469)
-#3467 := (iff #2879 #3466)
-#3464 := (iff #2876 #3461)
-#3456 := (and #210 #3453)
-#3462 := (iff #3456 #3461)
-#3463 := [rewrite]: #3462
-#3457 := (iff #2876 #3456)
-#3454 := (iff #2873 #3453)
-#3451 := (iff #2870 #3448)
-#3443 := (and #1400 #3440)
-#3449 := (iff #3443 #3448)
-#3450 := [rewrite]: #3449
-#3444 := (iff #2870 #3443)
-#3441 := (iff #2867 #3440)
-#3438 := (iff #2864 #3435)
-#3429 := (and #3352 #3426)
-#3436 := (iff #3429 #3435)
-#3437 := [rewrite]: #3436
-#3430 := (iff #2864 #3429)
-#3427 := (iff #2861 #3426)
-#3424 := (iff #2856 #3421)
-#3415 := (and #3397 #2249 #2813 #3412)
-#3422 := (iff #3415 #3421)
-#3423 := [rewrite]: #3422
-#3416 := (iff #2856 #3415)
-#3413 := (iff #2847 #3412)
-#3410 := (iff #2844 #3401)
-#3402 := (not #3401)
-#3405 := (not #3402)
-#3408 := (iff #3405 #3401)
-#3409 := [rewrite]: #3408
-#3406 := (iff #2844 #3405)
-#3403 := (iff #2841 #3402)
-#3404 := [rewrite]: #3403
-#3407 := [monotonicity #3404]: #3406
-#3411 := [trans #3407 #3409]: #3410
-#3414 := [quant-intro #3411]: #3413
-#3398 := (iff #1438 #3397)
-#3395 := (iff #1435 #3392)
-#3378 := (or #213 #1144)
-#3389 := (or #3378 #1431)
-#3393 := (iff #3389 #3392)
-#3394 := [rewrite]: #3393
-#3390 := (iff #1435 #3389)
-#3387 := (iff #1428 #3378)
-#3379 := (not #3378)
-#3382 := (not #3379)
-#3385 := (iff #3382 #3378)
-#3386 := [rewrite]: #3385
-#3383 := (iff #1428 #3382)
-#3380 := (iff #1423 #3379)
-#3381 := [rewrite]: #3380
-#3384 := [monotonicity #3381]: #3383
-#3388 := [trans #3384 #3386]: #3387
-#3391 := [monotonicity #3388]: #3390
-#3396 := [trans #3391 #3394]: #3395
-#3399 := [quant-intro #3396]: #3398
-#3417 := [monotonicity #3399 #3414]: #3416
-#3425 := [trans #3417 #3423]: #3424
-#3376 := (iff #2807 #3375)
-#3373 := (iff #2804 #3370)
-#3356 := (or #2217 #3355)
-#3367 := (or #3356 #2799)
-#3371 := (iff #3367 #3370)
-#3372 := [rewrite]: #3371
-#3368 := (iff #2804 #3367)
-#3365 := (iff #2786 #3356)
-#3357 := (not #3356)
-#3360 := (not #3357)
-#3363 := (iff #3360 #3356)
-#3364 := [rewrite]: #3363
-#3361 := (iff #2786 #3360)
-#3358 := (iff #2783 #3357)
-#3359 := [rewrite]: #3358
-#3362 := [monotonicity #3359]: #3361
-#3366 := [trans #3362 #3364]: #3365
-#3369 := [monotonicity #3366]: #3368
-#3374 := [trans #3369 #3372]: #3373
-#3377 := [monotonicity #3374]: #3376
-#3428 := [monotonicity #3377 #3425]: #3427
-#3353 := (iff #1414 #3352)
-#3350 := (iff #1411 #3347)
-#3333 := (or #206 #3332)
-#3344 := (or #3333 #1406)
-#3348 := (iff #3344 #3347)
-#3349 := [rewrite]: #3348
-#3345 := (iff #1411 #3344)
-#3342 := (iff #716 #3333)
-#3334 := (not #3333)
-#3337 := (not #3334)
-#3340 := (iff #3337 #3333)
-#3341 := [rewrite]: #3340
-#3338 := (iff #716 #3337)
-#3335 := (iff #215 #3334)
-#3336 := [rewrite]: #3335
-#3339 := [monotonicity #3336]: #3338
-#3343 := [trans #3339 #3341]: #3342
-#3346 := [monotonicity #3343]: #3345
-#3351 := [trans #3346 #3349]: #3350
-#3354 := [quant-intro #3351]: #3353
-#3431 := [monotonicity #3354 #3428]: #3430
-#3439 := [trans #3431 #3437]: #3438
-#3330 := (iff #2780 #3329)
-#3327 := (iff #2777 #3324)
-#3310 := (or #3309 #2188)
-#3321 := (or #3310 #2772)
-#3325 := (iff #3321 #3324)
-#3326 := [rewrite]: #3325
-#3322 := (iff #2777 #3321)
-#3319 := (iff #2761 #3310)
-#3311 := (not #3310)
-#3314 := (not #3311)
-#3317 := (iff #3314 #3310)
-#3318 := [rewrite]: #3317
-#3315 := (iff #2761 #3314)
-#3312 := (iff #2758 #3311)
-#3313 := [rewrite]: #3312
-#3316 := [monotonicity #3313]: #3315
-#3320 := [trans #3316 #3318]: #3319
-#3323 := [monotonicity #3320]: #3322
-#3328 := [trans #3323 #3326]: #3327
-#3331 := [monotonicity #3328]: #3330
-#3442 := [monotonicity #3331 #3439]: #3441
-#3445 := [monotonicity #3442]: #3444
-#3452 := [trans #3445 #3450]: #3451
-#3455 := [monotonicity #3452]: #3454
-#3458 := [monotonicity #3455]: #3457
-#3465 := [trans #3458 #3463]: #3464
-#3468 := [monotonicity #3465]: #3467
-#3471 := [monotonicity #3468]: #3470
-#3478 := [trans #3471 #3476]: #3477
-#3481 := [monotonicity #3478]: #3480
-#3484 := [monotonicity #3481]: #3483
-#3491 := [trans #3484 #3489]: #3490
-#3494 := [monotonicity #3491]: #3493
-#3307 := (iff #1528 #3306)
-#3304 := (iff #1525 #3301)
-#3298 := (or #3281 #1522)
-#3302 := (iff #3298 #3301)
-#3303 := [rewrite]: #3302
-#3299 := (iff #1525 #3298)
-#3296 := (iff #1517 #3281)
-#3291 := (not #3282)
-#3294 := (iff #3291 #3281)
-#3295 := [rewrite]: #3294
-#3292 := (iff #1517 #3291)
-#3283 := (iff #1364 #3282)
-#3284 := [rewrite]: #3283
-#3293 := [monotonicity #3284]: #3292
-#3297 := [trans #3293 #3295]: #3296
-#3300 := [monotonicity #3297]: #3299
-#3305 := [trans #3300 #3303]: #3304
-#3308 := [quant-intro #3305]: #3307
-#3289 := (iff #1375 #3288)
-#3286 := (iff #1370 #3285)
-#3287 := [monotonicity #3284]: #3286
-#3290 := [quant-intro #3287]: #3289
-#3497 := [monotonicity #3290 #3308 #3494]: #3496
-#3506 := [trans #3497 #3504]: #3505
-#3279 := (iff #2724 #3276)
-#3270 := (and #140 #145 #506 #509 #3144 #3267)
-#3277 := (iff #3270 #3276)
-#3278 := [rewrite]: #3277
-#3271 := (iff #2724 #3270)
-#3268 := (iff #2718 #3267)
-#3265 := (iff #2715 #3262)
-#3256 := (and #3196 #3253)
-#3263 := (iff #3256 #3262)
-#3264 := [rewrite]: #3263
-#3257 := (iff #2715 #3256)
-#3254 := (iff #2710 #3253)
-#3251 := (iff #2704 #3250)
-#3248 := (iff #2701 #3245)
-#3231 := (or #2045 #2049)
-#3242 := (or #3231 #2696)
-#3246 := (iff #3242 #3245)
-#3247 := [rewrite]: #3246
-#3243 := (iff #2701 #3242)
-#3240 := (iff #2684 #3231)
-#3232 := (not #3231)
-#3235 := (not #3232)
-#3238 := (iff #3235 #3231)
-#3239 := [rewrite]: #3238
-#3236 := (iff #2684 #3235)
-#3233 := (iff #2681 #3232)
-#3234 := [rewrite]: #3233
-#3237 := [monotonicity #3234]: #3236
-#3241 := [trans #3237 #3239]: #3240
-#3244 := [monotonicity #3241]: #3243
-#3249 := [trans #3244 #3247]: #3248
-#3252 := [monotonicity #3249]: #3251
-#3229 := (iff #2069 #3226)
-#3221 := (and #2059 #3218)
-#3227 := (iff #3221 #3226)
-#3228 := [rewrite]: #3227
-#3222 := (iff #2069 #3221)
-#3219 := (iff #1300 #3218)
-#3216 := (iff #1297 #3213)
-#3199 := (or #1144 #1253)
-#3210 := (or #3199 #1294)
-#3214 := (iff #3210 #3213)
-#3215 := [rewrite]: #3214
-#3211 := (iff #1297 #3210)
-#3208 := (iff #1291 #3199)
-#3200 := (not #3199)
-#3203 := (not #3200)
-#3206 := (iff #3203 #3199)
-#3207 := [rewrite]: #3206
-#3204 := (iff #1291 #3203)
-#3201 := (iff #1288 #3200)
-#3202 := [rewrite]: #3201
-#3205 := [monotonicity #3202]: #3204
-#3209 := [trans #3205 #3207]: #3208
-#3212 := [monotonicity #3209]: #3211
-#3217 := [trans #3212 #3215]: #3216
-#3220 := [quant-intro #3217]: #3219
-#3223 := [monotonicity #3220]: #3222
-#3230 := [trans #3223 #3228]: #3229
-#3255 := [monotonicity #3230 #3252]: #3254
-#3197 := (iff #2678 #3196)
-#3194 := (iff #2675 #3191)
-#3172 := (or #65 #1253)
-#3188 := (or #3172 #3185)
-#3192 := (iff #3188 #3191)
-#3193 := [rewrite]: #3192
-#3189 := (iff #2675 #3188)
-#3186 := (iff #2672 #3185)
-#3187 := [rewrite]: #3186
-#3181 := (iff #1260 #3172)
-#3173 := (not #3172)
-#3176 := (not #3173)
-#3179 := (iff #3176 #3172)
-#3180 := [rewrite]: #3179
-#3177 := (iff #1260 #3176)
-#3174 := (iff #1257 #3173)
-#3175 := [rewrite]: #3174
-#3178 := [monotonicity #3175]: #3177
-#3182 := [trans #3178 #3180]: #3181
-#3190 := [monotonicity #3182 #3187]: #3189
-#3195 := [trans #3190 #3193]: #3194
-#3198 := [quant-intro #3195]: #3197
-#3258 := [monotonicity #3198 #3255]: #3257
-#3266 := [trans #3258 #3264]: #3265
-#3170 := (iff #2636 #3167)
-#3162 := (and #1984 #2596 #3159)
-#3168 := (iff #3162 #3167)
-#3169 := [rewrite]: #3168
-#3163 := (iff #2636 #3162)
-#3160 := (iff #2630 #3159)
-#3157 := (iff #2627 #3148)
-#3149 := (not #3148)
-#3152 := (not #3149)
-#3155 := (iff #3152 #3148)
-#3156 := [rewrite]: #3155
-#3153 := (iff #2627 #3152)
-#3150 := (iff #2624 #3149)
-#3151 := [rewrite]: #3150
-#3154 := [monotonicity #3151]: #3153
-#3158 := [trans #3154 #3156]: #3157
-#3161 := [quant-intro #3158]: #3160
-#3164 := [monotonicity #3161]: #3163
-#3171 := [trans #3164 #3169]: #3170
-#3269 := [monotonicity #3171 #3266]: #3268
-#3145 := (iff #2088 #3144)
-#3142 := (iff #2085 #3133)
-#3134 := (not #3133)
-#3137 := (not #3134)
-#3140 := (iff #3137 #3133)
-#3141 := [rewrite]: #3140
-#3138 := (iff #2085 #3137)
-#3135 := (iff #1325 #3134)
-#3136 := [rewrite]: #3135
-#3139 := [monotonicity #3136]: #3138
-#3143 := [trans #3139 #3141]: #3142
-#3146 := [quant-intro #3143]: #3145
-#3272 := [monotonicity #3146 #3269]: #3271
-#3280 := [trans #3272 #3278]: #3279
-#3509 := [monotonicity #3280 #3506]: #3508
-#3131 := (iff #2582 #3130)
-#3128 := (iff #2579 #3125)
-#3105 := (or #65 #1208)
-#3122 := (or #3105 #3119)
-#3126 := (iff #3122 #3125)
-#3127 := [rewrite]: #3126
-#3123 := (iff #2579 #3122)
-#3120 := (iff #2576 #3119)
-#3121 := [rewrite]: #3120
-#3114 := (iff #1215 #3105)
-#3106 := (not #3105)
-#3109 := (not #3106)
-#3112 := (iff #3109 #3105)
-#3113 := [rewrite]: #3112
-#3110 := (iff #1215 #3109)
-#3107 := (iff #1212 #3106)
-#3108 := [rewrite]: #3107
-#3111 := [monotonicity #3108]: #3110
-#3115 := [trans #3111 #3113]: #3114
-#3124 := [monotonicity #3115 #3121]: #3123
-#3129 := [trans #3124 #3127]: #3128
-#3132 := [quant-intro #3129]: #3131
-#3103 := (iff #2542 #3102)
-#3100 := (iff #2539 #3097)
-#3077 := (or #65 #1170)
-#3094 := (or #3077 #3091)
-#3098 := (iff #3094 #3097)
-#3099 := [rewrite]: #3098
-#3095 := (iff #2539 #3094)
-#3092 := (iff #2536 #3091)
-#3093 := [rewrite]: #3092
-#3086 := (iff #1177 #3077)
-#3078 := (not #3077)
-#3081 := (not #3078)
-#3084 := (iff #3081 #3077)
-#3085 := [rewrite]: #3084
-#3082 := (iff #1177 #3081)
-#3079 := (iff #1174 #3078)
-#3080 := [rewrite]: #3079
-#3083 := [monotonicity #3080]: #3082
-#3087 := [trans #3083 #3085]: #3086
-#3096 := [monotonicity #3087 #3093]: #3095
-#3101 := [trans #3096 #3099]: #3100
-#3104 := [quant-intro #3101]: #3103
-#3075 := (iff #1619 #3074)
-#3072 := (iff #1616 #3069)
-#3055 := (or #111 #3054)
-#3066 := (or #3055 #1224)
-#3070 := (iff #3066 #3069)
-#3071 := [rewrite]: #3070
-#3067 := (iff #1616 #3066)
-#3064 := (iff #454 #3055)
-#3056 := (not #3055)
-#3059 := (not #3056)
-#3062 := (iff #3059 #3055)
-#3063 := [rewrite]: #3062
-#3060 := (iff #454 #3059)
-#3057 := (iff #114 #3056)
-#3058 := [rewrite]: #3057
-#3061 := [monotonicity #3058]: #3060
-#3065 := [trans #3061 #3063]: #3064
-#3068 := [monotonicity #3065]: #3067
-#3073 := [trans #3068 #3071]: #3072
-#3076 := [quant-intro #3073]: #3075
-#3052 := (iff #1608 #3051)
-#3049 := (iff #1605 #3046)
-#3032 := (or #112 #1144)
-#3043 := (or #3032 #1602)
+#2450 := (or #2433 #2445)
+#2453 := (not #2450)
+#1852 := (not #176)
+#1862 := (and #1852 #1201)
+#2459 := (or #1862 #2453)
+#2403 := (not #2398)
+#2421 := (and #2403 #2416)
+#2424 := (or #1161 #2421)
+#2427 := (forall (vars (?x49 T2)) #2424)
+#2464 := (and #2427 #2459)
+#1772 := (not #1771)
+#2373 := (and #1772 #2370)
+#2376 := (not #2373)
+#2379 := (forall (vars (?x50 T2)) #2376)
+#2345 := (not #2342)
+#1777 := (not #1776)
+#2385 := (and #1777 #2345 #2379)
+#2467 := (or #2385 #2464)
+#1751 := (not #1132)
+#1754 := (forall (vars (?x48 T2)) #1751)
+#2473 := (and #149 #154 #530 #533 #1754 #2467)
+#2649 := (or #2473 #2644)
+#2330 := (not #2325)
+#2333 := (and #1717 #2313 #2330)
+#2336 := (or #1063 #2333)
+#2339 := (forall (vars (?x46 T2)) #2336)
+#2275 := (not #2270)
+#2293 := (and #1690 #2275 #2288)
+#2296 := (or #1002 #2293)
+#2299 := (forall (vars (?x37 T2)) #2296)
+#2658 := (and #109 #1051 #1098 #1104 #2299 #2339 #2649)
+#1646 := (not #1645)
+#2245 := (and #74 #1646 #2242)
+#2248 := (not #2245)
+#2251 := (forall (vars (?x38 T2)) #2248)
+#2217 := (not #2214)
+#1651 := (not #1650)
+#2257 := (and #1651 #2217 #2251)
+#2663 := (or #2257 #2658)
+#2666 := (and #987 #2663)
+#1619 := (not #1618)
+#2195 := (and #1619 #1620)
+#2198 := (not #2195)
+#2208 := (or #2198 #2205)
+#2211 := (not #2208)
+#2669 := (or #2211 #2666)
+#2672 := (and #961 #2669)
+#1590 := (not #1589)
+#2170 := (and #1588 #1590)
+#2173 := (not #2170)
+#2189 := (or #2173 #2184)
+#2192 := (not #2189)
+#2675 := (or #2192 #2672)
+#2678 := (and #946 #2675)
+#2681 := (or #1569 #2678)
+#2684 := (and #81 #2681)
+#2687 := (or #940 #2684)
+#3334 := (iff #2687 #3333)
+#3331 := (iff #2684 #3328)
+#3323 := (and #81 #3320)
+#3329 := (iff #3323 #3328)
+#3330 := [rewrite]: #3329
+#3324 := (iff #2684 #3323)
+#3321 := (iff #2681 #3320)
+#3318 := (iff #2678 #3315)
+#3310 := (and #946 #3307)
+#3316 := (iff #3310 #3315)
+#3317 := [rewrite]: #3316
+#3311 := (iff #2678 #3310)
+#3308 := (iff #2675 #3307)
+#3305 := (iff #2672 #3302)
+#3296 := (and #2706 #3293)
+#3303 := (iff #3296 #3302)
+#3304 := [rewrite]: #3303
+#3297 := (iff #2672 #3296)
+#3294 := (iff #2669 #3293)
+#3291 := (iff #2666 #3288)
+#3282 := (and #2751 #3279)
+#3289 := (iff #3282 #3288)
+#3290 := [rewrite]: #3289
+#3283 := (iff #2666 #3282)
+#3280 := (iff #2663 #3279)
+#3277 := (iff #2658 #3274)
+#3263 := (and #109 #2798 #2821 #1104 #2849 #2877 #3260)
+#3275 := (iff #3263 #3274)
+#3276 := [rewrite]: #3275
+#3264 := (iff #2658 #3263)
+#3261 := (iff #2649 #3260)
+#3258 := (iff #2644 #3255)
+#3246 := (and #184 #192 #3051 #3057 #1278 #1283 #1891 #1893 #3243)
+#3256 := (iff #3246 #3255)
+#3257 := [rewrite]: #3256
+#3247 := (iff #2644 #3246)
+#3244 := (iff #2638 #3243)
+#3241 := (iff #2635 #3238)
+#3233 := (and #1316 #3230)
+#3239 := (iff #3233 #3238)
+#3240 := [rewrite]: #3239
+#3234 := (iff #2635 #3233)
+#3231 := (iff #2632 #3230)
+#3228 := (iff #2629 #3225)
+#3220 := (and #705 #3217)
+#3226 := (iff #3220 #3225)
+#3227 := [rewrite]: #3226
+#3221 := (iff #2629 #3220)
+#3218 := (iff #2626 #3217)
+#3215 := (iff #2623 #3212)
+#3207 := (and #220 #3204)
+#3213 := (iff #3207 #3212)
+#3214 := [rewrite]: #3213
+#3208 := (iff #2623 #3207)
+#3205 := (iff #2620 #3204)
+#3202 := (iff #2617 #3199)
+#3194 := (and #1330 #3191)
+#3200 := (iff #3194 #3199)
+#3201 := [rewrite]: #3200
+#3195 := (iff #2617 #3194)
+#3192 := (iff #2614 #3191)
+#3189 := (iff #2611 #3186)
+#3180 := (and #3103 #3177)
+#3187 := (iff #3180 #3186)
+#3188 := [rewrite]: #3187
+#3181 := (iff #2611 #3180)
+#3178 := (iff #2608 #3177)
+#3175 := (iff #2603 #3172)
+#3166 := (and #3148 #2045 #2560 #3163)
+#3173 := (iff #3166 #3172)
+#3174 := [rewrite]: #3173
+#3167 := (iff #2603 #3166)
+#3164 := (iff #2594 #3163)
+#3161 := (iff #2591 #3152)
+#3153 := (not #3152)
+#3156 := (not #3153)
+#3159 := (iff #3156 #3152)
+#3160 := [rewrite]: #3159
+#3157 := (iff #2591 #3156)
+#3154 := (iff #2588 #3153)
+#3155 := [rewrite]: #3154
+#3158 := [monotonicity #3155]: #3157
+#3162 := [trans #3158 #3160]: #3161
+#3165 := [quant-intro #3162]: #3164
+#3149 := (iff #1368 #3148)
+#3146 := (iff #1365 #3143)
+#3129 := (or #223 #969)
+#3140 := (or #3129 #1361)
+#3144 := (iff #3140 #3143)
+#3145 := [rewrite]: #3144
+#3141 := (iff #1365 #3140)
+#3138 := (iff #1358 #3129)
+#3130 := (not #3129)
+#3133 := (not #3130)
+#3136 := (iff #3133 #3129)
+#3137 := [rewrite]: #3136
+#3134 := (iff #1358 #3133)
+#3131 := (iff #1353 #3130)
+#3132 := [rewrite]: #3131
+#3135 := [monotonicity #3132]: #3134
+#3139 := [trans #3135 #3137]: #3138
+#3142 := [monotonicity #3139]: #3141
+#3147 := [trans #3142 #3145]: #3146
+#3150 := [quant-intro #3147]: #3149
+#3168 := [monotonicity #3150 #3165]: #3167
+#3176 := [trans #3168 #3174]: #3175
+#3127 := (iff #2554 #3126)
+#3124 := (iff #2551 #3121)
+#3107 := (or #2013 #3106)
+#3118 := (or #3107 #2546)
+#3122 := (iff #3118 #3121)
+#3123 := [rewrite]: #3122
+#3119 := (iff #2551 #3118)
+#3116 := (iff #2533 #3107)
+#3108 := (not #3107)
+#3111 := (not #3108)
+#3114 := (iff #3111 #3107)
+#3115 := [rewrite]: #3114
+#3112 := (iff #2533 #3111)
+#3109 := (iff #2530 #3108)
+#3110 := [rewrite]: #3109
+#3113 := [monotonicity #3110]: #3112
+#3117 := [trans #3113 #3115]: #3116
+#3120 := [monotonicity #3117]: #3119
+#3125 := [trans #3120 #3123]: #3124
+#3128 := [monotonicity #3125]: #3127
+#3179 := [monotonicity #3128 #3176]: #3178
+#3104 := (iff #1344 #3103)
+#3101 := (iff #1341 #3098)
+#3084 := (or #216 #3083)
+#3095 := (or #3084 #1336)
+#3099 := (iff #3095 #3098)
+#3100 := [rewrite]: #3099
+#3096 := (iff #1341 #3095)
+#3093 := (iff #711 #3084)
+#3085 := (not #3084)
+#3088 := (not #3085)
+#3091 := (iff #3088 #3084)
+#3092 := [rewrite]: #3091
+#3089 := (iff #711 #3088)
+#3086 := (iff #225 #3085)
+#3087 := [rewrite]: #3086
+#3090 := [monotonicity #3087]: #3089
+#3094 := [trans #3090 #3092]: #3093
+#3097 := [monotonicity #3094]: #3096
+#3102 := [trans #3097 #3100]: #3101
+#3105 := [quant-intro #3102]: #3104
+#3182 := [monotonicity #3105 #3179]: #3181
+#3190 := [trans #3182 #3188]: #3189
+#3081 := (iff #2527 #3080)
+#3078 := (iff #2524 #3075)
+#3061 := (or #3060 #1984)
+#3072 := (or #3061 #2519)
+#3076 := (iff #3072 #3075)
+#3077 := [rewrite]: #3076
+#3073 := (iff #2524 #3072)
+#3070 := (iff #2508 #3061)
+#3062 := (not #3061)
+#3065 := (not #3062)
+#3068 := (iff #3065 #3061)
+#3069 := [rewrite]: #3068
+#3066 := (iff #2508 #3065)
+#3063 := (iff #2505 #3062)
+#3064 := [rewrite]: #3063
+#3067 := [monotonicity #3064]: #3066
+#3071 := [trans #3067 #3069]: #3070
+#3074 := [monotonicity #3071]: #3073
+#3079 := [trans #3074 #3077]: #3078
+#3082 := [monotonicity #3079]: #3081
+#3193 := [monotonicity #3082 #3190]: #3192
+#3196 := [monotonicity #3193]: #3195
+#3203 := [trans #3196 #3201]: #3202
+#3206 := [monotonicity #3203]: #3205
+#3209 := [monotonicity #3206]: #3208
+#3216 := [trans #3209 #3214]: #3215
+#3219 := [monotonicity #3216]: #3218
+#3222 := [monotonicity #3219]: #3221
+#3229 := [trans #3222 #3227]: #3228
+#3232 := [monotonicity #3229]: #3231
+#3235 := [monotonicity #3232]: #3234
+#3242 := [trans #3235 #3240]: #3241
+#3245 := [monotonicity #3242]: #3244
+#3058 := (iff #1268 #3057)
+#3055 := (iff #1263 #3054)
+#3034 := (iff #1242 #3033)
+#3035 := [rewrite]: #3034
+#3056 := [monotonicity #3035]: #3055
+#3059 := [quant-intro #3056]: #3058
+#3052 := (iff #1257 #3051)
+#3049 := (iff #1254 #3046)
+#3043 := (or #3032 #1251)
#3047 := (iff #3043 #3046)
#3048 := [rewrite]: #3047
-#3044 := (iff #1605 #3043)
-#3041 := (iff #1599 #3032)
-#3033 := (not #3032)
+#3044 := (iff #1254 #3043)
+#3041 := (iff #1245 #3032)
#3036 := (not #3033)
#3039 := (iff #3036 #3032)
#3040 := [rewrite]: #3039
-#3037 := (iff #1599 #3036)
-#3034 := (iff #1594 #3033)
-#3035 := [rewrite]: #3034
+#3037 := (iff #1245 #3036)
#3038 := [monotonicity #3035]: #3037
#3042 := [trans #3038 #3040]: #3041
#3045 := [monotonicity #3042]: #3044
#3050 := [trans #3045 #3048]: #3049
#3053 := [quant-intro #3050]: #3052
-#3512 := [monotonicity #3053 #3076 #3104 #3132 #3509]: #3511
-#3523 := [trans #3512 #3521]: #3522
-#3030 := (iff #2498 #3027)
-#3022 := (and #1878 #2458 #3019)
-#3028 := (iff #3022 #3027)
+#3248 := [monotonicity #3053 #3059 #3245]: #3247
+#3259 := [trans #3248 #3257]: #3258
+#3030 := (iff #2473 #3027)
+#3017 := (and #149 #154 #530 #533 #2891 #3014)
+#3028 := (iff #3017 #3027)
#3029 := [rewrite]: #3028
-#3023 := (iff #2498 #3022)
-#3020 := (iff #2492 #3019)
-#3017 := (iff #2489 #3008)
-#3009 := (not #3008)
-#3012 := (not #3009)
-#3015 := (iff #3012 #3008)
-#3016 := [rewrite]: #3015
-#3013 := (iff #2489 #3012)
-#3010 := (iff #2486 #3009)
+#3018 := (iff #2473 #3017)
+#3015 := (iff #2467 #3014)
+#3012 := (iff #2464 #3009)
+#3003 := (and #2943 #3000)
+#3010 := (iff #3003 #3009)
#3011 := [rewrite]: #3010
-#3014 := [monotonicity #3011]: #3013
-#3018 := [trans #3014 #3016]: #3017
-#3021 := [quant-intro #3018]: #3020
-#3024 := [monotonicity #3021]: #3023
-#3031 := [trans #3024 #3029]: #3030
-#3526 := [monotonicity #3031 #3523]: #3525
-#3005 := (iff #1162 #3004)
-#3002 := (iff #1159 #2999)
-#2985 := (or #75 #1144)
-#2996 := (or #2985 #1155)
-#3000 := (iff #2996 #2999)
-#3001 := [rewrite]: #3000
-#2997 := (iff #1159 #2996)
-#2994 := (iff #1151 #2985)
-#2986 := (not #2985)
-#2989 := (not #2986)
-#2992 := (iff #2989 #2985)
-#2993 := [rewrite]: #2992
-#2990 := (iff #1151 #2989)
-#2987 := (iff #1148 #2986)
-#2988 := [rewrite]: #2987
-#2991 := [monotonicity #2988]: #2990
-#2995 := [trans #2991 #2993]: #2994
-#2998 := [monotonicity #2995]: #2997
-#3003 := [trans #2998 #3001]: #3002
-#3006 := [quant-intro #3003]: #3005
-#3529 := [monotonicity #3006 #3526]: #3528
-#3537 := [trans #3529 #3535]: #3536
-#2983 := (iff #2452 #2982)
-#2980 := (iff #2449 #2977)
-#2963 := (or #1845 #2962)
-#2974 := (or #2963 #2446)
-#2978 := (iff #2974 #2977)
-#2979 := [rewrite]: #2978
-#2975 := (iff #2449 #2974)
-#2972 := (iff #2439 #2963)
-#2964 := (not #2963)
-#2967 := (not #2964)
-#2970 := (iff #2967 #2963)
-#2971 := [rewrite]: #2970
-#2968 := (iff #2439 #2967)
-#2965 := (iff #2436 #2964)
-#2966 := [rewrite]: #2965
-#2969 := [monotonicity #2966]: #2968
-#2973 := [trans #2969 #2971]: #2972
-#2976 := [monotonicity #2973]: #2975
-#2981 := [trans #2976 #2979]: #2980
+#3004 := (iff #2464 #3003)
+#3001 := (iff #2459 #3000)
+#2998 := (iff #2453 #2997)
+#2995 := (iff #2450 #2992)
+#2978 := (or #1838 #1842)
+#2989 := (or #2978 #2445)
+#2993 := (iff #2989 #2992)
+#2994 := [rewrite]: #2993
+#2990 := (iff #2450 #2989)
+#2987 := (iff #2433 #2978)
+#2979 := (not #2978)
+#2982 := (not #2979)
+#2985 := (iff #2982 #2978)
+#2986 := [rewrite]: #2985
+#2983 := (iff #2433 #2982)
+#2980 := (iff #2430 #2979)
+#2981 := [rewrite]: #2980
#2984 := [monotonicity #2981]: #2983
-#3540 := [monotonicity #2984 #3537]: #3539
-#2960 := (iff #1136 #2959)
-#2957 := (iff #1133 #2954)
-#2940 := (or #74 #2408)
-#2951 := (or #2940 #1129)
-#2955 := (iff #2951 #2954)
-#2956 := [rewrite]: #2955
-#2952 := (iff #1133 #2951)
-#2949 := (iff #430 #2940)
-#2941 := (not #2940)
-#2944 := (not #2941)
-#2947 := (iff #2944 #2940)
-#2948 := [rewrite]: #2947
-#2945 := (iff #430 #2944)
-#2942 := (iff #82 #2941)
-#2943 := [rewrite]: #2942
-#2946 := [monotonicity #2943]: #2945
-#2950 := [trans #2946 #2948]: #2949
-#2953 := [monotonicity #2950]: #2952
-#2958 := [trans #2953 #2956]: #2957
-#2961 := [quant-intro #2958]: #2960
-#3543 := [monotonicity #2961 #3540]: #3542
-#3551 := [trans #3543 #3549]: #3550
-#2111 := (iff #2433 #1948)
-#1825 := (iff #2430 #2132)
-#1799 := (or #2058 #1816)
-#2197 := (or #1799 #2425)
-#2133 := (iff #2197 #2132)
-#1824 := [rewrite]: #2133
-#2171 := (iff #2430 #2197)
-#1855 := (iff #2414 #1799)
-#1755 := (not #1799)
-#1921 := (not #1755)
-#2152 := (iff #1921 #1799)
-#1854 := [rewrite]: #2152
-#1922 := (iff #2414 #1921)
-#1756 := (iff #2411 #1755)
-#1800 := [rewrite]: #1756
-#2151 := [monotonicity #1800]: #1922
-#2196 := [trans #2151 #1854]: #1855
-#2172 := [monotonicity #2196]: #2171
-#1947 := [trans #2172 #1824]: #1825
-#2112 := [monotonicity #1947]: #2111
-#3554 := [monotonicity #2112 #3551]: #3553
-#3557 := [monotonicity #3554]: #3556
-#3564 := [trans #3557 #3562]: #3563
-#3567 := [monotonicity #3564]: #3566
-#3570 := [monotonicity #3567]: #3569
-#3577 := [trans #3570 #3575]: #3576
-#3580 := [monotonicity #3577]: #3579
-#2241 := (+ #2240 #2238)
-#2242 := (= #2241 0::int)
-#2245 := (and #206 #2244 #2242)
-#2262 := (not #2245)
-#2265 := (forall (vars (?x76 T2)) #2262)
-#2250 := (= ?x75!20 uf_11)
-#2251 := (not #2250)
-#2252 := (and #2251 #2249)
-#2253 := (not #2252)
-#2259 := (not #2253)
-#2269 := (and #2259 #2265)
-#2274 := (and #1438 #2269)
-#2208 := (* -1::int #2207)
-#2210 := (+ #2209 #2208)
-#2213 := (+ #2212 #2210)
-#2214 := (>= #2213 0::int)
-#2220 := (and #2219 #2218)
-#2221 := (not #2220)
-#2222 := (or #2221 #2214)
-#2223 := (not #2222)
-#2278 := (or #2223 #2274)
-#2282 := (and #1414 #2278)
-#2183 := (* -1::int #2182)
-#2185 := (+ #2184 #2183)
-#2186 := (>= #2185 0::int)
-#2190 := (and #2189 #2187)
-#2191 := (not #2190)
-#2192 := (or #2191 #2186)
-#2193 := (not #2192)
-#2286 := (or #2193 #2282)
-#2290 := (and #1400 #2286)
-#2294 := (or #2168 #2290)
-#2162 := (not #1394)
-#2298 := (and #2162 #2294)
-#2302 := (or #1394 #2298)
-#2306 := (and #710 #2302)
-#2144 := (= #2143 #2142)
-#2147 := (or #2146 #2144)
-#2148 := (not #2147)
-#2310 := (or #2148 #2306)
-#2314 := (and #1386 #2310)
-#2125 := (* -1::int #2124)
-#2127 := (+ #2126 #2125)
-#2128 := (>= #2127 0::int)
-#2129 := (not #2128)
-#2318 := (or #2129 #2314)
-#2108 := (and #2107 #2105)
-#2097 := (not #888)
-#2338 := (and #181 #2097 #2108 #1375 #2318 #1528 #1541 #1549)
-#2036 := (* -1::int #2035)
-#2038 := (+ #2037 #2036)
-#2041 := (+ #2040 #2038)
-#2042 := (>= #2041 0::int)
-#2051 := (and #2050 #2046)
-#2052 := (not #2051)
-#2053 := (or #2052 #2042)
-#2054 := (not #2053)
-#2073 := (or #2054 #2069)
-#2012 := (+ #2011 #1251)
-#2015 := (+ #2014 #2012)
-#2016 := (= #2015 0::int)
-#2017 := (>= #2012 0::int)
-#2018 := (not #2017)
-#2019 := (and #2018 #2016)
-#2024 := (or #1260 #2019)
-#2027 := (forall (vars (?x49 T2)) #2024)
-#2077 := (and #2027 #2073)
-#1976 := (+ #1975 #1973)
-#1977 := (= #1976 0::int)
-#1980 := (and #1979 #1977)
-#1996 := (not #1980)
-#1999 := (forall (vars (?x50 T2)) #1996)
-#1985 := (= ?x49!8 uf_11)
-#1986 := (not #1985)
-#1987 := (and #1986 #1984)
-#1988 := (not #1987)
-#1993 := (not #1988)
-#2003 := (and #1993 #1999)
-#2081 := (or #2003 #2077)
-#1967 := (not #614)
-#1964 := (not #632)
-#1961 := (not #605)
-#1958 := (not #623)
-#2091 := (and #1958 #1961 #1964 #1967 #2081 #2088)
-#2342 := (or #2091 #2338)
-#1936 := (+ #1935 #1206)
-#1937 := (>= #1936 0::int)
-#1938 := (not #1937)
-#1941 := (+ #1940 #1936)
-#1942 := (= #1941 0::int)
-#1944 := (and #1943 #1942 #1938)
-#1949 := (or #1215 #1944)
-#1952 := (forall (vars (?x46 T2)) #1949)
-#1910 := (+ #1168 #1909)
-#1912 := (+ #1911 #1910)
-#1913 := (= #1912 0::int)
-#1914 := (+ #1911 #1168)
-#1915 := (>= #1914 0::int)
-#1916 := (not #1915)
-#1918 := (and #1917 #1916 #1913)
-#1923 := (or #1177 #1918)
-#1926 := (forall (vars (?x37 T2)) #1923)
-#1902 := (not #1636)
-#2367 := (and #1902 #1926 #1952 #2342 #1608 #1619 #1628)
-#1868 := (+ #1867 #1865)
-#1869 := (+ #66 #1868)
-#1870 := (= #1869 0::int)
-#1874 := (and #74 #1873 #1870)
-#1890 := (not #1874)
-#1893 := (forall (vars (?x38 T2)) #1890)
-#1879 := (= ?x37!5 uf_11)
-#1880 := (not #1879)
-#1881 := (and #1880 #1878)
-#1882 := (not #1881)
-#1887 := (not #1882)
-#1897 := (and #1887 #1893)
-#2371 := (or #1897 #2367)
-#2375 := (and #1162 #2371)
-#1839 := (+ #1838 #1836)
-#1841 := (+ #1840 #1839)
-#1842 := (>= #1841 0::int)
-#1848 := (and #1847 #1846)
-#1849 := (not #1848)
-#1850 := (or #1849 #1842)
-#1851 := (not #1850)
-#2379 := (or #1851 #2375)
-#2383 := (and #1136 #2379)
-#1811 := (* -1::int #1810)
-#1813 := (+ #1812 #1811)
-#1814 := (>= #1813 0::int)
-#1818 := (and #1817 #1815)
-#1819 := (not #1818)
-#1820 := (or #1819 #1814)
-#1821 := (not #1820)
-#2387 := (or #1821 #2383)
-#2391 := (and #1121 #2387)
-#2395 := (or #1796 #2391)
-#1751 := (not #1115)
-#2399 := (and #1751 #2395)
-#2403 := (or #1115 #2399)
-#2938 := (iff #2403 #2937)
-#2935 := (iff #2399 #2934)
-#2932 := (iff #2395 #2931)
-#2929 := (iff #2391 #2928)
-#2926 := (iff #2387 #2925)
-#2923 := (iff #2383 #2922)
-#2920 := (iff #2379 #2919)
-#2917 := (iff #2375 #2916)
-#2914 := (iff #2371 #2913)
-#2911 := (iff #2367 #2908)
-#2905 := (and #106 #2542 #2582 #2902 #1608 #1619 #1628)
-#2909 := (iff #2905 #2908)
-#2910 := [rewrite]: #2909
-#2906 := (iff #2367 #2905)
-#2903 := (iff #2342 #2902)
-#2900 := (iff #2338 #2897)
-#2894 := (and #181 #189 #2108 #1375 #2891 #1528 #1541 #1549)
-#2898 := (iff #2894 #2897)
-#2899 := [rewrite]: #2898
-#2895 := (iff #2338 #2894)
-#2892 := (iff #2318 #2891)
-#2889 := (iff #2314 #2888)
-#2886 := (iff #2310 #2885)
-#2883 := (iff #2306 #2882)
-#2880 := (iff #2302 #2879)
-#2877 := (iff #2298 #2876)
-#2874 := (iff #2294 #2873)
-#2871 := (iff #2290 #2870)
-#2868 := (iff #2286 #2867)
-#2865 := (iff #2282 #2864)
-#2862 := (iff #2278 #2861)
-#2859 := (iff #2274 #2856)
-#2819 := (and #2249 #2813)
-#2850 := (and #2819 #2847)
-#2853 := (and #1438 #2850)
-#2857 := (iff #2853 #2856)
-#2858 := [rewrite]: #2857
-#2854 := (iff #2274 #2853)
-#2851 := (iff #2269 #2850)
-#2848 := (iff #2265 #2847)
-#2845 := (iff #2262 #2844)
-#2842 := (iff #2245 #2841)
-#2839 := (iff #2242 #2838)
-#2836 := (= #2241 #2835)
-#2837 := [rewrite]: #2836
-#2840 := [monotonicity #2837]: #2839
-#2843 := [monotonicity #2840]: #2842
-#2846 := [monotonicity #2843]: #2845
-#2849 := [quant-intro #2846]: #2848
-#2832 := (iff #2259 #2819)
-#2824 := (not #2819)
-#2827 := (not #2824)
-#2830 := (iff #2827 #2819)
-#2831 := [rewrite]: #2830
-#2828 := (iff #2259 #2827)
-#2825 := (iff #2253 #2824)
-#2822 := (iff #2252 #2819)
-#2816 := (and #2813 #2249)
-#2820 := (iff #2816 #2819)
-#2821 := [rewrite]: #2820
-#2817 := (iff #2252 #2816)
-#2814 := (iff #2251 #2813)
-#2811 := (iff #2250 #2810)
-#2812 := [rewrite]: #2811
+#2988 := [trans #2984 #2986]: #2987
+#2991 := [monotonicity #2988]: #2990
+#2996 := [trans #2991 #2994]: #2995
+#2999 := [monotonicity #2996]: #2998
+#2976 := (iff #1862 #2973)
+#2968 := (and #1852 #2965)
+#2974 := (iff #2968 #2973)
+#2975 := [rewrite]: #2974
+#2969 := (iff #1862 #2968)
+#2966 := (iff #1201 #2965)
+#2963 := (iff #1198 #2960)
+#2946 := (or #969 #1154)
+#2957 := (or #2946 #1195)
+#2961 := (iff #2957 #2960)
+#2962 := [rewrite]: #2961
+#2958 := (iff #1198 #2957)
+#2955 := (iff #1192 #2946)
+#2947 := (not #2946)
+#2950 := (not #2947)
+#2953 := (iff #2950 #2946)
+#2954 := [rewrite]: #2953
+#2951 := (iff #1192 #2950)
+#2948 := (iff #1189 #2947)
+#2949 := [rewrite]: #2948
+#2952 := [monotonicity #2949]: #2951
+#2956 := [trans #2952 #2954]: #2955
+#2959 := [monotonicity #2956]: #2958
+#2964 := [trans #2959 #2962]: #2963
+#2967 := [quant-intro #2964]: #2966
+#2970 := [monotonicity #2967]: #2969
+#2977 := [trans #2970 #2975]: #2976
+#3002 := [monotonicity #2977 #2999]: #3001
+#2944 := (iff #2427 #2943)
+#2941 := (iff #2424 #2938)
+#2919 := (or #65 #1154)
+#2935 := (or #2919 #2932)
+#2939 := (iff #2935 #2938)
+#2940 := [rewrite]: #2939
+#2936 := (iff #2424 #2935)
+#2933 := (iff #2421 #2932)
+#2934 := [rewrite]: #2933
+#2928 := (iff #1161 #2919)
+#2920 := (not #2919)
+#2923 := (not #2920)
+#2926 := (iff #2923 #2919)
+#2927 := [rewrite]: #2926
+#2924 := (iff #1161 #2923)
+#2921 := (iff #1158 #2920)
+#2922 := [rewrite]: #2921
+#2925 := [monotonicity #2922]: #2924
+#2929 := [trans #2925 #2927]: #2928
+#2937 := [monotonicity #2929 #2934]: #2936
+#2942 := [trans #2937 #2940]: #2941
+#2945 := [quant-intro #2942]: #2944
+#3005 := [monotonicity #2945 #3002]: #3004
+#3013 := [trans #3005 #3011]: #3012
+#2917 := (iff #2385 #2914)
+#2909 := (and #1777 #2345 #2906)
+#2915 := (iff #2909 #2914)
+#2916 := [rewrite]: #2915
+#2910 := (iff #2385 #2909)
+#2907 := (iff #2379 #2906)
+#2904 := (iff #2376 #2895)
+#2896 := (not #2895)
+#2899 := (not #2896)
+#2902 := (iff #2899 #2895)
+#2903 := [rewrite]: #2902
+#2900 := (iff #2376 #2899)
+#2897 := (iff #2373 #2896)
+#2898 := [rewrite]: #2897
+#2901 := [monotonicity #2898]: #2900
+#2905 := [trans #2901 #2903]: #2904
+#2908 := [quant-intro #2905]: #2907
+#2911 := [monotonicity #2908]: #2910
+#2918 := [trans #2911 #2916]: #2917
+#3016 := [monotonicity #2918 #3013]: #3015
+#2892 := (iff #1754 #2891)
+#2889 := (iff #1751 #2880)
+#2881 := (not #2880)
+#2884 := (not #2881)
+#2887 := (iff #2884 #2880)
+#2888 := [rewrite]: #2887
+#2885 := (iff #1751 #2884)
+#2882 := (iff #1132 #2881)
+#2883 := [rewrite]: #2882
+#2886 := [monotonicity #2883]: #2885
+#2890 := [trans #2886 #2888]: #2889
+#2893 := [quant-intro #2890]: #2892
+#3019 := [monotonicity #2893 #3016]: #3018
+#3031 := [trans #3019 #3029]: #3030
+#3262 := [monotonicity #3031 #3259]: #3261
+#2878 := (iff #2339 #2877)
+#2875 := (iff #2336 #2872)
+#2852 := (or #65 #1056)
+#2869 := (or #2852 #2866)
+#2873 := (iff #2869 #2872)
+#2874 := [rewrite]: #2873
+#2870 := (iff #2336 #2869)
+#2867 := (iff #2333 #2866)
+#2868 := [rewrite]: #2867
+#2861 := (iff #1063 #2852)
+#2853 := (not #2852)
+#2856 := (not #2853)
+#2859 := (iff #2856 #2852)
+#2860 := [rewrite]: #2859
+#2857 := (iff #1063 #2856)
+#2854 := (iff #1060 #2853)
+#2855 := [rewrite]: #2854
+#2858 := [monotonicity #2855]: #2857
+#2862 := [trans #2858 #2860]: #2861
+#2871 := [monotonicity #2862 #2868]: #2870
+#2876 := [trans #2871 #2874]: #2875
+#2879 := [quant-intro #2876]: #2878
+#2850 := (iff #2299 #2849)
+#2847 := (iff #2296 #2844)
+#2824 := (or #65 #995)
+#2841 := (or #2824 #2838)
+#2845 := (iff #2841 #2844)
+#2846 := [rewrite]: #2845
+#2842 := (iff #2296 #2841)
+#2839 := (iff #2293 #2838)
+#2840 := [rewrite]: #2839
+#2833 := (iff #1002 #2824)
+#2825 := (not #2824)
+#2828 := (not #2825)
+#2831 := (iff #2828 #2824)
+#2832 := [rewrite]: #2831
+#2829 := (iff #1002 #2828)
+#2826 := (iff #999 #2825)
+#2827 := [rewrite]: #2826
+#2830 := [monotonicity #2827]: #2829
+#2834 := [trans #2830 #2832]: #2833
+#2843 := [monotonicity #2834 #2840]: #2842
+#2848 := [trans #2843 #2846]: #2847
+#2851 := [quant-intro #2848]: #2850
+#2822 := (iff #1098 #2821)
+#2819 := (iff #1095 #2816)
+#2802 := (or #114 #2801)
+#2813 := (or #2802 #1069)
+#2817 := (iff #2813 #2816)
+#2818 := [rewrite]: #2817
+#2814 := (iff #1095 #2813)
+#2811 := (iff #456 #2802)
+#2803 := (not #2802)
+#2806 := (not #2803)
+#2809 := (iff #2806 #2802)
+#2810 := [rewrite]: #2809
+#2807 := (iff #456 #2806)
+#2804 := (iff #117 #2803)
+#2805 := [rewrite]: #2804
+#2808 := [monotonicity #2805]: #2807
+#2812 := [trans #2808 #2810]: #2811
#2815 := [monotonicity #2812]: #2814
-#2818 := [monotonicity #2815]: #2817
-#2823 := [trans #2818 #2821]: #2822
-#2826 := [monotonicity #2823]: #2825
-#2829 := [monotonicity #2826]: #2828
-#2833 := [trans #2829 #2831]: #2832
-#2852 := [monotonicity #2833 #2849]: #2851
-#2855 := [monotonicity #2852]: #2854
-#2860 := [trans #2855 #2858]: #2859
-#2808 := (iff #2223 #2807)
-#2805 := (iff #2222 #2804)
-#2802 := (iff #2214 #2799)
-#2789 := (+ #2209 #2212)
-#2790 := (+ #2208 #2789)
-#2793 := (>= #2790 0::int)
-#2800 := (iff #2793 #2799)
-#2801 := [rewrite]: #2800
-#2794 := (iff #2214 #2793)
-#2791 := (= #2213 #2790)
-#2792 := [rewrite]: #2791
-#2795 := [monotonicity #2792]: #2794
-#2803 := [trans #2795 #2801]: #2802
-#2787 := (iff #2221 #2786)
-#2784 := (iff #2220 #2783)
-#2785 := [rewrite]: #2784
-#2788 := [monotonicity #2785]: #2787
-#2806 := [monotonicity #2788 #2803]: #2805
-#2809 := [monotonicity #2806]: #2808
-#2863 := [monotonicity #2809 #2860]: #2862
-#2866 := [monotonicity #2863]: #2865
-#2781 := (iff #2193 #2780)
-#2778 := (iff #2192 #2777)
-#2775 := (iff #2186 #2772)
-#2764 := (+ #2183 #2184)
-#2767 := (>= #2764 0::int)
-#2773 := (iff #2767 #2772)
-#2774 := [rewrite]: #2773
-#2768 := (iff #2186 #2767)
-#2765 := (= #2185 #2764)
-#2766 := [rewrite]: #2765
-#2769 := [monotonicity #2766]: #2768
-#2776 := [trans #2769 #2774]: #2775
-#2762 := (iff #2191 #2761)
-#2759 := (iff #2190 #2758)
-#2760 := [rewrite]: #2759
-#2763 := [monotonicity #2760]: #2762
-#2779 := [monotonicity #2763 #2776]: #2778
-#2782 := [monotonicity #2779]: #2781
-#2869 := [monotonicity #2782 #2866]: #2868
-#2872 := [monotonicity #2869]: #2871
-#2875 := [monotonicity #2872]: #2874
-#2756 := (iff #2162 #210)
-#2757 := [rewrite]: #2756
-#2878 := [monotonicity #2757 #2875]: #2877
-#2881 := [monotonicity #2878]: #2880
-#2884 := [monotonicity #2881]: #2883
-#2754 := (iff #2148 #2753)
-#2751 := (iff #2147 #2750)
-#2748 := (iff #2144 #2747)
-#2749 := [rewrite]: #2748
-#2752 := [monotonicity #2749]: #2751
-#2755 := [monotonicity #2752]: #2754
-#2887 := [monotonicity #2755 #2884]: #2886
-#2890 := [monotonicity #2887]: #2889
-#2745 := (iff #2129 #2744)
-#2742 := (iff #2128 #2739)
-#2731 := (+ #2125 #2126)
-#2734 := (>= #2731 0::int)
-#2740 := (iff #2734 #2739)
-#2741 := [rewrite]: #2740
-#2735 := (iff #2128 #2734)
-#2732 := (= #2127 #2731)
-#2733 := [rewrite]: #2732
-#2736 := [monotonicity #2733]: #2735
-#2743 := [trans #2736 #2741]: #2742
-#2746 := [monotonicity #2743]: #2745
-#2893 := [monotonicity #2746 #2890]: #2892
-#2729 := (iff #2097 #189)
-#2730 := [rewrite]: #2729
-#2896 := [monotonicity #2730 #2893]: #2895
-#2901 := [trans #2896 #2899]: #2900
-#2727 := (iff #2091 #2724)
-#2721 := (and #140 #145 #506 #509 #2718 #2088)
+#2820 := [trans #2815 #2818]: #2819
+#2823 := [quant-intro #2820]: #2822
+#2799 := (iff #1051 #2798)
+#2796 := (iff #1048 #2793)
+#2779 := (or #115 #969)
+#2790 := (or #2779 #1043)
+#2794 := (iff #2790 #2793)
+#2795 := [rewrite]: #2794
+#2791 := (iff #1048 #2790)
+#2788 := (iff #1039 #2779)
+#2780 := (not #2779)
+#2783 := (not #2780)
+#2786 := (iff #2783 #2779)
+#2787 := [rewrite]: #2786
+#2784 := (iff #1039 #2783)
+#2781 := (iff #1034 #2780)
+#2782 := [rewrite]: #2781
+#2785 := [monotonicity #2782]: #2784
+#2789 := [trans #2785 #2787]: #2788
+#2792 := [monotonicity #2789]: #2791
+#2797 := [trans #2792 #2795]: #2796
+#2800 := [quant-intro #2797]: #2799
+#3265 := [monotonicity #2800 #2823 #2851 #2879 #3262]: #3264
+#3278 := [trans #3265 #3276]: #3277
+#2777 := (iff #2257 #2774)
+#2769 := (and #1651 #2217 #2766)
+#2775 := (iff #2769 #2774)
+#2776 := [rewrite]: #2775
+#2770 := (iff #2257 #2769)
+#2767 := (iff #2251 #2766)
+#2764 := (iff #2248 #2755)
+#2756 := (not #2755)
+#2759 := (not #2756)
+#2762 := (iff #2759 #2755)
+#2763 := [rewrite]: #2762
+#2760 := (iff #2248 #2759)
+#2757 := (iff #2245 #2756)
+#2758 := [rewrite]: #2757
+#2761 := [monotonicity #2758]: #2760
+#2765 := [trans #2761 #2763]: #2764
+#2768 := [quant-intro #2765]: #2767
+#2771 := [monotonicity #2768]: #2770
+#2778 := [trans #2771 #2776]: #2777
+#3281 := [monotonicity #2778 #3278]: #3280
+#2752 := (iff #987 #2751)
+#2749 := (iff #984 #2746)
+#2732 := (or #75 #969)
+#2743 := (or #2732 #980)
+#2747 := (iff #2743 #2746)
+#2748 := [rewrite]: #2747
+#2744 := (iff #984 #2743)
+#2741 := (iff #976 #2732)
+#2733 := (not #2732)
+#2736 := (not #2733)
+#2739 := (iff #2736 #2732)
+#2740 := [rewrite]: #2739
+#2737 := (iff #976 #2736)
+#2734 := (iff #973 #2733)
+#2735 := [rewrite]: #2734
+#2738 := [monotonicity #2735]: #2737
+#2742 := [trans #2738 #2740]: #2741
+#2745 := [monotonicity #2742]: #2744
+#2750 := [trans #2745 #2748]: #2749
+#2753 := [quant-intro #2750]: #2752
+#3284 := [monotonicity #2753 #3281]: #3283
+#3292 := [trans #3284 #3290]: #3291
+#2730 := (iff #2211 #2729)
+#2727 := (iff #2208 #2724)
+#2710 := (or #1618 #2709)
+#2721 := (or #2710 #2205)
#2725 := (iff #2721 #2724)
#2726 := [rewrite]: #2725
-#2722 := (iff #2091 #2721)
-#2719 := (iff #2081 #2718)
-#2716 := (iff #2077 #2715)
-#2713 := (iff #2073 #2710)
-#2707 := (or #2704 #2069)
-#2711 := (iff #2707 #2710)
-#2712 := [rewrite]: #2711
-#2708 := (iff #2073 #2707)
-#2705 := (iff #2054 #2704)
-#2702 := (iff #2053 #2701)
-#2699 := (iff #2042 #2696)
-#2687 := (+ #2037 #2040)
-#2688 := (+ #2036 #2687)
-#2691 := (>= #2688 0::int)
-#2697 := (iff #2691 #2696)
-#2698 := [rewrite]: #2697
-#2692 := (iff #2042 #2691)
-#2689 := (= #2041 #2688)
-#2690 := [rewrite]: #2689
+#2722 := (iff #2208 #2721)
+#2719 := (iff #2198 #2710)
+#2711 := (not #2710)
+#2714 := (not #2711)
+#2717 := (iff #2714 #2710)
+#2718 := [rewrite]: #2717
+#2715 := (iff #2198 #2714)
+#2712 := (iff #2195 #2711)
+#2713 := [rewrite]: #2712
+#2716 := [monotonicity #2713]: #2715
+#2720 := [trans #2716 #2718]: #2719
+#2723 := [monotonicity #2720]: #2722
+#2728 := [trans #2723 #2726]: #2727
+#2731 := [monotonicity #2728]: #2730
+#3295 := [monotonicity #2731 #3292]: #3294
+#2707 := (iff #961 #2706)
+#2704 := (iff #958 #2701)
+#1721 := (or #74 #1898)
+#2698 := (or #1721 #954)
+#2702 := (iff #2698 #2701)
+#2703 := [rewrite]: #2702
+#2699 := (iff #958 #2698)
+#2696 := (iff #432 #1721)
+#1722 := (not #1721)
+#2691 := (not #1722)
+#2694 := (iff #2691 #1721)
+#2695 := [rewrite]: #2694
+#2692 := (iff #432 #2691)
+#2167 := (iff #85 #1722)
+#2690 := [rewrite]: #2167
#2693 := [monotonicity #2690]: #2692
-#2700 := [trans #2693 #2698]: #2699
-#2685 := (iff #2052 #2684)
-#2682 := (iff #2051 #2681)
-#2683 := [rewrite]: #2682
-#2686 := [monotonicity #2683]: #2685
-#2703 := [monotonicity #2686 #2700]: #2702
-#2706 := [monotonicity #2703]: #2705
-#2709 := [monotonicity #2706]: #2708
-#2714 := [trans #2709 #2712]: #2713
-#2679 := (iff #2027 #2678)
-#2676 := (iff #2024 #2675)
-#2673 := (iff #2019 #2672)
-#2670 := (iff #2016 #2667)
-#2657 := (+ #2011 #2014)
-#2658 := (+ #1251 #2657)
-#2661 := (= #2658 0::int)
-#2668 := (iff #2661 #2667)
-#2669 := [rewrite]: #2668
-#2662 := (iff #2016 #2661)
-#2659 := (= #2015 #2658)
+#2697 := [trans #2693 #2695]: #2696
+#2700 := [monotonicity #2697]: #2699
+#2705 := [trans #2700 #2703]: #2704
+#2708 := [quant-intro #2705]: #2707
+#3298 := [monotonicity #2708 #3295]: #3297
+#3306 := [trans #3298 #3304]: #3305
+#2023 := (iff #2192 #2022)
+#1694 := (iff #2189 #1948)
+#1598 := (or #1597 #1589)
+#1572 := (or #1598 #2184)
+#1850 := (iff #1572 #1948)
+#1851 := [rewrite]: #1850
+#1573 := (iff #2189 #1572)
+#1992 := (iff #2173 #1598)
+#1967 := (not #1598)
+#1929 := (not #1967)
+#1815 := (iff #1929 #1598)
+#1816 := [rewrite]: #1815
+#1627 := (iff #2173 #1929)
+#1968 := (iff #2170 #1967)
+#1928 := [rewrite]: #1968
+#1628 := [monotonicity #1928]: #1627
+#1993 := [trans #1628 #1816]: #1992
+#1947 := [monotonicity #1993]: #1573
+#1695 := [trans #1947 #1851]: #1694
+#1897 := [monotonicity #1695]: #2023
+#3309 := [monotonicity #1897 #3306]: #3308
+#3312 := [monotonicity #3309]: #3311
+#3319 := [trans #3312 #3317]: #3318
+#3322 := [monotonicity #3319]: #3321
+#3325 := [monotonicity #3322]: #3324
+#3332 := [trans #3325 #3330]: #3331
+#3335 := [monotonicity #3332]: #3334
+#2037 := (+ #2036 #2034)
+#2038 := (= #2037 0::int)
+#2041 := (and #216 #2040 #2038)
+#2058 := (not #2041)
+#2061 := (forall (vars (?x76 T2)) #2058)
+#2046 := (= ?x75!20 uf_11)
+#2047 := (not #2046)
+#2048 := (and #2047 #2045)
+#2049 := (not #2048)
+#2055 := (not #2049)
+#2065 := (and #2055 #2061)
+#2070 := (and #1368 #2065)
+#2004 := (* -1::int #2003)
+#2006 := (+ #2005 #2004)
+#2009 := (+ #2008 #2006)
+#2010 := (>= #2009 0::int)
+#2016 := (and #2015 #2014)
+#2017 := (not #2016)
+#2018 := (or #2017 #2010)
+#2019 := (not #2018)
+#2074 := (or #2019 #2070)
+#2078 := (and #1344 #2074)
+#1979 := (* -1::int #1978)
+#1981 := (+ #1980 #1979)
+#1982 := (>= #1981 0::int)
+#1986 := (and #1985 #1983)
+#1987 := (not #1986)
+#1988 := (or #1987 #1982)
+#1989 := (not #1988)
+#2082 := (or #1989 #2078)
+#2086 := (and #1330 #2082)
+#2090 := (or #1964 #2086)
+#1958 := (not #1324)
+#2094 := (and #1958 #2090)
+#2098 := (or #1324 #2094)
+#2102 := (and #705 #2098)
+#1940 := (= #1939 #1938)
+#1943 := (or #1942 #1940)
+#1944 := (not #1943)
+#2106 := (or #1944 #2102)
+#2110 := (and #1316 #2106)
+#1921 := (* -1::int #1920)
+#1923 := (+ #1922 #1921)
+#1924 := (>= #1923 0::int)
+#1925 := (not #1924)
+#2114 := (or #1925 #2110)
+#1894 := (and #1893 #1891)
+#1913 := (and #184 #192 #1894 #1257 #1268 #1278 #1283)
+#2118 := (and #1913 #2114)
+#1829 := (* -1::int #1828)
+#1831 := (+ #1830 #1829)
+#1834 := (+ #1833 #1831)
+#1835 := (>= #1834 0::int)
+#1844 := (and #1843 #1839)
+#1845 := (not #1844)
+#1846 := (or #1845 #1835)
+#1847 := (not #1846)
+#1866 := (or #1847 #1862)
+#1805 := (+ #1804 #1152)
+#1808 := (+ #1807 #1805)
+#1809 := (= #1808 0::int)
+#1810 := (>= #1805 0::int)
+#1811 := (not #1810)
+#1812 := (and #1811 #1809)
+#1817 := (or #1161 #1812)
+#1820 := (forall (vars (?x49 T2)) #1817)
+#1870 := (and #1820 #1866)
+#1769 := (+ #1768 #1766)
+#1770 := (= #1769 0::int)
+#1773 := (and #1772 #1770)
+#1789 := (not #1773)
+#1792 := (forall (vars (?x50 T2)) #1789)
+#1778 := (= ?x49!8 uf_11)
+#1779 := (not #1778)
+#1780 := (and #1779 #1777)
+#1781 := (not #1780)
+#1786 := (not #1781)
+#1796 := (and #1786 #1792)
+#1874 := (or #1796 #1870)
+#1757 := (and #149 #154 #530 #533 #1754)
+#1878 := (and #1757 #1874)
+#2122 := (or #1878 #2118)
+#1710 := (+ #1709 #1054)
+#1711 := (>= #1710 0::int)
+#1712 := (not #1711)
+#1715 := (+ #1714 #1710)
+#1716 := (= #1715 0::int)
+#1718 := (and #1717 #1716 #1712)
+#1723 := (or #1063 #1718)
+#1726 := (forall (vars (?x46 T2)) #1723)
+#1683 := (+ #993 #1682)
+#1685 := (+ #1684 #1683)
+#1686 := (= #1685 0::int)
+#1687 := (+ #1684 #993)
+#1688 := (>= #1687 0::int)
+#1689 := (not #1688)
+#1691 := (and #1690 #1689 #1686)
+#1696 := (or #1002 #1691)
+#1699 := (forall (vars (?x37 T2)) #1696)
+#1737 := (and #109 #1699 #1051 #1726 #1098 #1104)
+#2126 := (and #1737 #2122)
+#1641 := (+ #1640 #1638)
+#1642 := (+ #66 #1641)
+#1643 := (= #1642 0::int)
+#1647 := (and #74 #1646 #1643)
+#1664 := (not #1647)
+#1667 := (forall (vars (?x38 T2)) #1664)
+#1652 := (= ?x37!5 uf_11)
+#1653 := (not #1652)
+#1654 := (and #1653 #1651)
+#1655 := (not #1654)
+#1661 := (not #1655)
+#1671 := (and #1661 #1667)
+#2130 := (or #1671 #2126)
+#2134 := (and #987 #2130)
+#1612 := (+ #1611 #1609)
+#1614 := (+ #1613 #1612)
+#1615 := (>= #1614 0::int)
+#1621 := (and #1620 #1619)
+#1622 := (not #1621)
+#1623 := (or #1622 #1615)
+#1624 := (not #1623)
+#2138 := (or #1624 #2134)
+#2142 := (and #961 #2138)
+#1584 := (* -1::int #1583)
+#1586 := (+ #1585 #1584)
+#1587 := (>= #1586 0::int)
+#1591 := (and #1590 #1588)
+#1592 := (not #1591)
+#1593 := (or #1592 #1587)
+#1594 := (not #1593)
+#2146 := (or #1594 #2142)
+#2150 := (and #946 #2146)
+#2154 := (or #1569 #2150)
+#1523 := (not #940)
+#2158 := (and #1523 #2154)
+#2162 := (or #940 #2158)
+#2688 := (iff #2162 #2687)
+#2685 := (iff #2158 #2684)
+#2682 := (iff #2154 #2681)
+#2679 := (iff #2150 #2678)
+#2676 := (iff #2146 #2675)
+#2673 := (iff #2142 #2672)
+#2670 := (iff #2138 #2669)
+#2667 := (iff #2134 #2666)
+#2664 := (iff #2130 #2663)
+#2661 := (iff #2126 #2658)
+#2652 := (and #109 #2299 #1051 #2339 #1098 #1104)
+#2655 := (and #2652 #2649)
+#2659 := (iff #2655 #2658)
#2660 := [rewrite]: #2659
-#2663 := [monotonicity #2660]: #2662
-#2671 := [trans #2663 #2669]: #2670
-#2655 := (iff #2018 #2654)
-#2652 := (iff #2017 #2649)
-#2641 := (+ #1251 #2011)
-#2644 := (>= #2641 0::int)
-#2650 := (iff #2644 #2649)
-#2651 := [rewrite]: #2650
-#2645 := (iff #2017 #2644)
-#2642 := (= #2012 #2641)
-#2643 := [rewrite]: #2642
-#2646 := [monotonicity #2643]: #2645
-#2653 := [trans #2646 #2651]: #2652
-#2656 := [monotonicity #2653]: #2655
-#2674 := [monotonicity #2656 #2671]: #2673
-#2677 := [monotonicity #2674]: #2676
-#2680 := [quant-intro #2677]: #2679
-#2717 := [monotonicity #2680 #2714]: #2716
-#2639 := (iff #2003 #2636)
-#2602 := (and #1984 #2596)
-#2633 := (and #2602 #2630)
-#2637 := (iff #2633 #2636)
-#2638 := [rewrite]: #2637
-#2634 := (iff #2003 #2633)
-#2631 := (iff #1999 #2630)
-#2628 := (iff #1996 #2627)
-#2625 := (iff #1980 #2624)
-#2622 := (iff #1977 #2621)
-#2619 := (= #1976 #2618)
-#2620 := [rewrite]: #2619
-#2623 := [monotonicity #2620]: #2622
-#2626 := [monotonicity #2623]: #2625
-#2629 := [monotonicity #2626]: #2628
-#2632 := [quant-intro #2629]: #2631
-#2615 := (iff #1993 #2602)
-#2607 := (not #2602)
-#2610 := (not #2607)
-#2613 := (iff #2610 #2602)
-#2614 := [rewrite]: #2613
-#2611 := (iff #1993 #2610)
-#2608 := (iff #1988 #2607)
-#2605 := (iff #1987 #2602)
-#2599 := (and #2596 #1984)
-#2603 := (iff #2599 #2602)
-#2604 := [rewrite]: #2603
-#2600 := (iff #1987 #2599)
-#2597 := (iff #1986 #2596)
-#2594 := (iff #1985 #2593)
-#2595 := [rewrite]: #2594
-#2598 := [monotonicity #2595]: #2597
-#2601 := [monotonicity #2598]: #2600
-#2606 := [trans #2601 #2604]: #2605
-#2609 := [monotonicity #2606]: #2608
-#2612 := [monotonicity #2609]: #2611
-#2616 := [trans #2612 #2614]: #2615
-#2635 := [monotonicity #2616 #2632]: #2634
-#2640 := [trans #2635 #2638]: #2639
-#2720 := [monotonicity #2640 #2717]: #2719
-#2591 := (iff #1967 #509)
-#2592 := [rewrite]: #2591
-#2589 := (iff #1964 #506)
-#2590 := [rewrite]: #2589
-#2587 := (iff #1961 #145)
-#2588 := [rewrite]: #2587
-#2585 := (iff #1958 #140)
-#2586 := [rewrite]: #2585
-#2723 := [monotonicity #2586 #2588 #2590 #2592 #2720]: #2722
-#2728 := [trans #2723 #2726]: #2727
-#2904 := [monotonicity #2728 #2901]: #2903
-#2583 := (iff #1952 #2582)
-#2580 := (iff #1949 #2579)
-#2577 := (iff #1944 #2576)
-#2574 := (iff #1938 #2573)
-#2571 := (iff #1937 #2568)
-#2561 := (+ #1206 #1935)
-#2564 := (>= #2561 0::int)
-#2569 := (iff #2564 #2568)
-#2570 := [rewrite]: #2569
-#2565 := (iff #1937 #2564)
-#2562 := (= #1936 #2561)
-#2563 := [rewrite]: #2562
-#2566 := [monotonicity #2563]: #2565
-#2572 := [trans #2566 #2570]: #2571
-#2575 := [monotonicity #2572]: #2574
-#2559 := (iff #1942 #2556)
-#2545 := (+ #1935 #1940)
-#2546 := (+ #1206 #2545)
-#2549 := (= #2546 0::int)
-#2557 := (iff #2549 #2556)
-#2558 := [rewrite]: #2557
-#2550 := (iff #1942 #2549)
-#2547 := (= #1941 #2546)
+#2656 := (iff #2126 #2655)
+#2650 := (iff #2122 #2649)
+#2647 := (iff #2118 #2644)
+#2641 := (and #1913 #2638)
+#2645 := (iff #2641 #2644)
+#2646 := [rewrite]: #2645
+#2642 := (iff #2118 #2641)
+#2639 := (iff #2114 #2638)
+#2636 := (iff #2110 #2635)
+#2633 := (iff #2106 #2632)
+#2630 := (iff #2102 #2629)
+#2627 := (iff #2098 #2626)
+#2624 := (iff #2094 #2623)
+#2621 := (iff #2090 #2620)
+#2618 := (iff #2086 #2617)
+#2615 := (iff #2082 #2614)
+#2612 := (iff #2078 #2611)
+#2609 := (iff #2074 #2608)
+#2606 := (iff #2070 #2603)
+#2566 := (and #2045 #2560)
+#2597 := (and #2566 #2594)
+#2600 := (and #1368 #2597)
+#2604 := (iff #2600 #2603)
+#2605 := [rewrite]: #2604
+#2601 := (iff #2070 #2600)
+#2598 := (iff #2065 #2597)
+#2595 := (iff #2061 #2594)
+#2592 := (iff #2058 #2591)
+#2589 := (iff #2041 #2588)
+#2586 := (iff #2038 #2585)
+#2583 := (= #2037 #2582)
+#2584 := [rewrite]: #2583
+#2587 := [monotonicity #2584]: #2586
+#2590 := [monotonicity #2587]: #2589
+#2593 := [monotonicity #2590]: #2592
+#2596 := [quant-intro #2593]: #2595
+#2579 := (iff #2055 #2566)
+#2571 := (not #2566)
+#2574 := (not #2571)
+#2577 := (iff #2574 #2566)
+#2578 := [rewrite]: #2577
+#2575 := (iff #2055 #2574)
+#2572 := (iff #2049 #2571)
+#2569 := (iff #2048 #2566)
+#2563 := (and #2560 #2045)
+#2567 := (iff #2563 #2566)
+#2568 := [rewrite]: #2567
+#2564 := (iff #2048 #2563)
+#2561 := (iff #2047 #2560)
+#2558 := (iff #2046 #2557)
+#2559 := [rewrite]: #2558
+#2562 := [monotonicity #2559]: #2561
+#2565 := [monotonicity #2562]: #2564
+#2570 := [trans #2565 #2568]: #2569
+#2573 := [monotonicity #2570]: #2572
+#2576 := [monotonicity #2573]: #2575
+#2580 := [trans #2576 #2578]: #2579
+#2599 := [monotonicity #2580 #2596]: #2598
+#2602 := [monotonicity #2599]: #2601
+#2607 := [trans #2602 #2605]: #2606
+#2555 := (iff #2019 #2554)
+#2552 := (iff #2018 #2551)
+#2549 := (iff #2010 #2546)
+#2536 := (+ #2005 #2008)
+#2537 := (+ #2004 #2536)
+#2540 := (>= #2537 0::int)
+#2547 := (iff #2540 #2546)
#2548 := [rewrite]: #2547
-#2551 := [monotonicity #2548]: #2550
-#2560 := [trans #2551 #2558]: #2559
-#2578 := [monotonicity #2560 #2575]: #2577
-#2581 := [monotonicity #2578]: #2580
-#2584 := [quant-intro #2581]: #2583
-#2543 := (iff #1926 #2542)
-#2540 := (iff #1923 #2539)
-#2537 := (iff #1918 #2536)
-#2534 := (iff #1913 #2531)
-#2521 := (+ #1909 #1911)
-#2522 := (+ #1168 #2521)
-#2525 := (= #2522 0::int)
-#2532 := (iff #2525 #2531)
-#2533 := [rewrite]: #2532
-#2526 := (iff #1913 #2525)
-#2523 := (= #1912 #2522)
-#2524 := [rewrite]: #2523
-#2527 := [monotonicity #2524]: #2526
-#2535 := [trans #2527 #2533]: #2534
-#2519 := (iff #1916 #2518)
-#2516 := (iff #1915 #2513)
-#2505 := (+ #1168 #1911)
-#2508 := (>= #2505 0::int)
-#2514 := (iff #2508 #2513)
-#2515 := [rewrite]: #2514
-#2509 := (iff #1915 #2508)
-#2506 := (= #1914 #2505)
+#2541 := (iff #2010 #2540)
+#2538 := (= #2009 #2537)
+#2539 := [rewrite]: #2538
+#2542 := [monotonicity #2539]: #2541
+#2550 := [trans #2542 #2548]: #2549
+#2534 := (iff #2017 #2533)
+#2531 := (iff #2016 #2530)
+#2532 := [rewrite]: #2531
+#2535 := [monotonicity #2532]: #2534
+#2553 := [monotonicity #2535 #2550]: #2552
+#2556 := [monotonicity #2553]: #2555
+#2610 := [monotonicity #2556 #2607]: #2609
+#2613 := [monotonicity #2610]: #2612
+#2528 := (iff #1989 #2527)
+#2525 := (iff #1988 #2524)
+#2522 := (iff #1982 #2519)
+#2511 := (+ #1979 #1980)
+#2514 := (>= #2511 0::int)
+#2520 := (iff #2514 #2519)
+#2521 := [rewrite]: #2520
+#2515 := (iff #1982 #2514)
+#2512 := (= #1981 #2511)
+#2513 := [rewrite]: #2512
+#2516 := [monotonicity #2513]: #2515
+#2523 := [trans #2516 #2521]: #2522
+#2509 := (iff #1987 #2508)
+#2506 := (iff #1986 #2505)
#2507 := [rewrite]: #2506
#2510 := [monotonicity #2507]: #2509
-#2517 := [trans #2510 #2515]: #2516
-#2520 := [monotonicity #2517]: #2519
-#2538 := [monotonicity #2520 #2535]: #2537
-#2541 := [monotonicity #2538]: #2540
-#2544 := [quant-intro #2541]: #2543
-#2503 := (iff #1902 #106)
+#2526 := [monotonicity #2510 #2523]: #2525
+#2529 := [monotonicity #2526]: #2528
+#2616 := [monotonicity #2529 #2613]: #2615
+#2619 := [monotonicity #2616]: #2618
+#2622 := [monotonicity #2619]: #2621
+#2503 := (iff #1958 #220)
#2504 := [rewrite]: #2503
-#2907 := [monotonicity #2504 #2544 #2584 #2904]: #2906
-#2912 := [trans #2907 #2910]: #2911
-#2501 := (iff #1897 #2498)
-#2464 := (and #1878 #2458)
-#2495 := (and #2464 #2492)
-#2499 := (iff #2495 #2498)
-#2500 := [rewrite]: #2499
-#2496 := (iff #1897 #2495)
-#2493 := (iff #1893 #2492)
-#2490 := (iff #1890 #2489)
-#2487 := (iff #1874 #2486)
-#2484 := (iff #1870 #2483)
-#2481 := (= #1869 #2480)
-#2482 := [rewrite]: #2481
-#2485 := [monotonicity #2482]: #2484
-#2488 := [monotonicity #2485]: #2487
-#2491 := [monotonicity #2488]: #2490
-#2494 := [quant-intro #2491]: #2493
-#2477 := (iff #1887 #2464)
-#2469 := (not #2464)
-#2472 := (not #2469)
-#2475 := (iff #2472 #2464)
-#2476 := [rewrite]: #2475
-#2473 := (iff #1887 #2472)
-#2470 := (iff #1882 #2469)
-#2467 := (iff #1881 #2464)
-#2461 := (and #2458 #1878)
-#2465 := (iff #2461 #2464)
-#2466 := [rewrite]: #2465
-#2462 := (iff #1881 #2461)
-#2459 := (iff #1880 #2458)
-#2456 := (iff #1879 #2455)
-#2457 := [rewrite]: #2456
-#2460 := [monotonicity #2457]: #2459
-#2463 := [monotonicity #2460]: #2462
-#2468 := [trans #2463 #2466]: #2467
-#2471 := [monotonicity #2468]: #2470
-#2474 := [monotonicity #2471]: #2473
-#2478 := [trans #2474 #2476]: #2477
-#2497 := [monotonicity #2478 #2494]: #2496
-#2502 := [trans #2497 #2500]: #2501
-#2915 := [monotonicity #2502 #2912]: #2914
-#2918 := [monotonicity #2915]: #2917
-#2453 := (iff #1851 #2452)
-#2450 := (iff #1850 #2449)
-#2447 := (iff #1842 #2446)
-#2444 := (= #1841 #2443)
-#2445 := [rewrite]: #2444
-#2448 := [monotonicity #2445]: #2447
-#2440 := (iff #1849 #2439)
-#2437 := (iff #1848 #2436)
-#2438 := [rewrite]: #2437
-#2441 := [monotonicity #2438]: #2440
-#2451 := [monotonicity #2441 #2448]: #2450
-#2454 := [monotonicity #2451]: #2453
-#2921 := [monotonicity #2454 #2918]: #2920
-#2924 := [monotonicity #2921]: #2923
-#2434 := (iff #1821 #2433)
-#2431 := (iff #1820 #2430)
-#2428 := (iff #1814 #2425)
-#2417 := (+ #1811 #1812)
-#2420 := (>= #2417 0::int)
-#2426 := (iff #2420 #2425)
-#2427 := [rewrite]: #2426
-#2421 := (iff #1814 #2420)
-#2418 := (= #1813 #2417)
-#2419 := [rewrite]: #2418
-#2422 := [monotonicity #2419]: #2421
-#2429 := [trans #2422 #2427]: #2428
-#2415 := (iff #1819 #2414)
-#2412 := (iff #1818 #2411)
-#2413 := [rewrite]: #2412
-#2416 := [monotonicity #2413]: #2415
-#2432 := [monotonicity #2416 #2429]: #2431
+#2625 := [monotonicity #2504 #2622]: #2624
+#2628 := [monotonicity #2625]: #2627
+#2631 := [monotonicity #2628]: #2630
+#2501 := (iff #1944 #2500)
+#2498 := (iff #1943 #2497)
+#2495 := (iff #1940 #2494)
+#2496 := [rewrite]: #2495
+#2499 := [monotonicity #2496]: #2498
+#2502 := [monotonicity #2499]: #2501
+#2634 := [monotonicity #2502 #2631]: #2633
+#2637 := [monotonicity #2634]: #2636
+#2492 := (iff #1925 #2491)
+#2489 := (iff #1924 #2486)
+#2478 := (+ #1921 #1922)
+#2481 := (>= #2478 0::int)
+#2487 := (iff #2481 #2486)
+#2488 := [rewrite]: #2487
+#2482 := (iff #1924 #2481)
+#2479 := (= #1923 #2478)
+#2480 := [rewrite]: #2479
+#2483 := [monotonicity #2480]: #2482
+#2490 := [trans #2483 #2488]: #2489
+#2493 := [monotonicity #2490]: #2492
+#2640 := [monotonicity #2493 #2637]: #2639
+#2643 := [monotonicity #2640]: #2642
+#2648 := [trans #2643 #2646]: #2647
+#2476 := (iff #1878 #2473)
+#2470 := (and #1757 #2467)
+#2474 := (iff #2470 #2473)
+#2475 := [rewrite]: #2474
+#2471 := (iff #1878 #2470)
+#2468 := (iff #1874 #2467)
+#2465 := (iff #1870 #2464)
+#2462 := (iff #1866 #2459)
+#2456 := (or #2453 #1862)
+#2460 := (iff #2456 #2459)
+#2461 := [rewrite]: #2460
+#2457 := (iff #1866 #2456)
+#2454 := (iff #1847 #2453)
+#2451 := (iff #1846 #2450)
+#2448 := (iff #1835 #2445)
+#2436 := (+ #1830 #1833)
+#2437 := (+ #1829 #2436)
+#2440 := (>= #2437 0::int)
+#2446 := (iff #2440 #2445)
+#2447 := [rewrite]: #2446
+#2441 := (iff #1835 #2440)
+#2438 := (= #1834 #2437)
+#2439 := [rewrite]: #2438
+#2442 := [monotonicity #2439]: #2441
+#2449 := [trans #2442 #2447]: #2448
+#2434 := (iff #1845 #2433)
+#2431 := (iff #1844 #2430)
+#2432 := [rewrite]: #2431
#2435 := [monotonicity #2432]: #2434
-#2927 := [monotonicity #2435 #2924]: #2926
-#2930 := [monotonicity #2927]: #2929
-#2933 := [monotonicity #2930]: #2932
-#2409 := (iff #1751 #78)
-#2410 := [rewrite]: #2409
-#2936 := [monotonicity #2410 #2933]: #2935
-#2939 := [monotonicity #2936]: #2938
-#1725 := (not #1689)
-#2404 := (~ #1725 #2403)
-#2400 := (not #1686)
-#2401 := (~ #2400 #2399)
-#2396 := (not #1683)
-#2397 := (~ #2396 #2395)
-#2392 := (not #1680)
-#2393 := (~ #2392 #2391)
-#2388 := (not #1677)
-#2389 := (~ #2388 #2387)
-#2384 := (not #1674)
-#2385 := (~ #2384 #2383)
-#2380 := (not #1671)
-#2381 := (~ #2380 #2379)
-#2376 := (not #1668)
-#2377 := (~ #2376 #2375)
-#2372 := (not #1665)
-#2373 := (~ #2372 #2371)
-#2368 := (not #1660)
-#2369 := (~ #2368 #2367)
-#2364 := (not #1631)
-#2365 := (~ #2364 #1628)
-#2362 := (~ #1628 #1628)
-#2360 := (~ #1625 #1625)
-#2361 := [refl]: #2360
-#2363 := [nnf-pos #2361]: #2362
-#2366 := [nnf-neg #2363]: #2365
-#2357 := (not #1622)
-#2358 := (~ #2357 #1619)
-#2355 := (~ #1619 #1619)
-#2353 := (~ #1616 #1616)
-#2354 := [refl]: #2353
-#2356 := [nnf-pos #2354]: #2355
-#2359 := [nnf-neg #2356]: #2358
-#2350 := (not #1611)
-#2351 := (~ #2350 #1608)
-#2348 := (~ #1608 #1608)
-#2346 := (~ #1605 #1605)
-#2347 := [refl]: #2346
-#2349 := [nnf-pos #2347]: #2348
-#2352 := [nnf-neg #2349]: #2351
-#2343 := (not #1588)
-#2344 := (~ #2343 #2342)
-#2339 := (not #1583)
-#2340 := (~ #2339 #2338)
-#2336 := (~ #1549 #1549)
-#2337 := [refl]: #2336
-#2333 := (not #1544)
-#2334 := (~ #2333 #1541)
-#2331 := (~ #1541 #1541)
-#2329 := (~ #1538 #1538)
-#2330 := [refl]: #2329
-#2332 := [nnf-pos #2330]: #2331
-#2335 := [nnf-neg #2332]: #2334
-#2326 := (not #1531)
-#2327 := (~ #2326 #1528)
-#2324 := (~ #1528 #1528)
-#2322 := (~ #1525 #1525)
-#2323 := [refl]: #2322
-#2325 := [nnf-pos #2323]: #2324
-#2328 := [nnf-neg #2325]: #2327
-#2319 := (not #1514)
-#2320 := (~ #2319 #2318)
-#2315 := (not #1511)
-#2316 := (~ #2315 #2314)
-#2311 := (not #1508)
-#2312 := (~ #2311 #2310)
-#2307 := (not #1505)
-#2308 := (~ #2307 #2306)
-#2303 := (not #1502)
-#2304 := (~ #2303 #2302)
-#2299 := (not #1499)
-#2300 := (~ #2299 #2298)
-#2295 := (not #1496)
-#2296 := (~ #2295 #2294)
-#2291 := (not #1493)
-#2292 := (~ #2291 #2290)
-#2287 := (not #1490)
-#2288 := (~ #2287 #2286)
-#2283 := (not #1487)
-#2284 := (~ #2283 #2282)
-#2279 := (not #1484)
-#2280 := (~ #2279 #2278)
-#2275 := (not #1481)
-#2276 := (~ #2275 #2274)
-#2256 := (not #1478)
-#2272 := (~ #2256 #2269)
-#2246 := (exists (vars (?x76 T2)) #2245)
-#2254 := (or #2253 #2246)
-#2255 := (not #2254)
-#2270 := (~ #2255 #2269)
-#2266 := (not #2246)
-#2267 := (~ #2266 #2265)
-#2263 := (~ #2262 #2262)
-#2264 := [refl]: #2263
-#2268 := [nnf-neg #2264]: #2267
-#2260 := (~ #2259 #2259)
-#2261 := [refl]: #2260
-#2271 := [nnf-neg #2261 #2268]: #2270
-#2257 := (~ #2256 #2255)
-#2258 := [sk]: #2257
-#2273 := [trans #2258 #2271]: #2272
-#2232 := (not #1441)
-#2233 := (~ #2232 #1438)
-#2230 := (~ #1438 #1438)
-#2228 := (~ #1435 #1435)
-#2229 := [refl]: #2228
-#2231 := [nnf-pos #2229]: #2230
-#2234 := [nnf-neg #2231]: #2233
-#2277 := [nnf-neg #2234 #2273]: #2276
-#2224 := (~ #1441 #2223)
-#2225 := [sk]: #2224
-#2281 := [nnf-neg #2225 #2277]: #2280
-#2202 := (not #1417)
-#2203 := (~ #2202 #1414)
-#2200 := (~ #1414 #1414)
-#2198 := (~ #1411 #1411)
-#2199 := [refl]: #2198
-#2201 := [nnf-pos #2199]: #2200
-#2204 := [nnf-neg #2201]: #2203
-#2285 := [nnf-neg #2204 #2281]: #2284
-#2194 := (~ #1417 #2193)
-#2195 := [sk]: #2194
-#2289 := [nnf-neg #2195 #2285]: #2288
-#2177 := (not #1403)
-#2178 := (~ #2177 #1400)
-#2175 := (~ #1400 #1400)
-#2173 := (~ #1397 #1397)
-#2174 := [refl]: #2173
-#2176 := [nnf-pos #2174]: #2175
-#2179 := [nnf-neg #2176]: #2178
-#2293 := [nnf-neg #2179 #2289]: #2292
-#2169 := (~ #1403 #2168)
-#2170 := [sk]: #2169
-#2297 := [nnf-neg #2170 #2293]: #2296
-#2163 := (~ #2162 #2162)
-#2164 := [refl]: #2163
-#2301 := [nnf-neg #2164 #2297]: #2300
-#2160 := (~ #1394 #1394)
-#2161 := [refl]: #2160
-#2305 := [nnf-neg #2161 #2301]: #2304
-#2157 := (not #846)
-#2158 := (~ #2157 #710)
-#2155 := (~ #710 #710)
-#2153 := (~ #705 #705)
-#2154 := [refl]: #2153
-#2156 := [nnf-pos #2154]: #2155
-#2159 := [nnf-neg #2156]: #2158
-#2309 := [nnf-neg #2159 #2305]: #2308
-#2149 := (~ #846 #2148)
-#2150 := [sk]: #2149
-#2313 := [nnf-neg #2150 #2309]: #2312
-#2138 := (not #1389)
-#2139 := (~ #2138 #1386)
-#2136 := (~ #1386 #1386)
-#2134 := (~ #1381 #1381)
-#2135 := [refl]: #2134
-#2137 := [nnf-pos #2135]: #2136
-#2140 := [nnf-neg #2137]: #2139
-#2317 := [nnf-neg #2140 #2313]: #2316
-#2130 := (~ #1389 #2129)
-#2131 := [sk]: #2130
-#2321 := [nnf-neg #2131 #2317]: #2320
-#2120 := (not #1378)
-#2121 := (~ #2120 #1375)
-#2118 := (~ #1375 #1375)
-#2116 := (~ #1370 #1370)
-#2117 := [refl]: #2116
-#2119 := [nnf-pos #2117]: #2118
-#2122 := [nnf-neg #2119]: #2121
-#2113 := (not #1559)
-#2114 := (~ #2113 #2108)
-#2109 := (~ #1328 #2108)
-#2110 := [sk]: #2109
-#2115 := [nnf-neg #2110]: #2114
-#2098 := (~ #2097 #2097)
-#2099 := [refl]: #2098
-#2095 := (~ #181 #181)
-#2096 := [refl]: #2095
-#2341 := [nnf-neg #2096 #2099 #2115 #2122 #2321 #2328 #2335 #2337]: #2340
-#2092 := (not #1346)
-#2093 := (~ #2092 #2091)
-#2089 := (~ #1559 #2088)
-#2086 := (~ #2085 #2085)
-#2087 := [refl]: #2086
-#2090 := [nnf-neg #2087]: #2089
-#2082 := (not #1322)
-#2083 := (~ #2082 #2081)
-#2078 := (not #1319)
-#2079 := (~ #2078 #2077)
-#2074 := (not #1316)
-#2075 := (~ #2074 #2073)
-#2070 := (not #1311)
-#2071 := (~ #2070 #2069)
-#2066 := (not #1303)
-#2067 := (~ #2066 #1300)
-#2064 := (~ #1300 #1300)
-#2062 := (~ #1297 #1297)
-#2063 := [refl]: #2062
-#2065 := [nnf-pos #2063]: #2064
-#2068 := [nnf-neg #2065]: #2067
-#2060 := (~ #2059 #2059)
-#2061 := [refl]: #2060
-#2072 := [nnf-neg #2061 #2068]: #2071
-#2055 := (~ #1303 #2054)
-#2056 := [sk]: #2055
-#2076 := [nnf-neg #2056 #2072]: #2075
-#2030 := (not #1285)
-#2031 := (~ #2030 #2027)
-#2028 := (~ #1282 #2027)
-#2025 := (~ #1279 #2024)
-#2020 := (~ #1276 #2019)
+#2452 := [monotonicity #2435 #2449]: #2451
+#2455 := [monotonicity #2452]: #2454
+#2458 := [monotonicity #2455]: #2457
+#2463 := [trans #2458 #2461]: #2462
+#2428 := (iff #1820 #2427)
+#2425 := (iff #1817 #2424)
+#2422 := (iff #1812 #2421)
+#2419 := (iff #1809 #2416)
+#2406 := (+ #1804 #1807)
+#2407 := (+ #1152 #2406)
+#2410 := (= #2407 0::int)
+#2417 := (iff #2410 #2416)
+#2418 := [rewrite]: #2417
+#2411 := (iff #1809 #2410)
+#2408 := (= #1808 #2407)
+#2409 := [rewrite]: #2408
+#2412 := [monotonicity #2409]: #2411
+#2420 := [trans #2412 #2418]: #2419
+#2404 := (iff #1811 #2403)
+#2401 := (iff #1810 #2398)
+#2390 := (+ #1152 #1804)
+#2393 := (>= #2390 0::int)
+#2399 := (iff #2393 #2398)
+#2400 := [rewrite]: #2399
+#2394 := (iff #1810 #2393)
+#2391 := (= #1805 #2390)
+#2392 := [rewrite]: #2391
+#2395 := [monotonicity #2392]: #2394
+#2402 := [trans #2395 #2400]: #2401
+#2405 := [monotonicity #2402]: #2404
+#2423 := [monotonicity #2405 #2420]: #2422
+#2426 := [monotonicity #2423]: #2425
+#2429 := [quant-intro #2426]: #2428
+#2466 := [monotonicity #2429 #2463]: #2465
+#2388 := (iff #1796 #2385)
+#2351 := (and #1777 #2345)
+#2382 := (and #2351 #2379)
+#2386 := (iff #2382 #2385)
+#2387 := [rewrite]: #2386
+#2383 := (iff #1796 #2382)
+#2380 := (iff #1792 #2379)
+#2377 := (iff #1789 #2376)
+#2374 := (iff #1773 #2373)
+#2371 := (iff #1770 #2370)
+#2368 := (= #1769 #2367)
+#2369 := [rewrite]: #2368
+#2372 := [monotonicity #2369]: #2371
+#2375 := [monotonicity #2372]: #2374
+#2378 := [monotonicity #2375]: #2377
+#2381 := [quant-intro #2378]: #2380
+#2364 := (iff #1786 #2351)
+#2356 := (not #2351)
+#2359 := (not #2356)
+#2362 := (iff #2359 #2351)
+#2363 := [rewrite]: #2362
+#2360 := (iff #1786 #2359)
+#2357 := (iff #1781 #2356)
+#2354 := (iff #1780 #2351)
+#2348 := (and #2345 #1777)
+#2352 := (iff #2348 #2351)
+#2353 := [rewrite]: #2352
+#2349 := (iff #1780 #2348)
+#2346 := (iff #1779 #2345)
+#2343 := (iff #1778 #2342)
+#2344 := [rewrite]: #2343
+#2347 := [monotonicity #2344]: #2346
+#2350 := [monotonicity #2347]: #2349
+#2355 := [trans #2350 #2353]: #2354
+#2358 := [monotonicity #2355]: #2357
+#2361 := [monotonicity #2358]: #2360
+#2365 := [trans #2361 #2363]: #2364
+#2384 := [monotonicity #2365 #2381]: #2383
+#2389 := [trans #2384 #2387]: #2388
+#2469 := [monotonicity #2389 #2466]: #2468
+#2472 := [monotonicity #2469]: #2471
+#2477 := [trans #2472 #2475]: #2476
+#2651 := [monotonicity #2477 #2648]: #2650
+#2653 := (iff #1737 #2652)
+#2340 := (iff #1726 #2339)
+#2337 := (iff #1723 #2336)
+#2334 := (iff #1718 #2333)
+#2331 := (iff #1712 #2330)
+#2328 := (iff #1711 #2325)
+#2318 := (+ #1054 #1709)
+#2321 := (>= #2318 0::int)
+#2326 := (iff #2321 #2325)
+#2327 := [rewrite]: #2326
+#2322 := (iff #1711 #2321)
+#2319 := (= #1710 #2318)
+#2320 := [rewrite]: #2319
+#2323 := [monotonicity #2320]: #2322
+#2329 := [trans #2323 #2327]: #2328
+#2332 := [monotonicity #2329]: #2331
+#2316 := (iff #1716 #2313)
+#2302 := (+ #1709 #1714)
+#2303 := (+ #1054 #2302)
+#2306 := (= #2303 0::int)
+#2314 := (iff #2306 #2313)
+#2315 := [rewrite]: #2314
+#2307 := (iff #1716 #2306)
+#2304 := (= #1715 #2303)
+#2305 := [rewrite]: #2304
+#2308 := [monotonicity #2305]: #2307
+#2317 := [trans #2308 #2315]: #2316
+#2335 := [monotonicity #2317 #2332]: #2334
+#2338 := [monotonicity #2335]: #2337
+#2341 := [quant-intro #2338]: #2340
+#2300 := (iff #1699 #2299)
+#2297 := (iff #1696 #2296)
+#2294 := (iff #1691 #2293)
+#2291 := (iff #1686 #2288)
+#2278 := (+ #1682 #1684)
+#2279 := (+ #993 #2278)
+#2282 := (= #2279 0::int)
+#2289 := (iff #2282 #2288)
+#2290 := [rewrite]: #2289
+#2283 := (iff #1686 #2282)
+#2280 := (= #1685 #2279)
+#2281 := [rewrite]: #2280
+#2284 := [monotonicity #2281]: #2283
+#2292 := [trans #2284 #2290]: #2291
+#2276 := (iff #1689 #2275)
+#2273 := (iff #1688 #2270)
+#2262 := (+ #993 #1684)
+#2265 := (>= #2262 0::int)
+#2271 := (iff #2265 #2270)
+#2272 := [rewrite]: #2271
+#2266 := (iff #1688 #2265)
+#2263 := (= #1687 #2262)
+#2264 := [rewrite]: #2263
+#2267 := [monotonicity #2264]: #2266
+#2274 := [trans #2267 #2272]: #2273
+#2277 := [monotonicity #2274]: #2276
+#2295 := [monotonicity #2277 #2292]: #2294
+#2298 := [monotonicity #2295]: #2297
+#2301 := [quant-intro #2298]: #2300
+#2654 := [monotonicity #2301 #2341]: #2653
+#2657 := [monotonicity #2654 #2651]: #2656
+#2662 := [trans #2657 #2660]: #2661
+#2260 := (iff #1671 #2257)
+#2223 := (and #1651 #2217)
+#2254 := (and #2223 #2251)
+#2258 := (iff #2254 #2257)
+#2259 := [rewrite]: #2258
+#2255 := (iff #1671 #2254)
+#2252 := (iff #1667 #2251)
+#2249 := (iff #1664 #2248)
+#2246 := (iff #1647 #2245)
+#2243 := (iff #1643 #2242)
+#2240 := (= #1642 #2239)
+#2241 := [rewrite]: #2240
+#2244 := [monotonicity #2241]: #2243
+#2247 := [monotonicity #2244]: #2246
+#2250 := [monotonicity #2247]: #2249
+#2253 := [quant-intro #2250]: #2252
+#2236 := (iff #1661 #2223)
+#2228 := (not #2223)
+#2231 := (not #2228)
+#2234 := (iff #2231 #2223)
+#2235 := [rewrite]: #2234
+#2232 := (iff #1661 #2231)
+#2229 := (iff #1655 #2228)
+#2226 := (iff #1654 #2223)
+#2220 := (and #2217 #1651)
+#2224 := (iff #2220 #2223)
+#2225 := [rewrite]: #2224
+#2221 := (iff #1654 #2220)
+#2218 := (iff #1653 #2217)
+#2215 := (iff #1652 #2214)
+#2216 := [rewrite]: #2215
+#2219 := [monotonicity #2216]: #2218
+#2222 := [monotonicity #2219]: #2221
+#2227 := [trans #2222 #2225]: #2226
+#2230 := [monotonicity #2227]: #2229
+#2233 := [monotonicity #2230]: #2232
+#2237 := [trans #2233 #2235]: #2236
+#2256 := [monotonicity #2237 #2253]: #2255
+#2261 := [trans #2256 #2259]: #2260
+#2665 := [monotonicity #2261 #2662]: #2664
+#2668 := [monotonicity #2665]: #2667
+#2212 := (iff #1624 #2211)
+#2209 := (iff #1623 #2208)
+#2206 := (iff #1615 #2205)
+#2203 := (= #1614 #2202)
+#2204 := [rewrite]: #2203
+#2207 := [monotonicity #2204]: #2206
+#2199 := (iff #1622 #2198)
+#2196 := (iff #1621 #2195)
+#2197 := [rewrite]: #2196
+#2200 := [monotonicity #2197]: #2199
+#2210 := [monotonicity #2200 #2207]: #2209
+#2213 := [monotonicity #2210]: #2212
+#2671 := [monotonicity #2213 #2668]: #2670
+#2674 := [monotonicity #2671]: #2673
+#2193 := (iff #1594 #2192)
+#2190 := (iff #1593 #2189)
+#2187 := (iff #1587 #2184)
+#2176 := (+ #1584 #1585)
+#2179 := (>= #2176 0::int)
+#2185 := (iff #2179 #2184)
+#2186 := [rewrite]: #2185
+#2180 := (iff #1587 #2179)
+#2177 := (= #1586 #2176)
+#2178 := [rewrite]: #2177
+#2181 := [monotonicity #2178]: #2180
+#2188 := [trans #2181 #2186]: #2187
+#2174 := (iff #1592 #2173)
+#2171 := (iff #1591 #2170)
+#2172 := [rewrite]: #2171
+#2175 := [monotonicity #2172]: #2174
+#2191 := [monotonicity #2175 #2188]: #2190
+#2194 := [monotonicity #2191]: #2193
+#2677 := [monotonicity #2194 #2674]: #2676
+#2680 := [monotonicity #2677]: #2679
+#2683 := [monotonicity #2680]: #2682
+#2168 := (iff #1523 #81)
+#2169 := [rewrite]: #2168
+#2686 := [monotonicity #2169 #2683]: #2685
+#2689 := [monotonicity #2686]: #2688
+#1496 := (not #1480)
+#2163 := (~ #1496 #2162)
+#2159 := (not #1477)
+#2160 := (~ #2159 #2158)
+#2155 := (not #1474)
+#2156 := (~ #2155 #2154)
+#2151 := (not #1471)
+#2152 := (~ #2151 #2150)
+#2147 := (not #1468)
+#2148 := (~ #2147 #2146)
+#2143 := (not #1465)
+#2144 := (~ #2143 #2142)
+#2139 := (not #1462)
+#2140 := (~ #2139 #2138)
+#2135 := (not #1459)
+#2136 := (~ #2135 #2134)
+#2131 := (not #1456)
+#2132 := (~ #2131 #2130)
+#2127 := (not #1453)
+#2128 := (~ #2127 #2126)
+#2123 := (not #1450)
+#2124 := (~ #2123 #2122)
+#2119 := (not #1447)
+#2120 := (~ #2119 #2118)
+#2115 := (not #1444)
+#2116 := (~ #2115 #2114)
+#2111 := (not #1441)
+#2112 := (~ #2111 #2110)
+#2107 := (not #1438)
+#2108 := (~ #2107 #2106)
+#2103 := (not #1435)
+#2104 := (~ #2103 #2102)
+#2099 := (not #1432)
+#2100 := (~ #2099 #2098)
+#2095 := (not #1429)
+#2096 := (~ #2095 #2094)
+#2091 := (not #1426)
+#2092 := (~ #2091 #2090)
+#2087 := (not #1423)
+#2088 := (~ #2087 #2086)
+#2083 := (not #1420)
+#2084 := (~ #2083 #2082)
+#2079 := (not #1417)
+#2080 := (~ #2079 #2078)
+#2075 := (not #1414)
+#2076 := (~ #2075 #2074)
+#2071 := (not #1411)
+#2072 := (~ #2071 #2070)
+#2052 := (not #1408)
+#2068 := (~ #2052 #2065)
+#2042 := (exists (vars (?x76 T2)) #2041)
+#2050 := (or #2049 #2042)
+#2051 := (not #2050)
+#2066 := (~ #2051 #2065)
+#2062 := (not #2042)
+#2063 := (~ #2062 #2061)
+#2059 := (~ #2058 #2058)
+#2060 := [refl]: #2059
+#2064 := [nnf-neg #2060]: #2063
+#2056 := (~ #2055 #2055)
+#2057 := [refl]: #2056
+#2067 := [nnf-neg #2057 #2064]: #2066
+#2053 := (~ #2052 #2051)
+#2054 := [sk]: #2053
+#2069 := [trans #2054 #2067]: #2068
+#2028 := (not #1371)
+#2029 := (~ #2028 #1368)
+#2026 := (~ #1368 #1368)
+#2024 := (~ #1365 #1365)
+#2025 := [refl]: #2024
+#2027 := [nnf-pos #2025]: #2026
+#2030 := [nnf-neg #2027]: #2029
+#2073 := [nnf-neg #2030 #2069]: #2072
+#2020 := (~ #1371 #2019)
#2021 := [sk]: #2020
-#2008 := (~ #1260 #1260)
-#2009 := [refl]: #2008
-#2026 := [monotonicity #2009 #2021]: #2025
-#2029 := [nnf-pos #2026]: #2028
-#2032 := [nnf-neg #2029]: #2031
-#2080 := [nnf-neg #2032 #2076]: #2079
-#2006 := (~ #1285 #2003)
-#1981 := (exists (vars (?x50 T2)) #1980)
-#1989 := (or #1988 #1981)
-#1990 := (not #1989)
-#2004 := (~ #1990 #2003)
-#2000 := (not #1981)
-#2001 := (~ #2000 #1999)
-#1997 := (~ #1996 #1996)
-#1998 := [refl]: #1997
-#2002 := [nnf-neg #1998]: #2001
-#1994 := (~ #1993 #1993)
+#2077 := [nnf-neg #2021 #2073]: #2076
+#1998 := (not #1347)
+#1999 := (~ #1998 #1344)
+#1996 := (~ #1344 #1344)
+#1994 := (~ #1341 #1341)
#1995 := [refl]: #1994
-#2005 := [nnf-neg #1995 #2002]: #2004
-#1991 := (~ #1285 #1990)
-#1992 := [sk]: #1991
-#2007 := [trans #1992 #2005]: #2006
-#2084 := [nnf-neg #2007 #2080]: #2083
-#1968 := (~ #1967 #1967)
-#1969 := [refl]: #1968
-#1965 := (~ #1964 #1964)
-#1966 := [refl]: #1965
-#1962 := (~ #1961 #1961)
-#1963 := [refl]: #1962
+#1997 := [nnf-pos #1995]: #1996
+#2000 := [nnf-neg #1997]: #1999
+#2081 := [nnf-neg #2000 #2077]: #2080
+#1990 := (~ #1347 #1989)
+#1991 := [sk]: #1990
+#2085 := [nnf-neg #1991 #2081]: #2084
+#1973 := (not #1333)
+#1974 := (~ #1973 #1330)
+#1971 := (~ #1330 #1330)
+#1969 := (~ #1327 #1327)
+#1970 := [refl]: #1969
+#1972 := [nnf-pos #1970]: #1971
+#1975 := [nnf-neg #1972]: #1974
+#2089 := [nnf-neg #1975 #2085]: #2088
+#1965 := (~ #1333 #1964)
+#1966 := [sk]: #1965
+#2093 := [nnf-neg #1966 #2089]: #2092
#1959 := (~ #1958 #1958)
#1960 := [refl]: #1959
-#2094 := [nnf-neg #1960 #1963 #1966 #1969 #2084 #2090]: #2093
-#2345 := [nnf-neg #2094 #2341]: #2344
-#1955 := (not #1248)
-#1956 := (~ #1955 #1952)
-#1953 := (~ #1245 #1952)
-#1950 := (~ #1242 #1949)
-#1945 := (~ #1239 #1944)
+#2097 := [nnf-neg #1960 #2093]: #2096
+#1956 := (~ #1324 #1324)
+#1957 := [refl]: #1956
+#2101 := [nnf-neg #1957 #2097]: #2100
+#1953 := (not #814)
+#1954 := (~ #1953 #705)
+#1951 := (~ #705 #705)
+#1949 := (~ #700 #700)
+#1950 := [refl]: #1949
+#1952 := [nnf-pos #1950]: #1951
+#1955 := [nnf-neg #1952]: #1954
+#2105 := [nnf-neg #1955 #2101]: #2104
+#1945 := (~ #814 #1944)
#1946 := [sk]: #1945
-#1932 := (~ #1215 #1215)
-#1933 := [refl]: #1932
-#1951 := [monotonicity #1933 #1946]: #1950
-#1954 := [nnf-pos #1951]: #1953
-#1957 := [nnf-neg #1954]: #1956
-#1929 := (not #1639)
-#1930 := (~ #1929 #1926)
-#1927 := (~ #1203 #1926)
-#1924 := (~ #1200 #1923)
-#1919 := (~ #1197 #1918)
-#1920 := [sk]: #1919
-#1905 := (~ #1177 #1177)
-#1906 := [refl]: #1905
-#1925 := [monotonicity #1906 #1920]: #1924
-#1928 := [nnf-pos #1925]: #1927
-#1931 := [nnf-neg #1928]: #1930
-#1903 := (~ #1902 #1902)
+#2109 := [nnf-neg #1946 #2105]: #2108
+#1934 := (not #1319)
+#1935 := (~ #1934 #1316)
+#1932 := (~ #1316 #1316)
+#1930 := (~ #1312 #1312)
+#1931 := [refl]: #1930
+#1933 := [nnf-pos #1931]: #1932
+#1936 := [nnf-neg #1933]: #1935
+#2113 := [nnf-neg #1936 #2109]: #2112
+#1926 := (~ #1319 #1925)
+#1927 := [sk]: #1926
+#2117 := [nnf-neg #1927 #2113]: #2116
+#1916 := (not #1309)
+#1917 := (~ #1916 #1913)
+#1914 := (~ #1304 #1913)
+#1911 := (~ #1283 #1283)
+#1912 := [refl]: #1911
+#1909 := (~ #1278 #1278)
+#1907 := (~ #1275 #1275)
+#1908 := [refl]: #1907
+#1910 := [nnf-pos #1908]: #1909
+#1905 := (~ #1268 #1268)
+#1903 := (~ #1263 #1263)
#1904 := [refl]: #1903
-#2370 := [nnf-neg #1904 #1931 #1957 #2345 #2352 #2359 #2366]: #2369
-#1900 := (~ #1639 #1897)
-#1875 := (exists (vars (?x38 T2)) #1874)
-#1883 := (or #1882 #1875)
-#1884 := (not #1883)
-#1898 := (~ #1884 #1897)
-#1894 := (not #1875)
-#1895 := (~ #1894 #1893)
-#1891 := (~ #1890 #1890)
-#1892 := [refl]: #1891
-#1896 := [nnf-neg #1892]: #1895
-#1888 := (~ #1887 #1887)
-#1889 := [refl]: #1888
-#1899 := [nnf-neg #1889 #1896]: #1898
-#1885 := (~ #1639 #1884)
-#1886 := [sk]: #1885
-#1901 := [trans #1886 #1899]: #1900
-#2374 := [nnf-neg #1901 #2370]: #2373
-#1860 := (not #1165)
-#1861 := (~ #1860 #1162)
-#1858 := (~ #1162 #1162)
-#1856 := (~ #1159 #1159)
-#1857 := [refl]: #1856
-#1859 := [nnf-pos #1857]: #1858
-#1862 := [nnf-neg #1859]: #1861
-#2378 := [nnf-neg #1862 #2374]: #2377
-#1852 := (~ #1165 #1851)
-#1853 := [sk]: #1852
-#2382 := [nnf-neg #1853 #2378]: #2381
-#1830 := (not #1139)
-#1831 := (~ #1830 #1136)
-#1828 := (~ #1136 #1136)
-#1826 := (~ #1133 #1133)
-#1827 := [refl]: #1826
-#1829 := [nnf-pos #1827]: #1828
-#1832 := [nnf-neg #1829]: #1831
-#2386 := [nnf-neg #1832 #2382]: #2385
-#1822 := (~ #1139 #1821)
-#1823 := [sk]: #1822
-#2390 := [nnf-neg #1823 #2386]: #2389
-#1805 := (not #1124)
-#1806 := (~ #1805 #1121)
-#1803 := (~ #1121 #1121)
-#1801 := (~ #1120 #1120)
+#1906 := [nnf-pos #1904]: #1905
+#1901 := (~ #1257 #1257)
+#1899 := (~ #1254 #1254)
+#1900 := [refl]: #1899
+#1902 := [nnf-pos #1900]: #1901
+#1895 := (~ #1135 #1894)
+#1896 := [sk]: #1895
+#1884 := (~ #192 #192)
+#1885 := [refl]: #1884
+#1882 := (~ #184 #184)
+#1883 := [refl]: #1882
+#1915 := [monotonicity #1883 #1885 #1896 #1902 #1906 #1910 #1912]: #1914
+#1918 := [nnf-neg #1915]: #1917
+#2121 := [nnf-neg #1918 #2117]: #2120
+#1879 := (not #1226)
+#1880 := (~ #1879 #1878)
+#1875 := (not #1223)
+#1876 := (~ #1875 #1874)
+#1871 := (not #1220)
+#1872 := (~ #1871 #1870)
+#1867 := (not #1217)
+#1868 := (~ #1867 #1866)
+#1863 := (not #1212)
+#1864 := (~ #1863 #1862)
+#1859 := (not #1204)
+#1860 := (~ #1859 #1201)
+#1857 := (~ #1201 #1201)
+#1855 := (~ #1198 #1198)
+#1856 := [refl]: #1855
+#1858 := [nnf-pos #1856]: #1857
+#1861 := [nnf-neg #1858]: #1860
+#1853 := (~ #1852 #1852)
+#1854 := [refl]: #1853
+#1865 := [nnf-neg #1854 #1861]: #1864
+#1848 := (~ #1204 #1847)
+#1849 := [sk]: #1848
+#1869 := [nnf-neg #1849 #1865]: #1868
+#1823 := (not #1186)
+#1824 := (~ #1823 #1820)
+#1821 := (~ #1183 #1820)
+#1818 := (~ #1180 #1817)
+#1813 := (~ #1177 #1812)
+#1814 := [sk]: #1813
+#1801 := (~ #1161 #1161)
#1802 := [refl]: #1801
-#1804 := [nnf-pos #1802]: #1803
-#1807 := [nnf-neg #1804]: #1806
-#2394 := [nnf-neg #1807 #2390]: #2393
-#1797 := (~ #1124 #1796)
-#1798 := [sk]: #1797
-#2398 := [nnf-neg #1798 #2394]: #2397
-#1752 := (~ #1751 #1751)
-#1792 := [refl]: #1752
-#2402 := [nnf-neg #1792 #2398]: #2401
-#1790 := (~ #1115 #1115)
+#1819 := [monotonicity #1802 #1814]: #1818
+#1822 := [nnf-pos #1819]: #1821
+#1825 := [nnf-neg #1822]: #1824
+#1873 := [nnf-neg #1825 #1869]: #1872
+#1799 := (~ #1186 #1796)
+#1774 := (exists (vars (?x50 T2)) #1773)
+#1782 := (or #1781 #1774)
+#1783 := (not #1782)
+#1797 := (~ #1783 #1796)
+#1793 := (not #1774)
+#1794 := (~ #1793 #1792)
+#1790 := (~ #1789 #1789)
#1791 := [refl]: #1790
-#2405 := [nnf-neg #1791 #2402]: #2404
-#1726 := [not-or-elim #1722]: #1725
-#2406 := [mp~ #1726 #2405]: #2403
-#2407 := [mp #2406 #2939]: #2937
-#3581 := [mp #2407 #3580]: #3578
-#4510 := [mp #3581 #4509]: #4507
-#10108 := [unit-resolution #4510 #5109]: #4504
-#3836 := (or #4501 #4495)
-#3679 := [def-axiom]: #3836
-#10111 := [unit-resolution #3679 #10108]: #4495
-#10112 := (or #4498 #4492)
-#3753 := (* -1::int #1794)
-#3720 := (+ uf_9 #3753)
-#3722 := (<= #3720 0::int)
-#3827 := (= uf_9 #1794)
-#3801 := (= uf_11 ?x27!0)
-#3650 := (not #3801)
-#3649 := (= #1794 0::int)
-#4542 := (not #3649)
-#4541 := [hypothesis]: #1796
-#4593 := (or #4542 #1795)
-#4594 := [th-lemma]: #4593
-#4595 := [unit-resolution #4594 #4541]: #4542
-#3660 := (or #3659 #3649 #3650)
-#3816 := (= ?x27!0 uf_11)
-#3651 := (not #3816)
-#3652 := (or #3651 #3649)
-#3661 := (or #3659 #3652)
-#4532 := (iff #3661 #3660)
-#3674 := (or #3649 #3650)
-#4533 := (or #3659 #3674)
-#4536 := (iff #4533 #3660)
-#4537 := [rewrite]: #4536
-#4534 := (iff #3661 #4533)
-#3672 := (iff #3652 #3674)
-#4518 := (or #3650 #3649)
-#3658 := (iff #4518 #3674)
-#3655 := [rewrite]: #3658
-#3673 := (iff #3652 #4518)
-#3653 := (iff #3651 #3650)
-#3799 := (iff #3816 #3801)
-#3802 := [rewrite]: #3799
-#4517 := [monotonicity #3802]: #3653
-#3665 := [monotonicity #4517]: #3673
-#3656 := [trans #3665 #3655]: #3672
-#4535 := [monotonicity #3656]: #4534
-#4538 := [trans #4535 #4537]: #4532
-#3657 := [quant-inst]: #3661
-#4539 := [mp #3657 #4538]: #3660
-#4596 := [unit-resolution #4539 #4516 #4595]: #3650
-#3784 := (or #3801 #3827)
-#4132 := (forall (vars (?x25 T2)) (:pat #4131) #419)
-#4135 := (iff #424 #4132)
-#4133 := (iff #419 #419)
-#4134 := [refl]: #4133
-#4136 := [quant-intro #4134]: #4135
-#1749 := (~ #424 #424)
-#1787 := (~ #419 #419)
+#1795 := [nnf-neg #1791]: #1794
+#1787 := (~ #1786 #1786)
#1788 := [refl]: #1787
-#1750 := [nnf-pos #1788]: #1749
-#1724 := [not-or-elim #1722]: #424
-#1789 := [mp~ #1724 #1750]: #424
-#4137 := [mp #1789 #4136]: #4132
-#3787 := (not #4132)
-#3788 := (or #3787 #3801 #3827)
-#3819 := (or #3816 #3827)
-#3789 := (or #3787 #3819)
-#3751 := (iff #3789 #3788)
-#3791 := (or #3787 #3784)
-#3742 := (iff #3791 #3788)
-#3749 := [rewrite]: #3742
-#3748 := (iff #3789 #3791)
-#3786 := (iff #3819 #3784)
-#3800 := [monotonicity #3802]: #3786
-#3750 := [monotonicity #3800]: #3748
-#3752 := [trans #3750 #3749]: #3751
-#3790 := [quant-inst]: #3789
-#3743 := [mp #3790 #3752]: #3788
-#4597 := [unit-resolution #3743 #4137]: #3784
-#4598 := [unit-resolution #4597 #4596]: #3827
-#4599 := (not #3827)
-#4600 := (or #4599 #3722)
-#4601 := [th-lemma]: #4600
-#4581 := [unit-resolution #4601 #4598]: #3722
-#4540 := (<= #1794 0::int)
-#4582 := (or #4540 #1795)
-#4583 := [th-lemma]: #4582
-#4584 := [unit-resolution #4583 #4541]: #4540
-#349 := (<= uf_9 0::int)
-#350 := (not #349)
+#1798 := [nnf-neg #1788 #1795]: #1797
+#1784 := (~ #1186 #1783)
+#1785 := [sk]: #1784
+#1800 := [trans #1785 #1798]: #1799
+#1877 := [nnf-neg #1800 #1873]: #1876
+#1760 := (not #1149)
+#1761 := (~ #1760 #1757)
+#1758 := (~ #1144 #1757)
+#1755 := (~ #1138 #1754)
+#1752 := (~ #1751 #1751)
+#1753 := [refl]: #1752
+#1756 := [nnf-neg #1753]: #1755
+#1749 := (~ #533 #533)
+#1750 := [refl]: #1749
+#1747 := (~ #530 #530)
+#1748 := [refl]: #1747
+#1745 := (~ #154 #154)
+#1746 := [refl]: #1745
+#1743 := (~ #149 #149)
+#1744 := [refl]: #1743
+#1759 := [monotonicity #1744 #1746 #1748 #1750 #1756]: #1758
+#1762 := [nnf-neg #1759]: #1761
+#1881 := [nnf-neg #1762 #1877]: #1880
+#2125 := [nnf-neg #1881 #2121]: #2124
+#1740 := (not #1129)
+#1741 := (~ #1740 #1737)
+#1738 := (~ #1124 #1737)
+#1735 := (~ #1104 #1104)
+#1733 := (~ #1101 #1101)
+#1734 := [refl]: #1733
+#1736 := [nnf-pos #1734]: #1735
+#1731 := (~ #1098 #1098)
+#1729 := (~ #1095 #1095)
+#1730 := [refl]: #1729
+#1732 := [nnf-pos #1730]: #1731
+#1727 := (~ #1090 #1726)
+#1724 := (~ #1087 #1723)
+#1719 := (~ #1084 #1718)
+#1720 := [sk]: #1719
+#1706 := (~ #1063 #1063)
+#1707 := [refl]: #1706
+#1725 := [monotonicity #1707 #1720]: #1724
+#1728 := [nnf-pos #1725]: #1727
+#1704 := (~ #1051 #1051)
+#1702 := (~ #1048 #1048)
+#1703 := [refl]: #1702
+#1705 := [nnf-pos #1703]: #1704
+#1700 := (~ #1028 #1699)
+#1697 := (~ #1025 #1696)
+#1692 := (~ #1022 #1691)
+#1693 := [sk]: #1692
+#1678 := (~ #1002 #1002)
+#1679 := [refl]: #1678
+#1698 := [monotonicity #1679 #1693]: #1697
+#1701 := [nnf-pos #1698]: #1700
+#1676 := (~ #109 #109)
+#1677 := [refl]: #1676
+#1739 := [monotonicity #1677 #1701 #1705 #1728 #1732 #1736]: #1738
+#1742 := [nnf-neg #1739]: #1741
+#2129 := [nnf-neg #1742 #2125]: #2128
+#1658 := (not #1028)
+#1674 := (~ #1658 #1671)
+#1648 := (exists (vars (?x38 T2)) #1647)
+#1656 := (or #1655 #1648)
+#1657 := (not #1656)
+#1672 := (~ #1657 #1671)
+#1668 := (not #1648)
+#1669 := (~ #1668 #1667)
+#1665 := (~ #1664 #1664)
+#1666 := [refl]: #1665
+#1670 := [nnf-neg #1666]: #1669
+#1662 := (~ #1661 #1661)
+#1663 := [refl]: #1662
+#1673 := [nnf-neg #1663 #1670]: #1672
+#1659 := (~ #1658 #1657)
+#1660 := [sk]: #1659
+#1675 := [trans #1660 #1673]: #1674
+#2133 := [nnf-neg #1675 #2129]: #2132
+#1633 := (not #990)
+#1634 := (~ #1633 #987)
+#1631 := (~ #987 #987)
+#1629 := (~ #984 #984)
+#1630 := [refl]: #1629
+#1632 := [nnf-pos #1630]: #1631
+#1635 := [nnf-neg #1632]: #1634
+#2137 := [nnf-neg #1635 #2133]: #2136
+#1625 := (~ #990 #1624)
+#1626 := [sk]: #1625
+#2141 := [nnf-neg #1626 #2137]: #2140
+#1603 := (not #964)
+#1604 := (~ #1603 #961)
+#1601 := (~ #961 #961)
+#1599 := (~ #958 #958)
+#1600 := [refl]: #1599
+#1602 := [nnf-pos #1600]: #1601
+#1605 := [nnf-neg #1602]: #1604
+#2145 := [nnf-neg #1605 #2141]: #2144
+#1595 := (~ #964 #1594)
+#1596 := [sk]: #1595
+#2149 := [nnf-neg #1596 #2145]: #2148
+#1578 := (not #949)
+#1579 := (~ #1578 #946)
+#1576 := (~ #946 #946)
+#1574 := (~ #945 #945)
+#1575 := [refl]: #1574
+#1577 := [nnf-pos #1575]: #1576
+#1580 := [nnf-neg #1577]: #1579
+#2153 := [nnf-neg #1580 #2149]: #2152
+#1570 := (~ #949 #1569)
+#1571 := [sk]: #1570
+#2157 := [nnf-neg #1571 #2153]: #2156
+#1524 := (~ #1523 #1523)
+#1565 := [refl]: #1524
+#2161 := [nnf-neg #1565 #2157]: #2160
+#1563 := (~ #940 #940)
+#1564 := [refl]: #1563
+#2164 := [nnf-neg #1564 #2161]: #2163
+#1497 := [not-or-elim #1491]: #1496
+#2165 := [mp~ #1497 #2164]: #2162
+#2166 := [mp #2165 #2689]: #2687
+#3336 := [mp #2166 #3335]: #3333
+#4271 := [mp #3336 #4270]: #4268
+#6202 := [unit-resolution #4271 #4994]: #4265
+#3591 := (or #4262 #4256)
+#3434 := [def-axiom]: #3591
+#6203 := [unit-resolution #3434 #6202]: #4256
+#6204 := (or #4259 #4253)
+#333 := (<= uf_9 0::int)
+#334 := (not #333)
#54 := (< 0::int uf_9)
-#351 := (iff #54 #350)
-#352 := [rewrite]: #351
-#345 := [asserted]: #54
-#353 := [mp #345 #352]: #350
-#4585 := [th-lemma #353 #4584 #4581]: false
-#4580 := [lemma #4585]: #1795
-#3831 := (or #4498 #1796 #4492)
-#3832 := [def-axiom]: #3831
-#10113 := [unit-resolution #3832 #4580]: #10112
-#10114 := [unit-resolution #10113 #10111]: #4492
-#3855 := (or #4489 #4483)
-#3856 := [def-axiom]: #3855
-#10107 := [unit-resolution #3856 #10114]: #4483
-#3850 := (or #4486 #1948 #4480)
-#3851 := [def-axiom]: #3850
-#10115 := [unit-resolution #3851 #10107]: #4483
-#10116 := [unit-resolution #10115 #10110]: #4480
-#3876 := (or #4477 #4471)
-#3877 := [def-axiom]: #3876
-#10117 := [unit-resolution #3877 #10116]: #4471
-#3872 := (or #4474 #2982 #4468)
-#3873 := [def-axiom]: #3872
-#10118 := [unit-resolution #3873 #10117 #10109]: #4468
-#3860 := (or #4465 #4459)
-#3861 := [def-axiom]: #3860
-#10119 := [unit-resolution #3861 #10118]: #4459
-#10121 := (or #4462 #4456)
-#4588 := [hypothesis]: #4176
-#4058 := (or #4173 #2458)
-#4059 := [def-axiom]: #4058
-#4725 := [unit-resolution #4059 #4588]: #2458
-#4673 := (= uf_9 #1866)
-#4816 := (not #4673)
-#3725 := (or #4173 #1878)
-#4057 := [def-axiom]: #3725
-#4726 := [unit-resolution #4057 #4588]: #1878
-#4826 := (or #4816 #1877)
-#4827 := [th-lemma]: #4826
-#4828 := [unit-resolution #4827 #4726]: #4816
-#4847 := (or #4673 #2455)
-#4817 := [hypothesis]: #4816
-#4818 := [hypothesis]: #2458
-#4730 := (or #3787 #2455 #4673)
-#4674 := (or #1879 #4673)
-#4731 := (or #3787 #4674)
-#4717 := (iff #4731 #4730)
-#4727 := (or #2455 #4673)
-#4733 := (or #3787 #4727)
-#4715 := (iff #4733 #4730)
-#4716 := [rewrite]: #4715
-#4713 := (iff #4731 #4733)
-#4728 := (iff #4674 #4727)
-#4729 := [monotonicity #2457]: #4728
-#4714 := [monotonicity #4729]: #4713
-#4712 := [trans #4714 #4716]: #4717
-#4732 := [quant-inst]: #4731
-#4718 := [mp #4732 #4712]: #4730
-#4819 := [unit-resolution #4718 #4137 #4818 #4817]: false
-#4849 := [lemma #4819]: #4847
-#4829 := [unit-resolution #4849 #4828 #4725]: false
-#4830 := [lemma #4829]: #4173
-#3894 := (or #4462 #4176 #4456)
-#3904 := [def-axiom]: #3894
-#10122 := [unit-resolution #3904 #4830]: #10121
-#10123 := [unit-resolution #10122 #10119]: #4456
-#3889 := (or #4453 #4447)
-#3848 := [def-axiom]: #3889
-#10335 := [unit-resolution #3848 #10123]: #4447
-#3682 := (not #2696)
-#3925 := (or #4453 #106)
-#3921 := [def-axiom]: #3925
-#10124 := [unit-resolution #3921 #10123]: #106
-#8213 := (= #161 #105)
-#4974 := [hypothesis]: #4289
-#3741 := (or #4286 #509)
-#4023 := [def-axiom]: #3741
-#4975 := [unit-resolution #4023 #4974]: #509
-#8228 := [symm #4975]: #142
-#8026 := [monotonicity #8228]: #8213
-#4825 := [trans #8026 #10124]: #162
-#3701 := (or #4262 #2059)
-#3702 := [def-axiom]: #3701
-#7196 := [unit-resolution #3702 #4825]: #4262
-#4028 := (or #4286 #4280)
-#4017 := [def-axiom]: #4028
-#8815 := [unit-resolution #4017 #4974]: #4280
-#10558 := (or #4240 #614)
-#8833 := (?x47!7 ?x49!8)
-#8906 := (uf_4 uf_19 #8833)
-#8925 := (* -1::int #8906)
-#8828 := (uf_4 uf_14 #8833)
-#9957 := (+ #8828 #8925)
-#9963 := (>= #9957 0::int)
-#9956 := (= #8828 #8906)
-#10507 := (= #8906 #8828)
-#6339 := [hypothesis]: #509
-#10506 := [symm #6339]: #142
-#10508 := [monotonicity #10506]: #10507
-#10509 := [symm #10508]: #9956
-#10510 := (not #9956)
-#10505 := (or #10510 #9963)
-#10511 := [th-lemma]: #10505
-#10512 := [unit-resolution #10511 #10509]: #9963
-#8834 := (* -1::int #8828)
-#8675 := (uf_4 uf_14 ?x49!8)
-#8835 := (+ #8675 #8834)
-#8836 := (<= #8835 0::int)
-#8878 := (not #8836)
-#8859 := (up_6 uf_15 #8833)
-#8860 := (not #8859)
-#8837 := (uf_1 #8833 ?x49!8)
-#8838 := (uf_10 #8837)
-#8854 := (* -1::int #8838)
-#8855 := (+ #8834 #8854)
-#8856 := (+ #8675 #8855)
-#8857 := (= #8856 0::int)
-#8858 := (not #8857)
-#8843 := (or #8836 #8858 #8860)
-#8846 := (not #8843)
-#8810 := (* -1::int #8675)
-#8823 := (+ uf_9 #8810)
-#8811 := (<= #8823 0::int)
-#9028 := (not #8811)
-#10513 := [hypothesis]: #4243
-#3711 := (or #4240 #1984)
-#3716 := [def-axiom]: #3711
-#10514 := [unit-resolution #3716 #10513]: #1984
-#9024 := (+ #1971 #8810)
-#9021 := (>= #9024 0::int)
-#9023 := (= #1971 #8675)
-#10515 := (= #8675 #1971)
-#10516 := [monotonicity #6339]: #10515
-#10517 := [symm #10516]: #9023
-#10518 := (not #9023)
-#10519 := (or #10518 #9021)
-#10520 := [th-lemma]: #10519
-#10521 := [unit-resolution #10520 #10517]: #9021
-#9020 := (not #9021)
-#9029 := (or #9028 #9020 #1983)
-#9025 := [hypothesis]: #1984
-#9022 := [hypothesis]: #8811
-#9026 := [hypothesis]: #9021
-#9027 := [th-lemma #9026 #9022 #9025]: false
-#8827 := [lemma #9027]: #9029
-#10522 := [unit-resolution #8827 #10521 #10514]: #9028
-#10532 := (or #8811 #8846)
-#4052 := (or #4240 #2596)
-#3712 := [def-axiom]: #4052
-#10523 := [unit-resolution #3712 #10513]: #2596
-#3914 := (or #4453 #4213)
-#3882 := [def-axiom]: #3914
-#10531 := [unit-resolution #3882 #10123]: #4213
-#8851 := (or #4218 #2593 #8811 #8846)
-#8861 := (or #8860 #8858 #8836)
-#8862 := (not #8861)
-#8842 := (or #1985 #8811 #8862)
-#8864 := (or #4218 #8842)
-#8870 := (iff #8864 #8851)
-#8848 := (or #2593 #8811 #8846)
-#8866 := (or #4218 #8848)
-#8863 := (iff #8866 #8851)
-#8869 := [rewrite]: #8863
-#8867 := (iff #8864 #8866)
-#8849 := (iff #8842 #8848)
-#8841 := (iff #8862 #8846)
-#8844 := (iff #8861 #8843)
-#8845 := [rewrite]: #8844
-#8847 := [monotonicity #8845]: #8841
-#8850 := [monotonicity #2595 #8847]: #8849
-#8868 := [monotonicity #8850]: #8867
-#8871 := [trans #8868 #8869]: #8870
-#8865 := [quant-inst]: #8864
-#8872 := [mp #8865 #8871]: #8851
-#10533 := [unit-resolution #8872 #10531 #10523]: #10532
-#10534 := [unit-resolution #10533 #10522]: #8846
-#8876 := (or #8843 #8878)
-#8879 := [def-axiom]: #8876
-#10535 := [unit-resolution #8879 #10534]: #8878
-#8920 := (+ #1971 #8925)
-#8937 := (<= #8920 0::int)
-#8982 := (+ #8854 #8925)
-#8983 := (+ #1971 #8982)
-#9001 := (= #8983 0::int)
-#9173 := (<= #8983 0::int)
-#9960 := (<= #9957 0::int)
-#10536 := (or #10510 #9960)
-#10537 := [th-lemma]: #10536
-#10538 := [unit-resolution #10537 #10509]: #9960
-#8873 := (<= #8856 0::int)
-#8880 := (or #8843 #8857)
-#8881 := [def-axiom]: #8880
-#10539 := [unit-resolution #8881 #10534]: #8857
-#10540 := (or #8858 #8873)
-#10541 := [th-lemma]: #10540
-#10542 := [unit-resolution #10541 #10539]: #8873
-#9019 := (<= #9024 0::int)
-#10543 := (or #10518 #9019)
-#10544 := [th-lemma]: #10543
-#10545 := [unit-resolution #10544 #10517]: #9019
-#10185 := (not #9960)
-#10187 := (not #8873)
-#10186 := (not #9019)
-#10188 := (or #9173 #10186 #10187 #10185)
-#10148 := [hypothesis]: #9960
-#10149 := [hypothesis]: #8873
-#10151 := [hypothesis]: #9019
-#10152 := (not #9173)
-#10153 := [hypothesis]: #10152
-#10154 := [th-lemma #10153 #10151 #10149 #10148]: false
-#10189 := [lemma #10154]: #10188
-#10546 := [unit-resolution #10189 #10545 #10542 #10538]: #9173
-#9157 := (>= #8983 0::int)
-#8877 := (>= #8856 0::int)
-#10547 := (or #8858 #8877)
-#10548 := [th-lemma]: #10547
-#10549 := [unit-resolution #10548 #10539]: #8877
-#10528 := (not #9963)
-#10096 := (not #8877)
-#10529 := (or #9157 #9020 #10096 #10528)
-#10524 := [hypothesis]: #9963
-#10027 := [hypothesis]: #8877
-#10525 := (not #9157)
-#10526 := [hypothesis]: #10525
-#10527 := [th-lemma #10526 #9026 #10027 #10524]: false
-#10530 := [lemma #10527]: #10529
-#10550 := [unit-resolution #10530 #10521 #10549 #10512]: #9157
-#10551 := (or #9001 #10152 #10525)
-#10552 := [th-lemma]: #10551
-#10553 := [unit-resolution #10552 #10550 #10546]: #9001
-#9000 := (not #9001)
-#9007 := (or #8937 #9000)
-#4053 := (or #4240 #4232)
-#3696 := [def-axiom]: #4053
-#10554 := [unit-resolution #3696 #10513]: #4232
-#9111 := (or #4237 #8937 #9000)
-#8904 := (+ #1972 #8838)
-#8907 := (+ #8906 #8904)
-#8908 := (= #8907 0::int)
-#8909 := (not #8908)
-#8910 := (+ #8906 #1972)
-#8914 := (>= #8910 0::int)
-#8915 := (or #8914 #8909)
-#9120 := (or #4237 #8915)
-#9170 := (iff #9120 #9111)
-#9167 := (or #4237 #9007)
-#9113 := (iff #9167 #9111)
-#9169 := [rewrite]: #9113
-#9168 := (iff #9120 #9167)
-#9008 := (iff #8915 #9007)
-#9005 := (iff #8909 #9000)
-#9004 := (iff #8908 #9001)
-#8975 := (+ #8838 #8906)
-#8976 := (+ #1972 #8975)
-#8979 := (= #8976 0::int)
-#9002 := (iff #8979 #9001)
-#9003 := [rewrite]: #9002
-#8980 := (iff #8908 #8979)
-#8977 := (= #8907 #8976)
-#8978 := [rewrite]: #8977
-#8981 := [monotonicity #8978]: #8980
-#8999 := [trans #8981 #9003]: #9004
-#9006 := [monotonicity #8999]: #9005
-#8946 := (iff #8914 #8937)
-#8916 := (+ #1972 #8906)
-#8922 := (>= #8916 0::int)
-#8938 := (iff #8922 #8937)
-#8945 := [rewrite]: #8938
-#8923 := (iff #8914 #8922)
-#8918 := (= #8910 #8916)
-#8921 := [rewrite]: #8918
-#8924 := [monotonicity #8921]: #8923
-#8947 := [trans #8924 #8945]: #8946
-#9009 := [monotonicity #8947 #9006]: #9008
-#9112 := [monotonicity #9009]: #9168
-#9171 := [trans #9112 #9169]: #9170
-#9121 := [quant-inst]: #9120
-#9172 := [mp #9121 #9171]: #9111
-#10555 := [unit-resolution #9172 #10554]: #9007
-#10556 := [unit-resolution #10555 #10553]: #8937
-#10557 := [th-lemma #10521 #10556 #10535 #10512]: false
-#10559 := [lemma #10557]: #10558
-#8829 := [unit-resolution #10559 #4975]: #4240
-#4030 := (or #4283 #4243 #4277)
-#4034 := [def-axiom]: #4030
-#8830 := [unit-resolution #4034 #8829 #8815]: #4277
-#3761 := (or #4274 #4268)
-#3654 := [def-axiom]: #3761
-#8831 := [unit-resolution #3654 #8830]: #4268
-#4036 := (or #4271 #4265 #3250)
-#3758 := [def-axiom]: #4036
-#8832 := [unit-resolution #3758 #8831 #7196]: #3250
-#4045 := (or #3245 #3682)
-#4047 := [def-axiom]: #4045
-#8875 := [unit-resolution #4047 #8832]: #3682
-#3852 := (or #4453 #4188)
-#3907 := [def-axiom]: #3852
-#10131 := [unit-resolution #3907 #10123]: #4188
-#4042 := (or #3245 #2046)
-#4043 := [def-axiom]: #4042
-#8882 := [unit-resolution #4043 #8832]: #2046
-#4038 := (or #3245 #2050)
-#4044 := [def-axiom]: #4038
-#8883 := [unit-resolution #4044 #8832]: #2050
-#4959 := (or #4286 #2045 #4193 #2049 #2696)
-#4978 := (uf_4 uf_14 ?x53!11)
-#4972 := (= #2035 #4978)
-#4976 := (= #4978 #2035)
-#4971 := [monotonicity #4975]: #4976
-#4977 := [symm #4971]: #4972
-#4979 := (* -1::int #4978)
-#6252 := (+ #2035 #4979)
-#6267 := (<= #6252 0::int)
-#6377 := (not #6267)
-#6280 := [hypothesis]: #3682
-#6333 := [hypothesis]: #2050
-#6336 := [hypothesis]: #4188
-#6338 := [hypothesis]: #2046
-#4027 := (or #4286 #4222)
-#4024 := [def-axiom]: #4027
-#4930 := [unit-resolution #4024 #4974]: #4222
-#6383 := (or #6377 #2045 #4227 #4193 #2049 #2696 #614)
-#5295 := (uf_4 uf_14 ?x54!10)
-#5296 := (* -1::int #5295)
-#5291 := (+ uf_9 #5296)
-#5297 := (<= #5291 0::int)
-#5298 := (up_6 uf_15 ?x54!10)
-#5736 := (not #5298)
-#5668 := (+ #4979 #5295)
-#5669 := (+ #2040 #5668)
-#5661 := (>= #5669 0::int)
-#6283 := (not #5661)
-#6285 := (+ #2037 #5296)
-#6297 := (>= #6285 0::int)
-#6284 := (= #2037 #5295)
-#6298 := (= #5295 #2037)
-#6296 := [monotonicity #6339]: #6298
-#6299 := [symm #6296]: #6284
-#6300 := (not #6284)
-#6301 := (or #6300 #6297)
-#6330 := [th-lemma]: #6301
-#6331 := [unit-resolution #6330 #6299]: #6297
-#6281 := [hypothesis]: #6267
-#6378 := (not #6297)
-#6379 := (or #6283 #6377 #2696 #6378)
-#6279 := [hypothesis]: #6297
-#6276 := [hypothesis]: #5661
-#6282 := [th-lemma #6276 #6281 #6280 #6279]: false
-#6380 := [lemma #6282]: #6379
-#6332 := [unit-resolution #6380 #6281 #6280 #6331]: #6283
-#6337 := (or #5736 #5661)
-#5758 := (or #4193 #2049 #5736 #5661)
-#5694 := (+ #5295 #4979)
-#5695 := (+ #2040 #5694)
-#5735 := (>= #5695 0::int)
-#5667 := (or #5736 #2049 #5735)
-#5763 := (or #4193 #5667)
-#6028 := (iff #5763 #5758)
-#5759 := (or #2049 #5736 #5661)
-#5765 := (or #4193 #5759)
-#5998 := (iff #5765 #5758)
-#5999 := [rewrite]: #5998
-#5766 := (iff #5763 #5765)
-#5762 := (iff #5667 #5759)
-#5687 := (or #5736 #2049 #5661)
-#5760 := (iff #5687 #5759)
-#5761 := [rewrite]: #5760
-#5690 := (iff #5667 #5687)
-#5688 := (iff #5735 #5661)
-#5670 := (= #5695 #5669)
-#5671 := [rewrite]: #5670
-#5689 := [monotonicity #5671]: #5688
-#5691 := [monotonicity #5689]: #5690
-#5757 := [trans #5691 #5761]: #5762
-#5767 := [monotonicity #5757]: #5766
-#6029 := [trans #5767 #5999]: #6028
-#5764 := [quant-inst]: #5763
-#6030 := [mp #5764 #6029]: #5758
-#6266 := [unit-resolution #6030 #6336 #6333]: #6337
-#6278 := [unit-resolution #6266 #6332]: #5736
-#5300 := (or #5297 #5298)
-#6257 := [hypothesis]: #4222
-#5326 := (or #4227 #5297 #5298)
-#5299 := (or #5298 #5297)
-#5327 := (or #4227 #5299)
-#5333 := (iff #5327 #5326)
-#5329 := (or #4227 #5300)
-#5331 := (iff #5329 #5326)
-#5332 := [rewrite]: #5331
-#5324 := (iff #5327 #5329)
-#5301 := (iff #5299 #5300)
-#5325 := [rewrite]: #5301
-#5330 := [monotonicity #5325]: #5324
-#5334 := [trans #5330 #5332]: #5333
-#5328 := [quant-inst]: #5327
-#5382 := [mp #5328 #5334]: #5326
-#6381 := [unit-resolution #5382 #6257]: #5300
-#6376 := [unit-resolution #6381 #6278]: #5297
-#6382 := [th-lemma #6331 #6376 #6338]: false
-#6384 := [lemma #6382]: #6383
-#4955 := [unit-resolution #6384 #4930 #6338 #6336 #6333 #6280 #4975]: #6377
-#4956 := (not #4972)
-#4957 := (or #4956 #6267)
-#4958 := [th-lemma]: #4957
-#4929 := [unit-resolution #4958 #4955 #4977]: false
-#4954 := [lemma #4929]: #4959
-#8884 := [unit-resolution #4954 #8883 #8882 #10131 #4974 #8875]: false
-#8887 := [lemma #8884]: #4286
-#3923 := (or #4450 #4289 #4444)
-#3924 := [def-axiom]: #3923
-#10456 := [unit-resolution #3924 #8887 #10335]: #4444
-#3945 := (or #4441 #189)
-#3931 := [def-axiom]: #3945
-#10469 := [unit-resolution #3931 #10456]: #189
-#10470 := [symm #10469]: #7202
-#13610 := (= #11533 #188)
-#13435 := (= #10571 uf_22)
-#10572 := (= uf_22 #10571)
+#335 := (iff #54 #334)
+#336 := [rewrite]: #335
+#329 := [asserted]: #54
+#337 := [mp #329 #336]: #334
+#4524 := (* -1::int #1567)
+#4525 := (+ uf_9 #4524)
+#4534 := (<= #4525 0::int)
+#3581 := (= uf_9 #1567)
+#3568 := (= uf_11 ?x27!0)
+#4428 := (not #3568)
+#4541 := (= #1567 0::int)
+#4864 := (not #4541)
+#4858 := [hypothesis]: #1569
+#4865 := (or #4864 #1568)
+#4866 := [th-lemma]: #4865
+#4867 := [unit-resolution #4866 #4858]: #4864
+#4873 := (or #4872 #4428 #4541)
+#3582 := (= ?x27!0 uf_11)
+#4542 := (not #3582)
+#4427 := (or #4542 #4541)
+#4874 := (or #4872 #4427)
+#4860 := (iff #4874 #4873)
+#4550 := (or #4428 #4541)
+#4876 := (or #4872 #4550)
+#4879 := (iff #4876 #4873)
+#4859 := [rewrite]: #4879
+#4877 := (iff #4874 #4876)
+#4553 := (iff #4427 #4550)
+#4362 := (iff #4542 #4428)
+#3573 := (iff #3582 #3568)
+#3558 := [rewrite]: #3573
+#4538 := [monotonicity #3558]: #4362
+#4871 := [monotonicity #4538]: #4553
+#4878 := [monotonicity #4871]: #4877
+#4861 := [trans #4878 #4859]: #4860
+#4875 := [quant-inst]: #4874
+#4862 := [mp #4875 #4861]: #4873
+#4868 := [unit-resolution #4862 #3898 #4867]: #4428
+#3560 := (or #3568 #3581)
+#3887 := (forall (vars (?x25 T2)) (:pat #3886) #403)
+#3890 := (iff #408 #3887)
+#3888 := (iff #403 #403)
+#3889 := [refl]: #3888
+#3891 := [quant-intro #3889]: #3890
+#1519 := (~ #408 #408)
+#1557 := (~ #403 #403)
+#1558 := [refl]: #1557
+#1520 := [nnf-pos #1558]: #1519
+#1494 := [and-elim #1492]: #408
+#1559 := [mp~ #1494 #1520]: #408
+#3892 := [mp #1559 #3891]: #3887
+#3569 := (not #3887)
+#4289 := (or #3569 #3568 #3581)
+#3578 := (or #3582 #3581)
+#4290 := (or #3569 #3578)
+#4364 := (iff #4290 #4289)
+#4292 := (or #3569 #3560)
+#4308 := (iff #4292 #4289)
+#4363 := [rewrite]: #4308
+#4307 := (iff #4290 #4292)
+#3561 := (iff #3578 #3560)
+#3565 := [monotonicity #3558]: #3561
+#4288 := [monotonicity #3565]: #4307
+#4365 := [trans #4288 #4363]: #4364
+#4291 := [quant-inst]: #4290
+#4366 := [mp #4291 #4365]: #4289
+#4883 := [unit-resolution #4366 #3892]: #3560
+#4884 := [unit-resolution #4883 #4868]: #3581
+#4909 := (not #3581)
+#4910 := (or #4909 #4534)
+#4911 := [th-lemma]: #4910
+#4912 := [unit-resolution #4911 #4884]: #4534
+#4863 := (<= #1567 0::int)
+#4913 := (or #4863 #1568)
+#4908 := [th-lemma]: #4913
+#4914 := [unit-resolution #4908 #4858]: #4863
+#4915 := [th-lemma #4914 #4912 #337]: false
+#4916 := [lemma #4915]: #1568
+#3586 := (or #4259 #1569 #4253)
+#3587 := [def-axiom]: #3586
+#6205 := [unit-resolution #3587 #4916]: #6204
+#6206 := [unit-resolution #6205 #6203]: #4253
+#3610 := (or #4250 #4244)
+#3611 := [def-axiom]: #3610
+#6207 := [unit-resolution #3611 #6206]: #4244
+#5458 := [hypothesis]: #1588
+#3880 := (forall (vars (?x26 T2)) (:pat #3879) #75)
+#3883 := (iff #76 #3880)
+#3881 := (iff #75 #75)
+#3882 := [refl]: #3881
+#3884 := [quant-intro #3882]: #3883
+#1517 := (~ #76 #76)
+#1554 := (~ #75 #75)
+#1555 := [refl]: #1554
+#1518 := [nnf-pos #1555]: #1517
+#1493 := [and-elim #1492]: #76
+#1556 := [mp~ #1493 #1518]: #76
+#3885 := [mp #1556 #3884]: #3880
+#4999 := (not #3880)
+#5000 := (or #4999 #1597)
+#5001 := [quant-inst]: #5000
+#5459 := [unit-resolution #5001 #3885 #5458]: false
+#5460 := [lemma #5459]: #1597
+#3403 := (or #1948 #1588)
+#3488 := [def-axiom]: #3403
+#6208 := [unit-resolution #3488 #5460]: #1948
+#3605 := (or #4247 #2022 #4241)
+#3606 := [def-axiom]: #3605
+#6209 := [unit-resolution #3606 #6208 #6207]: #4241
+#3631 := (or #4238 #4232)
+#3632 := [def-axiom]: #3631
+#6210 := [unit-resolution #3632 #6209]: #4232
+#5548 := [hypothesis]: #1620
+#5094 := (or #4999 #2709)
+#5095 := [quant-inst]: #5094
+#5549 := [unit-resolution #5095 #3885 #5548]: false
+#5553 := [lemma #5549]: #2709
+#3474 := (or #2724 #1620)
+#3809 := [def-axiom]: #3474
+#6211 := [unit-resolution #3809 #5553]: #2724
+#3627 := (or #4235 #2729 #4229)
+#3628 := [def-axiom]: #3627
+#6212 := [unit-resolution #3628 #6211 #6210]: #4229
+#3615 := (or #4226 #4220)
+#3616 := [def-axiom]: #3615
+#6213 := [unit-resolution #3616 #6212]: #4220
+#5484 := (= uf_9 #1639)
+#5561 := (not #5484)
+#5559 := [hypothesis]: #3937
+#3480 := (or #3934 #1651)
+#3812 := [def-axiom]: #3480
+#5560 := [unit-resolution #3812 #5559]: #1651
+#5562 := (or #5561 #1650)
+#5563 := [th-lemma]: #5562
+#5564 := [unit-resolution #5563 #5560]: #5561
+#3813 := (or #3934 #2217)
+#3814 := [def-axiom]: #3813
+#5565 := [unit-resolution #3814 #5559]: #2217
+#5504 := (or #3569 #2214 #5484)
+#5485 := (or #1652 #5484)
+#5505 := (or #3569 #5485)
+#5519 := (iff #5505 #5504)
+#5501 := (or #2214 #5484)
+#5506 := (or #3569 #5501)
+#5509 := (iff #5506 #5504)
+#5510 := [rewrite]: #5509
+#5507 := (iff #5505 #5506)
+#5502 := (iff #5485 #5501)
+#5503 := [monotonicity #2216]: #5502
+#5508 := [monotonicity #5503]: #5507
+#5521 := [trans #5508 #5510]: #5519
+#5500 := [quant-inst]: #5505
+#5522 := [mp #5500 #5521]: #5504
+#5566 := [unit-resolution #5522 #3892 #5565 #5564]: false
+#5567 := [lemma #5566]: #3934
+#3649 := (or #4223 #3937 #4217)
+#3659 := [def-axiom]: #3649
+#6214 := [unit-resolution #3659 #5567 #6213]: #4217
+#5359 := (or #4214 #4205)
+#4367 := (?x47!7 ?x49!8)
+#4714 := (uf_4 uf_19 #4367)
+#4751 := (* -1::int #4714)
+#4368 := (uf_4 uf_14 #4367)
+#5011 := (+ #4368 #4751)
+#5015 := (<= #5011 0::int)
+#5010 := (= #4368 #4714)
+#5211 := (= #4714 #4368)
+#5206 := [hypothesis]: #4202
+#5207 := [hypothesis]: #4217
+#3644 := (or #4214 #4208)
+#3603 := [def-axiom]: #3644
+#5208 := [unit-resolution #3603 #5207]: #4208
+#3678 := (or #4211 #4050 #4205)
+#3679 := [def-axiom]: #3678
+#5209 := [unit-resolution #3679 #5208 #5206]: #4050
+#3496 := (or #4047 #533)
+#3778 := [def-axiom]: #3496
+#5210 := [unit-resolution #3778 #5209]: #533
+#5205 := [symm #5210]: #151
+#5212 := [monotonicity #5205]: #5211
+#5213 := [symm #5212]: #5010
+#5214 := (not #5010)
+#5215 := (or #5214 #5015)
+#5268 := [th-lemma]: #5215
+#5269 := [unit-resolution #5268 #5213]: #5015
+#4369 := (* -1::int #4368)
+#4353 := (uf_1 #4367 ?x49!8)
+#4354 := (uf_10 #4353)
+#4355 := (* -1::int #4354)
+#4432 := (+ #4355 #4369)
+#4301 := (uf_4 uf_14 ?x49!8)
+#4433 := (+ #4301 #4432)
+#4520 := (<= #4433 0::int)
+#4436 := (= #4433 0::int)
+#4418 := (not #4436)
+#4359 := (up_6 uf_15 #4367)
+#4360 := (not #4359)
+#4351 := (+ #4301 #4369)
+#4352 := (<= #4351 0::int)
+#4423 := (or #4352 #4360 #4418)
+#4442 := (not #4423)
+#4302 := (* -1::int #4301)
+#4303 := (+ uf_9 #4302)
+#4283 := (<= #4303 0::int)
+#4675 := (not #4283)
+#3783 := (or #4047 #4041)
+#3772 := [def-axiom]: #3783
+#5270 := [unit-resolution #3772 #5209]: #4041
+#3680 := (or #4214 #109)
+#3676 := [def-axiom]: #3680
+#5271 := [unit-resolution #3676 #5207]: #109
+#5272 := (= #175 #108)
+#5273 := [monotonicity #5205]: #5272
+#5274 := [trans #5273 #5271]: #176
+#3456 := (or #4023 #1852)
+#3457 := [def-axiom]: #3456
+#5275 := [unit-resolution #3457 #5274]: #4023
+#3782 := (or #4047 #3983)
+#3779 := [def-axiom]: #3782
+#5276 := [unit-resolution #3779 #5209]: #3983
+#3607 := (or #4214 #3949)
+#3662 := [def-axiom]: #3607
+#5256 := [unit-resolution #3662 #5207]: #3949
+#4685 := (or #2992 #3954 #3988 #3023)
+#4370 := (uf_4 uf_14 ?x54!10)
+#4371 := (* -1::int #4370)
+#4552 := (+ #1830 #4371)
+#4554 := (>= #4552 0::int)
+#4551 := (= #1830 #4370)
+#4633 := (= #4370 #1830)
+#4632 := [hypothesis]: #533
+#4634 := [monotonicity #4632]: #4633
+#4635 := [symm #4634]: #4551
+#4636 := (not #4551)
+#4637 := (or #4636 #4554)
+#4638 := [th-lemma]: #4637
+#4620 := [unit-resolution #4638 #4635]: #4554
+#3503 := (uf_4 uf_14 ?x53!11)
+#3504 := (* -1::int #3503)
+#4545 := (+ #1828 #3504)
+#4549 := (<= #4545 0::int)
+#4544 := (= #1828 #3503)
+#4621 := (= #3503 #1828)
+#4622 := [monotonicity #4632]: #4621
+#4623 := [symm #4622]: #4544
+#4624 := (not #4544)
+#4619 := (or #4624 #4549)
+#4625 := [th-lemma]: #4619
+#4626 := [unit-resolution #4625 #4623]: #4549
+#3437 := (not #2445)
+#4627 := [hypothesis]: #2997
+#3800 := (or #2992 #3437)
+#3802 := [def-axiom]: #3800
+#4628 := [unit-resolution #3802 #4627]: #3437
+#4822 := [hypothesis]: #3983
+#4826 := [hypothesis]: #3949
+#3793 := (or #2992 #1843)
+#3799 := [def-axiom]: #3793
+#4629 := [unit-resolution #3799 #4627]: #1843
+#4372 := (+ uf_9 #4371)
+#4373 := (<= #4372 0::int)
+#4820 := (not #4373)
+#3797 := (or #2992 #1839)
+#3798 := [def-axiom]: #3797
+#4680 := [unit-resolution #3798 #4627]: #1839
+#4816 := (not #4554)
+#4681 := (or #4820 #1838 #4816)
+#4682 := [th-lemma]: #4681
+#4683 := [unit-resolution #4682 #4680 #4620]: #4820
+#4815 := (not #4549)
+#4830 := (or #4373 #3954 #1842 #3988 #2445 #4815 #4816)
+#4446 := (+ #3504 #4370)
+#4447 := (+ #1833 #4446)
+#4450 := (>= #4447 0::int)
+#4814 := (not #4450)
+#4811 := [hypothesis]: #4554
+#4812 := [hypothesis]: #4549
+#4813 := [hypothesis]: #3437
+#4817 := (or #4814 #4815 #2445 #4816)
+#4818 := [th-lemma]: #4817
+#4819 := [unit-resolution #4818 #4813 #4812 #4811]: #4814
+#4374 := (up_6 uf_15 ?x54!10)
+#4821 := [hypothesis]: #4820
+#4376 := (or #4373 #4374)
+#4379 := (or #3988 #4373 #4374)
+#4375 := (or #4374 #4373)
+#4380 := (or #3988 #4375)
+#4387 := (iff #4380 #4379)
+#4382 := (or #3988 #4376)
+#4385 := (iff #4382 #4379)
+#4386 := [rewrite]: #4385
+#4383 := (iff #4380 #4382)
+#4377 := (iff #4375 #4376)
+#4378 := [rewrite]: #4377
+#4384 := [monotonicity #4378]: #4383
+#4388 := [trans #4384 #4386]: #4387
+#4381 := [quant-inst]: #4380
+#4389 := [mp #4381 #4388]: #4379
+#4823 := [unit-resolution #4389 #4822]: #4376
+#4824 := [unit-resolution #4823 #4821]: #4374
+#4444 := (not #4374)
+#4827 := (or #4444 #4450)
+#4825 := [hypothesis]: #1843
+#4461 := (or #3954 #1842 #4444 #4450)
+#4439 := (+ #4370 #3504)
+#4440 := (+ #1833 #4439)
+#4443 := (>= #4440 0::int)
+#4445 := (or #4444 #1842 #4443)
+#4462 := (or #3954 #4445)
+#4469 := (iff #4462 #4461)
+#4456 := (or #1842 #4444 #4450)
+#4464 := (or #3954 #4456)
+#4467 := (iff #4464 #4461)
+#4468 := [rewrite]: #4467
+#4465 := (iff #4462 #4464)
+#4459 := (iff #4445 #4456)
+#4453 := (or #4444 #1842 #4450)
+#4457 := (iff #4453 #4456)
+#4458 := [rewrite]: #4457
+#4454 := (iff #4445 #4453)
+#4451 := (iff #4443 #4450)
+#4448 := (= #4440 #4447)
+#4449 := [rewrite]: #4448
+#4452 := [monotonicity #4449]: #4451
+#4455 := [monotonicity #4452]: #4454
+#4460 := [trans #4455 #4458]: #4459
+#4466 := [monotonicity #4460]: #4465
+#4470 := [trans #4466 #4468]: #4469
+#4463 := [quant-inst]: #4462
+#4471 := [mp #4463 #4470]: #4461
+#4828 := [unit-resolution #4471 #4826 #4825]: #4827
+#4829 := [unit-resolution #4828 #4824 #4819]: false
+#4831 := [lemma #4829]: #4830
+#4684 := [unit-resolution #4831 #4683 #4629 #4826 #4822 #4628 #4626 #4620]: false
+#4686 := [lemma #4684]: #4685
+#5257 := [unit-resolution #4686 #5256 #5276 #5210]: #2992
+#3791 := (or #4032 #4026 #2997)
+#3513 := [def-axiom]: #3791
+#5258 := [unit-resolution #3513 #5257 #5275]: #4032
+#3516 := (or #4035 #4029)
+#3409 := [def-axiom]: #3516
+#5259 := [unit-resolution #3409 #5258]: #4035
+#3785 := (or #4044 #4004 #4038)
+#3789 := [def-axiom]: #3785
+#5260 := [unit-resolution #3789 #5259 #5270]: #4004
+#3466 := (or #4001 #1777)
+#3471 := [def-axiom]: #3466
+#5255 := [unit-resolution #3471 #5260]: #1777
+#4669 := (+ #1764 #4302)
+#4671 := (>= #4669 0::int)
+#4668 := (= #1764 #4301)
+#5261 := (= #4301 #1764)
+#5262 := [monotonicity #5210]: #5261
+#5263 := [symm #5262]: #4668
+#5264 := (not #4668)
+#5265 := (or #5264 #4671)
+#5280 := [th-lemma]: #5265
+#5281 := [unit-resolution #5280 #5263]: #4671
+#4676 := (not #4671)
+#4677 := (or #4675 #4676 #1776)
+#4672 := [hypothesis]: #1777
+#4667 := [hypothesis]: #4283
+#4673 := [hypothesis]: #4671
+#4674 := [th-lemma #4673 #4667 #4672]: false
+#4692 := [lemma #4674]: #4677
+#5302 := [unit-resolution #4692 #5281 #5255]: #4675
+#5305 := (or #4283 #4442)
+#3807 := (or #4001 #2345)
+#3467 := [def-axiom]: #3807
+#5303 := [unit-resolution #3467 #5260]: #2345
+#3669 := (or #4214 #3974)
+#3637 := [def-axiom]: #3669
+#5304 := [unit-resolution #3637 #5207]: #3974
+#4496 := (or #3979 #2342 #4283 #4442)
+#4350 := (+ #4369 #4355)
+#4356 := (+ #4301 #4350)
+#4357 := (= #4356 0::int)
+#4358 := (not #4357)
+#4429 := (or #4360 #4358 #4352)
+#4430 := (not #4429)
+#4431 := (or #1778 #4283 #4430)
+#4502 := (or #3979 #4431)
+#4517 := (iff #4502 #4496)
+#4499 := (or #2342 #4283 #4442)
+#4504 := (or #3979 #4499)
+#4514 := (iff #4504 #4496)
+#4516 := [rewrite]: #4514
+#4505 := (iff #4502 #4504)
+#4500 := (iff #4431 #4499)
+#4497 := (iff #4430 #4442)
+#4426 := (iff #4429 #4423)
+#4421 := (or #4360 #4418 #4352)
+#4424 := (iff #4421 #4423)
+#4425 := [rewrite]: #4424
+#4416 := (iff #4429 #4421)
+#4419 := (iff #4358 #4418)
+#4437 := (iff #4357 #4436)
+#4434 := (= #4356 #4433)
+#4435 := [rewrite]: #4434
+#4417 := [monotonicity #4435]: #4437
+#4420 := [monotonicity #4417]: #4419
+#4422 := [monotonicity #4420]: #4416
+#4441 := [trans #4422 #4425]: #4426
+#4498 := [monotonicity #4441]: #4497
+#4501 := [monotonicity #2344 #4498]: #4500
+#4506 := [monotonicity #4501]: #4505
+#4518 := [trans #4506 #4516]: #4517
+#4503 := [quant-inst]: #4502
+#4519 := [mp #4503 #4518]: #4496
+#5306 := [unit-resolution #4519 #5304 #5303]: #5305
+#5301 := [unit-resolution #5306 #5302]: #4442
+#4531 := (or #4423 #4436)
+#4533 := [def-axiom]: #4531
+#5307 := [unit-resolution #4533 #5301]: #4436
+#5308 := (or #4418 #4520)
+#5309 := [th-lemma]: #5308
+#5310 := [unit-resolution #5309 #5307]: #4520
+#4670 := (<= #4669 0::int)
+#5311 := (or #5264 #4670)
+#5337 := [th-lemma]: #5311
+#5338 := [unit-resolution #5337 #5263]: #4670
+#5016 := (>= #5011 0::int)
+#5339 := (or #5214 #5016)
+#5340 := [th-lemma]: #5339
+#5341 := [unit-resolution #5340 #5213]: #5016
+#4521 := (not #4352)
+#4522 := (or #4423 #4521)
+#4523 := [def-axiom]: #4522
+#5336 := [unit-resolution #4523 #5301]: #4521
+#3808 := (or #4001 #3993)
+#3451 := [def-axiom]: #3808
+#5342 := [unit-resolution #3451 #5260]: #3993
+#4766 := (+ #4355 #4751)
+#4772 := (+ #1764 #4766)
+#4799 := (>= #4772 0::int)
+#4515 := (>= #4433 0::int)
+#5343 := (or #4418 #4515)
+#5344 := [th-lemma]: #5343
+#5345 := [unit-resolution #5344 #5307]: #4515
+#5402 := (not #5016)
+#5346 := (not #4515)
+#5355 := (or #4799 #4676 #5346 #5402)
+#5356 := [th-lemma]: #5355
+#5357 := [unit-resolution #5356 #5281 #5345 #5341]: #4799
+#5412 := (not #5015)
+#5411 := (not #4520)
+#5410 := (not #4670)
+#5417 := (not #4799)
+#5424 := (or #5417 #3998 #4676 #4352 #5402 #5410 #5411 #5412)
+#4752 := (+ #1764 #4751)
+#4753 := (<= #4752 0::int)
+#5401 := (not #4753)
+#5399 := [hypothesis]: #5016
+#5400 := [hypothesis]: #4521
+#5403 := (or #5401 #4676 #4352 #5402)
+#5404 := [th-lemma]: #5403
+#5405 := [unit-resolution #5404 #4673 #5400 #5399]: #5401
+#4773 := (= #4772 0::int)
+#5406 := [hypothesis]: #4799
+#4798 := (<= #4772 0::int)
+#5407 := [hypothesis]: #5015
+#5408 := [hypothesis]: #4520
+#5409 := [hypothesis]: #4670
+#5413 := (or #4798 #5410 #5411 #5412)
+#5414 := [th-lemma]: #5413
+#5415 := [unit-resolution #5414 #5409 #5408 #5407]: #4798
+#5416 := (not #4798)
+#5418 := (or #4773 #5416 #5417)
+#5419 := [th-lemma]: #5418
+#5420 := [unit-resolution #5419 #5415 #5406]: #4773
+#4784 := (not #4773)
+#4786 := (or #4753 #4784)
+#5421 := [hypothesis]: #3993
+#4792 := (or #3998 #4753 #4784)
+#4693 := (+ #1765 #4354)
+#4715 := (+ #4714 #4693)
+#4716 := (= #4715 0::int)
+#4717 := (not #4716)
+#4718 := (+ #4714 #1765)
+#4713 := (>= #4718 0::int)
+#4719 := (or #4713 #4717)
+#4795 := (or #3998 #4719)
+#4793 := (iff #4795 #4792)
+#4630 := (or #3998 #4786)
+#4679 := (iff #4630 #4792)
+#4788 := [rewrite]: #4679
+#4631 := (iff #4795 #4630)
+#4787 := (iff #4719 #4786)
+#4782 := (iff #4717 #4784)
+#4776 := (iff #4716 #4773)
+#4757 := (+ #4354 #4714)
+#4758 := (+ #1765 #4757)
+#4769 := (= #4758 0::int)
+#4774 := (iff #4769 #4773)
+#4775 := [rewrite]: #4774
+#4770 := (iff #4716 #4769)
+#4767 := (= #4715 #4758)
+#4768 := [rewrite]: #4767
+#4771 := [monotonicity #4768]: #4770
+#4783 := [trans #4771 #4775]: #4776
+#4785 := [monotonicity #4783]: #4782
+#4755 := (iff #4713 #4753)
+#4720 := (+ #1765 #4714)
+#4723 := (>= #4720 0::int)
+#4748 := (iff #4723 #4753)
+#4754 := [rewrite]: #4748
+#4749 := (iff #4713 #4723)
+#4721 := (= #4718 #4720)
+#4722 := [rewrite]: #4721
+#4750 := [monotonicity #4722]: #4749
+#4756 := [trans #4750 #4754]: #4755
+#4791 := [monotonicity #4756 #4785]: #4787
+#4678 := [monotonicity #4791]: #4631
+#4794 := [trans #4678 #4788]: #4793
+#4796 := [quant-inst]: #4795
+#4797 := [mp #4796 #4794]: #4792
+#5422 := [unit-resolution #4797 #5421]: #4786
+#5423 := [unit-resolution #5422 #5420 #5405]: false
+#5425 := [lemma #5423]: #5424
+#5358 := [unit-resolution #5425 #5357 #5342 #5281 #5336 #5341 #5338 #5310 #5269]: false
+#5354 := [lemma #5358]: #5359
+#6215 := [unit-resolution #5354 #6214]: #4205
+#3700 := (or #4202 #192)
+#3686 := [def-axiom]: #3700
+#7417 := [unit-resolution #3686 #6215]: #192
+#7413 := [symm #7417]: #7418
+#8694 := (= #8048 #191)
+#9757 := (= #6904 uf_22)
+#6905 := (= uf_22 #6904)
+#12 := (uf_1 #10 #11)
+#3825 := (pattern #12)
#13 := (uf_3 #12)
-#309 := (= #11 #13)
-#4071 := (forall (vars (?x2 T2) (?x3 T2)) (:pat #4070) #309)
-#313 := (forall (vars (?x2 T2) (?x3 T2)) #309)
-#4074 := (iff #313 #4071)
-#4072 := (iff #309 #309)
-#4073 := [refl]: #4072
-#4075 := [quant-intro #4073]: #4074
-#1730 := (~ #313 #313)
-#1762 := (~ #309 #309)
-#1763 := [refl]: #1762
-#1728 := [nnf-pos #1763]: #1730
+#293 := (= #11 #13)
+#3826 := (forall (vars (?x2 T2) (?x3 T2)) (:pat #3825) #293)
+#297 := (forall (vars (?x2 T2) (?x3 T2)) #293)
+#3829 := (iff #297 #3826)
+#3827 := (iff #293 #293)
+#3828 := [refl]: #3827
+#3830 := [quant-intro #3828]: #3829
+#1500 := (~ #297 #297)
+#1532 := (~ #293 #293)
+#1533 := [refl]: #1532
+#1498 := [nnf-pos #1533]: #1500
#14 := (= #13 #11)
#15 := (forall (vars (?x2 T2) (?x3 T2)) #14)
-#314 := (iff #15 #313)
-#311 := (iff #14 #309)
-#312 := [rewrite]: #311
-#315 := [quant-intro #312]: #314
-#308 := [asserted]: #15
-#318 := [mp #308 #315]: #313
-#1764 := [mp~ #318 #1728]: #313
-#4076 := [mp #1764 #4075]: #4071
-#7845 := (not #4071)
-#10578 := (or #7845 #10572)
-#10579 := [quant-inst]: #10578
-#13434 := [unit-resolution #10579 #4076]: #10572
-#13436 := [symm #13434]: #13435
-#13611 := [monotonicity #13436]: #13610
-#13613 := [trans #13611 #10470]: #13612
-#27317 := [monotonicity #13613 #27305]: #27316
-#27319 := [symm #27317]: #27318
-#27321 := [monotonicity #27319]: #27320
-#27315 := [hypothesis]: #16890
-#27322 := [mp #27315 #27321]: #27198
-#27164 := (= #10571 #19932)
-#25982 := (up_6 uf_15 #19932)
-#27170 := (or #25982 #27164)
-#27175 := (iff #27162 #27170)
-#30 := (:var 1 T5)
-#20 := (:var 2 T2)
-#29 := (:var 3 T4)
-#31 := (uf_7 #29 #20 #30)
-#32 := (up_6 #31 #11)
-#4090 := (pattern #32)
-#35 := (up_6 #29 #11)
-#34 := (= #30 uf_8)
-#24 := (= #11 #20)
-#36 := (ite #24 #34 #35)
-#37 := (iff #32 #36)
-#4091 := (forall (vars (?x10 T4) (?x11 T2) (?x12 T5) (?x13 T2)) (:pat #4090) #37)
-#38 := (forall (vars (?x10 T4) (?x11 T2) (?x12 T5) (?x13 T2)) #37)
-#4094 := (iff #38 #4091)
-#4092 := (iff #37 #37)
-#4093 := [refl]: #4092
-#4095 := [quant-intro #4093]: #4094
-#1735 := (~ #38 #38)
-#1771 := (~ #37 #37)
-#1772 := [refl]: #1771
-#1736 := [nnf-pos #1772]: #1735
-#325 := [asserted]: #38
-#1773 := [mp~ #325 #1736]: #38
-#4096 := [mp #1773 #4095]: #4091
-#6627 := (not #4091)
-#27178 := (or #6627 #27175)
-#3770 := (= uf_8 uf_8)
-#27158 := (= #19932 #10571)
-#27159 := (ite #27158 #3770 #25982)
-#27163 := (iff #27162 #27159)
-#27179 := (or #6627 #27163)
-#27181 := (iff #27179 #27178)
-#27183 := (iff #27178 #27178)
-#27184 := [rewrite]: #27183
-#27176 := (iff #27163 #27175)
-#27173 := (iff #27159 #27170)
-#27167 := (ite #27164 true #25982)
-#27171 := (iff #27167 #27170)
-#27172 := [rewrite]: #27171
-#27168 := (iff #27159 #27167)
-#3773 := (iff #3770 true)
-#3762 := [rewrite]: #3773
-#27165 := (iff #27158 #27164)
-#27166 := [rewrite]: #27165
-#27169 := [monotonicity #27166 #3762]: #27168
-#27174 := [trans #27169 #27172]: #27173
-#27177 := [monotonicity #27174]: #27176
-#27182 := [monotonicity #27177]: #27181
-#27185 := [trans #27182 #27184]: #27181
-#27180 := [quant-inst]: #27179
-#27186 := [mp #27180 #27185]: #27178
-#27285 := [unit-resolution #27186 #4096]: #27175
-#27195 := (not #27175)
-#27312 := (or #27195 #27162)
-#5924 := (up_6 uf_15 #5912)
-#27308 := (iff #5924 #25982)
-#27306 := (iff #25982 #5924)
-#27307 := [monotonicity #27305]: #27306
-#27309 := [symm #27307]: #27308
-#5925 := (not #5924)
-#5917 := (uf_1 #5912 ?x75!20)
-#5918 := (uf_10 #5917)
-#5919 := (* -1::int #5918)
-#5913 := (uf_4 uf_14 #5912)
-#5914 := (* -1::int #5913)
-#5920 := (+ #5914 #5919)
-#5650 := (uf_4 uf_14 ?x75!20)
-#5921 := (+ #5650 #5920)
-#5922 := (= #5921 0::int)
-#5923 := (not #5922)
-#5915 := (+ #5650 #5914)
-#5916 := (<= #5915 0::int)
-#5931 := (or #5916 #5923 #5925)
-#5934 := (not #5931)
-#5680 := (* -1::int #5650)
-#5928 := (+ uf_9 #5680)
-#5929 := (<= #5928 0::int)
-#22674 := (not #5929)
-#6611 := [hypothesis]: #3329
-#4009 := (not #2772)
-#4010 := (or #3324 #4009)
-#4004 := [def-axiom]: #4010
-#6612 := [unit-resolution #4004 #6611]: #4009
-#13689 := (or #3324 #2772)
-#6525 := (uf_1 uf_22 ?x68!16)
-#6526 := (uf_10 #6525)
-#6551 := (+ #2770 #6526)
-#6552 := (+ #182 #6551)
-#13122 := (<= #6552 0::int)
-#6555 := (= #6552 0::int)
-#6492 := (uf_4 uf_14 ?x68!16)
-#6509 := (* -1::int #6492)
-#6544 := (+ #6509 #6526)
-#6545 := (+ #182 #6544)
-#6546 := (>= #6545 0::int)
-#6530 := (* -1::int #6526)
-#6534 := (+ uf_9 #6530)
-#6535 := (<= #6534 0::int)
-#6581 := (or #6535 #6546)
-#6584 := (not #6581)
-#6578 := (= #2184 #6492)
-#9424 := (not #6578)
-#6510 := (+ #2184 #6509)
-#13145 := (>= #6510 0::int)
-#13450 := (not #13145)
-#13018 := (= ?x67!17 #10571)
-#6294 := (up_6 uf_15 ?x67!17)
-#13025 := (or #6294 #13018)
-#13020 := (up_6 #11533 ?x67!17)
-#13030 := (iff #13020 #13025)
-#12984 := (or #6627 #13030)
-#13019 := (ite #13018 #3770 #6294)
-#13021 := (iff #13020 #13019)
-#12985 := (or #6627 #13021)
-#12987 := (iff #12985 #12984)
-#12983 := (iff #12984 #12984)
-#12989 := [rewrite]: #12983
-#13031 := (iff #13021 #13030)
-#13028 := (iff #13019 #13025)
-#13022 := (ite #13018 true #6294)
-#13026 := (iff #13022 #13025)
-#13027 := [rewrite]: #13026
-#13023 := (iff #13019 #13022)
-#13024 := [monotonicity #3762]: #13023
-#13029 := [trans #13024 #13027]: #13028
-#13032 := [monotonicity #13029]: #13031
-#12988 := [monotonicity #13032]: #12987
-#12990 := [trans #12988 #12989]: #12987
-#12986 := [quant-inst]: #12985
-#12991 := [mp #12986 #12990]: #12984
-#13478 := [unit-resolution #12991 #4096]: #13030
-#11491 := (iff #2187 #13020)
-#13479 := (iff #13020 #2187)
-#11490 := [monotonicity #13613]: #13479
-#13624 := [symm #11490]: #11491
-#3864 := (or #3324 #2187)
-#3865 := [def-axiom]: #3864
-#6614 := [unit-resolution #3865 #6611]: #2187
-#13625 := [mp #6614 #13624]: #13020
-#13051 := (not #13020)
-#13048 := (not #13030)
-#13052 := (or #13048 #13051 #13025)
-#13053 := [def-axiom]: #13052
-#13626 := [unit-resolution #13053 #13625 #13478]: #13025
-#11064 := [hypothesis]: #4009
-#6322 := (+ #182 #6509)
-#6323 := (<= #6322 0::int)
-#3838 := (up_6 uf_15 ?x68!16)
-#12920 := (not #3838)
-#12850 := (= ?x68!16 #10571)
-#12856 := (or #3838 #12850)
-#12925 := (not #12856)
-#12852 := (up_6 #11533 ?x68!16)
-#12885 := (iff #12852 #12856)
-#12882 := (or #6627 #12885)
-#12851 := (ite #12850 #3770 #3838)
-#12853 := (iff #12852 #12851)
-#12888 := (or #6627 #12853)
-#12890 := (iff #12888 #12882)
-#12892 := (iff #12882 #12882)
-#12917 := [rewrite]: #12892
-#12886 := (iff #12853 #12885)
-#12883 := (iff #12851 #12856)
-#12848 := (ite #12850 true #3838)
-#12857 := (iff #12848 #12856)
-#12858 := [rewrite]: #12857
-#12854 := (iff #12851 #12848)
-#12855 := [monotonicity #3762]: #12854
-#12884 := [trans #12855 #12858]: #12883
-#12887 := [monotonicity #12884]: #12886
-#12891 := [monotonicity #12887]: #12890
-#12918 := [trans #12891 #12917]: #12890
-#12889 := [quant-inst]: #12888
-#12919 := [mp #12889 #12918]: #12882
-#13654 := [unit-resolution #12919 #4096]: #12885
-#12955 := (not #12852)
-#13653 := (iff #2189 #12955)
-#13657 := (iff #2188 #12852)
-#13655 := (iff #12852 #2188)
-#13656 := [monotonicity #13613]: #13655
-#13658 := [symm #13656]: #13657
-#13659 := [monotonicity #13658]: #13653
-#4007 := (or #3324 #2189)
-#4008 := [def-axiom]: #4007
-#6613 := [unit-resolution #4008 #6611]: #2189
-#13660 := [mp #6613 #13659]: #12955
-#12952 := (not #12885)
-#12953 := (or #12952 #12852 #12925)
-#12954 := [def-axiom]: #12953
-#13661 := [unit-resolution #12954 #13660 #13654]: #12925
-#12921 := (or #12856 #12920)
-#12916 := [def-axiom]: #12921
-#13662 := [unit-resolution #12916 #13661]: #12920
-#6327 := (or #3838 #6323)
-#3927 := (or #4441 #4292)
-#3928 := [def-axiom]: #3927
-#10883 := [unit-resolution #3928 #10456]: #4292
-#13176 := (or #4297 #3838 #6323)
-#6340 := (+ #6492 #1357)
-#6341 := (>= #6340 0::int)
-#6342 := (or #3838 #6341)
-#13174 := (or #4297 #6342)
-#13184 := (iff #13174 #13176)
-#13180 := (or #4297 #6327)
-#13182 := (iff #13180 #13176)
-#13183 := [rewrite]: #13182
-#13173 := (iff #13174 #13180)
-#6328 := (iff #6342 #6327)
-#6325 := (iff #6341 #6323)
-#6343 := (+ #1357 #6492)
-#6346 := (>= #6343 0::int)
-#6321 := (iff #6346 #6323)
-#6324 := [rewrite]: #6321
-#6347 := (iff #6341 #6346)
-#6344 := (= #6340 #6343)
-#6345 := [rewrite]: #6344
-#6348 := [monotonicity #6345]: #6347
-#6326 := [trans #6348 #6324]: #6325
-#6329 := [monotonicity #6326]: #6328
-#13181 := [monotonicity #6329]: #13173
-#13185 := [trans #13181 #13183]: #13184
-#13179 := [quant-inst]: #13174
-#13187 := [mp #13179 #13185]: #13176
-#13663 := [unit-resolution #13187 #10883]: #6327
-#10648 := [unit-resolution #13663 #13662]: #6323
-#13045 := (not #13025)
-#13449 := (not #6323)
-#13468 := (or #13450 #2772 #13449 #13045)
-#4615 := (uf_24 uf_22)
-#4656 := (* -1::int #4615)
-#6243 := (+ #2182 #4656)
-#13424 := (<= #6243 0::int)
-#13423 := (= #2182 #4615)
-#6295 := (= ?x67!17 uf_22)
-#13432 := [hypothesis]: #13025
-#10379 := (not #6294)
-#10904 := (uf_4 uf_14 #10571)
-#10931 := (* -1::int #10904)
-#6265 := (uf_4 uf_14 ?x67!17)
-#13376 := (+ #6265 #10931)
-#13377 := (<= #13376 0::int)
-#13553 := (not #13377)
-#13446 := [hypothesis]: #6323
-#5447 := (* -1::int #6265)
-#5547 := (+ #2182 #5447)
-#5548 := (<= #5547 0::int)
-#3804 := (or #4441 #4435)
-#3915 := [def-axiom]: #3804
-#10913 := [unit-resolution #3915 #10456]: #4435
-#3936 := (or #4441 #4302)
-#3909 := [def-axiom]: #3936
-#10462 := [unit-resolution #3909 #10456]: #4302
-#3910 := (or #4441 #4310)
-#3911 := [def-axiom]: #3910
-#10914 := [unit-resolution #3911 #10456]: #4310
-#6880 := (or #2739 #4315 #4307)
-#6572 := (uf_1 uf_22 ?x61!13)
-#6573 := (uf_10 #6572)
-#6656 := (+ #2125 #6573)
-#6657 := (+ #182 #6656)
-#6677 := (>= #6657 0::int)
-#6659 := (= #6657 0::int)
-#6633 := (* -1::int #6573)
-#6629 := (+ uf_9 #6633)
-#6637 := (<= #6629 0::int)
-#6714 := (not #6637)
-#6647 := (+ #2737 #6573)
-#6642 := (+ #182 #6647)
-#6648 := (>= #6642 0::int)
-#6685 := (or #6637 #6648)
-#6687 := (not #6685)
-#6682 := (= #2124 #2126)
-#6838 := (not #6682)
-#6822 := [hypothesis]: #2744
-#6841 := (or #6838 #2739)
-#6842 := [th-lemma]: #6841
-#6837 := [unit-resolution #6842 #6822]: #6838
-#6843 := [hypothesis]: #4302
-#6692 := (or #4307 #6682 #6687)
-#6634 := (+ #1357 #6633)
-#6635 := (+ #2126 #6634)
-#6636 := (<= #6635 0::int)
-#6678 := (or #6637 #6636)
-#6680 := (not #6678)
-#6681 := (= #2126 #2124)
-#6679 := (or #6681 #6680)
-#6693 := (or #4307 #6679)
-#6710 := (iff #6693 #6692)
-#6690 := (or #6682 #6687)
-#6695 := (or #4307 #6690)
-#6708 := (iff #6695 #6692)
-#6709 := [rewrite]: #6708
-#6706 := (iff #6693 #6695)
-#6653 := (iff #6679 #6690)
-#6688 := (iff #6680 #6687)
-#6686 := (iff #6678 #6685)
-#6651 := (iff #6636 #6648)
-#6639 := (+ #2126 #6633)
-#6640 := (+ #1357 #6639)
-#6644 := (<= #6640 0::int)
-#6649 := (iff #6644 #6648)
-#6650 := [rewrite]: #6649
-#6645 := (iff #6636 #6644)
-#6641 := (= #6635 #6640)
-#6643 := [rewrite]: #6641
-#6646 := [monotonicity #6643]: #6645
-#6652 := [trans #6646 #6650]: #6651
-#6654 := [monotonicity #6652]: #6686
-#6689 := [monotonicity #6654]: #6688
-#6683 := (iff #6681 #6682)
-#6684 := [rewrite]: #6683
-#6691 := [monotonicity #6684 #6689]: #6653
-#6707 := [monotonicity #6691]: #6706
-#6711 := [trans #6707 #6709]: #6710
-#6694 := [quant-inst]: #6693
-#6712 := [mp #6694 #6711]: #6692
-#6844 := [unit-resolution #6712 #6843 #6837]: #6687
-#6715 := (or #6685 #6714)
-#6716 := [def-axiom]: #6715
-#6845 := [unit-resolution #6716 #6844]: #6714
-#6717 := (not #6648)
-#6718 := (or #6685 #6717)
-#6719 := [def-axiom]: #6718
-#6846 := [unit-resolution #6719 #6844]: #6717
-#6663 := (or #6637 #6648 #6659)
-#6847 := [hypothesis]: #4310
-#6665 := (or #4315 #6637 #6648 #6659)
-#6631 := (+ #6573 #2125)
-#6632 := (+ #182 #6631)
-#6630 := (= #6632 0::int)
-#6638 := (or #6637 #6636 #6630)
-#6666 := (or #4315 #6638)
-#6674 := (iff #6666 #6665)
-#6669 := (or #4315 #6663)
-#6671 := (iff #6669 #6665)
-#6672 := [rewrite]: #6671
-#6667 := (iff #6666 #6669)
-#6661 := (iff #6638 #6663)
-#6660 := (iff #6630 #6659)
-#6655 := (= #6632 #6657)
-#6658 := [rewrite]: #6655
-#6662 := [monotonicity #6658]: #6660
-#6664 := [monotonicity #6652 #6662]: #6661
-#6670 := [monotonicity #6664]: #6667
-#6675 := [trans #6670 #6672]: #6674
-#6668 := [quant-inst]: #6666
-#6673 := [mp #6668 #6675]: #6665
-#6871 := [unit-resolution #6673 #6847]: #6663
-#6872 := [unit-resolution #6871 #6846 #6845]: #6659
-#6873 := (not #6659)
-#6874 := (or #6873 #6677)
-#6875 := [th-lemma]: #6874
-#6870 := [unit-resolution #6875 #6872]: #6677
-#6713 := (>= #2738 0::int)
-#6876 := (or #6713 #2739)
-#6877 := [th-lemma]: #6876
-#6878 := [unit-resolution #6877 #6822]: #6713
-#6879 := [th-lemma #6878 #6846 #6870]: false
-#6904 := [lemma #6879]: #6880
-#10915 := [unit-resolution #6904 #10914 #10462]: #2739
-#3942 := (or #4438 #2744 #4432)
-#3943 := [def-axiom]: #3942
-#10916 := [unit-resolution #3943 #10915 #10913]: #4432
-#3955 := (or #4429 #4318)
-#3956 := [def-axiom]: #3955
-#10924 := [unit-resolution #3956 #10916]: #4318
-#10586 := (or #4323 #5548)
-#5540 := (+ #6265 #2183)
-#5541 := (>= #5540 0::int)
-#10587 := (or #4323 #5541)
-#10591 := (iff #10587 #10586)
-#10593 := (iff #10586 #10586)
-#10594 := [rewrite]: #10593
-#5574 := (iff #5541 #5548)
-#5542 := (+ #2183 #6265)
-#5539 := (>= #5542 0::int)
-#5549 := (iff #5539 #5548)
-#5573 := [rewrite]: #5549
-#5545 := (iff #5541 #5539)
-#5543 := (= #5540 #5542)
-#5544 := [rewrite]: #5543
-#5546 := [monotonicity #5544]: #5545
-#5575 := [trans #5546 #5573]: #5574
-#10592 := [monotonicity #5575]: #10591
-#10595 := [trans #10592 #10594]: #10591
-#10590 := [quant-inst]: #10587
-#10596 := [mp #10590 #10595]: #10586
-#11152 := [unit-resolution #10596 #10924]: #5548
-#10639 := (+ #182 #10931)
-#10640 := (>= #10639 0::int)
-#10940 := (= #182 #10904)
-#13544 := (= #10904 #182)
-#13545 := [monotonicity #13436]: #13544
-#13546 := [symm #13545]: #10940
-#13547 := (not #10940)
-#13548 := (or #13547 #10640)
-#13549 := [th-lemma]: #13548
-#13550 := [unit-resolution #13549 #13546]: #10640
-#13447 := [hypothesis]: #13145
-#13420 := (not #10640)
-#11843 := (not #5548)
-#13451 := (or #13553 #11843 #2772 #13449 #13420 #13450)
-#13452 := [th-lemma]: #13451
-#13453 := [unit-resolution #13452 #13447 #13550 #11152 #13446 #11064]: #13553
-#13675 := (or #10379 #13377)
-#13664 := [hypothesis]: #13553
-#10929 := (up_6 uf_15 #10571)
-#13669 := (not #10929)
-#13670 := (iff #181 #13669)
-#13667 := (iff #180 #10929)
-#13665 := (iff #10929 #180)
-#13666 := [monotonicity #13436]: #13665
-#13668 := [symm #13666]: #13667
-#13671 := [monotonicity #13668]: #13670
-#3944 := (or #4441 #181)
-#3939 := [def-axiom]: #3944
-#10457 := [unit-resolution #3939 #10456]: #181
-#13672 := [mp #10457 #13671]: #13669
-#13673 := [hypothesis]: #6294
-#3888 := (or #4453 #4197)
-#3912 := [def-axiom]: #3888
-#10562 := [unit-resolution #3912 #10123]: #4197
-#13578 := (or #4202 #10379 #10929 #13377)
-#13339 := (+ #10904 #5447)
-#13340 := (>= #13339 0::int)
-#13371 := (or #10929 #10379 #13340)
-#13580 := (or #4202 #13371)
-#13591 := (iff #13580 #13578)
-#13395 := (or #10379 #10929 #13377)
-#13586 := (or #4202 #13395)
-#13589 := (iff #13586 #13578)
-#13590 := [rewrite]: #13589
-#13587 := (iff #13580 #13586)
-#13408 := (iff #13371 #13395)
-#13400 := (or #10929 #10379 #13377)
-#13404 := (iff #13400 #13395)
-#13407 := [rewrite]: #13404
-#13405 := (iff #13371 #13400)
-#13398 := (iff #13340 #13377)
-#13372 := (+ #5447 #10904)
-#13375 := (>= #13372 0::int)
-#13378 := (iff #13375 #13377)
-#13379 := [rewrite]: #13378
-#13369 := (iff #13340 #13375)
-#13373 := (= #13339 #13372)
-#13374 := [rewrite]: #13373
-#13370 := [monotonicity #13374]: #13369
-#13399 := [trans #13370 #13379]: #13398
-#13406 := [monotonicity #13399]: #13405
-#13409 := [trans #13406 #13407]: #13408
-#13588 := [monotonicity #13409]: #13587
-#13458 := [trans #13588 #13590]: #13591
-#13581 := [quant-inst]: #13580
-#13472 := [mp #13581 #13458]: #13578
-#13674 := [unit-resolution #13472 #10562 #13673 #13672 #13664]: false
-#13676 := [lemma #13674]: #13675
-#13448 := [unit-resolution #13676 #13453]: #10379
-#13046 := (or #13045 #6294 #13018)
-#13047 := [def-axiom]: #13046
-#13425 := [unit-resolution #13047 #13448 #13432]: #13018
-#13469 := [trans #13425 #13436]: #6295
-#13470 := [monotonicity #13469]: #13423
-#13463 := (not #13423)
-#13471 := (or #13463 #13424)
-#13527 := [th-lemma]: #13471
-#13464 := [unit-resolution #13527 #13470]: #13424
-#4857 := (+ #182 #4656)
-#4858 := (>= #4857 0::int)
-#9945 := (or #4323 #4858)
-#9946 := [quant-inst]: #9945
-#10925 := [unit-resolution #9946 #10924]: #4858
-#13467 := [th-lemma #11064 #13446 #13447 #10925 #13464]: false
-#13461 := [lemma #13467]: #13468
-#10666 := [unit-resolution #13461 #10648 #11064 #13626]: #13450
-#11458 := (or #9424 #13145)
-#13497 := [th-lemma]: #11458
-#13498 := [unit-resolution #13497 #10666]: #9424
-#6587 := (or #6578 #6584)
-#13124 := (or #4307 #6578 #6584)
-#6531 := (+ #1357 #6530)
-#6532 := (+ #6492 #6531)
-#6533 := (<= #6532 0::int)
-#6574 := (or #6535 #6533)
-#6575 := (not #6574)
-#6576 := (= #6492 #2184)
-#6577 := (or #6576 #6575)
-#13125 := (or #4307 #6577)
-#13142 := (iff #13125 #13124)
-#13138 := (or #4307 #6587)
-#13141 := (iff #13138 #13124)
-#13136 := [rewrite]: #13141
-#13139 := (iff #13125 #13138)
-#6588 := (iff #6577 #6587)
-#6585 := (iff #6575 #6584)
-#6582 := (iff #6574 #6581)
-#6549 := (iff #6533 #6546)
-#6537 := (+ #6492 #6530)
-#6538 := (+ #1357 #6537)
-#6541 := (<= #6538 0::int)
-#6547 := (iff #6541 #6546)
-#6548 := [rewrite]: #6547
-#6542 := (iff #6533 #6541)
-#6539 := (= #6532 #6538)
-#6540 := [rewrite]: #6539
-#6543 := [monotonicity #6540]: #6542
-#6550 := [trans #6543 #6548]: #6549
-#6583 := [monotonicity #6550]: #6582
-#6586 := [monotonicity #6583]: #6585
-#6579 := (iff #6576 #6578)
-#6580 := [rewrite]: #6579
-#6589 := [monotonicity #6580 #6586]: #6588
-#13140 := [monotonicity #6589]: #13139
-#13143 := [trans #13140 #13136]: #13142
-#13137 := [quant-inst]: #13125
-#13144 := [mp #13137 #13143]: #13124
-#13476 := [unit-resolution #13144 #10462]: #6587
-#13477 := [unit-resolution #13476 #13498]: #6584
-#13474 := (or #6581 #6555)
-#13523 := (not #6555)
-#13524 := [hypothesis]: #13523
-#13146 := (not #6535)
-#13528 := [hypothesis]: #6584
-#13166 := (or #6581 #13146)
-#13167 := [def-axiom]: #13166
-#13454 := [unit-resolution #13167 #13528]: #13146
-#13168 := (not #6546)
-#13169 := (or #6581 #13168)
-#13170 := [def-axiom]: #13169
-#13455 := [unit-resolution #13170 #13528]: #13168
-#6558 := (or #6535 #6546 #6555)
-#13101 := (or #4315 #6535 #6546 #6555)
-#6527 := (+ #6526 #2770)
-#6528 := (+ #182 #6527)
-#6529 := (= #6528 0::int)
-#6536 := (or #6535 #6533 #6529)
-#13102 := (or #4315 #6536)
-#13120 := (iff #13102 #13101)
-#13104 := (or #4315 #6558)
-#13118 := (iff #13104 #13101)
-#13119 := [rewrite]: #13118
-#13116 := (iff #13102 #13104)
-#6559 := (iff #6536 #6558)
-#6556 := (iff #6529 #6555)
-#6553 := (= #6528 #6552)
-#6554 := [rewrite]: #6553
-#6557 := [monotonicity #6554]: #6556
-#6560 := [monotonicity #6550 #6557]: #6559
-#13117 := [monotonicity #6560]: #13116
-#13115 := [trans #13117 #13119]: #13120
-#13103 := [quant-inst]: #13102
-#13121 := [mp #13103 #13115]: #13101
-#13456 := [unit-resolution #13121 #10914]: #6558
-#13457 := [unit-resolution #13456 #13455 #13454 #13524]: false
-#13475 := [lemma #13457]: #13474
-#13577 := [unit-resolution #13475 #13477]: #6555
-#13677 := (or #13523 #13122)
-#13678 := [th-lemma]: #13677
-#13679 := [unit-resolution #13678 #13577]: #13122
-#13013 := (uf_1 #10571 ?x68!16)
-#13014 := (uf_10 #13013)
-#13149 := (* -1::int #13014)
-#13526 := (+ #6526 #13149)
-#13530 := (>= #13526 0::int)
-#13525 := (= #6526 #13014)
-#13534 := (= #13014 #6526)
-#13532 := (= #13013 #6525)
-#13533 := [monotonicity #13436]: #13532
-#13535 := [monotonicity #13533]: #13534
-#13536 := [symm #13535]: #13525
-#13537 := (not #13525)
-#13538 := (or #13537 #13530)
-#13539 := [th-lemma]: #13538
-#13540 := [unit-resolution #13539 #13536]: #13530
-#13499 := (<= #13014 0::int)
-#13500 := (not #13499)
-#12922 := (not #12850)
-#12923 := (or #12856 #12922)
-#12924 := [def-axiom]: #12923
-#13680 := [unit-resolution #12924 #13661]: #12922
+#298 := (iff #15 #297)
+#295 := (iff #14 #293)
+#296 := [rewrite]: #295
+#299 := [quant-intro #296]: #298
+#292 := [asserted]: #15
+#302 := [mp #292 #299]: #297
+#1534 := [mp~ #302 #1498]: #297
+#3831 := [mp #1534 #3830]: #3826
+#5007 := (not #3826)
+#6321 := (or #5007 #6905)
+#6319 := [quant-inst]: #6321
+#9756 := [unit-resolution #6319 #3831]: #6905
+#9758 := [symm #9756]: #9757
+#8603 := [monotonicity #9758]: #8694
+#8592 := [trans #8603 #7413]: #8591
+#30523 := [monotonicity #8592]: #30522
+#30525 := [symm #30523]: #30524
+#30527 := [monotonicity #30525]: #30526
+#30518 := (not #10311)
+#30519 := [hypothesis]: #30518
+#10337 := (or #9640 #10311)
+#9157 := (= #185 #1939)
+#9114 := (= #1939 #185)
+#8831 := (= ?x63!14 uf_22)
+#8074 := (= ?x63!14 #6904)
+#8212 := (uf_1 #6904 ?x63!14)
+#8213 := (uf_10 #8212)
+#8269 := (<= #8213 0::int)
+#7863 := (up_6 uf_15 ?x63!14)
+#8055 := (or #7863 #8074)
+#8073 := (up_6 #8048 ?x63!14)
+#8128 := (iff #8055 #8073)
+#8076 := (or #6010 #8128)
+#8075 := (ite #8074 #5970 #7863)
+#8051 := (iff #8073 #8075)
+#8160 := (or #6010 #8051)
+#8176 := (iff #8160 #8076)
+#8178 := (iff #8076 #8076)
+#8204 := [rewrite]: #8178
+#8140 := (iff #8051 #8128)
+#8126 := (iff #8073 #8055)
+#8138 := (iff #8126 #8128)
+#8139 := [rewrite]: #8138
+#8122 := (iff #8051 #8126)
+#8124 := (iff #8075 #8055)
+#8052 := (ite #8074 true #7863)
+#8123 := (iff #8052 #8055)
+#8118 := [rewrite]: #8123
+#8053 := (iff #8075 #8052)
+#8054 := [monotonicity #5976]: #8053
+#8125 := [trans #8054 #8118]: #8124
+#8127 := [monotonicity #8125]: #8122
+#8158 := [trans #8127 #8139]: #8140
+#8177 := [monotonicity #8158]: #8176
+#8205 := [trans #8177 #8204]: #8176
+#8173 := [quant-inst]: #8160
+#8206 := [mp #8173 #8205]: #8076
+#8606 := [unit-resolution #8206 #3851]: #8128
+#8980 := (iff #1941 #8073)
+#8593 := (iff #8073 #1941)
+#8588 := [monotonicity #8592]: #8593
+#8983 := [symm #8588]: #8980
+#8453 := [hypothesis]: #2500
+#3769 := (or #2497 #1941)
+#3770 := [def-axiom]: #3769
+#8590 := [unit-resolution #3770 #8453]: #1941
+#9018 := [mp #8590 #8983]: #8073
+#8157 := (not #8073)
+#7280 := (not #8128)
+#7285 := (or #7280 #8055 #8157)
+#7639 := [def-axiom]: #7285
+#9019 := [unit-resolution #7639 #9018 #8606]: #8055
+#7891 := (uf_1 uf_22 ?x63!14)
+#7892 := (uf_10 #7891)
+#7831 := (* -1::int #1939)
+#7910 := (+ #7831 #7892)
+#7911 := (+ #185 #7910)
+#7912 := (>= #7911 0::int)
+#7886 := (not #7912)
+#7896 := (* -1::int #7892)
+#7900 := (+ uf_9 #7896)
+#7901 := (<= #7900 0::int)
+#7943 := (or #7901 #7912)
+#7946 := (not #7943)
+#3775 := (not #2494)
+#3776 := (or #2497 #3775)
+#3771 := [def-axiom]: #3776
+#8506 := [unit-resolution #3771 #8453]: #3775
+#3665 := (or #4202 #4071)
+#3666 := [def-axiom]: #3665
+#6216 := [unit-resolution #3666 #6215]: #4071
+#7928 := (or #4076 #2494 #7946)
+#7897 := (+ #1235 #7896)
+#7898 := (+ #1939 #7897)
+#7899 := (<= #7898 0::int)
+#7940 := (or #7901 #7899)
+#7941 := (not #7940)
+#7942 := (or #1940 #7941)
+#7929 := (or #4076 #7942)
+#7591 := (iff #7929 #7928)
+#7949 := (or #2494 #7946)
+#7931 := (or #4076 #7949)
+#7590 := (iff #7931 #7928)
+#7585 := [rewrite]: #7590
+#7855 := (iff #7929 #7931)
+#7950 := (iff #7942 #7949)
+#7947 := (iff #7941 #7946)
+#7944 := (iff #7940 #7943)
+#7915 := (iff #7899 #7912)
+#7903 := (+ #1939 #7896)
+#7904 := (+ #1235 #7903)
+#7907 := (<= #7904 0::int)
+#7913 := (iff #7907 #7912)
+#7914 := [rewrite]: #7913
+#7908 := (iff #7899 #7907)
+#7905 := (= #7898 #7904)
+#7906 := [rewrite]: #7905
+#7909 := [monotonicity #7906]: #7908
+#7916 := [trans #7909 #7914]: #7915
+#7945 := [monotonicity #7916]: #7944
+#7948 := [monotonicity #7945]: #7947
+#7951 := [monotonicity #2496 #7948]: #7950
+#7856 := [monotonicity #7951]: #7855
+#7592 := [trans #7856 #7585]: #7591
+#7930 := [quant-inst]: #7929
+#7582 := [mp #7930 #7592]: #7928
+#9020 := [unit-resolution #7582 #6216 #8506]: #7946
+#7887 := (or #7943 #7886)
+#7888 := [def-axiom]: #7887
+#9017 := [unit-resolution #7888 #9020]: #7886
+#6768 := (not #8055)
+#9798 := (or #8269 #7912 #6768)
+#8339 := (uf_3 #8212)
+#9705 := (uf_4 uf_14 #8339)
+#9706 := (* -1::int #9705)
+#8358 := (uf_4 uf_14 #6904)
+#9707 := (+ #8358 #9706)
+#9708 := (>= #9707 0::int)
+#9709 := (up_6 uf_15 #8339)
+#9753 := (iff #7863 #9709)
+#9751 := (iff #9709 #7863)
+#9749 := (= #8339 ?x63!14)
+#8340 := (= ?x63!14 #8339)
+#8389 := (or #5007 #8340)
+#8383 := [quant-inst]: #8389
+#9748 := [unit-resolution #8383 #3831]: #8340
+#9750 := [symm #9748]: #9749
+#9752 := [monotonicity #9750]: #9751
+#9754 := [symm #9752]: #9753
+#9739 := [hypothesis]: #8055
+#8159 := (not #8074)
+#8299 := (= #8213 0::int)
+#9741 := (not #8299)
+#8270 := (not #8269)
+#9740 := [hypothesis]: #8270
+#9742 := (or #9741 #8269)
+#9743 := [th-lemma]: #9742
+#9744 := [unit-resolution #9743 #9740]: #9741
+#8309 := (or #8159 #8299)
#56 := (uf_10 #12)
-#385 := (<= #56 0::int)
-#386 := (not #385)
+#57 := (= #56 0::int)
#55 := (= #10 #11)
-#389 := (or #55 #386)
-#4118 := (forall (vars (?x22 T2) (?x23 T2)) (:pat #4070) #389)
-#392 := (forall (vars (?x22 T2) (?x23 T2)) #389)
-#4121 := (iff #392 #4118)
-#4119 := (iff #389 #389)
-#4120 := [refl]: #4119
-#4122 := [quant-intro #4120]: #4121
-#1745 := (~ #392 #392)
-#1744 := (~ #389 #389)
-#1782 := [refl]: #1744
-#1746 := [nnf-pos #1782]: #1745
-#61 := (< 0::int #56)
#60 := (not #55)
-#62 := (implies #60 #61)
-#63 := (forall (vars (?x22 T2) (?x23 T2)) #62)
-#395 := (iff #63 #392)
-#379 := (or #55 #61)
-#382 := (forall (vars (?x22 T2) (?x23 T2)) #379)
-#393 := (iff #382 #392)
-#390 := (iff #379 #389)
-#387 := (iff #61 #386)
-#388 := [rewrite]: #387
-#391 := [monotonicity #388]: #390
-#394 := [quant-intro #391]: #393
-#383 := (iff #63 #382)
-#380 := (iff #62 #379)
-#381 := [rewrite]: #380
-#384 := [quant-intro #381]: #383
-#396 := [trans #384 #394]: #395
-#378 := [asserted]: #63
-#397 := [mp #378 #396]: #392
-#1783 := [mp~ #397 #1746]: #392
-#4123 := [mp #1783 #4122]: #4118
-#7140 := (not #4118)
-#13508 := (or #7140 #12850 #13500)
-#13501 := (= #10571 ?x68!16)
-#13502 := (or #13501 #13500)
-#13509 := (or #7140 #13502)
-#13516 := (iff #13509 #13508)
-#13505 := (or #12850 #13500)
-#13511 := (or #7140 #13505)
-#13514 := (iff #13511 #13508)
-#13515 := [rewrite]: #13514
-#13512 := (iff #13509 #13511)
-#13506 := (iff #13502 #13505)
-#13503 := (iff #13501 #12850)
-#13504 := [rewrite]: #13503
-#13507 := [monotonicity #13504]: #13506
-#13513 := [monotonicity #13507]: #13512
-#13517 := [trans #13513 #13515]: #13516
-#13510 := [quant-inst]: #13509
-#13518 := [mp #13510 #13517]: #13508
-#13681 := [unit-resolution #13518 #4123 #13680]: #13500
-#13552 := (not #13122)
-#13554 := (or #13552 #2772 #13553 #12850)
-#13531 := [hypothesis]: #13377
-#13541 := [hypothesis]: #12922
-#13542 := [unit-resolution #13518 #4123 #13541]: #13500
-#13543 := [hypothesis]: #13122
-#13551 := [th-lemma #13550 #11152 #11064 #13543 #13542 #13540 #13531]: false
-#13555 := [lemma #13551]: #13554
-#13682 := [unit-resolution #13555 #13679 #11064 #13680]: #13553
-#13683 := [unit-resolution #13676 #13682]: #10379
-#13684 := [unit-resolution #13047 #13683 #13626]: #13018
-#13685 := [trans #13684 #13436]: #6295
-#13686 := [monotonicity #13685]: #13423
-#13687 := [unit-resolution #13527 #13686]: #13424
-#13688 := [th-lemma #13687 #11064 #10925 #13681 #13540 #13679]: false
-#13690 := [lemma #13688]: #13689
-#13952 := [unit-resolution #13690 #6612 #6611]: false
-#13940 := [lemma #13952]: #3324
-#7565 := (uf_1 uf_22 ?x63!14)
-#8498 := (uf_2 #7565)
-#8617 := (up_6 uf_15 #8498)
-#8671 := (iff #8617 #180)
-#8652 := (iff #180 #8617)
-#8499 := (= uf_22 #8498)
-#8505 := (or #8504 #8499)
-#8506 := [quant-inst]: #8505
-#8783 := [unit-resolution #8506 #4082]: #8499
-#10566 := [monotonicity #8783]: #8652
-#10567 := [symm #10566]: #8671
-#7566 := (uf_10 #7565)
-#6801 := (* -1::int #2143)
-#7652 := (+ #6801 #7566)
-#7653 := (+ #182 #7652)
-#7726 := (>= #7653 0::int)
-#7921 := (not #7726)
-#7567 := (* -1::int #7566)
-#7569 := (+ uf_9 #7567)
-#7570 := (<= #7569 0::int)
-#7725 := (or #7570 #7726)
-#7720 := (not #7725)
-#4020 := (not #2747)
-#10460 := [hypothesis]: #2753
-#4021 := (or #2750 #4020)
-#4016 := [def-axiom]: #4021
-#10461 := [unit-resolution #4016 #10460]: #4020
-#7787 := (or #4307 #2747 #7720)
-#7561 := (+ #1357 #7567)
-#7568 := (+ #2143 #7561)
-#7563 := (<= #7568 0::int)
-#7571 := (or #7570 #7563)
-#7369 := (not #7571)
-#7462 := (or #2144 #7369)
-#7783 := (or #4307 #7462)
-#6906 := (iff #7783 #7787)
-#7785 := (or #2747 #7720)
-#7788 := (or #4307 #7785)
-#7813 := (iff #7788 #7787)
-#7809 := [rewrite]: #7813
-#7793 := (iff #7783 #7788)
-#7786 := (iff #7462 #7785)
-#7715 := (iff #7369 #7720)
-#7718 := (iff #7571 #7725)
-#7716 := (iff #7563 #7726)
-#7463 := (+ #2143 #7567)
-#7469 := (+ #1357 #7463)
-#7370 := (<= #7469 0::int)
-#7751 := (iff #7370 #7726)
-#7752 := [rewrite]: #7751
-#7458 := (iff #7563 #7370)
-#7470 := (= #7568 #7469)
-#7587 := [rewrite]: #7470
-#7459 := [monotonicity #7587]: #7458
-#7717 := [trans #7459 #7752]: #7716
-#7719 := [monotonicity #7717]: #7718
-#7721 := [monotonicity #7719]: #7715
-#7784 := [monotonicity #2749 #7721]: #7786
-#7812 := [monotonicity #7784]: #7793
-#7392 := [trans #7812 #7809]: #6906
-#7789 := [quant-inst]: #7783
-#7572 := [mp #7789 #7392]: #7787
-#10464 := [unit-resolution #7572 #10462 #10461]: #7720
-#7960 := (or #7725 #7921)
-#7961 := [def-axiom]: #7960
-#10466 := [unit-resolution #7961 #10464]: #7921
-#8277 := (= uf_22 ?x63!14)
-#8348 := (not #8277)
-#9109 := (or #8348 #7726)
-#7497 := (+ #182 #6801)
-#8974 := (>= #7497 0::int)
-#8973 := (= #182 #2143)
-#9039 := (= #2143 #182)
-#8218 := (= ?x63!14 uf_22)
-#9035 := [hypothesis]: #8277
-#9036 := [symm #9035]: #8218
-#9040 := [monotonicity #9036]: #9039
-#9072 := [symm #9040]: #8973
-#9073 := (not #8973)
-#9074 := (or #9073 #8974)
-#9070 := [th-lemma]: #9074
-#9066 := [unit-resolution #9070 #9072]: #8974
-#10210 := [hypothesis]: #7921
-#8497 := (>= #7566 0::int)
-#8487 := (= #7566 0::int)
-#8488 := (or #8348 #8487)
-#57 := (= #56 0::int)
-#369 := (or #60 #57)
-#4112 := (forall (vars (?x20 T2) (?x21 T2)) (:pat #4070) #369)
-#372 := (forall (vars (?x20 T2) (?x21 T2)) #369)
-#4115 := (iff #372 #4112)
-#4113 := (iff #369 #369)
-#4114 := [refl]: #4113
-#4116 := [quant-intro #4114]: #4115
-#1741 := (~ #372 #372)
-#1780 := (~ #369 #369)
-#1781 := [refl]: #1780
-#1742 := [nnf-pos #1781]: #1741
+#353 := (or #60 #57)
+#3867 := (forall (vars (?x20 T2) (?x21 T2)) (:pat #3825) #353)
+#356 := (forall (vars (?x20 T2) (?x21 T2)) #353)
+#3870 := (iff #356 #3867)
+#3868 := (iff #353 #353)
+#3869 := [refl]: #3868
+#3871 := [quant-intro #3869]: #3870
+#1511 := (~ #356 #356)
+#1550 := (~ #353 #353)
+#1551 := [refl]: #1550
+#1512 := [nnf-pos #1551]: #1511
#58 := (implies #55 #57)
#59 := (forall (vars (?x20 T2) (?x21 T2)) #58)
-#375 := (iff #59 #372)
-#348 := (= 0::int #56)
-#359 := (or #60 #348)
-#364 := (forall (vars (?x20 T2) (?x21 T2)) #359)
-#373 := (iff #364 #372)
-#370 := (iff #359 #369)
-#367 := (iff #348 #57)
-#368 := [rewrite]: #367
-#371 := [monotonicity #368]: #370
-#374 := [quant-intro #371]: #373
-#365 := (iff #59 #364)
-#362 := (iff #58 #359)
-#356 := (implies #55 #348)
-#360 := (iff #356 #359)
-#361 := [rewrite]: #360
-#357 := (iff #58 #356)
-#354 := (iff #57 #348)
-#355 := [rewrite]: #354
-#358 := [monotonicity #355]: #357
-#363 := [trans #358 #361]: #362
-#366 := [quant-intro #363]: #365
-#376 := [trans #366 #374]: #375
-#346 := [asserted]: #59
-#377 := [mp #346 #376]: #372
-#1743 := [mp~ #377 #1742]: #372
-#4117 := [mp #1743 #4116]: #4112
-#7157 := (not #4112)
-#8491 := (or #7157 #8348 #8487)
-#8492 := (or #7157 #8488)
-#8494 := (iff #8492 #8491)
-#8495 := [rewrite]: #8494
-#8493 := [quant-inst]: #8492
-#8496 := [mp #8493 #8495]: #8491
-#9075 := [unit-resolution #8496 #4117]: #8488
-#9076 := [unit-resolution #9075 #9035]: #8487
-#9104 := (not #8487)
-#9105 := (or #9104 #8497)
-#9071 := [th-lemma]: #9105
-#9106 := [unit-resolution #9071 #9076]: #8497
-#9103 := [th-lemma #9106 #10210 #9066]: false
-#9110 := [lemma #9103]: #9109
-#10467 := [unit-resolution #9110 #10466]: #8348
-#7253 := (up_6 uf_15 ?x63!14)
-#8282 := (or #7253 #8277)
-#8262 := (up_6 #188 ?x63!14)
-#8295 := (iff #8262 #8282)
-#8298 := (or #6627 #8295)
-#8219 := (ite #8218 #3770 #7253)
-#8276 := (iff #8262 #8219)
-#8293 := (or #6627 #8276)
-#8300 := (iff #8293 #8298)
-#8302 := (iff #8298 #8298)
-#8303 := [rewrite]: #8302
-#8296 := (iff #8276 #8295)
-#8285 := (iff #8219 #8282)
-#8280 := (ite #8277 true #7253)
-#8283 := (iff #8280 #8282)
-#8284 := [rewrite]: #8283
-#8275 := (iff #8219 #8280)
-#8278 := (iff #8218 #8277)
-#8279 := [rewrite]: #8278
-#8281 := [monotonicity #8279 #3762]: #8275
-#8294 := [trans #8281 #8284]: #8285
-#8297 := [monotonicity #8294]: #8296
-#8301 := [monotonicity #8297]: #8300
-#8335 := [trans #8301 #8303]: #8300
-#8299 := [quant-inst]: #8293
-#8336 := [mp #8299 #8335]: #8298
-#8642 := [unit-resolution #8336 #4096]: #8295
-#8763 := (iff #2145 #8262)
-#8755 := (iff #8262 #2145)
-#10471 := [monotonicity #10470]: #8755
-#10474 := [symm #10471]: #8763
-#4014 := (or #2750 #2145)
-#4015 := [def-axiom]: #4014
-#10468 := [unit-resolution #4015 #10460]: #2145
-#10560 := [mp #10468 #10474]: #8262
-#8379 := (not #8262)
-#8376 := (not #8295)
-#8380 := (or #8376 #8379 #8282)
-#8375 := [def-axiom]: #8380
-#10561 := [unit-resolution #8375 #10560 #8642]: #8282
-#8351 := (not #8282)
-#10563 := (or #8277 #7726 #8617 #8351)
-#10231 := (or #8277 #7726 #4202 #8617 #8351)
-#8489 := (uf_3 #7565)
-#10159 := (uf_4 uf_14 #8489)
-#10160 := (* -1::int #10159)
-#8614 := (uf_4 uf_14 #8498)
-#10161 := (+ #8614 #10160)
-#10162 := (>= #10161 0::int)
-#10163 := (up_6 uf_15 #8489)
-#10201 := (iff #7253 #10163)
-#10199 := (iff #10163 #7253)
-#10197 := (= #8489 ?x63!14)
-#8490 := (= ?x63!14 #8489)
-#8500 := (or #7845 #8490)
-#8501 := [quant-inst]: #8500
-#10196 := [unit-resolution #8501 #4076]: #8490
-#10198 := [symm #10196]: #10197
-#10200 := [monotonicity #10198]: #10199
-#10202 := [symm #10200]: #10201
-#10193 := [hypothesis]: #8282
-#10194 := [hypothesis]: #8348
-#8352 := (or #8351 #7253 #8277)
-#8353 := [def-axiom]: #8352
-#10195 := [unit-resolution #8353 #10194 #10193]: #7253
-#10203 := [mp #10195 #10202]: #10163
-#10164 := (not #10163)
-#10207 := (or #10162 #10164)
-#10204 := (not #8617)
-#10205 := [hypothesis]: #10204
-#10206 := [hypothesis]: #4197
-#10169 := (or #4202 #8617 #10162 #10164)
-#10165 := (or #8617 #10164 #10162)
-#10170 := (or #4202 #10165)
-#10177 := (iff #10170 #10169)
-#10166 := (or #8617 #10162 #10164)
-#10172 := (or #4202 #10166)
-#10175 := (iff #10172 #10169)
-#10176 := [rewrite]: #10175
-#10173 := (iff #10170 #10172)
-#10167 := (iff #10165 #10166)
-#10168 := [rewrite]: #10167
-#10174 := [monotonicity #10168]: #10173
-#10178 := [trans #10174 #10176]: #10177
-#10171 := [quant-inst]: #10170
-#10179 := [mp #10171 #10178]: #10169
-#10208 := [unit-resolution #10179 #10206 #10205]: #10207
-#10209 := [unit-resolution #10208 #10203]: #10162
-#8246 := (<= #7566 0::int)
-#8247 := (not #8246)
-#8249 := (or #8247 #8277)
-#8252 := (or #7140 #8247 #8277)
-#8248 := (or #8277 #8247)
-#8253 := (or #7140 #8248)
-#8311 := (iff #8253 #8252)
-#8304 := (or #7140 #8249)
-#8307 := (iff #8304 #8252)
-#8308 := [rewrite]: #8307
-#8305 := (iff #8253 #8304)
-#8250 := (iff #8248 #8249)
-#8251 := [rewrite]: #8250
-#8306 := [monotonicity #8251]: #8305
-#8436 := [trans #8306 #8308]: #8311
-#8255 := [quant-inst]: #8253
-#8486 := [mp #8255 #8436]: #8252
-#10211 := [unit-resolution #8486 #4123]: #8249
-#10212 := [unit-resolution #10211 #10194]: #8247
-#10213 := (or #8497 #8246)
-#10214 := [th-lemma]: #10213
-#10215 := [unit-resolution #10214 #10212]: #8497
-#10184 := (+ #2143 #10160)
-#10191 := (<= #10184 0::int)
-#10183 := (= #2143 #10159)
-#10216 := (= #10159 #2143)
-#10217 := [monotonicity #10198]: #10216
-#10218 := [symm #10217]: #10183
-#10219 := (not #10183)
-#10220 := (or #10219 #10191)
-#10221 := [th-lemma]: #10220
-#10222 := [unit-resolution #10221 #10218]: #10191
-#8625 := (* -1::int #8614)
-#8892 := (+ #182 #8625)
-#8905 := (>= #8892 0::int)
-#8891 := (= #182 #8614)
-#10223 := (= #8614 #182)
-#8787 := (= #8498 uf_22)
-#8788 := [symm #8783]: #8787
-#10224 := [monotonicity #8788]: #10223
-#10225 := [symm #10224]: #8891
-#10226 := (not #8891)
-#10227 := (or #10226 #8905)
-#10228 := [th-lemma]: #10227
-#10229 := [unit-resolution #10228 #10225]: #8905
-#10230 := [th-lemma #10229 #10222 #10215 #10210 #10209]: false
-#10232 := [lemma #10230]: #10231
-#10564 := [unit-resolution #10232 #10562]: #10563
-#10565 := [unit-resolution #10564 #10561 #10467 #10466]: #8617
-#10568 := [mp #10565 #10567]: #180
-#10569 := [unit-resolution #10457 #10568]: false
-#10570 := [lemma #10569]: #2750
-#3957 := (or #4429 #4423)
-#3958 := [def-axiom]: #3957
-#10917 := [unit-resolution #3958 #10916]: #4423
-#3953 := (or #4426 #2753 #4420)
-#3954 := [def-axiom]: #3953
-#10918 := [unit-resolution #3954 #10917]: #4423
-#10919 := [unit-resolution #10918 #10570]: #4420
-#11065 := (or #4417 #4396)
-#3926 := (or #4453 #4180)
-#3906 := [def-axiom]: #3926
-#9067 := [unit-resolution #3906 #10123]: #4180
-#7840 := (or #4417 #888 #4315 #4307 #4185 #4396 #1636)
-#4734 := (>= #182 0::int)
-#7810 := [hypothesis]: #4180
-#6390 := (or #4185 #4734)
-#6391 := [quant-inst]: #6390
-#7811 := [unit-resolution #6391 #7810]: #4734
-#7291 := (uf_1 uf_22 ?x65!15)
-#7292 := (uf_10 #7291)
-#4945 := (* -1::int #2166)
-#7344 := (+ #4945 #7292)
-#7345 := (+ #182 #7344)
-#7389 := (<= #7345 0::int)
-#7362 := (= #7345 0::int)
-#6919 := (uf_4 uf_14 ?x65!15)
-#7124 := (* -1::int #6919)
-#7338 := (+ #7124 #7292)
-#7339 := (+ #182 #7338)
-#7340 := (>= #7339 0::int)
-#7434 := (not #7340)
-#7295 := (* -1::int #7292)
-#7299 := (+ uf_9 #7295)
-#7350 := (<= #7299 0::int)
-#7440 := (or #7340 #7350)
-#7445 := (not #7440)
-#6927 := (= #2166 #6919)
-#8448 := (not #6927)
-#7125 := (+ #2166 #7124)
-#7959 := (>= #7125 0::int)
-#7967 := (not #7959)
-#7578 := (>= #6919 0::int)
-#7581 := (or #4185 #7578)
-#7576 := [quant-inst]: #7581
-#7815 := [unit-resolution #7576 #7810]: #7578
-#7816 := [hypothesis]: #4393
-#7817 := [hypothesis]: #4420
-#3962 := (or #4417 #4411)
-#3966 := [def-axiom]: #3962
-#7818 := [unit-resolution #3966 #7817]: #4411
-#4798 := (= #105 #209)
-#7837 := (iff #4798 #210)
-#7836 := [commutativity]: #1392
-#7829 := (iff #4798 #713)
-#7814 := [hypothesis]: #106
-#7835 := [monotonicity #7814]: #7829
-#7838 := [trans #7835 #7836]: #7837
-#4810 := (<= #105 0::int)
-#7819 := (or #1636 #4810)
-#7830 := [th-lemma]: #7819
-#7831 := [unit-resolution #7830 #7814]: #4810
-#7201 := [hypothesis]: #189
-#3964 := (or #4417 #4327)
-#3965 := [def-axiom]: #3964
-#7832 := [unit-resolution #3965 #7817]: #4327
-#7250 := (not #4734)
-#7249 := (not #4810)
-#7251 := (or #4798 #7249 #7250 #888 #4332 #4307)
-#4756 := (uf_1 uf_22 uf_11)
-#4757 := (uf_10 #4756)
-#7072 := (<= #4757 0::int)
-#7073 := (not #7072)
-#4695 := (= uf_11 uf_22)
-#6917 := (not #4695)
-#4739 := (up_6 uf_15 uf_11)
-#7422 := (or #4695 #4739)
-#6926 := (not #7422)
-#7417 := (up_6 #188 uf_11)
-#7427 := (iff #7417 #7422)
-#3826 := (or #6627 #7427)
-#7416 := (ite #4695 #3770 #4739)
-#7418 := (iff #7417 #7416)
-#6907 := (or #6627 #7418)
-#6903 := (iff #6907 #3826)
-#6910 := (iff #3826 #3826)
-#6911 := [rewrite]: #6910
-#7428 := (iff #7418 #7427)
-#7425 := (iff #7416 #7422)
-#7419 := (ite #4695 true #4739)
-#7423 := (iff #7419 #7422)
-#7424 := [rewrite]: #7423
-#7420 := (iff #7416 #7419)
-#7421 := [monotonicity #3762]: #7420
-#7426 := [trans #7421 #7424]: #7425
-#7429 := [monotonicity #7426]: #7428
-#6909 := [monotonicity #7429]: #6903
-#6912 := [trans #6909 #6911]: #6903
-#6908 := [quant-inst]: #6907
-#6913 := [mp #6908 #6912]: #3826
-#7172 := [unit-resolution #6913 #4096]: #7427
-#6931 := (not #7417)
-#4884 := (up_6 uf_23 uf_11)
-#4885 := (not #4884)
-#7258 := (iff #4885 #6931)
-#7256 := (iff #4884 #7417)
-#7204 := (iff #7417 #4884)
-#7203 := [symm #7201]: #7202
-#7205 := [monotonicity #7203]: #7204
-#7257 := [symm #7205]: #7256
-#7259 := [monotonicity #7257]: #7258
-#7173 := (not #4798)
-#7198 := [hypothesis]: #7173
-#4887 := (or #4798 #4885)
-#7199 := [hypothesis]: #4327
-#6803 := (or #4332 #4798 #4885)
-#4886 := (or #4885 #4798)
-#6818 := (or #4332 #4886)
-#6809 := (iff #6818 #6803)
-#6817 := (or #4332 #4887)
-#6807 := (iff #6817 #6803)
-#6808 := [rewrite]: #6807
-#6820 := (iff #6818 #6817)
-#4888 := (iff #4886 #4887)
-#4889 := [rewrite]: #4888
-#6806 := [monotonicity #4889]: #6820
-#6810 := [trans #6806 #6808]: #6809
-#6819 := [quant-inst]: #6818
-#6805 := [mp #6819 #6810]: #6803
-#7200 := [unit-resolution #6805 #7199]: #4887
-#7195 := [unit-resolution #7200 #7198]: #4885
-#7260 := [mp #7195 #7259]: #6931
-#6929 := (not #7427)
-#6930 := (or #6929 #7417 #6926)
-#6925 := [def-axiom]: #6930
-#7261 := [unit-resolution #6925 #7260 #7172]: #6926
-#6918 := (or #7422 #6917)
-#6916 := [def-axiom]: #6918
-#7262 := [unit-resolution #6916 #7261]: #6917
-#7075 := (or #4695 #7073)
-#7078 := (or #7140 #4695 #7073)
-#4693 := (= uf_22 uf_11)
-#7074 := (or #4693 #7073)
-#7079 := (or #7140 #7074)
-#7067 := (iff #7079 #7078)
-#7063 := (or #7140 #7075)
-#7066 := (iff #7063 #7078)
-#7061 := [rewrite]: #7066
-#7064 := (iff #7079 #7063)
-#7076 := (iff #7074 #7075)
-#4696 := (iff #4693 #4695)
-#4697 := [rewrite]: #4696
-#7077 := [monotonicity #4697]: #7076
-#7065 := [monotonicity #7077]: #7064
-#7068 := [trans #7065 #7061]: #7067
-#7062 := [quant-inst]: #7079
-#7069 := [mp #7062 #7068]: #7078
-#7263 := [unit-resolution #7069 #4123]: #7075
-#7264 := [unit-resolution #7263 #7262]: #7073
-#4761 := (* -1::int #4757)
-#4762 := (+ #1357 #4761)
-#4763 := (+ #105 #4762)
-#4764 := (<= #4763 0::int)
-#6619 := (not #4764)
-#4765 := (+ uf_9 #4761)
-#4766 := (<= #4765 0::int)
-#4800 := (or #4764 #4766)
-#4803 := (not #4800)
-#4806 := (or #4798 #4803)
-#6415 := (or #4307 #4798 #4803)
-#4796 := (or #4766 #4764)
-#4797 := (not #4796)
-#4799 := (or #4798 #4797)
-#6444 := (or #4307 #4799)
-#6449 := (iff #6444 #6415)
-#6446 := (or #4307 #4806)
-#6443 := (iff #6446 #6415)
-#6448 := [rewrite]: #6443
-#6441 := (iff #6444 #6446)
-#4807 := (iff #4799 #4806)
-#4804 := (iff #4797 #4803)
-#4801 := (iff #4796 #4800)
-#4802 := [rewrite]: #4801
-#4805 := [monotonicity #4802]: #4804
-#4808 := [monotonicity #4805]: #4807
-#6447 := [monotonicity #4808]: #6441
-#6450 := [trans #6447 #6448]: #6449
-#6445 := [quant-inst]: #6444
-#6451 := [mp #6445 #6450]: #6415
-#7244 := [unit-resolution #6451 #6843]: #4806
-#7245 := [unit-resolution #7244 #7198]: #4803
-#6620 := (or #4800 #6619)
-#6621 := [def-axiom]: #6620
-#7246 := [unit-resolution #6621 #7245]: #6619
-#7247 := [hypothesis]: #4734
-#7248 := [hypothesis]: #4810
-#7243 := [th-lemma #7248 #7247 #7246 #7264]: false
-#7252 := [lemma #7243]: #7251
-#7833 := [unit-resolution #7252 #7832 #7811 #7201 #7831 #6843]: #4798
-#7834 := [mp #7833 #7838]: #210
-#3961 := (or #4414 #1394 #4408)
-#3963 := [def-axiom]: #3961
-#7839 := [unit-resolution #3963 #7834 #7818]: #4408
-#3968 := (or #4405 #4399)
-#3970 := [def-axiom]: #3968
-#7850 := [unit-resolution #3970 #7839]: #4399
-#3982 := (or #4402 #2168 #4396)
-#3976 := [def-axiom]: #3982
-#7851 := [unit-resolution #3976 #7850 #7816]: #2168
-#8005 := (not #7578)
-#8006 := (or #7967 #2167 #8005)
-#7964 := [hypothesis]: #7578
-#7965 := [hypothesis]: #7959
-#8460 := [hypothesis]: #2168
-#7966 := [th-lemma #8460 #7965 #7964]: false
-#8007 := [lemma #7966]: #8006
-#7852 := [unit-resolution #8007 #7851 #7815]: #7967
-#7885 := (or #8448 #7959)
-#7881 := [hypothesis]: #7967
-#7882 := [hypothesis]: #6927
-#7886 := [th-lemma]: #7885
-#7887 := [unit-resolution #7886 #7882 #7881]: false
-#7888 := [lemma #7887]: #7885
-#7879 := [unit-resolution #7888 #7852]: #8448
-#7448 := (or #6927 #7445)
-#7451 := (or #4307 #6927 #7445)
-#7296 := (+ #1357 #7295)
-#7297 := (+ #6919 #7296)
-#7298 := (<= #7297 0::int)
-#7393 := (or #7350 #7298)
-#7394 := (not #7393)
-#6920 := (= #6919 #2166)
-#7395 := (or #6920 #7394)
-#7452 := (or #4307 #7395)
-#7432 := (iff #7452 #7451)
-#7454 := (or #4307 #7448)
-#7457 := (iff #7454 #7451)
-#7431 := [rewrite]: #7457
-#7455 := (iff #7452 #7454)
-#7449 := (iff #7395 #7448)
-#7446 := (iff #7394 #7445)
-#7443 := (iff #7393 #7440)
-#7396 := (or #7350 #7340)
-#7441 := (iff #7396 #7440)
-#7442 := [rewrite]: #7441
-#7397 := (iff #7393 #7396)
-#7337 := (iff #7298 #7340)
-#7352 := (+ #6919 #7295)
-#7353 := (+ #1357 #7352)
-#7356 := (<= #7353 0::int)
-#7341 := (iff #7356 #7340)
-#7342 := [rewrite]: #7341
-#7357 := (iff #7298 #7356)
-#7354 := (= #7297 #7353)
-#7355 := [rewrite]: #7354
-#7358 := [monotonicity #7355]: #7357
-#7343 := [trans #7358 #7342]: #7337
-#7439 := [monotonicity #7343]: #7397
-#7444 := [trans #7439 #7442]: #7443
-#7447 := [monotonicity #7444]: #7446
-#6928 := (iff #6920 #6927)
-#6932 := [rewrite]: #6928
-#7450 := [monotonicity #6932 #7447]: #7449
-#7456 := [monotonicity #7450]: #7455
-#7430 := [trans #7456 #7431]: #7432
-#7453 := [quant-inst]: #7452
-#7433 := [mp #7453 #7430]: #7451
-#8450 := [unit-resolution #7433 #6843]: #7448
-#7883 := [unit-resolution #8450 #7879]: #7445
-#7435 := (or #7440 #7434)
-#7436 := [def-axiom]: #7435
-#7878 := [unit-resolution #7436 #7883]: #7434
-#7437 := (not #7350)
-#7438 := (or #7440 #7437)
-#7500 := [def-axiom]: #7438
-#7884 := [unit-resolution #7500 #7883]: #7437
-#7402 := (or #7340 #7350 #7362)
-#7407 := (or #4315 #7340 #7350 #7362)
-#7293 := (+ #7292 #4945)
-#7294 := (+ #182 #7293)
-#7289 := (= #7294 0::int)
-#7351 := (or #7350 #7298 #7289)
-#7408 := (or #4315 #7351)
-#7415 := (iff #7408 #7407)
-#7410 := (or #4315 #7402)
-#7413 := (iff #7410 #7407)
-#7414 := [rewrite]: #7413
-#7411 := (iff #7408 #7410)
-#7405 := (iff #7351 #7402)
-#7399 := (or #7350 #7340 #7362)
-#7403 := (iff #7399 #7402)
-#7404 := [rewrite]: #7403
-#7400 := (iff #7351 #7399)
-#7363 := (iff #7289 #7362)
-#7346 := (= #7294 #7345)
-#7347 := [rewrite]: #7346
-#7398 := [monotonicity #7347]: #7363
-#7401 := [monotonicity #7343 #7398]: #7400
-#7406 := [trans #7401 #7404]: #7405
-#7412 := [monotonicity #7406]: #7411
-#7390 := [trans #7412 #7414]: #7415
-#7409 := [quant-inst]: #7408
-#7391 := [mp #7409 #7390]: #7407
-#8454 := [unit-resolution #7391 #6847]: #7402
-#7713 := [unit-resolution #8454 #7884 #7878]: #7362
-#8456 := (not #7362)
-#8457 := (or #8456 #7389)
-#8458 := [th-lemma]: #8457
-#7714 := [unit-resolution #8458 #7713]: #7389
-#7790 := (>= #7292 0::int)
-#7701 := (<= #7292 0::int)
-#7702 := (not #7701)
-#7603 := (= uf_22 ?x65!15)
-#7464 := (not #7603)
-#7505 := (up_6 uf_15 ?x65!15)
-#7628 := (or #7505 #7603)
-#4543 := (not #7628)
-#7584 := (up_6 #188 ?x65!15)
-#7611 := (iff #7584 #7628)
-#7612 := (or #6627 #7611)
-#7597 := (= ?x65!15 uf_22)
-#7598 := (ite #7597 #3770 #7505)
-#7599 := (iff #7584 #7598)
-#7607 := (or #6627 #7599)
-#7614 := (iff #7607 #7612)
-#7467 := (iff #7612 #7612)
-#7468 := [rewrite]: #7467
-#7627 := (iff #7599 #7611)
-#7609 := (iff #7598 #7628)
-#7606 := (ite #7603 true #7505)
-#7608 := (iff #7606 #7628)
-#7586 := [rewrite]: #7608
-#7602 := (iff #7598 #7606)
-#7604 := (iff #7597 #7603)
-#7605 := [rewrite]: #7604
-#7620 := [monotonicity #7605 #3762]: #7602
-#7610 := [trans #7620 #7586]: #7609
-#7585 := [monotonicity #7610]: #7627
-#7615 := [monotonicity #7585]: #7614
-#7460 := [trans #7615 #7468]: #7614
-#7613 := [quant-inst]: #7607
-#7461 := [mp #7613 #7460]: #7612
-#7705 := [unit-resolution #7461 #4096]: #7611
-#7575 := (not #7584)
-#4948 := (up_6 uf_23 ?x65!15)
-#4949 := (not #4948)
-#7912 := (iff #4949 #7575)
-#7910 := (iff #4948 #7584)
-#7908 := (iff #7584 #4948)
-#7909 := [monotonicity #7203]: #7908
-#7911 := [symm #7909]: #7910
-#7907 := [monotonicity #7911]: #7912
-#6933 := (or #4949 #6927)
-#6943 := (or #4332 #4949 #6927)
-#6921 := (or #4949 #6920)
-#6944 := (or #4332 #6921)
-#7042 := (iff #6944 #6943)
-#7039 := (or #4332 #6933)
-#7036 := (iff #7039 #6943)
-#7037 := [rewrite]: #7036
-#7040 := (iff #6944 #7039)
-#6934 := (iff #6921 #6933)
-#6935 := [monotonicity #6932]: #6934
-#7041 := [monotonicity #6935]: #7040
-#7043 := [trans #7041 #7037]: #7042
-#7038 := [quant-inst]: #6944
-#7044 := [mp #7038 #7043]: #6943
-#7700 := [unit-resolution #7044 #7832]: #6933
-#7880 := [unit-resolution #7700 #7879]: #4949
-#7913 := [mp #7880 #7907]: #7575
-#7471 := (not #7611)
-#7574 := (or #7471 #7584 #4543)
-#3825 := [def-axiom]: #7574
-#7914 := [unit-resolution #3825 #7913 #7705]: #4543
-#7465 := (or #7628 #7464)
-#7466 := [def-axiom]: #7465
-#7915 := [unit-resolution #7466 #7914]: #7464
-#7703 := (or #7603 #7702)
-#7750 := (or #7140 #7603 #7702)
-#7754 := (or #7140 #7703)
-#7757 := (iff #7754 #7750)
-#7758 := [rewrite]: #7757
-#7756 := [quant-inst]: #7754
-#7759 := [mp #7756 #7758]: #7750
-#7916 := [unit-resolution #7759 #4123]: #7703
-#7917 := [unit-resolution #7916 #7915]: #7702
-#7661 := (or #7790 #7701)
-#7662 := [th-lemma]: #7661
-#7711 := [unit-resolution #7662 #7917]: #7790
-#7712 := [th-lemma #7711 #7851 #7714 #7811]: false
-#7918 := [lemma #7712]: #7840
-#11066 := [unit-resolution #7918 #10914 #10462 #9067 #10469 #10124]: #11065
-#11067 := [unit-resolution #11066 #10919]: #4396
-#3987 := (or #4393 #4387)
-#3988 := [def-axiom]: #3987
-#27286 := [unit-resolution #3988 #11067]: #4387
-#3986 := (or #4390 #3329 #4384)
-#3978 := [def-axiom]: #3986
-#27289 := [unit-resolution #3978 #27286]: #4387
-#27290 := [unit-resolution #27289 #13940]: #4384
-#3900 := (or #4381 #4375)
-#3901 := [def-axiom]: #3900
-#27291 := [unit-resolution #3901 #27290]: #4375
-#27292 := (or #4378 #4372)
-#7000 := (uf_1 uf_22 ?x71!19)
-#7001 := (uf_10 #7000)
-#6954 := (uf_4 uf_14 ?x71!19)
-#6984 := (* -1::int #6954)
-#7019 := (+ #6984 #7001)
-#7020 := (+ #182 #7019)
-#7021 := (>= #7020 0::int)
-#18941 := (not #7021)
-#7005 := (* -1::int #7001)
-#19226 := (+ #2212 #7005)
-#19230 := (>= #19226 0::int)
-#19225 := (= #2212 #7001)
-#19092 := (= #2211 #7000)
-#5993 := (= ?x72!18 uf_22)
-#15489 := (= ?x72!18 #10571)
-#5992 := (up_6 uf_15 ?x72!18)
-#6705 := (not #5992)
-#5963 := (uf_4 uf_14 ?x72!18)
-#7080 := (+ #5963 #6984)
-#7081 := (+ #2212 #7080)
-#7082 := (>= #7081 0::int)
-#19649 := (not #7082)
-#6985 := (+ #2207 #6984)
-#6986 := (<= #6985 0::int)
-#18752 := (or #4323 #6986)
-#6976 := (+ #6954 #2208)
-#6977 := (>= #6976 0::int)
-#18750 := (or #4323 #6977)
-#18782 := (iff #18750 #18752)
-#18807 := (iff #18752 #18752)
-#18838 := [rewrite]: #18807
-#6989 := (iff #6977 #6986)
-#6978 := (+ #2208 #6954)
-#6981 := (>= #6978 0::int)
-#6987 := (iff #6981 #6986)
-#6988 := [rewrite]: #6987
-#6982 := (iff #6977 #6981)
-#6979 := (= #6976 #6978)
-#6980 := [rewrite]: #6979
-#6983 := [monotonicity #6980]: #6982
-#6990 := [trans #6983 #6988]: #6989
-#18837 := [monotonicity #6990]: #18782
-#18839 := [trans #18837 #18838]: #18782
-#18781 := [quant-inst]: #18750
-#18834 := [mp #18781 #18839]: #18752
-#19630 := [unit-resolution #18834 #10924]: #6986
-#3844 := (not #2799)
-#19636 := [hypothesis]: #3375
-#3845 := (or #3370 #3844)
-#3998 := [def-axiom]: #3845
-#19637 := [unit-resolution #3998 #19636]: #3844
-#6055 := (* -1::int #5963)
-#6056 := (+ #2209 #6055)
-#17904 := (>= #6056 0::int)
-#5968 := (= #2209 #5963)
-#4013 := (or #3370 #2219)
-#3842 := [def-axiom]: #4013
-#19638 := [unit-resolution #3842 #19636]: #2219
-#10920 := [unit-resolution #3965 #10919]: #4327
-#16950 := (or #4332 #3355 #5968)
-#5964 := (= #5963 #2209)
-#5967 := (or #3355 #5964)
-#16945 := (or #4332 #5967)
-#17123 := (iff #16945 #16950)
-#5971 := (or #3355 #5968)
-#16952 := (or #4332 #5971)
-#16954 := (iff #16952 #16950)
-#16815 := [rewrite]: #16954
-#16953 := (iff #16945 #16952)
-#5972 := (iff #5967 #5971)
-#5969 := (iff #5964 #5968)
-#5970 := [rewrite]: #5969
-#5973 := [monotonicity #5970]: #5972
-#16949 := [monotonicity #5973]: #16953
-#17124 := [trans #16949 #16815]: #17123
-#16951 := [quant-inst]: #16945
-#17131 := [mp #16951 #17124]: #16950
-#19604 := [unit-resolution #17131 #10920 #19638]: #5968
-#19640 := (not #5968)
-#19647 := (or #19640 #17904)
-#19648 := [th-lemma]: #19647
-#19646 := [unit-resolution #19648 #19604]: #17904
-#19668 := (not #6986)
-#19667 := (not #17904)
-#19669 := (or #19649 #19667 #19668 #2799)
-#19674 := [th-lemma]: #19669
-#19675 := [unit-resolution #19674 #19646 #19637 #19630]: #19649
-#19673 := (or #6705 #7082)
-#4012 := (or #3370 #2218)
-#4006 := [def-axiom]: #4012
-#19664 := [unit-resolution #4006 #19636]: #2218
-#18969 := (or #4193 #2217 #6705 #7082)
-#7083 := (or #6705 #2217 #7082)
-#18974 := (or #4193 #7083)
-#18875 := (iff #18974 #18969)
-#7084 := (or #2217 #6705 #7082)
-#18976 := (or #4193 #7084)
-#18944 := (iff #18976 #18969)
-#18874 := [rewrite]: #18944
-#18977 := (iff #18974 #18976)
-#7085 := (iff #7083 #7084)
-#7086 := [rewrite]: #7085
-#18943 := [monotonicity #7086]: #18977
-#18873 := [trans #18943 #18874]: #18875
-#18975 := [quant-inst]: #18974
-#18942 := [mp #18975 #18873]: #18969
-#18871 := [unit-resolution #18942 #10131 #19664]: #19673
-#19651 := [unit-resolution #18871 #19675]: #6705
-#15504 := (or #5992 #15489)
-#15499 := (up_6 #11533 ?x72!18)
-#15509 := (iff #15499 #15504)
-#17157 := (or #6627 #15509)
-#15490 := (ite #15489 #3770 #5992)
-#15500 := (iff #15499 #15490)
-#17130 := (or #6627 #15500)
-#17160 := (iff #17130 #17157)
-#17164 := (iff #17157 #17157)
-#17122 := [rewrite]: #17164
-#15510 := (iff #15500 #15509)
-#15507 := (iff #15490 #15504)
-#15501 := (ite #15489 true #5992)
-#15505 := (iff #15501 #15504)
-#15506 := [rewrite]: #15505
-#15502 := (iff #15490 #15501)
-#15503 := [monotonicity #3762]: #15502
-#15508 := [trans #15503 #15506]: #15507
-#15511 := [monotonicity #15508]: #15510
-#17158 := [monotonicity #15511]: #17160
-#17165 := [trans #17158 #17122]: #17160
-#17159 := [quant-inst]: #17130
-#17166 := [mp #17159 #17165]: #17157
-#19656 := [unit-resolution #17166 #4096]: #15509
-#19655 := (iff #2219 #15499)
-#19650 := (iff #15499 #2219)
-#19657 := [monotonicity #13613]: #19650
-#19665 := [symm #19657]: #19655
-#19666 := [mp #19638 #19665]: #15499
-#17221 := (not #15499)
-#17223 := (not #15509)
-#17226 := (or #17223 #17221 #15504)
-#17227 := [def-axiom]: #17226
-#19660 := [unit-resolution #17227 #19666 #19656]: #15504
-#17201 := (not #15504)
-#17222 := (or #17201 #5992 #15489)
-#17217 := [def-axiom]: #17222
-#19635 := [unit-resolution #17217 #19660 #19651]: #15489
-#19038 := [trans #19635 #13436]: #5993
-#19241 := [monotonicity #19038]: #19092
-#19355 := [monotonicity #19241]: #19225
-#19410 := (not #19225)
-#18872 := (or #19410 #19230)
-#19416 := [th-lemma]: #18872
-#19513 := [unit-resolution #19416 #19355]: #19230
-#7165 := (uf_2 #2211)
-#7171 := (uf_4 uf_14 #7165)
-#7185 := (* -1::int #7171)
-#7186 := (+ #182 #7185)
-#7187 := (<= #7186 0::int)
-#19218 := (= #182 #7171)
-#19098 := (= #7171 #182)
-#19114 := (= #7165 uf_22)
-#19115 := (= #7165 #10571)
-#19577 := (= #7165 ?x72!18)
-#7166 := (= ?x72!18 #7165)
-#19027 := (or #8504 #7166)
-#19028 := [quant-inst]: #19027
-#19514 := [unit-resolution #19028 #4082]: #7166
-#19578 := [symm #19514]: #19577
-#19116 := [trans #19578 #19635]: #19115
-#19205 := [trans #19116 #13436]: #19114
-#19206 := [monotonicity #19205]: #19098
-#19213 := [symm #19206]: #19218
-#19214 := (not #19218)
-#19207 := (or #19214 #7187)
-#19215 := [th-lemma]: #19207
-#19583 := [unit-resolution #19215 #19213]: #7187
-#19048 := (+ #5963 #7185)
-#19031 := (>= #19048 0::int)
-#19047 := (= #5963 #7171)
-#19592 := [monotonicity #19514]: #19047
-#19803 := (not #19047)
-#19804 := (or #19803 #19031)
-#19805 := [th-lemma]: #19804
-#19806 := [unit-resolution #19805 #19592]: #19031
-#19929 := (not #19230)
-#19808 := (not #19031)
-#19807 := (not #7187)
-#19809 := (or #18941 #19668 #2799 #19807 #19808 #19667 #19929)
-#19810 := [th-lemma]: #19809
-#19811 := [unit-resolution #19810 #19630 #19646 #19637 #19806 #19583 #19513]: #18941
-#7009 := (+ uf_9 #7005)
-#7010 := (<= #7009 0::int)
-#18939 := (not #7010)
-#19930 := (or #18939 #19929 #2217)
-#19925 := [hypothesis]: #2218
-#19926 := [hypothesis]: #7010
-#19927 := [hypothesis]: #19230
-#19928 := [th-lemma #19927 #19926 #19925]: false
-#19931 := [lemma #19928]: #19930
-#19837 := [unit-resolution #19931 #19513 #19664]: #18939
-#7026 := (+ #2208 #7001)
-#7027 := (+ #182 #7026)
-#7030 := (= #7027 0::int)
-#19841 := (not #7030)
-#18789 := (>= #7027 0::int)
-#19838 := (not #18789)
-#19839 := (or #19838 #2799 #19807 #19808 #19667 #19929)
-#19835 := [th-lemma]: #19839
-#19836 := [unit-resolution #19835 #19646 #19637 #19806 #19583 #19513]: #19838
-#19842 := (or #19841 #18789)
-#19843 := [th-lemma]: #19842
-#19840 := [unit-resolution #19843 #19836]: #19841
-#7033 := (or #7010 #7021 #7030)
-#18840 := (or #4315 #7010 #7021 #7030)
-#7002 := (+ #7001 #2208)
-#7003 := (+ #182 #7002)
-#7004 := (= #7003 0::int)
-#7006 := (+ #1357 #7005)
-#7007 := (+ #6954 #7006)
-#7008 := (<= #7007 0::int)
-#7011 := (or #7010 #7008 #7004)
-#18841 := (or #4315 #7011)
-#18786 := (iff #18841 #18840)
-#18900 := (or #4315 #7033)
-#18780 := (iff #18900 #18840)
-#18785 := [rewrite]: #18780
-#18784 := (iff #18841 #18900)
-#7034 := (iff #7011 #7033)
-#7031 := (iff #7004 #7030)
-#7028 := (= #7003 #7027)
-#7029 := [rewrite]: #7028
-#7032 := [monotonicity #7029]: #7031
-#7024 := (iff #7008 #7021)
-#7012 := (+ #6954 #7005)
-#7013 := (+ #1357 #7012)
-#7016 := (<= #7013 0::int)
-#7022 := (iff #7016 #7021)
-#7023 := [rewrite]: #7022
-#7017 := (iff #7008 #7016)
-#7014 := (= #7007 #7013)
-#7015 := [rewrite]: #7014
-#7018 := [monotonicity #7015]: #7017
-#7025 := [trans #7018 #7023]: #7024
-#7035 := [monotonicity #7025 #7032]: #7034
-#18779 := [monotonicity #7035]: #18784
-#18783 := [trans #18779 #18785]: #18786
-#18905 := [quant-inst]: #18841
-#18787 := [mp #18905 #18783]: #18840
-#19844 := [unit-resolution #18787 #10914]: #7033
-#19845 := [unit-resolution #19844 #19840 #19837 #19811]: false
-#19871 := [lemma #19845]: #3370
-#3897 := (or #4378 #3375 #4372)
-#3898 := [def-axiom]: #3897
-#27293 := [unit-resolution #3898 #19871]: #27292
-#27294 := [unit-resolution #27293 #27291]: #4372
-#4002 := (or #4369 #2249)
-#4000 := [def-axiom]: #4002
-#27295 := [unit-resolution #4000 #27294]: #2249
-#5681 := (+ #2236 #5680)
-#18550 := (>= #5681 0::int)
-#5655 := (= #2236 #5650)
-#3846 := (or #4369 #4361)
-#3994 := [def-axiom]: #3846
-#27296 := [unit-resolution #3994 #27294]: #4361
-#18659 := (or #5655 #4366)
-#15149 := (uf_10 #15148)
-#15175 := (* -1::int #15149)
-#11814 := (uf_24 #10571)
-#11812 := (* -1::int #11814)
-#15176 := (+ #11812 #15175)
-#15177 := (+ #2236 #15176)
-#15830 := (>= #15177 0::int)
-#5696 := (uf_1 uf_22 ?x75!20)
-#5697 := (uf_10 #5696)
-#15921 := (+ #5697 #15175)
-#15923 := (>= #15921 0::int)
-#15920 := (= #5697 #15149)
-#18516 := (= #15149 #5697)
-#18514 := (= #15148 #5696)
-#18515 := [monotonicity #13436]: #18514
-#18517 := [monotonicity #18515]: #18516
-#18518 := [symm #18517]: #15920
-#18513 := (not #15920)
-#18519 := (or #18513 #15923)
-#18520 := [th-lemma]: #18519
-#18521 := [unit-resolution #18520 #18518]: #15923
-#11766 := (+ #4615 #11812)
-#8751 := (>= #11766 0::int)
-#8693 := (= #4615 #11814)
-#18522 := (= #11814 #4615)
-#18523 := [monotonicity #13436]: #18522
-#18530 := [symm #18523]: #8693
-#18531 := (not #8693)
-#18529 := (or #18531 #8751)
-#18532 := [th-lemma]: #18529
-#18533 := [unit-resolution #18532 #18530]: #8751
-#5722 := (+ #2237 #5697)
-#5723 := (+ #182 #5722)
-#15649 := (<= #5723 0::int)
-#5726 := (= #5723 0::int)
-#5701 := (* -1::int #5697)
-#5705 := (+ uf_9 #5701)
-#5706 := (<= #5705 0::int)
-#15666 := (not #5706)
-#5715 := (+ #5680 #5697)
-#5716 := (+ #182 #5715)
-#5717 := (>= #5716 0::int)
-#5748 := (or #5706 #5717)
-#5751 := (not #5748)
-#18607 := (not #5655)
-#18534 := [hypothesis]: #18607
-#5754 := (or #5655 #5751)
-#15653 := (or #4307 #5655 #5751)
-#5702 := (+ #1357 #5701)
-#5703 := (+ #5650 #5702)
-#5704 := (<= #5703 0::int)
-#5745 := (or #5706 #5704)
-#5746 := (not #5745)
-#5651 := (= #5650 #2236)
-#5747 := (or #5651 #5746)
-#15654 := (or #4307 #5747)
-#15663 := (iff #15654 #15653)
-#15656 := (or #4307 #5754)
-#15659 := (iff #15656 #15653)
-#15660 := [rewrite]: #15659
-#15657 := (iff #15654 #15656)
-#5755 := (iff #5747 #5754)
-#5752 := (iff #5746 #5751)
-#5749 := (iff #5745 #5748)
-#5720 := (iff #5704 #5717)
-#5708 := (+ #5650 #5701)
-#5709 := (+ #1357 #5708)
-#5712 := (<= #5709 0::int)
-#5718 := (iff #5712 #5717)
-#5719 := [rewrite]: #5718
-#5713 := (iff #5704 #5712)
-#5710 := (= #5703 #5709)
-#5711 := [rewrite]: #5710
-#5714 := [monotonicity #5711]: #5713
-#5721 := [trans #5714 #5719]: #5720
-#5750 := [monotonicity #5721]: #5749
-#5753 := [monotonicity #5750]: #5752
-#5656 := (iff #5651 #5655)
-#5657 := [rewrite]: #5656
-#5756 := [monotonicity #5657 #5753]: #5755
-#15658 := [monotonicity #5756]: #15657
-#15664 := [trans #15658 #15660]: #15663
-#15655 := [quant-inst]: #15654
-#15665 := [mp #15655 #15664]: #15653
-#18538 := [unit-resolution #15665 #10462]: #5754
-#18539 := [unit-resolution #18538 #18534]: #5751
-#15667 := (or #5748 #15666)
-#15697 := [def-axiom]: #15667
-#18542 := [unit-resolution #15697 #18539]: #15666
-#15813 := (not #5717)
-#15814 := (or #5748 #15813)
-#15815 := [def-axiom]: #15814
-#18543 := [unit-resolution #15815 #18539]: #15813
-#5729 := (or #5706 #5717 #5726)
-#15441 := (or #4315 #5706 #5717 #5726)
-#5698 := (+ #5697 #2237)
-#5699 := (+ #182 #5698)
-#5700 := (= #5699 0::int)
-#5707 := (or #5706 #5704 #5700)
-#15442 := (or #4315 #5707)
-#15615 := (iff #15442 #15441)
-#15466 := (or #4315 #5729)
-#15594 := (iff #15466 #15441)
-#15595 := [rewrite]: #15594
-#15497 := (iff #15442 #15466)
-#5730 := (iff #5707 #5729)
-#5727 := (iff #5700 #5726)
-#5724 := (= #5699 #5723)
-#5725 := [rewrite]: #5724
-#5728 := [monotonicity #5725]: #5727
-#5731 := [monotonicity #5721 #5728]: #5730
-#15498 := [monotonicity #5731]: #15497
-#15638 := [trans #15498 #15595]: #15615
-#15465 := [quant-inst]: #15442
-#15639 := [mp #15465 #15638]: #15441
-#18579 := [unit-resolution #15639 #10914]: #5729
-#18580 := [unit-resolution #18579 #18543 #18542]: #5726
-#18581 := (not #5726)
-#18582 := (or #18581 #15649)
-#18583 := [th-lemma]: #18582
-#18584 := [unit-resolution #18583 #18580]: #15649
-#18588 := (not #15923)
-#18587 := (not #4858)
-#18586 := (not #8751)
-#18585 := (not #15649)
-#18589 := (or #15830 #18585 #18586 #18587 #18588)
-#18590 := [th-lemma]: #18589
-#18591 := [unit-resolution #18590 #18584 #18533 #10925 #18521]: #15830
-#15829 := (<= #15177 0::int)
-#15922 := (<= #15921 0::int)
-#18592 := (or #18513 #15922)
-#18593 := [th-lemma]: #18592
-#18594 := [unit-resolution #18593 #18518]: #15922
-#9376 := (<= #4857 0::int)
-#4616 := (= #182 #4615)
-#4865 := (up_6 uf_23 uf_22)
-#3772 := (up_6 #188 uf_22)
-#10910 := (iff #3772 #4865)
-#10908 := (iff #4865 #3772)
-#10909 := [monotonicity #10469]: #10908
-#10911 := [symm #10909]: #10910
+#359 := (iff #59 #356)
+#332 := (= 0::int #56)
+#343 := (or #60 #332)
+#348 := (forall (vars (?x20 T2) (?x21 T2)) #343)
+#357 := (iff #348 #356)
+#354 := (iff #343 #353)
+#351 := (iff #332 #57)
+#352 := [rewrite]: #351
+#355 := [monotonicity #352]: #354
+#358 := [quant-intro #355]: #357
+#349 := (iff #59 #348)
+#346 := (iff #58 #343)
+#340 := (implies #55 #332)
+#344 := (iff #340 #343)
+#345 := [rewrite]: #344
+#341 := (iff #58 #340)
+#338 := (iff #57 #332)
+#339 := [rewrite]: #338
+#342 := [monotonicity #339]: #341
+#347 := [trans #342 #345]: #346
+#350 := [quant-intro #347]: #349
+#360 := [trans #350 #358]: #359
+#330 := [asserted]: #59
+#361 := [mp #330 #360]: #356
+#1513 := [mp~ #361 #1512]: #356
+#3872 := [mp #1513 #3871]: #3867
+#8154 := (not #3867)
+#8319 := (or #8154 #8159 #8299)
+#8250 := (= #6904 ?x63!14)
+#8305 := (not #8250)
+#8306 := (or #8305 #8299)
+#8324 := (or #8154 #8306)
+#8351 := (iff #8324 #8319)
+#8326 := (or #8154 #8309)
+#8342 := (iff #8326 #8319)
+#8343 := [rewrite]: #8342
+#8336 := (iff #8324 #8326)
+#8317 := (iff #8306 #8309)
+#8307 := (iff #8305 #8159)
+#8252 := (iff #8250 #8074)
+#8253 := [rewrite]: #8252
+#8308 := [monotonicity #8253]: #8307
+#8318 := [monotonicity #8308]: #8317
+#8337 := [monotonicity #8318]: #8336
+#8352 := [trans #8337 #8343]: #8351
+#8325 := [quant-inst]: #8324
+#8384 := [mp #8325 #8352]: #8319
+#9745 := [unit-resolution #8384 #3872]: #8309
+#9746 := [unit-resolution #9745 #9744]: #8159
+#6665 := (or #6768 #7863 #8074)
+#6767 := [def-axiom]: #6665
+#9747 := [unit-resolution #6767 #9746 #9739]: #7863
+#9755 := [mp #9747 #9754]: #9709
+#9710 := (not #9709)
+#9767 := (or #9708 #9710)
+#8353 := (up_6 uf_15 #6904)
+#9763 := (not #8353)
+#9764 := (iff #184 #9763)
+#9761 := (iff #183 #8353)
+#9759 := (iff #8353 #183)
+#9760 := [monotonicity #9758]: #9759
+#9762 := [symm #9760]: #9761
+#9765 := [monotonicity #9762]: #9764
+#3699 := (or #4202 #184)
+#3694 := [def-axiom]: #3699
+#7018 := [unit-resolution #3694 #6215]: #184
+#9766 := [mp #7018 #9765]: #9763
+#3643 := (or #4214 #3958)
+#3667 := [def-axiom]: #3643
+#7019 := [unit-resolution #3667 #6214]: #3958
+#9715 := (or #3963 #8353 #9708 #9710)
+#9711 := (or #8353 #9710 #9708)
+#9716 := (or #3963 #9711)
+#9723 := (iff #9716 #9715)
+#9712 := (or #8353 #9708 #9710)
+#9718 := (or #3963 #9712)
+#9721 := (iff #9718 #9715)
+#9722 := [rewrite]: #9721
+#9719 := (iff #9716 #9718)
+#9713 := (iff #9711 #9712)
+#9714 := [rewrite]: #9713
+#9720 := [monotonicity #9714]: #9719
+#9724 := [trans #9720 #9722]: #9723
+#9717 := [quant-inst]: #9716
+#9725 := [mp #9717 #9724]: #9715
+#9768 := [unit-resolution #9725 #7019 #9766]: #9767
+#9769 := [unit-resolution #9768 #9755]: #9708
+#9770 := [hypothesis]: #7886
+#8219 := (* -1::int #8213)
+#8387 := (+ #7892 #8219)
+#8393 := (>= #8387 0::int)
+#8386 := (= #7892 #8213)
+#9773 := (= #8213 #7892)
+#9771 := (= #8212 #7891)
+#9772 := [monotonicity #9758]: #9771
+#9774 := [monotonicity #9772]: #9773
+#9775 := [symm #9774]: #8386
+#9776 := (not #8386)
+#9777 := (or #9776 #8393)
+#9778 := [th-lemma]: #9777
+#9779 := [unit-resolution #9778 #9775]: #8393
+#8385 := (>= #8213 0::int)
+#9780 := (or #8385 #8269)
+#9781 := [th-lemma]: #9780
+#9782 := [unit-resolution #9781 #9740]: #8385
+#9730 := (+ #1939 #9706)
+#9737 := (<= #9730 0::int)
+#9729 := (= #1939 #9705)
+#9783 := (= #9705 #1939)
+#9784 := [monotonicity #9750]: #9783
+#9785 := [symm #9784]: #9729
+#9786 := (not #9729)
+#9787 := (or #9786 #9737)
+#9788 := [th-lemma]: #9787
+#9789 := [unit-resolution #9788 #9785]: #9737
+#8369 := (* -1::int #8358)
+#9504 := (+ #185 #8369)
+#9300 := (>= #9504 0::int)
+#9503 := (= #185 #8358)
+#9790 := (= #8358 #185)
+#9791 := [monotonicity #9758]: #9790
+#9792 := [symm #9791]: #9503
+#9793 := (not #9503)
+#9794 := (or #9793 #9300)
+#9795 := [th-lemma]: #9794
+#9796 := [unit-resolution #9795 #9792]: #9300
+#9797 := [th-lemma #9796 #9789 #9782 #9779 #9770 #9769]: false
+#9799 := [lemma #9797]: #9798
+#9021 := [unit-resolution #9799 #9017 #9019]: #8269
+#8254 := (or #8074 #8270)
+#369 := (<= #56 0::int)
+#370 := (not #369)
+#373 := (or #55 #370)
+#3873 := (forall (vars (?x22 T2) (?x23 T2)) (:pat #3825) #373)
+#376 := (forall (vars (?x22 T2) (?x23 T2)) #373)
+#3876 := (iff #376 #3873)
+#3874 := (iff #373 #373)
+#3875 := [refl]: #3874
+#3877 := [quant-intro #3875]: #3876
+#1515 := (~ #376 #376)
+#1514 := (~ #373 #373)
+#1552 := [refl]: #1514
+#1516 := [nnf-pos #1552]: #1515
+#61 := (< 0::int #56)
+#62 := (implies #60 #61)
+#63 := (forall (vars (?x22 T2) (?x23 T2)) #62)
+#379 := (iff #63 #376)
+#363 := (or #55 #61)
+#366 := (forall (vars (?x22 T2) (?x23 T2)) #363)
+#377 := (iff #366 #376)
+#374 := (iff #363 #373)
+#371 := (iff #61 #370)
+#372 := [rewrite]: #371
+#375 := [monotonicity #372]: #374
+#378 := [quant-intro #375]: #377
+#367 := (iff #63 #366)
+#364 := (iff #62 #363)
+#365 := [rewrite]: #364
+#368 := [quant-intro #365]: #367
+#380 := [trans #368 #378]: #379
+#362 := [asserted]: #63
+#381 := [mp #362 #380]: #376
+#1553 := [mp~ #381 #1516]: #376
+#3878 := [mp #1553 #3877]: #3873
+#7324 := (not #3873)
+#8256 := (or #7324 #8074 #8270)
+#8251 := (or #8250 #8270)
+#8300 := (or #7324 #8251)
+#8316 := (iff #8300 #8256)
+#8302 := (or #7324 #8254)
+#8320 := (iff #8302 #8256)
+#8321 := [rewrite]: #8320
+#8303 := (iff #8300 #8302)
+#8249 := (iff #8251 #8254)
+#8255 := [monotonicity #8253]: #8249
+#8304 := [monotonicity #8255]: #8303
+#8322 := [trans #8304 #8321]: #8316
+#8301 := [quant-inst]: #8300
+#8323 := [mp #8301 #8322]: #8256
+#9022 := [unit-resolution #8323 #3878]: #8254
+#9013 := [unit-resolution #9022 #9021]: #8074
+#8832 := [trans #9013 #9758]: #8831
+#9155 := [monotonicity #8832]: #9114
+#9158 := [symm #9155]: #9157
+#9154 := (= #1938 #185)
+#5830 := (uf_24 uf_22)
+#9156 := (= #5830 #185)
+#5879 := (= #185 #5830)
+#5833 := (uf_10 #5832)
+#5849 := (>= #5833 0::int)
+#5837 := (* -1::int #5833)
+#5841 := (+ uf_9 #5837)
+#5842 := (<= #5841 0::int)
+#5881 := (or #5842 #5849)
+#6879 := (= #5833 0::int)
+decl uf_2 :: (-> T1 T2)
+#5488 := (uf_1 uf_22 uf_11)
+#8349 := (uf_2 #5488)
+#8700 := (uf_1 #8349 #8349)
+#8701 := (uf_10 #8700)
+#9648 := (= #8701 0::int)
+#9665 := (or #8154 #9648)
+#9649 := (= #8349 #8349)
+#9650 := (not #9649)
+#9651 := (or #9650 #9648)
+#9666 := (or #8154 #9651)
+#9668 := (iff #9666 #9665)
+#9670 := (iff #9665 #9665)
+#9671 := [rewrite]: #9670
+#9663 := (iff #9651 #9648)
+#9658 := (or false #9648)
+#9661 := (iff #9658 #9648)
+#9662 := [rewrite]: #9661
+#9659 := (iff #9651 #9658)
+#9656 := (iff #9650 false)
+#9654 := (iff #9650 #4945)
+#9652 := (iff #9649 true)
+#9653 := [rewrite]: #9652
+#9655 := [monotonicity #9653]: #9654
+#9657 := [trans #9655 #4943]: #9656
+#9660 := [monotonicity #9657]: #9659
+#9664 := [trans #9660 #9662]: #9663
+#9669 := [monotonicity #9664]: #9668
+#9672 := [trans #9669 #9671]: #9668
+#9667 := [quant-inst]: #9666
+#9673 := [mp #9667 #9672]: #9665
+#9675 := [unit-resolution #9673 #3872]: #9648
+#9684 := (= #5833 #8701)
+#9682 := (= #5832 #8700)
+#9680 := (= #8700 #5832)
+#9048 := (= #8349 uf_22)
+#8350 := (= uf_22 #8349)
+#16 := (uf_2 #12)
+#301 := (= #10 #16)
+#3832 := (forall (vars (?x4 T2) (?x5 T2)) (:pat #3825) #301)
+#305 := (forall (vars (?x4 T2) (?x5 T2)) #301)
+#3835 := (iff #305 #3832)
+#3833 := (iff #301 #301)
+#3834 := [refl]: #3833
+#3836 := [quant-intro #3834]: #3835
+#1501 := (~ #305 #305)
+#1535 := (~ #301 #301)
+#1536 := [refl]: #1535
+#1502 := [nnf-pos #1536]: #1501
+#17 := (= #16 #10)
+#18 := (forall (vars (?x4 T2) (?x5 T2)) #17)
+#306 := (iff #18 #305)
+#303 := (iff #17 #301)
+#304 := [rewrite]: #303
+#307 := [quant-intro #304]: #306
+#300 := [asserted]: #18
+#310 := [mp #300 #307]: #305
+#1537 := [mp~ #310 #1502]: #305
+#3837 := [mp #1537 #3836]: #3832
+#8156 := (not #3832)
+#8355 := (or #8156 #8350)
+#8356 := [quant-inst]: #8355
+#9047 := [unit-resolution #8356 #3837]: #8350
+#9049 := [symm #9047]: #9048
+#9681 := [monotonicity #9049 #9049]: #9680
+#9683 := [symm #9681]: #9682
+#9685 := [monotonicity #9683]: #9684
+#9686 := [trans #9685 #9675]: #6879
+#5907 := (not #5849)
+#9687 := [hypothesis]: #5907
+#9688 := (not #6879)
+#9689 := (or #9688 #5849)
+#9690 := [th-lemma]: #9689
+#9691 := [unit-resolution #9690 #9687 #9686]: false
+#9692 := [lemma #9691]: #5849
+#5908 := (or #5881 #5907)
+#5909 := [def-axiom]: #5908
+#8604 := [unit-resolution #5909 #9692]: #5881
+#5884 := (not #5881)
+#5887 := (or #5879 #5884)
+#5890 := (or #4076 #5879 #5884)
+#5838 := (+ #1235 #5837)
+#5839 := (+ #185 #5838)
+#5840 := (<= #5839 0::int)
+#5877 := (or #5842 #5840)
+#5878 := (not #5877)
+#5880 := (or #5879 #5878)
+#5891 := (or #4076 #5880)
+#5898 := (iff #5891 #5890)
+#5893 := (or #4076 #5887)
+#5896 := (iff #5893 #5890)
+#5897 := [rewrite]: #5896
+#5894 := (iff #5891 #5893)
+#5888 := (iff #5880 #5887)
+#5885 := (iff #5878 #5884)
+#5882 := (iff #5877 #5881)
+#5852 := (iff #5840 #5849)
+#5846 := (<= #5837 0::int)
+#5850 := (iff #5846 #5849)
+#5851 := [rewrite]: #5850
+#5847 := (iff #5840 #5846)
+#5844 := (= #5839 #5837)
+#5845 := [rewrite]: #5844
+#5848 := [monotonicity #5845]: #5847
+#5853 := [trans #5848 #5851]: #5852
+#5883 := [monotonicity #5853]: #5882
+#5886 := [monotonicity #5883]: #5885
+#5889 := [monotonicity #5886]: #5888
+#5895 := [monotonicity #5889]: #5894
+#5899 := [trans #5895 #5897]: #5898
+#5892 := [quant-inst]: #5891
+#5900 := [mp #5892 #5899]: #5890
+#8599 := [unit-resolution #5900 #6216]: #5887
+#8605 := [unit-resolution #8599 #8604]: #5879
+#9120 := [symm #8605]: #9156
+#9121 := (= #1938 #5830)
+#9149 := [monotonicity #8832]: #9121
+#9182 := [trans #9149 #9120]: #9154
+#9183 := [trans #9182 #9158]: #2494
+#9203 := [unit-resolution #8506 #9183]: false
+#9204 := [lemma #9203]: #2497
+#3559 := (or #4202 #4196)
+#3670 := [def-axiom]: #3559
+#7035 := [unit-resolution #3670 #6215]: #4196
+#6020 := (uf_1 uf_22 ?x61!13)
+#6021 := (uf_10 #6020)
+#6082 := (+ #1921 #6021)
+#6083 := (+ #185 #6082)
+#6104 := (>= #6083 0::int)
+#6086 := (= #6083 0::int)
+#6022 := (* -1::int #6021)
+#6026 := (+ uf_9 #6022)
+#6027 := (<= #6026 0::int)
+#6070 := (not #6027)
+#6042 := (+ #2484 #6021)
+#6043 := (+ #185 #6042)
+#6044 := (>= #6043 0::int)
+#6049 := (or #6027 #6044)
+#6052 := (not #6049)
+#6032 := (= #1920 #1922)
+#6196 := (not #6032)
+#6195 := [hypothesis]: #2491
+#6199 := (or #6196 #2486)
+#6200 := [th-lemma]: #6199
+#6201 := [unit-resolution #6200 #6195]: #6196
+#6058 := (or #4076 #6032 #6052)
+#6023 := (+ #1235 #6022)
+#6024 := (+ #1922 #6023)
+#6025 := (<= #6024 0::int)
+#6028 := (or #6027 #6025)
+#6029 := (not #6028)
+#6030 := (= #1922 #1920)
+#6031 := (or #6030 #6029)
+#6059 := (or #4076 #6031)
+#6066 := (iff #6059 #6058)
+#6055 := (or #6032 #6052)
+#6061 := (or #4076 #6055)
+#6064 := (iff #6061 #6058)
+#6065 := [rewrite]: #6064
+#6062 := (iff #6059 #6061)
+#6056 := (iff #6031 #6055)
+#6053 := (iff #6029 #6052)
+#6050 := (iff #6028 #6049)
+#6047 := (iff #6025 #6044)
+#6035 := (+ #1922 #6022)
+#6036 := (+ #1235 #6035)
+#6039 := (<= #6036 0::int)
+#6045 := (iff #6039 #6044)
+#6046 := [rewrite]: #6045
+#6040 := (iff #6025 #6039)
+#6037 := (= #6024 #6036)
+#6038 := [rewrite]: #6037
+#6041 := [monotonicity #6038]: #6040
+#6048 := [trans #6041 #6046]: #6047
+#6051 := [monotonicity #6048]: #6050
+#6054 := [monotonicity #6051]: #6053
+#6033 := (iff #6030 #6032)
+#6034 := [rewrite]: #6033
+#6057 := [monotonicity #6034 #6054]: #6056
+#6063 := [monotonicity #6057]: #6062
+#6067 := [trans #6063 #6065]: #6066
+#6060 := [quant-inst]: #6059
+#6068 := [mp #6060 #6067]: #6058
+#6217 := [unit-resolution #6068 #6216 #6201]: #6052
+#6071 := (or #6049 #6070)
+#6072 := [def-axiom]: #6071
+#6218 := [unit-resolution #6072 #6217]: #6070
+#6073 := (not #6044)
+#6074 := (or #6049 #6073)
+#6075 := [def-axiom]: #6074
+#6219 := [unit-resolution #6075 #6217]: #6073
+#6089 := (or #6027 #6044 #6086)
+#3691 := (or #4202 #4063)
+#3664 := [def-axiom]: #3691
+#6220 := [unit-resolution #3664 #6215]: #4063
+#6092 := (or #4068 #6027 #6044 #6086)
+#6078 := (+ #6021 #1921)
+#6079 := (+ #185 #6078)
+#6080 := (= #6079 0::int)
+#6081 := (or #6027 #6025 #6080)
+#6093 := (or #4068 #6081)
+#6100 := (iff #6093 #6092)
+#6095 := (or #4068 #6089)
+#6098 := (iff #6095 #6092)
+#6099 := [rewrite]: #6098
+#6096 := (iff #6093 #6095)
+#6090 := (iff #6081 #6089)
+#6087 := (iff #6080 #6086)
+#6084 := (= #6079 #6083)
+#6085 := [rewrite]: #6084
+#6088 := [monotonicity #6085]: #6087
+#6091 := [monotonicity #6048 #6088]: #6090
+#6097 := [monotonicity #6091]: #6096
+#6101 := [trans #6097 #6099]: #6100
+#6094 := [quant-inst]: #6093
+#6102 := [mp #6094 #6101]: #6092
+#6221 := [unit-resolution #6102 #6220]: #6089
+#6222 := [unit-resolution #6221 #6219 #6218]: #6086
+#6223 := (not #6086)
+#6224 := (or #6223 #6104)
+#6225 := [th-lemma]: #6224
+#6226 := [unit-resolution #6225 #6222]: #6104
+#6069 := (>= #2485 0::int)
+#6227 := (or #6069 #2486)
+#6228 := [th-lemma]: #6227
+#6229 := [unit-resolution #6228 #6195]: #6069
+#6230 := [th-lemma #6229 #6219 #6226]: false
+#6231 := [lemma #6230]: #2486
+#3697 := (or #4199 #2491 #4193)
+#3698 := [def-axiom]: #3697
+#7036 := [unit-resolution #3698 #6231 #7035]: #4193
+#3712 := (or #4190 #4184)
+#3713 := [def-axiom]: #3712
+#13484 := [unit-resolution #3713 #7036]: #4184
+#3708 := (or #4187 #2500 #4181)
+#3709 := [def-axiom]: #3708
+#13485 := [unit-resolution #3709 #13484]: #4184
+#13486 := [unit-resolution #13485 #9204]: #4181
+#3719 := (or #4178 #4088)
+#3720 := [def-axiom]: #3719
+#13487 := [unit-resolution #3720 #13486]: #4088
+#21789 := (or #4093 #9640 #10311)
+#10307 := (= #10301 #2032)
+#10310 := (or #9640 #10307)
+#22142 := (or #4093 #10310)
+#16210 := (iff #22142 #21789)
+#22191 := (or #4093 #10337)
+#22697 := (iff #22191 #21789)
+#17223 := [rewrite]: #22697
+#22172 := (iff #22142 #22191)
+#10338 := (iff #10310 #10337)
+#10335 := (iff #10307 #10311)
+#10336 := [rewrite]: #10335
+#10339 := [monotonicity #10336]: #10338
+#22553 := [monotonicity #10339]: #22172
+#20892 := [trans #22553 #17223]: #16210
+#19383 := [quant-inst]: #22142
+#19478 := [mp #19383 #20892]: #21789
+#30520 := [unit-resolution #19478 #13487]: #10337
+#30521 := [unit-resolution #30520 #30519]: #9640
+#30528 := [mp #30521 #30527]: #30208
+#30205 := (not #14006)
+#30206 := (or #30205 #13996 #30202)
+#30207 := [def-axiom]: #30206
+#30529 := [unit-resolution #30207 #30528 #30517]: #30202
+#30200 := (or #14001 #30199)
+#30201 := [def-axiom]: #30200
+#30530 := [unit-resolution #30201 #30529]: #30199
+#30279 := (or #13994 #30274)
+#30282 := (or #7324 #13994 #30274)
+#30275 := (= #6904 ?x75!20)
+#30276 := (or #30275 #30274)
+#30283 := (or #7324 #30276)
+#30290 := (iff #30283 #30282)
+#30285 := (or #7324 #30279)
+#30288 := (iff #30285 #30282)
+#30289 := [rewrite]: #30288
+#30286 := (iff #30283 #30285)
+#30280 := (iff #30276 #30279)
+#30277 := (iff #30275 #13994)
+#30278 := [rewrite]: #30277
+#30281 := [monotonicity #30278]: #30280
+#30287 := [monotonicity #30281]: #30286
+#30291 := [trans #30287 #30289]: #30290
+#30284 := [quant-inst]: #30283
+#30292 := [mp #30284 #30291]: #30282
+#30531 := [unit-resolution #30292 #3878]: #30279
+#30532 := [unit-resolution #30531 #30530]: #30274
+#10732 := (* -1::int #10701)
+#10440 := (uf_1 uf_22 ?x75!20)
+#10441 := (uf_10 #10440)
+#17909 := (+ #10441 #10732)
+#23144 := (>= #17909 0::int)
+#15426 := (= #10441 #10701)
+#30535 := (= #10701 #10441)
+#30533 := (= #10700 #10440)
+#30534 := [monotonicity #9758]: #30533
+#30536 := [monotonicity #30534]: #30535
+#30537 := [symm #30536]: #15426
+#30538 := (not #15426)
+#30539 := (or #30538 #23144)
+#30540 := [th-lemma]: #30539
+#30541 := [unit-resolution #30540 #30537]: #23144
+#5831 := (* -1::int #5830)
+#5854 := (+ #5831 #5833)
+#5855 := (+ #185 #5854)
+#5876 := (>= #5855 0::int)
+#5901 := (+ #185 #5831)
+#5903 := (>= #5901 0::int)
+#3710 := (or #4190 #4079)
+#3711 := [def-axiom]: #3710
+#7037 := [unit-resolution #3711 #7036]: #4079
+#6127 := (or #4084 #5903)
+#6128 := [quant-inst]: #6127
+#6900 := [unit-resolution #6128 #7037]: #5903
+#30542 := (not #5903)
+#30543 := (or #5876 #5907 #30542)
+#30544 := [th-lemma]: #30543
+#30545 := [unit-resolution #30544 #9692 #6900]: #5876
+#6249 := (<= #5833 0::int)
+#6242 := (or #8154 #6879)
+#5994 := (= uf_22 uf_22)
+#6880 := (not #5994)
+#6881 := (or #6880 #6879)
+#6243 := (or #8154 #6881)
+#6245 := (iff #6243 #6242)
+#6241 := (iff #6242 #6242)
+#6246 := [rewrite]: #6241
+#6891 := (iff #6881 #6879)
+#6886 := (or false #6879)
+#6889 := (iff #6886 #6879)
+#6890 := [rewrite]: #6889
+#6887 := (iff #6881 #6886)
+#6884 := (iff #6880 false)
+#6882 := (iff #6880 #4945)
+#5997 := (iff #5994 true)
+#5998 := [rewrite]: #5997
+#6883 := [monotonicity #5998]: #6882
+#6885 := [trans #6883 #4943]: #6884
+#6888 := [monotonicity #6885]: #6887
+#6892 := [trans #6888 #6890]: #6891
+#6240 := [monotonicity #6892]: #6245
+#6247 := [trans #6240 #6246]: #6245
+#6244 := [quant-inst]: #6243
+#6248 := [mp #6244 #6247]: #6242
+#22629 := [unit-resolution #6248 #3872]: #6879
+#22630 := (or #9688 #6249)
+#22631 := [th-lemma]: #22630
+#22632 := [unit-resolution #22631 #22629]: #6249
+#7858 := (+ #5830 #7883)
+#7614 := (>= #7858 0::int)
+#8012 := (= #5830 #8027)
+#23496 := (= #8027 #5830)
+#23497 := [monotonicity #9758]: #23496
+#23498 := [symm #23497]: #8012
+#23499 := (not #8012)
+#30546 := (or #23499 #7614)
+#30547 := [th-lemma]: #30546
+#30548 := [unit-resolution #30547 #23498]: #7614
+#10675 := (+ #2033 #10441)
+#10676 := (+ #185 #10675)
+#22250 := (<= #10676 0::int)
+#10649 := (= #10676 0::int)
+#10442 := (* -1::int #10441)
+#10469 := (+ uf_9 #10442)
+#10470 := (<= #10469 0::int)
+#21628 := (not #10470)
+#10503 := (+ #10403 #10441)
+#10504 := (+ #185 #10503)
+#10499 := (>= #10504 0::int)
+#10509 := (or #10470 #10499)
+#10515 := (not #10509)
+#10520 := (or #10311 #10515)
+#15425 := (or #4076 #10311 #10515)
+#10443 := (+ #1235 #10442)
+#10467 := (+ #10301 #10443)
+#10468 := (<= #10467 0::int)
+#10471 := (or #10470 #10468)
+#10466 := (not #10471)
+#10472 := (or #10307 #10466)
+#22687 := (or #4076 #10472)
+#19499 := (iff #22687 #15425)
+#21215 := (or #4076 #10520)
+#21268 := (iff #21215 #15425)
+#21592 := [rewrite]: #21268
+#21283 := (iff #22687 #21215)
+#10522 := (iff #10472 #10520)
+#10518 := (iff #10466 #10515)
+#10516 := (iff #10471 #10509)
+#10507 := (iff #10468 #10499)
+#10473 := (+ #10301 #10442)
+#10474 := (+ #1235 #10473)
+#10500 := (<= #10474 0::int)
+#10505 := (iff #10500 #10499)
+#10506 := [rewrite]: #10505
+#10501 := (iff #10468 #10500)
+#10475 := (= #10467 #10474)
+#10476 := [rewrite]: #10475
+#10502 := [monotonicity #10476]: #10501
+#10508 := [trans #10502 #10506]: #10507
+#10517 := [monotonicity #10508]: #10516
+#10519 := [monotonicity #10517]: #10518
+#10523 := [monotonicity #10336 #10519]: #10522
+#21447 := [monotonicity #10523]: #21283
+#21457 := [trans #21447 #21592]: #19499
+#20648 := [quant-inst]: #22687
+#21455 := [mp #20648 #21457]: #15425
+#30549 := [unit-resolution #21455 #6216]: #10520
+#30550 := [unit-resolution #30549 #30519]: #10515
+#22684 := (or #10509 #21628)
+#20894 := [def-axiom]: #22684
+#30551 := [unit-resolution #20894 #30550]: #21628
+#22445 := (not #10499)
+#21601 := (or #10509 #22445)
+#19500 := [def-axiom]: #21601
+#30552 := [unit-resolution #19500 #30550]: #22445
+#10680 := (or #10470 #10499 #10649)
+#22909 := (or #4068 #10470 #10499 #10649)
+#10641 := (+ #10441 #2033)
+#10642 := (+ #185 #10641)
+#10643 := (= #10642 0::int)
+#10650 := (or #10470 #10468 #10643)
+#22907 := (or #4068 #10650)
+#21602 := (iff #22907 #22909)
+#22692 := (or #4068 #10680)
+#22690 := (iff #22692 #22909)
+#22915 := [rewrite]: #22690
+#22680 := (iff #22907 #22692)
+#10681 := (iff #10650 #10680)
+#10679 := (iff #10643 #10649)
+#10677 := (= #10642 #10676)
+#10678 := [rewrite]: #10677
+#10674 := [monotonicity #10678]: #10679
+#10682 := [monotonicity #10508 #10674]: #10681
+#22900 := [monotonicity #10682]: #22680
+#21458 := [trans #22900 #22915]: #21602
+#22696 := [quant-inst]: #22907
+#21404 := [mp #22696 #21458]: #22909
+#30553 := [unit-resolution #21404 #6220]: #10680
+#30554 := [unit-resolution #30553 #30552 #30551]: #10649
+#30555 := (not #10649)
+#30556 := (or #30555 #22250)
+#30557 := [th-lemma]: #30556
+#30558 := [unit-resolution #30557 #30554]: #22250
+#30563 := (not #23144)
+#30562 := (not #5876)
+#23491 := (not #6249)
+#30561 := (not #7614)
+#30560 := (not #22250)
+#30564 := (or #30559 #30560 #30561 #23491 #30562 #30273 #30563)
+#30565 := [th-lemma]: #30564
+#30566 := [unit-resolution #30565 #30558 #30548 #22632 #30545 #30541 #30532]: #30559
+#10720 := (+ #7883 #10732)
+#10721 := (+ #2032 #10720)
+#10722 := (= #10721 0::int)
+#17953 := (>= #10721 0::int)
+#30567 := (or #17953 #30560 #30561 #23491 #30562 #30563)
+#30568 := [th-lemma]: #30567
+#30569 := [unit-resolution #30568 #30558 #30548 #22632 #30545 #30541]: #17953
+#17886 := (<= #10721 0::int)
+#17849 := (<= #17909 0::int)
+#30570 := (or #30538 #17849)
+#30571 := [th-lemma]: #30570
+#30572 := [unit-resolution #30571 #30537]: #17849
+#5875 := (<= #5855 0::int)
+#5902 := (<= #5901 0::int)
+#23487 := (not #5879)
+#23488 := (or #23487 #5902)
+#23489 := [th-lemma]: #23488
+#23490 := [unit-resolution #23489 #8605]: #5902
+#23492 := (not #5902)
+#23493 := (or #5875 #23491 #23492)
+#23494 := [th-lemma]: #23493
+#23495 := [unit-resolution #23494 #23490 #22632]: #5875
+#7859 := (<= #7858 0::int)
+#23500 := (or #23499 #7859)
+#23501 := [th-lemma]: #23500
+#23502 := [unit-resolution #23501 #23498]: #7859
+#22934 := (>= #10676 0::int)
+#30573 := (or #30555 #22934)
+#30574 := [th-lemma]: #30573
+#30575 := [unit-resolution #30574 #30554]: #22934
+#30579 := (not #17849)
+#30578 := (not #5875)
+#30577 := (not #7859)
+#30576 := (not #22934)
+#30580 := (or #17886 #30576 #30577 #5907 #30578 #30579)
+#30581 := [th-lemma]: #30580
+#30582 := [unit-resolution #30581 #30575 #23502 #23495 #9692 #30572]: #17886
+#30584 := (not #17953)
+#30583 := (not #17886)
+#30585 := (or #10722 #30583 #30584)
+#30586 := [th-lemma]: #30585
+#30587 := [unit-resolution #30586 #30582 #30569]: #10722
+#10725 := (not #10722)
+#30589 := (or #10697 #10725)
+#8020 := (up_6 uf_23 #6904)
+#5971 := (up_6 #191 uf_22)
+#29386 := (iff #5971 #8020)
+#29384 := (iff #8020 #5971)
+#29385 := [monotonicity #7417 #9758]: #29384
+#29387 := [symm #29385]: #29386
#46 := (:var 0 T5)
#45 := (:var 2 T4)
#47 := (uf_7 #45 #10 #46)
-#4105 := (pattern #47)
-#335 := (= uf_8 #46)
+#3860 := (pattern #47)
+#319 := (= uf_8 #46)
#48 := (up_6 #47 #10)
-#339 := (iff #48 #335)
-#4106 := (forall (vars (?x17 T4) (?x18 T2) (?x19 T5)) (:pat #4105) #339)
-#342 := (forall (vars (?x17 T4) (?x18 T2) (?x19 T5)) #339)
-#4109 := (iff #342 #4106)
-#4107 := (iff #339 #339)
-#4108 := [refl]: #4107
-#4110 := [quant-intro #4108]: #4109
-#1739 := (~ #342 #342)
-#1777 := (~ #339 #339)
-#1778 := [refl]: #1777
-#1740 := [nnf-pos #1778]: #1739
+#323 := (iff #48 #319)
+#3861 := (forall (vars (?x17 T4) (?x18 T2) (?x19 T5)) (:pat #3860) #323)
+#326 := (forall (vars (?x17 T4) (?x18 T2) (?x19 T5)) #323)
+#3864 := (iff #326 #3861)
+#3862 := (iff #323 #323)
+#3863 := [refl]: #3862
+#3865 := [quant-intro #3863]: #3864
+#1509 := (~ #326 #326)
+#1547 := (~ #323 #323)
+#1548 := [refl]: #1547
+#1510 := [nnf-pos #1548]: #1509
#49 := (= #46 uf_8)
#50 := (iff #48 #49)
#51 := (forall (vars (?x17 T4) (?x18 T2) (?x19 T5)) #50)
-#343 := (iff #51 #342)
-#340 := (iff #50 #339)
-#337 := (iff #49 #335)
-#338 := [rewrite]: #337
-#341 := [monotonicity #338]: #340
-#344 := [quant-intro #341]: #343
-#334 := [asserted]: #51
-#347 := [mp #334 #344]: #342
-#1779 := [mp~ #347 #1740]: #342
-#4111 := [mp #1779 #4110]: #4106
-#8930 := (not #4106)
-#8932 := (or #8930 #3772)
-#3771 := (iff #3772 #3770)
-#8926 := (or #8930 #3771)
-#8933 := (iff #8926 #8932)
-#8935 := (iff #8932 #8932)
-#8936 := [rewrite]: #8935
-#3757 := (iff #3771 #3772)
-#3763 := (iff #3772 true)
-#3765 := (iff #3763 #3772)
-#3766 := [rewrite]: #3765
-#3764 := (iff #3771 #3763)
-#3756 := [monotonicity #3762]: #3764
-#3767 := [trans #3756 #3766]: #3757
-#8934 := [monotonicity #3767]: #8933
-#8931 := [trans #8934 #8936]: #8933
-#8927 := [quant-inst]: #8926
-#9499 := [mp #8927 #8931]: #8932
-#10907 := [unit-resolution #9499 #4111]: #3772
-#10912 := [mp #10907 #10911]: #4865
-#4866 := (not #4865)
-#4870 := (or #4616 #4866)
-#10280 := (or #4332 #4616 #4866)
-#4869 := (or #4866 #4616)
-#10281 := (or #4332 #4869)
-#10339 := (iff #10281 #10280)
-#10282 := (or #4332 #4870)
-#10337 := (iff #10282 #10280)
-#10338 := [rewrite]: #10337
-#10283 := (iff #10281 #10282)
-#4871 := (iff #4869 #4870)
-#4872 := [rewrite]: #4871
-#10336 := [monotonicity #4872]: #10283
-#10341 := [trans #10336 #10338]: #10339
-#10279 := [quant-inst]: #10281
-#10342 := [mp #10279 #10341]: #10280
-#10921 := [unit-resolution #10342 #10920]: #4870
-#10922 := [unit-resolution #10921 #10912]: #4616
-#9432 := (not #4616)
-#9433 := (or #9432 #9376)
-#9434 := [th-lemma]: #9433
-#10923 := [unit-resolution #9434 #10922]: #9376
-#11761 := (<= #11766 0::int)
-#18595 := (or #18531 #11761)
-#18596 := [th-lemma]: #18595
-#18597 := [unit-resolution #18596 #18530]: #11761
-#15650 := (>= #5723 0::int)
-#18571 := (or #18581 #15650)
-#18572 := [th-lemma]: #18571
-#18570 := [unit-resolution #18572 #18580]: #15650
-#18576 := (not #15922)
-#18575 := (not #9376)
-#18574 := (not #11761)
-#18573 := (not #15650)
-#18577 := (or #15829 #18573 #18574 #18575 #18576)
-#18578 := [th-lemma]: #18577
-#17865 := [unit-resolution #18578 #18570 #18597 #10923 #18594]: #15829
-#15178 := (= #15177 0::int)
-#15183 := (not #15178)
-#15085 := (+ #2236 #11812)
-#15163 := (<= #15085 0::int)
-#18640 := (not #15163)
-#17249 := (uf_3 #5917)
-#18036 := (uf_1 #10571 #17249)
-#18037 := (uf_10 #18036)
-#18039 := (* -1::int #18037)
-#18208 := (+ #5697 #18039)
-#18250 := (>= #18208 0::int)
-#18205 := (= #5697 #18037)
-#18060 := (= #18037 #5697)
-#18054 := (= #18036 #5696)
-#17915 := (= #17249 ?x75!20)
-#17250 := (= ?x75!20 #17249)
-#17253 := (or #7845 #17250)
-#17254 := [quant-inst]: #17253
-#17866 := [unit-resolution #17254 #4076]: #17250
-#17916 := [symm #17866]: #17915
-#18059 := [monotonicity #13436 #17916]: #18054
-#18063 := [monotonicity #18059]: #18060
-#18064 := [symm #18063]: #18205
-#18065 := (not #18205)
-#18068 := (or #18065 #18250)
-#18126 := [th-lemma]: #18068
-#18127 := [unit-resolution #18126 #18064]: #18250
-#18137 := (<= #18037 0::int)
-#18138 := (not #18137)
-#18556 := (= #10571 #17249)
-#17987 := (not #18556)
-#18551 := (up_6 uf_15 #17249)
-#18562 := (or #18551 #18556)
-#18015 := (not #18562)
-#18554 := (up_6 #11533 #17249)
-#18567 := (iff #18554 #18562)
-#17962 := (or #6627 #18567)
-#18552 := (= #17249 #10571)
-#18553 := (ite #18552 #3770 #18551)
-#18555 := (iff #18554 #18553)
-#17963 := (or #6627 #18555)
-#17965 := (iff #17963 #17962)
-#17981 := (iff #17962 #17962)
-#17982 := [rewrite]: #17981
-#18568 := (iff #18555 #18567)
-#18565 := (iff #18553 #18562)
-#18559 := (ite #18556 true #18551)
-#18563 := (iff #18559 #18562)
-#18564 := [rewrite]: #18563
-#18560 := (iff #18553 #18559)
-#18557 := (iff #18552 #18556)
-#18558 := [rewrite]: #18557
-#18561 := [monotonicity #18558 #3762]: #18560
-#18566 := [trans #18561 #18564]: #18565
-#18569 := [monotonicity #18566]: #18568
-#17980 := [monotonicity #18569]: #17965
-#17983 := [trans #17980 #17982]: #17965
-#17964 := [quant-inst]: #17963
-#17984 := [mp #17964 #17983]: #17962
-#18426 := [unit-resolution #17984 #4096]: #18567
-#18020 := (not #18554)
-#5037 := (up_6 uf_23 ?x75!20)
-#5038 := (not #5037)
-#18541 := (iff #5038 #18020)
-#18441 := (iff #5037 #18554)
-#18430 := (iff #18554 #5037)
-#18431 := [monotonicity #13613 #17916]: #18430
-#18540 := [symm #18431]: #18441
-#18544 := [monotonicity #18540]: #18541
-#5658 := (or #5038 #5655)
-#15006 := (or #4332 #5038 #5655)
-#5654 := (or #5038 #5651)
-#15007 := (or #4332 #5654)
-#15427 := (iff #15007 #15006)
-#15242 := (or #4332 #5658)
-#15290 := (iff #15242 #15006)
-#15291 := [rewrite]: #15290
-#15251 := (iff #15007 #15242)
-#5659 := (iff #5654 #5658)
-#5660 := [monotonicity #5657]: #5659
-#15252 := [monotonicity #5660]: #15251
-#15428 := [trans #15252 #15291]: #15427
-#15241 := [quant-inst]: #15007
-#15429 := [mp #15241 #15428]: #15006
-#18605 := [unit-resolution #15429 #10920]: #5658
-#18427 := [unit-resolution #18605 #18534]: #5038
-#18545 := [mp #18427 #18544]: #18020
-#18018 := (not #18567)
-#18019 := (or #18018 #18554 #18015)
-#18014 := [def-axiom]: #18019
-#18546 := [unit-resolution #18014 #18545 #18426]: #18015
-#17988 := (or #18562 #17987)
-#17989 := [def-axiom]: #17988
-#18549 := [unit-resolution #17989 #18546]: #17987
-#18153 := (or #18138 #18556)
-#18155 := (or #7140 #18138 #18556)
-#18152 := (or #18556 #18138)
-#18156 := (or #7140 #18152)
-#18170 := (iff #18156 #18155)
-#18162 := (or #7140 #18153)
-#18164 := (iff #18162 #18155)
-#18165 := [rewrite]: #18164
-#18160 := (iff #18156 #18162)
-#18151 := (iff #18152 #18153)
-#18154 := [rewrite]: #18151
-#18163 := [monotonicity #18154]: #18160
-#18171 := [trans #18163 #18165]: #18170
-#18161 := [quant-inst]: #18156
-#18169 := [mp #18161 #18171]: #18155
-#18638 := [unit-resolution #18169 #4123]: #18153
-#18639 := [unit-resolution #18638 #18549]: #18138
-#18641 := (not #18250)
-#18642 := (or #18640 #18585 #18586 #18587 #18137 #18641)
-#18643 := [th-lemma]: #18642
-#18644 := [unit-resolution #18643 #18584 #18533 #10925 #18639 #18127]: #18640
-#18651 := (or #15163 #15183)
-#11817 := (up_6 uf_23 #10571)
-#18647 := (iff #3772 #11817)
-#18645 := (iff #11817 #3772)
-#18646 := [monotonicity #10469 #13436]: #18645
-#18648 := [symm #18646]: #18647
-#18649 := [mp #10907 #18648]: #11817
-#18650 := [hypothesis]: #4361
-#11821 := (not #11817)
-#15818 := (or #4366 #11821 #15163 #15183)
-#15150 := (+ #2237 #15149)
-#15151 := (+ #11814 #15150)
-#15152 := (= #15151 0::int)
-#15153 := (not #15152)
-#15154 := (+ #11814 #2237)
-#15155 := (>= #15154 0::int)
-#15156 := (or #11821 #15155 #15153)
-#15819 := (or #4366 #15156)
-#15826 := (iff #15819 #15818)
-#15186 := (or #11821 #15163 #15183)
-#15821 := (or #4366 #15186)
-#15824 := (iff #15821 #15818)
-#15825 := [rewrite]: #15824
-#15822 := (iff #15819 #15821)
-#15187 := (iff #15156 #15186)
-#15184 := (iff #15153 #15183)
-#15181 := (iff #15152 #15178)
-#15168 := (+ #11814 #15149)
-#15169 := (+ #2237 #15168)
-#15172 := (= #15169 0::int)
-#15179 := (iff #15172 #15178)
-#15180 := [rewrite]: #15179
-#15173 := (iff #15152 #15172)
-#15170 := (= #15151 #15169)
-#15171 := [rewrite]: #15170
-#15174 := [monotonicity #15171]: #15173
-#15182 := [trans #15174 #15180]: #15181
-#15185 := [monotonicity #15182]: #15184
-#15166 := (iff #15155 #15163)
-#15157 := (+ #2237 #11814)
-#15160 := (>= #15157 0::int)
-#15164 := (iff #15160 #15163)
-#15165 := [rewrite]: #15164
-#15161 := (iff #15155 #15160)
-#15158 := (= #15154 #15157)
-#15159 := [rewrite]: #15158
-#15162 := [monotonicity #15159]: #15161
-#15167 := [trans #15162 #15165]: #15166
-#15188 := [monotonicity #15167 #15185]: #15187
-#15823 := [monotonicity #15188]: #15822
-#15827 := [trans #15823 #15825]: #15826
-#15820 := [quant-inst]: #15819
-#15828 := [mp #15820 #15827]: #15818
-#18652 := [unit-resolution #15828 #18650 #18649]: #18651
-#18653 := [unit-resolution #18652 #18644]: #15183
-#18655 := (not #15830)
-#18654 := (not #15829)
-#18656 := (or #15178 #18654 #18655)
-#18657 := [th-lemma]: #18656
-#18658 := [unit-resolution #18657 #18653 #17865 #18591]: false
-#18660 := [lemma #18658]: #18659
-#27297 := [unit-resolution #18660 #27296]: #5655
-#18608 := (or #18607 #18550)
-#18609 := [th-lemma]: #18608
-#27298 := [unit-resolution #18609 #27297]: #18550
-#22669 := (not #18550)
-#22675 := (or #22674 #22669 #2248)
-#22670 := [hypothesis]: #2249
-#22671 := [hypothesis]: #18550
-#22672 := [hypothesis]: #5929
-#22673 := [th-lemma #22672 #22671 #22670]: false
-#22676 := [lemma #22673]: #22675
-#27299 := [unit-resolution #22676 #27298 #27295]: #22674
-#4003 := (or #4369 #2813)
-#3885 := [def-axiom]: #4003
-#27300 := [unit-resolution #3885 #27294]: #2813
-#16375 := (or #4218 #2810 #5929 #5934)
-#5926 := (or #5925 #5923 #5916)
-#5927 := (not #5926)
-#5930 := (or #2250 #5929 #5927)
-#16398 := (or #4218 #5930)
-#16553 := (iff #16398 #16375)
-#5937 := (or #2810 #5929 #5934)
-#16114 := (or #4218 #5937)
-#16503 := (iff #16114 #16375)
-#16391 := [rewrite]: #16503
-#16534 := (iff #16398 #16114)
-#5938 := (iff #5930 #5937)
-#5935 := (iff #5927 #5934)
-#5932 := (iff #5926 #5931)
-#5933 := [rewrite]: #5932
-#5936 := [monotonicity #5933]: #5935
-#5939 := [monotonicity #2812 #5936]: #5938
-#16550 := [monotonicity #5939]: #16534
-#16502 := [trans #16550 #16391]: #16553
-#16396 := [quant-inst]: #16398
-#16533 := [mp #16396 #16502]: #16375
-#27301 := [unit-resolution #16533 #10531 #27300 #27299]: #5934
-#16782 := (or #5931 #5924)
-#16643 := [def-axiom]: #16782
-#27302 := [unit-resolution #16643 #27301]: #5924
-#27310 := [mp #27302 #27309]: #25982
-#25983 := (not #25982)
-#27187 := (or #27170 #25983)
-#27188 := [def-axiom]: #27187
-#27311 := [unit-resolution #27188 #27310]: #27170
-#27192 := (not #27170)
-#27196 := (or #27195 #27162 #27192)
-#27197 := [def-axiom]: #27196
-#27313 := [unit-resolution #27197 #27311]: #27312
-#27314 := [unit-resolution #27313 #27285]: #27162
-#27323 := [unit-resolution #27314 #27322]: false
-#27324 := [lemma #27323]: #16889
-#16887 := (uf_24 #5912)
-#16906 := (* -1::int #16887)
-#17099 := (+ #2236 #16906)
-#17100 := (<= #17099 0::int)
-#22648 := (not #17100)
-#15991 := (not #5916)
-#16678 := (or #5931 #15991)
-#16501 := [def-axiom]: #16678
-#22641 := [unit-resolution #16501 #27301]: #15991
-#16907 := (+ #5913 #16906)
-#16908 := (>= #16907 0::int)
-#16989 := (or #4323 #16908)
-#16027 := [quant-inst]: #16989
-#22647 := [unit-resolution #16027 #10924]: #16908
-#22633 := (not #16908)
-#23110 := (or #22648 #5916 #22669 #22633)
-#22643 := [th-lemma]: #23110
-#18144 := [unit-resolution #22643 #27298 #22647 #22641]: #22648
-#17066 := (+ #5919 #16906)
-#17067 := (+ #2236 #17066)
-#17110 := (= #17067 0::int)
-#20033 := (>= #17067 0::int)
-#9449 := (>= #5921 0::int)
-#16586 := (or #5931 #5922)
-#16707 := [def-axiom]: #16586
-#16131 := [unit-resolution #16707 #27301]: #5922
-#21450 := (or #5923 #9449)
-#21454 := [th-lemma]: #21450
-#23041 := [unit-resolution #21454 #16131]: #9449
-#23063 := (not #9449)
-#23122 := (or #20033 #23063 #22669 #22633)
-#23065 := [th-lemma]: #23122
-#23044 := [unit-resolution #23065 #23041 #22647 #27298]: #20033
-#17068 := (<= #17067 0::int)
-#22638 := (<= #16907 0::int)
-#16888 := (= #5913 #16887)
-#16892 := (or #16888 #16890)
-#17827 := (or #4332 #16888 #16890)
-#16891 := (or #16890 #16888)
-#18121 := (or #4332 #16891)
-#16026 := (iff #18121 #17827)
-#18876 := (or #4332 #16892)
-#19063 := (iff #18876 #17827)
-#19159 := [rewrite]: #19063
-#19153 := (iff #18121 #18876)
-#16893 := (iff #16891 #16892)
-#16894 := [rewrite]: #16893
-#17533 := [monotonicity #16894]: #19153
-#19171 := [trans #17533 #19159]: #16026
-#19129 := [quant-inst]: #18121
-#17856 := [mp #19129 #19171]: #17827
-#18906 := [unit-resolution #17856 #10920]: #16892
-#23139 := [unit-resolution #18906 #27324]: #16888
-#23113 := (not #16888)
-#23108 := (or #23113 #22638)
-#23136 := [th-lemma]: #23108
-#23104 := [unit-resolution #23136 #23139]: #22638
-#5682 := (<= #5681 0::int)
-#20530 := (not #5682)
-#20531 := [hypothesis]: #20530
-#20442 := (or #4323 #5682)
-#5672 := (+ #5650 #2237)
-#5673 := (>= #5672 0::int)
-#20459 := (or #4323 #5673)
-#20466 := (iff #20459 #20442)
-#20473 := (iff #20442 #20442)
-#20476 := [rewrite]: #20473
-#5685 := (iff #5673 #5682)
-#5674 := (+ #2237 #5650)
-#5677 := (>= #5674 0::int)
-#5683 := (iff #5677 #5682)
-#5684 := [rewrite]: #5683
-#5678 := (iff #5673 #5677)
-#5675 := (= #5672 #5674)
-#5676 := [rewrite]: #5675
-#5679 := [monotonicity #5676]: #5678
-#5686 := [trans #5679 #5684]: #5685
-#20472 := [monotonicity #5686]: #20466
-#20477 := [trans #20472 #20476]: #20466
-#20465 := [quant-inst]: #20459
-#20527 := [mp #20465 #20477]: #20442
-#20526 := [unit-resolution #20527 #10924 #20531]: false
-#20532 := [lemma #20526]: #5682
-#15989 := (<= #5921 0::int)
-#23166 := (or #5923 #15989)
-#23129 := [th-lemma]: #23166
-#23111 := [unit-resolution #23129 #16131]: #15989
-#23142 := (not #22638)
-#23025 := (not #15989)
-#23268 := (or #17068 #23025 #20530 #23142)
-#23150 := [th-lemma]: #23268
-#23164 := [unit-resolution #23150 #23111 #20532 #23104]: #17068
-#23270 := (not #20033)
-#23269 := (not #17068)
-#23148 := (or #17110 #23269 #23270)
-#23146 := [th-lemma]: #23148
-#23271 := [unit-resolution #23146 #23164 #23044]: #17110
-#17115 := (not #17110)
-#17118 := (or #16890 #17100 #17115)
-#19968 := (or #4366 #16890 #17100 #17115)
-#17087 := (+ #2237 #5918)
-#17088 := (+ #16887 #17087)
-#17089 := (= #17088 0::int)
-#17090 := (not #17089)
-#17051 := (+ #16887 #2237)
-#17091 := (>= #17051 0::int)
-#17092 := (or #16890 #17091 #17090)
-#19969 := (or #4366 #17092)
-#19939 := (iff #19969 #19968)
-#19962 := (or #4366 #17118)
-#20004 := (iff #19962 #19968)
-#19963 := [rewrite]: #20004
-#20227 := (iff #19969 #19962)
-#17119 := (iff #17092 #17118)
-#17116 := (iff #17090 #17115)
-#17113 := (iff #17089 #17110)
-#17059 := (+ #5918 #16887)
-#17060 := (+ #2237 #17059)
-#17107 := (= #17060 0::int)
-#17111 := (iff #17107 #17110)
-#17112 := [rewrite]: #17111
-#17108 := (iff #17089 #17107)
-#17105 := (= #17088 #17060)
-#17106 := [rewrite]: #17105
-#17109 := [monotonicity #17106]: #17108
-#17114 := [trans #17109 #17112]: #17113
-#17117 := [monotonicity #17114]: #17116
-#17103 := (iff #17091 #17100)
-#17093 := (+ #2237 #16887)
-#17096 := (>= #17093 0::int)
-#17101 := (iff #17096 #17100)
-#17102 := [rewrite]: #17101
-#17097 := (iff #17091 #17096)
-#17094 := (= #17051 #17093)
-#17095 := [rewrite]: #17094
-#17098 := [monotonicity #17095]: #17097
-#17104 := [trans #17098 #17102]: #17103
-#17120 := [monotonicity #17104 #17117]: #17119
-#20114 := [monotonicity #17120]: #20227
-#20212 := [trans #20114 #19963]: #19939
-#19967 := [quant-inst]: #19969
-#20005 := [mp #19967 #20212]: #19968
-#23170 := [unit-resolution #20005 #27296]: #17118
-[unit-resolution #23170 #23271 #18144 #27324]: false
+#327 := (iff #51 #326)
+#324 := (iff #50 #323)
+#321 := (iff #49 #319)
+#322 := [rewrite]: #321
+#325 := [monotonicity #322]: #324
+#328 := [quant-intro #325]: #327
+#318 := [asserted]: #51
+#331 := [mp #318 #328]: #326
+#1549 := [mp~ #331 #1510]: #326
+#3866 := [mp #1549 #3865]: #3861
+#5984 := (not #3861)
+#5985 := (or #5984 #5971)
+#5974 := (iff #5971 #5970)
+#5986 := (or #5984 #5974)
+#5988 := (iff #5986 #5985)
+#5990 := (iff #5985 #5985)
+#5991 := [rewrite]: #5990
+#5982 := (iff #5974 #5971)
+#5977 := (iff #5971 true)
+#5980 := (iff #5977 #5971)
+#5981 := [rewrite]: #5980
+#5978 := (iff #5974 #5977)
+#5979 := [monotonicity #5976]: #5978
+#5983 := [trans #5979 #5981]: #5982
+#5989 := [monotonicity #5983]: #5988
+#5992 := [trans #5989 #5991]: #5988
+#5987 := [quant-inst]: #5986
+#5993 := [mp #5987 #5992]: #5985
+#29383 := [unit-resolution #5993 #3866]: #5971
+#29388 := [mp #29383 #29387]: #8020
+#30588 := [hypothesis]: #4122
+#8023 := (not #8020)
+#23013 := (or #4127 #8023 #10697 #10725)
+#10702 := (+ #2033 #10701)
+#10703 := (+ #8027 #10702)
+#10704 := (= #10703 0::int)
+#10717 := (not #10704)
+#10718 := (+ #8027 #2033)
+#10716 := (>= #10718 0::int)
+#10711 := (or #8023 #10716 #10717)
+#22922 := (or #4127 #10711)
+#22994 := (iff #22922 #23013)
+#10742 := (or #8023 #10697 #10725)
+#22980 := (or #4127 #10742)
+#22955 := (iff #22980 #23013)
+#22999 := [rewrite]: #22955
+#22906 := (iff #22922 #22980)
+#10798 := (iff #10711 #10742)
+#10726 := (iff #10717 #10725)
+#10719 := (iff #10704 #10722)
+#10746 := (+ #8027 #10701)
+#10747 := (+ #2033 #10746)
+#10750 := (= #10747 0::int)
+#10715 := (iff #10750 #10722)
+#10723 := [rewrite]: #10715
+#10730 := (iff #10704 #10750)
+#10748 := (= #10703 #10747)
+#10749 := [rewrite]: #10748
+#10731 := [monotonicity #10749]: #10730
+#10724 := [trans #10731 #10723]: #10719
+#10727 := [monotonicity #10724]: #10726
+#10694 := (iff #10716 #10697)
+#10712 := (+ #2033 #8027)
+#10695 := (>= #10712 0::int)
+#10698 := (iff #10695 #10697)
+#10699 := [rewrite]: #10698
+#10710 := (iff #10716 #10695)
+#10713 := (= #10718 #10712)
+#10714 := [rewrite]: #10713
+#10696 := [monotonicity #10714]: #10710
+#10745 := [trans #10696 #10699]: #10694
+#10796 := [monotonicity #10745 #10727]: #10798
+#21369 := [monotonicity #10796]: #22906
+#17898 := [trans #21369 #22999]: #22994
+#22936 := [quant-inst]: #22922
+#17881 := [mp #22936 #17898]: #23013
+#30590 := [unit-resolution #17881 #30588 #29388]: #30589
+#30591 := [unit-resolution #30590 #30587 #30566]: false
+#30593 := [lemma #30591]: #30592
+#25195 := [unit-resolution #30593 #25346]: #10311
+#35020 := (or #30518 #7777)
+#35021 := [th-lemma]: #35020
+#25253 := [unit-resolution #35021 #25195]: #7777
+#33354 := (not #7777)
+#33351 := (or #33430 #33354 #2044)
+#33352 := [hypothesis]: #11140
+#38137 := [hypothesis]: #2045
+#33346 := [hypothesis]: #7777
+#33353 := [th-lemma #33346 #38137 #33352]: false
+#33355 := [lemma #33353]: #33351
+#25263 := [unit-resolution #33355 #25253 #25348]: #33430
+#25380 := (or #11140 #11168)
+#3758 := (or #4130 #2560)
+#3640 := [def-axiom]: #3758
+#25347 := [unit-resolution #3640 #25286]: #2560
+#10947 := [unit-resolution #3637 #6214]: #3974
+#19098 := (or #3979 #2557 #11140 #11168)
+#11132 := (or #11137 #11135 #11146)
+#11138 := (not #11132)
+#11141 := (or #2046 #11140 #11138)
+#24415 := (or #3979 #11141)
+#24431 := (iff #24415 #19098)
+#11165 := (or #2557 #11140 #11168)
+#24132 := (or #3979 #11165)
+#24604 := (iff #24132 #19098)
+#24448 := [rewrite]: #24604
+#24350 := (iff #24415 #24132)
+#11171 := (iff #11141 #11165)
+#11169 := (iff #11138 #11168)
+#11166 := (iff #11132 #11142)
+#11167 := [rewrite]: #11166
+#11170 := [monotonicity #11167]: #11169
+#11172 := [monotonicity #2559 #11170]: #11171
+#24430 := [monotonicity #11172]: #24350
+#24397 := [trans #24430 #24448]: #24431
+#24390 := [quant-inst]: #24415
+#24306 := [mp #24390 #24397]: #19098
+#25352 := [unit-resolution #24306 #10947 #25347]: #25380
+#25354 := [unit-resolution #25352 #25263]: #11168
+#24534 := (or #11142 #11136)
+#24551 := [def-axiom]: #24534
+#25381 := [unit-resolution #24551 #25354]: #11136
+#31877 := (up_6 uf_23 #11097)
+#31878 := (not #31877)
+#35062 := (or #11142 #31878 #4127)
+#24524 := (not #11146)
+#38825 := [hypothesis]: #11168
+#24595 := (or #11142 #24524)
+#24514 := [def-axiom]: #24595
+#38826 := [unit-resolution #24514 #38825]: #24524
+#31875 := (uf_24 #11097)
+#31894 := (* -1::int #31875)
+#32091 := (+ #2032 #31894)
+#32092 := (<= #32091 0::int)
+#33436 := [hypothesis]: #31877
+#32104 := (+ #11149 #31894)
+#32105 := (+ #2032 #32104)
+#32106 := (= #32105 0::int)
+#35038 := (iff #11134 #32106)
+#35034 := (iff #32106 #11134)
+#35015 := (= #32105 #11133)
+#35029 := (= #11133 #32105)
+#31895 := (+ #11143 #31894)
+#33428 := (<= #31895 0::int)
+#31876 := (= #11143 #31875)
+#31880 := (or #31876 #31878)
+#31423 := (or #4093 #31876 #31878)
+#31879 := (or #31878 #31876)
+#31434 := (or #4093 #31879)
+#31443 := (iff #31434 #31423)
+#31441 := (or #4093 #31880)
+#31431 := (iff #31441 #31423)
+#31440 := [rewrite]: #31431
+#31433 := (iff #31434 #31441)
+#31881 := (iff #31879 #31880)
+#31882 := [rewrite]: #31881
+#31442 := [monotonicity #31882]: #31433
+#31436 := [trans #31442 #31440]: #31443
+#31435 := [quant-inst]: #31434
+#31444 := [mp #31435 #31436]: #31423
+#33899 := [unit-resolution #31444 #13487]: #31880
+#33904 := [unit-resolution #33899 #33436]: #31876
+#23171 := (not #31876)
+#16263 := (or #23171 #33428)
+#31645 := [th-lemma]: #16263
+#31947 := [unit-resolution #31645 #33904]: #33428
+#31896 := (>= #31895 0::int)
+#31469 := (or #4084 #31896)
+#31470 := [quant-inst]: #31469
+#29796 := [unit-resolution #31470 #7037]: #31896
+#10405 := (<= #10404 0::int)
+#25597 := [unit-resolution #30593 #30588]: #10311
+#34536 := (or #30518 #10405)
+#34894 := [th-lemma]: #34536
+#34869 := [unit-resolution #34894 #25597]: #10405
+#35016 := [unit-resolution #35021 #25597]: #7777
+#35035 := [th-lemma #35016 #34869 #29796 #31947]: #35029
+#35036 := [symm #35035]: #35015
+#35037 := [monotonicity #35036]: #35034
+#35039 := [symm #35037]: #35038
+#24631 := (or #11142 #11134)
+#18952 := [def-axiom]: #24631
+#33451 := [unit-resolution #18952 #38825]: #11134
+#35064 := [mp #33451 #35039]: #32106
+#32111 := (not #32106)
+#32114 := (or #31878 #32092 #32111)
+#31607 := (or #4127 #31878 #32092 #32111)
+#32078 := (+ #2033 #11148)
+#32079 := (+ #31875 #32078)
+#32080 := (= #32079 0::int)
+#32081 := (not #32080)
+#32082 := (+ #31875 #2033)
+#32083 := (>= #32082 0::int)
+#32084 := (or #31878 #32083 #32081)
+#31608 := (or #4127 #32084)
+#31636 := (iff #31608 #31607)
+#31565 := (or #4127 #32114)
+#31606 := (iff #31565 #31607)
+#31635 := [rewrite]: #31606
+#31633 := (iff #31608 #31565)
+#32115 := (iff #32084 #32114)
+#32112 := (iff #32081 #32111)
+#32109 := (iff #32080 #32106)
+#32097 := (+ #11148 #31875)
+#32098 := (+ #2033 #32097)
+#32101 := (= #32098 0::int)
+#32107 := (iff #32101 #32106)
+#32108 := [rewrite]: #32107
+#32102 := (iff #32080 #32101)
+#32099 := (= #32079 #32098)
+#32100 := [rewrite]: #32099
+#32103 := [monotonicity #32100]: #32102
+#32110 := [trans #32103 #32108]: #32109
+#32113 := [monotonicity #32110]: #32112
+#32095 := (iff #32083 #32092)
+#32085 := (+ #2033 #31875)
+#32088 := (>= #32085 0::int)
+#32093 := (iff #32088 #32092)
+#32094 := [rewrite]: #32093
+#32089 := (iff #32083 #32088)
+#32086 := (= #32082 #32085)
+#32087 := [rewrite]: #32086
+#32090 := [monotonicity #32087]: #32089
+#32096 := [trans #32090 #32094]: #32095
+#32116 := [monotonicity #32096 #32113]: #32115
+#31634 := [monotonicity #32116]: #31633
+#17991 := [trans #31634 #31635]: #31636
+#31598 := [quant-inst]: #31608
+#23660 := [mp #31598 #17991]: #31607
+#35065 := [unit-resolution #23660 #30588]: #32114
+#35066 := [unit-resolution #35065 #35064 #33436]: #32092
+#35067 := [th-lemma #35016 #29796 #35066 #38826]: false
+#35063 := [lemma #35067]: #35062
+#25287 := [unit-resolution #35063 #25354 #25346]: #31878
+#38842 := (or #31877 #11137)
+#36185 := (up_6 #8048 #11097)
+#38820 := (not #36185)
+#38838 := (iff #31878 #38820)
+#38836 := (iff #31877 #36185)
+#38834 := (iff #36185 #31877)
+#38835 := [monotonicity #8592]: #38834
+#38837 := [symm #38835]: #38836
+#38839 := [monotonicity #38837]: #38838
+#38833 := [hypothesis]: #31878
+#38840 := [mp #38833 #38839]: #38820
+#36187 := (= #6904 #11097)
+#36206 := (or #11136 #36187)
+#36190 := (iff #36185 #36206)
+#38800 := (or #6010 #36190)
+#36182 := (= #11097 #6904)
+#36183 := (ite #36182 #5970 #11136)
+#36186 := (iff #36185 #36183)
+#38801 := (or #6010 #36186)
+#38803 := (iff #38801 #38800)
+#38805 := (iff #38800 #38800)
+#38806 := [rewrite]: #38805
+#36191 := (iff #36186 #36190)
+#36209 := (iff #36183 #36206)
+#36203 := (ite #36187 true #11136)
+#36207 := (iff #36203 #36206)
+#36208 := [rewrite]: #36207
+#36204 := (iff #36183 #36203)
+#36188 := (iff #36182 #36187)
+#36202 := [rewrite]: #36188
+#36205 := [monotonicity #36202 #5976]: #36204
+#36210 := [trans #36205 #36208]: #36209
+#36192 := [monotonicity #36210]: #36191
+#38804 := [monotonicity #36192]: #38803
+#38807 := [trans #38804 #38806]: #38803
+#38802 := [quant-inst]: #38801
+#38808 := [mp #38802 #38807]: #38800
+#38827 := [unit-resolution #38808 #3851]: #36190
+#38817 := (not #36190)
+#38830 := (or #38817 #36185)
+#38828 := [hypothesis]: #11136
+#38809 := (or #36206 #11137)
+#38810 := [def-axiom]: #38809
+#38829 := [unit-resolution #38810 #38828]: #36206
+#38814 := (not #36206)
+#38818 := (or #38817 #36185 #38814)
+#38819 := [def-axiom]: #38818
+#38831 := [unit-resolution #38819 #38829]: #38830
+#38832 := [unit-resolution #38831 #38827]: #36185
+#38841 := [unit-resolution #38832 #38840]: false
+#38843 := [lemma #38841]: #38842
+#25282 := [unit-resolution #38843 #25287 #25381]: false
+#25366 := [lemma #25282]: #4130
+#5604 := (= #108 #219)
+#25448 := (iff #5604 #220)
+#25447 := [commutativity]: #1322
+#25445 := (iff #5604 #708)
+#6918 := [unit-resolution #3676 #6214]: #109
+#25446 := [monotonicity #6918]: #25445
+#25449 := [trans #25446 #25447]: #25448
+#5555 := (uf_10 #5488)
+#5558 := (* -1::int #5555)
+#5571 := (+ uf_9 #5558)
+#5572 := (<= #5571 0::int)
+#5568 := (+ #1235 #5558)
+#5569 := (+ #108 #5568)
+#5570 := (<= #5569 0::int)
+#5606 := (or #5570 #5572)
+#5486 := (+ #108 #1235)
+#7088 := (<= #5486 0::int)
+#3577 := (<= #108 0::int)
+#6919 := (or #3266 #3577)
+#6920 := [th-lemma]: #6919
+#6921 := [unit-resolution #6920 #6918]: #3577
+#5969 := (>= #185 0::int)
+#3681 := (or #4214 #3941)
+#3661 := [def-axiom]: #3681
+#7734 := [unit-resolution #3661 #6214]: #3941
+#5972 := (or #3946 #5969)
+#5973 := [quant-inst]: #5972
+#7735 := [unit-resolution #5973 #7734]: #5969
+#5140 := (not #3577)
+#25432 := (not #5969)
+#25433 := (or #7088 #25432 #5140)
+#25434 := [th-lemma]: #25433
+#25435 := [unit-resolution #25434 #7735 #6921]: #7088
+#8448 := (+ #5555 #5837)
+#8450 := (>= #8448 0::int)
+#8695 := (uf_1 #6904 uf_11)
+#8696 := (uf_10 #8695)
+#9032 := (= #8696 0::int)
+#22996 := (not #9032)
+#15055 := (>= #8696 0::int)
+#22691 := (not #15055)
+#10009 := (not #8450)
+#23412 := [hypothesis]: #10009
+#22699 := (or #8450 #22691)
+#8632 := (* -1::int #8696)
+#15323 := (+ #5555 #8632)
+#15325 := (>= #15323 0::int)
+#9629 := (= #5555 #8696)
+#23372 := (= #8696 #5555)
+#23413 := (= #8695 #5488)
+#23441 := [monotonicity #9758]: #23413
+#23440 := [monotonicity #23441]: #23372
+#23446 := [symm #23440]: #9629
+#23447 := (not #9629)
+#23449 := (or #23447 #15325)
+#23379 := [th-lemma]: #23449
+#22669 := [unit-resolution #23379 #23446]: #15325
+#22695 := (not #15325)
+#23448 := (or #8450 #23491 #22691 #22695)
+#22698 := [th-lemma]: #23448
+#22700 := [unit-resolution #22698 #22669 #22632]: #22699
+#22883 := [unit-resolution #22700 #23412]: #22691
+#23043 := (or #22996 #15055)
+#23450 := [th-lemma]: #23043
+#23034 := [unit-resolution #23450 #22883]: #22996
+#8929 := (= uf_11 #6904)
+#8788 := (<= #8696 0::int)
+#8449 := (<= #8448 0::int)
+#10435 := (or #8450 #8449)
+#10434 := [th-lemma]: #10435
+#23407 := [unit-resolution #10434 #23412]: #8449
+#10438 := (not #8449)
+#23443 := (or #10438 #8788)
+#22944 := (or #10438 #23491 #8788 #22695)
+#23411 := [th-lemma]: #22944
+#23409 := [unit-resolution #23411 #22669 #22632]: #23443
+#23442 := [unit-resolution #23409 #23407]: #8788
+#8794 := (not #8788)
+#8933 := (or #8794 #8929)
+#12371 := (or #7324 #8794 #8929)
+#8795 := (= #6904 uf_11)
+#8796 := (or #8795 #8794)
+#12524 := (or #7324 #8796)
+#13058 := (iff #12524 #12371)
+#12579 := (or #7324 #8933)
+#13029 := (iff #12579 #12371)
+#13057 := [rewrite]: #13029
+#12620 := (iff #12524 #12579)
+#8965 := (iff #8796 #8933)
+#8931 := (or #8929 #8794)
+#8934 := (iff #8931 #8933)
+#8964 := [rewrite]: #8934
+#8932 := (iff #8796 #8931)
+#8924 := (iff #8795 #8929)
+#8930 := [rewrite]: #8924
+#8928 := [monotonicity #8930]: #8932
+#8966 := [trans #8928 #8964]: #8965
+#12543 := [monotonicity #8966]: #12620
+#13028 := [trans #12543 #13057]: #13058
+#12621 := [quant-inst]: #12524
+#13055 := [mp #12621 #13028]: #12371
+#23444 := [unit-resolution #13055 #3878]: #8933
+#20374 := [unit-resolution #23444 #23442]: #8929
+#9008 := (not #8929)
+#9009 := (or #9008 #9032)
+#12859 := (or #8154 #9008 #9032)
+#9033 := (not #8795)
+#9007 := (or #9033 #9032)
+#13122 := (or #8154 #9007)
+#13134 := (iff #13122 #12859)
+#12878 := (or #8154 #9009)
+#13147 := (iff #12878 #12859)
+#13150 := [rewrite]: #13147
+#13131 := (iff #13122 #12878)
+#9010 := (iff #9007 #9009)
+#9006 := (iff #9033 #9008)
+#9031 := [monotonicity #8930]: #9006
+#9011 := [monotonicity #9031]: #9010
+#13120 := [monotonicity #9011]: #13131
+#13149 := [trans #13120 #13150]: #13134
+#13056 := [quant-inst]: #13122
+#13151 := [mp #13056 #13149]: #12859
+#23445 := [unit-resolution #13151 #3872]: #9009
+#11092 := [unit-resolution #23445 #20374 #23034]: false
+#17852 := [lemma #11092]: #8450
+#25436 := (not #7088)
+#25439 := (or #5570 #10009 #25436)
+#25437 := (or #5570 #5907 #10009 #25436)
+#25438 := [th-lemma]: #25437
+#25440 := [unit-resolution #25438 #9692]: #25439
+#25441 := [unit-resolution #25440 #17852 #25435]: #5570
+#5629 := (not #5570)
+#5630 := (or #5606 #5629)
+#5631 := [def-axiom]: #5630
+#25442 := [unit-resolution #5631 #25441]: #5606
+#5609 := (not #5606)
+#5612 := (or #5604 #5609)
+#5615 := (or #4076 #5604 #5609)
+#5602 := (or #5572 #5570)
+#5603 := (not #5602)
+#5605 := (or #5604 #5603)
+#5616 := (or #4076 #5605)
+#5623 := (iff #5616 #5615)
+#5618 := (or #4076 #5612)
+#5621 := (iff #5618 #5615)
+#5622 := [rewrite]: #5621
+#5619 := (iff #5616 #5618)
+#5613 := (iff #5605 #5612)
+#5610 := (iff #5603 #5609)
+#5607 := (iff #5602 #5606)
+#5608 := [rewrite]: #5607
+#5611 := [monotonicity #5608]: #5610
+#5614 := [monotonicity #5611]: #5613
+#5620 := [monotonicity #5614]: #5619
+#5624 := [trans #5620 #5622]: #5623
+#5617 := [quant-inst]: #5616
+#5625 := [mp #5617 #5624]: #5615
+#25443 := [unit-resolution #5625 #6216]: #5612
+#25444 := [unit-resolution #25443 #25442]: #5604
+#25450 := [mp #25444 #25449]: #220
+#3717 := (or #4178 #4172)
+#3721 := [def-axiom]: #3717
+#13489 := [unit-resolution #3721 #13486]: #4172
+#3716 := (or #4175 #1324 #4169)
+#3718 := [def-axiom]: #3716
+#25451 := [unit-resolution #3718 #13489]: #4172
+#25452 := [unit-resolution #25451 #25450]: #4169
+#3723 := (or #4166 #4160)
+#3725 := [def-axiom]: #3723
+#25453 := [unit-resolution #3725 #25452]: #4160
+#25454 := (or #4163 #4157)
+#8769 := (or #4093 #1963)
+#8407 := (uf_4 uf_14 ?x65!15)
+#8699 := (>= #8407 0::int)
+#8702 := (or #3946 #8699)
+#8703 := [quant-inst]: #8702
+#8764 := [unit-resolution #8703 #7734]: #8699
+#9046 := [hypothesis]: #1964
+#8438 := (* -1::int #8407)
+#8439 := (+ #1962 #8438)
+#8763 := (>= #8439 0::int)
+#8412 := (= #1962 #8407)
+#9060 := [hypothesis]: #4088
+#9082 := (or #8412 #1963 #4093)
+#8454 := (uf_1 uf_22 ?x65!15)
+#8455 := (uf_10 #8454)
+#8429 := (* -1::int #1962)
+#8511 := (+ #8429 #8455)
+#8512 := (+ #185 #8511)
+#8532 := (<= #8512 0::int)
+#8515 := (= #8512 0::int)
+#8456 := (* -1::int #8455)
+#8460 := (+ uf_9 #8456)
+#8461 := (<= #8460 0::int)
+#8499 := (not #8461)
+#8472 := (+ #8438 #8455)
+#8473 := (+ #185 #8472)
+#8474 := (>= #8473 0::int)
+#8479 := (or #8461 #8474)
+#8482 := (not #8479)
+#9034 := (not #8412)
+#9035 := [hypothesis]: #9034
+#8485 := (or #8412 #8482)
+#8488 := (or #4076 #8412 #8482)
+#8457 := (+ #1235 #8456)
+#8458 := (+ #8407 #8457)
+#8459 := (<= #8458 0::int)
+#8462 := (or #8461 #8459)
+#8463 := (not #8462)
+#8408 := (= #8407 #1962)
+#8464 := (or #8408 #8463)
+#8489 := (or #4076 #8464)
+#8496 := (iff #8489 #8488)
+#8491 := (or #4076 #8485)
+#8494 := (iff #8491 #8488)
+#8495 := [rewrite]: #8494
+#8492 := (iff #8489 #8491)
+#8486 := (iff #8464 #8485)
+#8483 := (iff #8463 #8482)
+#8480 := (iff #8462 #8479)
+#8477 := (iff #8459 #8474)
+#8465 := (+ #8407 #8456)
+#8466 := (+ #1235 #8465)
+#8469 := (<= #8466 0::int)
+#8475 := (iff #8469 #8474)
+#8476 := [rewrite]: #8475
+#8470 := (iff #8459 #8469)
+#8467 := (= #8458 #8466)
+#8468 := [rewrite]: #8467
+#8471 := [monotonicity #8468]: #8470
+#8478 := [trans #8471 #8476]: #8477
+#8481 := [monotonicity #8478]: #8480
+#8484 := [monotonicity #8481]: #8483
+#8413 := (iff #8408 #8412)
+#8414 := [rewrite]: #8413
+#8487 := [monotonicity #8414 #8484]: #8486
+#8493 := [monotonicity #8487]: #8492
+#8497 := [trans #8493 #8495]: #8496
+#8490 := [quant-inst]: #8489
+#8498 := [mp #8490 #8497]: #8488
+#9036 := [unit-resolution #8498 #6216]: #8485
+#9037 := [unit-resolution #9036 #9035]: #8482
+#8500 := (or #8479 #8499)
+#8501 := [def-axiom]: #8500
+#9038 := [unit-resolution #8501 #9037]: #8499
+#8502 := (not #8474)
+#8503 := (or #8479 #8502)
+#8504 := [def-axiom]: #8503
+#9039 := [unit-resolution #8504 #9037]: #8502
+#8518 := (or #8461 #8474 #8515)
+#8521 := (or #4068 #8461 #8474 #8515)
+#8507 := (+ #8455 #8429)
+#8508 := (+ #185 #8507)
+#8509 := (= #8508 0::int)
+#8510 := (or #8461 #8459 #8509)
+#8522 := (or #4068 #8510)
+#8529 := (iff #8522 #8521)
+#8524 := (or #4068 #8518)
+#8527 := (iff #8524 #8521)
+#8528 := [rewrite]: #8527
+#8525 := (iff #8522 #8524)
+#8519 := (iff #8510 #8518)
+#8516 := (iff #8509 #8515)
+#8513 := (= #8508 #8512)
+#8514 := [rewrite]: #8513
+#8517 := [monotonicity #8514]: #8516
+#8520 := [monotonicity #8478 #8517]: #8519
+#8526 := [monotonicity #8520]: #8525
+#8530 := [trans #8526 #8528]: #8529
+#8523 := [quant-inst]: #8522
+#8531 := [mp #8523 #8530]: #8521
+#9040 := [unit-resolution #8531 #6220]: #8518
+#9041 := [unit-resolution #9040 #9039 #9038]: #8515
+#9042 := (not #8515)
+#9043 := (or #9042 #8532)
+#9044 := [th-lemma]: #9043
+#9045 := [unit-resolution #9044 #9041]: #8532
+#8903 := (uf_1 #8349 ?x65!15)
+#8904 := (uf_10 #8903)
+#8910 := (* -1::int #8904)
+#8982 := (+ #8455 #8910)
+#8986 := (>= #8982 0::int)
+#8981 := (= #8455 #8904)
+#9052 := (= #8904 #8455)
+#9050 := (= #8903 #8454)
+#9051 := [monotonicity #9049]: #9050
+#9053 := [monotonicity #9051]: #9052
+#9054 := [symm #9053]: #8981
+#9055 := (not #8981)
+#9056 := (or #9055 #8986)
+#9057 := [th-lemma]: #9056
+#9058 := [unit-resolution #9057 #9054]: #8986
+#8974 := (>= #8904 0::int)
+#8935 := (<= #8904 0::int)
+#8936 := (not #8935)
+#8859 := (= ?x65!15 #8349)
+#8887 := (not #8859)
+#8612 := (up_6 uf_15 ?x65!15)
+#8867 := (or #8612 #8859)
+#8890 := (not #8867)
+#8861 := (uf_7 uf_15 #8349 uf_8)
+#8862 := (up_6 #8861 ?x65!15)
+#8872 := (iff #8862 #8867)
+#8875 := (or #6010 #8872)
+#8860 := (ite #8859 #5970 #8612)
+#8863 := (iff #8862 #8860)
+#8876 := (or #6010 #8863)
+#8878 := (iff #8876 #8875)
+#8880 := (iff #8875 #8875)
+#8881 := [rewrite]: #8880
+#8873 := (iff #8863 #8872)
+#8870 := (iff #8860 #8867)
+#8864 := (ite #8859 true #8612)
+#8868 := (iff #8864 #8867)
+#8869 := [rewrite]: #8868
+#8865 := (iff #8860 #8864)
+#8866 := [monotonicity #5976]: #8865
+#8871 := [trans #8866 #8869]: #8870
+#8874 := [monotonicity #8871]: #8873
+#8879 := [monotonicity #8874]: #8878
+#8882 := [trans #8879 #8881]: #8878
+#8877 := [quant-inst]: #8876
+#8883 := [mp #8877 #8882]: #8875
+#9059 := [unit-resolution #8883 #3851]: #8872
+#8896 := (not #8862)
+#8409 := (up_6 uf_23 ?x65!15)
+#8410 := (not #8409)
+#9071 := (iff #8410 #8896)
+#9069 := (iff #8409 #8862)
+#9067 := (iff #8862 #8409)
+#9065 := (= #8861 uf_23)
+#9063 := (= #8861 #191)
+#9064 := [monotonicity #9049]: #9063
+#9066 := [trans #9064 #7413]: #9065
+#9068 := [monotonicity #9066]: #9067
+#9070 := [symm #9068]: #9069
+#9072 := [monotonicity #9070]: #9071
+#8415 := (or #8410 #8412)
+#8418 := (or #4093 #8410 #8412)
+#8411 := (or #8410 #8408)
+#8419 := (or #4093 #8411)
+#8426 := (iff #8419 #8418)
+#8421 := (or #4093 #8415)
+#8424 := (iff #8421 #8418)
+#8425 := [rewrite]: #8424
+#8422 := (iff #8419 #8421)
+#8416 := (iff #8411 #8415)
+#8417 := [monotonicity #8414]: #8416
+#8423 := [monotonicity #8417]: #8422
+#8427 := [trans #8423 #8425]: #8426
+#8420 := [quant-inst]: #8419
+#8428 := [mp #8420 #8427]: #8418
+#9061 := [unit-resolution #8428 #9060]: #8415
+#9062 := [unit-resolution #9061 #9035]: #8410
+#9073 := [mp #9062 #9072]: #8896
+#8893 := (not #8872)
+#8894 := (or #8893 #8862 #8890)
+#8895 := [def-axiom]: #8894
+#9074 := [unit-resolution #8895 #9073 #9059]: #8890
+#8888 := (or #8867 #8887)
+#8889 := [def-axiom]: #8888
+#9075 := [unit-resolution #8889 #9074]: #8887
+#8941 := (or #8859 #8936)
+#8944 := (or #7324 #8859 #8936)
+#8937 := (= #8349 ?x65!15)
+#8938 := (or #8937 #8936)
+#8945 := (or #7324 #8938)
+#8952 := (iff #8945 #8944)
+#8947 := (or #7324 #8941)
+#8950 := (iff #8947 #8944)
+#8951 := [rewrite]: #8950
+#8948 := (iff #8945 #8947)
+#8942 := (iff #8938 #8941)
+#8939 := (iff #8937 #8859)
+#8940 := [rewrite]: #8939
+#8943 := [monotonicity #8940]: #8942
+#8949 := [monotonicity #8943]: #8948
+#8953 := [trans #8949 #8951]: #8952
+#8946 := [quant-inst]: #8945
+#8954 := [mp #8946 #8953]: #8944
+#9076 := [unit-resolution #8954 #3878]: #8941
+#9077 := [unit-resolution #9076 #9075]: #8936
+#9078 := (or #8974 #8935)
+#9079 := [th-lemma]: #9078
+#9080 := [unit-resolution #9079 #9077]: #8974
+#9081 := [th-lemma #9080 #9058 #9046 #9045 #7735]: false
+#9083 := [lemma #9081]: #9082
+#8765 := [unit-resolution #9083 #9060 #9046]: #8412
+#8766 := (or #9034 #8763)
+#8767 := [th-lemma]: #8766
+#8762 := [unit-resolution #8767 #8765]: #8763
+#8768 := [th-lemma #8762 #9046 #8764]: false
+#8770 := [lemma #8768]: #8769
+#13488 := [unit-resolution #8770 #13487]: #1963
+#3737 := (or #4163 #1964 #4157)
+#3731 := [def-axiom]: #3737
+#25455 := [unit-resolution #3731 #13488]: #25454
+#25456 := [unit-resolution #25455 #25453]: #4157
+#3742 := (or #4154 #4148)
+#3743 := [def-axiom]: #3742
+#29389 := [unit-resolution #3743 #25456]: #4148
+#29390 := (or #4151 #4145)
+#6820 := (+ #1978 #5831)
+#6821 := (<= #6820 0::int)
+#6819 := (= #1978 #5830)
+#6333 := (= uf_22 ?x67!17)
+#6328 := (up_6 uf_15 ?x67!17)
+#6339 := (or #6328 #6333)
+#6331 := (up_6 #191 ?x67!17)
+#6344 := (iff #6331 #6339)
+#6347 := (or #6010 #6344)
+#6329 := (= ?x67!17 uf_22)
+#6330 := (ite #6329 #5970 #6328)
+#6332 := (iff #6331 #6330)
+#6348 := (or #6010 #6332)
+#6350 := (iff #6348 #6347)
+#6352 := (iff #6347 #6347)
+#6353 := [rewrite]: #6352
+#6345 := (iff #6332 #6344)
+#6342 := (iff #6330 #6339)
+#6336 := (ite #6333 true #6328)
+#6340 := (iff #6336 #6339)
+#6341 := [rewrite]: #6340
+#6337 := (iff #6330 #6336)
+#6334 := (iff #6329 #6333)
+#6335 := [rewrite]: #6334
+#6338 := [monotonicity #6335 #5976]: #6337
+#6343 := [trans #6338 #6341]: #6342
+#6346 := [monotonicity #6343]: #6345
+#6351 := [monotonicity #6346]: #6350
+#6354 := [trans #6351 #6353]: #6350
+#6349 := [quant-inst]: #6348
+#6355 := [mp #6349 #6354]: #6347
+#7414 := [unit-resolution #6355 #3851]: #6344
+#7421 := (iff #1983 #6331)
+#7419 := (iff #6331 #1983)
+#7420 := [monotonicity #7413]: #7419
+#7422 := [symm #7420]: #7421
+#7415 := [hypothesis]: #3080
+#3619 := (or #3075 #1983)
+#3620 := [def-axiom]: #3619
+#7416 := [unit-resolution #3620 #7415]: #1983
+#7423 := [mp #7416 #7422]: #6331
+#6368 := (not #6331)
+#6365 := (not #6344)
+#6369 := (or #6365 #6368 #6339)
+#6370 := [def-axiom]: #6369
+#7454 := [unit-resolution #6370 #7423 #7414]: #6339
+#6356 := (not #6328)
+#6308 := (uf_4 uf_14 ?x67!17)
+#6386 := (* -1::int #6308)
+#6605 := (+ #185 #6386)
+#6961 := (>= #6605 0::int)
+#7041 := (not #6961)
+#6266 := (= uf_22 ?x68!16)
+#6292 := (not #6266)
+#6257 := (up_6 uf_15 ?x68!16)
+#6272 := (or #6257 #6266)
+#6295 := (not #6272)
+#6264 := (up_6 #191 ?x68!16)
+#6277 := (iff #6264 #6272)
+#6280 := (or #6010 #6277)
+#6260 := (= ?x68!16 uf_22)
+#6263 := (ite #6260 #5970 #6257)
+#6265 := (iff #6264 #6263)
+#6281 := (or #6010 #6265)
+#6283 := (iff #6281 #6280)
+#6285 := (iff #6280 #6280)
+#6286 := [rewrite]: #6285
+#6278 := (iff #6265 #6277)
+#6275 := (iff #6263 #6272)
+#6269 := (ite #6266 true #6257)
+#6273 := (iff #6269 #6272)
+#6274 := [rewrite]: #6273
+#6270 := (iff #6263 #6269)
+#6267 := (iff #6260 #6266)
+#6268 := [rewrite]: #6267
+#6271 := [monotonicity #6268 #5976]: #6270
+#6276 := [trans #6271 #6274]: #6275
+#6279 := [monotonicity #6276]: #6278
+#6284 := [monotonicity #6279]: #6283
+#6287 := [trans #6284 #6286]: #6283
+#6282 := [quant-inst]: #6281
+#6288 := [mp #6282 #6287]: #6280
+#7455 := [unit-resolution #6288 #3851]: #6277
+#6301 := (not #6264)
+#7491 := (iff #1985 #6301)
+#7461 := (iff #1984 #6264)
+#7459 := (iff #6264 #1984)
+#7460 := [monotonicity #7413]: #7459
+#7490 := [symm #7460]: #7461
+#7492 := [monotonicity #7490]: #7491
+#3762 := (or #3075 #1985)
+#3763 := [def-axiom]: #3762
+#7458 := [unit-resolution #3763 #7415]: #1985
+#7493 := [mp #7458 #7492]: #6301
+#6298 := (not #6277)
+#6299 := (or #6298 #6264 #6295)
+#6300 := [def-axiom]: #6299
+#7494 := [unit-resolution #6300 #7493 #7455]: #6295
+#6293 := (or #6272 #6292)
+#6294 := [def-axiom]: #6293
+#7495 := [unit-resolution #6294 #7494]: #6292
+#6483 := (uf_4 uf_14 ?x68!16)
+#6494 := (* -1::int #6483)
+#6696 := (+ #185 #6494)
+#6697 := (<= #6696 0::int)
+#6289 := (not #6257)
+#6290 := (or #6272 #6289)
+#6291 := [def-axiom]: #6290
+#7496 := [unit-resolution #6291 #7494]: #6289
+#6702 := (or #6257 #6697)
+#3682 := (or #4202 #4053)
+#3683 := [def-axiom]: #3682
+#7497 := [unit-resolution #3683 #6215]: #4053
+#6705 := (or #4058 #6257 #6697)
+#6685 := (+ #6483 #1235)
+#6686 := (>= #6685 0::int)
+#6689 := (or #6257 #6686)
+#6706 := (or #4058 #6689)
+#6713 := (iff #6706 #6705)
+#6708 := (or #4058 #6702)
+#6711 := (iff #6708 #6705)
+#6712 := [rewrite]: #6711
+#6709 := (iff #6706 #6708)
+#6703 := (iff #6689 #6702)
+#6700 := (iff #6686 #6697)
+#6690 := (+ #1235 #6483)
+#6693 := (>= #6690 0::int)
+#6698 := (iff #6693 #6697)
+#6699 := [rewrite]: #6698
+#6694 := (iff #6686 #6693)
+#6691 := (= #6685 #6690)
+#6692 := [rewrite]: #6691
+#6695 := [monotonicity #6692]: #6694
+#6701 := [trans #6695 #6699]: #6700
+#6704 := [monotonicity #6701]: #6703
+#6710 := [monotonicity #6704]: #6709
+#6714 := [trans #6710 #6712]: #6713
+#6707 := [quant-inst]: #6706
+#6715 := [mp #6707 #6714]: #6705
+#7498 := [unit-resolution #6715 #7497]: #6702
+#7499 := [unit-resolution #7498 #7496]: #6697
+#3764 := (not #2519)
+#3765 := (or #3075 #3764)
+#3759 := [def-axiom]: #3765
+#7500 := [unit-resolution #3759 #7415]: #3764
+#7343 := (not #6697)
+#7344 := (or #7041 #2519 #7343 #6266)
+#7098 := [hypothesis]: #6697
+#7034 := [hypothesis]: #3764
+#6387 := (+ #1978 #6386)
+#6388 := (<= #6387 0::int)
+#6393 := (or #4084 #6388)
+#6376 := (+ #6308 #1979)
+#6377 := (>= #6376 0::int)
+#6394 := (or #4084 #6377)
+#6396 := (iff #6394 #6393)
+#6398 := (iff #6393 #6393)
+#6399 := [rewrite]: #6398
+#6391 := (iff #6377 #6388)
+#6380 := (+ #1979 #6308)
+#6383 := (>= #6380 0::int)
+#6389 := (iff #6383 #6388)
+#6390 := [rewrite]: #6389
+#6384 := (iff #6377 #6383)
+#6381 := (= #6376 #6380)
+#6382 := [rewrite]: #6381
+#6385 := [monotonicity #6382]: #6384
+#6392 := [trans #6385 #6390]: #6391
+#6397 := [monotonicity #6392]: #6396
+#6400 := [trans #6397 #6399]: #6396
+#6395 := [quant-inst]: #6394
+#6401 := [mp #6395 #6400]: #6393
+#7038 := [unit-resolution #6401 #7037]: #6388
+#7039 := [hypothesis]: #6961
+#6495 := (+ #1980 #6494)
+#6559 := (>= #6495 0::int)
+#6522 := (= #1980 #6483)
+#6510 := (uf_1 uf_22 ?x68!16)
+#6511 := (uf_10 #6510)
+#6532 := (+ #6494 #6511)
+#6533 := (+ #185 #6532)
+#6534 := (>= #6533 0::int)
+#6512 := (* -1::int #6511)
+#6516 := (+ uf_9 #6512)
+#6517 := (<= #6516 0::int)
+#6539 := (or #6517 #6534)
+#6572 := (+ #2517 #6511)
+#6573 := (+ #185 #6572)
+#6576 := (= #6573 0::int)
+#6784 := (not #6576)
+#6593 := (<= #6573 0::int)
+#7348 := (not #6593)
+#7173 := [hypothesis]: #6292
+#7349 := (or #7348 #7041 #2519 #6266)
+#6828 := (<= #6511 0::int)
+#6829 := (not #6828)
+#7097 := (or #7324 #6266 #6829)
+#6830 := (or #6266 #6829)
+#7326 := (or #7324 #6830)
+#7328 := (iff #7326 #7097)
+#7339 := [rewrite]: #7328
+#7327 := [quant-inst]: #7326
+#7340 := [mp #7327 #7339]: #7097
+#7323 := [unit-resolution #7340 #3878 #7173]: #6829
+#7346 := [hypothesis]: #6593
+#7347 := [th-lemma #7346 #7039 #7038 #7034 #7323]: false
+#7350 := [lemma #7347]: #7349
+#7178 := [unit-resolution #7350 #7039 #7034 #7173]: #7348
+#7179 := (or #6784 #6593)
+#7309 := [th-lemma]: #7179
+#7321 := [unit-resolution #7309 #7178]: #6784
+#6791 := (or #6539 #6576)
+#6785 := [hypothesis]: #6784
+#6560 := (not #6517)
+#6542 := (not #6539)
+#6786 := [hypothesis]: #6542
+#6561 := (or #6539 #6560)
+#6562 := [def-axiom]: #6561
+#6787 := [unit-resolution #6562 #6786]: #6560
+#6563 := (not #6534)
+#6564 := (or #6539 #6563)
+#6565 := [def-axiom]: #6564
+#6788 := [unit-resolution #6565 #6786]: #6563
+#6579 := (or #6517 #6534 #6576)
+#6582 := (or #4068 #6517 #6534 #6576)
+#6568 := (+ #6511 #2517)
+#6569 := (+ #185 #6568)
+#6570 := (= #6569 0::int)
+#6513 := (+ #1235 #6512)
+#6514 := (+ #6483 #6513)
+#6515 := (<= #6514 0::int)
+#6571 := (or #6517 #6515 #6570)
+#6583 := (or #4068 #6571)
+#6590 := (iff #6583 #6582)
+#6585 := (or #4068 #6579)
+#6588 := (iff #6585 #6582)
+#6589 := [rewrite]: #6588
+#6586 := (iff #6583 #6585)
+#6580 := (iff #6571 #6579)
+#6577 := (iff #6570 #6576)
+#6574 := (= #6569 #6573)
+#6575 := [rewrite]: #6574
+#6578 := [monotonicity #6575]: #6577
+#6537 := (iff #6515 #6534)
+#6525 := (+ #6483 #6512)
+#6526 := (+ #1235 #6525)
+#6529 := (<= #6526 0::int)
+#6535 := (iff #6529 #6534)
+#6536 := [rewrite]: #6535
+#6530 := (iff #6515 #6529)
+#6527 := (= #6514 #6526)
+#6528 := [rewrite]: #6527
+#6531 := [monotonicity #6528]: #6530
+#6538 := [trans #6531 #6536]: #6537
+#6581 := [monotonicity #6538 #6578]: #6580
+#6587 := [monotonicity #6581]: #6586
+#6591 := [trans #6587 #6589]: #6590
+#6584 := [quant-inst]: #6583
+#6592 := [mp #6584 #6591]: #6582
+#6789 := [unit-resolution #6592 #6220]: #6579
+#6790 := [unit-resolution #6789 #6788 #6787 #6785]: false
+#6792 := [lemma #6790]: #6791
+#7320 := [unit-resolution #6792 #7321]: #6539
+#6545 := (or #6522 #6542)
+#6548 := (or #4076 #6522 #6542)
+#6518 := (or #6517 #6515)
+#6519 := (not #6518)
+#6520 := (= #6483 #1980)
+#6521 := (or #6520 #6519)
+#6549 := (or #4076 #6521)
+#6556 := (iff #6549 #6548)
+#6551 := (or #4076 #6545)
+#6554 := (iff #6551 #6548)
+#6555 := [rewrite]: #6554
+#6552 := (iff #6549 #6551)
+#6546 := (iff #6521 #6545)
+#6543 := (iff #6519 #6542)
+#6540 := (iff #6518 #6539)
+#6541 := [monotonicity #6538]: #6540
+#6544 := [monotonicity #6541]: #6543
+#6523 := (iff #6520 #6522)
+#6524 := [rewrite]: #6523
+#6547 := [monotonicity #6524 #6544]: #6546
+#6553 := [monotonicity #6547]: #6552
+#6557 := [trans #6553 #6555]: #6556
+#6550 := [quant-inst]: #6549
+#6558 := [mp #6550 #6557]: #6548
+#7322 := [unit-resolution #6558 #6216]: #6545
+#7325 := [unit-resolution #7322 #7320]: #6522
+#7123 := (not #6522)
+#7124 := (or #7123 #6559)
+#7167 := [th-lemma]: #7124
+#7341 := [unit-resolution #7167 #7325]: #6559
+#7342 := [th-lemma #7341 #7039 #7038 #7034 #7098]: false
+#7345 := [lemma #7342]: #7344
+#7501 := [unit-resolution #7345 #7500 #7499 #7495]: #7041
+#7021 := (or #6356 #6961)
+#6969 := [hypothesis]: #7041
+#7017 := [hypothesis]: #6328
+#6874 := (or #3963 #183 #6356 #6961)
+#6962 := (or #183 #6356 #6961)
+#6875 := (or #3963 #6962)
+#6877 := (iff #6875 #6874)
+#6878 := [rewrite]: #6877
+#6876 := [quant-inst]: #6875
+#6893 := [mp #6876 #6878]: #6874
+#7020 := [unit-resolution #6893 #7019 #7018 #7017 #6969]: false
+#7016 := [lemma #7020]: #7021
+#7502 := [unit-resolution #7016 #7501]: #6356
+#6362 := (not #6339)
+#6363 := (or #6362 #6328 #6333)
+#6364 := [def-axiom]: #6363
+#7503 := [unit-resolution #6364 #7502 #7454]: #6333
+#6359 := (not #6333)
+#7337 := (or #6359 #6819)
+#6999 := [hypothesis]: #6333
+#7334 := [symm #6999]: #6329
+#7335 := [monotonicity #7334]: #6819
+#6870 := (not #6819)
+#7333 := [hypothesis]: #6870
+#7336 := [unit-resolution #7333 #7335]: false
+#7338 := [lemma #7336]: #7337
+#7504 := [unit-resolution #7338 #7503]: #6819
+#6898 := (or #6870 #6821)
+#6899 := [th-lemma]: #6898
+#7505 := [unit-resolution #6899 #7504]: #6821
+#7488 := (or #7348 #2519 #6359 #6266)
+#7463 := [unit-resolution #7338 #6999]: #6819
+#7464 := [unit-resolution #6899 #7463]: #6821
+#7380 := (uf_1 ?x67!17 ?x68!16)
+#7381 := (uf_10 #7380)
+#7385 := (* -1::int #7381)
+#7457 := (+ #6511 #7385)
+#7462 := (>= #7457 0::int)
+#7456 := (= #6511 #7381)
+#7467 := (= #7381 #6511)
+#7465 := (= #7380 #6510)
+#7466 := [monotonicity #7334]: #7465
+#7468 := [monotonicity #7466]: #7467
+#7469 := [symm #7468]: #7456
+#7470 := (not #7456)
+#7471 := (or #7470 #7462)
+#7472 := [th-lemma]: #7471
+#7473 := [unit-resolution #7472 #7469]: #7462
+#7424 := (<= #7381 0::int)
+#7425 := (not #7424)
+#7428 := (= ?x68!16 ?x67!17)
+#7481 := (not #7428)
+#7482 := (iff #6292 #7481)
+#7479 := (iff #6266 #7428)
+#7477 := (iff #7428 #6266)
+#7476 := [commutativity]: #6267
+#7474 := (iff #7428 #6260)
+#7475 := [monotonicity #7334]: #7474
+#7478 := [trans #7475 #7476]: #7477
+#7480 := [symm #7478]: #7479
+#7483 := [monotonicity #7480]: #7482
+#7484 := [mp #7173 #7483]: #7481
+#7434 := (or #7425 #7428)
+#7439 := (or #7324 #7425 #7428)
+#7426 := (= ?x67!17 ?x68!16)
+#7427 := (or #7426 #7425)
+#7440 := (or #7324 #7427)
+#7447 := (iff #7440 #7439)
+#7442 := (or #7324 #7434)
+#7445 := (iff #7442 #7439)
+#7446 := [rewrite]: #7445
+#7443 := (iff #7440 #7442)
+#7437 := (iff #7427 #7434)
+#7431 := (or #7428 #7425)
+#7435 := (iff #7431 #7434)
+#7436 := [rewrite]: #7435
+#7432 := (iff #7427 #7431)
+#7429 := (iff #7426 #7428)
+#7430 := [rewrite]: #7429
+#7433 := [monotonicity #7430]: #7432
+#7438 := [trans #7433 #7436]: #7437
+#7444 := [monotonicity #7438]: #7443
+#7448 := [trans #7444 #7446]: #7447
+#7441 := [quant-inst]: #7440
+#7449 := [mp #7441 #7448]: #7439
+#7485 := [unit-resolution #7449 #3878]: #7434
+#7486 := [unit-resolution #7485 #7484]: #7425
+#7487 := [th-lemma #6900 #7486 #7473 #7346 #7034 #7464]: false
+#7489 := [lemma #7487]: #7488
+#7506 := [unit-resolution #7489 #7500 #7503 #7495]: #7348
+#7507 := [unit-resolution #7309 #7506]: #6784
+#7508 := [unit-resolution #6792 #7507]: #6539
+#7509 := [unit-resolution #7322 #7508]: #6522
+#7510 := [unit-resolution #7167 #7509]: #6559
+#7511 := [th-lemma #6900 #7499 #7510 #7500 #7505]: false
+#7512 := [lemma #7511]: #3075
+#3741 := (or #4151 #3080 #4145)
+#3733 := [def-axiom]: #3741
+#29391 := [unit-resolution #3733 #7512]: #29390
+#29392 := [unit-resolution #29391 #29389]: #4145
+#3655 := (or #4142 #4136)
+#3656 := [def-axiom]: #3655
+#30174 := [unit-resolution #3656 #29392]: #4136
+#3652 := (or #4139 #3126 #4133)
+#3653 := [def-axiom]: #3652
+#30181 := [unit-resolution #3653 #30174]: #4136
+#30177 := [unit-resolution #30181 #25366]: #3126
+#3767 := (or #3121 #2014)
+#3761 := [def-axiom]: #3767
+#30178 := [unit-resolution #3761 #30177]: #2014
+#19559 := (+ #2008 #18070)
+#21560 := (>= #19559 0::int)
+#21561 := (= #2008 #18066)
+#18705 := (= #2007 #18065)
+#11539 := (= ?x71!19 #11538)
+#16935 := (or #5007 #11539)
+#16865 := [quant-inst]: #16935
+#22732 := [unit-resolution #16865 #3831]: #11539
+#22885 := (= ?x72!18 uf_22)
+#10926 := (= ?x72!18 #6904)
+#10925 := (up_6 uf_15 ?x72!18)
+#10950 := (not #10925)
+#11332 := (uf_4 uf_14 ?x71!19)
+#11368 := (* -1::int #11332)
+#10897 := (uf_4 uf_14 ?x72!18)
+#11464 := (+ #10897 #11368)
+#11465 := (+ #2008 #11464)
+#11466 := (>= #11465 0::int)
+#20486 := (not #11466)
+#3599 := (not #2546)
+#3600 := (or #3121 #3599)
+#3753 := [def-axiom]: #3600
+#30156 := [unit-resolution #3753 #30177]: #3599
+#11082 := (* -1::int #10897)
+#11083 := (+ #2005 #11082)
+#22657 := (>= #11083 0::int)
+#10908 := (= #2005 #10897)
+#3768 := (or #3121 #2015)
+#3597 := [def-axiom]: #3768
+#30163 := [unit-resolution #3597 #30177]: #2015
+#10911 := (or #3106 #10908)
+#19109 := (not #10908)
+#26554 := [hypothesis]: #19109
+#22817 := [hypothesis]: #2015
+#10589 := (or #4093 #3106 #10908)
+#10898 := (= #10897 #2005)
+#10907 := (or #3106 #10898)
+#10544 := (or #4093 #10907)
+#10543 := (iff #10544 #10589)
+#10433 := (or #4093 #10911)
+#10439 := (iff #10433 #10589)
+#10639 := [rewrite]: #10439
+#10592 := (iff #10544 #10433)
+#10912 := (iff #10907 #10911)
+#10909 := (iff #10898 #10908)
+#10910 := [rewrite]: #10909
+#10913 := [monotonicity #10910]: #10912
+#10593 := [monotonicity #10913]: #10592
+#10545 := [trans #10593 #10639]: #10543
+#10590 := [quant-inst]: #10544
+#10836 := [mp #10590 #10545]: #10589
+#24736 := [unit-resolution #10836 #13487 #22817 #26554]: false
+#24651 := [lemma #24736]: #10911
+#30132 := [unit-resolution #24651 #30163]: #10908
+#18810 := (or #19109 #22657)
+#19214 := [th-lemma]: #18810
+#30138 := [unit-resolution #19214 #30132]: #22657
+#22794 := (not #22657)
+#30135 := (or #20486 #22794 #2546)
+#11369 := (+ #2003 #11368)
+#11370 := (<= #11369 0::int)
+#20553 := (not #11370)
+#25610 := [hypothesis]: #20553
+#15168 := (or #4084 #11370)
+#11360 := (+ #11332 #2004)
+#11361 := (>= #11360 0::int)
+#15207 := (or #4084 #11361)
+#15251 := (iff #15207 #15168)
+#15083 := (iff #15168 #15168)
+#15254 := [rewrite]: #15083
+#11373 := (iff #11361 #11370)
+#11362 := (+ #2004 #11332)
+#11365 := (>= #11362 0::int)
+#11371 := (iff #11365 #11370)
+#11372 := [rewrite]: #11371
+#11366 := (iff #11361 #11365)
+#11363 := (= #11360 #11362)
+#11364 := [rewrite]: #11363
+#11367 := [monotonicity #11364]: #11366
+#11374 := [trans #11367 #11372]: #11373
+#15205 := [monotonicity #11374]: #15251
+#15289 := [trans #15205 #15254]: #15251
+#15255 := [quant-inst]: #15207
+#15163 := [mp #15255 #15289]: #15168
+#25611 := [unit-resolution #15163 #7037 #25610]: false
+#25612 := [lemma #25611]: #11370
+#20554 := (or #20486 #22794 #20553 #2546)
+#20520 := [th-lemma]: #20554
+#30137 := [unit-resolution #20520 #25612]: #30135
+#30131 := [unit-resolution #30137 #30138 #30156]: #20486
+#18604 := [unit-resolution #3662 #6214]: #3949
+#16306 := (or #3954 #2013 #10950 #11466)
+#11467 := (or #10950 #2013 #11466)
+#16314 := (or #3954 #11467)
+#16528 := (iff #16314 #16306)
+#11468 := (or #2013 #10950 #11466)
+#16531 := (or #3954 #11468)
+#16527 := (iff #16531 #16306)
+#16599 := [rewrite]: #16527
+#16302 := (iff #16314 #16531)
+#11469 := (iff #11467 #11468)
+#11470 := [rewrite]: #11469
+#16598 := [monotonicity #11470]: #16302
+#16273 := [trans #16598 #16599]: #16528
+#16530 := [quant-inst]: #16314
+#16623 := [mp #16530 #16273]: #16306
+#30186 := [unit-resolution #16623 #18604 #30178 #30131]: #10950
+#10933 := (or #10925 #10926)
+#10928 := (up_6 #8048 ?x72!18)
+#10938 := (iff #10928 #10933)
+#10919 := (or #6010 #10938)
+#10927 := (ite #10926 #5970 #10925)
+#10929 := (iff #10928 #10927)
+#9877 := (or #6010 #10929)
+#10920 := (iff #9877 #10919)
+#9876 := (iff #10919 #10919)
+#10914 := [rewrite]: #9876
+#10939 := (iff #10929 #10938)
+#10936 := (iff #10927 #10933)
+#10930 := (ite #10926 true #10925)
+#10934 := (iff #10930 #10933)
+#10935 := [rewrite]: #10934
+#10931 := (iff #10927 #10930)
+#10932 := [monotonicity #5976]: #10931
+#10937 := [trans #10932 #10935]: #10936
+#10940 := [monotonicity #10937]: #10939
+#10127 := [monotonicity #10940]: #10920
+#10900 := [trans #10127 #10914]: #10920
+#10921 := [quant-inst]: #9877
+#9879 := [mp #10921 #10900]: #10919
+#20934 := [unit-resolution #9879 #3851]: #10938
+#18735 := (iff #2015 #10928)
+#20942 := (iff #10928 #2015)
+#18734 := [monotonicity #8592]: #20942
+#22886 := [symm #18734]: #18735
+#30136 := [mp #30163 #22886]: #10928
+#11095 := (not #10928)
+#10096 := (not #10938)
+#11478 := (or #10096 #11095 #10933)
+#11383 := [def-axiom]: #11478
+#30170 := [unit-resolution #11383 #30136 #20934]: #10933
+#10013 := (not #10933)
+#11234 := (or #10013 #10925 #10926)
+#11235 := [def-axiom]: #11234
+#30187 := [unit-resolution #11235 #30170 #30186]: #10926
+#30204 := [trans #30187 #9758]: #22885
+#30209 := [monotonicity #30204 #22732]: #18705
+#30203 := [monotonicity #30209]: #21561
+#21225 := (not #21561)
+#21229 := (or #21225 #21560)
+#20945 := [th-lemma]: #21229
+#30210 := [unit-resolution #20945 #30203]: #21560
+#22878 := (not #21560)
+#22827 := (or #17499 #22878 #2013)
+#22811 := [hypothesis]: #2014
+#22788 := [hypothesis]: #18075
+#22810 := [hypothesis]: #21560
+#22879 := [th-lemma #22810 #22788 #22811]: false
+#22888 := [lemma #22879]: #22827
+#30169 := [unit-resolution #22888 #30210 #30178]: #17499
+#11646 := (uf_4 uf_14 #11538)
+#11659 := (* -1::int #11646)
+#18084 := (+ #11659 #18066)
+#18085 := (+ #185 #18084)
+#18086 := (>= #18085 0::int)
+#17430 := (not #18086)
+#11384 := (uf_1 uf_22 ?x71!19)
+#11385 := (uf_10 #11384)
+#11402 := (+ #11368 #11385)
+#11403 := (+ #185 #11402)
+#11404 := (>= #11403 0::int)
+#15413 := (not #11404)
+#11547 := (uf_2 #2007)
+#11552 := (uf_4 uf_14 #11547)
+#11565 := (* -1::int #11552)
+#11566 := (+ #185 #11565)
+#11567 := (<= #11566 0::int)
+#11557 := (up_6 uf_15 #11547)
+#28137 := (not #11557)
+#30214 := (iff #10950 #28137)
+#30212 := (iff #10925 #11557)
+#30211 := (iff #11557 #10925)
+#28262 := (= #11547 ?x72!18)
+#11548 := (= ?x72!18 #11547)
+#16957 := (or #8156 #11548)
+#16964 := [quant-inst]: #16957
+#22805 := [unit-resolution #16964 #3837]: #11548
+#30153 := [symm #22805]: #28262
+#30171 := [monotonicity #30153]: #30211
+#30134 := [symm #30171]: #30212
+#30240 := [monotonicity #30134]: #30214
+#30237 := [mp #30186 #30240]: #28137
+#11572 := (or #11557 #11567)
+#26052 := (or #4058 #11557 #11567)
+#11555 := (+ #11552 #1235)
+#11556 := (>= #11555 0::int)
+#11558 := (or #11557 #11556)
+#24959 := (or #4058 #11558)
+#26166 := (iff #24959 #26052)
+#24944 := (or #4058 #11572)
+#26155 := (iff #24944 #26052)
+#24069 := [rewrite]: #26155
+#24260 := (iff #24959 #24944)
+#11573 := (iff #11558 #11572)
+#11570 := (iff #11556 #11567)
+#11559 := (+ #1235 #11552)
+#11562 := (>= #11559 0::int)
+#11568 := (iff #11562 #11567)
+#11569 := [rewrite]: #11568
+#11563 := (iff #11556 #11562)
+#11560 := (= #11555 #11559)
+#11561 := [rewrite]: #11560
+#11564 := [monotonicity #11561]: #11563
+#11571 := [trans #11564 #11569]: #11570
+#11574 := [monotonicity #11571]: #11573
+#24990 := [monotonicity #11574]: #24260
+#24992 := [trans #24990 #24069]: #26166
+#24970 := [quant-inst]: #24959
+#26321 := [mp #24970 #24992]: #26052
+#30243 := [unit-resolution #26321 #7497]: #11572
+#30133 := [unit-resolution #30243 #30237]: #11567
+#22774 := (not #11567)
+#22822 := (or #15413 #2546 #22878 #22774 #22794)
+#22808 := [hypothesis]: #11404
+#22807 := [hypothesis]: #22657
+#17228 := (+ #10897 #11565)
+#17121 := (>= #17228 0::int)
+#17229 := (= #10897 #11552)
+#22789 := [monotonicity #22805]: #17229
+#22790 := (not #17229)
+#22787 := (or #22790 #17121)
+#22771 := [th-lemma]: #22787
+#22781 := [unit-resolution #22771 #22789]: #17121
+#22802 := [hypothesis]: #11567
+#22779 := [hypothesis]: #3599
+#22829 := [unit-resolution #15163 #7037]: #11370
+#18268 := (+ #11385 #18070)
+#18269 := (<= #18268 0::int)
+#17117 := (= #11385 #18066)
+#22830 := (= #11384 #18065)
+#22818 := [monotonicity #22732]: #22830
+#22816 := [monotonicity #22818]: #17117
+#22804 := (not #17117)
+#22797 := (or #22804 #18269)
+#22821 := [th-lemma]: #22797
+#22813 := [unit-resolution #22821 #22816]: #18269
+#22819 := [th-lemma #22813 #22829 #22779 #22810 #22802 #22781 #22807 #22808]: false
+#18715 := [lemma #22819]: #22822
+#30242 := [unit-resolution #18715 #30210 #30138 #30133 #30156]: #15413
+#30247 := (or #17430 #11404)
+#18270 := (>= #18268 0::int)
+#22014 := (uf_1 #6904 ?x72!18)
+#23089 := (uf_3 #22014)
+#25777 := (uf_1 #23089 ?x71!19)
+#25872 := (uf_3 #25777)
+#27572 := (= #25872 #11538)
+#26806 := (= #11538 #25872)
+#26804 := (= #2007 #25777)
+#23115 := (= ?x72!18 #23089)
+#17768 := (or #5007 #23115)
+#17877 := [quant-inst]: #17768
+#26799 := [unit-resolution #17877 #3831]: #23115
+#27568 := [monotonicity #26799]: #26804
+#27571 := [monotonicity #27568]: #26806
+#27573 := [symm #27571]: #27572
+#25873 := (= ?x71!19 #25872)
+#26816 := (not #25873)
+#27263 := [hypothesis]: #26816
+#27265 := (or #5007 #25873)
+#27266 := [quant-inst]: #27265
+#27264 := [unit-resolution #27266 #3831 #27263]: false
+#27269 := [lemma #27264]: #25873
+#27574 := [trans #27269 #27573]: #11539
+#28371 := [monotonicity #27574]: #22830
+#28372 := [monotonicity #28371]: #17117
+#28370 := [hypothesis]: #22804
+#28373 := [unit-resolution #28370 #28372]: false
+#28374 := [lemma #28373]: #17117
+#20936 := (or #22804 #18270)
+#20938 := [th-lemma]: #20936
+#30213 := [unit-resolution #20938 #28374]: #18270
+#21441 := (not #18270)
+#30244 := (or #17430 #11404 #21441)
+#18023 := (+ #11332 #11659)
+#18024 := (<= #18023 0::int)
+#17225 := (= #11332 #11646)
+#26817 := (or #26816 #17225)
+#26812 := (= #11646 #11332)
+#26810 := (= #11538 ?x71!19)
+#26808 := (= #25872 ?x71!19)
+#26794 := [hypothesis]: #25873
+#26809 := [symm #26794]: #26808
+#26802 := (= #25777 #2007)
+#26800 := (= #23089 ?x72!18)
+#26801 := [symm #26799]: #26800
+#26803 := [monotonicity #26801]: #26802
+#26805 := [symm #26803]: #26804
+#26807 := [monotonicity #26805]: #26806
+#26811 := [trans #26807 #26809]: #26810
+#26813 := [monotonicity #26811]: #26812
+#26814 := [symm #26813]: #17225
+#21448 := (not #17225)
+#26793 := [hypothesis]: #21448
+#26815 := [unit-resolution #26793 #26814]: false
+#26818 := [lemma #26815]: #26817
+#30239 := [unit-resolution #26818 #27269]: #17225
+#21454 := (or #21448 #18024)
+#20933 := [th-lemma]: #21454
+#30238 := [unit-resolution #20933 #30239]: #18024
+#22247 := (not #18024)
+#22226 := (or #17430 #22247 #11404 #21441)
+#22243 := [th-lemma]: #22226
+#30246 := [unit-resolution #22243 #30238]: #30244
+#30265 := [unit-resolution #30246 #30213]: #30247
+#30266 := [unit-resolution #30265 #30242]: #17430
+#25738 := (or #18075 #18086 #18113)
+#26503 := [hypothesis]: #17430
+#26498 := [hypothesis]: #17499
+#18095 := (= #18092 0::int)
+#22231 := (not #18095)
+#25748 := (not #18113)
+#25749 := [hypothesis]: #25748
+#22367 := (or #22231 #18113)
+#22447 := [th-lemma]: #22367
+#25750 := [unit-resolution #22447 #25749]: #22231
+#26505 := (or #18095 #18086 #18075)
+#26497 := [hypothesis]: #22231
+#17115 := (or #4068 #18075 #18086 #18095)
+#18067 := (+ #18066 #18049)
+#18068 := (+ #185 #18067)
+#18069 := (= #18068 0::int)
+#18071 := (+ #1235 #18070)
+#18072 := (+ #11646 #18071)
+#18073 := (<= #18072 0::int)
+#18076 := (or #18075 #18073 #18069)
+#17396 := (or #4068 #18076)
+#17345 := (iff #17396 #17115)
+#18098 := (or #18075 #18086 #18095)
+#17361 := (or #4068 #18098)
+#17360 := (iff #17361 #17115)
+#17359 := [rewrite]: #17360
+#17363 := (iff #17396 #17361)
+#18099 := (iff #18076 #18098)
+#18096 := (iff #18069 #18095)
+#18093 := (= #18068 #18092)
+#18094 := [rewrite]: #18093
+#18097 := [monotonicity #18094]: #18096
+#18089 := (iff #18073 #18086)
+#18077 := (+ #11646 #18070)
+#18078 := (+ #1235 #18077)
+#18081 := (<= #18078 0::int)
+#18087 := (iff #18081 #18086)
+#18088 := [rewrite]: #18087
+#18082 := (iff #18073 #18081)
+#18079 := (= #18072 #18078)
+#18080 := [rewrite]: #18079
+#18083 := [monotonicity #18080]: #18082
+#18090 := [trans #18083 #18088]: #18089
+#18100 := [monotonicity #18090 #18097]: #18099
+#17346 := [monotonicity #18100]: #17363
+#17364 := [trans #17346 #17359]: #17345
+#17381 := [quant-inst]: #17396
+#17317 := [mp #17381 #17364]: #17115
+#26504 := [unit-resolution #17317 #6220 #26503 #26498 #26497]: false
+#26506 := [lemma #26504]: #26505
+#25751 := [unit-resolution #26506 #25750 #26498 #26503]: false
+#25739 := [lemma #25751]: #25738
+#30267 := [unit-resolution #25739 #30266 #30169]: #18113
+#25874 := (uf_2 #25777)
+#25875 := (= #23089 #25874)
+#28275 := (not #25875)
+#28360 := [hypothesis]: #28275
+#28362 := (or #8156 #25875)
+#28363 := [quant-inst]: #28362
+#28361 := [unit-resolution #28363 #3837 #28360]: false
+#28364 := [lemma #28361]: #25875
+#28276 := (or #28275 #17229)
+#28264 := (= #11552 #10897)
+#28260 := (= #11547 #23089)
+#28258 := (= #25874 #23089)
+#28251 := [hypothesis]: #25875
+#28259 := [symm #28251]: #28258
+#28256 := (= #11547 #25874)
+#28257 := [monotonicity #27568]: #28256
+#28261 := [trans #28257 #28259]: #28260
+#28263 := [trans #28261 #26801]: #28262
+#28265 := [monotonicity #28263]: #28264
+#28273 := [symm #28265]: #17229
+#28210 := [hypothesis]: #22790
+#28274 := [unit-resolution #28210 #28273]: false
+#28277 := [lemma #28274]: #28276
+#26555 := [unit-resolution #28277 #28364]: #17229
+#30245 := [unit-resolution #22771 #26555]: #17121
+#18258 := (+ #2003 #18049)
+#18261 := (<= #18258 0::int)
+#17226 := (= #2003 #18030)
+#27788 := [monotonicity #27574]: #17226
+#21452 := (not #17226)
+#27567 := [hypothesis]: #21452
+#27789 := [unit-resolution #27567 #27788]: false
+#27872 := [lemma #27789]: #17226
+#22901 := (or #21452 #18261)
+#22182 := [th-lemma]: #22901
+#30269 := [unit-resolution #22182 #27872]: #18261
+[th-lemma #30269 #30245 #30138 #30133 #30156 #30210 #30267]: false
unsat
--- a/src/HOL/Boogie/Examples/cert/Boogie_max Fri Dec 11 15:06:12 2009 +0100
+++ b/src/HOL/Boogie/Examples/cert/Boogie_max Fri Dec 11 15:35:29 2009 +0100
@@ -14,6 +14,6 @@
(uf_3 Int Int)
(uf_1 Int)
)
-:assumption (not (implies true (implies (< 0 uf_1) (implies true (implies (= uf_2 (uf_3 0)) (implies (and (<= 0 0) (and (<= 0 0) (and (<= 1 1) (<= 1 1)))) (and (forall (?x1 Int) (implies (and (<= 0 ?x1) (< ?x1 1)) (<= (uf_3 ?x1) uf_2))) (implies (forall (?x2 Int) (implies (and (<= 0 ?x2) (< ?x2 1)) (<= (uf_3 ?x2) uf_2))) (and (= (uf_3 0) uf_2) (implies (= (uf_3 0) uf_2) (implies true (implies (and (<= 0 uf_4) (<= 1 uf_5)) (implies (forall (?x3 Int) (implies (and (<= 0 ?x3) (< ?x3 uf_5)) (<= (uf_3 ?x3) uf_6))) (implies (= (uf_3 uf_4) uf_6) (implies (and (<= 0 uf_4) (<= 1 uf_5)) (and (implies true (implies (and (<= 0 uf_4) (<= 1 uf_5)) (implies (<= uf_1 uf_5) (implies (and (<= 0 uf_4) (<= 1 uf_5)) (implies true (implies (= uf_7 uf_4) (implies (= uf_8 uf_6) (implies (= uf_9 uf_5) (implies true (and (exists (?x4 Int) (implies (and (<= 0 ?x4) (< ?x4 uf_1)) (= (uf_3 ?x4) uf_8))) (implies (exists (?x5 Int) (implies (and (<= 0 ?x5) (< ?x5 uf_1)) (= (uf_3 ?x5) uf_8))) (and (forall (?x6 Int) (implies (and (<= 0 ?x6) (< ?x6 uf_1)) (<= (uf_3 ?x6) uf_8))) (implies (forall (?x7 Int) (implies (and (<= 0 ?x7) (< ?x7 uf_1)) (<= (uf_3 ?x7) uf_8))) true))))))))))))) (implies true (implies (and (<= 0 uf_4) (<= 1 uf_5)) (implies (< uf_5 uf_1) (implies (and (<= 0 uf_4) (<= 1 uf_5)) (and (implies true (implies (and (<= 0 uf_4) (<= 1 uf_5)) (implies (< uf_6 (uf_3 uf_5)) (implies (= uf_10 (uf_3 uf_5)) (implies (and (<= 1 uf_5) (<= 1 uf_5)) (implies true (implies (= uf_11 uf_5) (implies (= uf_12 uf_10) (implies true (implies (and (<= 0 uf_11) (<= 1 uf_5)) (implies (= uf_13 (+ uf_5 1)) (implies (and (<= 0 uf_11) (<= 2 uf_13)) (implies true (and (forall (?x8 Int) (implies (and (<= 0 ?x8) (< ?x8 uf_13)) (<= (uf_3 ?x8) uf_12))) (implies (forall (?x9 Int) (implies (and (<= 0 ?x9) (< ?x9 uf_13)) (<= (uf_3 ?x9) uf_12))) (and (= (uf_3 uf_11) uf_12) (implies (= (uf_3 uf_11) uf_12) (implies false true)))))))))))))))))) (implies true (implies (and (<= 0 uf_4) (<= 1 uf_5)) (implies (<= (uf_3 uf_5) uf_6) (implies (and (<= 0 uf_4) (<= 1 uf_5)) (implies true (implies (= uf_11 uf_4) (implies (= uf_12 uf_6) (implies true (implies (and (<= 0 uf_11) (<= 1 uf_5)) (implies (= uf_13 (+ uf_5 1)) (implies (and (<= 0 uf_11) (<= 2 uf_13)) (implies true (and (forall (?x10 Int) (implies (and (<= 0 ?x10) (< ?x10 uf_13)) (<= (uf_3 ?x10) uf_12))) (implies (forall (?x11 Int) (implies (and (<= 0 ?x11) (< ?x11 uf_13)) (<= (uf_3 ?x11) uf_12))) (and (= (uf_3 uf_11) uf_12) (implies (= (uf_3 uf_11) uf_12) (implies false true))))))))))))))))))))))))))))))))))))))
+:assumption (not (implies (and true (and (< 0 uf_1) (and (= uf_2 (uf_3 0)) (and (<= 0 0) (and (<= 0 0) (and (<= 1 1) (<= 1 1))))))) (and (forall (?x1 Int) (implies (and (<= 0 ?x1) (< ?x1 1)) (<= (uf_3 ?x1) uf_2))) (implies (forall (?x2 Int) (implies (and (<= 0 ?x2) (< ?x2 1)) (<= (uf_3 ?x2) uf_2))) (and (= (uf_3 0) uf_2) (implies (and (= (uf_3 0) uf_2) (and true (and (and (<= 0 uf_4) (<= 1 uf_5)) (and (forall (?x3 Int) (implies (and (<= 0 ?x3) (< ?x3 uf_5)) (<= (uf_3 ?x3) uf_6))) (and (= (uf_3 uf_4) uf_6) (and (<= 0 uf_4) (<= 1 uf_5))))))) (and (implies (and true (and (and (<= 0 uf_4) (<= 1 uf_5)) (and (<= uf_1 uf_5) (and (and (<= 0 uf_4) (<= 1 uf_5)) (and (= uf_7 uf_4) (and (= uf_8 uf_6) (= uf_9 uf_5))))))) (and (exists (?x4 Int) (implies (and (<= 0 ?x4) (< ?x4 uf_1)) (= (uf_3 ?x4) uf_8))) (implies (exists (?x5 Int) (implies (and (<= 0 ?x5) (< ?x5 uf_1)) (= (uf_3 ?x5) uf_8))) (forall (?x6 Int) (implies (and (<= 0 ?x6) (< ?x6 uf_1)) (<= (uf_3 ?x6) uf_8)))))) (implies (and true (and (and (<= 0 uf_4) (<= 1 uf_5)) (and (< uf_5 uf_1) (and (<= 0 uf_4) (<= 1 uf_5))))) (and (implies (and true (and (and (<= 0 uf_4) (<= 1 uf_5)) (and (< uf_6 (uf_3 uf_5)) (and (= uf_10 (uf_3 uf_5)) (and (and (<= 1 uf_5) (<= 1 uf_5)) (and (= uf_11 uf_5) (and (= uf_12 uf_10) (and true (and (and (<= 0 uf_11) (<= 1 uf_5)) (and (= uf_13 (+ uf_5 1)) (and (<= 0 uf_11) (<= 2 uf_13)))))))))))) (and (forall (?x7 Int) (implies (and (<= 0 ?x7) (< ?x7 uf_13)) (<= (uf_3 ?x7) uf_12))) (implies (forall (?x8 Int) (implies (and (<= 0 ?x8) (< ?x8 uf_13)) (<= (uf_3 ?x8) uf_12))) (= (uf_3 uf_11) uf_12)))) (implies (and true (and (and (<= 0 uf_4) (<= 1 uf_5)) (and (<= (uf_3 uf_5) uf_6) (and (and (<= 0 uf_4) (<= 1 uf_5)) (and (= uf_11 uf_4) (and (= uf_12 uf_6) (and true (and (and (<= 0 uf_11) (<= 1 uf_5)) (and (= uf_13 (+ uf_5 1)) (and (<= 0 uf_11) (<= 2 uf_13))))))))))) (and (forall (?x9 Int) (implies (and (<= 0 ?x9) (< ?x9 uf_13)) (<= (uf_3 ?x9) uf_12))) (implies (forall (?x10 Int) (implies (and (<= 0 ?x10) (< ?x10 uf_13)) (<= (uf_3 ?x10) uf_12))) (= (uf_3 uf_11) uf_12)))))))))))))
:formula true
)
--- a/src/HOL/Boogie/Examples/cert/Boogie_max.proof Fri Dec 11 15:06:12 2009 +0100
+++ b/src/HOL/Boogie/Examples/cert/Boogie_max.proof Fri Dec 11 15:35:29 2009 +0100
@@ -1,2291 +1,1995 @@
#2 := false
-#4 := 0::int
decl uf_3 :: (-> int int)
-#8 := (uf_3 0::int)
-#714 := -1::int
-#2117 := (* -1::int #8)
+decl ?x1!0 :: int
+#870 := ?x1!0
+#871 := (uf_3 ?x1!0)
decl uf_2 :: int
#7 := uf_2
-#2122 := (+ uf_2 #2117)
-#2118 := (>= #2122 0::int)
-#9 := (= uf_2 #8)
+#1833 := (= uf_2 #871)
+#4 := 0::int
+#8 := (uf_3 0::int)
+#1842 := (= #8 #871)
+#1841 := (= #871 #8)
+#1860 := (= ?x1!0 0::int)
+#11 := 1::int
+#873 := (>= ?x1!0 1::int)
+#874 := (not #873)
+#875 := (>= ?x1!0 0::int)
+#913 := (not #875)
+#516 := -1::int
+#827 := (* -1::int #871)
+#828 := (+ uf_2 #827)
+#872 := (>= #828 0::int)
+#1208 := (or #872 #873 #913)
+#1213 := (not #1208)
+decl ?x4!1 :: int
+#900 := ?x4!1
+#908 := (uf_3 ?x4!1)
+decl uf_8 :: int
+#49 := uf_8
+#1055 := (= uf_8 #908)
+#905 := (>= ?x4!1 0::int)
+#1286 := (not #905)
+#901 := (* -1::int ?x4!1)
decl uf_1 :: int
#5 := uf_1
-#1032 := (<= uf_1 0::int)
-decl uf_6 :: int
-#32 := uf_6
-#989 := (* -1::int uf_6)
-#16 := (:var 0 int)
-#20 := (uf_3 #16)
-#990 := (+ #20 #989)
-#991 := (<= #990 0::int)
+#902 := (+ uf_1 #901)
+#903 := (<= #902 0::int)
+#1301 := (or #903 #1286 #1055)
+#1332 := (not #1301)
+decl ?x6!2 :: int
+#918 := ?x6!2
+#919 := (uf_3 ?x6!2)
+#1079 := (* -1::int #919)
+#1080 := (+ uf_8 #1079)
+#1081 := (>= #1080 0::int)
+#926 := (>= ?x6!2 0::int)
+#1306 := (not #926)
+#922 := (* -1::int ?x6!2)
+#923 := (+ uf_1 #922)
+#924 := (<= #923 0::int)
+#1474 := (or #924 #1306 #1081 #1332)
+#1477 := (not #1474)
+#19 := (:var 0 int)
+#23 := (uf_3 #19)
+#1986 := (pattern #23)
+#604 := (* -1::int #19)
+#605 := (+ uf_1 #604)
+#606 := (<= #605 0::int)
+#505 := (>= #19 0::int)
+#1216 := (not #505)
+#61 := (= #23 uf_8)
+#1275 := (or #61 #1216 #606)
+#1280 := (not #1275)
+#2047 := (forall (vars (?x4 int)) (:pat #1986) #1280)
+#2052 := (or #2047 #1477)
+#2055 := (not #2052)
decl uf_5 :: int
-#27 := uf_5
-#784 := (* -1::int uf_5)
-#979 := (+ #16 #784)
-#978 := (>= #979 0::int)
-#980 := (not #978)
-#703 := (>= #16 0::int)
-#983 := (and #703 #980)
-#986 := (not #983)
-#994 := (or #986 #991)
-#997 := (forall (vars (?x3 int)) #994)
-#1000 := (not #997)
-#67 := (uf_3 uf_5)
-#882 := (* -1::int #67)
-#883 := (+ uf_6 #882)
-#881 := (>= #883 0::int)
-#880 := (not #881)
-decl uf_11 :: int
-#72 := uf_11
-#816 := (>= uf_11 0::int)
-#11 := 1::int
-#733 := (>= uf_5 1::int)
-#871 := (and #733 #816)
-#874 := (not #871)
-decl uf_13 :: int
-#78 := uf_13
-#828 := (* -1::int uf_13)
-#865 := (+ uf_5 #828)
-#864 := (= #865 -1::int)
-#868 := (not #864)
-decl uf_12 :: int
-#74 := uf_12
-#839 := (* -1::int uf_12)
-#840 := (+ #20 #839)
-#841 := (<= #840 0::int)
-#829 := (+ #16 #828)
-#827 := (>= #829 0::int)
-#830 := (not #827)
-#833 := (and #703 #830)
-#836 := (not #833)
-#844 := (or #836 #841)
-#847 := (forall (vars (?x8 int)) #844)
-#850 := (not #847)
-#89 := (uf_3 uf_11)
-#332 := (= uf_12 #89)
-#856 := (or #332 #850)
-#861 := (and #847 #856)
-#81 := 2::int
-#819 := (>= uf_13 2::int)
-#821 := (and #816 #819)
-#824 := (not #821)
+#30 := uf_5
+#533 := (* -1::int uf_5)
+#583 := (+ uf_1 #533)
+#584 := (<= #583 0::int)
+#650 := (not #584)
+#557 := (>= uf_5 1::int)
+#1349 := (not #557)
decl uf_4 :: int
-#25 := uf_4
-#730 := (>= uf_4 0::int)
-#735 := (and #730 #733)
-#738 := (not #735)
-#474 := (= uf_6 uf_12)
-#480 := (not #474)
-#471 := (= uf_4 uf_11)
-#489 := (not #471)
-#944 := (or #489 #480 #738 #824 #861 #868 #874 #880)
-#877 := (not #733)
-decl uf_10 :: int
-#69 := uf_10
-#313 := (= uf_10 uf_12)
-#407 := (not #313)
-#310 := (= uf_5 uf_11)
-#416 := (not #310)
-#305 := (= #67 uf_10)
-#441 := (not #305)
-#920 := (or #441 #416 #407 #877 #738 #824 #861 #868 #874 #881)
-#949 := (and #920 #944)
-#785 := (+ uf_1 #784)
-#786 := (<= #785 0::int)
-#970 := (or #738 #786 #949)
-#789 := (not #786)
-decl uf_8 :: int
-#41 := uf_8
-#767 := (* -1::int uf_8)
-#768 := (+ #20 #767)
-#769 := (<= #768 0::int)
-#741 := (* -1::int #16)
-#742 := (+ uf_1 #741)
-#743 := (<= #742 0::int)
-#744 := (not #743)
-#747 := (and #703 #744)
-#750 := (not #747)
-#772 := (or #750 #769)
-#775 := (forall (vars (?x6 int)) #772)
-#47 := (= #20 uf_8)
-#756 := (or #47 #750)
-#761 := (exists (vars (?x4 int)) #756)
-#764 := (not #761)
-#778 := (or #764 #775)
-#781 := (and #761 #778)
+#28 := uf_4
+#555 := (>= uf_4 0::int)
+#1348 := (not #555)
decl uf_9 :: int
-#43 := uf_9
-#189 := (= uf_5 uf_9)
-#241 := (not #189)
-#186 := (= uf_6 uf_8)
-#250 := (not #186)
+#51 := uf_9
+#216 := (= uf_5 uf_9)
+#1347 := (not #216)
+decl uf_6 :: int
+#35 := uf_6
+#213 := (= uf_6 uf_8)
+#1346 := (not #213)
decl uf_7 :: int
-#39 := uf_7
-#183 := (= uf_4 uf_7)
-#259 := (not #183)
-#810 := (or #259 #250 #241 #738 #781 #789)
-#975 := (and #810 #970)
-#36 := (uf_3 uf_4)
-#180 := (= uf_6 #36)
-#583 := (not #180)
-#616 := (not #9)
-#1018 := (or #616 #583 #738 #975 #1000)
-#1023 := (and #9 #1018)
-#717 := (* -1::int #20)
-#718 := (+ uf_2 #717)
-#716 := (>= #718 0::int)
-#706 := (>= #16 1::int)
-#704 := (not #706)
-#708 := (and #703 #704)
-#711 := (not #708)
-#720 := (or #711 #716)
-#723 := (forall (vars (?x1 int)) #720)
-#726 := (not #723)
-#1026 := (or #726 #1023)
-#1029 := (and #723 #1026)
-#1052 := (or #616 #1029 #1032)
-#1057 := (not #1052)
+#47 := uf_7
+#210 := (= uf_4 uf_7)
+#1345 := (not #210)
+#2058 := (or #1345 #1346 #1347 #1348 #1349 #650 #2055)
+#2061 := (not #2058)
+decl ?x7!3 :: int
+#954 := ?x7!3
+#955 := (uf_3 ?x7!3)
+#1138 := (* -1::int #955)
+decl uf_12 :: int
+#81 := uf_12
+#1139 := (+ uf_12 #1138)
+#1140 := (>= #1139 0::int)
+#1116 := (* -1::int ?x7!3)
+decl uf_13 :: int
+#85 := uf_13
+#1117 := (+ uf_13 #1116)
+#1118 := (<= #1117 0::int)
+#961 := (>= ?x7!3 0::int)
+#1389 := (not #961)
+#1404 := (or #1389 #1118 #1140)
+#1409 := (not #1404)
+#733 := (* -1::int uf_12)
+#734 := (+ #23 #733)
+#735 := (<= #734 0::int)
+#674 := (* -1::int uf_13)
+#723 := (+ #19 #674)
+#722 := (>= #723 0::int)
+#1371 := (or #1216 #722 #735)
+#2003 := (forall (vars (?x7 int)) (:pat #1986) #1371)
+#2008 := (not #2003)
+decl uf_11 :: int
+#79 := uf_11
+#106 := (uf_3 uf_11)
+#358 := (= uf_12 #106)
+#2011 := (or #358 #2008)
+#2014 := (not #2011)
+#2017 := (or #2014 #1409)
+#2020 := (not #2017)
+#74 := (uf_3 uf_5)
+#679 := (* -1::int #74)
+#680 := (+ uf_6 #679)
+#678 := (>= #680 0::int)
+#681 := (not #678)
+#675 := (+ uf_5 #674)
+#673 := (= #675 -1::int)
+#1423 := (not #673)
+#88 := 2::int
+#671 := (>= uf_13 2::int)
+#1422 := (not #671)
+#668 := (>= uf_11 0::int)
+#1421 := (not #668)
+#385 := (= uf_6 uf_12)
+#1435 := (not #385)
+#382 := (= uf_4 uf_11)
+#1434 := (not #382)
+#2029 := (or #1434 #1435 #1348 #1349 #1421 #1422 #1423 #681 #2020)
+#2032 := (not #2029)
+decl uf_10 :: int
+#76 := uf_10
+#296 := (= uf_10 uf_12)
+#1420 := (not #296)
+#293 := (= uf_5 uf_11)
+#1419 := (not #293)
+#288 := (= #74 uf_10)
+#1418 := (not #288)
+#2023 := (or #1418 #1419 #1420 #1348 #1349 #1421 #1422 #1423 #678 #2020)
+#2026 := (not #2023)
+#2035 := (or #2026 #2032)
+#2038 := (not #2035)
+#2041 := (or #1348 #1349 #584 #2038)
+#2044 := (not #2041)
+#2064 := (or #2044 #2061)
+#2067 := (not #2064)
+#543 := (* -1::int uf_6)
+#544 := (+ #23 #543)
+#545 := (<= #544 0::int)
+#534 := (+ #19 #533)
+#532 := (>= #534 0::int)
+#1253 := (or #1216 #532 #545)
+#1995 := (forall (vars (?x3 int)) (:pat #1986) #1253)
+#2000 := (not #1995)
+#519 := (* -1::int #23)
+#520 := (+ uf_2 #519)
+#518 := (>= #520 0::int)
+#508 := (>= #19 1::int)
+#1231 := (or #1216 #508 #518)
+#1987 := (forall (vars (?x1 int)) (:pat #1986) #1231)
+#1992 := (not #1987)
+#39 := (uf_3 uf_4)
+#183 := (= uf_6 #39)
+#1461 := (not #183)
+#2070 := (or #1461 #1348 #1349 #1992 #2000 #2067)
+#1123 := (not #1118)
+#2073 := (not #2070)
+#2289 := [hypothesis]: #2073
+#1885 := (or #2070 #2064)
+#1887 := [def-axiom]: #1885
+#2290 := [unit-resolution #1887 #2289]: #2064
+#1898 := (or #2070 #1995)
+#1884 := [def-axiom]: #1898
+#2227 := [unit-resolution #1884 #2289]: #1995
+#1907 := (or #2070 #183)
+#1891 := [def-axiom]: #1907
+#2228 := [unit-resolution #1891 #2289]: #183
+#2142 := (or #2058 #1461 #2000)
+#2082 := (uf_3 uf_7)
+#2086 := (= uf_8 #2082)
+#2136 := (= #39 #2082)
+#2134 := (= #2082 #39)
+#48 := (= uf_7 uf_4)
+#2119 := [hypothesis]: #2061
+#1824 := (or #2058 #210)
+#1825 := [def-axiom]: #1824
+#2130 := [unit-resolution #1825 #2119]: #210
+#2131 := [symm #2130]: #48
+#2135 := [monotonicity #2131]: #2134
+#2137 := [symm #2135]: #2136
+#2138 := (= uf_8 #39)
+#2132 := [hypothesis]: #183
+#50 := (= uf_8 uf_6)
+#1826 := (or #2058 #213)
+#1827 := [def-axiom]: #1826
+#2121 := [unit-resolution #1827 #2119]: #213
+#2133 := [symm #2121]: #50
+#2139 := [trans #2133 #2132]: #2138
+#2140 := [trans #2139 #2137]: #2086
+#2114 := (not #2086)
+#2080 := (>= uf_7 0::int)
+#2081 := (not #2080)
+#1577 := (* -1::int uf_7)
+#1575 := (+ uf_1 #1577)
+#1578 := (<= #1575 0::int)
+#2092 := (or #1578 #2081 #2086)
+#2097 := (not #2092)
+#1911 := (or #2058 #2052)
+#1912 := [def-axiom]: #1911
+#2120 := [unit-resolution #1912 #2119]: #2052
+#630 := (* -1::int uf_8)
+#1747 := (+ uf_6 #630)
+#1737 := (<= #1747 0::int)
+#2122 := (or #1346 #1737)
+#2123 := [th-lemma]: #2122
+#2124 := [unit-resolution #2123 #2121]: #1737
+#1630 := [hypothesis]: #1995
+#1899 := (or #2058 #584)
+#1901 := [def-axiom]: #1899
+#2125 := [unit-resolution #1901 #2119]: #584
+#1656 := (not #1737)
+#1615 := (or #1474 #650 #2000 #1656)
+#1642 := [hypothesis]: #584
+#1724 := (+ uf_5 #922)
+#1725 := (<= #1724 0::int)
+#1717 := (+ uf_6 #1079)
+#1718 := (>= #1717 0::int)
+#1645 := (not #1718)
+#1652 := [hypothesis]: #1737
+#1810 := (not #1081)
+#1621 := [hypothesis]: #1477
+#1811 := (or #1474 #1810)
+#1770 := [def-axiom]: #1811
+#1623 := [unit-resolution #1770 #1621]: #1810
+#1639 := (or #1645 #1081 #1656)
+#1653 := [hypothesis]: #1810
+#1654 := [hypothesis]: #1718
+#1655 := [th-lemma #1654 #1653 #1652]: false
+#1641 := [lemma #1655]: #1639
+#1628 := [unit-resolution #1641 #1623 #1652]: #1645
+#1631 := (or #1718 #1725)
+#1927 := (or #1474 #926)
+#1809 := [def-axiom]: #1927
+#1629 := [unit-resolution #1809 #1621]: #926
+#1699 := (or #2000 #1306 #1718 #1725)
+#1729 := (+ #919 #543)
+#1730 := (<= #1729 0::int)
+#1733 := (+ ?x6!2 #533)
+#1734 := (>= #1733 0::int)
+#1738 := (or #1306 #1734 #1730)
+#1707 := (or #2000 #1738)
+#1686 := (iff #1707 #1699)
+#1702 := (or #1306 #1718 #1725)
+#1693 := (or #2000 #1702)
+#1695 := (iff #1693 #1699)
+#1697 := [rewrite]: #1695
+#1694 := (iff #1707 #1693)
+#1705 := (iff #1738 #1702)
+#1698 := (or #1306 #1725 #1718)
+#1703 := (iff #1698 #1702)
+#1704 := [rewrite]: #1703
+#1700 := (iff #1738 #1698)
+#1709 := (iff #1730 #1718)
+#1711 := (+ #543 #919)
+#1714 := (<= #1711 0::int)
+#1719 := (iff #1714 #1718)
+#1720 := [rewrite]: #1719
+#1715 := (iff #1730 #1714)
+#1712 := (= #1729 #1711)
+#1713 := [rewrite]: #1712
+#1716 := [monotonicity #1713]: #1715
+#1721 := [trans #1716 #1720]: #1709
+#1708 := (iff #1734 #1725)
+#1731 := (+ #533 ?x6!2)
+#1735 := (>= #1731 0::int)
+#1723 := (iff #1735 #1725)
+#1726 := [rewrite]: #1723
+#1736 := (iff #1734 #1735)
+#1739 := (= #1733 #1731)
+#1732 := [rewrite]: #1739
+#1722 := [monotonicity #1732]: #1736
+#1710 := [trans #1722 #1726]: #1708
+#1701 := [monotonicity #1710 #1721]: #1700
+#1706 := [trans #1701 #1704]: #1705
+#1696 := [monotonicity #1706]: #1694
+#1687 := [trans #1696 #1697]: #1686
+#1692 := [quant-inst]: #1707
+#1688 := [mp #1692 #1687]: #1699
+#1632 := [unit-resolution #1688 #1630 #1629]: #1631
+#1633 := [unit-resolution #1632 #1628]: #1725
+#925 := (not #924)
+#1926 := (or #1474 #925)
+#1924 := [def-axiom]: #1926
+#1622 := [unit-resolution #1924 #1621]: #925
+#1634 := [th-lemma #1622 #1633 #1642]: false
+#1617 := [lemma #1634]: #1615
+#2126 := [unit-resolution #1617 #2125 #1630 #2124]: #1474
+#1815 := (or #2055 #2047 #1477)
+#1823 := [def-axiom]: #1815
+#2127 := [unit-resolution #1823 #2126 #2120]: #2047
+#1919 := (not #2047)
+#2100 := (or #1919 #2097)
+#2083 := (= #2082 uf_8)
+#2084 := (or #2083 #2081 #1578)
+#2085 := (not #2084)
+#2101 := (or #1919 #2085)
+#2103 := (iff #2101 #2100)
+#2105 := (iff #2100 #2100)
+#2106 := [rewrite]: #2105
+#2098 := (iff #2085 #2097)
+#2095 := (iff #2084 #2092)
+#2089 := (or #2086 #2081 #1578)
+#2093 := (iff #2089 #2092)
+#2094 := [rewrite]: #2093
+#2090 := (iff #2084 #2089)
+#2087 := (iff #2083 #2086)
+#2088 := [rewrite]: #2087
+#2091 := [monotonicity #2088]: #2090
+#2096 := [trans #2091 #2094]: #2095
+#2099 := [monotonicity #2096]: #2098
+#2104 := [monotonicity #2099]: #2103
+#2107 := [trans #2104 #2106]: #2103
+#2102 := [quant-inst]: #2101
+#2108 := [mp #2102 #2107]: #2100
+#2128 := [unit-resolution #2108 #2127]: #2097
+#2115 := (or #2092 #2114)
+#2116 := [def-axiom]: #2115
+#2129 := [unit-resolution #2116 #2128]: #2114
+#2141 := [unit-resolution #2129 #2140]: false
+#2143 := [lemma #2141]: #2142
+#2229 := [unit-resolution #2143 #2228 #2227]: #2058
+#1906 := (or #2067 #2044 #2061)
+#1900 := [def-axiom]: #1906
+#2299 := [unit-resolution #1900 #2229 #2290]: #2044
+#1934 := (or #2041 #2035)
+#1928 := [def-axiom]: #1934
+#2376 := [unit-resolution #1928 #2299]: #2035
+#2374 := (or #2029 #2000 #1461)
+#2334 := [hypothesis]: #2032
+#1952 := (or #2029 #2017)
+#1941 := [def-axiom]: #1952
+#2335 := [unit-resolution #1941 #2334]: #2017
+#2226 := (= #39 #106)
+#2340 := (= #106 #39)
+#112 := (= uf_11 uf_4)
+#1678 := (or #2029 #382)
+#1679 := [def-axiom]: #1678
+#2336 := [unit-resolution #1679 #2334]: #382
+#2337 := [symm #2336]: #112
+#2341 := [monotonicity #2337]: #2340
+#2342 := [symm #2341]: #2226
+#2343 := (= uf_12 #39)
+#113 := (= uf_12 uf_6)
+#1953 := (or #2029 #385)
+#1957 := [def-axiom]: #1953
+#2338 := [unit-resolution #1957 #2334]: #385
+#2339 := [symm #2338]: #113
+#2344 := [trans #2339 #2132]: #2343
+#2345 := [trans #2344 #2342]: #358
+#970 := (not #358)
+#1643 := (or #2011 #970)
+#1978 := [def-axiom]: #1643
+#2346 := [unit-resolution #1978 #2345]: #2011
+#1977 := (or #2020 #2014 #1409)
+#1620 := [def-axiom]: #1977
+#2347 := [unit-resolution #1620 #2346 #2335]: #1409
+#1981 := (or #1404 #1123)
+#1982 := [def-axiom]: #1981
+#2348 := [unit-resolution #1982 #2347]: #1123
+#1664 := (>= #675 -1::int)
+#1665 := (or #2029 #673)
+#1947 := [def-axiom]: #1665
+#2349 := [unit-resolution #1947 #2334]: #673
+#2350 := (or #1423 #1664)
+#2351 := [th-lemma]: #2350
+#2352 := [unit-resolution #2351 #2349]: #1664
+#2241 := (+ uf_5 #1116)
+#2311 := (>= #2241 0::int)
+#2367 := (not #2311)
+#2300 := (= uf_5 ?x7!3)
+#2331 := (not #2300)
+#2310 := (= #74 #955)
+#2315 := (not #2310)
+#2314 := (+ #74 #1138)
+#2316 := (>= #2314 0::int)
+#2320 := (not #2316)
+#1951 := (or #2029 #678)
+#1948 := [def-axiom]: #1951
+#2353 := [unit-resolution #1948 #2334]: #678
+#1983 := (not #1140)
+#1984 := (or #1404 #1983)
+#1979 := [def-axiom]: #1984
+#2354 := [unit-resolution #1979 #2347]: #1983
+#1579 := (+ uf_6 #733)
+#1955 := (<= #1579 0::int)
+#2355 := (or #1435 #1955)
+#2356 := [th-lemma]: #2355
+#2357 := [unit-resolution #2356 #2338]: #1955
+#2296 := (not #1955)
+#2321 := (or #2320 #2296 #1140 #681)
+#2317 := [hypothesis]: #678
+#2292 := [hypothesis]: #1983
+#2291 := [hypothesis]: #1955
+#2318 := [hypothesis]: #2316
+#2319 := [th-lemma #2318 #2291 #2292 #2317]: false
+#2322 := [lemma #2319]: #2321
+#2358 := [unit-resolution #2322 #2357 #2354 #2353]: #2320
+#2323 := (or #2315 #2316)
+#2324 := [th-lemma]: #2323
+#2359 := [unit-resolution #2324 #2358]: #2315
+#2332 := (or #2331 #2310)
+#2328 := [hypothesis]: #2300
+#2329 := [monotonicity #2328]: #2310
+#2327 := [hypothesis]: #2315
+#2330 := [unit-resolution #2327 #2329]: false
+#2333 := [lemma #2330]: #2332
+#2360 := [unit-resolution #2333 #2359]: #2331
+#2370 := (or #2300 #2367)
+#2242 := (<= #2241 0::int)
+#2253 := (+ uf_6 #1138)
+#2254 := (>= #2253 0::int)
+#2295 := (not #2254)
+#2297 := (or #2295 #1140 #2296)
+#2293 := [hypothesis]: #2254
+#2294 := [th-lemma #2293 #2292 #2291]: false
+#2298 := [lemma #2294]: #2297
+#2361 := [unit-resolution #2298 #2354 #2357]: #2295
+#2363 := (or #2242 #2254)
+#1648 := (or #1404 #961)
+#1649 := [def-axiom]: #1648
+#2362 := [unit-resolution #1649 #2347]: #961
+#2262 := (or #2000 #1389 #2242 #2254)
+#2230 := (+ #955 #543)
+#2231 := (<= #2230 0::int)
+#2232 := (+ ?x7!3 #533)
+#2233 := (>= #2232 0::int)
+#2234 := (or #1389 #2233 #2231)
+#2263 := (or #2000 #2234)
+#2270 := (iff #2263 #2262)
+#2259 := (or #1389 #2242 #2254)
+#2265 := (or #2000 #2259)
+#2268 := (iff #2265 #2262)
+#2269 := [rewrite]: #2268
+#2266 := (iff #2263 #2265)
+#2260 := (iff #2234 #2259)
+#2257 := (iff #2231 #2254)
+#2247 := (+ #543 #955)
+#2250 := (<= #2247 0::int)
+#2255 := (iff #2250 #2254)
+#2256 := [rewrite]: #2255
+#2251 := (iff #2231 #2250)
+#2248 := (= #2230 #2247)
+#2249 := [rewrite]: #2248
+#2252 := [monotonicity #2249]: #2251
+#2258 := [trans #2252 #2256]: #2257
+#2245 := (iff #2233 #2242)
+#2235 := (+ #533 ?x7!3)
+#2238 := (>= #2235 0::int)
+#2243 := (iff #2238 #2242)
+#2244 := [rewrite]: #2243
+#2239 := (iff #2233 #2238)
+#2236 := (= #2232 #2235)
+#2237 := [rewrite]: #2236
+#2240 := [monotonicity #2237]: #2239
+#2246 := [trans #2240 #2244]: #2245
+#2261 := [monotonicity #2246 #2258]: #2260
+#2267 := [monotonicity #2261]: #2266
+#2271 := [trans #2267 #2269]: #2270
+#2264 := [quant-inst]: #2263
+#2272 := [mp #2264 #2271]: #2262
+#2364 := [unit-resolution #2272 #1630 #2362]: #2363
+#2365 := [unit-resolution #2364 #2361]: #2242
+#2366 := (not #2242)
+#2368 := (or #2300 #2366 #2367)
+#2369 := [th-lemma]: #2368
+#2371 := [unit-resolution #2369 #2365]: #2370
+#2372 := [unit-resolution #2371 #2360]: #2367
+#2373 := [th-lemma #2372 #2352 #2348]: false
+#2375 := [lemma #2373]: #2374
+#2377 := [unit-resolution #2375 #2227 #2228]: #2029
+#1940 := (or #2038 #2026 #2032)
+#1946 := [def-axiom]: #1940
+#2378 := [unit-resolution #1946 #2377 #2376]: #2026
+#1970 := (or #2023 #2017)
+#1973 := [def-axiom]: #1970
+#2379 := [unit-resolution #1973 #2378]: #2017
+#2388 := (= #74 #106)
+#2384 := (= #106 #74)
+#80 := (= uf_11 uf_5)
+#1625 := (or #2023 #293)
+#1626 := [def-axiom]: #1625
+#2382 := [unit-resolution #1626 #2378]: #293
+#2383 := [symm #2382]: #80
+#2385 := [monotonicity #2383]: #2384
+#2389 := [symm #2385]: #2388
+#2390 := (= uf_12 #74)
+#77 := (= uf_10 #74)
+#1961 := (or #2023 #288)
+#1624 := [def-axiom]: #1961
+#2381 := [unit-resolution #1624 #2378]: #288
+#2387 := [symm #2381]: #77
+#82 := (= uf_12 uf_10)
+#1627 := (or #2023 #296)
+#1963 := [def-axiom]: #1627
+#2380 := [unit-resolution #1963 #2378]: #296
+#2386 := [symm #2380]: #82
+#2391 := [trans #2386 #2387]: #2390
+#2392 := [trans #2391 #2389]: #358
+#2393 := [unit-resolution #1978 #2392]: #2011
+#2394 := [unit-resolution #1620 #2393 #2379]: #1409
+#2395 := [unit-resolution #1982 #2394]: #1123
+#2157 := (+ #74 #733)
+#2158 := (<= #2157 0::int)
+#2156 := (= #74 uf_12)
+#2396 := [trans #2381 #2380]: #2156
+#2397 := (not #2156)
+#2398 := (or #2397 #2158)
+#2399 := [th-lemma]: #2398
+#2400 := [unit-resolution #2399 #2396]: #2158
+#1612 := (or #2023 #681)
+#1972 := [def-axiom]: #1612
+#2401 := [unit-resolution #1972 #2378]: #681
+#2402 := [unit-resolution #1979 #2394]: #1983
+#2403 := (not #2158)
+#2404 := (or #2295 #1140 #2403 #678)
+#2405 := [th-lemma]: #2404
+#2406 := [unit-resolution #2405 #2402 #2401 #2400]: #2295
+#2407 := [unit-resolution #1649 #2394]: #961
+#2408 := [unit-resolution #2272 #2227 #2407 #2406]: #2242
+#2409 := (or #2320 #1140 #2403)
+#2410 := [th-lemma]: #2409
+#2411 := [unit-resolution #2410 #2402 #2400]: #2320
+#2412 := [unit-resolution #2324 #2411]: #2315
+#2413 := [unit-resolution #2333 #2412]: #2331
+#2414 := [unit-resolution #2369 #2413 #2408]: #2367
+#1971 := (or #2023 #673)
+#1611 := [def-axiom]: #1971
+#2415 := [unit-resolution #1611 #2378]: #673
+#2416 := [unit-resolution #2351 #2415]: #1664
+#2417 := [th-lemma #2416 #2414 #2395]: false
+#2418 := [lemma #2417]: #2070
+#2076 := (or #1213 #2073)
+#1283 := (forall (vars (?x4 int)) #1280)
+#1480 := (or #1283 #1477)
+#1483 := (not #1480)
+#1486 := (or #1345 #1346 #1347 #1348 #1349 #650 #1483)
+#1489 := (not #1486)
+#1376 := (forall (vars (?x7 int)) #1371)
+#1382 := (not #1376)
+#1383 := (or #358 #1382)
+#1384 := (not #1383)
+#1412 := (or #1384 #1409)
+#1424 := (not #1412)
+#1436 := (or #1434 #1435 #1348 #1349 #1421 #1422 #1423 #681 #1424)
+#1437 := (not #1436)
+#1425 := (or #1418 #1419 #1420 #1348 #1349 #1421 #1422 #1423 #678 #1424)
+#1426 := (not #1425)
+#1442 := (or #1426 #1437)
+#1448 := (not #1442)
+#1449 := (or #1348 #1349 #584 #1448)
+#1450 := (not #1449)
+#1495 := (or #1450 #1489)
+#1500 := (not #1495)
+#1258 := (forall (vars (?x3 int)) #1253)
+#1463 := (not #1258)
+#1236 := (forall (vars (?x1 int)) #1231)
+#1462 := (not #1236)
+#1503 := (or #1461 #1348 #1349 #1462 #1463 #1500)
+#1506 := (not #1503)
+#1509 := (or #1213 #1506)
+#2077 := (iff #1509 #2076)
+#2074 := (iff #1506 #2073)
+#2071 := (iff #1503 #2070)
+#2068 := (iff #1500 #2067)
+#2065 := (iff #1495 #2064)
+#2062 := (iff #1489 #2061)
+#2059 := (iff #1486 #2058)
+#2056 := (iff #1483 #2055)
+#2053 := (iff #1480 #2052)
+#2050 := (iff #1283 #2047)
+#2048 := (iff #1280 #1280)
+#2049 := [refl]: #2048
+#2051 := [quant-intro #2049]: #2050
+#2054 := [monotonicity #2051]: #2053
+#2057 := [monotonicity #2054]: #2056
+#2060 := [monotonicity #2057]: #2059
+#2063 := [monotonicity #2060]: #2062
+#2045 := (iff #1450 #2044)
+#2042 := (iff #1449 #2041)
+#2039 := (iff #1448 #2038)
+#2036 := (iff #1442 #2035)
+#2033 := (iff #1437 #2032)
+#2030 := (iff #1436 #2029)
+#2021 := (iff #1424 #2020)
+#2018 := (iff #1412 #2017)
+#2015 := (iff #1384 #2014)
+#2012 := (iff #1383 #2011)
+#2009 := (iff #1382 #2008)
+#2006 := (iff #1376 #2003)
+#2004 := (iff #1371 #1371)
+#2005 := [refl]: #2004
+#2007 := [quant-intro #2005]: #2006
+#2010 := [monotonicity #2007]: #2009
+#2013 := [monotonicity #2010]: #2012
+#2016 := [monotonicity #2013]: #2015
+#2019 := [monotonicity #2016]: #2018
+#2022 := [monotonicity #2019]: #2021
+#2031 := [monotonicity #2022]: #2030
+#2034 := [monotonicity #2031]: #2033
+#2027 := (iff #1426 #2026)
+#2024 := (iff #1425 #2023)
+#2025 := [monotonicity #2022]: #2024
+#2028 := [monotonicity #2025]: #2027
+#2037 := [monotonicity #2028 #2034]: #2036
+#2040 := [monotonicity #2037]: #2039
+#2043 := [monotonicity #2040]: #2042
+#2046 := [monotonicity #2043]: #2045
+#2066 := [monotonicity #2046 #2063]: #2065
+#2069 := [monotonicity #2066]: #2068
+#2001 := (iff #1463 #2000)
+#1998 := (iff #1258 #1995)
+#1996 := (iff #1253 #1253)
+#1997 := [refl]: #1996
+#1999 := [quant-intro #1997]: #1998
+#2002 := [monotonicity #1999]: #2001
+#1993 := (iff #1462 #1992)
+#1990 := (iff #1236 #1987)
+#1988 := (iff #1231 #1231)
+#1989 := [refl]: #1988
+#1991 := [quant-intro #1989]: #1990
+#1994 := [monotonicity #1991]: #1993
+#2072 := [monotonicity #1994 #2002 #2069]: #2071
+#2075 := [monotonicity #2072]: #2074
+#2078 := [monotonicity #2075]: #2077
+#1126 := (and #961 #1123)
+#1129 := (not #1126)
+#1145 := (or #1129 #1140)
+#1148 := (not #1145)
+#724 := (not #722)
+#727 := (and #505 #724)
+#730 := (not #727)
+#738 := (or #730 #735)
+#741 := (forall (vars (?x7 int)) #738)
+#980 := (and #970 #741)
+#1154 := (or #980 #1148)
+#1172 := (and #382 #385 #555 #557 #668 #671 #673 #678 #1154)
+#1162 := (and #288 #293 #296 #555 #557 #668 #671 #673 #681 #1154)
+#1177 := (or #1162 #1172)
+#1183 := (and #555 #557 #650 #1177)
+#1067 := (and #925 #926)
+#1070 := (not #1067)
+#1086 := (or #1070 #1081)
+#1089 := (not #1086)
+#904 := (not #903)
+#1058 := (and #904 #905)
+#1061 := (not #1058)
+#1064 := (or #1055 #1061)
+#1092 := (and #1064 #1089)
+#607 := (not #606)
+#610 := (and #505 #607)
+#613 := (not #610)
+#619 := (or #61 #613)
+#894 := (not #619)
+#897 := (forall (vars (?x4 int)) #894)
+#1095 := (or #897 #1092)
+#1101 := (and #210 #213 #216 #555 #557 #584 #1095)
+#1188 := (or #1101 #1183)
+#531 := (not #532)
+#537 := (and #505 #531)
+#540 := (not #537)
+#548 := (or #540 #545)
+#551 := (forall (vars (?x3 int)) #548)
+#506 := (not #508)
+#510 := (and #505 #506)
+#513 := (not #510)
+#522 := (or #513 #518)
+#525 := (forall (vars (?x1 int)) #522)
+#1194 := (and #183 #525 #551 #555 #557 #1188)
+#1036 := (and #874 #875)
+#1039 := (not #1036)
+#1045 := (or #872 #1039)
+#1050 := (not #1045)
+#1199 := (or #1050 #1194)
+#1512 := (iff #1199 #1509)
+#1321 := (or #924 #1306 #1081)
+#1333 := (or #1332 #1321)
+#1334 := (not #1333)
+#1339 := (or #1283 #1334)
+#1350 := (not #1339)
+#1351 := (or #1345 #1346 #1347 #1348 #1349 #650 #1350)
+#1352 := (not #1351)
+#1455 := (or #1352 #1450)
+#1464 := (not #1455)
+#1465 := (or #1461 #1348 #1349 #1462 #1463 #1464)
+#1466 := (not #1465)
+#1471 := (or #1213 #1466)
+#1510 := (iff #1471 #1509)
+#1507 := (iff #1466 #1506)
+#1504 := (iff #1465 #1503)
+#1501 := (iff #1464 #1500)
+#1498 := (iff #1455 #1495)
+#1492 := (or #1489 #1450)
+#1496 := (iff #1492 #1495)
+#1497 := [rewrite]: #1496
+#1493 := (iff #1455 #1492)
+#1490 := (iff #1352 #1489)
+#1487 := (iff #1351 #1486)
+#1484 := (iff #1350 #1483)
+#1481 := (iff #1339 #1480)
+#1478 := (iff #1334 #1477)
+#1475 := (iff #1333 #1474)
+#1476 := [rewrite]: #1475
+#1479 := [monotonicity #1476]: #1478
+#1482 := [monotonicity #1479]: #1481
+#1485 := [monotonicity #1482]: #1484
+#1488 := [monotonicity #1485]: #1487
+#1491 := [monotonicity #1488]: #1490
+#1494 := [monotonicity #1491]: #1493
+#1499 := [trans #1494 #1497]: #1498
+#1502 := [monotonicity #1499]: #1501
+#1505 := [monotonicity #1502]: #1504
+#1508 := [monotonicity #1505]: #1507
+#1511 := [monotonicity #1508]: #1510
+#1472 := (iff #1199 #1471)
+#1469 := (iff #1194 #1466)
+#1458 := (and #183 #1236 #1258 #555 #557 #1455)
+#1467 := (iff #1458 #1466)
+#1468 := [rewrite]: #1467
+#1459 := (iff #1194 #1458)
+#1456 := (iff #1188 #1455)
+#1453 := (iff #1183 #1450)
+#1445 := (and #555 #557 #650 #1442)
+#1451 := (iff #1445 #1450)
+#1452 := [rewrite]: #1451
+#1446 := (iff #1183 #1445)
+#1443 := (iff #1177 #1442)
+#1440 := (iff #1172 #1437)
+#1431 := (and #382 #385 #555 #557 #668 #671 #673 #678 #1412)
+#1438 := (iff #1431 #1437)
+#1439 := [rewrite]: #1438
+#1432 := (iff #1172 #1431)
+#1413 := (iff #1154 #1412)
+#1410 := (iff #1148 #1409)
+#1407 := (iff #1145 #1404)
+#1390 := (or #1389 #1118)
+#1401 := (or #1390 #1140)
+#1405 := (iff #1401 #1404)
+#1406 := [rewrite]: #1405
+#1402 := (iff #1145 #1401)
+#1399 := (iff #1129 #1390)
+#1391 := (not #1390)
+#1394 := (not #1391)
+#1397 := (iff #1394 #1390)
+#1398 := [rewrite]: #1397
+#1395 := (iff #1129 #1394)
+#1392 := (iff #1126 #1391)
+#1393 := [rewrite]: #1392
+#1396 := [monotonicity #1393]: #1395
+#1400 := [trans #1396 #1398]: #1399
+#1403 := [monotonicity #1400]: #1402
+#1408 := [trans #1403 #1406]: #1407
+#1411 := [monotonicity #1408]: #1410
+#1387 := (iff #980 #1384)
+#1379 := (and #970 #1376)
+#1385 := (iff #1379 #1384)
+#1386 := [rewrite]: #1385
+#1380 := (iff #980 #1379)
+#1377 := (iff #741 #1376)
+#1374 := (iff #738 #1371)
+#1357 := (or #1216 #722)
+#1368 := (or #1357 #735)
+#1372 := (iff #1368 #1371)
+#1373 := [rewrite]: #1372
+#1369 := (iff #738 #1368)
+#1366 := (iff #730 #1357)
+#1358 := (not #1357)
+#1361 := (not #1358)
+#1364 := (iff #1361 #1357)
+#1365 := [rewrite]: #1364
+#1362 := (iff #730 #1361)
+#1359 := (iff #727 #1358)
+#1360 := [rewrite]: #1359
+#1363 := [monotonicity #1360]: #1362
+#1367 := [trans #1363 #1365]: #1366
+#1370 := [monotonicity #1367]: #1369
+#1375 := [trans #1370 #1373]: #1374
+#1378 := [quant-intro #1375]: #1377
+#1381 := [monotonicity #1378]: #1380
+#1388 := [trans #1381 #1386]: #1387
+#1414 := [monotonicity #1388 #1411]: #1413
+#1433 := [monotonicity #1414]: #1432
+#1441 := [trans #1433 #1439]: #1440
+#1429 := (iff #1162 #1426)
+#1415 := (and #288 #293 #296 #555 #557 #668 #671 #673 #681 #1412)
+#1427 := (iff #1415 #1426)
+#1428 := [rewrite]: #1427
+#1416 := (iff #1162 #1415)
+#1417 := [monotonicity #1414]: #1416
+#1430 := [trans #1417 #1428]: #1429
+#1444 := [monotonicity #1430 #1441]: #1443
+#1447 := [monotonicity #1444]: #1446
+#1454 := [trans #1447 #1452]: #1453
+#1355 := (iff #1101 #1352)
+#1342 := (and #210 #213 #216 #555 #557 #584 #1339)
+#1353 := (iff #1342 #1352)
+#1354 := [rewrite]: #1353
+#1343 := (iff #1101 #1342)
+#1340 := (iff #1095 #1339)
+#1337 := (iff #1092 #1334)
+#1326 := (not #1321)
+#1329 := (and #1301 #1326)
+#1335 := (iff #1329 #1334)
+#1336 := [rewrite]: #1335
+#1330 := (iff #1092 #1329)
+#1327 := (iff #1089 #1326)
+#1324 := (iff #1086 #1321)
+#1307 := (or #924 #1306)
+#1318 := (or #1307 #1081)
+#1322 := (iff #1318 #1321)
+#1323 := [rewrite]: #1322
+#1319 := (iff #1086 #1318)
+#1316 := (iff #1070 #1307)
+#1308 := (not #1307)
+#1311 := (not #1308)
+#1314 := (iff #1311 #1307)
+#1315 := [rewrite]: #1314
+#1312 := (iff #1070 #1311)
+#1309 := (iff #1067 #1308)
+#1310 := [rewrite]: #1309
+#1313 := [monotonicity #1310]: #1312
+#1317 := [trans #1313 #1315]: #1316
+#1320 := [monotonicity #1317]: #1319
+#1325 := [trans #1320 #1323]: #1324
+#1328 := [monotonicity #1325]: #1327
+#1304 := (iff #1064 #1301)
+#1287 := (or #903 #1286)
+#1298 := (or #1055 #1287)
+#1302 := (iff #1298 #1301)
+#1303 := [rewrite]: #1302
+#1299 := (iff #1064 #1298)
+#1296 := (iff #1061 #1287)
+#1288 := (not #1287)
+#1291 := (not #1288)
+#1294 := (iff #1291 #1287)
+#1295 := [rewrite]: #1294
+#1292 := (iff #1061 #1291)
+#1289 := (iff #1058 #1288)
+#1290 := [rewrite]: #1289
+#1293 := [monotonicity #1290]: #1292
+#1297 := [trans #1293 #1295]: #1296
+#1300 := [monotonicity #1297]: #1299
+#1305 := [trans #1300 #1303]: #1304
+#1331 := [monotonicity #1305 #1328]: #1330
+#1338 := [trans #1331 #1336]: #1337
+#1284 := (iff #897 #1283)
+#1281 := (iff #894 #1280)
+#1278 := (iff #619 #1275)
+#1261 := (or #1216 #606)
+#1272 := (or #61 #1261)
+#1276 := (iff #1272 #1275)
+#1277 := [rewrite]: #1276
+#1273 := (iff #619 #1272)
+#1270 := (iff #613 #1261)
+#1262 := (not #1261)
+#1265 := (not #1262)
+#1268 := (iff #1265 #1261)
+#1269 := [rewrite]: #1268
+#1266 := (iff #613 #1265)
+#1263 := (iff #610 #1262)
+#1264 := [rewrite]: #1263
+#1267 := [monotonicity #1264]: #1266
+#1271 := [trans #1267 #1269]: #1270
+#1274 := [monotonicity #1271]: #1273
+#1279 := [trans #1274 #1277]: #1278
+#1282 := [monotonicity #1279]: #1281
+#1285 := [quant-intro #1282]: #1284
+#1341 := [monotonicity #1285 #1338]: #1340
+#1344 := [monotonicity #1341]: #1343
+#1356 := [trans #1344 #1354]: #1355
+#1457 := [monotonicity #1356 #1454]: #1456
+#1259 := (iff #551 #1258)
+#1256 := (iff #548 #1253)
+#1239 := (or #1216 #532)
+#1250 := (or #1239 #545)
+#1254 := (iff #1250 #1253)
+#1255 := [rewrite]: #1254
+#1251 := (iff #548 #1250)
+#1248 := (iff #540 #1239)
+#1240 := (not #1239)
+#1243 := (not #1240)
+#1246 := (iff #1243 #1239)
+#1247 := [rewrite]: #1246
+#1244 := (iff #540 #1243)
+#1241 := (iff #537 #1240)
+#1242 := [rewrite]: #1241
+#1245 := [monotonicity #1242]: #1244
+#1249 := [trans #1245 #1247]: #1248
+#1252 := [monotonicity #1249]: #1251
+#1257 := [trans #1252 #1255]: #1256
+#1260 := [quant-intro #1257]: #1259
+#1237 := (iff #525 #1236)
+#1234 := (iff #522 #1231)
+#1217 := (or #1216 #508)
+#1228 := (or #1217 #518)
+#1232 := (iff #1228 #1231)
+#1233 := [rewrite]: #1232
+#1229 := (iff #522 #1228)
+#1226 := (iff #513 #1217)
+#1218 := (not #1217)
+#1221 := (not #1218)
+#1224 := (iff #1221 #1217)
+#1225 := [rewrite]: #1224
+#1222 := (iff #513 #1221)
+#1219 := (iff #510 #1218)
+#1220 := [rewrite]: #1219
+#1223 := [monotonicity #1220]: #1222
+#1227 := [trans #1223 #1225]: #1226
+#1230 := [monotonicity #1227]: #1229
+#1235 := [trans #1230 #1233]: #1234
+#1238 := [quant-intro #1235]: #1237
+#1460 := [monotonicity #1238 #1260 #1457]: #1459
+#1470 := [trans #1460 #1468]: #1469
+#1214 := (iff #1050 #1213)
+#1211 := (iff #1045 #1208)
+#914 := (or #873 #913)
+#1205 := (or #872 #914)
+#1209 := (iff #1205 #1208)
+#1210 := [rewrite]: #1209
+#1206 := (iff #1045 #1205)
+#1203 := (iff #1039 #914)
+#882 := (not #914)
+#935 := (not #882)
+#1035 := (iff #935 #914)
+#1202 := [rewrite]: #1035
+#968 := (iff #1039 #935)
+#883 := (iff #1036 #882)
+#934 := [rewrite]: #883
+#969 := [monotonicity #934]: #968
+#1204 := [trans #969 #1202]: #1203
+#1207 := [monotonicity #1204]: #1206
+#1212 := [trans #1207 #1210]: #1211
+#1215 := [monotonicity #1212]: #1214
+#1473 := [monotonicity #1215 #1470]: #1472
+#1513 := [trans #1473 #1511]: #1512
+#832 := (and #183 #551 #555 #557)
+#778 := (and #382 #385 #555 #557 #668 #671 #673 #678)
+#783 := (not #778)
+#992 := (not #783)
+#956 := (+ #955 #733)
+#957 := (<= #956 0::int)
+#958 := (+ ?x7!3 #674)
+#959 := (>= #958 0::int)
+#960 := (not #959)
+#962 := (and #961 #960)
+#963 := (not #962)
+#964 := (or #963 #957)
+#965 := (not #964)
+#984 := (or #965 #980)
+#995 := (and #984 #992)
+#714 := (and #288 #293 #296 #555 #557 #668 #671 #673 #681)
+#719 := (not #714)
+#951 := (not #719)
+#988 := (and #951 #984)
+#999 := (or #988 #995)
+#659 := (and #555 #557 #650)
+#664 := (not #659)
+#948 := (not #664)
+#1003 := (and #948 #999)
+#920 := (+ #919 #630)
+#921 := (<= #920 0::int)
+#927 := (and #926 #925)
+#928 := (not #927)
+#929 := (or #928 #921)
+#930 := (not #929)
+#906 := (and #905 #904)
+#907 := (not #906)
+#909 := (= #908 uf_8)
+#910 := (or #909 #907)
+#936 := (and #910 #930)
+#940 := (or #897 #936)
+#596 := (and #210 #213 #216 #555 #557 #584)
+#601 := (not #596)
+#891 := (not #601)
+#944 := (and #891 #940)
+#1007 := (or #944 #1003)
+#1026 := (and #525 #1007 #832)
+#876 := (and #875 #874)
+#877 := (not #876)
+#878 := (or #877 #872)
+#879 := (not #878)
+#1030 := (or #879 #1026)
+#1200 := (iff #1030 #1199)
+#1197 := (iff #1026 #1194)
+#1191 := (and #525 #1188 #832)
+#1195 := (iff #1191 #1194)
+#1196 := [rewrite]: #1195
+#1192 := (iff #1026 #1191)
+#1189 := (iff #1007 #1188)
+#1186 := (iff #1003 #1183)
+#1180 := (and #659 #1177)
+#1184 := (iff #1180 #1183)
+#1185 := [rewrite]: #1184
+#1181 := (iff #1003 #1180)
+#1178 := (iff #999 #1177)
+#1175 := (iff #995 #1172)
+#1169 := (and #1154 #778)
+#1173 := (iff #1169 #1172)
+#1174 := [rewrite]: #1173
+#1170 := (iff #995 #1169)
+#1167 := (iff #992 #778)
+#1168 := [rewrite]: #1167
+#1157 := (iff #984 #1154)
+#1151 := (or #1148 #980)
+#1155 := (iff #1151 #1154)
+#1156 := [rewrite]: #1155
+#1152 := (iff #984 #1151)
+#1149 := (iff #965 #1148)
+#1146 := (iff #964 #1145)
+#1143 := (iff #957 #1140)
+#1132 := (+ #733 #955)
+#1135 := (<= #1132 0::int)
+#1141 := (iff #1135 #1140)
+#1142 := [rewrite]: #1141
+#1136 := (iff #957 #1135)
+#1133 := (= #956 #1132)
+#1134 := [rewrite]: #1133
+#1137 := [monotonicity #1134]: #1136
+#1144 := [trans #1137 #1142]: #1143
+#1130 := (iff #963 #1129)
+#1127 := (iff #962 #1126)
+#1124 := (iff #960 #1123)
+#1121 := (iff #959 #1118)
+#1110 := (+ #674 ?x7!3)
+#1113 := (>= #1110 0::int)
+#1119 := (iff #1113 #1118)
+#1120 := [rewrite]: #1119
+#1114 := (iff #959 #1113)
+#1111 := (= #958 #1110)
+#1112 := [rewrite]: #1111
+#1115 := [monotonicity #1112]: #1114
+#1122 := [trans #1115 #1120]: #1121
+#1125 := [monotonicity #1122]: #1124
+#1128 := [monotonicity #1125]: #1127
+#1131 := [monotonicity #1128]: #1130
+#1147 := [monotonicity #1131 #1144]: #1146
+#1150 := [monotonicity #1147]: #1149
+#1153 := [monotonicity #1150]: #1152
+#1158 := [trans #1153 #1156]: #1157
+#1171 := [monotonicity #1158 #1168]: #1170
+#1176 := [trans #1171 #1174]: #1175
+#1165 := (iff #988 #1162)
+#1159 := (and #714 #1154)
+#1163 := (iff #1159 #1162)
+#1164 := [rewrite]: #1163
+#1160 := (iff #988 #1159)
+#1108 := (iff #951 #714)
+#1109 := [rewrite]: #1108
+#1161 := [monotonicity #1109 #1158]: #1160
+#1166 := [trans #1161 #1164]: #1165
+#1179 := [monotonicity #1166 #1176]: #1178
+#1106 := (iff #948 #659)
+#1107 := [rewrite]: #1106
+#1182 := [monotonicity #1107 #1179]: #1181
+#1187 := [trans #1182 #1185]: #1186
+#1104 := (iff #944 #1101)
+#1098 := (and #596 #1095)
+#1102 := (iff #1098 #1101)
+#1103 := [rewrite]: #1102
+#1099 := (iff #944 #1098)
+#1096 := (iff #940 #1095)
+#1093 := (iff #936 #1092)
+#1090 := (iff #930 #1089)
+#1087 := (iff #929 #1086)
+#1084 := (iff #921 #1081)
+#1073 := (+ #630 #919)
+#1076 := (<= #1073 0::int)
+#1082 := (iff #1076 #1081)
+#1083 := [rewrite]: #1082
+#1077 := (iff #921 #1076)
+#1074 := (= #920 #1073)
+#1075 := [rewrite]: #1074
+#1078 := [monotonicity #1075]: #1077
+#1085 := [trans #1078 #1083]: #1084
+#1071 := (iff #928 #1070)
+#1068 := (iff #927 #1067)
+#1069 := [rewrite]: #1068
+#1072 := [monotonicity #1069]: #1071
+#1088 := [monotonicity #1072 #1085]: #1087
+#1091 := [monotonicity #1088]: #1090
+#1065 := (iff #910 #1064)
+#1062 := (iff #907 #1061)
+#1059 := (iff #906 #1058)
+#1060 := [rewrite]: #1059
+#1063 := [monotonicity #1060]: #1062
+#1056 := (iff #909 #1055)
+#1057 := [rewrite]: #1056
+#1066 := [monotonicity #1057 #1063]: #1065
+#1094 := [monotonicity #1066 #1091]: #1093
+#1097 := [monotonicity #1094]: #1096
+#1053 := (iff #891 #596)
+#1054 := [rewrite]: #1053
+#1100 := [monotonicity #1054 #1097]: #1099
+#1105 := [trans #1100 #1103]: #1104
+#1190 := [monotonicity #1105 #1187]: #1189
+#1193 := [monotonicity #1190]: #1192
+#1198 := [trans #1193 #1196]: #1197
+#1051 := (iff #879 #1050)
+#1048 := (iff #878 #1045)
+#1042 := (or #1039 #872)
+#1046 := (iff #1042 #1045)
+#1047 := [rewrite]: #1046
+#1043 := (iff #878 #1042)
+#1040 := (iff #877 #1039)
+#1037 := (iff #876 #1036)
+#1038 := [rewrite]: #1037
+#1041 := [monotonicity #1038]: #1040
+#1044 := [monotonicity #1041]: #1043
+#1049 := [trans #1044 #1047]: #1048
+#1052 := [monotonicity #1049]: #1051
+#1201 := [monotonicity #1052 #1198]: #1200
+#837 := (not #832)
+#744 := (not #741)
+#750 := (or #358 #744)
+#755 := (and #741 #750)
+#786 := (or #755 #783)
+#758 := (or #719 #755)
+#789 := (and #758 #786)
+#792 := (or #664 #789)
+#631 := (+ #23 #630)
+#632 := (<= #631 0::int)
+#635 := (or #613 #632)
+#638 := (forall (vars (?x6 int)) #635)
+#624 := (exists (vars (?x4 int)) #619)
+#627 := (not #624)
+#641 := (or #627 #638)
+#644 := (and #624 #641)
+#647 := (or #601 #644)
+#795 := (and #647 #792)
+#528 := (not #525)
+#858 := (or #528 #795 #837)
+#863 := (and #525 #858)
+#866 := (not #863)
+#1031 := (~ #866 #1030)
+#1027 := (not #858)
+#1028 := (~ #1027 #1026)
+#1023 := (not #837)
+#1024 := (~ #1023 #832)
+#1021 := (~ #832 #832)
+#1019 := (~ #557 #557)
+#1020 := [refl]: #1019
+#1017 := (~ #555 #555)
+#1018 := [refl]: #1017
+#1015 := (~ #551 #551)
+#1013 := (~ #548 #548)
+#1014 := [refl]: #1013
+#1016 := [nnf-pos #1014]: #1015
+#1011 := (~ #183 #183)
+#1012 := [refl]: #1011
+#1022 := [monotonicity #1012 #1016 #1018 #1020]: #1021
+#1025 := [nnf-neg #1022]: #1024
+#1008 := (not #795)
+#1009 := (~ #1008 #1007)
+#1004 := (not #792)
+#1005 := (~ #1004 #1003)
+#1000 := (not #789)
+#1001 := (~ #1000 #999)
+#996 := (not #786)
+#997 := (~ #996 #995)
+#993 := (~ #992 #992)
+#994 := [refl]: #993
+#985 := (not #755)
+#986 := (~ #985 #984)
+#981 := (not #750)
+#982 := (~ #981 #980)
+#977 := (not #744)
+#978 := (~ #977 #741)
+#975 := (~ #741 #741)
+#973 := (~ #738 #738)
+#974 := [refl]: #973
+#976 := [nnf-pos #974]: #975
+#979 := [nnf-neg #976]: #978
+#971 := (~ #970 #970)
+#972 := [refl]: #971
+#983 := [nnf-neg #972 #979]: #982
+#966 := (~ #744 #965)
+#967 := [sk]: #966
+#987 := [nnf-neg #967 #983]: #986
+#998 := [nnf-neg #987 #994]: #997
+#989 := (not #758)
+#990 := (~ #989 #988)
+#952 := (~ #951 #951)
+#953 := [refl]: #952
+#991 := [nnf-neg #953 #987]: #990
+#1002 := [nnf-neg #991 #998]: #1001
+#949 := (~ #948 #948)
+#950 := [refl]: #949
+#1006 := [nnf-neg #950 #1002]: #1005
+#945 := (not #647)
+#946 := (~ #945 #944)
+#941 := (not #644)
+#942 := (~ #941 #940)
+#937 := (not #641)
+#938 := (~ #937 #936)
+#931 := (not #638)
+#932 := (~ #931 #930)
+#933 := [sk]: #932
+#915 := (not #627)
+#916 := (~ #915 #910)
+#911 := (~ #624 #910)
+#912 := [sk]: #911
+#917 := [nnf-neg #912]: #916
+#939 := [nnf-neg #917 #933]: #938
+#898 := (~ #627 #897)
+#895 := (~ #894 #894)
+#896 := [refl]: #895
+#899 := [nnf-neg #896]: #898
+#943 := [nnf-neg #899 #939]: #942
+#892 := (~ #891 #891)
+#893 := [refl]: #892
+#947 := [nnf-neg #893 #943]: #946
+#1010 := [nnf-neg #947 #1006]: #1009
+#888 := (not #528)
+#889 := (~ #888 #525)
+#886 := (~ #525 #525)
+#884 := (~ #522 #522)
+#885 := [refl]: #884
+#887 := [nnf-pos #885]: #886
+#890 := [nnf-neg #887]: #889
+#1029 := [nnf-neg #890 #1010 #1025]: #1028
+#880 := (~ #528 #879)
+#881 := [sk]: #880
+#1032 := [nnf-neg #881 #1029]: #1031
+#9 := (= uf_2 #8)
+#575 := (and #9 #183 #551 #555 #557)
+#580 := (not #575)
+#798 := (or #580 #795)
+#801 := (and #9 #798)
+#804 := (or #528 #801)
+#807 := (and #525 #804)
+#822 := (not #807)
+#867 := (iff #822 #866)
+#864 := (iff #807 #863)
+#861 := (iff #804 #858)
+#843 := (or #795 #837)
+#855 := (or #528 #843)
+#859 := (iff #855 #858)
+#860 := [rewrite]: #859
+#856 := (iff #804 #855)
+#853 := (iff #801 #843)
#1 := true
-#91 := (implies false true)
-#90 := (= #89 uf_12)
-#92 := (implies #90 #91)
-#93 := (and #90 #92)
-#86 := (<= #20 uf_12)
-#84 := (< #16 uf_13)
-#17 := (<= 0::int #16)
-#85 := (and #17 #84)
-#87 := (implies #85 #86)
-#88 := (forall (vars (?x8 int)) #87)
-#94 := (implies #88 #93)
-#95 := (and #88 #94)
-#96 := (implies true #95)
-#82 := (<= 2::int uf_13)
-#76 := (<= 0::int uf_11)
-#83 := (and #76 #82)
-#97 := (implies #83 #96)
-#79 := (+ uf_5 1::int)
-#80 := (= uf_13 #79)
-#98 := (implies #80 #97)
-#28 := (<= 1::int uf_5)
-#77 := (and #76 #28)
-#99 := (implies #77 #98)
-#100 := (implies true #99)
-#111 := (= uf_12 uf_6)
-#112 := (implies #111 #100)
-#110 := (= uf_11 uf_4)
-#113 := (implies #110 #112)
-#114 := (implies true #113)
-#26 := (<= 0::int uf_4)
-#29 := (and #26 #28)
-#115 := (implies #29 #114)
-#109 := (<= #67 uf_6)
-#116 := (implies #109 #115)
-#117 := (implies #29 #116)
-#118 := (implies true #117)
-#75 := (= uf_12 uf_10)
-#101 := (implies #75 #100)
-#73 := (= uf_11 uf_5)
-#102 := (implies #73 #101)
-#103 := (implies true #102)
-#71 := (and #28 #28)
-#104 := (implies #71 #103)
-#70 := (= uf_10 #67)
-#105 := (implies #70 #104)
-#68 := (< uf_6 #67)
-#106 := (implies #68 #105)
-#107 := (implies #29 #106)
-#108 := (implies true #107)
-#119 := (and #108 #118)
-#120 := (implies #29 #119)
-#66 := (< uf_5 uf_1)
-#121 := (implies #66 #120)
-#122 := (implies #29 #121)
-#123 := (implies true #122)
-#50 := (<= #20 uf_8)
-#45 := (< #16 uf_1)
-#46 := (and #17 #45)
-#51 := (implies #46 #50)
-#52 := (forall (vars (?x6 int)) #51)
-#53 := (implies #52 true)
-#54 := (and #52 #53)
-#48 := (implies #46 #47)
-#49 := (exists (vars (?x4 int)) #48)
-#55 := (implies #49 #54)
-#56 := (and #49 #55)
-#57 := (implies true #56)
-#44 := (= uf_9 uf_5)
-#58 := (implies #44 #57)
-#42 := (= uf_8 uf_6)
-#59 := (implies #42 #58)
-#40 := (= uf_7 uf_4)
-#60 := (implies #40 #59)
-#61 := (implies true #60)
-#62 := (implies #29 #61)
-#38 := (<= uf_1 uf_5)
-#63 := (implies #38 #62)
-#64 := (implies #29 #63)
-#65 := (implies true #64)
-#124 := (and #65 #123)
-#125 := (implies #29 #124)
-#37 := (= #36 uf_6)
-#126 := (implies #37 #125)
-#33 := (<= #20 uf_6)
-#30 := (< #16 uf_5)
-#31 := (and #17 #30)
-#34 := (implies #31 #33)
-#35 := (forall (vars (?x3 int)) #34)
-#127 := (implies #35 #126)
-#128 := (implies #29 #127)
-#129 := (implies true #128)
-#24 := (= #8 uf_2)
-#130 := (implies #24 #129)
-#131 := (and #24 #130)
-#21 := (<= #20 uf_2)
-#18 := (< #16 1::int)
-#19 := (and #17 #18)
-#22 := (implies #19 #21)
-#23 := (forall (vars (?x1 int)) #22)
-#132 := (implies #23 #131)
-#133 := (and #23 #132)
+#848 := (and true #843)
+#851 := (iff #848 #843)
+#852 := [rewrite]: #851
+#849 := (iff #801 #848)
+#846 := (iff #798 #843)
+#840 := (or #837 #795)
+#844 := (iff #840 #843)
+#845 := [rewrite]: #844
+#841 := (iff #798 #840)
+#838 := (iff #580 #837)
+#835 := (iff #575 #832)
+#829 := (and true #183 #551 #555 #557)
+#833 := (iff #829 #832)
+#834 := [rewrite]: #833
+#830 := (iff #575 #829)
+#826 := (iff #9 true)
+#479 := (<= uf_1 0::int)
+#480 := (not #479)
+#495 := (and #9 #480)
+#500 := (not #495)
+#810 := (or #500 #807)
+#813 := (not #810)
+#107 := (= #106 uf_12)
+#103 := (<= #23 uf_12)
+#101 := (< #19 uf_13)
+#20 := (<= 0::int #19)
+#102 := (and #20 #101)
+#104 := (implies #102 #103)
+#105 := (forall (vars (?x7 int)) #104)
+#108 := (implies #105 #107)
+#109 := (and #105 #108)
+#89 := (<= 2::int uf_13)
+#83 := (<= 0::int uf_11)
+#90 := (and #83 #89)
+#86 := (+ uf_5 1::int)
+#87 := (= uf_13 #86)
+#91 := (and #87 #90)
+#31 := (<= 1::int uf_5)
+#84 := (and #83 #31)
+#92 := (and #84 #91)
+#93 := (and true #92)
+#114 := (and #113 #93)
+#115 := (and #112 #114)
+#29 := (<= 0::int uf_4)
+#32 := (and #29 #31)
+#116 := (and #32 #115)
+#111 := (<= #74 uf_6)
+#117 := (and #111 #116)
+#118 := (and #32 #117)
+#119 := (and true #118)
+#120 := (implies #119 #109)
+#94 := (and #82 #93)
+#95 := (and #80 #94)
+#78 := (and #31 #31)
+#96 := (and #78 #95)
+#97 := (and #77 #96)
+#75 := (< uf_6 #74)
+#98 := (and #75 #97)
+#99 := (and #32 #98)
+#100 := (and true #99)
+#110 := (implies #100 #109)
+#121 := (and #110 #120)
+#70 := (< uf_5 uf_1)
+#71 := (and #70 #32)
+#72 := (and #32 #71)
+#73 := (and true #72)
+#122 := (implies #73 #121)
+#64 := (<= #23 uf_8)
+#59 := (< #19 uf_1)
+#60 := (and #20 #59)
+#65 := (implies #60 #64)
+#66 := (forall (vars (?x6 int)) #65)
+#62 := (implies #60 #61)
+#63 := (exists (vars (?x4 int)) #62)
+#67 := (implies #63 #66)
+#68 := (and #63 #67)
+#52 := (= uf_9 uf_5)
+#53 := (and #50 #52)
+#54 := (and #48 #53)
+#55 := (and #32 #54)
+#46 := (<= uf_1 uf_5)
+#56 := (and #46 #55)
+#57 := (and #32 #56)
+#58 := (and true #57)
+#69 := (implies #58 #68)
+#123 := (and #69 #122)
+#40 := (= #39 uf_6)
+#41 := (and #40 #32)
+#36 := (<= #23 uf_6)
+#33 := (< #19 uf_5)
+#34 := (and #20 #33)
+#37 := (implies #34 #36)
+#38 := (forall (vars (?x3 int)) #37)
+#42 := (and #38 #41)
+#43 := (and #32 #42)
+#44 := (and true #43)
+#27 := (= #8 uf_2)
+#45 := (and #27 #44)
+#124 := (implies #45 #123)
+#125 := (and #27 #124)
+#24 := (<= #23 uf_2)
+#21 := (< #19 1::int)
+#22 := (and #20 #21)
+#25 := (implies #22 #24)
+#26 := (forall (vars (?x1 int)) #25)
+#126 := (implies #26 #125)
+#127 := (and #26 #126)
#12 := (<= 1::int 1::int)
#13 := (and #12 #12)
#10 := (<= 0::int 0::int)
#14 := (and #10 #13)
#15 := (and #10 #14)
-#134 := (implies #15 #133)
-#135 := (implies #9 #134)
-#136 := (implies true #135)
+#16 := (and #9 #15)
#6 := (< 0::int uf_1)
-#137 := (implies #6 #136)
-#138 := (implies true #137)
-#139 := (not #138)
-#1060 := (iff #139 #1057)
-#325 := (not #85)
-#326 := (or #325 #86)
-#329 := (forall (vars (?x8 int)) #326)
-#354 := (not #329)
-#355 := (or #354 #332)
-#360 := (and #329 #355)
-#373 := (not #83)
-#374 := (or #373 #360)
-#319 := (+ 1::int uf_5)
-#322 := (= uf_13 #319)
-#382 := (not #322)
-#383 := (or #382 #374)
-#316 := (and #28 #76)
-#391 := (not #316)
-#392 := (or #391 #383)
-#481 := (or #392 #480)
-#490 := (or #489 #481)
-#275 := (not #29)
-#505 := (or #275 #490)
-#513 := (not #109)
-#514 := (or #513 #505)
-#522 := (or #275 #514)
-#408 := (or #407 #392)
-#417 := (or #416 #408)
-#432 := (not #28)
-#433 := (or #432 #417)
-#442 := (or #441 #433)
-#450 := (not #68)
-#451 := (or #450 #442)
-#459 := (or #275 #451)
-#534 := (and #459 #522)
-#540 := (or #275 #534)
-#548 := (not #66)
-#549 := (or #548 #540)
-#557 := (or #275 #549)
-#192 := (not #46)
-#199 := (or #192 #50)
-#202 := (forall (vars (?x6 int)) #199)
-#193 := (or #192 #47)
-#196 := (exists (vars (?x4 int)) #193)
-#222 := (not #196)
-#223 := (or #222 #202)
-#228 := (and #196 #223)
-#242 := (or #241 #228)
-#251 := (or #250 #242)
-#260 := (or #259 #251)
-#276 := (or #275 #260)
-#284 := (not #38)
-#285 := (or #284 #276)
-#293 := (or #275 #285)
-#569 := (and #293 #557)
-#575 := (or #275 #569)
-#584 := (or #583 #575)
-#173 := (not #31)
-#174 := (or #173 #33)
-#177 := (forall (vars (?x3 int)) #174)
-#592 := (not #177)
-#593 := (or #592 #584)
-#601 := (or #275 #593)
-#617 := (or #616 #601)
-#622 := (and #9 #617)
-#164 := (not #19)
-#165 := (or #164 #21)
-#168 := (forall (vars (?x1 int)) #165)
-#628 := (not #168)
-#629 := (or #628 #622)
-#634 := (and #168 #629)
-#158 := (and #10 #12)
-#161 := (and #10 #158)
-#640 := (not #161)
-#641 := (or #640 #634)
-#649 := (or #616 #641)
-#664 := (not #6)
-#665 := (or #664 #649)
-#677 := (not #665)
-#1058 := (iff #677 #1057)
-#1055 := (iff #665 #1052)
-#1043 := (or false #1029)
-#1046 := (or #616 #1043)
-#1049 := (or #1032 #1046)
-#1053 := (iff #1049 #1052)
-#1054 := [rewrite]: #1053
-#1050 := (iff #665 #1049)
-#1047 := (iff #649 #1046)
-#1044 := (iff #641 #1043)
-#1030 := (iff #634 #1029)
-#1027 := (iff #629 #1026)
-#1024 := (iff #622 #1023)
-#1021 := (iff #617 #1018)
-#1003 := (or #738 #975)
-#1006 := (or #583 #1003)
-#1009 := (or #1000 #1006)
-#1012 := (or #738 #1009)
-#1015 := (or #616 #1012)
-#1019 := (iff #1015 #1018)
-#1020 := [rewrite]: #1019
-#1016 := (iff #617 #1015)
-#1013 := (iff #601 #1012)
-#1010 := (iff #593 #1009)
-#1007 := (iff #584 #1006)
-#1004 := (iff #575 #1003)
-#976 := (iff #569 #975)
-#973 := (iff #557 #970)
-#961 := (or #738 #949)
-#964 := (or #786 #961)
-#967 := (or #738 #964)
-#971 := (iff #967 #970)
-#972 := [rewrite]: #971
-#968 := (iff #557 #967)
-#965 := (iff #549 #964)
-#962 := (iff #540 #961)
-#950 := (iff #534 #949)
-#947 := (iff #522 #944)
-#893 := (or #824 #861)
-#896 := (or #868 #893)
-#899 := (or #874 #896)
-#929 := (or #899 #480)
-#932 := (or #489 #929)
-#935 := (or #738 #932)
-#938 := (or #880 #935)
-#941 := (or #738 #938)
-#945 := (iff #941 #944)
-#946 := [rewrite]: #945
-#942 := (iff #522 #941)
-#939 := (iff #514 #938)
-#936 := (iff #505 #935)
-#933 := (iff #490 #932)
-#930 := (iff #481 #929)
-#900 := (iff #392 #899)
-#897 := (iff #383 #896)
-#894 := (iff #374 #893)
-#862 := (iff #360 #861)
-#859 := (iff #355 #856)
-#853 := (or #850 #332)
-#857 := (iff #853 #856)
-#858 := [rewrite]: #857
-#854 := (iff #355 #853)
-#851 := (iff #354 #850)
-#848 := (iff #329 #847)
-#845 := (iff #326 #844)
-#842 := (iff #86 #841)
-#843 := [rewrite]: #842
-#837 := (iff #325 #836)
-#834 := (iff #85 #833)
-#831 := (iff #84 #830)
-#832 := [rewrite]: #831
-#701 := (iff #17 #703)
-#702 := [rewrite]: #701
-#835 := [monotonicity #702 #832]: #834
-#838 := [monotonicity #835]: #837
-#846 := [monotonicity #838 #843]: #845
-#849 := [quant-intro #846]: #848
-#852 := [monotonicity #849]: #851
-#855 := [monotonicity #852]: #854
-#860 := [trans #855 #858]: #859
-#863 := [monotonicity #849 #860]: #862
-#825 := (iff #373 #824)
-#822 := (iff #83 #821)
-#818 := (iff #82 #819)
-#820 := [rewrite]: #818
-#815 := (iff #76 #816)
-#817 := [rewrite]: #815
-#823 := [monotonicity #817 #820]: #822
-#826 := [monotonicity #823]: #825
-#895 := [monotonicity #826 #863]: #894
-#869 := (iff #382 #868)
-#866 := (iff #322 #864)
-#867 := [rewrite]: #866
-#870 := [monotonicity #867]: #869
-#898 := [monotonicity #870 #895]: #897
-#875 := (iff #391 #874)
-#872 := (iff #316 #871)
-#732 := (iff #28 #733)
-#734 := [rewrite]: #732
-#873 := [monotonicity #734 #817]: #872
-#876 := [monotonicity #873]: #875
-#901 := [monotonicity #876 #898]: #900
-#931 := [monotonicity #901]: #930
-#934 := [monotonicity #931]: #933
-#739 := (iff #275 #738)
-#736 := (iff #29 #735)
-#729 := (iff #26 #730)
-#731 := [rewrite]: #729
-#737 := [monotonicity #731 #734]: #736
-#740 := [monotonicity #737]: #739
-#937 := [monotonicity #740 #934]: #936
-#927 := (iff #513 #880)
-#925 := (iff #109 #881)
-#926 := [rewrite]: #925
-#928 := [monotonicity #926]: #927
-#940 := [monotonicity #928 #937]: #939
-#943 := [monotonicity #740 #940]: #942
-#948 := [trans #943 #946]: #947
-#923 := (iff #459 #920)
-#902 := (or #407 #899)
-#905 := (or #416 #902)
-#908 := (or #877 #905)
-#911 := (or #441 #908)
-#914 := (or #881 #911)
-#917 := (or #738 #914)
-#921 := (iff #917 #920)
-#922 := [rewrite]: #921
-#918 := (iff #459 #917)
-#915 := (iff #451 #914)
-#912 := (iff #442 #911)
-#909 := (iff #433 #908)
-#906 := (iff #417 #905)
-#903 := (iff #408 #902)
-#904 := [monotonicity #901]: #903
-#907 := [monotonicity #904]: #906
-#878 := (iff #432 #877)
-#879 := [monotonicity #734]: #878
-#910 := [monotonicity #879 #907]: #909
-#913 := [monotonicity #910]: #912
-#891 := (iff #450 #881)
-#886 := (not #880)
-#889 := (iff #886 #881)
-#890 := [rewrite]: #889
-#887 := (iff #450 #886)
-#884 := (iff #68 #880)
-#885 := [rewrite]: #884
-#888 := [monotonicity #885]: #887
-#892 := [trans #888 #890]: #891
-#916 := [monotonicity #892 #913]: #915
-#919 := [monotonicity #740 #916]: #918
-#924 := [trans #919 #922]: #923
-#951 := [monotonicity #924 #948]: #950
-#963 := [monotonicity #740 #951]: #962
-#959 := (iff #548 #786)
-#954 := (not #789)
-#957 := (iff #954 #786)
-#958 := [rewrite]: #957
-#955 := (iff #548 #954)
-#952 := (iff #66 #789)
-#953 := [rewrite]: #952
-#956 := [monotonicity #953]: #955
-#960 := [trans #956 #958]: #959
-#966 := [monotonicity #960 #963]: #965
-#969 := [monotonicity #740 #966]: #968
-#974 := [trans #969 #972]: #973
-#813 := (iff #293 #810)
-#792 := (or #241 #781)
-#795 := (or #250 #792)
-#798 := (or #259 #795)
-#801 := (or #738 #798)
-#804 := (or #789 #801)
-#807 := (or #738 #804)
-#811 := (iff #807 #810)
-#812 := [rewrite]: #811
-#808 := (iff #293 #807)
-#805 := (iff #285 #804)
-#802 := (iff #276 #801)
-#799 := (iff #260 #798)
-#796 := (iff #251 #795)
-#793 := (iff #242 #792)
-#782 := (iff #228 #781)
-#779 := (iff #223 #778)
-#776 := (iff #202 #775)
-#773 := (iff #199 #772)
-#770 := (iff #50 #769)
-#771 := [rewrite]: #770
-#751 := (iff #192 #750)
-#748 := (iff #46 #747)
-#745 := (iff #45 #744)
-#746 := [rewrite]: #745
-#749 := [monotonicity #702 #746]: #748
-#752 := [monotonicity #749]: #751
-#774 := [monotonicity #752 #771]: #773
-#777 := [quant-intro #774]: #776
-#765 := (iff #222 #764)
-#762 := (iff #196 #761)
-#759 := (iff #193 #756)
-#753 := (or #750 #47)
-#757 := (iff #753 #756)
-#758 := [rewrite]: #757
-#754 := (iff #193 #753)
-#755 := [monotonicity #752]: #754
-#760 := [trans #755 #758]: #759
-#763 := [quant-intro #760]: #762
-#766 := [monotonicity #763]: #765
-#780 := [monotonicity #766 #777]: #779
-#783 := [monotonicity #763 #780]: #782
-#794 := [monotonicity #783]: #793
-#797 := [monotonicity #794]: #796
-#800 := [monotonicity #797]: #799
-#803 := [monotonicity #740 #800]: #802
-#790 := (iff #284 #789)
-#787 := (iff #38 #786)
-#788 := [rewrite]: #787
-#791 := [monotonicity #788]: #790
-#806 := [monotonicity #791 #803]: #805
-#809 := [monotonicity #740 #806]: #808
-#814 := [trans #809 #812]: #813
-#977 := [monotonicity #814 #974]: #976
-#1005 := [monotonicity #740 #977]: #1004
-#1008 := [monotonicity #1005]: #1007
-#1001 := (iff #592 #1000)
-#998 := (iff #177 #997)
-#995 := (iff #174 #994)
-#992 := (iff #33 #991)
-#993 := [rewrite]: #992
-#987 := (iff #173 #986)
-#984 := (iff #31 #983)
-#981 := (iff #30 #980)
-#982 := [rewrite]: #981
-#985 := [monotonicity #702 #982]: #984
-#988 := [monotonicity #985]: #987
-#996 := [monotonicity #988 #993]: #995
-#999 := [quant-intro #996]: #998
-#1002 := [monotonicity #999]: #1001
-#1011 := [monotonicity #1002 #1008]: #1010
-#1014 := [monotonicity #740 #1011]: #1013
-#1017 := [monotonicity #1014]: #1016
-#1022 := [trans #1017 #1020]: #1021
-#1025 := [monotonicity #1022]: #1024
-#727 := (iff #628 #726)
-#724 := (iff #168 #723)
-#721 := (iff #165 #720)
-#715 := (iff #21 #716)
-#719 := [rewrite]: #715
-#712 := (iff #164 #711)
-#709 := (iff #19 #708)
-#705 := (iff #18 #704)
-#707 := [rewrite]: #705
-#710 := [monotonicity #702 #707]: #709
-#713 := [monotonicity #710]: #712
-#722 := [monotonicity #713 #719]: #721
-#725 := [quant-intro #722]: #724
-#728 := [monotonicity #725]: #727
-#1028 := [monotonicity #728 #1025]: #1027
-#1031 := [monotonicity #725 #1028]: #1030
-#699 := (iff #640 false)
-#694 := (not true)
-#697 := (iff #694 false)
-#698 := [rewrite]: #697
-#695 := (iff #640 #694)
-#692 := (iff #161 true)
-#684 := (and true true)
-#687 := (and true #684)
-#690 := (iff #687 true)
-#691 := [rewrite]: #690
-#688 := (iff #161 #687)
-#685 := (iff #158 #684)
-#682 := (iff #12 true)
+#17 := (and #6 #16)
+#18 := (and true #17)
+#128 := (implies #18 #127)
+#129 := (not #128)
+#816 := (iff #129 #813)
+#302 := (+ 1::int uf_5)
+#305 := (= uf_13 #302)
+#311 := (and #90 #305)
+#299 := (and #31 #83)
+#316 := (and #299 #311)
+#391 := (and #316 #385)
+#396 := (and #382 #391)
+#399 := (and #32 #396)
+#402 := (and #111 #399)
+#405 := (and #32 #402)
+#418 := (not #405)
+#351 := (not #102)
+#352 := (or #351 #103)
+#355 := (forall (vars (?x7 int)) #352)
+#364 := (not #355)
+#365 := (or #364 #358)
+#370 := (and #355 #365)
+#419 := (or #370 #418)
+#326 := (and #296 #316)
+#329 := (and #293 #326)
+#332 := (and #31 #329)
+#335 := (and #288 #332)
+#338 := (and #75 #335)
+#341 := (and #32 #338)
+#376 := (not #341)
+#377 := (or #376 #370)
+#424 := (and #377 #419)
+#275 := (and #32 #70)
+#278 := (and #32 #275)
+#430 := (not #278)
+#431 := (or #430 #424)
+#241 := (not #60)
+#248 := (or #241 #64)
+#251 := (forall (vars (?x6 int)) #248)
+#242 := (or #241 #61)
+#245 := (exists (vars (?x4 int)) #242)
+#257 := (not #245)
+#258 := (or #257 #251)
+#263 := (and #245 #258)
+#219 := (and #213 #216)
+#222 := (and #210 #219)
+#225 := (and #32 #222)
+#228 := (and #46 #225)
+#231 := (and #32 #228)
+#269 := (not #231)
+#270 := (or #269 #263)
+#436 := (and #270 #431)
+#189 := (and #32 #183)
+#176 := (not #34)
+#177 := (or #176 #36)
+#180 := (forall (vars (?x3 int)) #177)
+#194 := (and #180 #189)
+#197 := (and #32 #194)
+#207 := (and #9 #197)
+#442 := (not #207)
+#443 := (or #442 #436)
+#448 := (and #9 #443)
+#167 := (not #22)
+#168 := (or #167 #24)
+#171 := (forall (vars (?x1 int)) #168)
+#454 := (not #171)
+#455 := (or #454 #448)
+#460 := (and #171 #455)
+#148 := (and #10 #12)
+#151 := (and #10 #148)
+#154 := (and #9 #151)
+#157 := (and #6 #154)
+#466 := (not #157)
+#467 := (or #466 #460)
+#472 := (not #467)
+#814 := (iff #472 #813)
+#811 := (iff #467 #810)
+#808 := (iff #460 #807)
+#805 := (iff #455 #804)
+#802 := (iff #448 #801)
+#799 := (iff #443 #798)
+#796 := (iff #436 #795)
+#793 := (iff #431 #792)
+#790 := (iff #424 #789)
+#787 := (iff #419 #786)
+#784 := (iff #418 #783)
+#781 := (iff #405 #778)
+#684 := (and #668 #671)
+#687 := (and #684 #673)
+#690 := (and #557 #668)
+#693 := (and #690 #687)
+#763 := (and #693 #385)
+#766 := (and #382 #763)
+#560 := (and #555 #557)
+#769 := (and #560 #766)
+#772 := (and #678 #769)
+#775 := (and #560 #772)
+#779 := (iff #775 #778)
+#780 := [rewrite]: #779
+#776 := (iff #405 #775)
+#773 := (iff #402 #772)
+#770 := (iff #399 #769)
+#767 := (iff #396 #766)
+#764 := (iff #391 #763)
+#694 := (iff #316 #693)
+#688 := (iff #311 #687)
+#676 := (iff #305 #673)
+#677 := [rewrite]: #676
+#685 := (iff #90 #684)
+#670 := (iff #89 #671)
+#672 := [rewrite]: #670
+#667 := (iff #83 #668)
+#669 := [rewrite]: #667
+#686 := [monotonicity #669 #672]: #685
+#689 := [monotonicity #686 #677]: #688
+#691 := (iff #299 #690)
+#558 := (iff #31 #557)
+#559 := [rewrite]: #558
+#692 := [monotonicity #559 #669]: #691
+#695 := [monotonicity #692 #689]: #694
+#765 := [monotonicity #695]: #764
+#768 := [monotonicity #765]: #767
+#561 := (iff #32 #560)
+#554 := (iff #29 #555)
+#556 := [rewrite]: #554
+#562 := [monotonicity #556 #559]: #561
+#771 := [monotonicity #562 #768]: #770
+#761 := (iff #111 #678)
+#762 := [rewrite]: #761
+#774 := [monotonicity #762 #771]: #773
+#777 := [monotonicity #562 #774]: #776
+#782 := [trans #777 #780]: #781
+#785 := [monotonicity #782]: #784
+#756 := (iff #370 #755)
+#753 := (iff #365 #750)
+#747 := (or #744 #358)
+#751 := (iff #747 #750)
+#752 := [rewrite]: #751
+#748 := (iff #365 #747)
+#745 := (iff #364 #744)
+#742 := (iff #355 #741)
+#739 := (iff #352 #738)
+#736 := (iff #103 #735)
+#737 := [rewrite]: #736
+#731 := (iff #351 #730)
+#728 := (iff #102 #727)
+#725 := (iff #101 #724)
+#726 := [rewrite]: #725
+#503 := (iff #20 #505)
+#504 := [rewrite]: #503
+#729 := [monotonicity #504 #726]: #728
+#732 := [monotonicity #729]: #731
+#740 := [monotonicity #732 #737]: #739
+#743 := [quant-intro #740]: #742
+#746 := [monotonicity #743]: #745
+#749 := [monotonicity #746]: #748
+#754 := [trans #749 #752]: #753
+#757 := [monotonicity #743 #754]: #756
+#788 := [monotonicity #757 #785]: #787
+#759 := (iff #377 #758)
+#720 := (iff #376 #719)
+#717 := (iff #341 #714)
+#696 := (and #296 #693)
+#699 := (and #293 #696)
+#702 := (and #557 #699)
+#705 := (and #288 #702)
+#708 := (and #681 #705)
+#711 := (and #560 #708)
+#715 := (iff #711 #714)
+#716 := [rewrite]: #715
+#712 := (iff #341 #711)
+#709 := (iff #338 #708)
+#706 := (iff #335 #705)
+#703 := (iff #332 #702)
+#700 := (iff #329 #699)
+#697 := (iff #326 #696)
+#698 := [monotonicity #695]: #697
+#701 := [monotonicity #698]: #700
+#704 := [monotonicity #559 #701]: #703
+#707 := [monotonicity #704]: #706
+#682 := (iff #75 #681)
#683 := [rewrite]: #682
-#680 := (iff #10 true)
-#681 := [rewrite]: #680
-#686 := [monotonicity #681 #683]: #685
-#689 := [monotonicity #681 #686]: #688
-#693 := [trans #689 #691]: #692
-#696 := [monotonicity #693]: #695
-#700 := [trans #696 #698]: #699
-#1045 := [monotonicity #700 #1031]: #1044
-#1048 := [monotonicity #1045]: #1047
-#1041 := (iff #664 #1032)
-#1033 := (not #1032)
-#1036 := (not #1033)
-#1039 := (iff #1036 #1032)
-#1040 := [rewrite]: #1039
-#1037 := (iff #664 #1036)
-#1034 := (iff #6 #1033)
-#1035 := [rewrite]: #1034
-#1038 := [monotonicity #1035]: #1037
-#1042 := [trans #1038 #1040]: #1041
-#1051 := [monotonicity #1042 #1048]: #1050
-#1056 := [trans #1051 #1054]: #1055
-#1059 := [monotonicity #1056]: #1058
-#678 := (iff #139 #677)
-#675 := (iff #138 #665)
-#670 := (implies true #665)
-#673 := (iff #670 #665)
-#674 := [rewrite]: #673
-#671 := (iff #138 #670)
-#668 := (iff #137 #665)
-#661 := (implies #6 #649)
-#666 := (iff #661 #665)
-#667 := [rewrite]: #666
-#662 := (iff #137 #661)
-#659 := (iff #136 #649)
-#654 := (implies true #649)
-#657 := (iff #654 #649)
-#658 := [rewrite]: #657
-#655 := (iff #136 #654)
-#652 := (iff #135 #649)
-#646 := (implies #9 #641)
-#650 := (iff #646 #649)
-#651 := [rewrite]: #650
-#647 := (iff #135 #646)
-#644 := (iff #134 #641)
-#637 := (implies #161 #634)
-#642 := (iff #637 #641)
-#643 := [rewrite]: #642
-#638 := (iff #134 #637)
-#635 := (iff #133 #634)
-#632 := (iff #132 #629)
-#625 := (implies #168 #622)
-#630 := (iff #625 #629)
-#631 := [rewrite]: #630
-#626 := (iff #132 #625)
-#623 := (iff #131 #622)
-#620 := (iff #130 #617)
-#613 := (implies #9 #601)
-#618 := (iff #613 #617)
-#619 := [rewrite]: #618
-#614 := (iff #130 #613)
-#611 := (iff #129 #601)
-#606 := (implies true #601)
-#609 := (iff #606 #601)
-#610 := [rewrite]: #609
-#607 := (iff #129 #606)
-#604 := (iff #128 #601)
-#598 := (implies #29 #593)
-#602 := (iff #598 #601)
-#603 := [rewrite]: #602
-#599 := (iff #128 #598)
-#596 := (iff #127 #593)
-#589 := (implies #177 #584)
-#594 := (iff #589 #593)
-#595 := [rewrite]: #594
-#590 := (iff #127 #589)
-#587 := (iff #126 #584)
-#580 := (implies #180 #575)
-#585 := (iff #580 #584)
+#710 := [monotonicity #683 #707]: #709
+#713 := [monotonicity #562 #710]: #712
+#718 := [trans #713 #716]: #717
+#721 := [monotonicity #718]: #720
+#760 := [monotonicity #721 #757]: #759
+#791 := [monotonicity #760 #788]: #790
+#665 := (iff #430 #664)
+#662 := (iff #278 #659)
+#653 := (and #560 #650)
+#656 := (and #560 #653)
+#660 := (iff #656 #659)
+#661 := [rewrite]: #660
+#657 := (iff #278 #656)
+#654 := (iff #275 #653)
+#651 := (iff #70 #650)
+#652 := [rewrite]: #651
+#655 := [monotonicity #562 #652]: #654
+#658 := [monotonicity #562 #655]: #657
+#663 := [trans #658 #661]: #662
+#666 := [monotonicity #663]: #665
+#794 := [monotonicity #666 #791]: #793
+#648 := (iff #270 #647)
+#645 := (iff #263 #644)
+#642 := (iff #258 #641)
+#639 := (iff #251 #638)
+#636 := (iff #248 #635)
+#633 := (iff #64 #632)
+#634 := [rewrite]: #633
+#614 := (iff #241 #613)
+#611 := (iff #60 #610)
+#608 := (iff #59 #607)
+#609 := [rewrite]: #608
+#612 := [monotonicity #504 #609]: #611
+#615 := [monotonicity #612]: #614
+#637 := [monotonicity #615 #634]: #636
+#640 := [quant-intro #637]: #639
+#628 := (iff #257 #627)
+#625 := (iff #245 #624)
+#622 := (iff #242 #619)
+#616 := (or #613 #61)
+#620 := (iff #616 #619)
+#621 := [rewrite]: #620
+#617 := (iff #242 #616)
+#618 := [monotonicity #615]: #617
+#623 := [trans #618 #621]: #622
+#626 := [quant-intro #623]: #625
+#629 := [monotonicity #626]: #628
+#643 := [monotonicity #629 #640]: #642
+#646 := [monotonicity #626 #643]: #645
+#602 := (iff #269 #601)
+#599 := (iff #231 #596)
+#587 := (and #560 #222)
+#590 := (and #584 #587)
+#593 := (and #560 #590)
+#597 := (iff #593 #596)
+#598 := [rewrite]: #597
+#594 := (iff #231 #593)
+#591 := (iff #228 #590)
+#588 := (iff #225 #587)
+#589 := [monotonicity #562]: #588
+#585 := (iff #46 #584)
#586 := [rewrite]: #585
-#581 := (iff #126 #580)
-#578 := (iff #125 #575)
-#572 := (implies #29 #569)
+#592 := [monotonicity #586 #589]: #591
+#595 := [monotonicity #562 #592]: #594
+#600 := [trans #595 #598]: #599
+#603 := [monotonicity #600]: #602
+#649 := [monotonicity #603 #646]: #648
+#797 := [monotonicity #649 #794]: #796
+#581 := (iff #442 #580)
+#578 := (iff #207 #575)
+#563 := (and #560 #183)
+#566 := (and #551 #563)
+#569 := (and #560 #566)
+#572 := (and #9 #569)
#576 := (iff #572 #575)
#577 := [rewrite]: #576
-#573 := (iff #125 #572)
-#570 := (iff #124 #569)
-#567 := (iff #123 #557)
-#562 := (implies true #557)
-#565 := (iff #562 #557)
-#566 := [rewrite]: #565
-#563 := (iff #123 #562)
-#560 := (iff #122 #557)
-#554 := (implies #29 #549)
-#558 := (iff #554 #557)
-#559 := [rewrite]: #558
-#555 := (iff #122 #554)
-#552 := (iff #121 #549)
-#545 := (implies #66 #540)
-#550 := (iff #545 #549)
-#551 := [rewrite]: #550
-#546 := (iff #121 #545)
-#543 := (iff #120 #540)
-#537 := (implies #29 #534)
-#541 := (iff #537 #540)
-#542 := [rewrite]: #541
-#538 := (iff #120 #537)
-#535 := (iff #119 #534)
-#532 := (iff #118 #522)
-#527 := (implies true #522)
-#530 := (iff #527 #522)
-#531 := [rewrite]: #530
-#528 := (iff #118 #527)
-#525 := (iff #117 #522)
-#519 := (implies #29 #514)
-#523 := (iff #519 #522)
-#524 := [rewrite]: #523
-#520 := (iff #117 #519)
-#517 := (iff #116 #514)
-#510 := (implies #109 #505)
-#515 := (iff #510 #514)
-#516 := [rewrite]: #515
-#511 := (iff #116 #510)
-#508 := (iff #115 #505)
-#502 := (implies #29 #490)
-#506 := (iff #502 #505)
-#507 := [rewrite]: #506
-#503 := (iff #115 #502)
-#500 := (iff #114 #490)
-#495 := (implies true #490)
-#498 := (iff #495 #490)
-#499 := [rewrite]: #498
-#496 := (iff #114 #495)
-#493 := (iff #113 #490)
-#486 := (implies #471 #481)
-#491 := (iff #486 #490)
-#492 := [rewrite]: #491
-#487 := (iff #113 #486)
-#484 := (iff #112 #481)
-#477 := (implies #474 #392)
-#482 := (iff #477 #481)
-#483 := [rewrite]: #482
-#478 := (iff #112 #477)
-#402 := (iff #100 #392)
-#397 := (implies true #392)
-#400 := (iff #397 #392)
-#401 := [rewrite]: #400
-#398 := (iff #100 #397)
-#395 := (iff #99 #392)
-#388 := (implies #316 #383)
-#393 := (iff #388 #392)
-#394 := [rewrite]: #393
-#389 := (iff #99 #388)
-#386 := (iff #98 #383)
-#379 := (implies #322 #374)
-#384 := (iff #379 #383)
-#385 := [rewrite]: #384
-#380 := (iff #98 #379)
-#377 := (iff #97 #374)
-#370 := (implies #83 #360)
-#375 := (iff #370 #374)
-#376 := [rewrite]: #375
-#371 := (iff #97 #370)
-#368 := (iff #96 #360)
-#363 := (implies true #360)
-#366 := (iff #363 #360)
-#367 := [rewrite]: #366
-#364 := (iff #96 #363)
-#361 := (iff #95 #360)
-#358 := (iff #94 #355)
-#351 := (implies #329 #332)
-#356 := (iff #351 #355)
-#357 := [rewrite]: #356
-#352 := (iff #94 #351)
-#349 := (iff #93 #332)
-#344 := (and #332 true)
-#347 := (iff #344 #332)
-#348 := [rewrite]: #347
-#345 := (iff #93 #344)
-#342 := (iff #92 true)
-#337 := (implies #332 true)
-#340 := (iff #337 true)
-#341 := [rewrite]: #340
-#338 := (iff #92 #337)
-#335 := (iff #91 true)
-#336 := [rewrite]: #335
-#333 := (iff #90 #332)
-#334 := [rewrite]: #333
-#339 := [monotonicity #334 #336]: #338
-#343 := [trans #339 #341]: #342
-#346 := [monotonicity #334 #343]: #345
-#350 := [trans #346 #348]: #349
-#330 := (iff #88 #329)
-#327 := (iff #87 #326)
-#328 := [rewrite]: #327
-#331 := [quant-intro #328]: #330
-#353 := [monotonicity #331 #350]: #352
-#359 := [trans #353 #357]: #358
-#362 := [monotonicity #331 #359]: #361
-#365 := [monotonicity #362]: #364
-#369 := [trans #365 #367]: #368
-#372 := [monotonicity #369]: #371
-#378 := [trans #372 #376]: #377
-#323 := (iff #80 #322)
-#320 := (= #79 #319)
-#321 := [rewrite]: #320
-#324 := [monotonicity #321]: #323
-#381 := [monotonicity #324 #378]: #380
-#387 := [trans #381 #385]: #386
-#317 := (iff #77 #316)
-#318 := [rewrite]: #317
-#390 := [monotonicity #318 #387]: #389
-#396 := [trans #390 #394]: #395
-#399 := [monotonicity #396]: #398
-#403 := [trans #399 #401]: #402
-#475 := (iff #111 #474)
-#476 := [rewrite]: #475
-#479 := [monotonicity #476 #403]: #478
-#485 := [trans #479 #483]: #484
-#472 := (iff #110 #471)
-#473 := [rewrite]: #472
-#488 := [monotonicity #473 #485]: #487
-#494 := [trans #488 #492]: #493
-#497 := [monotonicity #494]: #496
-#501 := [trans #497 #499]: #500
-#504 := [monotonicity #501]: #503
-#509 := [trans #504 #507]: #508
-#512 := [monotonicity #509]: #511
-#518 := [trans #512 #516]: #517
-#521 := [monotonicity #518]: #520
-#526 := [trans #521 #524]: #525
-#529 := [monotonicity #526]: #528
-#533 := [trans #529 #531]: #532
-#469 := (iff #108 #459)
-#464 := (implies true #459)
-#467 := (iff #464 #459)
-#468 := [rewrite]: #467
-#465 := (iff #108 #464)
-#462 := (iff #107 #459)
-#456 := (implies #29 #451)
-#460 := (iff #456 #459)
-#461 := [rewrite]: #460
-#457 := (iff #107 #456)
-#454 := (iff #106 #451)
-#447 := (implies #68 #442)
-#452 := (iff #447 #451)
-#453 := [rewrite]: #452
-#448 := (iff #106 #447)
-#445 := (iff #105 #442)
-#438 := (implies #305 #433)
-#443 := (iff #438 #442)
-#444 := [rewrite]: #443
-#439 := (iff #105 #438)
-#436 := (iff #104 #433)
-#429 := (implies #28 #417)
-#434 := (iff #429 #433)
-#435 := [rewrite]: #434
-#430 := (iff #104 #429)
-#427 := (iff #103 #417)
-#422 := (implies true #417)
-#425 := (iff #422 #417)
-#426 := [rewrite]: #425
-#423 := (iff #103 #422)
-#420 := (iff #102 #417)
-#413 := (implies #310 #408)
-#418 := (iff #413 #417)
-#419 := [rewrite]: #418
-#414 := (iff #102 #413)
-#411 := (iff #101 #408)
-#404 := (implies #313 #392)
-#409 := (iff #404 #408)
-#410 := [rewrite]: #409
-#405 := (iff #101 #404)
-#314 := (iff #75 #313)
-#315 := [rewrite]: #314
-#406 := [monotonicity #315 #403]: #405
-#412 := [trans #406 #410]: #411
-#311 := (iff #73 #310)
-#312 := [rewrite]: #311
-#415 := [monotonicity #312 #412]: #414
-#421 := [trans #415 #419]: #420
-#424 := [monotonicity #421]: #423
-#428 := [trans #424 #426]: #427
-#308 := (iff #71 #28)
-#309 := [rewrite]: #308
-#431 := [monotonicity #309 #428]: #430
-#437 := [trans #431 #435]: #436
-#306 := (iff #70 #305)
-#307 := [rewrite]: #306
-#440 := [monotonicity #307 #437]: #439
-#446 := [trans #440 #444]: #445
-#449 := [monotonicity #446]: #448
-#455 := [trans #449 #453]: #454
-#458 := [monotonicity #455]: #457
-#463 := [trans #458 #461]: #462
-#466 := [monotonicity #463]: #465
-#470 := [trans #466 #468]: #469
-#536 := [monotonicity #470 #533]: #535
-#539 := [monotonicity #536]: #538
-#544 := [trans #539 #542]: #543
-#547 := [monotonicity #544]: #546
-#553 := [trans #547 #551]: #552
-#556 := [monotonicity #553]: #555
-#561 := [trans #556 #559]: #560
-#564 := [monotonicity #561]: #563
-#568 := [trans #564 #566]: #567
-#303 := (iff #65 #293)
-#298 := (implies true #293)
-#301 := (iff #298 #293)
-#302 := [rewrite]: #301
-#299 := (iff #65 #298)
-#296 := (iff #64 #293)
-#290 := (implies #29 #285)
-#294 := (iff #290 #293)
-#295 := [rewrite]: #294
-#291 := (iff #64 #290)
-#288 := (iff #63 #285)
-#281 := (implies #38 #276)
-#286 := (iff #281 #285)
-#287 := [rewrite]: #286
-#282 := (iff #63 #281)
-#279 := (iff #62 #276)
-#272 := (implies #29 #260)
-#277 := (iff #272 #276)
-#278 := [rewrite]: #277
-#273 := (iff #62 #272)
-#270 := (iff #61 #260)
-#265 := (implies true #260)
-#268 := (iff #265 #260)
-#269 := [rewrite]: #268
-#266 := (iff #61 #265)
-#263 := (iff #60 #260)
-#256 := (implies #183 #251)
-#261 := (iff #256 #260)
-#262 := [rewrite]: #261
-#257 := (iff #60 #256)
-#254 := (iff #59 #251)
-#247 := (implies #186 #242)
-#252 := (iff #247 #251)
-#253 := [rewrite]: #252
-#248 := (iff #59 #247)
-#245 := (iff #58 #242)
-#238 := (implies #189 #228)
-#243 := (iff #238 #242)
-#244 := [rewrite]: #243
-#239 := (iff #58 #238)
-#236 := (iff #57 #228)
-#231 := (implies true #228)
-#234 := (iff #231 #228)
-#235 := [rewrite]: #234
-#232 := (iff #57 #231)
-#229 := (iff #56 #228)
-#226 := (iff #55 #223)
-#219 := (implies #196 #202)
-#224 := (iff #219 #223)
-#225 := [rewrite]: #224
-#220 := (iff #55 #219)
-#217 := (iff #54 #202)
-#212 := (and #202 true)
-#215 := (iff #212 #202)
-#216 := [rewrite]: #215
-#213 := (iff #54 #212)
-#210 := (iff #53 true)
-#205 := (implies #202 true)
-#208 := (iff #205 true)
-#209 := [rewrite]: #208
-#206 := (iff #53 #205)
-#203 := (iff #52 #202)
-#200 := (iff #51 #199)
-#201 := [rewrite]: #200
-#204 := [quant-intro #201]: #203
-#207 := [monotonicity #204]: #206
-#211 := [trans #207 #209]: #210
-#214 := [monotonicity #204 #211]: #213
-#218 := [trans #214 #216]: #217
-#197 := (iff #49 #196)
-#194 := (iff #48 #193)
-#195 := [rewrite]: #194
-#198 := [quant-intro #195]: #197
-#221 := [monotonicity #198 #218]: #220
-#227 := [trans #221 #225]: #226
-#230 := [monotonicity #198 #227]: #229
-#233 := [monotonicity #230]: #232
-#237 := [trans #233 #235]: #236
-#190 := (iff #44 #189)
-#191 := [rewrite]: #190
-#240 := [monotonicity #191 #237]: #239
-#246 := [trans #240 #244]: #245
-#187 := (iff #42 #186)
-#188 := [rewrite]: #187
-#249 := [monotonicity #188 #246]: #248
-#255 := [trans #249 #253]: #254
-#184 := (iff #40 #183)
-#185 := [rewrite]: #184
-#258 := [monotonicity #185 #255]: #257
-#264 := [trans #258 #262]: #263
-#267 := [monotonicity #264]: #266
-#271 := [trans #267 #269]: #270
-#274 := [monotonicity #271]: #273
-#280 := [trans #274 #278]: #279
-#283 := [monotonicity #280]: #282
-#289 := [trans #283 #287]: #288
-#292 := [monotonicity #289]: #291
-#297 := [trans #292 #295]: #296
-#300 := [monotonicity #297]: #299
-#304 := [trans #300 #302]: #303
-#571 := [monotonicity #304 #568]: #570
+#573 := (iff #207 #572)
+#570 := (iff #197 #569)
+#567 := (iff #194 #566)
+#564 := (iff #189 #563)
+#565 := [monotonicity #562]: #564
+#552 := (iff #180 #551)
+#549 := (iff #177 #548)
+#546 := (iff #36 #545)
+#547 := [rewrite]: #546
+#541 := (iff #176 #540)
+#538 := (iff #34 #537)
+#535 := (iff #33 #531)
+#536 := [rewrite]: #535
+#539 := [monotonicity #504 #536]: #538
+#542 := [monotonicity #539]: #541
+#550 := [monotonicity #542 #547]: #549
+#553 := [quant-intro #550]: #552
+#568 := [monotonicity #553 #565]: #567
+#571 := [monotonicity #562 #568]: #570
#574 := [monotonicity #571]: #573
#579 := [trans #574 #577]: #578
-#181 := (iff #37 #180)
-#182 := [rewrite]: #181
-#582 := [monotonicity #182 #579]: #581
-#588 := [trans #582 #586]: #587
-#178 := (iff #35 #177)
-#175 := (iff #34 #174)
-#176 := [rewrite]: #175
-#179 := [quant-intro #176]: #178
-#591 := [monotonicity #179 #588]: #590
-#597 := [trans #591 #595]: #596
-#600 := [monotonicity #597]: #599
-#605 := [trans #600 #603]: #604
-#608 := [monotonicity #605]: #607
-#612 := [trans #608 #610]: #611
-#171 := (iff #24 #9)
-#172 := [rewrite]: #171
-#615 := [monotonicity #172 #612]: #614
-#621 := [trans #615 #619]: #620
-#624 := [monotonicity #172 #621]: #623
-#169 := (iff #23 #168)
-#166 := (iff #22 #165)
-#167 := [rewrite]: #166
-#170 := [quant-intro #167]: #169
-#627 := [monotonicity #170 #624]: #626
-#633 := [trans #627 #631]: #632
-#636 := [monotonicity #170 #633]: #635
-#162 := (iff #15 #161)
-#159 := (iff #14 #158)
-#156 := (iff #13 #12)
-#157 := [rewrite]: #156
-#160 := [monotonicity #157]: #159
-#163 := [monotonicity #160]: #162
-#639 := [monotonicity #163 #636]: #638
-#645 := [trans #639 #643]: #644
-#648 := [monotonicity #645]: #647
-#653 := [trans #648 #651]: #652
-#656 := [monotonicity #653]: #655
-#660 := [trans #656 #658]: #659
-#663 := [monotonicity #660]: #662
-#669 := [trans #663 #667]: #668
-#672 := [monotonicity #669]: #671
-#676 := [trans #672 #674]: #675
-#679 := [monotonicity #676]: #678
-#1061 := [trans #679 #1059]: #1060
-#155 := [asserted]: #139
-#1062 := [mp #155 #1061]: #1057
-#1063 := [not-or-elim #1062]: #9
-#2109 := (or #616 #2118)
-#2133 := [th-lemma]: #2109
-#2110 := [unit-resolution #2133 #1063]: #2118
-decl ?x1!0 :: int
-#1106 := ?x1!0
-#1107 := (uf_3 ?x1!0)
-#1104 := (* -1::int #1107)
-#1105 := (+ uf_2 #1104)
-#1108 := (>= #1105 0::int)
-#1847 := (not #1108)
-#1111 := (>= ?x1!0 0::int)
-#1229 := (not #1111)
-#1109 := (>= ?x1!0 1::int)
-#1494 := (or #1108 #1109 #1229)
-#1499 := (not #1494)
-decl ?x4!1 :: int
-#1148 := ?x4!1
-#1156 := (uf_3 ?x4!1)
-#1329 := (= uf_8 #1156)
-#1153 := (>= ?x4!1 0::int)
-#1572 := (not #1153)
-#1149 := (* -1::int ?x4!1)
-#1150 := (+ uf_1 #1149)
-#1151 := (<= #1150 0::int)
-#1587 := (or #1151 #1572 #1329)
-#1618 := (not #1587)
-decl ?x6!2 :: int
-#1166 := ?x6!2
-#1167 := (uf_3 ?x6!2)
-#1353 := (* -1::int #1167)
-#1354 := (+ uf_8 #1353)
-#1355 := (>= #1354 0::int)
-#1174 := (>= ?x6!2 0::int)
-#1592 := (not #1174)
-#1170 := (* -1::int ?x6!2)
-#1171 := (+ uf_1 #1170)
-#1172 := (<= #1171 0::int)
-#1749 := (or #1172 #1592 #1355 #1618)
-#1752 := (not #1749)
-#2262 := (pattern #20)
-#1502 := (not #703)
-#1561 := (or #47 #1502 #743)
-#1566 := (not #1561)
-#2323 := (forall (vars (?x4 int)) (:pat #2262) #1566)
-#2328 := (or #2323 #1752)
-#2331 := (not #2328)
-#1631 := (not #730)
-#2334 := (or #259 #250 #241 #1631 #877 #789 #2331)
-#2337 := (not #2334)
-decl ?x8!3 :: int
-#1215 := ?x8!3
-#1216 := (uf_3 ?x8!3)
-#1418 := (* -1::int #1216)
-#1419 := (+ uf_12 #1418)
-#1420 := (>= #1419 0::int)
-#1396 := (* -1::int ?x8!3)
-#1397 := (+ uf_13 #1396)
-#1398 := (<= #1397 0::int)
-#1222 := (>= ?x8!3 0::int)
-#1671 := (not #1222)
-#1686 := (or #1671 #1398 #1420)
-#1691 := (not #1686)
-#1653 := (or #1502 #827 #841)
-#2279 := (forall (vars (?x8 int)) (:pat #2262) #1653)
-#2284 := (not #2279)
-#2287 := (or #332 #2284)
-#2290 := (not #2287)
-#2293 := (or #2290 #1691)
-#2296 := (not #2293)
-#1701 := (not #819)
-#1700 := (not #816)
-#2305 := (or #489 #480 #1631 #877 #1700 #1701 #868 #880 #2296)
-#2308 := (not #2305)
-#2299 := (or #441 #416 #407 #1631 #877 #1700 #1701 #868 #881 #2296)
-#2302 := (not #2299)
-#2311 := (or #2302 #2308)
-#2314 := (not #2311)
-#2317 := (or #1631 #877 #786 #2314)
-#2320 := (not #2317)
-#2340 := (or #2320 #2337)
-#2343 := (not #2340)
-#1539 := (or #1502 #978 #991)
-#2271 := (forall (vars (?x3 int)) (:pat #2262) #1539)
-#2276 := (not #2271)
-#1517 := (or #1502 #706 #716)
-#2263 := (forall (vars (?x1 int)) (:pat #2262) #1517)
-#2268 := (not #2263)
-#2346 := (or #583 #1631 #877 #2268 #2276 #2343)
-#1403 := (not #1398)
-#2349 := (not #2346)
-#2592 := [hypothesis]: #2349
-#2176 := (or #2346 #180)
-#2183 := [def-axiom]: #2176
-#2593 := [unit-resolution #2183 #2592]: #180
-#2160 := (or #2346 #2340)
-#2161 := [def-axiom]: #2160
-#2594 := [unit-resolution #2161 #2592]: #2340
-#2169 := (or #2346 #2271)
-#2174 := [def-axiom]: #2169
-#2595 := [unit-resolution #2174 #2592]: #2271
-#2414 := (or #2334 #583 #2276)
-#1851 := (uf_3 uf_7)
-#2358 := (= uf_8 #1851)
-#2408 := (= #36 #1851)
-#2406 := (= #1851 #36)
-#2391 := [hypothesis]: #2337
-#2099 := (or #2334 #183)
-#2100 := [def-axiom]: #2099
-#2402 := [unit-resolution #2100 #2391]: #183
-#2403 := [symm #2402]: #40
-#2407 := [monotonicity #2403]: #2406
-#2409 := [symm #2407]: #2408
-#2410 := (= uf_8 #36)
-#2404 := [hypothesis]: #180
-#2101 := (or #2334 #186)
-#2102 := [def-axiom]: #2101
-#2393 := [unit-resolution #2102 #2391]: #186
-#2405 := [symm #2393]: #42
-#2411 := [trans #2405 #2404]: #2410
-#2412 := [trans #2411 #2409]: #2358
-#2386 := (not #2358)
-#1852 := (>= uf_7 0::int)
-#1853 := (not #1852)
-#1858 := (* -1::int uf_7)
-#1863 := (+ uf_1 #1858)
-#1850 := (<= #1863 0::int)
-#2364 := (or #1850 #1853 #2358)
-#2369 := (not #2364)
-#2177 := (or #2334 #2328)
-#2187 := [def-axiom]: #2177
-#2392 := [unit-resolution #2187 #2391]: #2328
-#2019 := (+ uf_6 #767)
-#2021 := (<= #2019 0::int)
-#2394 := (or #250 #2021)
-#2395 := [th-lemma]: #2394
-#2396 := [unit-resolution #2395 #2393]: #2021
-#1897 := [hypothesis]: #2271
-#2178 := (or #2334 #786)
-#2175 := [def-axiom]: #2178
-#2397 := [unit-resolution #2175 #2391]: #786
-#1929 := (not #2021)
-#1908 := (or #1749 #789 #2276 #1929)
-#1921 := [hypothesis]: #786
-#2008 := (+ uf_5 #1170)
-#2011 := (<= #2008 0::int)
-#1989 := (+ uf_6 #1353)
-#1990 := (>= #1989 0::int)
-#1928 := (not #1990)
-#1922 := [hypothesis]: #2021
-#2085 := (not #1355)
-#1932 := [hypothesis]: #1752
-#2086 := (or #1749 #2085)
-#2087 := [def-axiom]: #2086
-#1915 := [unit-resolution #2087 #1932]: #2085
-#1930 := (or #1928 #1355 #1929)
-#1923 := [hypothesis]: #2085
-#1914 := [hypothesis]: #1990
-#1927 := [th-lemma #1914 #1923 #1922]: false
-#1931 := [lemma #1927]: #1930
-#1917 := [unit-resolution #1931 #1915 #1922]: #1928
-#1899 := (or #1990 #2011)
-#2200 := (or #1749 #1174)
-#2203 := [def-axiom]: #2200
-#1918 := [unit-resolution #2203 #1932]: #1174
-#1979 := (or #2276 #1592 #1990 #2011)
-#2018 := (+ #1167 #989)
-#2023 := (<= #2018 0::int)
-#2013 := (+ ?x6!2 #784)
-#2003 := (>= #2013 0::int)
-#2005 := (or #1592 #2003 #2023)
-#1980 := (or #2276 #2005)
-#1970 := (iff #1980 #1979)
-#1997 := (or #1592 #1990 #2011)
-#1982 := (or #2276 #1997)
-#1968 := (iff #1982 #1979)
-#1969 := [rewrite]: #1968
-#1975 := (iff #1980 #1982)
-#1977 := (iff #2005 #1997)
-#1995 := (or #1592 #2011 #1990)
-#1974 := (iff #1995 #1997)
-#1976 := [rewrite]: #1974
-#1996 := (iff #2005 #1995)
-#1993 := (iff #2023 #1990)
-#1999 := (+ #989 #1167)
-#1986 := (<= #1999 0::int)
-#1991 := (iff #1986 #1990)
-#1992 := [rewrite]: #1991
-#1987 := (iff #2023 #1986)
-#2002 := (= #2018 #1999)
-#1984 := [rewrite]: #2002
-#1988 := [monotonicity #1984]: #1987
-#1994 := [trans #1988 #1992]: #1993
-#2000 := (iff #2003 #2011)
-#2006 := (+ #784 ?x6!2)
-#2014 := (>= #2006 0::int)
-#2012 := (iff #2014 #2011)
-#1998 := [rewrite]: #2012
-#2007 := (iff #2003 #2014)
-#2009 := (= #2013 #2006)
-#2010 := [rewrite]: #2009
-#2015 := [monotonicity #2010]: #2007
-#2001 := [trans #2015 #1998]: #2000
-#1985 := [monotonicity #2001 #1994]: #1996
-#1978 := [trans #1985 #1976]: #1977
-#1983 := [monotonicity #1978]: #1975
-#1972 := [trans #1983 #1969]: #1970
-#1981 := [quant-inst]: #1980
-#1971 := [mp #1981 #1972]: #1979
-#1904 := [unit-resolution #1971 #1897 #1918]: #1899
-#1905 := [unit-resolution #1904 #1917]: #2011
-#1173 := (not #1172)
-#2201 := (or #1749 #1173)
-#2202 := [def-axiom]: #2201
-#1906 := [unit-resolution #2202 #1932]: #1173
-#1907 := [th-lemma #1906 #1905 #1921]: false
-#1909 := [lemma #1907]: #1908
-#2398 := [unit-resolution #1909 #2397 #1897 #2396]: #1749
-#2098 := (or #2331 #2323 #1752)
-#2091 := [def-axiom]: #2098
-#2399 := [unit-resolution #2091 #2398 #2392]: #2323
-#2192 := (not #2323)
-#2372 := (or #2192 #2369)
-#1854 := (= #1851 uf_8)
-#2356 := (or #1854 #1853 #1850)
-#2357 := (not #2356)
-#2373 := (or #2192 #2357)
-#2375 := (iff #2373 #2372)
-#2377 := (iff #2372 #2372)
-#2378 := [rewrite]: #2377
-#2370 := (iff #2357 #2369)
-#2367 := (iff #2356 #2364)
-#2361 := (or #2358 #1853 #1850)
-#2365 := (iff #2361 #2364)
-#2366 := [rewrite]: #2365
-#2362 := (iff #2356 #2361)
-#2359 := (iff #1854 #2358)
-#2360 := [rewrite]: #2359
-#2363 := [monotonicity #2360]: #2362
-#2368 := [trans #2363 #2366]: #2367
-#2371 := [monotonicity #2368]: #2370
-#2376 := [monotonicity #2371]: #2375
-#2379 := [trans #2376 #2378]: #2375
-#2374 := [quant-inst]: #2373
-#2380 := [mp #2374 #2379]: #2372
-#2400 := [unit-resolution #2380 #2399]: #2369
-#2387 := (or #2364 #2386)
-#2388 := [def-axiom]: #2387
-#2401 := [unit-resolution #2388 #2400]: #2386
-#2413 := [unit-resolution #2401 #2412]: false
-#2415 := [lemma #2413]: #2414
-#2596 := [unit-resolution #2415 #2593 #2595]: #2334
-#2181 := (or #2343 #2320 #2337)
-#2182 := [def-axiom]: #2181
-#2615 := [unit-resolution #2182 #2596 #2594]: #2320
-#2209 := (or #2317 #2311)
-#2210 := [def-axiom]: #2209
-#2681 := [unit-resolution #2210 #2615]: #2311
-#2422 := (or #332 #2314 #583)
-#2416 := (= #67 #89)
-#2022 := (= #89 #67)
-#2381 := [hypothesis]: #2311
-#1231 := (not #332)
-#1874 := [hypothesis]: #1231
-#1868 := (or #2305 #332 #583)
-#1880 := (= #36 #89)
-#1862 := (= #89 #36)
-#1859 := [hypothesis]: #2308
-#2232 := (or #2305 #471)
-#1954 := [def-axiom]: #2232
-#1856 := [unit-resolution #1954 #1859]: #471
-#1857 := [symm #1856]: #110
-#1878 := [monotonicity #1857]: #1862
-#1877 := [symm #1878]: #1880
-#1876 := (= uf_12 #36)
-#1955 := (or #2305 #474)
-#2229 := [def-axiom]: #1955
-#1860 := [unit-resolution #2229 #1859]: #474
-#1861 := [symm #1860]: #111
-#1881 := [trans #1861 #2404]: #1876
-#1864 := [trans #1881 #1877]: #332
-#1867 := [unit-resolution #1874 #1864]: false
-#1871 := [lemma #1867]: #1868
-#2382 := [unit-resolution #1871 #1874 #2404]: #2305
-#2221 := (or #2314 #2302 #2308)
-#2216 := [def-axiom]: #2221
-#2383 := [unit-resolution #2216 #2382 #2381]: #2302
-#1900 := (or #2299 #310)
-#1901 := [def-axiom]: #1900
-#2389 := [unit-resolution #1901 #2383]: #310
-#2390 := [symm #2389]: #73
-#1872 := [monotonicity #2390]: #2022
-#2417 := [symm #1872]: #2416
-#2418 := (= uf_12 #67)
-#1896 := (or #2299 #305)
-#2237 := [def-axiom]: #1896
-#2385 := [unit-resolution #2237 #2383]: #305
-#1866 := [symm #2385]: #70
-#1902 := (or #2299 #313)
-#1903 := [def-axiom]: #1902
-#2384 := [unit-resolution #1903 #2383]: #313
-#1873 := [symm #2384]: #75
-#2419 := [trans #1873 #1866]: #2418
-#2420 := [trans #2419 #2417]: #332
-#2421 := [unit-resolution #1874 #2420]: false
-#2423 := [lemma #2421]: #2422
-#2682 := [unit-resolution #2423 #2681 #2593]: #332
-#1940 := (or #2287 #1231)
-#1919 := [def-axiom]: #1940
-#2683 := [unit-resolution #1919 #2682]: #2287
-#2679 := (or #2305 #2276 #2290)
-#2650 := [hypothesis]: #2287
-#2224 := (or #2305 #2293)
-#2228 := [def-axiom]: #2224
-#2651 := [unit-resolution #2228 #1859]: #2293
-#1912 := (or #2296 #2290 #1691)
-#2253 := [def-axiom]: #1912
-#2652 := [unit-resolution #2253 #2651 #2650]: #1691
-#1925 := (or #1686 #1403)
-#2257 := [def-axiom]: #1925
-#2653 := [unit-resolution #2257 #2652]: #1403
-#2482 := (+ uf_5 #1396)
-#2627 := (>= #2482 0::int)
-#2668 := (not #2627)
-#2616 := (= uf_5 ?x8!3)
-#2647 := (not #2616)
-#2626 := (= #67 #1216)
-#2631 := (not #2626)
-#2630 := (+ #67 #1418)
-#2632 := (>= #2630 0::int)
-#2636 := (not #2632)
-#2223 := (or #2305 #881)
-#2227 := [def-axiom]: #2223
-#2654 := [unit-resolution #2227 #1859]: #881
-#2258 := (not #1420)
-#2259 := (or #1686 #2258)
-#2260 := [def-axiom]: #2259
-#2655 := [unit-resolution #2260 #2652]: #2258
-#1961 := (+ uf_6 #839)
-#1855 := (<= #1961 0::int)
-#2656 := (or #480 #1855)
-#2657 := [th-lemma]: #2656
-#2658 := [unit-resolution #2657 #1860]: #1855
-#2606 := (not #1855)
-#2637 := (or #2636 #2606 #1420 #880)
-#2633 := [hypothesis]: #881
-#2598 := [hypothesis]: #2258
-#2600 := [hypothesis]: #1855
-#2634 := [hypothesis]: #2632
-#2635 := [th-lemma #2634 #2600 #2598 #2633]: false
-#2638 := [lemma #2635]: #2637
-#2659 := [unit-resolution #2638 #2658 #2655 #2654]: #2636
-#2639 := (or #2631 #2632)
-#2640 := [th-lemma]: #2639
-#2660 := [unit-resolution #2640 #2659]: #2631
-#2648 := (or #2647 #2626)
-#2644 := [hypothesis]: #2616
-#2645 := [monotonicity #2644]: #2626
-#2643 := [hypothesis]: #2631
-#2646 := [unit-resolution #2643 #2645]: false
-#2649 := [lemma #2646]: #2648
-#2661 := [unit-resolution #2649 #2660]: #2647
-#2671 := (or #2616 #2668)
-#2483 := (<= #2482 0::int)
-#2494 := (+ uf_6 #1418)
-#2495 := (>= #2494 0::int)
-#2612 := (not #2495)
-#2613 := (or #2612 #2606 #1420)
-#2610 := [hypothesis]: #2495
-#2611 := [th-lemma #2600 #2598 #2610]: false
-#2614 := [lemma #2611]: #2613
-#2662 := [unit-resolution #2614 #2658 #2655]: #2612
-#2664 := (or #2483 #2495)
-#2250 := (or #1686 #1222)
-#1924 := [def-axiom]: #2250
-#2663 := [unit-resolution #1924 #2652]: #1222
-#2503 := (or #2276 #1671 #2483 #2495)
-#2471 := (+ #1216 #989)
-#2472 := (<= #2471 0::int)
-#2473 := (+ ?x8!3 #784)
-#2474 := (>= #2473 0::int)
-#2475 := (or #1671 #2474 #2472)
-#2504 := (or #2276 #2475)
-#2511 := (iff #2504 #2503)
-#2500 := (or #1671 #2483 #2495)
-#2506 := (or #2276 #2500)
-#2509 := (iff #2506 #2503)
-#2510 := [rewrite]: #2509
-#2507 := (iff #2504 #2506)
-#2501 := (iff #2475 #2500)
-#2498 := (iff #2472 #2495)
-#2488 := (+ #989 #1216)
-#2491 := (<= #2488 0::int)
-#2496 := (iff #2491 #2495)
-#2497 := [rewrite]: #2496
-#2492 := (iff #2472 #2491)
-#2489 := (= #2471 #2488)
-#2490 := [rewrite]: #2489
-#2493 := [monotonicity #2490]: #2492
-#2499 := [trans #2493 #2497]: #2498
-#2486 := (iff #2474 #2483)
-#2476 := (+ #784 ?x8!3)
-#2479 := (>= #2476 0::int)
-#2484 := (iff #2479 #2483)
-#2485 := [rewrite]: #2484
-#2480 := (iff #2474 #2479)
-#2477 := (= #2473 #2476)
-#2478 := [rewrite]: #2477
-#2481 := [monotonicity #2478]: #2480
-#2487 := [trans #2481 #2485]: #2486
-#2502 := [monotonicity #2487 #2499]: #2501
-#2508 := [monotonicity #2502]: #2507
-#2512 := [trans #2508 #2510]: #2511
-#2505 := [quant-inst]: #2504
-#2513 := [mp #2505 #2512]: #2503
-#2665 := [unit-resolution #2513 #1897 #2663]: #2664
-#2666 := [unit-resolution #2665 #2662]: #2483
-#2667 := (not #2483)
-#2669 := (or #2616 #2667 #2668)
-#2670 := [th-lemma]: #2669
-#2672 := [unit-resolution #2670 #2666]: #2671
-#2673 := [unit-resolution #2672 #2661]: #2668
-#1936 := (>= #865 -1::int)
-#2226 := (or #2305 #864)
-#1941 := [def-axiom]: #2226
-#2674 := [unit-resolution #1941 #1859]: #864
-#2675 := (or #868 #1936)
-#2676 := [th-lemma]: #2675
-#2677 := [unit-resolution #2676 #2674]: #1936
-#2678 := [th-lemma #2677 #2673 #2653]: false
-#2680 := [lemma #2678]: #2679
-#2684 := [unit-resolution #2680 #2595 #2683]: #2305
-#2685 := [unit-resolution #2216 #2684 #2681]: #2302
-#2248 := (or #2299 #2293)
-#2246 := [def-axiom]: #2248
-#2686 := [unit-resolution #2246 #2685]: #2293
-#2687 := [unit-resolution #2253 #2686 #2683]: #1691
-#2688 := [unit-resolution #2257 #2687]: #1403
-#2464 := (+ #67 #839)
-#2465 := (<= #2464 0::int)
-#2463 := (= #67 uf_12)
-#2689 := [unit-resolution #1903 #2685]: #313
-#2690 := [unit-resolution #2237 #2685]: #305
-#2691 := [trans #2690 #2689]: #2463
-#2692 := (not #2463)
-#2693 := (or #2692 #2465)
-#2694 := [th-lemma]: #2693
-#2695 := [unit-resolution #2694 #2691]: #2465
-#1887 := (or #2299 #880)
-#1888 := [def-axiom]: #1887
-#2696 := [unit-resolution #1888 #2685]: #880
-#2697 := [unit-resolution #2260 #2687]: #2258
-#2698 := (not #2465)
-#2699 := (or #2612 #1420 #2698 #881)
-#2700 := [th-lemma]: #2699
-#2701 := [unit-resolution #2700 #2697 #2696 #2695]: #2612
-#2702 := [unit-resolution #1924 #2687]: #1222
-#2703 := [unit-resolution #2513 #2595 #2702]: #2664
-#2704 := [unit-resolution #2703 #2701]: #2483
-#2705 := (or #2636 #1420 #2698)
-#2706 := [th-lemma]: #2705
-#2707 := [unit-resolution #2706 #2697 #2695]: #2636
-#2708 := [unit-resolution #2640 #2707]: #2631
-#2709 := [unit-resolution #2649 #2708]: #2647
-#2710 := [unit-resolution #2670 #2709 #2704]: #2668
-#2245 := (or #2299 #864)
-#2247 := [def-axiom]: #2245
-#2711 := [unit-resolution #2247 #2685]: #864
-#2712 := [unit-resolution #2676 #2711]: #1936
-#2713 := [th-lemma #2712 #2710 #2688]: false
-#2714 := [lemma #2713]: #2346
-#2352 := (or #1499 #2349)
-#1569 := (forall (vars (?x4 int)) #1566)
-#1755 := (or #1569 #1752)
-#1758 := (not #1755)
-#1761 := (or #259 #250 #241 #1631 #877 #789 #1758)
-#1764 := (not #1761)
-#1658 := (forall (vars (?x8 int)) #1653)
-#1664 := (not #1658)
-#1665 := (or #332 #1664)
-#1666 := (not #1665)
-#1694 := (or #1666 #1691)
-#1702 := (not #1694)
-#1712 := (or #489 #480 #1631 #877 #1700 #1701 #868 #880 #1702)
-#1713 := (not #1712)
-#1703 := (or #441 #416 #407 #1631 #877 #1700 #1701 #868 #881 #1702)
-#1704 := (not #1703)
-#1718 := (or #1704 #1713)
-#1724 := (not #1718)
-#1725 := (or #1631 #877 #786 #1724)
-#1726 := (not #1725)
-#1770 := (or #1726 #1764)
-#1775 := (not #1770)
-#1544 := (forall (vars (?x3 int)) #1539)
-#1738 := (not #1544)
-#1522 := (forall (vars (?x1 int)) #1517)
-#1737 := (not #1522)
-#1778 := (or #583 #1631 #877 #1737 #1738 #1775)
-#1781 := (not #1778)
-#1784 := (or #1499 #1781)
-#2353 := (iff #1784 #2352)
-#2350 := (iff #1781 #2349)
-#2347 := (iff #1778 #2346)
-#2344 := (iff #1775 #2343)
-#2341 := (iff #1770 #2340)
-#2338 := (iff #1764 #2337)
-#2335 := (iff #1761 #2334)
-#2332 := (iff #1758 #2331)
-#2329 := (iff #1755 #2328)
-#2326 := (iff #1569 #2323)
-#2324 := (iff #1566 #1566)
-#2325 := [refl]: #2324
-#2327 := [quant-intro #2325]: #2326
-#2330 := [monotonicity #2327]: #2329
-#2333 := [monotonicity #2330]: #2332
-#2336 := [monotonicity #2333]: #2335
-#2339 := [monotonicity #2336]: #2338
-#2321 := (iff #1726 #2320)
-#2318 := (iff #1725 #2317)
-#2315 := (iff #1724 #2314)
-#2312 := (iff #1718 #2311)
-#2309 := (iff #1713 #2308)
-#2306 := (iff #1712 #2305)
-#2297 := (iff #1702 #2296)
-#2294 := (iff #1694 #2293)
-#2291 := (iff #1666 #2290)
-#2288 := (iff #1665 #2287)
-#2285 := (iff #1664 #2284)
-#2282 := (iff #1658 #2279)
-#2280 := (iff #1653 #1653)
-#2281 := [refl]: #2280
-#2283 := [quant-intro #2281]: #2282
-#2286 := [monotonicity #2283]: #2285
-#2289 := [monotonicity #2286]: #2288
-#2292 := [monotonicity #2289]: #2291
-#2295 := [monotonicity #2292]: #2294
-#2298 := [monotonicity #2295]: #2297
-#2307 := [monotonicity #2298]: #2306
-#2310 := [monotonicity #2307]: #2309
-#2303 := (iff #1704 #2302)
-#2300 := (iff #1703 #2299)
-#2301 := [monotonicity #2298]: #2300
-#2304 := [monotonicity #2301]: #2303
-#2313 := [monotonicity #2304 #2310]: #2312
-#2316 := [monotonicity #2313]: #2315
-#2319 := [monotonicity #2316]: #2318
-#2322 := [monotonicity #2319]: #2321
-#2342 := [monotonicity #2322 #2339]: #2341
-#2345 := [monotonicity #2342]: #2344
-#2277 := (iff #1738 #2276)
-#2274 := (iff #1544 #2271)
-#2272 := (iff #1539 #1539)
-#2273 := [refl]: #2272
-#2275 := [quant-intro #2273]: #2274
-#2278 := [monotonicity #2275]: #2277
-#2269 := (iff #1737 #2268)
-#2266 := (iff #1522 #2263)
-#2264 := (iff #1517 #1517)
-#2265 := [refl]: #2264
-#2267 := [quant-intro #2265]: #2266
-#2270 := [monotonicity #2267]: #2269
-#2348 := [monotonicity #2270 #2278 #2345]: #2347
-#2351 := [monotonicity #2348]: #2350
-#2354 := [monotonicity #2351]: #2353
-#1406 := (and #1222 #1403)
-#1409 := (not #1406)
-#1425 := (or #1409 #1420)
-#1428 := (not #1425)
-#1241 := (and #1231 #847)
-#1434 := (or #1241 #1428)
-#1458 := (and #471 #474 #730 #733 #816 #819 #864 #881 #1434)
-#1446 := (and #305 #310 #313 #730 #733 #816 #819 #864 #880 #1434)
-#1463 := (or #1446 #1458)
-#1469 := (and #730 #733 #789 #1463)
-#1341 := (and #1173 #1174)
-#1344 := (not #1341)
-#1360 := (or #1344 #1355)
-#1363 := (not #1360)
-#1152 := (not #1151)
-#1332 := (and #1152 #1153)
-#1335 := (not #1332)
-#1338 := (or #1329 #1335)
-#1366 := (and #1338 #1363)
-#1142 := (not #756)
-#1145 := (forall (vars (?x4 int)) #1142)
-#1369 := (or #1145 #1366)
-#1375 := (and #183 #186 #189 #730 #733 #786 #1369)
-#1474 := (or #1375 #1469)
-#1480 := (and #180 #723 #730 #733 #997 #1474)
-#1110 := (not #1109)
-#1302 := (and #1110 #1111)
-#1305 := (not #1302)
-#1311 := (or #1108 #1305)
-#1316 := (not #1311)
-#1485 := (or #1316 #1480)
-#1787 := (iff #1485 #1784)
-#1607 := (or #1172 #1592 #1355)
-#1619 := (or #1618 #1607)
-#1620 := (not #1619)
-#1625 := (or #1569 #1620)
-#1632 := (not #1625)
-#1633 := (or #259 #250 #241 #1631 #877 #789 #1632)
-#1634 := (not #1633)
-#1731 := (or #1634 #1726)
-#1739 := (not #1731)
-#1740 := (or #583 #1631 #877 #1737 #1738 #1739)
-#1741 := (not #1740)
-#1746 := (or #1499 #1741)
-#1785 := (iff #1746 #1784)
-#1782 := (iff #1741 #1781)
-#1779 := (iff #1740 #1778)
-#1776 := (iff #1739 #1775)
-#1773 := (iff #1731 #1770)
-#1767 := (or #1764 #1726)
-#1771 := (iff #1767 #1770)
-#1772 := [rewrite]: #1771
-#1768 := (iff #1731 #1767)
-#1765 := (iff #1634 #1764)
-#1762 := (iff #1633 #1761)
-#1759 := (iff #1632 #1758)
-#1756 := (iff #1625 #1755)
-#1753 := (iff #1620 #1752)
-#1750 := (iff #1619 #1749)
-#1751 := [rewrite]: #1750
-#1754 := [monotonicity #1751]: #1753
-#1757 := [monotonicity #1754]: #1756
-#1760 := [monotonicity #1757]: #1759
-#1763 := [monotonicity #1760]: #1762
-#1766 := [monotonicity #1763]: #1765
-#1769 := [monotonicity #1766]: #1768
-#1774 := [trans #1769 #1772]: #1773
-#1777 := [monotonicity #1774]: #1776
-#1780 := [monotonicity #1777]: #1779
-#1783 := [monotonicity #1780]: #1782
-#1786 := [monotonicity #1783]: #1785
-#1747 := (iff #1485 #1746)
-#1744 := (iff #1480 #1741)
-#1734 := (and #180 #1522 #730 #733 #1544 #1731)
-#1742 := (iff #1734 #1741)
-#1743 := [rewrite]: #1742
-#1735 := (iff #1480 #1734)
-#1732 := (iff #1474 #1731)
-#1729 := (iff #1469 #1726)
-#1721 := (and #730 #733 #789 #1718)
-#1727 := (iff #1721 #1726)
-#1728 := [rewrite]: #1727
-#1722 := (iff #1469 #1721)
-#1719 := (iff #1463 #1718)
-#1716 := (iff #1458 #1713)
-#1709 := (and #471 #474 #730 #733 #816 #819 #864 #881 #1694)
-#1714 := (iff #1709 #1713)
-#1715 := [rewrite]: #1714
-#1710 := (iff #1458 #1709)
-#1695 := (iff #1434 #1694)
-#1692 := (iff #1428 #1691)
-#1689 := (iff #1425 #1686)
-#1672 := (or #1671 #1398)
-#1683 := (or #1672 #1420)
-#1687 := (iff #1683 #1686)
-#1688 := [rewrite]: #1687
-#1684 := (iff #1425 #1683)
-#1681 := (iff #1409 #1672)
-#1673 := (not #1672)
-#1676 := (not #1673)
-#1679 := (iff #1676 #1672)
-#1680 := [rewrite]: #1679
-#1677 := (iff #1409 #1676)
-#1674 := (iff #1406 #1673)
-#1675 := [rewrite]: #1674
-#1678 := [monotonicity #1675]: #1677
-#1682 := [trans #1678 #1680]: #1681
-#1685 := [monotonicity #1682]: #1684
-#1690 := [trans #1685 #1688]: #1689
-#1693 := [monotonicity #1690]: #1692
-#1669 := (iff #1241 #1666)
-#1661 := (and #1231 #1658)
-#1667 := (iff #1661 #1666)
-#1668 := [rewrite]: #1667
-#1662 := (iff #1241 #1661)
-#1659 := (iff #847 #1658)
-#1656 := (iff #844 #1653)
-#1639 := (or #1502 #827)
-#1650 := (or #1639 #841)
-#1654 := (iff #1650 #1653)
-#1655 := [rewrite]: #1654
-#1651 := (iff #844 #1650)
-#1648 := (iff #836 #1639)
-#1640 := (not #1639)
-#1643 := (not #1640)
-#1646 := (iff #1643 #1639)
-#1647 := [rewrite]: #1646
-#1644 := (iff #836 #1643)
-#1641 := (iff #833 #1640)
-#1642 := [rewrite]: #1641
-#1645 := [monotonicity #1642]: #1644
-#1649 := [trans #1645 #1647]: #1648
-#1652 := [monotonicity #1649]: #1651
-#1657 := [trans #1652 #1655]: #1656
-#1660 := [quant-intro #1657]: #1659
-#1663 := [monotonicity #1660]: #1662
-#1670 := [trans #1663 #1668]: #1669
-#1696 := [monotonicity #1670 #1693]: #1695
-#1711 := [monotonicity #1696]: #1710
-#1717 := [trans #1711 #1715]: #1716
-#1707 := (iff #1446 #1704)
-#1697 := (and #305 #310 #313 #730 #733 #816 #819 #864 #880 #1694)
-#1705 := (iff #1697 #1704)
-#1706 := [rewrite]: #1705
-#1698 := (iff #1446 #1697)
-#1699 := [monotonicity #1696]: #1698
-#1708 := [trans #1699 #1706]: #1707
-#1720 := [monotonicity #1708 #1717]: #1719
-#1723 := [monotonicity #1720]: #1722
-#1730 := [trans #1723 #1728]: #1729
-#1637 := (iff #1375 #1634)
-#1628 := (and #183 #186 #189 #730 #733 #786 #1625)
-#1635 := (iff #1628 #1634)
-#1636 := [rewrite]: #1635
-#1629 := (iff #1375 #1628)
-#1626 := (iff #1369 #1625)
-#1623 := (iff #1366 #1620)
-#1612 := (not #1607)
-#1615 := (and #1587 #1612)
-#1621 := (iff #1615 #1620)
-#1622 := [rewrite]: #1621
-#1616 := (iff #1366 #1615)
-#1613 := (iff #1363 #1612)
-#1610 := (iff #1360 #1607)
-#1593 := (or #1172 #1592)
-#1604 := (or #1593 #1355)
-#1608 := (iff #1604 #1607)
-#1609 := [rewrite]: #1608
-#1605 := (iff #1360 #1604)
-#1602 := (iff #1344 #1593)
-#1594 := (not #1593)
-#1597 := (not #1594)
-#1600 := (iff #1597 #1593)
-#1601 := [rewrite]: #1600
-#1598 := (iff #1344 #1597)
-#1595 := (iff #1341 #1594)
-#1596 := [rewrite]: #1595
-#1599 := [monotonicity #1596]: #1598
-#1603 := [trans #1599 #1601]: #1602
-#1606 := [monotonicity #1603]: #1605
-#1611 := [trans #1606 #1609]: #1610
-#1614 := [monotonicity #1611]: #1613
-#1590 := (iff #1338 #1587)
-#1573 := (or #1151 #1572)
-#1584 := (or #1329 #1573)
-#1588 := (iff #1584 #1587)
-#1589 := [rewrite]: #1588
-#1585 := (iff #1338 #1584)
-#1582 := (iff #1335 #1573)
-#1574 := (not #1573)
-#1577 := (not #1574)
-#1580 := (iff #1577 #1573)
-#1581 := [rewrite]: #1580
-#1578 := (iff #1335 #1577)
-#1575 := (iff #1332 #1574)
-#1576 := [rewrite]: #1575
-#1579 := [monotonicity #1576]: #1578
-#1583 := [trans #1579 #1581]: #1582
-#1586 := [monotonicity #1583]: #1585
-#1591 := [trans #1586 #1589]: #1590
-#1617 := [monotonicity #1591 #1614]: #1616
-#1624 := [trans #1617 #1622]: #1623
-#1570 := (iff #1145 #1569)
-#1567 := (iff #1142 #1566)
-#1564 := (iff #756 #1561)
-#1547 := (or #1502 #743)
-#1558 := (or #47 #1547)
-#1562 := (iff #1558 #1561)
-#1563 := [rewrite]: #1562
-#1559 := (iff #756 #1558)
-#1556 := (iff #750 #1547)
-#1548 := (not #1547)
-#1551 := (not #1548)
-#1554 := (iff #1551 #1547)
-#1555 := [rewrite]: #1554
-#1552 := (iff #750 #1551)
-#1549 := (iff #747 #1548)
-#1550 := [rewrite]: #1549
-#1553 := [monotonicity #1550]: #1552
-#1557 := [trans #1553 #1555]: #1556
-#1560 := [monotonicity #1557]: #1559
-#1565 := [trans #1560 #1563]: #1564
-#1568 := [monotonicity #1565]: #1567
-#1571 := [quant-intro #1568]: #1570
-#1627 := [monotonicity #1571 #1624]: #1626
-#1630 := [monotonicity #1627]: #1629
-#1638 := [trans #1630 #1636]: #1637
-#1733 := [monotonicity #1638 #1730]: #1732
-#1545 := (iff #997 #1544)
-#1542 := (iff #994 #1539)
-#1525 := (or #1502 #978)
-#1536 := (or #1525 #991)
-#1540 := (iff #1536 #1539)
-#1541 := [rewrite]: #1540
-#1537 := (iff #994 #1536)
-#1534 := (iff #986 #1525)
-#1526 := (not #1525)
-#1529 := (not #1526)
-#1532 := (iff #1529 #1525)
-#1533 := [rewrite]: #1532
-#1530 := (iff #986 #1529)
-#1527 := (iff #983 #1526)
-#1528 := [rewrite]: #1527
-#1531 := [monotonicity #1528]: #1530
-#1535 := [trans #1531 #1533]: #1534
-#1538 := [monotonicity #1535]: #1537
-#1543 := [trans #1538 #1541]: #1542
-#1546 := [quant-intro #1543]: #1545
-#1523 := (iff #723 #1522)
-#1520 := (iff #720 #1517)
-#1503 := (or #1502 #706)
-#1514 := (or #1503 #716)
-#1518 := (iff #1514 #1517)
-#1519 := [rewrite]: #1518
-#1515 := (iff #720 #1514)
-#1512 := (iff #711 #1503)
-#1504 := (not #1503)
-#1507 := (not #1504)
-#1510 := (iff #1507 #1503)
-#1511 := [rewrite]: #1510
-#1508 := (iff #711 #1507)
-#1505 := (iff #708 #1504)
-#1506 := [rewrite]: #1505
-#1509 := [monotonicity #1506]: #1508
-#1513 := [trans #1509 #1511]: #1512
-#1516 := [monotonicity #1513]: #1515
-#1521 := [trans #1516 #1519]: #1520
-#1524 := [quant-intro #1521]: #1523
-#1736 := [monotonicity #1524 #1546 #1733]: #1735
-#1745 := [trans #1736 #1743]: #1744
-#1500 := (iff #1316 #1499)
-#1497 := (iff #1311 #1494)
-#1230 := (or #1109 #1229)
-#1491 := (or #1108 #1230)
-#1495 := (iff #1491 #1494)
-#1496 := [rewrite]: #1495
-#1492 := (iff #1311 #1491)
-#1489 := (iff #1305 #1230)
-#1182 := (not #1230)
-#1162 := (not #1182)
-#1301 := (iff #1162 #1230)
-#1488 := [rewrite]: #1301
-#1118 := (iff #1305 #1162)
-#1183 := (iff #1302 #1182)
-#1161 := [rewrite]: #1183
-#1119 := [monotonicity #1161]: #1118
-#1490 := [trans #1119 #1488]: #1489
-#1493 := [monotonicity #1490]: #1492
-#1498 := [trans #1493 #1496]: #1497
-#1501 := [monotonicity #1498]: #1500
-#1748 := [monotonicity #1501 #1745]: #1747
-#1788 := [trans #1748 #1786]: #1787
-#1252 := (not #874)
-#1249 := (not #868)
-#1217 := (+ #1216 #839)
-#1218 := (<= #1217 0::int)
-#1219 := (+ ?x8!3 #828)
-#1220 := (>= #1219 0::int)
-#1221 := (not #1220)
-#1223 := (and #1222 #1221)
-#1224 := (not #1223)
-#1225 := (or #1224 #1218)
-#1226 := (not #1225)
-#1245 := (or #1226 #1241)
-#1212 := (not #824)
-#1130 := (not #738)
-#1264 := (not #480)
-#1261 := (not #489)
-#1269 := (and #1261 #1264 #1130 #1212 #1245 #1249 #1252 #886)
-#1209 := (not #877)
-#1206 := (not #407)
-#1203 := (not #416)
-#1200 := (not #441)
-#1257 := (and #1200 #1203 #1206 #1209 #1130 #1212 #1245 #1249 #1252 #880)
-#1273 := (or #1257 #1269)
-#1277 := (and #1130 #789 #1273)
-#1168 := (+ #1167 #767)
-#1169 := (<= #1168 0::int)
-#1175 := (and #1174 #1173)
-#1176 := (not #1175)
-#1177 := (or #1176 #1169)
-#1178 := (not #1177)
-#1154 := (and #1153 #1152)
-#1155 := (not #1154)
-#1157 := (= #1156 uf_8)
-#1158 := (or #1157 #1155)
-#1184 := (and #1158 #1178)
-#1188 := (or #1145 #1184)
-#1139 := (not #241)
-#1136 := (not #250)
-#1133 := (not #259)
-#1194 := (and #1133 #1136 #1139 #1130 #1188 #954)
-#1281 := (or #1194 #1277)
-#1120 := (not #583)
-#1292 := (and #1120 #723 #1130 #1281 #997)
-#1112 := (and #1111 #1110)
-#1113 := (not #1112)
-#1114 := (or #1113 #1108)
-#1115 := (not #1114)
-#1296 := (or #1115 #1292)
-#1486 := (iff #1296 #1485)
-#1483 := (iff #1292 #1480)
-#1477 := (and #180 #723 #735 #1474 #997)
-#1481 := (iff #1477 #1480)
-#1482 := [rewrite]: #1481
-#1478 := (iff #1292 #1477)
-#1475 := (iff #1281 #1474)
-#1472 := (iff #1277 #1469)
-#1466 := (and #735 #789 #1463)
-#1470 := (iff #1466 #1469)
-#1471 := [rewrite]: #1470
-#1467 := (iff #1277 #1466)
-#1464 := (iff #1273 #1463)
-#1461 := (iff #1269 #1458)
-#1455 := (and #471 #474 #735 #821 #1434 #864 #871 #881)
-#1459 := (iff #1455 #1458)
-#1460 := [rewrite]: #1459
-#1456 := (iff #1269 #1455)
-#1441 := (iff #1252 #871)
-#1442 := [rewrite]: #1441
-#1439 := (iff #1249 #864)
-#1440 := [rewrite]: #1439
-#1437 := (iff #1245 #1434)
-#1431 := (or #1428 #1241)
-#1435 := (iff #1431 #1434)
-#1436 := [rewrite]: #1435
-#1432 := (iff #1245 #1431)
-#1429 := (iff #1226 #1428)
-#1426 := (iff #1225 #1425)
-#1423 := (iff #1218 #1420)
-#1412 := (+ #839 #1216)
-#1415 := (<= #1412 0::int)
-#1421 := (iff #1415 #1420)
-#1422 := [rewrite]: #1421
-#1416 := (iff #1218 #1415)
-#1413 := (= #1217 #1412)
-#1414 := [rewrite]: #1413
-#1417 := [monotonicity #1414]: #1416
-#1424 := [trans #1417 #1422]: #1423
-#1410 := (iff #1224 #1409)
-#1407 := (iff #1223 #1406)
-#1404 := (iff #1221 #1403)
-#1401 := (iff #1220 #1398)
-#1390 := (+ #828 ?x8!3)
-#1393 := (>= #1390 0::int)
-#1399 := (iff #1393 #1398)
-#1400 := [rewrite]: #1399
-#1394 := (iff #1220 #1393)
-#1391 := (= #1219 #1390)
-#1392 := [rewrite]: #1391
-#1395 := [monotonicity #1392]: #1394
-#1402 := [trans #1395 #1400]: #1401
-#1405 := [monotonicity #1402]: #1404
-#1408 := [monotonicity #1405]: #1407
-#1411 := [monotonicity #1408]: #1410
-#1427 := [monotonicity #1411 #1424]: #1426
-#1430 := [monotonicity #1427]: #1429
-#1433 := [monotonicity #1430]: #1432
-#1438 := [trans #1433 #1436]: #1437
-#1388 := (iff #1212 #821)
-#1389 := [rewrite]: #1388
-#1321 := (iff #1130 #735)
-#1322 := [rewrite]: #1321
-#1453 := (iff #1264 #474)
-#1454 := [rewrite]: #1453
-#1451 := (iff #1261 #471)
-#1452 := [rewrite]: #1451
-#1457 := [monotonicity #1452 #1454 #1322 #1389 #1438 #1440 #1442 #890]: #1456
-#1462 := [trans #1457 #1460]: #1461
-#1449 := (iff #1257 #1446)
-#1443 := (and #305 #310 #313 #733 #735 #821 #1434 #864 #871 #880)
-#1447 := (iff #1443 #1446)
-#1448 := [rewrite]: #1447
-#1444 := (iff #1257 #1443)
-#1386 := (iff #1209 #733)
-#1387 := [rewrite]: #1386
-#1384 := (iff #1206 #313)
-#1385 := [rewrite]: #1384
-#1382 := (iff #1203 #310)
-#1383 := [rewrite]: #1382
-#1380 := (iff #1200 #305)
-#1381 := [rewrite]: #1380
-#1445 := [monotonicity #1381 #1383 #1385 #1387 #1322 #1389 #1438 #1440 #1442]: #1444
-#1450 := [trans #1445 #1448]: #1449
-#1465 := [monotonicity #1450 #1462]: #1464
-#1468 := [monotonicity #1322 #1465]: #1467
-#1473 := [trans #1468 #1471]: #1472
-#1378 := (iff #1194 #1375)
-#1372 := (and #183 #186 #189 #735 #1369 #786)
-#1376 := (iff #1372 #1375)
-#1377 := [rewrite]: #1376
-#1373 := (iff #1194 #1372)
-#1370 := (iff #1188 #1369)
-#1367 := (iff #1184 #1366)
-#1364 := (iff #1178 #1363)
-#1361 := (iff #1177 #1360)
-#1358 := (iff #1169 #1355)
-#1347 := (+ #767 #1167)
-#1350 := (<= #1347 0::int)
-#1356 := (iff #1350 #1355)
-#1357 := [rewrite]: #1356
-#1351 := (iff #1169 #1350)
-#1348 := (= #1168 #1347)
-#1349 := [rewrite]: #1348
-#1352 := [monotonicity #1349]: #1351
-#1359 := [trans #1352 #1357]: #1358
-#1345 := (iff #1176 #1344)
-#1342 := (iff #1175 #1341)
-#1343 := [rewrite]: #1342
-#1346 := [monotonicity #1343]: #1345
-#1362 := [monotonicity #1346 #1359]: #1361
-#1365 := [monotonicity #1362]: #1364
-#1339 := (iff #1158 #1338)
-#1336 := (iff #1155 #1335)
-#1333 := (iff #1154 #1332)
-#1334 := [rewrite]: #1333
-#1337 := [monotonicity #1334]: #1336
-#1330 := (iff #1157 #1329)
-#1331 := [rewrite]: #1330
-#1340 := [monotonicity #1331 #1337]: #1339
-#1368 := [monotonicity #1340 #1365]: #1367
-#1371 := [monotonicity #1368]: #1370
-#1327 := (iff #1139 #189)
-#1328 := [rewrite]: #1327
-#1325 := (iff #1136 #186)
-#1326 := [rewrite]: #1325
-#1323 := (iff #1133 #183)
-#1324 := [rewrite]: #1323
-#1374 := [monotonicity #1324 #1326 #1328 #1322 #1371 #958]: #1373
-#1379 := [trans #1374 #1377]: #1378
-#1476 := [monotonicity #1379 #1473]: #1475
-#1319 := (iff #1120 #180)
-#1320 := [rewrite]: #1319
-#1479 := [monotonicity #1320 #1322 #1476]: #1478
-#1484 := [trans #1479 #1482]: #1483
-#1317 := (iff #1115 #1316)
-#1314 := (iff #1114 #1311)
-#1308 := (or #1305 #1108)
-#1312 := (iff #1308 #1311)
-#1313 := [rewrite]: #1312
-#1309 := (iff #1114 #1308)
-#1306 := (iff #1113 #1305)
-#1303 := (iff #1112 #1302)
-#1304 := [rewrite]: #1303
-#1307 := [monotonicity #1304]: #1306
-#1310 := [monotonicity #1307]: #1309
-#1315 := [trans #1310 #1313]: #1314
-#1318 := [monotonicity #1315]: #1317
-#1487 := [monotonicity #1318 #1484]: #1486
-#1092 := (or #583 #726 #738 #975 #1000)
-#1097 := (and #723 #1092)
-#1100 := (not #1097)
-#1297 := (~ #1100 #1296)
-#1293 := (not #1092)
-#1294 := (~ #1293 #1292)
-#1289 := (not #1000)
-#1290 := (~ #1289 #997)
-#1287 := (~ #997 #997)
-#1285 := (~ #994 #994)
-#1286 := [refl]: #1285
-#1288 := [nnf-pos #1286]: #1287
-#1291 := [nnf-neg #1288]: #1290
-#1282 := (not #975)
-#1283 := (~ #1282 #1281)
-#1278 := (not #970)
-#1279 := (~ #1278 #1277)
-#1274 := (not #949)
-#1275 := (~ #1274 #1273)
-#1270 := (not #944)
-#1271 := (~ #1270 #1269)
-#1267 := (~ #886 #886)
-#1268 := [refl]: #1267
-#1253 := (~ #1252 #1252)
-#1254 := [refl]: #1253
-#1250 := (~ #1249 #1249)
-#1251 := [refl]: #1250
-#1246 := (not #861)
-#1247 := (~ #1246 #1245)
-#1242 := (not #856)
-#1243 := (~ #1242 #1241)
-#1238 := (not #850)
-#1239 := (~ #1238 #847)
-#1236 := (~ #847 #847)
-#1234 := (~ #844 #844)
-#1235 := [refl]: #1234
-#1237 := [nnf-pos #1235]: #1236
-#1240 := [nnf-neg #1237]: #1239
-#1232 := (~ #1231 #1231)
-#1233 := [refl]: #1232
-#1244 := [nnf-neg #1233 #1240]: #1243
-#1227 := (~ #850 #1226)
-#1228 := [sk]: #1227
-#1248 := [nnf-neg #1228 #1244]: #1247
-#1213 := (~ #1212 #1212)
-#1214 := [refl]: #1213
-#1131 := (~ #1130 #1130)
-#1132 := [refl]: #1131
-#1265 := (~ #1264 #1264)
-#1266 := [refl]: #1265
-#1262 := (~ #1261 #1261)
-#1263 := [refl]: #1262
-#1272 := [nnf-neg #1263 #1266 #1132 #1214 #1248 #1251 #1254 #1268]: #1271
-#1258 := (not #920)
-#1259 := (~ #1258 #1257)
-#1255 := (~ #880 #880)
-#1256 := [refl]: #1255
-#1210 := (~ #1209 #1209)
-#1211 := [refl]: #1210
-#1207 := (~ #1206 #1206)
-#1208 := [refl]: #1207
-#1204 := (~ #1203 #1203)
-#1205 := [refl]: #1204
-#1201 := (~ #1200 #1200)
-#1202 := [refl]: #1201
-#1260 := [nnf-neg #1202 #1205 #1208 #1211 #1132 #1214 #1248 #1251 #1254 #1256]: #1259
-#1276 := [nnf-neg #1260 #1272]: #1275
-#1198 := (~ #789 #789)
-#1199 := [refl]: #1198
-#1280 := [nnf-neg #1132 #1199 #1276]: #1279
-#1195 := (not #810)
-#1196 := (~ #1195 #1194)
-#1192 := (~ #954 #954)
-#1193 := [refl]: #1192
-#1189 := (not #781)
-#1190 := (~ #1189 #1188)
-#1185 := (not #778)
-#1186 := (~ #1185 #1184)
-#1179 := (not #775)
-#1180 := (~ #1179 #1178)
-#1181 := [sk]: #1180
-#1163 := (not #764)
-#1164 := (~ #1163 #1158)
-#1159 := (~ #761 #1158)
-#1160 := [sk]: #1159
-#1165 := [nnf-neg #1160]: #1164
-#1187 := [nnf-neg #1165 #1181]: #1186
-#1146 := (~ #764 #1145)
-#1143 := (~ #1142 #1142)
-#1144 := [refl]: #1143
-#1147 := [nnf-neg #1144]: #1146
-#1191 := [nnf-neg #1147 #1187]: #1190
-#1140 := (~ #1139 #1139)
-#1141 := [refl]: #1140
-#1137 := (~ #1136 #1136)
-#1138 := [refl]: #1137
-#1134 := (~ #1133 #1133)
-#1135 := [refl]: #1134
-#1197 := [nnf-neg #1135 #1138 #1141 #1132 #1191 #1193]: #1196
-#1284 := [nnf-neg #1197 #1280]: #1283
-#1127 := (not #726)
-#1128 := (~ #1127 #723)
-#1125 := (~ #723 #723)
-#1123 := (~ #720 #720)
-#1124 := [refl]: #1123
-#1126 := [nnf-pos #1124]: #1125
-#1129 := [nnf-neg #1126]: #1128
-#1121 := (~ #1120 #1120)
-#1122 := [refl]: #1121
-#1295 := [nnf-neg #1122 #1129 #1132 #1284 #1291]: #1294
-#1116 := (~ #726 #1115)
-#1117 := [sk]: #1116
-#1298 := [nnf-neg #1117 #1295]: #1297
-#1064 := (not #1029)
-#1101 := (iff #1064 #1100)
-#1098 := (iff #1029 #1097)
-#1095 := (iff #1026 #1092)
-#1077 := (or #583 #738 #975 #1000)
-#1089 := (or #726 #1077)
-#1093 := (iff #1089 #1092)
-#1094 := [rewrite]: #1093
-#1090 := (iff #1026 #1089)
-#1087 := (iff #1023 #1077)
-#1082 := (and true #1077)
-#1085 := (iff #1082 #1077)
-#1086 := [rewrite]: #1085
-#1083 := (iff #1023 #1082)
-#1080 := (iff #1018 #1077)
-#1074 := (or false #583 #738 #975 #1000)
-#1078 := (iff #1074 #1077)
-#1079 := [rewrite]: #1078
-#1075 := (iff #1018 #1074)
-#1072 := (iff #616 false)
-#1070 := (iff #616 #694)
-#1069 := (iff #9 true)
-#1067 := [iff-true #1063]: #1069
-#1071 := [monotonicity #1067]: #1070
-#1073 := [trans #1071 #698]: #1072
-#1076 := [monotonicity #1073]: #1075
-#1081 := [trans #1076 #1079]: #1080
-#1084 := [monotonicity #1067 #1081]: #1083
-#1088 := [trans #1084 #1086]: #1087
-#1091 := [monotonicity #1088]: #1090
-#1096 := [trans #1091 #1094]: #1095
-#1099 := [monotonicity #1096]: #1098
-#1102 := [monotonicity #1099]: #1101
-#1065 := [not-or-elim #1062]: #1064
-#1103 := [mp #1065 #1102]: #1100
-#1299 := [mp~ #1103 #1298]: #1296
-#1300 := [mp #1299 #1487]: #1485
-#1789 := [mp #1300 #1788]: #1784
-#2355 := [mp #1789 #2354]: #2352
-#2111 := [unit-resolution #2355 #2714]: #1499
-#1933 := (or #1494 #1847)
-#1848 := [def-axiom]: #1933
-#2004 := [unit-resolution #1848 #2111]: #1847
-#2135 := (+ #8 #1104)
-#2136 := (>= #2135 0::int)
-#2134 := (= #8 #1107)
-#2112 := (= #1107 #8)
-#2113 := (= ?x1!0 0::int)
-#1934 := (or #1494 #1110)
-#1849 := [def-axiom]: #1934
-#2115 := [unit-resolution #1849 #2111]: #1110
-#1935 := (or #1494 #1111)
-#1926 := [def-axiom]: #1935
-#2116 := [unit-resolution #1926 #2111]: #1111
-#2108 := [th-lemma #2116 #2115]: #2113
-#2114 := [monotonicity #2108]: #2112
-#2082 := [symm #2114]: #2134
-#2089 := (not #2134)
-#2048 := (or #2089 #2136)
-#2079 := [th-lemma]: #2048
-#2081 := [unit-resolution #2079 #2082]: #2136
-[th-lemma #2081 #2004 #2110]: false
+#582 := [monotonicity #579]: #581
+#800 := [monotonicity #582 #797]: #799
+#803 := [monotonicity #800]: #802
+#529 := (iff #454 #528)
+#526 := (iff #171 #525)
+#523 := (iff #168 #522)
+#517 := (iff #24 #518)
+#521 := [rewrite]: #517
+#514 := (iff #167 #513)
+#511 := (iff #22 #510)
+#507 := (iff #21 #506)
+#509 := [rewrite]: #507
+#512 := [monotonicity #504 #509]: #511
+#515 := [monotonicity #512]: #514
+#524 := [monotonicity #515 #521]: #523
+#527 := [quant-intro #524]: #526
+#530 := [monotonicity #527]: #529
+#806 := [monotonicity #530 #803]: #805
+#809 := [monotonicity #527 #806]: #808
+#501 := (iff #466 #500)
+#498 := (iff #157 #495)
+#483 := (and true true)
+#486 := (and true #483)
+#489 := (and #9 #486)
+#492 := (and #480 #489)
+#496 := (iff #492 #495)
+#497 := [rewrite]: #496
+#493 := (iff #157 #492)
+#490 := (iff #154 #489)
+#487 := (iff #151 #486)
+#484 := (iff #148 #483)
+#477 := (iff #12 true)
+#478 := [rewrite]: #477
+#475 := (iff #10 true)
+#476 := [rewrite]: #475
+#485 := [monotonicity #476 #478]: #484
+#488 := [monotonicity #476 #485]: #487
+#491 := [monotonicity #488]: #490
+#481 := (iff #6 #480)
+#482 := [rewrite]: #481
+#494 := [monotonicity #482 #491]: #493
+#499 := [trans #494 #497]: #498
+#502 := [monotonicity #499]: #501
+#812 := [monotonicity #502 #809]: #811
+#815 := [monotonicity #812]: #814
+#473 := (iff #129 #472)
+#470 := (iff #128 #467)
+#463 := (implies #157 #460)
+#468 := (iff #463 #467)
+#469 := [rewrite]: #468
+#464 := (iff #128 #463)
+#461 := (iff #127 #460)
+#458 := (iff #126 #455)
+#451 := (implies #171 #448)
+#456 := (iff #451 #455)
+#457 := [rewrite]: #456
+#452 := (iff #126 #451)
+#449 := (iff #125 #448)
+#446 := (iff #124 #443)
+#439 := (implies #207 #436)
+#444 := (iff #439 #443)
+#445 := [rewrite]: #444
+#440 := (iff #124 #439)
+#437 := (iff #123 #436)
+#434 := (iff #122 #431)
+#427 := (implies #278 #424)
+#432 := (iff #427 #431)
+#433 := [rewrite]: #432
+#428 := (iff #122 #427)
+#425 := (iff #121 #424)
+#422 := (iff #120 #419)
+#415 := (implies #405 #370)
+#420 := (iff #415 #419)
+#421 := [rewrite]: #420
+#416 := (iff #120 #415)
+#371 := (iff #109 #370)
+#368 := (iff #108 #365)
+#361 := (implies #355 #358)
+#366 := (iff #361 #365)
+#367 := [rewrite]: #366
+#362 := (iff #108 #361)
+#359 := (iff #107 #358)
+#360 := [rewrite]: #359
+#356 := (iff #105 #355)
+#353 := (iff #104 #352)
+#354 := [rewrite]: #353
+#357 := [quant-intro #354]: #356
+#363 := [monotonicity #357 #360]: #362
+#369 := [trans #363 #367]: #368
+#372 := [monotonicity #357 #369]: #371
+#413 := (iff #119 #405)
+#408 := (and true #405)
+#411 := (iff #408 #405)
+#412 := [rewrite]: #411
+#409 := (iff #119 #408)
+#406 := (iff #118 #405)
+#403 := (iff #117 #402)
+#400 := (iff #116 #399)
+#397 := (iff #115 #396)
+#394 := (iff #114 #391)
+#388 := (and #385 #316)
+#392 := (iff #388 #391)
+#393 := [rewrite]: #392
+#389 := (iff #114 #388)
+#324 := (iff #93 #316)
+#319 := (and true #316)
+#322 := (iff #319 #316)
+#323 := [rewrite]: #322
+#320 := (iff #93 #319)
+#317 := (iff #92 #316)
+#314 := (iff #91 #311)
+#308 := (and #305 #90)
+#312 := (iff #308 #311)
+#313 := [rewrite]: #312
+#309 := (iff #91 #308)
+#306 := (iff #87 #305)
+#303 := (= #86 #302)
+#304 := [rewrite]: #303
+#307 := [monotonicity #304]: #306
+#310 := [monotonicity #307]: #309
+#315 := [trans #310 #313]: #314
+#300 := (iff #84 #299)
+#301 := [rewrite]: #300
+#318 := [monotonicity #301 #315]: #317
+#321 := [monotonicity #318]: #320
+#325 := [trans #321 #323]: #324
+#386 := (iff #113 #385)
+#387 := [rewrite]: #386
+#390 := [monotonicity #387 #325]: #389
+#395 := [trans #390 #393]: #394
+#383 := (iff #112 #382)
+#384 := [rewrite]: #383
+#398 := [monotonicity #384 #395]: #397
+#401 := [monotonicity #398]: #400
+#404 := [monotonicity #401]: #403
+#407 := [monotonicity #404]: #406
+#410 := [monotonicity #407]: #409
+#414 := [trans #410 #412]: #413
+#417 := [monotonicity #414 #372]: #416
+#423 := [trans #417 #421]: #422
+#380 := (iff #110 #377)
+#373 := (implies #341 #370)
+#378 := (iff #373 #377)
+#379 := [rewrite]: #378
+#374 := (iff #110 #373)
+#349 := (iff #100 #341)
+#344 := (and true #341)
+#347 := (iff #344 #341)
+#348 := [rewrite]: #347
+#345 := (iff #100 #344)
+#342 := (iff #99 #341)
+#339 := (iff #98 #338)
+#336 := (iff #97 #335)
+#333 := (iff #96 #332)
+#330 := (iff #95 #329)
+#327 := (iff #94 #326)
+#297 := (iff #82 #296)
+#298 := [rewrite]: #297
+#328 := [monotonicity #298 #325]: #327
+#294 := (iff #80 #293)
+#295 := [rewrite]: #294
+#331 := [monotonicity #295 #328]: #330
+#291 := (iff #78 #31)
+#292 := [rewrite]: #291
+#334 := [monotonicity #292 #331]: #333
+#289 := (iff #77 #288)
+#290 := [rewrite]: #289
+#337 := [monotonicity #290 #334]: #336
+#340 := [monotonicity #337]: #339
+#343 := [monotonicity #340]: #342
+#346 := [monotonicity #343]: #345
+#350 := [trans #346 #348]: #349
+#375 := [monotonicity #350 #372]: #374
+#381 := [trans #375 #379]: #380
+#426 := [monotonicity #381 #423]: #425
+#286 := (iff #73 #278)
+#281 := (and true #278)
+#284 := (iff #281 #278)
+#285 := [rewrite]: #284
+#282 := (iff #73 #281)
+#279 := (iff #72 #278)
+#276 := (iff #71 #275)
+#277 := [rewrite]: #276
+#280 := [monotonicity #277]: #279
+#283 := [monotonicity #280]: #282
+#287 := [trans #283 #285]: #286
+#429 := [monotonicity #287 #426]: #428
+#435 := [trans #429 #433]: #434
+#273 := (iff #69 #270)
+#266 := (implies #231 #263)
+#271 := (iff #266 #270)
+#272 := [rewrite]: #271
+#267 := (iff #69 #266)
+#264 := (iff #68 #263)
+#261 := (iff #67 #258)
+#254 := (implies #245 #251)
+#259 := (iff #254 #258)
+#260 := [rewrite]: #259
+#255 := (iff #67 #254)
+#252 := (iff #66 #251)
+#249 := (iff #65 #248)
+#250 := [rewrite]: #249
+#253 := [quant-intro #250]: #252
+#246 := (iff #63 #245)
+#243 := (iff #62 #242)
+#244 := [rewrite]: #243
+#247 := [quant-intro #244]: #246
+#256 := [monotonicity #247 #253]: #255
+#262 := [trans #256 #260]: #261
+#265 := [monotonicity #247 #262]: #264
+#239 := (iff #58 #231)
+#234 := (and true #231)
+#237 := (iff #234 #231)
+#238 := [rewrite]: #237
+#235 := (iff #58 #234)
+#232 := (iff #57 #231)
+#229 := (iff #56 #228)
+#226 := (iff #55 #225)
+#223 := (iff #54 #222)
+#220 := (iff #53 #219)
+#217 := (iff #52 #216)
+#218 := [rewrite]: #217
+#214 := (iff #50 #213)
+#215 := [rewrite]: #214
+#221 := [monotonicity #215 #218]: #220
+#211 := (iff #48 #210)
+#212 := [rewrite]: #211
+#224 := [monotonicity #212 #221]: #223
+#227 := [monotonicity #224]: #226
+#230 := [monotonicity #227]: #229
+#233 := [monotonicity #230]: #232
+#236 := [monotonicity #233]: #235
+#240 := [trans #236 #238]: #239
+#268 := [monotonicity #240 #265]: #267
+#274 := [trans #268 #272]: #273
+#438 := [monotonicity #274 #435]: #437
+#208 := (iff #45 #207)
+#205 := (iff #44 #197)
+#200 := (and true #197)
+#203 := (iff #200 #197)
+#204 := [rewrite]: #203
+#201 := (iff #44 #200)
+#198 := (iff #43 #197)
+#195 := (iff #42 #194)
+#192 := (iff #41 #189)
+#186 := (and #183 #32)
+#190 := (iff #186 #189)
+#191 := [rewrite]: #190
+#187 := (iff #41 #186)
+#184 := (iff #40 #183)
+#185 := [rewrite]: #184
+#188 := [monotonicity #185]: #187
+#193 := [trans #188 #191]: #192
+#181 := (iff #38 #180)
+#178 := (iff #37 #177)
+#179 := [rewrite]: #178
+#182 := [quant-intro #179]: #181
+#196 := [monotonicity #182 #193]: #195
+#199 := [monotonicity #196]: #198
+#202 := [monotonicity #199]: #201
+#206 := [trans #202 #204]: #205
+#174 := (iff #27 #9)
+#175 := [rewrite]: #174
+#209 := [monotonicity #175 #206]: #208
+#441 := [monotonicity #209 #438]: #440
+#447 := [trans #441 #445]: #446
+#450 := [monotonicity #175 #447]: #449
+#172 := (iff #26 #171)
+#169 := (iff #25 #168)
+#170 := [rewrite]: #169
+#173 := [quant-intro #170]: #172
+#453 := [monotonicity #173 #450]: #452
+#459 := [trans #453 #457]: #458
+#462 := [monotonicity #173 #459]: #461
+#165 := (iff #18 #157)
+#160 := (and true #157)
+#163 := (iff #160 #157)
+#164 := [rewrite]: #163
+#161 := (iff #18 #160)
+#158 := (iff #17 #157)
+#155 := (iff #16 #154)
+#152 := (iff #15 #151)
+#149 := (iff #14 #148)
+#146 := (iff #13 #12)
+#147 := [rewrite]: #146
+#150 := [monotonicity #147]: #149
+#153 := [monotonicity #150]: #152
+#156 := [monotonicity #153]: #155
+#159 := [monotonicity #156]: #158
+#162 := [monotonicity #159]: #161
+#166 := [trans #162 #164]: #165
+#465 := [monotonicity #166 #462]: #464
+#471 := [trans #465 #469]: #470
+#474 := [monotonicity #471]: #473
+#817 := [trans #474 #815]: #816
+#145 := [asserted]: #129
+#818 := [mp #145 #817]: #813
+#819 := [not-or-elim #818]: #495
+#820 := [and-elim #819]: #9
+#824 := [iff-true #820]: #826
+#831 := [monotonicity #824]: #830
+#836 := [trans #831 #834]: #835
+#839 := [monotonicity #836]: #838
+#842 := [monotonicity #839]: #841
+#847 := [trans #842 #845]: #846
+#850 := [monotonicity #824 #847]: #849
+#854 := [trans #850 #852]: #853
+#857 := [monotonicity #854]: #856
+#862 := [trans #857 #860]: #861
+#865 := [monotonicity #862]: #864
+#868 := [monotonicity #865]: #867
+#823 := [not-or-elim #818]: #822
+#869 := [mp #823 #868]: #866
+#1033 := [mp~ #869 #1032]: #1030
+#1034 := [mp #1033 #1201]: #1199
+#1514 := [mp #1034 #1513]: #1509
+#2079 := [mp #1514 #2078]: #2076
+#1835 := [unit-resolution #2079 #2418]: #1213
+#1658 := (or #1208 #874)
+#1659 := [def-axiom]: #1658
+#1728 := [unit-resolution #1659 #1835]: #874
+#1650 := (or #1208 #875)
+#1661 := [def-axiom]: #1650
+#1839 := [unit-resolution #1661 #1835]: #875
+#1840 := [th-lemma #1839 #1728]: #1860
+#1846 := [monotonicity #1840]: #1841
+#1843 := [symm #1846]: #1842
+#1813 := [trans #820 #1843]: #1833
+#1572 := (not #872)
+#1657 := (or #1208 #1572)
+#1573 := [def-axiom]: #1657
+#1772 := [unit-resolution #1573 #1835]: #1572
+#1806 := (not #1833)
+#1803 := (or #1806 #872)
+#1805 := [th-lemma]: #1803
+[unit-resolution #1805 #1772 #1813]: false
unsat
--- a/src/HOL/Boogie/Examples/cert/VCC_maximum Fri Dec 11 15:06:12 2009 +0100
+++ b/src/HOL/Boogie/Examples/cert/VCC_maximum Fri Dec 11 15:35:29 2009 +0100
@@ -687,6 +687,6 @@
:assumption (distinct uf_282)
:assumption (distinct uf_42)
:assumption (distinct uf_263 uf_14 uf_262 uf_261)
-:assumption (not (implies true (implies (and (<= 0 uf_283) (<= uf_283 uf_78)) (implies (and (<= 0 uf_284) (<= uf_284 uf_76)) (implies (and (<= 0 uf_285) (<= uf_285 uf_76)) (implies (< uf_286 1099511627776) (implies (< 0 uf_286) (implies (and (= (uf_27 uf_287 (uf_43 (uf_124 uf_7 uf_286) (uf_116 (uf_43 uf_7 uf_288)))) uf_9) (and (= (uf_25 uf_287 (uf_43 (uf_124 uf_7 uf_286) (uf_116 (uf_43 uf_7 uf_288)))) uf_26) (and (= (uf_48 (uf_43 (uf_124 uf_7 uf_286) (uf_116 (uf_43 uf_7 uf_288))) (uf_124 uf_7 uf_286)) uf_9) (and (= (uf_24 uf_287 (uf_43 (uf_124 uf_7 uf_286) (uf_116 (uf_43 uf_7 uf_288)))) uf_9) (and (not (= (uf_12 (uf_124 uf_7 uf_286)) uf_14)) (= (uf_23 (uf_124 uf_7 uf_286)) uf_9)))))) (implies true (implies (= (uf_203 uf_287) uf_9) (implies (and (= (uf_202 uf_281 uf_287) uf_9) (= (uf_55 uf_287) uf_9)) (implies (forall (?x771 T19) (< (uf_289 ?x771) uf_290) :pat { (uf_289 ?x771) }) (implies (and (up_291 uf_287 uf_281 uf_280 (uf_29 (uf_43 uf_7 uf_288)) (uf_6 uf_7)) (up_292 uf_287 uf_281 uf_280 (uf_43 uf_7 uf_288) (uf_6 uf_7))) (implies (up_291 uf_287 uf_281 uf_279 uf_286 uf_4) (implies (= uf_293 (uf_173 uf_287)) (implies (forall (?x772 T5) (iff (= (uf_294 uf_293 ?x772) uf_9) false) :pat { (uf_294 uf_293 ?x772) }) (implies (and (<= 0 uf_286) (<= uf_286 uf_76)) (and (= (uf_200 uf_287 (uf_43 (uf_124 uf_7 uf_286) (uf_116 (uf_43 uf_7 uf_288))) (uf_43 (uf_124 uf_7 uf_286) (uf_116 (uf_43 uf_7 uf_288))) uf_282) uf_9) (implies (= (uf_200 uf_287 (uf_43 (uf_124 uf_7 uf_286) (uf_116 (uf_43 uf_7 uf_288))) (uf_43 (uf_124 uf_7 uf_286) (uf_116 (uf_43 uf_7 uf_288))) uf_282) uf_9) (and (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_288) 0 uf_7) uf_7) uf_9) (= (uf_24 uf_287 (uf_66 (uf_43 uf_7 uf_288) 0 uf_7)) uf_9)) (implies (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_288) 0 uf_7) uf_7) uf_9) (= (uf_24 uf_287 (uf_66 (uf_43 uf_7 uf_288) 0 uf_7)) uf_9)) (and (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_288) 0 uf_7) uf_7) uf_9) (= (uf_68 uf_287 (uf_66 (uf_43 uf_7 uf_288) 0 uf_7)) uf_9)) (implies (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_288) 0 uf_7) uf_7) uf_9) (= (uf_68 uf_287 (uf_66 (uf_43 uf_7 uf_288) 0 uf_7)) uf_9)) (implies (= uf_295 (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) 0 uf_7))) (implies (up_291 uf_287 uf_278 uf_277 uf_295 uf_7) (implies (up_291 uf_287 uf_276 uf_275 0 uf_4) (implies (up_291 uf_287 uf_274 uf_273 1 uf_4) (implies (and (<= 1 1) (and (<= 1 1) (and (<= 0 0) (<= 0 0)))) (and (<= 1 uf_286) (implies (<= 1 uf_286) (and (forall (?x773 Int) (implies (and (<= 0 ?x773) (<= ?x773 uf_76)) (implies (< ?x773 1) (<= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) ?x773 uf_7)) uf_295)))) (implies (forall (?x774 Int) (implies (and (<= 0 ?x774) (<= ?x774 uf_76)) (implies (< ?x774 1) (<= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) ?x774 uf_7)) uf_295)))) (and (and (< 0 uf_286) (= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) 0 uf_7)) uf_295)) (implies (and (< 0 uf_286) (= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) 0 uf_7)) uf_295)) (implies true (implies (and (<= 0 uf_296) (<= uf_296 uf_78)) (implies (and (<= 0 uf_297) (<= uf_297 uf_76)) (implies (and (<= 0 uf_298) (<= uf_298 uf_76)) (implies (and (<= 1 uf_298) (<= 0 uf_297)) (implies (<= uf_298 uf_286) (implies (forall (?x775 Int) (implies (and (<= 0 ?x775) (<= ?x775 uf_76)) (implies (< ?x775 uf_298) (<= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) ?x775 uf_7)) uf_296)))) (implies (and (< uf_297 uf_286) (= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) uf_297 uf_7)) uf_296)) (implies (and (<= 1 uf_298) (<= 0 uf_297)) (and (implies true (implies (and (<= 1 uf_298) (<= 0 uf_297)) (implies (not true) (implies (and (<= 1 uf_298) (<= 0 uf_297)) (implies true (implies (and (<= 1 uf_298) (<= 0 uf_297)) (implies (and (= (uf_202 uf_272 uf_287) uf_9) (= (uf_55 uf_287) uf_9)) (implies (and (<= 1 uf_298) (<= 0 uf_297)) (implies true (implies (and (<= 1 uf_298) (<= 0 uf_297)) (and up_216 (implies up_216 (implies (and (<= 1 uf_298) (<= 0 uf_297)) (implies true (implies (and (<= 1 uf_298) (<= 0 uf_297)) (implies (and (<= 1 uf_298) (<= 0 uf_297)) (implies true (implies (= uf_299 uf_296) (implies (= uf_300 uf_298) (implies (= uf_301 uf_297) (implies (= uf_302 uf_296) (implies true (and (forall (?x776 Int) (implies (and (<= 0 ?x776) (<= ?x776 uf_76)) (implies (< ?x776 uf_286) (<= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) ?x776 uf_7)) uf_302)))) (implies (forall (?x777 Int) (implies (and (<= 0 ?x777) (<= ?x777 uf_76)) (implies (< ?x777 uf_286) (<= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) ?x777 uf_7)) uf_302)))) (and (exists (?x778 Int) (and (<= 0 ?x778) (and (<= ?x778 uf_76) (and (< ?x778 uf_286) (= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) ?x778 uf_7)) uf_302))))) (implies (exists (?x779 Int) (and (<= 0 ?x779) (and (<= ?x779 uf_76) (and (< ?x779 uf_286) (= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) ?x779 uf_7)) uf_302))))) true)))))))))))))))))))))))))) (implies true (implies (and (<= 1 uf_298) (<= 0 uf_297)) (implies true (implies (and (<= 1 uf_298) (<= 0 uf_297)) (implies true (implies (and (<= 1 uf_298) (<= 0 uf_297)) (implies (and (forall (?x780 T5) (implies (not (= (uf_12 (uf_13 (uf_25 uf_287 ?x780))) uf_261)) (not (= (uf_12 (uf_13 (uf_25 uf_287 ?x780))) uf_261))) :pat { (uf_40 (uf_41 uf_287) ?x780) }) (and (forall (?x781 T5) (implies (= (uf_68 uf_287 ?x781) uf_9) (and (= (uf_19 (uf_20 uf_287) ?x781) (uf_19 (uf_20 uf_287) ?x781)) (= (uf_68 uf_287 ?x781) uf_9))) :pat { (uf_19 (uf_20 uf_287) ?x781) }) (and (forall (?x782 T5) (implies (= (uf_68 uf_287 ?x782) uf_9) (and (= (uf_40 (uf_41 uf_287) ?x782) (uf_40 (uf_41 uf_287) ?x782)) (= (uf_68 uf_287 ?x782) uf_9))) :pat { (uf_40 (uf_41 uf_287) ?x782) }) (and (forall (?x783 T5) (implies (= (uf_68 uf_287 ?x783) uf_9) (and (= (uf_58 (uf_59 uf_287) ?x783) (uf_58 (uf_59 uf_287) ?x783)) (= (uf_68 uf_287 ?x783) uf_9))) :pat { (uf_58 (uf_59 uf_287) ?x783) }) (and (<= (uf_173 uf_287) (uf_173 uf_287)) (and (forall (?x784 T5) (<= (uf_172 uf_287 ?x784) (uf_172 uf_287 ?x784)) :pat { (uf_172 uf_287 ?x784) }) (= (uf_178 uf_287 uf_287) uf_9))))))) (implies (and (<= (uf_173 uf_287) (uf_173 uf_287)) (and (forall (?x785 T5) (<= (uf_172 uf_287 ?x785) (uf_172 uf_287 ?x785)) :pat { (uf_172 uf_287 ?x785) }) (= (uf_178 uf_287 uf_287) uf_9))) (implies (and (= (uf_202 uf_272 uf_287) uf_9) (= (uf_55 uf_287) uf_9)) (implies (up_291 uf_287 uf_272 uf_273 uf_298 uf_4) (implies (up_291 uf_287 uf_272 uf_275 uf_297 uf_4) (implies (up_291 uf_287 uf_272 uf_277 uf_296 uf_7) (implies (up_291 uf_287 uf_272 uf_279 uf_286 uf_4) (implies (and (up_291 uf_287 uf_272 uf_280 (uf_29 (uf_43 uf_7 uf_288)) (uf_6 uf_7)) (up_292 uf_287 uf_272 uf_280 (uf_43 uf_7 uf_288) (uf_6 uf_7))) (implies (and (= (uf_59 uf_287) (uf_59 uf_287)) (= (uf_41 uf_287) (uf_41 uf_287))) (implies (and (<= 1 uf_298) (<= 0 uf_297)) (and (implies true (implies (and (<= 1 uf_298) (<= 0 uf_297)) (implies (< uf_298 uf_286) (implies (and (<= 1 uf_298) (<= 0 uf_297)) (implies true (implies (and (<= 1 uf_298) (<= 0 uf_297)) (and (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7) uf_7) uf_9) (= (uf_24 uf_287 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7)) uf_9)) (implies (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7) uf_7) uf_9) (= (uf_24 uf_287 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7)) uf_9)) (and (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7) uf_7) uf_9) (= (uf_68 uf_287 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7)) uf_9)) (implies (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7) uf_7) uf_9) (= (uf_68 uf_287 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7)) uf_9)) (implies (and (<= 1 uf_298) (<= 0 uf_297)) (and (implies true (implies (and (<= 1 uf_298) (<= 0 uf_297)) (implies (< uf_296 (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7))) (implies (and (<= 1 uf_298) (<= 0 uf_297)) (implies true (implies (and (<= 1 uf_298) (<= 0 uf_297)) (and (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7) uf_7) uf_9) (= (uf_24 uf_287 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7)) uf_9)) (implies (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7) uf_7) uf_9) (= (uf_24 uf_287 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7)) uf_9)) (and (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7) uf_7) uf_9) (= (uf_68 uf_287 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7)) uf_9)) (implies (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7) uf_7) uf_9) (= (uf_68 uf_287 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7)) uf_9)) (implies (= uf_303 (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7))) (implies (up_291 uf_287 uf_271 uf_277 uf_303 uf_7) (implies (up_291 uf_287 uf_270 uf_275 uf_298 uf_4) (implies (and (<= 1 uf_298) (<= 1 uf_298)) (implies true (implies (= uf_304 uf_303) (implies (= uf_305 uf_298) (implies true (implies (and (<= 1 uf_298) (<= 0 uf_305)) (and (and (<= 0 (+ uf_298 1)) (<= (+ uf_298 1) uf_76)) (implies (and (<= 0 (+ uf_298 1)) (<= (+ uf_298 1) uf_76)) (implies (= uf_306 (+ uf_298 1)) (implies (up_291 uf_287 uf_269 uf_273 uf_306 uf_4) (implies (and (<= 2 uf_306) (<= 0 uf_305)) (implies true (and (<= uf_306 uf_286) (implies (<= uf_306 uf_286) (and (forall (?x786 Int) (implies (and (<= 0 ?x786) (<= ?x786 uf_76)) (implies (< ?x786 uf_306) (<= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) ?x786 uf_7)) uf_304)))) (implies (forall (?x787 Int) (implies (and (<= 0 ?x787) (<= ?x787 uf_76)) (implies (< ?x787 uf_306) (<= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) ?x787 uf_7)) uf_304)))) (and (and (< uf_305 uf_286) (= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) uf_305 uf_7)) uf_304)) (implies (and (< uf_305 uf_286) (= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) uf_305 uf_7)) uf_304)) (implies false true)))))))))))))))))))))))))))))))) (implies true (implies (and (<= 1 uf_298) (<= 0 uf_297)) (implies (<= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7)) uf_296) (implies (and (<= 1 uf_298) (<= 0 uf_297)) (implies true (implies (and (<= 1 uf_298) (<= 0 uf_297)) (implies (and (<= 1 uf_298) (<= 0 uf_297)) (implies true (implies (= uf_304 uf_296) (implies (= uf_305 uf_297) (implies true (implies (and (<= 1 uf_298) (<= 0 uf_305)) (and (and (<= 0 (+ uf_298 1)) (<= (+ uf_298 1) uf_76)) (implies (and (<= 0 (+ uf_298 1)) (<= (+ uf_298 1) uf_76)) (implies (= uf_306 (+ uf_298 1)) (implies (up_291 uf_287 uf_269 uf_273 uf_306 uf_4) (implies (and (<= 2 uf_306) (<= 0 uf_305)) (implies true (and (<= uf_306 uf_286) (implies (<= uf_306 uf_286) (and (forall (?x788 Int) (implies (and (<= 0 ?x788) (<= ?x788 uf_76)) (implies (< ?x788 uf_306) (<= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) ?x788 uf_7)) uf_304)))) (implies (forall (?x789 Int) (implies (and (<= 0 ?x789) (<= ?x789 uf_76)) (implies (< ?x789 uf_306) (<= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) ?x789 uf_7)) uf_304)))) (and (and (< uf_305 uf_286) (= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) uf_305 uf_7)) uf_304)) (implies (and (< uf_305 uf_286) (= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) uf_305 uf_7)) uf_304)) (implies false true))))))))))))))))))))))))))))))))))))) (implies true (implies (and (<= 1 uf_298) (<= 0 uf_297)) (implies (<= uf_286 uf_298) (implies (and (<= 1 uf_298) (<= 0 uf_297)) (implies true (implies (and (<= 1 uf_298) (<= 0 uf_297)) (implies (and (<= 1 uf_298) (<= 0 uf_297)) (implies true (implies (and (<= 1 uf_298) (<= 0 uf_297)) (and up_216 (implies up_216 (implies (and (<= 1 uf_298) (<= 0 uf_297)) (implies true (implies (and (<= 1 uf_298) (<= 0 uf_297)) (implies (and (<= 1 uf_298) (<= 0 uf_297)) (implies true (implies (= uf_299 uf_296) (implies (= uf_300 uf_298) (implies (= uf_301 uf_297) (implies (= uf_302 uf_296) (implies true (and (forall (?x790 Int) (implies (and (<= 0 ?x790) (<= ?x790 uf_76)) (implies (< ?x790 uf_286) (<= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) ?x790 uf_7)) uf_302)))) (implies (forall (?x791 Int) (implies (and (<= 0 ?x791) (<= ?x791 uf_76)) (implies (< ?x791 uf_286) (<= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) ?x791 uf_7)) uf_302)))) (and (exists (?x792 Int) (and (<= 0 ?x792) (and (<= ?x792 uf_76) (and (< ?x792 uf_286) (= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) ?x792 uf_7)) uf_302))))) (implies (exists (?x793 Int) (and (<= 0 ?x793) (and (<= ?x793 uf_76) (and (< ?x793 uf_286) (= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) ?x793 uf_7)) uf_302))))) true))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))
+:assumption (not (implies (and (and (<= 0 uf_283) (<= uf_283 uf_78)) (and (and (<= 0 uf_284) (<= uf_284 uf_76)) (and (and (<= 0 uf_285) (<= uf_285 uf_76)) (and (< uf_286 1099511627776) (and (< 0 uf_286) (and (and (= (uf_27 uf_287 (uf_43 (uf_124 uf_7 uf_286) (uf_116 (uf_43 uf_7 uf_288)))) uf_9) (and (= (uf_25 uf_287 (uf_43 (uf_124 uf_7 uf_286) (uf_116 (uf_43 uf_7 uf_288)))) uf_26) (and (= (uf_48 (uf_43 (uf_124 uf_7 uf_286) (uf_116 (uf_43 uf_7 uf_288))) (uf_124 uf_7 uf_286)) uf_9) (and (= (uf_24 uf_287 (uf_43 (uf_124 uf_7 uf_286) (uf_116 (uf_43 uf_7 uf_288)))) uf_9) (and (not (= (uf_12 (uf_124 uf_7 uf_286)) uf_14)) (= (uf_23 (uf_124 uf_7 uf_286)) uf_9)))))) (and (= (uf_203 uf_287) uf_9) (and (and (= (uf_202 uf_281 uf_287) uf_9) (= (uf_55 uf_287) uf_9)) (and (forall (?x771 T19) (< (uf_289 ?x771) uf_290) :pat { (uf_289 ?x771) }) (and (and (up_291 uf_287 uf_281 uf_280 (uf_29 (uf_43 uf_7 uf_288)) (uf_6 uf_7)) (up_292 uf_287 uf_281 uf_280 (uf_43 uf_7 uf_288) (uf_6 uf_7))) (and (up_291 uf_287 uf_281 uf_279 uf_286 uf_4) (and (= uf_293 (uf_173 uf_287)) (and (forall (?x772 T5) (iff (= (uf_294 uf_293 ?x772) uf_9) false) :pat { (uf_294 uf_293 ?x772) }) (and (<= 0 uf_286) (<= uf_286 uf_76))))))))))))))) (and (= (uf_200 uf_287 (uf_43 (uf_124 uf_7 uf_286) (uf_116 (uf_43 uf_7 uf_288))) (uf_43 (uf_124 uf_7 uf_286) (uf_116 (uf_43 uf_7 uf_288))) uf_282) uf_9) (implies (= (uf_200 uf_287 (uf_43 (uf_124 uf_7 uf_286) (uf_116 (uf_43 uf_7 uf_288))) (uf_43 (uf_124 uf_7 uf_286) (uf_116 (uf_43 uf_7 uf_288))) uf_282) uf_9) (and (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_288) 0 uf_7) uf_7) uf_9) (= (uf_24 uf_287 (uf_66 (uf_43 uf_7 uf_288) 0 uf_7)) uf_9)) (implies (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_288) 0 uf_7) uf_7) uf_9) (= (uf_24 uf_287 (uf_66 (uf_43 uf_7 uf_288) 0 uf_7)) uf_9)) (and (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_288) 0 uf_7) uf_7) uf_9) (= (uf_68 uf_287 (uf_66 (uf_43 uf_7 uf_288) 0 uf_7)) uf_9)) (implies (and (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_288) 0 uf_7) uf_7) uf_9) (= (uf_68 uf_287 (uf_66 (uf_43 uf_7 uf_288) 0 uf_7)) uf_9)) (and (= uf_295 (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) 0 uf_7))) (and (up_291 uf_287 uf_278 uf_277 uf_295 uf_7) (and (up_291 uf_287 uf_276 uf_275 0 uf_4) (and (up_291 uf_287 uf_274 uf_273 1 uf_4) (and (<= 1 1) (and (<= 1 1) (and (<= 0 0) (<= 0 0))))))))) (and (<= 1 uf_286) (implies (<= 1 uf_286) (and (forall (?x773 Int) (implies (and (and (<= 0 ?x773) (<= ?x773 uf_76)) (< ?x773 1)) (<= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) ?x773 uf_7)) uf_295))) (implies (forall (?x774 Int) (implies (and (and (<= 0 ?x774) (<= ?x774 uf_76)) (< ?x774 1)) (<= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) ?x774 uf_7)) uf_295))) (and (and (< 0 uf_286) (= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) 0 uf_7)) uf_295)) (implies (and (and (< 0 uf_286) (= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) 0 uf_7)) uf_295)) (and (and (<= 0 uf_296) (<= uf_296 uf_78)) (and (and (<= 0 uf_297) (<= uf_297 uf_76)) (and (and (<= 0 uf_298) (<= uf_298 uf_76)) (and (and (<= 1 uf_298) (<= 0 uf_297)) (and (<= uf_298 uf_286) (and (forall (?x775 Int) (implies (and (and (<= 0 ?x775) (<= ?x775 uf_76)) (< ?x775 uf_298)) (<= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) ?x775 uf_7)) uf_296))) (and (and (< uf_297 uf_286) (= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) uf_297 uf_7)) uf_296)) (and (<= 1 uf_298) (<= 0 uf_297)))))))))) (and (implies (and (and (<= 1 uf_298) (<= 0 uf_297)) (and (not true) (and (and (<= 1 uf_298) (<= 0 uf_297)) (and (and (<= 1 uf_298) (<= 0 uf_297)) (and (and (= (uf_202 uf_272 uf_287) uf_9) (= (uf_55 uf_287) uf_9)) (and (and (<= 1 uf_298) (<= 0 uf_297)) (and (<= 1 uf_298) (<= 0 uf_297)))))))) (and up_216 (implies (and up_216 (and (and (<= 1 uf_298) (<= 0 uf_297)) (and (and (<= 1 uf_298) (<= 0 uf_297)) (and (and (<= 1 uf_298) (<= 0 uf_297)) (and (= uf_299 uf_296) (and (= uf_300 uf_298) (and (= uf_301 uf_297) (= uf_302 uf_296)))))))) (and (forall (?x776 Int) (implies (and (and (<= 0 ?x776) (<= ?x776 uf_76)) (< ?x776 uf_286)) (<= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) ?x776 uf_7)) uf_302))) (implies (forall (?x777 Int) (implies (and (and (<= 0 ?x777) (<= ?x777 uf_76)) (< ?x777 uf_286)) (<= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) ?x777 uf_7)) uf_302))) (exists (?x778 Int) (and (<= 0 ?x778) (and (<= ?x778 uf_76) (and (< ?x778 uf_286) (= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) ?x778 uf_7)) uf_302)))))))))) (implies (and (and (<= 1 uf_298) (<= 0 uf_297)) (and (and (<= 1 uf_298) (<= 0 uf_297)) (and (and (<= 1 uf_298) (<= 0 uf_297)) (and (and (forall (?x779 T5) (implies (not (= (uf_12 (uf_13 (uf_25 uf_287 ?x779))) uf_261)) (not (= (uf_12 (uf_13 (uf_25 uf_287 ?x779))) uf_261))) :pat { (uf_40 (uf_41 uf_287) ?x779) }) (and (forall (?x780 T5) (implies (= (uf_68 uf_287 ?x780) uf_9) (and (= (uf_19 (uf_20 uf_287) ?x780) (uf_19 (uf_20 uf_287) ?x780)) (= (uf_68 uf_287 ?x780) uf_9))) :pat { (uf_19 (uf_20 uf_287) ?x780) }) (and (forall (?x781 T5) (implies (= (uf_68 uf_287 ?x781) uf_9) (and (= (uf_40 (uf_41 uf_287) ?x781) (uf_40 (uf_41 uf_287) ?x781)) (= (uf_68 uf_287 ?x781) uf_9))) :pat { (uf_40 (uf_41 uf_287) ?x781) }) (and (forall (?x782 T5) (implies (= (uf_68 uf_287 ?x782) uf_9) (and (= (uf_58 (uf_59 uf_287) ?x782) (uf_58 (uf_59 uf_287) ?x782)) (= (uf_68 uf_287 ?x782) uf_9))) :pat { (uf_58 (uf_59 uf_287) ?x782) }) (and (<= (uf_173 uf_287) (uf_173 uf_287)) (and (forall (?x783 T5) (<= (uf_172 uf_287 ?x783) (uf_172 uf_287 ?x783)) :pat { (uf_172 uf_287 ?x783) }) (= (uf_178 uf_287 uf_287) uf_9))))))) (and (and (<= (uf_173 uf_287) (uf_173 uf_287)) (and (forall (?x784 T5) (<= (uf_172 uf_287 ?x784) (uf_172 uf_287 ?x784)) :pat { (uf_172 uf_287 ?x784) }) (= (uf_178 uf_287 uf_287) uf_9))) (and (and (= (uf_202 uf_272 uf_287) uf_9) (= (uf_55 uf_287) uf_9)) (and (up_291 uf_287 uf_272 uf_273 uf_298 uf_4) (and (up_291 uf_287 uf_272 uf_275 uf_297 uf_4) (and (up_291 uf_287 uf_272 uf_277 uf_296 uf_7) (and (up_291 uf_287 uf_272 uf_279 uf_286 uf_4) (and (and (up_291 uf_287 uf_272 uf_280 (uf_29 (uf_43 uf_7 uf_288)) (uf_6 uf_7)) (up_292 uf_287 uf_272 uf_280 (uf_43 uf_7 uf_288) (uf_6 uf_7))) (and (and (= (uf_59 uf_287) (uf_59 uf_287)) (= (uf_41 uf_287) (uf_41 uf_287))) (and (<= 1 uf_298) (<= 0 uf_297)))))))))))))) (and (implies (and (and (<= 1 uf_298) (<= 0 uf_297)) (and (< uf_298 uf_286) (and (and (<= 1 uf_298) (<= 0 uf_297)) (and (<= 1 uf_298) (<= 0 uf_297))))) (and (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7) uf_7) uf_9) (= (uf_24 uf_287 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7)) uf_9)) (implies (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7) uf_7) uf_9) (= (uf_24 uf_287 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7)) uf_9)) (and (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7) uf_7) uf_9) (= (uf_68 uf_287 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7)) uf_9)) (implies (and (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7) uf_7) uf_9) (= (uf_68 uf_287 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7)) uf_9)) (and (<= 1 uf_298) (<= 0 uf_297))) (and (implies (and (and (<= 1 uf_298) (<= 0 uf_297)) (and (< uf_296 (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7))) (and (and (<= 1 uf_298) (<= 0 uf_297)) (and (<= 1 uf_298) (<= 0 uf_297))))) (and (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7) uf_7) uf_9) (= (uf_24 uf_287 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7)) uf_9)) (implies (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7) uf_7) uf_9) (= (uf_24 uf_287 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7)) uf_9)) (and (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7) uf_7) uf_9) (= (uf_68 uf_287 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7)) uf_9)) (implies (and (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7) uf_7) uf_9) (= (uf_68 uf_287 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7)) uf_9)) (and (= uf_303 (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7))) (and (up_291 uf_287 uf_271 uf_277 uf_303 uf_7) (and (up_291 uf_287 uf_270 uf_275 uf_298 uf_4) (and (and (<= 1 uf_298) (<= 1 uf_298)) (and (= uf_304 uf_303) (and (= uf_305 uf_298) (and (<= 1 uf_298) (<= 0 uf_305))))))))) (and (and (<= 0 (+ uf_298 1)) (<= (+ uf_298 1) uf_76)) (implies (and (and (<= 0 (+ uf_298 1)) (<= (+ uf_298 1) uf_76)) (and (= uf_306 (+ uf_298 1)) (and (up_291 uf_287 uf_269 uf_273 uf_306 uf_4) (and (<= 2 uf_306) (<= 0 uf_305))))) (and (<= uf_306 uf_286) (implies (<= uf_306 uf_286) (and (forall (?x785 Int) (implies (and (and (<= 0 ?x785) (<= ?x785 uf_76)) (< ?x785 uf_306)) (<= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) ?x785 uf_7)) uf_304))) (implies (forall (?x786 Int) (implies (and (and (<= 0 ?x786) (<= ?x786 uf_76)) (< ?x786 uf_306)) (<= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) ?x786 uf_7)) uf_304))) (and (< uf_305 uf_286) (= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) uf_305 uf_7)) uf_304))))))))))))) (implies (and (and (<= 1 uf_298) (<= 0 uf_297)) (and (<= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) uf_298 uf_7)) uf_296) (and (and (<= 1 uf_298) (<= 0 uf_297)) (and (and (<= 1 uf_298) (<= 0 uf_297)) (and (and (<= 1 uf_298) (<= 0 uf_297)) (and (= uf_304 uf_296) (and (= uf_305 uf_297) (and (<= 1 uf_298) (<= 0 uf_305))))))))) (and (and (<= 0 (+ uf_298 1)) (<= (+ uf_298 1) uf_76)) (implies (and (and (<= 0 (+ uf_298 1)) (<= (+ uf_298 1) uf_76)) (and (= uf_306 (+ uf_298 1)) (and (up_291 uf_287 uf_269 uf_273 uf_306 uf_4) (and (<= 2 uf_306) (<= 0 uf_305))))) (and (<= uf_306 uf_286) (implies (<= uf_306 uf_286) (and (forall (?x787 Int) (implies (and (and (<= 0 ?x787) (<= ?x787 uf_76)) (< ?x787 uf_306)) (<= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) ?x787 uf_7)) uf_304))) (implies (forall (?x788 Int) (implies (and (and (<= 0 ?x788) (<= ?x788 uf_76)) (< ?x788 uf_306)) (<= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) ?x788 uf_7)) uf_304))) (and (< uf_305 uf_286) (= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) uf_305 uf_7)) uf_304))))))))))))))) (implies (and (and (<= 1 uf_298) (<= 0 uf_297)) (and (<= uf_286 uf_298) (and (and (<= 1 uf_298) (<= 0 uf_297)) (and (and (<= 1 uf_298) (<= 0 uf_297)) (and (and (<= 1 uf_298) (<= 0 uf_297)) (and (<= 1 uf_298) (<= 0 uf_297))))))) (and up_216 (implies (and up_216 (and (and (<= 1 uf_298) (<= 0 uf_297)) (and (and (<= 1 uf_298) (<= 0 uf_297)) (and (and (<= 1 uf_298) (<= 0 uf_297)) (and (= uf_299 uf_296) (and (= uf_300 uf_298) (and (= uf_301 uf_297) (= uf_302 uf_296)))))))) (and (forall (?x789 Int) (implies (and (and (<= 0 ?x789) (<= ?x789 uf_76)) (< ?x789 uf_286)) (<= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) ?x789 uf_7)) uf_302))) (implies (forall (?x790 Int) (implies (and (and (<= 0 ?x790) (<= ?x790 uf_76)) (< ?x790 uf_286)) (<= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) ?x790 uf_7)) uf_302))) (exists (?x791 Int) (and (<= 0 ?x791) (and (<= ?x791 uf_76) (and (< ?x791 uf_286) (= (uf_110 uf_287 (uf_66 (uf_43 uf_7 uf_288) ?x791 uf_7)) uf_302)))))))))))))))))))))))))))
:formula true
)
--- a/src/HOL/Boogie/Examples/cert/VCC_maximum.proof Fri Dec 11 15:06:12 2009 +0100
+++ b/src/HOL/Boogie/Examples/cert/VCC_maximum.proof Fri Dec 11 15:35:29 2009 +0100
@@ -1,106 +1,309 @@
#2 := false
+decl ?x785!14 :: int
+#17756 := ?x785!14
+decl uf_298 :: int
+#3087 := uf_298
+#30181 := (= uf_298 ?x785!14)
+#121 := 0::int
+#4042 := -1::int
+#18020 := (* -1::int ?x785!14)
+decl uf_306 :: int
+#3256 := uf_306
+#18021 := (+ uf_306 #18020)
+#18022 := (<= #18021 0::int)
decl uf_110 :: (-> T4 T5 int)
decl uf_66 :: (-> T5 int T3 T5)
decl uf_7 :: T3
#10 := uf_7
-decl ?x786!14 :: int
-#18507 := ?x786!14
decl uf_43 :: (-> T3 int T5)
decl uf_288 :: int
#2978 := uf_288
#2979 := (uf_43 uf_7 uf_288)
-#18512 := (uf_66 #2979 ?x786!14 uf_7)
+#17764 := (uf_66 #2979 ?x785!14 uf_7)
decl uf_287 :: T4
#2976 := uf_287
-#18513 := (uf_110 uf_287 #18512)
-decl uf_298 :: int
-#3069 := uf_298
-#3188 := (uf_66 #2979 uf_298 uf_7)
-#3197 := (uf_110 uf_287 #3188)
-#30708 := (= #3197 #18513)
-#30713 := (not #30708)
-#121 := 0::int
-#4071 := -1::int
-#18834 := (* -1::int #18513)
-#30712 := (+ #3197 #18834)
-#30714 := (>= #30712 0::int)
-#30724 := (not #30714)
+#17765 := (uf_110 uf_287 #17764)
+#18007 := (* -1::int #17765)
decl uf_304 :: int
-#3204 := uf_304
-#13490 := (* -1::int uf_304)
-#30366 := (+ #3197 #13490)
-#30319 := (<= #30366 0::int)
-#30365 := (= #3197 uf_304)
-decl uf_303 :: int
-#3199 := uf_303
-#12389 := (= uf_303 uf_304)
-#18835 := (+ uf_304 #18834)
-#18836 := (>= #18835 0::int)
-#18821 := (* -1::int ?x786!14)
-decl uf_306 :: int
-#3214 := uf_306
-#18822 := (+ uf_306 #18821)
-#18823 := (<= #18822 0::int)
-#18509 := (>= ?x786!14 0::int)
-#22289 := (not #18509)
-#7500 := 4294967295::int
-#18508 := (<= ?x786!14 4294967295::int)
-#22288 := (not #18508)
-#22304 := (or #22288 #22289 #18823 #18836)
-#22309 := (not #22304)
+#3239 := uf_304
+#18008 := (+ uf_304 #18007)
+#18009 := (>= #18008 0::int)
+#17761 := (>= ?x785!14 0::int)
+#21511 := (not #17761)
+#7471 := 4294967295::int
+#17757 := (<= ?x785!14 4294967295::int)
+#21510 := (not #17757)
+#21526 := (or #21510 #21511 #18009 #18022)
+#21531 := (not #21526)
#161 := (:var 0 int)
-#3053 := (uf_66 #2979 #161 uf_7)
-#23211 := (pattern #3053)
-#15127 := (<= #161 4294967295::int)
-#19506 := (not #15127)
-#3054 := (uf_110 uf_287 #3053)
-#13491 := (+ #3054 #13490)
-#13492 := (<= #13491 0::int)
-#13471 := (* -1::int uf_306)
-#13479 := (+ #161 #13471)
-#13478 := (>= #13479 0::int)
-#4070 := (>= #161 0::int)
-#4992 := (not #4070)
-#22270 := (or #4992 #13478 #13492 #19506)
-#23228 := (forall (vars (?x786 int)) (:pat #23211) #22270)
-#23233 := (not #23228)
+#3072 := (uf_66 #2979 #161 uf_7)
+#22468 := (pattern #3072)
+#14378 := (<= #161 4294967295::int)
+#18695 := (not #14378)
+#13112 := (* -1::int uf_304)
+#3073 := (uf_110 uf_287 #3072)
+#13113 := (+ #3073 #13112)
+#13114 := (<= #13113 0::int)
+#13062 := (* -1::int uf_306)
+#13097 := (+ #161 #13062)
+#13096 := (>= #13097 0::int)
+#4041 := (>= #161 0::int)
+#4963 := (not #4041)
+#21492 := (or #4963 #13096 #13114 #18695)
+#22485 := (forall (vars (?x785 int)) (:pat #22468) #21492)
+#22490 := (not #22485)
decl uf_305 :: int
-#3206 := uf_305
-#13512 := (* -1::int uf_305)
+#3241 := uf_305
+#13126 := (* -1::int uf_305)
decl uf_286 :: int
#2973 := uf_286
-#13513 := (+ uf_286 #13512)
-#13514 := (<= #13513 0::int)
-#3226 := (uf_66 #2979 uf_305 uf_7)
-#3227 := (uf_110 uf_287 #3226)
-#12428 := (= uf_304 #3227)
-#22255 := (not #12428)
-#22256 := (or #22255 #13514)
-#22257 := (not #22256)
-#23236 := (or #22257 #23233)
-#23239 := (not #23236)
-#23242 := (or #23239 #22309)
-#23245 := (not #23242)
-#13472 := (+ uf_286 #13471)
-#13470 := (>= #13472 0::int)
-#13475 := (not #13470)
-#23248 := (or #13475 #23245)
-#23251 := (not #23248)
-#23254 := (or #13475 #23251)
-#23257 := (not #23254)
-#15819 := 4294967294::int
-#15820 := (<= uf_298 4294967294::int)
-#18490 := (not #15820)
-#13539 := (+ uf_298 #13471)
-#13538 := (= #13539 -1::int)
-#13542 := (not #13538)
+#13127 := (+ uf_286 #13126)
+#13128 := (<= #13127 0::int)
+#3271 := (uf_66 #2979 uf_305 uf_7)
+#3272 := (uf_110 uf_287 #3271)
+#12446 := (= uf_304 #3272)
+#21466 := (not #12446)
+#21467 := (or #21466 #13128)
+#21468 := (not #21467)
+#22493 := (or #21468 #22490)
+#3220 := (uf_66 #2979 uf_298 uf_7)
+#3230 := (uf_110 uf_287 #3220)
+#28730 := (= #3230 #3272)
+#28979 := (= #3272 #3230)
+#29112 := (= #3271 #3220)
+decl uf_144 :: (-> T3 T3)
+decl uf_124 :: (-> T3 int T3)
+#2977 := (uf_124 uf_7 uf_286)
+#24108 := (uf_144 #2977)
+decl uf_125 :: (-> T5 T5 int)
+decl uf_116 :: (-> T5 int)
+decl uf_28 :: (-> int T5)
+decl uf_29 :: (-> T5 int)
+#3013 := (uf_29 #2979)
+#23935 := (uf_28 #3013)
+#26317 := (uf_116 #23935)
+#26333 := (uf_43 #24108 #26317)
+decl uf_13 :: (-> T5 T3)
+#26868 := (uf_13 #26333)
+#27481 := (uf_66 #26333 uf_298 #26868)
+#27482 := (uf_125 #27481 #26333)
+#2980 := (uf_116 #2979)
+#26469 := (uf_43 #24108 #2980)
+#27555 := (uf_66 #26469 #27482 #24108)
+#27664 := (= #27555 #3220)
+#27655 := (= #24108 uf_7)
+#24109 := (= uf_7 #24108)
+#326 := (:var 1 T3)
+#1358 := (uf_124 #326 #161)
+#1592 := (pattern #1358)
+#1605 := (uf_144 #1358)
+#8271 := (= #326 #1605)
+#8275 := (forall (vars (?x388 T3) (?x389 int)) (:pat #1592) #8271)
+#16576 := (~ #8275 #8275)
+#16574 := (~ #8271 #8271)
+#16575 := [refl]: #16574
+#16577 := [nnf-pos #16575]: #16576
+#1606 := (= #1605 #326)
+#1607 := (forall (vars (?x388 T3) (?x389 int)) (:pat #1592) #1606)
+#8276 := (iff #1607 #8275)
+#8273 := (iff #1606 #8271)
+#8274 := [rewrite]: #8273
+#8277 := [quant-intro #8274]: #8276
+#8270 := [asserted]: #1607
+#8280 := [mp #8270 #8277]: #8275
+#16578 := [mp~ #8280 #16577]: #8275
+#24112 := (not #8275)
+#24113 := (or #24112 #24109)
+#24114 := [quant-inst]: #24113
+#27654 := [unit-resolution #24114 #16578]: #24109
+#27656 := [symm #27654]: #27655
+#27657 := (= #27482 uf_298)
+#26369 := (uf_13 #23935)
+#28212 := (uf_66 #23935 uf_298 #26369)
+#28213 := (uf_125 #28212 #23935)
+#29709 := (= #28213 uf_298)
+#28214 := (= uf_298 #28213)
+#15 := (:var 1 T5)
+#1390 := (uf_13 #15)
+#1391 := (uf_66 #15 #161 #1390)
+#1392 := (pattern #1391)
+#1393 := (uf_125 #1391 #15)
+#7663 := (= #161 #1393)
+#7667 := (forall (vars (?x319 T5) (?x320 int)) (:pat #1392) #7663)
+#16288 := (~ #7667 #7667)
+#16286 := (~ #7663 #7663)
+#16287 := [refl]: #16286
+#16289 := [nnf-pos #16287]: #16288
+#1394 := (= #1393 #161)
+#1395 := (forall (vars (?x319 T5) (?x320 int)) (:pat #1392) #1394)
+#7668 := (iff #1395 #7667)
+#7665 := (iff #1394 #7663)
+#7666 := [rewrite]: #7665
+#7669 := [quant-intro #7666]: #7668
+#7662 := [asserted]: #1395
+#7672 := [mp #7662 #7669]: #7667
+#16290 := [mp~ #7672 #16289]: #7667
+#26376 := (not #7667)
+#28217 := (or #26376 #28214)
+#28218 := [quant-inst]: #28217
+#28873 := [unit-resolution #28218 #16290]: #28214
+#29710 := [symm #28873]: #29709
+#28080 := (= #27482 #28213)
+#29457 := (= #28213 #27482)
+#27795 := (= #23935 #26333)
+#27716 := (= #2979 #26333)
+#27718 := (= #26333 #2979)
+#27783 := (= #26317 uf_288)
+#27660 := (= #2980 uf_288)
+#24227 := (= uf_288 #2980)
+#2697 := (uf_43 #326 #161)
+#22405 := (pattern #2697)
+#2698 := (uf_116 #2697)
+#11208 := (= #161 #2698)
+#22406 := (forall (vars (?x718 T3) (?x719 int)) (:pat #22405) #11208)
+#11212 := (forall (vars (?x718 T3) (?x719 int)) #11208)
+#22409 := (iff #11212 #22406)
+#22407 := (iff #11208 #11208)
+#22408 := [refl]: #22407
+#22410 := [quant-intro #22408]: #22409
+#17483 := (~ #11212 #11212)
+#17481 := (~ #11208 #11208)
+#17482 := [refl]: #17481
+#17484 := [nnf-pos #17482]: #17483
+#2699 := (= #2698 #161)
+#2700 := (forall (vars (?x718 T3) (?x719 int)) #2699)
+#11213 := (iff #2700 #11212)
+#11210 := (iff #2699 #11208)
+#11211 := [rewrite]: #11210
+#11214 := [quant-intro #11211]: #11213
+#11207 := [asserted]: #2700
+#11217 := [mp #11207 #11214]: #11212
+#17485 := [mp~ #11217 #17484]: #11212
+#22411 := [mp #17485 #22410]: #22406
+#24181 := (not #22406)
+#24232 := (or #24181 #24227)
+#24233 := [quant-inst]: #24232
+#27659 := [unit-resolution #24233 #22411]: #24227
+#27661 := [symm #27659]: #27660
+#27710 := (= #26317 #2980)
+#27708 := (= #23935 #2979)
+#23936 := (= #2979 #23935)
+#23 := (:var 0 T5)
+#93 := (uf_29 #23)
+#22316 := (pattern #93)
+#94 := (uf_28 #93)
+#3540 := (= #23 #94)
+#22317 := (forall (vars (?x14 T5)) (:pat #22316) #3540)
+#3543 := (forall (vars (?x14 T5)) #3540)
+#22318 := (iff #3543 #22317)
+#22320 := (iff #22317 #22317)
+#22321 := [rewrite]: #22320
+#22319 := [rewrite]: #22318
+#22322 := [trans #22319 #22321]: #22318
+#15525 := (~ #3543 #3543)
+#15515 := (~ #3540 #3540)
+#15516 := [refl]: #15515
+#15581 := [nnf-pos #15516]: #15525
+#95 := (= #94 #23)
+#96 := (forall (vars (?x14 T5)) #95)
+#3544 := (iff #96 #3543)
+#3541 := (iff #95 #3540)
+#3542 := [rewrite]: #3541
+#3545 := [quant-intro #3542]: #3544
+#3539 := [asserted]: #96
+#3548 := [mp #3539 #3545]: #3543
+#15582 := [mp~ #3548 #15581]: #3543
+#22323 := [mp #15582 #22322]: #22317
+#23939 := (not #22317)
+#23940 := (or #23939 #23936)
+#23941 := [quant-inst]: #23940
+#27707 := [unit-resolution #23941 #22323]: #23936
+#27709 := [symm #27707]: #27708
+#27711 := [monotonicity #27709]: #27710
+#27784 := [trans #27711 #27661]: #27783
+#27789 := [monotonicity #27656 #27784]: #27718
+#27790 := [symm #27789]: #27716
+#27796 := [trans #27709 #27790]: #27795
+#29455 := (= #28212 #27481)
+#29453 := (= #3220 #27481)
+#29451 := (= #27481 #3220)
+#29423 := (= #26868 uf_7)
+#24223 := (uf_13 #2979)
+#27846 := (= #24223 uf_7)
+#24224 := (= uf_7 #24223)
+#2701 := (uf_13 #2697)
+#11216 := (= #326 #2701)
+#22412 := (forall (vars (?x720 T3) (?x721 int)) (:pat #22405) #11216)
+#11220 := (forall (vars (?x720 T3) (?x721 int)) #11216)
+#22415 := (iff #11220 #22412)
+#22413 := (iff #11216 #11216)
+#22414 := [refl]: #22413
+#22416 := [quant-intro #22414]: #22415
+#17488 := (~ #11220 #11220)
+#17486 := (~ #11216 #11216)
+#17487 := [refl]: #17486
+#17489 := [nnf-pos #17487]: #17488
+#2702 := (= #2701 #326)
+#2703 := (forall (vars (?x720 T3) (?x721 int)) #2702)
+#11221 := (iff #2703 #11220)
+#11218 := (iff #2702 #11216)
+#11219 := [rewrite]: #11218
+#11222 := [quant-intro #11219]: #11221
+#11215 := [asserted]: #2703
+#11225 := [mp #11215 #11222]: #11220
+#17490 := [mp~ #11225 #17489]: #11220
+#22417 := [mp #17490 #22416]: #22412
+#24175 := (not #22412)
+#24229 := (or #24175 #24224)
+#24230 := [quant-inst]: #24229
+#27843 := [unit-resolution #24230 #22417]: #24224
+#27847 := [symm #27843]: #27846
+#29421 := (= #26868 #24223)
+#29422 := [monotonicity #27789]: #29421
+#29424 := [trans #29422 #27847]: #29423
+#29452 := [monotonicity #27789 #29424]: #29451
+#29454 := [symm #29452]: #29453
+#29449 := (= #28212 #3220)
+#27848 := (= #26369 uf_7)
+#27844 := (= #26369 #24223)
+#27845 := [monotonicity #27709]: #27844
+#27849 := [trans #27845 #27847]: #27848
+#29450 := [monotonicity #27709 #27849]: #29449
+#29456 := [trans #29450 #29454]: #29455
+#29458 := [monotonicity #29456 #27796]: #29457
+#29441 := [symm #29458]: #28080
+#29711 := [trans #29441 #29710]: #27657
+#27662 := (= #26469 #2979)
+#27663 := [monotonicity #27656 #27661]: #27662
+#29712 := [monotonicity #27663 #29711 #27656]: #27664
+#29208 := (= #3271 #27555)
+#29074 := (= uf_305 #27482)
+#28951 := (= uf_305 #28213)
+#3242 := (= uf_305 uf_298)
+#12383 := (= uf_298 uf_305)
+#22496 := (not #22493)
+#22499 := (or #22496 #21531)
+#22502 := (not #22499)
+#13090 := (+ uf_286 #13062)
+#13089 := (>= #13090 0::int)
+#13093 := (not #13089)
+#22505 := (or #13093 #22502)
+#22508 := (not #22505)
+#22511 := (or #13093 #22508)
+#22514 := (not #22511)
+#15126 := 4294967294::int
+#15127 := (<= uf_298 4294967294::int)
+#17745 := (not #15127)
+#13063 := (+ uf_298 #13062)
+#13061 := (= #13063 -1::int)
+#21556 := (not #13061)
#892 := 2::int
-#13462 := (>= uf_306 2::int)
-#22332 := (not #13462)
-#13454 := (>= uf_298 -1::int)
-#18487 := (not #13454)
-#13445 := (>= uf_305 0::int)
-#22331 := (not #13445)
+#13059 := (>= uf_306 2::int)
+#21555 := (not #13059)
+#13051 := (>= uf_298 -1::int)
+#17742 := (not #13051)
+#13017 := (>= uf_305 0::int)
+#21554 := (not #13017)
decl up_291 :: (-> T4 T1 T1 int T3 bool)
decl uf_4 :: T3
#7 := uf_4
@@ -108,208 +311,5085 @@
#2946 := uf_273
decl uf_269 :: T1
#2942 := uf_269
-#3216 := (up_291 uf_287 uf_269 uf_273 uf_306 uf_4)
-#12493 := (not #3216)
-#23260 := (or #12493 #22331 #18487 #22332 #13542 #18490 #23257)
-#23263 := (not #23260)
-#23266 := (or #18487 #18490 #23263)
-#23269 := (not #23266)
+#3258 := (up_291 uf_287 uf_269 uf_273 uf_306 uf_4)
+#21553 := (not #3258)
+#22517 := (or #21553 #21554 #17742 #21555 #21556 #17745 #22514)
+#22520 := (not #22517)
+#22523 := (or #17742 #17745 #22520)
+#22526 := (not #22523)
#4 := 1::int
-#13436 := (>= uf_298 1::int)
-#13576 := (not #13436)
-#12392 := (= uf_298 uf_305)
-#12539 := (not #12392)
-#12548 := (not #12389)
-#12384 := (= #3197 uf_303)
-#12591 := (not #12384)
+#12913 := (>= uf_298 1::int)
+#21575 := (not #12913)
+#21574 := (not #12383)
+decl uf_303 :: int
+#3234 := uf_303
+#12380 := (= uf_303 uf_304)
+#21573 := (not #12380)
+#12375 := (= #3230 uf_303)
+#21572 := (not #12375)
decl uf_68 :: (-> T4 T5 T2)
-#3194 := (uf_68 uf_287 #3188)
+#3226 := (uf_68 uf_287 #3220)
decl uf_9 :: T2
#19 := uf_9
-#12378 := (= uf_9 #3194)
-#18458 := (not #12378)
+#12355 := (= uf_9 #3226)
+#17730 := (not #12355)
decl uf_48 :: (-> T5 T3 T2)
-#3189 := (uf_48 #3188 uf_7)
-#12369 := (= uf_9 #3189)
-#18449 := (not #12369)
+#3221 := (uf_48 #3220 uf_7)
+#12346 := (= uf_9 #3221)
+#17721 := (not #12346)
decl uf_275 :: T1
#2948 := uf_275
decl uf_270 :: T1
#2943 := uf_270
-#3202 := (up_291 uf_287 uf_270 uf_275 uf_298 uf_4)
-#12573 := (not #3202)
+#3237 := (up_291 uf_287 uf_270 uf_275 uf_298 uf_4)
+#21571 := (not #3237)
decl uf_277 :: T1
#2950 := uf_277
decl uf_271 :: T1
#2944 := uf_271
-#3201 := (up_291 uf_287 uf_271 uf_277 uf_303 uf_7)
-#12582 := (not #3201)
-#23272 := (or #12582 #12573 #18449 #18458 #12591 #12548 #12539 #13576 #22331 #23269)
-#23275 := (not #23272)
-#13629 := (* -1::int #3197)
+#3236 := (up_291 uf_287 uf_271 uf_277 uf_303 uf_7)
+#21570 := (not #3236)
+#22529 := (or #21570 #21571 #17721 #17730 #21572 #21573 #21574 #21575 #21554 #22526)
+#22532 := (not #22529)
+#12998 := (* -1::int #3230)
decl uf_296 :: int
-#3061 := uf_296
-#13630 := (+ uf_296 #13629)
-#13628 := (>= #13630 0::int)
-#13627 := (not #13628)
+#3079 := uf_296
+#12999 := (+ uf_296 #12998)
+#12997 := (>= #12999 0::int)
+#12996 := (not #12997)
decl uf_297 :: int
-#3065 := uf_297
-#13433 := (>= uf_297 0::int)
-#22372 := (not #13433)
-#12671 := (= uf_297 uf_305)
-#12677 := (not #12671)
-#12668 := (= uf_296 uf_304)
-#12686 := (not #12668)
-#23302 := (or #12686 #12677 #22372 #13576 #22331 #13627 #23269)
-#23305 := (not #23302)
-#23278 := (or #18449 #18458 #23275)
-#23281 := (not #23278)
+#3083 := uf_297
+#12910 := (>= uf_297 0::int)
+#21602 := (not #12910)
+#12524 := (= uf_297 uf_305)
+#21614 := (not #12524)
+#12521 := (= uf_296 uf_304)
+#21613 := (not #12521)
+#22559 := (or #21613 #21614 #21602 #21575 #12996 #21554 #22526)
+#22562 := (not #22559)
+#22535 := (or #17721 #17730 #22532)
+#22538 := (not #22535)
decl uf_24 :: (-> T4 T5 T2)
-#3191 := (uf_24 uf_287 #3188)
-#12372 := (= uf_9 #3191)
-#18452 := (not #12372)
-#23284 := (or #18449 #18452 #23281)
-#23287 := (not #23284)
-#23290 := (or #18449 #18452 #23287)
-#23293 := (not #23290)
-#23296 := (or #22372 #13576 #13628 #23293)
-#23299 := (not #23296)
-#23308 := (or #23299 #23305)
-#23311 := (not #23308)
-#23314 := (or #18449 #18458 #22372 #13576 #23311)
-#23317 := (not #23314)
+#3223 := (uf_24 uf_287 #3220)
+#12349 := (= uf_9 #3223)
+#17724 := (not #12349)
+#22541 := (or #17721 #17724 #22538)
+#22544 := (not #22541)
+#22547 := (or #17721 #17724 #22544)
+#22550 := (not #22547)
+#22553 := (or #21602 #21575 #12997 #22550)
+#22556 := (not #22553)
+#22565 := (or #22556 #22562)
+#22568 := (not #22565)
+#22571 := (or #17721 #17730 #21602 #21575 #22568)
+#22574 := (not #22571)
decl uf_25 :: (-> T4 T5 T5)
decl uf_135 :: (-> T14 T5)
decl uf_58 :: (-> T13 T5 T14)
decl uf_59 :: (-> T4 T13)
-#3157 := (uf_59 uf_287)
-#27840 := (uf_58 #3157 #3188)
-#29300 := (uf_135 #27840)
-#29303 := (uf_25 uf_287 #29300)
+#3175 := (uf_59 uf_287)
+#27095 := (uf_58 #3175 #3220)
+#28578 := (uf_135 #27095)
+#28581 := (uf_25 uf_287 #28578)
decl uf_26 :: T5
#77 := uf_26
-#29304 := (= uf_26 #29303)
+#28582 := (= uf_26 #28581)
decl uf_210 :: (-> T4 T5 T2)
-#29301 := (uf_210 uf_287 #29300)
-#29302 := (= uf_9 #29301)
-#29360 := (or #29302 #29304)
-#29363 := (not #29360)
+#28579 := (uf_210 uf_287 #28578)
+#28580 := (= uf_9 #28579)
+#28638 := (or #28580 #28582)
+#28641 := (not #28638)
decl uf_136 :: (-> T14 T2)
-#29313 := (uf_136 #27840)
-#29314 := (= uf_9 #29313)
-#29315 := (not #29314)
+#28591 := (uf_136 #27095)
+#28592 := (= uf_9 #28591)
+#28593 := (not #28592)
decl uf_27 :: (-> T4 T5 T2)
-#29310 := (uf_27 uf_287 #29300)
-#29311 := (= uf_9 #29310)
-#29312 := (not #29311)
-#29354 := (or #29312 #29315)
-#29357 := (not #29354)
+#28588 := (uf_27 uf_287 #28578)
+#28589 := (= uf_9 #28588)
+#28590 := (not #28589)
+#28632 := (or #28590 #28593)
+#28635 := (not #28632)
decl uf_12 :: (-> T3 T8)
-decl uf_13 :: (-> T5 T3)
-#28096 := (uf_13 #3188)
-#29318 := (uf_12 #28096)
+#27351 := (uf_13 #3220)
+#28596 := (uf_12 #27351)
decl uf_14 :: T8
#28 := uf_14
-#29336 := (= uf_14 #29318)
-#29351 := (not #29336)
-#29307 := (uf_13 #29300)
-#29308 := (uf_12 #29307)
-#29309 := (= uf_14 #29308)
-#29369 := (or #29309 #29351 #29357 #29363)
-#29374 := (not #29369)
-#29325 := (uf_25 uf_287 #3188)
-#29326 := (= uf_26 #29325)
-#29323 := (uf_210 uf_287 #3188)
-#29324 := (= uf_9 #29323)
-#29339 := (or #29324 #29326)
-#29342 := (not #29339)
-#29345 := (or #29336 #29342)
-#29348 := (not #29345)
-#29377 := (or #29348 #29374)
-#29380 := (not #29377)
-#29383 := (or #18452 #29380)
-#29386 := (not #29383)
-#29389 := (iff #12378 #29386)
-#29961 := (not #29389)
-#30156 := [hypothesis]: #29961
-#23 := (:var 0 T5)
+#28614 := (= uf_14 #28596)
+#28629 := (not #28614)
+#28585 := (uf_13 #28578)
+#28586 := (uf_12 #28585)
+#28587 := (= uf_14 #28586)
+#28647 := (or #28587 #28629 #28635 #28641)
+#28652 := (not #28647)
+#28603 := (uf_25 uf_287 #3220)
+#28604 := (= uf_26 #28603)
+#28601 := (uf_210 uf_287 #3220)
+#28602 := (= uf_9 #28601)
+#28617 := (or #28602 #28604)
+#28620 := (not #28617)
+#28623 := (or #28614 #28620)
+#28626 := (not #28623)
+#28655 := (or #28626 #28652)
+#28658 := (not #28655)
+#28661 := (or #17724 #28658)
+#28664 := (not #28661)
+decl up_67 :: (-> T14 bool)
+decl uf_143 :: (-> T3 int)
+#24110 := (uf_143 #2977)
+#26391 := (uf_124 #24108 #24110)
+#26392 := (uf_43 #26391 #2980)
+#27485 := (uf_66 #26392 #27482 #24108)
+#27491 := (uf_58 #3175 #27485)
+#27497 := (up_67 #27491)
+#27498 := (not #27497)
+#27494 := (uf_135 #27491)
+#27495 := (= #26392 #27494)
+#27496 := (not #27495)
+#27492 := (uf_136 #27491)
+#27493 := (= uf_9 #27492)
+#27488 := (uf_24 uf_287 #27485)
+#27489 := (= uf_9 #27488)
+#27490 := (not #27489)
+#27519 := (or #27490 #27493 #27496 #27498)
+#27522 := (not #27519)
+#30005 := [hypothesis]: #27519
+#27503 := (>= #27482 0::int)
+#28457 := (* -1::int #28213)
+#28081 := (+ #27482 #28457)
+#28082 := (>= #28081 0::int)
+#29442 := (not #28080)
+#29440 := (or #29442 #28082)
+#29443 := [th-lemma]: #29440
+#29444 := [unit-resolution #29443 #29441]: #28082
+#13392 := (>= uf_298 0::int)
+decl ?x776!15 :: int
+#17865 := ?x776!15
+#18168 := (* -1::int ?x776!15)
+#18169 := (+ uf_286 #18168)
+#18170 := (<= #18169 0::int)
+#17873 := (uf_66 #2979 ?x776!15 uf_7)
+#17874 := (uf_110 uf_287 #17873)
+#18155 := (* -1::int #17874)
+decl uf_302 :: int
+#3128 := uf_302
+#18156 := (+ uf_302 #18155)
+#18157 := (>= #18156 0::int)
+#17870 := (>= ?x776!15 0::int)
+#21709 := (not #17870)
+#17866 := (<= ?x776!15 4294967295::int)
+#21708 := (not #17866)
+#21724 := (or #21708 #21709 #18157 #18170)
+#21729 := (not #21724)
+#12727 := (* -1::int uf_286)
+#13271 := (+ #161 #12727)
+#13270 := (>= #13271 0::int)
+#3142 := (= #3073 uf_302)
+#21682 := (not #3142)
+#21683 := (or #21682 #4963 #13270 #18695)
+#22609 := (forall (vars (?x778 int)) (:pat #22468) #21683)
+#22614 := (not #22609)
+#13286 := (* -1::int uf_302)
+#13287 := (+ #3073 #13286)
+#13288 := (<= #13287 0::int)
+#21674 := (or #4963 #13270 #13288 #18695)
+#22601 := (forall (vars (?x776 int)) (:pat #22468) #21674)
+#22606 := (not #22601)
+#22617 := (or #22606 #22614)
+#22620 := (not #22617)
+#22623 := (or #22620 #21729)
+#22626 := (not #22623)
+#12965 := (* -1::int uf_298)
+#12966 := (+ uf_286 #12965)
+#12967 := (<= #12966 0::int)
+#12968 := (not #12967)
+#12122 := (= uf_296 uf_302)
+#21741 := (not #12122)
+decl uf_301 :: int
+#3126 := uf_301
+#12119 := (= uf_297 uf_301)
+#21740 := (not #12119)
+decl uf_300 :: int
+#3124 := uf_300
+#12116 := (= uf_298 uf_300)
+#21739 := (not #12116)
+decl uf_299 :: int
+#3122 := uf_299
+#12113 := (= uf_296 uf_299)
+#21738 := (not #12113)
+#22629 := (or #21738 #21739 #21740 #21741 #21602 #21575 #12968 #22626)
+#22632 := (not #22629)
+#22577 := (or #17721 #17730 #22574)
+#22580 := (not #22577)
+#22583 := (or #17721 #17724 #22580)
+#22586 := (not #22583)
+#22589 := (or #17721 #17724 #22586)
+#22592 := (not #22589)
+#22595 := (or #21602 #21575 #12967 #22592)
+#22598 := (not #22595)
+#22635 := (or #22598 #22632)
+#22638 := (not #22635)
+#13378 := (* -1::int uf_296)
+#13379 := (+ #3073 #13378)
+#13380 := (<= #13379 0::int)
+#13363 := (+ #161 #12965)
+#13362 := (>= #13363 0::int)
+#21458 := (or #4963 #13362 #13380 #18695)
+#22477 := (forall (vars (?x775 int)) (:pat #22468) #21458)
+#22482 := (not #22477)
+#1331 := 255::int
+#15107 := (<= uf_296 255::int)
+#21770 := (not #15107)
+#15096 := (<= uf_297 4294967295::int)
+#21769 := (not #15096)
+#15085 := (<= uf_298 4294967295::int)
+#21768 := (not #15085)
+#13402 := (>= uf_296 0::int)
+#21766 := (not #13402)
+#21765 := (not #13392)
+#13389 := (>= #12966 0::int)
+#21764 := (not #13389)
+#13356 := (* -1::int uf_297)
+#13357 := (+ uf_286 #13356)
+#13358 := (<= #13357 0::int)
+#12740 := (<= uf_286 0::int)
+decl uf_178 :: (-> T4 T4 T2)
+#3187 := (uf_178 uf_287 uf_287)
+#12261 := (= uf_9 #3187)
+#21763 := (not #12261)
+decl uf_202 :: (-> T1 T4 T2)
+decl uf_272 :: T1
+#2945 := uf_272
+#3113 := (uf_202 uf_272 uf_287)
+#12067 := (= uf_9 #3113)
+#21762 := (not #12067)
+#3100 := (uf_66 #2979 uf_297 uf_7)
+#3101 := (uf_110 uf_287 #3100)
+#12030 := (= uf_296 #3101)
+#21761 := (not #12030)
+decl up_292 :: (-> T4 T1 T1 T5 T3 bool)
+decl uf_6 :: (-> T3 T3)
+#11 := (uf_6 uf_7)
+decl uf_280 :: T1
+#2953 := uf_280
+#3200 := (up_292 uf_287 uf_272 uf_280 #2979 #11)
+#21760 := (not #3200)
+#3199 := (up_291 uf_287 uf_272 uf_280 #3013 #11)
+#21759 := (not #3199)
+decl uf_279 :: T1
+#2952 := uf_279
+#3198 := (up_291 uf_287 uf_272 uf_279 uf_286 uf_4)
+#21758 := (not #3198)
+#3197 := (up_291 uf_287 uf_272 uf_277 uf_296 uf_7)
+#21757 := (not #3197)
+#3196 := (up_291 uf_287 uf_272 uf_275 uf_297 uf_4)
+#21756 := (not #3196)
+#3195 := (up_291 uf_287 uf_272 uf_273 uf_298 uf_4)
+#21755 := (not #3195)
+#3044 := (uf_66 #2979 0::int uf_7)
+#3054 := (uf_110 uf_287 #3044)
+decl uf_295 :: int
+#3053 := uf_295
+#3055 := (= uf_295 #3054)
+#17676 := (not #3055)
+#22641 := (or #17676 #21755 #21756 #21757 #21758 #21759 #21760 #21761 #21762 #21763 #12740 #21602 #21575 #13358 #21764 #21765 #21766 #21768 #21769 #21770 #22482 #22638)
+#22644 := (not #22641)
+#22647 := (or #17676 #12740 #22644)
+#22650 := (not #22647)
+#12886 := (* -1::int #3073)
+#12887 := (+ uf_295 #12886)
+#12885 := (>= #12887 0::int)
+#12869 := (>= #161 1::int)
+#21436 := (or #4963 #12869 #12885 #18695)
+#22469 := (forall (vars (?x773 int)) (:pat #22468) #21436)
+#22474 := (not #22469)
+#22653 := (or #22474 #22650)
+#22656 := (not #22653)
+decl ?x773!13 :: int
+#17651 := ?x773!13
+#17658 := (uf_66 #2979 ?x773!13 uf_7)
+#17659 := (uf_110 uf_287 #17658)
+#17660 := (* -1::int #17659)
+#17661 := (+ uf_295 #17660)
+#17662 := (>= #17661 0::int)
+#17655 := (>= ?x773!13 0::int)
+#21399 := (not #17655)
+#17653 := (>= ?x773!13 1::int)
+#17652 := (<= ?x773!13 4294967295::int)
+#21398 := (not #17652)
+#21414 := (or #21398 #17653 #21399 #17662)
+#21419 := (not #21414)
+#22659 := (or #21419 #22656)
+#22662 := (not #22659)
+#12863 := (>= uf_286 1::int)
+#12866 := (not #12863)
+#22665 := (or #12866 #22662)
+#22668 := (not #22665)
+#22671 := (or #12866 #22668)
+#22674 := (not #22671)
+#3050 := (uf_68 uf_287 #3044)
+#11979 := (= uf_9 #3050)
+#17640 := (not #11979)
+#3045 := (uf_48 #3044 uf_7)
+#11970 := (= uf_9 #3045)
+#17631 := (not #11970)
+decl uf_274 :: T1
+#2947 := uf_274
+#3058 := (up_291 uf_287 uf_274 uf_273 1::int uf_4)
+#21813 := (not #3058)
+decl uf_276 :: T1
+#2949 := uf_276
+#3057 := (up_291 uf_287 uf_276 uf_275 0::int uf_4)
+#21812 := (not #3057)
+decl uf_278 :: T1
+#2951 := uf_278
+#3056 := (up_291 uf_287 uf_278 uf_277 uf_295 uf_7)
+#21811 := (not #3056)
+#22677 := (or #17676 #21811 #21812 #21813 #17631 #17640 #22674)
+#22680 := (not #22677)
+#22683 := (or #17631 #17640 #22680)
+#22686 := (not #22683)
+#3047 := (uf_24 uf_287 #3044)
+#11973 := (= uf_9 #3047)
+#17634 := (not #11973)
+#22689 := (or #17631 #17634 #22686)
+#22692 := (not #22689)
+#26286 := (uf_13 #3044)
+#26287 := (= #24108 #26286)
+#26260 := (uf_48 #3044 #24108)
+#26261 := (= uf_9 #26260)
+#26289 := (iff #26261 #26287)
+#233 := (:var 0 T3)
+#2666 := (uf_48 #15 #233)
+#2667 := (pattern #2666)
+#11138 := (= uf_9 #2666)
+#9039 := (= #233 #1390)
+#11159 := (iff #9039 #11138)
+#22399 := (forall (vars (?x712 T5) (?x713 T3)) (:pat #2667) #11159)
+#11164 := (forall (vars (?x712 T5) (?x713 T3)) #11159)
+#22402 := (iff #11164 #22399)
+#22400 := (iff #11159 #11159)
+#22401 := [refl]: #22400
+#22403 := [quant-intro #22401]: #22402
+#17468 := (~ #11164 #11164)
+#17466 := (~ #11159 #11159)
+#17467 := [refl]: #17466
+#17469 := [nnf-pos #17467]: #17468
+#1890 := (= #1390 #233)
+#2668 := (= #2666 uf_9)
+#2673 := (iff #2668 #1890)
+#2674 := (forall (vars (?x712 T5) (?x713 T3)) #2673)
+#11165 := (iff #2674 #11164)
+#11162 := (iff #2673 #11159)
+#11155 := (iff #11138 #9039)
+#11160 := (iff #11155 #11159)
+#11161 := [rewrite]: #11160
+#11157 := (iff #2673 #11155)
+#9040 := (iff #1890 #9039)
+#9041 := [rewrite]: #9040
+#11140 := (iff #2668 #11138)
+#11141 := [rewrite]: #11140
+#11158 := [monotonicity #11141 #9041]: #11157
+#11163 := [trans #11158 #11161]: #11162
+#11166 := [quant-intro #11163]: #11165
+#11154 := [asserted]: #2674
+#11169 := [mp #11154 #11166]: #11164
+#17470 := [mp~ #11169 #17469]: #11164
+#22404 := [mp #17470 #22403]: #22399
+#25427 := (not #22399)
+#26306 := (or #25427 #26289)
+#26288 := (iff #26287 #26261)
+#26307 := (or #25427 #26288)
+#26363 := (iff #26307 #26306)
+#26462 := (iff #26306 #26306)
+#26453 := [rewrite]: #26462
+#26290 := (iff #26288 #26289)
+#26291 := [rewrite]: #26290
+#26459 := [monotonicity #26291]: #26363
+#26518 := [trans #26459 #26453]: #26363
+#26300 := [quant-inst]: #26307
+#26463 := [mp #26300 #26518]: #26306
+#26613 := [unit-resolution #26463 #22404]: #26289
+#26268 := (not #26261)
+#26569 := (iff #17631 #26268)
+#26567 := (iff #11970 #26261)
+#26657 := (iff #26261 #11970)
+#26615 := (= #26260 #3045)
+#26616 := [monotonicity #27656]: #26615
+#26658 := [monotonicity #26616]: #26657
+#26568 := [symm #26658]: #26567
+#26617 := [monotonicity #26568]: #26569
+#26614 := [hypothesis]: #17631
+#26618 := [mp #26614 #26617]: #26268
+#26660 := (= #24223 #26286)
+#26622 := (= #26286 #24223)
+#26620 := (= #3044 #2979)
+#26566 := (= #3044 #26333)
+#26311 := (uf_66 #23935 0::int #24108)
+#26336 := (= #26311 #26333)
+#26339 := (not #26336)
+decl uf_138 :: (-> T5 T5 T2)
+#26312 := (uf_138 #26311 #23935)
+#26313 := (= uf_9 #26312)
+#26314 := (not #26313)
+#26345 := (or #26314 #26339)
+#26350 := (not #26345)
+#247 := (:var 1 int)
+#21 := (:var 2 T5)
+#1576 := (uf_66 #21 #247 #233)
+#1577 := (pattern #1576)
+#1578 := (uf_138 #1576 #21)
+#8220 := (= uf_9 #1578)
+#20314 := (not #8220)
+decl uf_139 :: (-> T3 int)
+#1581 := (uf_139 #233)
+#1582 := (* #247 #1581)
+#1580 := (uf_116 #21)
+#1583 := (+ #1580 #1582)
+#1584 := (uf_43 #233 #1583)
+#1585 := (= #1576 #1584)
+#20313 := (not #1585)
+#20315 := (or #20313 #20314)
+#20316 := (not #20315)
+#20319 := (forall (vars (?x375 T5) (?x376 int) (?x377 T3)) (:pat #1577) #20316)
+#8226 := (and #1585 #8220)
+#8231 := (forall (vars (?x375 T5) (?x376 int) (?x377 T3)) (:pat #1577) #8226)
+#20320 := (iff #8231 #20319)
+#20317 := (iff #8226 #20316)
+#20318 := [rewrite]: #20317
+#20321 := [quant-intro #20318]: #20320
+#16546 := (~ #8231 #8231)
+#16544 := (~ #8226 #8226)
+#16545 := [refl]: #16544
+#16547 := [nnf-pos #16545]: #16546
+#1579 := (= #1578 uf_9)
+#1586 := (and #1579 #1585)
+#1587 := (forall (vars (?x375 T5) (?x376 int) (?x377 T3)) (:pat #1577) #1586)
+#8232 := (iff #1587 #8231)
+#8229 := (iff #1586 #8226)
+#8223 := (and #8220 #1585)
+#8227 := (iff #8223 #8226)
+#8228 := [rewrite]: #8227
+#8224 := (iff #1586 #8223)
+#8221 := (iff #1579 #8220)
+#8222 := [rewrite]: #8221
+#8225 := [monotonicity #8222]: #8224
+#8230 := [trans #8225 #8228]: #8229
+#8233 := [quant-intro #8230]: #8232
+#8219 := [asserted]: #1587
+#8236 := [mp #8219 #8233]: #8231
+#16548 := [mp~ #8236 #16547]: #8231
+#20322 := [mp #16548 #20321]: #20319
+#26353 := (not #20319)
+#26293 := (or #26353 #26350)
+#26315 := (uf_139 #24108)
+#26316 := (* 0::int #26315)
+#26318 := (+ #26317 #26316)
+#26319 := (uf_43 #24108 #26318)
+#26320 := (= #26311 #26319)
+#26321 := (not #26320)
+#26322 := (or #26321 #26314)
+#26323 := (not #26322)
+#26294 := (or #26353 #26323)
+#26295 := (iff #26294 #26293)
+#26310 := (iff #26293 #26293)
+#26297 := [rewrite]: #26310
+#26351 := (iff #26323 #26350)
+#26348 := (iff #26322 #26345)
+#26342 := (or #26339 #26314)
+#26346 := (iff #26342 #26345)
+#26347 := [rewrite]: #26346
+#26343 := (iff #26322 #26342)
+#26340 := (iff #26321 #26339)
+#26337 := (iff #26320 #26336)
+#26334 := (= #26319 #26333)
+#26331 := (= #26318 #26317)
+#26326 := (+ #26317 0::int)
+#26329 := (= #26326 #26317)
+#26330 := [rewrite]: #26329
+#26327 := (= #26318 #26326)
+#26324 := (= #26316 0::int)
+#26325 := [rewrite]: #26324
+#26328 := [monotonicity #26325]: #26327
+#26332 := [trans #26328 #26330]: #26331
+#26335 := [monotonicity #26332]: #26334
+#26338 := [monotonicity #26335]: #26337
+#26341 := [monotonicity #26338]: #26340
+#26344 := [monotonicity #26341]: #26343
+#26349 := [trans #26344 #26347]: #26348
+#26352 := [monotonicity #26349]: #26351
+#26296 := [monotonicity #26352]: #26295
+#26298 := [trans #26296 #26297]: #26295
+#26292 := [quant-inst]: #26294
+#26299 := [mp #26292 #26298]: #26293
+#27785 := [unit-resolution #26299 #20322]: #26350
+#26385 := (or #26345 #26336)
+#26458 := [def-axiom]: #26385
+#27786 := [unit-resolution #26458 #27785]: #26336
+#27787 := (= #3044 #26311)
+#27788 := [monotonicity #27707 #27654]: #27787
+#26619 := [trans #27788 #27786]: #26566
+#26621 := [trans #26619 #27789]: #26620
+#26609 := [monotonicity #26621]: #26622
+#26661 := [symm #26609]: #26660
+#26662 := (= #24108 #24223)
+#26645 := [trans #27656 #27843]: #26662
+#26647 := [trans #26645 #26661]: #26287
+#26503 := (not #26287)
+#26502 := (not #26289)
+#26504 := (or #26502 #26261 #26503)
+#26505 := [def-axiom]: #26504
+#26659 := [unit-resolution #26505 #26647 #26618 #26613]: false
+#26648 := [lemma #26659]: #11970
+#28096 := (or #17631 #22692)
+#22695 := (or #17631 #17634 #22692)
+#22698 := (not #22695)
+decl uf_200 :: (-> T4 T5 T5 T16 T2)
+decl uf_282 :: T16
+#2957 := uf_282
+#2981 := (uf_43 #2977 #2980)
+#3042 := (uf_200 uf_287 #2981 #2981 uf_282)
+#11967 := (= uf_9 #3042)
+#12702 := (not #11967)
+#22701 := (or #12702 #22698)
+#22704 := (not #22701)
+#24010 := (uf_116 #2981)
+#25399 := (uf_43 #2977 #24010)
+#25810 := (uf_13 #25399)
+#26056 := (uf_12 #25810)
+#26059 := (= uf_14 #26056)
+#26178 := (not #26059)
+#2990 := (uf_12 #2977)
+#11873 := (= uf_14 #2990)
+#11876 := (not #11873)
+#26179 := (iff #11876 #26178)
+#26176 := (iff #11873 #26059)
+#26195 := (iff #26059 #11873)
+#26171 := (= #26056 #2990)
+#26169 := (= #25810 #2977)
+#24228 := (uf_13 #2981)
+#27920 := (= #24228 #2977)
+#24231 := (= #2977 #24228)
+#24236 := (or #24175 #24231)
+#24237 := [quant-inst]: #24236
+#27857 := [unit-resolution #24237 #22417]: #24231
+#27921 := [symm #27857]: #27920
+#26162 := (= #25810 #24228)
+#27876 := (= #25399 #2981)
+#25406 := (= #2981 #25399)
+#2986 := (uf_48 #2981 #2977)
+#11867 := (= uf_9 #2986)
+decl uf_283 :: int
+#2961 := uf_283
+#12768 := (* -1::int uf_283)
+decl uf_78 :: int
+#429 := uf_78
+#12769 := (+ uf_78 #12768)
+#12767 := (>= #12769 0::int)
+#12765 := (>= uf_283 0::int)
+decl uf_284 :: int
+#2965 := uf_284
+#12760 := (* -1::int uf_284)
+decl uf_76 :: int
+#409 := uf_76
+#12761 := (+ uf_76 #12760)
+#12759 := (>= #12761 0::int)
+#12757 := (>= uf_284 0::int)
+decl uf_285 :: int
+#2969 := uf_285
+#12752 := (* -1::int uf_285)
+#12753 := (+ uf_76 #12752)
+#12751 := (>= #12753 0::int)
+#12749 := (>= uf_285 0::int)
+#1042 := 1099511627776::int
+#12744 := (>= uf_286 1099511627776::int)
+#12745 := (not #12744)
+#12741 := (not #12740)
+decl uf_289 :: (-> T19 int)
+#3007 := (:var 0 T19)
+#3008 := (uf_289 #3007)
+#3009 := (pattern #3008)
+decl uf_290 :: int
+#3010 := uf_290
+#12733 := (* -1::int uf_290)
+#12734 := (+ #3008 #12733)
+#12732 := (>= #12734 0::int)
+#12731 := (not #12732)
+#12737 := (forall (vars (?x771 T19)) (:pat #3009) #12731)
+#12728 := (+ uf_76 #12727)
+#12726 := (>= #12728 0::int)
+#12724 := (>= uf_286 0::int)
+decl uf_294 :: (-> int T5 T2)
+decl uf_293 :: int
+#3018 := uf_293
+#3021 := (uf_294 uf_293 #23)
+#3022 := (pattern #3021)
+#11909 := (= uf_9 #3021)
+#11915 := (not #11909)
+#11920 := (forall (vars (?x772 T5)) (:pat #3022) #11915)
+decl uf_55 :: (-> T4 T2)
+#3004 := (uf_55 uf_287)
+#11903 := (= uf_9 #3004)
+decl uf_281 :: T1
+#2954 := uf_281
+#3002 := (uf_202 uf_281 uf_287)
+#11900 := (= uf_9 #3002)
+decl uf_203 :: (-> T4 T2)
+#3000 := (uf_203 uf_287)
+#11897 := (= uf_9 #3000)
+decl uf_23 :: (-> T3 T2)
+#2993 := (uf_23 #2977)
+#11879 := (= uf_9 #2993)
+#2988 := (uf_24 uf_287 #2981)
+#11870 := (= uf_9 #2988)
+#2984 := (uf_25 uf_287 #2981)
+#11864 := (= uf_26 #2984)
+#2982 := (uf_27 uf_287 #2981)
+#11860 := (= uf_9 #2982)
+decl uf_173 :: (-> T4 int)
+#3019 := (uf_173 uf_287)
+#3020 := (= uf_293 #3019)
+#3017 := (up_291 uf_287 uf_281 uf_279 uf_286 uf_4)
+#3015 := (up_292 uf_287 uf_281 uf_280 #2979 #11)
+#3014 := (up_291 uf_287 uf_281 uf_280 #3013 #11)
+#12823 := (and #3014 #3015 #3017 #3020 #11860 #11864 #11867 #11870 #11876 #11879 #11897 #11900 #11903 #11920 #12724 #12726 #12737 #12741 #12745 #12749 #12751 #12757 #12759 #12765 #12767)
+#13406 := (+ uf_78 #13378)
+#13405 := (>= #13406 0::int)
+#13399 := (+ uf_76 #13356)
+#13398 := (>= #13399 0::int)
+#13055 := (+ uf_76 #12965)
+#13395 := (>= #13055 0::int)
+#13364 := (not #13362)
+#4353 := (* -1::int uf_76)
+#4354 := (+ #161 #4353)
+#4355 := (<= #4354 0::int)
+#13370 := (and #4041 #4355 #13364)
+#13375 := (not #13370)
+#13383 := (or #13375 #13380)
+#13386 := (forall (vars (?x775 int)) #13383)
+#13359 := (not #13358)
+#13448 := (and #3055 #12030 #12741 #12910 #12913 #13359 #13386 #13389 #13392 #13395 #13398 #13402 #13405)
+#13453 := (not #13448)
+#13342 := (and #12910 #12913 #12967)
+#13347 := (not #13342)
+#13272 := (not #13270)
+#13306 := (and #3142 #4041 #4355 #13272)
+#13311 := (exists (vars (?x778 int)) #13306)
+#13278 := (and #4041 #4355 #13272)
+#13283 := (not #13278)
+#13291 := (or #13283 #13288)
+#13294 := (forall (vars (?x776 int)) #13291)
+#13314 := (not #13294)
+#13320 := (or #13314 #13311)
+#13325 := (and #13294 #13320)
+decl up_216 :: bool
+#2482 := up_216
+#13262 := (and up_216 #12113 #12116 #12119 #12122 #12910 #12913)
+#13267 := (not #13262)
+#13328 := (or #13267 #13325)
+#13331 := (and up_216 #13328)
+#13350 := (or #13331 #13347)
+#13211 := (and #12521 #12524 #12910 #12913 #12997 #13017)
+#13216 := (not #13211)
+#13129 := (not #13128)
+#13135 := (and #12446 #13129)
+#13098 := (not #13096)
+#13104 := (and #4041 #4355 #13098)
+#13109 := (not #13104)
+#13117 := (or #13109 #13114)
+#13120 := (forall (vars (?x785 int)) #13117)
+#13123 := (not #13120)
+#13140 := (or #13123 #13135)
+#13143 := (and #13120 #13140)
+#13146 := (or #13093 #13143)
+#13149 := (and #13089 #13146)
+#13054 := (>= #13055 1::int)
+#13081 := (and #3258 #13017 #13051 #13054 #13059 #13061)
+#13086 := (not #13081)
+#13152 := (or #13086 #13149)
+#13158 := (and #13051 #13054 #13152)
+#13219 := (or #13158 #13216)
+#13043 := (and #3236 #3237 #12346 #12355 #12375 #12380 #12383 #12913 #13017)
+#13048 := (not #13043)
+#13163 := (or #13048 #13158)
+#13169 := (and #12346 #12355 #13163)
+#12352 := (and #12346 #12349)
+#12503 := (not #12352)
+#13174 := (or #12503 #13169)
+#13180 := (and #12346 #12349 #13174)
+#13008 := (and #12910 #12913 #12996)
+#13013 := (not #13008)
+#13185 := (or #13013 #13180)
+#13222 := (and #13185 #13219)
+#12988 := (and #12346 #12355 #12910 #12913)
+#12993 := (not #12988)
+#13225 := (or #12993 #13222)
+#13231 := (and #12346 #12355 #13225)
+#13236 := (or #12503 #13231)
+#13242 := (and #12346 #12349 #13236)
+#12977 := (and #12910 #12913 #12968)
+#12982 := (not #12977)
+#13247 := (or #12982 #13242)
+#13353 := (and #13247 #13350)
+#12957 := (and #3195 #3196 #3197 #3198 #3199 #3200 #11903 #12067 #12261 #12910 #12913)
+#12962 := (not #12957)
+#13462 := (or #12962 #13353 #13453)
+#13470 := (and #3055 #12741 #13462)
+#12870 := (not #12869)
+#12876 := (and #4041 #4355 #12870)
+#12881 := (not #12876)
+#12889 := (or #12881 #12885)
+#12892 := (forall (vars (?x773 int)) #12889)
+#12895 := (not #12892)
+#13475 := (or #12895 #13470)
+#13478 := (and #12892 #13475)
+#13481 := (or #12866 #13478)
+#13484 := (and #12863 #13481)
+#12855 := (and #3055 #3056 #3057 #3058 #11970 #11979)
+#12860 := (not #12855)
+#13487 := (or #12860 #13484)
+#13493 := (and #11970 #11979 #13487)
+#11976 := (and #11970 #11973)
+#12690 := (not #11976)
+#13498 := (or #12690 #13493)
+#13504 := (and #11970 #11973 #13498)
+#13509 := (or #12702 #13504)
+#13512 := (and #11967 #13509)
+#12828 := (not #12823)
+#13515 := (or #12828 #13512)
+#13518 := (not #13515)
+#3137 := (< #161 uf_286)
+#3143 := (and #3137 #3142)
+#411 := (<= #161 uf_76)
+#3144 := (and #411 #3143)
+#285 := (<= 0::int #161)
+#3145 := (and #285 #3144)
+#3146 := (exists (vars (?x778 int)) #3145)
+#3139 := (<= #3073 uf_302)
+#412 := (and #285 #411)
+#3138 := (and #412 #3137)
+#3140 := (implies #3138 #3139)
+#3141 := (forall (vars (?x776 int)) #3140)
+#3147 := (implies #3141 #3146)
+#3148 := (and #3141 #3147)
+#3129 := (= uf_302 uf_296)
+#3127 := (= uf_301 uf_297)
+#3130 := (and #3127 #3129)
+#3125 := (= uf_300 uf_298)
+#3131 := (and #3125 #3130)
+#3123 := (= uf_299 uf_296)
+#3132 := (and #3123 #3131)
+#3084 := (<= 0::int uf_297)
+#3091 := (<= 1::int uf_298)
+#3092 := (and #3091 #3084)
+#3133 := (and #3092 #3132)
+#3134 := (and #3092 #3133)
+#3135 := (and #3092 #3134)
+#3136 := (and up_216 #3135)
+#3149 := (implies #3136 #3148)
+#3150 := (and up_216 #3149)
+#3116 := (and #3092 #3092)
+#3304 := (and #3092 #3116)
+#3305 := (and #3092 #3304)
+#3303 := (<= uf_286 uf_298)
+#3306 := (and #3303 #3305)
+#3307 := (and #3092 #3306)
+#3308 := (implies #3307 #3150)
+#3273 := (= #3272 uf_304)
+#3270 := (< uf_305 uf_286)
+#3274 := (and #3270 #3273)
+#3267 := (<= #3073 uf_304)
+#3265 := (< #161 uf_306)
+#3266 := (and #412 #3265)
+#3268 := (implies #3266 #3267)
+#3269 := (forall (vars (?x785 int)) #3268)
+#3275 := (implies #3269 #3274)
+#3276 := (and #3269 #3275)
+#3264 := (<= uf_306 uf_286)
+#3277 := (implies #3264 #3276)
+#3278 := (and #3264 #3277)
+#3243 := (<= 0::int uf_305)
+#3259 := (<= 2::int uf_306)
+#3260 := (and #3259 #3243)
+#3261 := (and #3258 #3260)
+#3252 := (+ uf_298 1::int)
+#3257 := (= uf_306 #3252)
+#3262 := (and #3257 #3261)
+#3254 := (<= #3252 uf_76)
+#3253 := (<= 0::int #3252)
+#3255 := (and #3253 #3254)
+#3263 := (and #3255 #3262)
+#3279 := (implies #3263 #3278)
+#3280 := (and #3255 #3279)
+#3244 := (and #3091 #3243)
+#3288 := (= uf_305 uf_297)
+#3289 := (and #3288 #3244)
+#3287 := (= uf_304 uf_296)
+#3290 := (and #3287 #3289)
+#3291 := (and #3092 #3290)
+#3292 := (and #3092 #3291)
+#3293 := (and #3092 #3292)
+#3286 := (<= #3230 uf_296)
+#3294 := (and #3286 #3293)
+#3295 := (and #3092 #3294)
+#3296 := (implies #3295 #3280)
+#3245 := (and #3242 #3244)
+#3240 := (= uf_304 uf_303)
+#3246 := (and #3240 #3245)
+#3238 := (and #3091 #3091)
+#3247 := (and #3238 #3246)
+#3248 := (and #3237 #3247)
+#3249 := (and #3236 #3248)
+#3235 := (= uf_303 #3230)
+#3250 := (and #3235 #3249)
+#3227 := (= #3226 uf_9)
+#3222 := (= #3221 uf_9)
+#3228 := (and #3222 #3227)
+#3251 := (and #3228 #3250)
+#3281 := (implies #3251 #3280)
+#3282 := (and #3228 #3281)
+#3224 := (= #3223 uf_9)
+#3225 := (and #3222 #3224)
+#3283 := (implies #3225 #3282)
+#3284 := (and #3225 #3283)
+#3231 := (< uf_296 #3230)
+#3232 := (and #3231 #3116)
+#3233 := (and #3092 #3232)
+#3285 := (implies #3233 #3284)
+#3297 := (and #3285 #3296)
+#3229 := (and #3228 #3092)
+#3298 := (implies #3229 #3297)
+#3299 := (and #3228 #3298)
+#3300 := (implies #3225 #3299)
+#3301 := (and #3225 #3300)
+#3217 := (< uf_298 uf_286)
+#3218 := (and #3217 #3116)
+#3219 := (and #3092 #3218)
+#3302 := (implies #3219 #3301)
+#3309 := (and #3302 #3308)
+decl uf_41 :: (-> T4 T12)
+#3152 := (uf_41 uf_287)
+#3203 := (= #3152 #3152)
+#3202 := (= #3175 #3175)
+#3204 := (and #3202 #3203)
+#3205 := (and #3204 #3092)
+#3201 := (and #3199 #3200)
+#3206 := (and #3201 #3205)
+#3207 := (and #3198 #3206)
+#3208 := (and #3197 #3207)
+#3209 := (and #3196 #3208)
+#3210 := (and #3195 #3209)
+#3005 := (= #3004 uf_9)
+#3114 := (= #3113 uf_9)
+#3115 := (and #3114 #3005)
+#3211 := (and #3115 #3210)
+#3188 := (= #3187 uf_9)
+decl uf_172 :: (-> T4 T5 int)
+#3183 := (uf_172 uf_287 #23)
+#3184 := (pattern #3183)
+#3185 := (<= #3183 #3183)
+#3186 := (forall (vars (?x783 T5)) (:pat #3184) #3185)
+#3189 := (and #3186 #3188)
+#3182 := (<= #3019 #3019)
+#3190 := (and #3182 #3189)
+#3212 := (and #3190 #3211)
+#3176 := (uf_58 #3175 #23)
+#3177 := (pattern #3176)
+#3165 := (uf_68 uf_287 #23)
+#3166 := (= #3165 uf_9)
+#3178 := (= #3176 #3176)
+#3179 := (and #3178 #3166)
+#3180 := (implies #3166 #3179)
+#3181 := (forall (vars (?x782 T5)) (:pat #3177) #3180)
+#3191 := (and #3181 #3190)
+decl uf_40 :: (-> T12 T5 T11)
+#3153 := (uf_40 #3152 #23)
+#3154 := (pattern #3153)
+#3171 := (= #3153 #3153)
+#3172 := (and #3171 #3166)
+#3173 := (implies #3166 #3172)
+#3174 := (forall (vars (?x781 T5)) (:pat #3154) #3173)
+#3192 := (and #3174 #3191)
+decl uf_19 :: (-> T9 T5 int)
+decl uf_20 :: (-> T4 T9)
+#3162 := (uf_20 uf_287)
+#3163 := (uf_19 #3162 #23)
+#3164 := (pattern #3163)
+#3167 := (= #3163 #3163)
+#3168 := (and #3167 #3166)
+#3169 := (implies #3166 #3168)
+#3170 := (forall (vars (?x780 T5)) (:pat #3164) #3169)
+#3193 := (and #3170 #3192)
+decl uf_261 :: T8
+#2837 := uf_261
+#3155 := (uf_25 uf_287 #23)
+#3156 := (uf_13 #3155)
+#3157 := (uf_12 #3156)
+#3158 := (= #3157 uf_261)
+#3159 := (not #3158)
+#3160 := (implies #3159 #3159)
+#3161 := (forall (vars (?x779 T5)) (:pat #3154) #3160)
+#3194 := (and #3161 #3193)
+#3213 := (and #3194 #3212)
+#3214 := (and #3092 #3213)
+#3215 := (and #3092 #3214)
+#3216 := (and #3092 #3215)
+#3310 := (implies #3216 #3309)
+#3117 := (and #3115 #3116)
+#3118 := (and #3092 #3117)
+#3119 := (and #3092 #3118)
+#1 := true
+#3112 := (not true)
+#3120 := (and #3112 #3119)
+#3121 := (and #3092 #3120)
+#3151 := (implies #3121 #3150)
+#3311 := (and #3151 #3310)
+#3102 := (= #3101 uf_296)
+#3099 := (< uf_297 uf_286)
+#3103 := (and #3099 #3102)
+#3104 := (and #3103 #3092)
+#3096 := (<= #3073 uf_296)
+#3094 := (< #161 uf_298)
+#3095 := (and #412 #3094)
+#3097 := (implies #3095 #3096)
+#3098 := (forall (vars (?x775 int)) #3097)
+#3105 := (and #3098 #3104)
+#3093 := (<= uf_298 uf_286)
+#3106 := (and #3093 #3105)
+#3107 := (and #3092 #3106)
+#3089 := (<= uf_298 uf_76)
+#3088 := (<= 0::int uf_298)
+#3090 := (and #3088 #3089)
+#3108 := (and #3090 #3107)
+#3085 := (<= uf_297 uf_76)
+#3086 := (and #3084 #3085)
+#3109 := (and #3086 #3108)
+#3081 := (<= uf_296 uf_78)
+#3080 := (<= 0::int uf_296)
+#3082 := (and #3080 #3081)
+#3110 := (and #3082 #3109)
+#3077 := (= #3054 uf_295)
+#2975 := (< 0::int uf_286)
+#3078 := (and #2975 #3077)
+#3111 := (and #3078 #3110)
+#3312 := (implies #3111 #3311)
+#3313 := (and #3078 #3312)
+#3074 := (<= #3073 uf_295)
+#3070 := (< #161 1::int)
+#3071 := (and #412 #3070)
+#3075 := (implies #3071 #3074)
+#3076 := (forall (vars (?x773 int)) #3075)
+#3314 := (implies #3076 #3313)
+#3315 := (and #3076 #3314)
+#3069 := (<= 1::int uf_286)
+#3316 := (implies #3069 #3315)
+#3317 := (and #3069 #3316)
+#3060 := (<= 0::int 0::int)
+#3061 := (and #3060 #3060)
+#3059 := (<= 1::int 1::int)
+#3062 := (and #3059 #3061)
+#3063 := (and #3059 #3062)
+#3064 := (and #3058 #3063)
+#3065 := (and #3057 #3064)
+#3066 := (and #3056 #3065)
+#3067 := (and #3055 #3066)
+#3051 := (= #3050 uf_9)
+#3046 := (= #3045 uf_9)
+#3052 := (and #3046 #3051)
+#3068 := (and #3052 #3067)
+#3318 := (implies #3068 #3317)
+#3319 := (and #3052 #3318)
+#3048 := (= #3047 uf_9)
+#3049 := (and #3046 #3048)
+#3320 := (implies #3049 #3319)
+#3321 := (and #3049 #3320)
+#3043 := (= #3042 uf_9)
+#3322 := (implies #3043 #3321)
+#3323 := (and #3043 #3322)
+#3027 := (<= uf_286 uf_76)
+#3026 := (<= 0::int uf_286)
+#3028 := (and #3026 #3027)
+#3023 := (= #3021 uf_9)
+#3024 := (iff #3023 false)
+#3025 := (forall (vars (?x772 T5)) (:pat #3022) #3024)
+#3029 := (and #3025 #3028)
+#3030 := (and #3020 #3029)
+#3031 := (and #3017 #3030)
+#3016 := (and #3014 #3015)
+#3032 := (and #3016 #3031)
+#3011 := (< #3008 uf_290)
+#3012 := (forall (vars (?x771 T19)) (:pat #3009) #3011)
+#3033 := (and #3012 #3032)
+#3003 := (= #3002 uf_9)
+#3006 := (and #3003 #3005)
+#3034 := (and #3006 #3033)
+#3001 := (= #3000 uf_9)
+#3035 := (and #3001 #3034)
+#2994 := (= #2993 uf_9)
+#2991 := (= #2990 uf_14)
+#2992 := (not #2991)
+#2995 := (and #2992 #2994)
+#2989 := (= #2988 uf_9)
+#2996 := (and #2989 #2995)
+#2987 := (= #2986 uf_9)
+#2997 := (and #2987 #2996)
+#2985 := (= #2984 uf_26)
+#2998 := (and #2985 #2997)
+#2983 := (= #2982 uf_9)
+#2999 := (and #2983 #2998)
+#3036 := (and #2999 #3035)
+#3037 := (and #2975 #3036)
+#2974 := (< uf_286 1099511627776::int)
+#3038 := (and #2974 #3037)
+#2971 := (<= uf_285 uf_76)
+#2970 := (<= 0::int uf_285)
+#2972 := (and #2970 #2971)
+#3039 := (and #2972 #3038)
+#2967 := (<= uf_284 uf_76)
+#2966 := (<= 0::int uf_284)
+#2968 := (and #2966 #2967)
+#3040 := (and #2968 #3039)
+#2963 := (<= uf_283 uf_78)
+#2962 := (<= 0::int uf_283)
+#2964 := (and #2962 #2963)
+#3041 := (and #2964 #3040)
+#3324 := (implies #3041 #3323)
+#3325 := (not #3324)
+#13521 := (iff #3325 #13518)
+#12020 := (and #3084 #3091)
+#12605 := (and #3303 #12020)
+#12608 := (and #12020 #12605)
+#12614 := (not #12608)
+#12146 := (not #3138)
+#12147 := (or #12146 #3139)
+#12150 := (forall (vars (?x776 int)) #12147)
+#12156 := (not #12150)
+#12157 := (or #3146 #12156)
+#12162 := (and #12150 #12157)
+#12125 := (and #12119 #12122)
+#12128 := (and #12116 #12125)
+#12131 := (and #12113 #12128)
+#12134 := (and #12020 #12131)
+#12137 := (and #12020 #12134)
+#12140 := (and #12020 #12137)
+#12143 := (and up_216 #12140)
+#12168 := (not #12143)
+#12169 := (or #12168 #12162)
+#12174 := (and up_216 #12169)
+#12615 := (or #12174 #12614)
+#12530 := (and #3244 #12524)
+#12535 := (and #12521 #12530)
+#12538 := (and #12020 #12535)
+#12541 := (and #12020 #12538)
+#12544 := (and #12020 #12541)
+#12547 := (and #3286 #12544)
+#12550 := (and #12020 #12547)
+#12556 := (not #12550)
+#12449 := (and #3270 #12446)
+#12439 := (not #3266)
+#12440 := (or #12439 #3267)
+#12443 := (forall (vars (?x785 int)) #12440)
+#12455 := (not #12443)
+#12456 := (or #12455 #12449)
+#12461 := (and #12443 #12456)
+#12467 := (not #3264)
+#12468 := (or #12467 #12461)
+#12473 := (and #3264 #12468)
+#12427 := (and #3243 #3259)
+#12430 := (and #3258 #12427)
+#12412 := (+ 1::int uf_298)
+#12424 := (= uf_306 #12412)
+#12433 := (and #12424 #12430)
+#12418 := (<= #12412 uf_76)
+#12415 := (<= 0::int #12412)
+#12421 := (and #12415 #12418)
+#12436 := (and #12421 #12433)
+#12479 := (not #12436)
+#12480 := (or #12479 #12473)
+#12485 := (and #12421 #12480)
+#12557 := (or #12485 #12556)
+#12389 := (and #3244 #12383)
+#12394 := (and #12380 #12389)
+#12397 := (and #3091 #12394)
+#12400 := (and #3237 #12397)
+#12403 := (and #3236 #12400)
+#12406 := (and #12375 #12403)
+#12358 := (and #12346 #12355)
+#12409 := (and #12358 #12406)
+#12491 := (not #12409)
+#12492 := (or #12491 #12485)
+#12497 := (and #12358 #12492)
+#12504 := (or #12503 #12497)
+#12509 := (and #12352 #12504)
+#12369 := (and #3231 #12020)
+#12372 := (and #12020 #12369)
+#12515 := (not #12372)
+#12516 := (or #12515 #12509)
+#12562 := (and #12516 #12557)
+#12364 := (and #12020 #12358)
+#12568 := (not #12364)
+#12569 := (or #12568 #12562)
+#12574 := (and #12358 #12569)
+#12580 := (or #12503 #12574)
+#12585 := (and #12352 #12580)
+#12340 := (and #3217 #12020)
+#12343 := (and #12020 #12340)
+#12591 := (not #12343)
+#12592 := (or #12591 #12585)
+#12620 := (and #12592 #12615)
+#12307 := (and #3201 #12020)
+#12310 := (and #3198 #12307)
+#12313 := (and #3197 #12310)
+#12316 := (and #3196 #12313)
+#12319 := (and #3195 #12316)
+#12073 := (and #11903 #12067)
+#12322 := (and #12073 #12319)
+#12264 := (and #3186 #12261)
+#12267 := (and #3182 #12264)
+#12325 := (and #12267 #12322)
+#12328 := (and #12267 #12325)
+#12331 := (and #12020 #12328)
+#12334 := (and #12020 #12331)
+#12337 := (and #12020 #12334)
+#12626 := (not #12337)
+#12627 := (or #12626 #12620)
+#12033 := (and #3099 #12030)
+#12039 := (and #12020 #12033)
+#12023 := (not #3095)
+#12024 := (or #12023 #3096)
+#12027 := (forall (vars (?x775 int)) #12024)
+#12044 := (and #12027 #12039)
+#12047 := (and #3093 #12044)
+#12050 := (and #12020 #12047)
+#12053 := (and #3090 #12050)
+#12056 := (and #3086 #12053)
+#12059 := (and #3082 #12056)
+#12017 := (and #2975 #3055)
+#12062 := (and #12017 #12059)
+#12642 := (not #12062)
+#12643 := (or #12642 #12627)
+#12648 := (and #12017 #12643)
+#12008 := (not #3071)
+#12009 := (or #12008 #3074)
+#12012 := (forall (vars (?x773 int)) #12009)
+#12654 := (not #12012)
+#12655 := (or #12654 #12648)
+#12660 := (and #12012 #12655)
+#12666 := (not #3069)
+#12667 := (or #12666 #12660)
+#12672 := (and #3069 #12667)
+#11987 := (and #3059 #3060)
+#11990 := (and #3059 #11987)
+#11993 := (and #3058 #11990)
+#11996 := (and #3057 #11993)
+#11999 := (and #3056 #11996)
+#12002 := (and #3055 #11999)
+#11982 := (and #11970 #11979)
+#12005 := (and #11982 #12002)
+#12678 := (not #12005)
+#12679 := (or #12678 #12672)
+#12684 := (and #11982 #12679)
+#12691 := (or #12690 #12684)
+#12696 := (and #11976 #12691)
+#12703 := (or #12702 #12696)
+#12708 := (and #11967 #12703)
+#11926 := (and #3028 #11920)
+#11931 := (and #3020 #11926)
+#11934 := (and #3017 #11931)
+#11937 := (and #3016 #11934)
+#11940 := (and #3012 #11937)
+#11906 := (and #11900 #11903)
+#11943 := (and #11906 #11940)
+#11946 := (and #11897 #11943)
+#11882 := (and #11876 #11879)
+#11885 := (and #11870 #11882)
+#11888 := (and #11867 #11885)
+#11891 := (and #11864 #11888)
+#11894 := (and #11860 #11891)
+#11949 := (and #11894 #11946)
+#11952 := (and #2975 #11949)
+#11955 := (and #2974 #11952)
+#11958 := (and #2972 #11955)
+#11961 := (and #2968 #11958)
+#11964 := (and #2964 #11961)
+#12714 := (not #11964)
+#12715 := (or #12714 #12708)
+#12720 := (not #12715)
+#13519 := (iff #12720 #13518)
+#13516 := (iff #12715 #13515)
+#13513 := (iff #12708 #13512)
+#13510 := (iff #12703 #13509)
+#13507 := (iff #12696 #13504)
+#13501 := (and #11976 #13498)
+#13505 := (iff #13501 #13504)
+#13506 := [rewrite]: #13505
+#13502 := (iff #12696 #13501)
+#13499 := (iff #12691 #13498)
+#13496 := (iff #12684 #13493)
+#13490 := (and #11982 #13487)
+#13494 := (iff #13490 #13493)
+#13495 := [rewrite]: #13494
+#13491 := (iff #12684 #13490)
+#13488 := (iff #12679 #13487)
+#13485 := (iff #12672 #13484)
+#13482 := (iff #12667 #13481)
+#13479 := (iff #12660 #13478)
+#13476 := (iff #12655 #13475)
+#13473 := (iff #12648 #13470)
+#13442 := (and #12741 #3055)
+#13467 := (and #13442 #13462)
+#13471 := (iff #13467 #13470)
+#13472 := [rewrite]: #13471
+#13468 := (iff #12648 #13467)
+#13465 := (iff #12643 #13462)
+#13456 := (or #12962 #13353)
+#13459 := (or #13453 #13456)
+#13463 := (iff #13459 #13462)
+#13464 := [rewrite]: #13463
+#13460 := (iff #12643 #13459)
+#13457 := (iff #12627 #13456)
+#13354 := (iff #12620 #13353)
+#13351 := (iff #12615 #13350)
+#13348 := (iff #12614 #13347)
+#13345 := (iff #12608 #13342)
+#12915 := (and #12910 #12913)
+#13336 := (and #12967 #12915)
+#13339 := (and #12915 #13336)
+#13343 := (iff #13339 #13342)
+#13344 := [rewrite]: #13343
+#13340 := (iff #12608 #13339)
+#13337 := (iff #12605 #13336)
+#12916 := (iff #12020 #12915)
+#12912 := (iff #3091 #12913)
+#12914 := [rewrite]: #12912
+#12909 := (iff #3084 #12910)
+#12911 := [rewrite]: #12909
+#12917 := [monotonicity #12911 #12914]: #12916
+#13334 := (iff #3303 #12967)
+#13335 := [rewrite]: #13334
+#13338 := [monotonicity #13335 #12917]: #13337
+#13341 := [monotonicity #12917 #13338]: #13340
+#13346 := [trans #13341 #13344]: #13345
+#13349 := [monotonicity #13346]: #13348
+#13332 := (iff #12174 #13331)
+#13329 := (iff #12169 #13328)
+#13326 := (iff #12162 #13325)
+#13323 := (iff #12157 #13320)
+#13317 := (or #13311 #13314)
+#13321 := (iff #13317 #13320)
+#13322 := [rewrite]: #13321
+#13318 := (iff #12157 #13317)
+#13315 := (iff #12156 #13314)
+#13295 := (iff #12150 #13294)
+#13292 := (iff #12147 #13291)
+#13289 := (iff #3139 #13288)
+#13290 := [rewrite]: #13289
+#13284 := (iff #12146 #13283)
+#13281 := (iff #3138 #13278)
+#4362 := (and #4041 #4355)
+#13275 := (and #4362 #13272)
+#13279 := (iff #13275 #13278)
+#13280 := [rewrite]: #13279
+#13276 := (iff #3138 #13275)
+#13273 := (iff #3137 #13272)
+#13274 := [rewrite]: #13273
+#4363 := (iff #412 #4362)
+#4356 := (iff #411 #4355)
+#4357 := [rewrite]: #4356
+#4039 := (iff #285 #4041)
+#4040 := [rewrite]: #4039
+#4364 := [monotonicity #4040 #4357]: #4363
+#13277 := [monotonicity #4364 #13274]: #13276
+#13282 := [trans #13277 #13280]: #13281
+#13285 := [monotonicity #13282]: #13284
+#13293 := [monotonicity #13285 #13290]: #13292
+#13296 := [quant-intro #13293]: #13295
+#13316 := [monotonicity #13296]: #13315
+#13312 := (iff #3146 #13311)
+#13309 := (iff #3145 #13306)
+#13297 := (and #13272 #3142)
+#13300 := (and #4355 #13297)
+#13303 := (and #4041 #13300)
+#13307 := (iff #13303 #13306)
+#13308 := [rewrite]: #13307
+#13304 := (iff #3145 #13303)
+#13301 := (iff #3144 #13300)
+#13298 := (iff #3143 #13297)
+#13299 := [monotonicity #13274]: #13298
+#13302 := [monotonicity #4357 #13299]: #13301
+#13305 := [monotonicity #4040 #13302]: #13304
+#13310 := [trans #13305 #13308]: #13309
+#13313 := [quant-intro #13310]: #13312
+#13319 := [monotonicity #13313 #13316]: #13318
+#13324 := [trans #13319 #13322]: #13323
+#13327 := [monotonicity #13296 #13324]: #13326
+#13268 := (iff #12168 #13267)
+#13265 := (iff #12143 #13262)
+#13250 := (and #12915 #12131)
+#13253 := (and #12915 #13250)
+#13256 := (and #12915 #13253)
+#13259 := (and up_216 #13256)
+#13263 := (iff #13259 #13262)
+#13264 := [rewrite]: #13263
+#13260 := (iff #12143 #13259)
+#13257 := (iff #12140 #13256)
+#13254 := (iff #12137 #13253)
+#13251 := (iff #12134 #13250)
+#13252 := [monotonicity #12917]: #13251
+#13255 := [monotonicity #12917 #13252]: #13254
+#13258 := [monotonicity #12917 #13255]: #13257
+#13261 := [monotonicity #13258]: #13260
+#13266 := [trans #13261 #13264]: #13265
+#13269 := [monotonicity #13266]: #13268
+#13330 := [monotonicity #13269 #13327]: #13329
+#13333 := [monotonicity #13330]: #13332
+#13352 := [monotonicity #13333 #13349]: #13351
+#13248 := (iff #12592 #13247)
+#13245 := (iff #12585 #13242)
+#13239 := (and #12352 #13236)
+#13243 := (iff #13239 #13242)
+#13244 := [rewrite]: #13243
+#13240 := (iff #12585 #13239)
+#13237 := (iff #12580 #13236)
+#13234 := (iff #12574 #13231)
+#13228 := (and #12358 #13225)
+#13232 := (iff #13228 #13231)
+#13233 := [rewrite]: #13232
+#13229 := (iff #12574 #13228)
+#13226 := (iff #12569 #13225)
+#13223 := (iff #12562 #13222)
+#13220 := (iff #12557 #13219)
+#13217 := (iff #12556 #13216)
+#13214 := (iff #12550 #13211)
+#13019 := (and #12913 #13017)
+#13190 := (and #13019 #12524)
+#13193 := (and #12521 #13190)
+#13196 := (and #12915 #13193)
+#13199 := (and #12915 #13196)
+#13202 := (and #12915 #13199)
+#13205 := (and #12997 #13202)
+#13208 := (and #12915 #13205)
+#13212 := (iff #13208 #13211)
+#13213 := [rewrite]: #13212
+#13209 := (iff #12550 #13208)
+#13206 := (iff #12547 #13205)
+#13203 := (iff #12544 #13202)
+#13200 := (iff #12541 #13199)
+#13197 := (iff #12538 #13196)
+#13194 := (iff #12535 #13193)
+#13191 := (iff #12530 #13190)
+#13020 := (iff #3244 #13019)
+#13016 := (iff #3243 #13017)
+#13018 := [rewrite]: #13016
+#13021 := [monotonicity #12914 #13018]: #13020
+#13192 := [monotonicity #13021]: #13191
+#13195 := [monotonicity #13192]: #13194
+#13198 := [monotonicity #12917 #13195]: #13197
+#13201 := [monotonicity #12917 #13198]: #13200
+#13204 := [monotonicity #12917 #13201]: #13203
+#13188 := (iff #3286 #12997)
+#13189 := [rewrite]: #13188
+#13207 := [monotonicity #13189 #13204]: #13206
+#13210 := [monotonicity #12917 #13207]: #13209
+#13215 := [trans #13210 #13213]: #13214
+#13218 := [monotonicity #13215]: #13217
+#13161 := (iff #12485 #13158)
+#13075 := (and #13051 #13054)
+#13155 := (and #13075 #13152)
+#13159 := (iff #13155 #13158)
+#13160 := [rewrite]: #13159
+#13156 := (iff #12485 #13155)
+#13153 := (iff #12480 #13152)
+#13150 := (iff #12473 #13149)
+#13147 := (iff #12468 #13146)
+#13144 := (iff #12461 #13143)
+#13141 := (iff #12456 #13140)
+#13138 := (iff #12449 #13135)
+#13132 := (and #13129 #12446)
+#13136 := (iff #13132 #13135)
+#13137 := [rewrite]: #13136
+#13133 := (iff #12449 #13132)
+#13130 := (iff #3270 #13129)
+#13131 := [rewrite]: #13130
+#13134 := [monotonicity #13131]: #13133
+#13139 := [trans #13134 #13137]: #13138
+#13124 := (iff #12455 #13123)
+#13121 := (iff #12443 #13120)
+#13118 := (iff #12440 #13117)
+#13115 := (iff #3267 #13114)
+#13116 := [rewrite]: #13115
+#13110 := (iff #12439 #13109)
+#13107 := (iff #3266 #13104)
+#13101 := (and #4362 #13098)
+#13105 := (iff #13101 #13104)
+#13106 := [rewrite]: #13105
+#13102 := (iff #3266 #13101)
+#13099 := (iff #3265 #13098)
+#13100 := [rewrite]: #13099
+#13103 := [monotonicity #4364 #13100]: #13102
+#13108 := [trans #13103 #13106]: #13107
+#13111 := [monotonicity #13108]: #13110
+#13119 := [monotonicity #13111 #13116]: #13118
+#13122 := [quant-intro #13119]: #13121
+#13125 := [monotonicity #13122]: #13124
+#13142 := [monotonicity #13125 #13139]: #13141
+#13145 := [monotonicity #13122 #13142]: #13144
+#13094 := (iff #12467 #13093)
+#13091 := (iff #3264 #13089)
+#13092 := [rewrite]: #13091
+#13095 := [monotonicity #13092]: #13094
+#13148 := [monotonicity #13095 #13145]: #13147
+#13151 := [monotonicity #13092 #13148]: #13150
+#13087 := (iff #12479 #13086)
+#13084 := (iff #12436 #13081)
+#13066 := (and #13017 #13059)
+#13069 := (and #3258 #13066)
+#13072 := (and #13061 #13069)
+#13078 := (and #13075 #13072)
+#13082 := (iff #13078 #13081)
+#13083 := [rewrite]: #13082
+#13079 := (iff #12436 #13078)
+#13073 := (iff #12433 #13072)
+#13070 := (iff #12430 #13069)
+#13067 := (iff #12427 #13066)
+#13058 := (iff #3259 #13059)
+#13060 := [rewrite]: #13058
+#13068 := [monotonicity #13018 #13060]: #13067
+#13071 := [monotonicity #13068]: #13070
+#13064 := (iff #12424 #13061)
+#13065 := [rewrite]: #13064
+#13074 := [monotonicity #13065 #13071]: #13073
+#13076 := (iff #12421 #13075)
+#13056 := (iff #12418 #13054)
+#13057 := [rewrite]: #13056
+#13052 := (iff #12415 #13051)
+#13053 := [rewrite]: #13052
+#13077 := [monotonicity #13053 #13057]: #13076
+#13080 := [monotonicity #13077 #13074]: #13079
+#13085 := [trans #13080 #13083]: #13084
+#13088 := [monotonicity #13085]: #13087
+#13154 := [monotonicity #13088 #13151]: #13153
+#13157 := [monotonicity #13077 #13154]: #13156
+#13162 := [trans #13157 #13160]: #13161
+#13221 := [monotonicity #13162 #13218]: #13220
+#13186 := (iff #12516 #13185)
+#13183 := (iff #12509 #13180)
+#13177 := (and #12352 #13174)
+#13181 := (iff #13177 #13180)
+#13182 := [rewrite]: #13181
+#13178 := (iff #12509 #13177)
+#13175 := (iff #12504 #13174)
+#13172 := (iff #12497 #13169)
+#13166 := (and #12358 #13163)
+#13170 := (iff #13166 #13169)
+#13171 := [rewrite]: #13170
+#13167 := (iff #12497 #13166)
+#13164 := (iff #12492 #13163)
+#13049 := (iff #12491 #13048)
+#13046 := (iff #12409 #13043)
+#13022 := (and #13019 #12383)
+#13025 := (and #12380 #13022)
+#13028 := (and #12913 #13025)
+#13031 := (and #3237 #13028)
+#13034 := (and #3236 #13031)
+#13037 := (and #12375 #13034)
+#13040 := (and #12358 #13037)
+#13044 := (iff #13040 #13043)
+#13045 := [rewrite]: #13044
+#13041 := (iff #12409 #13040)
+#13038 := (iff #12406 #13037)
+#13035 := (iff #12403 #13034)
+#13032 := (iff #12400 #13031)
+#13029 := (iff #12397 #13028)
+#13026 := (iff #12394 #13025)
+#13023 := (iff #12389 #13022)
+#13024 := [monotonicity #13021]: #13023
+#13027 := [monotonicity #13024]: #13026
+#13030 := [monotonicity #12914 #13027]: #13029
+#13033 := [monotonicity #13030]: #13032
+#13036 := [monotonicity #13033]: #13035
+#13039 := [monotonicity #13036]: #13038
+#13042 := [monotonicity #13039]: #13041
+#13047 := [trans #13042 #13045]: #13046
+#13050 := [monotonicity #13047]: #13049
+#13165 := [monotonicity #13050 #13162]: #13164
+#13168 := [monotonicity #13165]: #13167
+#13173 := [trans #13168 #13171]: #13172
+#13176 := [monotonicity #13173]: #13175
+#13179 := [monotonicity #13176]: #13178
+#13184 := [trans #13179 #13182]: #13183
+#13014 := (iff #12515 #13013)
+#13011 := (iff #12372 #13008)
+#13002 := (and #12996 #12915)
+#13005 := (and #12915 #13002)
+#13009 := (iff #13005 #13008)
+#13010 := [rewrite]: #13009
+#13006 := (iff #12372 #13005)
+#13003 := (iff #12369 #13002)
+#13000 := (iff #3231 #12996)
+#13001 := [rewrite]: #13000
+#13004 := [monotonicity #13001 #12917]: #13003
+#13007 := [monotonicity #12917 #13004]: #13006
+#13012 := [trans #13007 #13010]: #13011
+#13015 := [monotonicity #13012]: #13014
+#13187 := [monotonicity #13015 #13184]: #13186
+#13224 := [monotonicity #13187 #13221]: #13223
+#12994 := (iff #12568 #12993)
+#12991 := (iff #12364 #12988)
+#12985 := (and #12915 #12358)
+#12989 := (iff #12985 #12988)
+#12990 := [rewrite]: #12989
+#12986 := (iff #12364 #12985)
+#12987 := [monotonicity #12917]: #12986
+#12992 := [trans #12987 #12990]: #12991
+#12995 := [monotonicity #12992]: #12994
+#13227 := [monotonicity #12995 #13224]: #13226
+#13230 := [monotonicity #13227]: #13229
+#13235 := [trans #13230 #13233]: #13234
+#13238 := [monotonicity #13235]: #13237
+#13241 := [monotonicity #13238]: #13240
+#13246 := [trans #13241 #13244]: #13245
+#12983 := (iff #12591 #12982)
+#12980 := (iff #12343 #12977)
+#12971 := (and #12968 #12915)
+#12974 := (and #12915 #12971)
+#12978 := (iff #12974 #12977)
+#12979 := [rewrite]: #12978
+#12975 := (iff #12343 #12974)
+#12972 := (iff #12340 #12971)
+#12969 := (iff #3217 #12968)
+#12970 := [rewrite]: #12969
+#12973 := [monotonicity #12970 #12917]: #12972
+#12976 := [monotonicity #12917 #12973]: #12975
+#12981 := [trans #12976 #12979]: #12980
+#12984 := [monotonicity #12981]: #12983
+#13249 := [monotonicity #12984 #13246]: #13248
+#13355 := [monotonicity #13249 #13352]: #13354
+#12963 := (iff #12626 #12962)
+#12960 := (iff #12337 #12957)
+#12918 := (and #3201 #12915)
+#12921 := (and #3198 #12918)
+#12924 := (and #3197 #12921)
+#12927 := (and #3196 #12924)
+#12930 := (and #3195 #12927)
+#12933 := (and #12073 #12930)
+#12936 := (and true #12261)
+#12939 := (and true #12936)
+#12942 := (and #12939 #12933)
+#12945 := (and #12939 #12942)
+#12948 := (and #12915 #12945)
+#12951 := (and #12915 #12948)
+#12954 := (and #12915 #12951)
+#12958 := (iff #12954 #12957)
+#12959 := [rewrite]: #12958
+#12955 := (iff #12337 #12954)
+#12952 := (iff #12334 #12951)
+#12949 := (iff #12331 #12948)
+#12946 := (iff #12328 #12945)
+#12943 := (iff #12325 #12942)
+#12934 := (iff #12322 #12933)
+#12931 := (iff #12319 #12930)
+#12928 := (iff #12316 #12927)
+#12925 := (iff #12313 #12924)
+#12922 := (iff #12310 #12921)
+#12919 := (iff #12307 #12918)
+#12920 := [monotonicity #12917]: #12919
+#12923 := [monotonicity #12920]: #12922
+#12926 := [monotonicity #12923]: #12925
+#12929 := [monotonicity #12926]: #12928
+#12932 := [monotonicity #12929]: #12931
+#12935 := [monotonicity #12932]: #12934
+#12940 := (iff #12267 #12939)
+#12937 := (iff #12264 #12936)
+#12905 := (iff #3186 true)
+#12900 := (forall (vars (?x783 T5)) (:pat #3184) true)
+#12903 := (iff #12900 true)
+#12904 := [elim-unused]: #12903
+#12901 := (iff #3186 #12900)
+#12898 := (iff #3185 true)
+#12899 := [rewrite]: #12898
+#12902 := [quant-intro #12899]: #12901
+#12906 := [trans #12902 #12904]: #12905
+#12938 := [monotonicity #12906]: #12937
+#12907 := (iff #3182 true)
+#12908 := [rewrite]: #12907
+#12941 := [monotonicity #12908 #12938]: #12940
+#12944 := [monotonicity #12941 #12935]: #12943
+#12947 := [monotonicity #12941 #12944]: #12946
+#12950 := [monotonicity #12917 #12947]: #12949
+#12953 := [monotonicity #12917 #12950]: #12952
+#12956 := [monotonicity #12917 #12953]: #12955
+#12961 := [trans #12956 #12959]: #12960
+#12964 := [monotonicity #12961]: #12963
+#13458 := [monotonicity #12964 #13355]: #13457
+#13454 := (iff #12642 #13453)
+#13451 := (iff #12062 #13448)
+#13409 := (and #13359 #12030)
+#13412 := (and #12915 #13409)
+#13415 := (and #13386 #13412)
+#13418 := (and #13389 #13415)
+#13421 := (and #12915 #13418)
+#13424 := (and #13392 #13395)
+#13427 := (and #13424 #13421)
+#13430 := (and #12910 #13398)
+#13433 := (and #13430 #13427)
+#13436 := (and #13402 #13405)
+#13439 := (and #13436 #13433)
+#13445 := (and #13442 #13439)
+#13449 := (iff #13445 #13448)
+#13450 := [rewrite]: #13449
+#13446 := (iff #12062 #13445)
+#13440 := (iff #12059 #13439)
+#13434 := (iff #12056 #13433)
+#13428 := (iff #12053 #13427)
+#13422 := (iff #12050 #13421)
+#13419 := (iff #12047 #13418)
+#13416 := (iff #12044 #13415)
+#13413 := (iff #12039 #13412)
+#13410 := (iff #12033 #13409)
+#13360 := (iff #3099 #13359)
+#13361 := [rewrite]: #13360
+#13411 := [monotonicity #13361]: #13410
+#13414 := [monotonicity #12917 #13411]: #13413
+#13387 := (iff #12027 #13386)
+#13384 := (iff #12024 #13383)
+#13381 := (iff #3096 #13380)
+#13382 := [rewrite]: #13381
+#13376 := (iff #12023 #13375)
+#13373 := (iff #3095 #13370)
+#13367 := (and #4362 #13364)
+#13371 := (iff #13367 #13370)
+#13372 := [rewrite]: #13371
+#13368 := (iff #3095 #13367)
+#13365 := (iff #3094 #13364)
+#13366 := [rewrite]: #13365
+#13369 := [monotonicity #4364 #13366]: #13368
+#13374 := [trans #13369 #13372]: #13373
+#13377 := [monotonicity #13374]: #13376
+#13385 := [monotonicity #13377 #13382]: #13384
+#13388 := [quant-intro #13385]: #13387
+#13417 := [monotonicity #13388 #13414]: #13416
+#13390 := (iff #3093 #13389)
+#13391 := [rewrite]: #13390
+#13420 := [monotonicity #13391 #13417]: #13419
+#13423 := [monotonicity #12917 #13420]: #13422
+#13425 := (iff #3090 #13424)
+#13396 := (iff #3089 #13395)
+#13397 := [rewrite]: #13396
+#13393 := (iff #3088 #13392)
+#13394 := [rewrite]: #13393
+#13426 := [monotonicity #13394 #13397]: #13425
+#13429 := [monotonicity #13426 #13423]: #13428
+#13431 := (iff #3086 #13430)
+#13400 := (iff #3085 #13398)
+#13401 := [rewrite]: #13400
+#13432 := [monotonicity #12911 #13401]: #13431
+#13435 := [monotonicity #13432 #13429]: #13434
+#13437 := (iff #3082 #13436)
+#13407 := (iff #3081 #13405)
+#13408 := [rewrite]: #13407
+#13403 := (iff #3080 #13402)
+#13404 := [rewrite]: #13403
+#13438 := [monotonicity #13404 #13408]: #13437
+#13441 := [monotonicity #13438 #13435]: #13440
+#13443 := (iff #12017 #13442)
+#12742 := (iff #2975 #12741)
+#12743 := [rewrite]: #12742
+#13444 := [monotonicity #12743]: #13443
+#13447 := [monotonicity #13444 #13441]: #13446
+#13452 := [trans #13447 #13450]: #13451
+#13455 := [monotonicity #13452]: #13454
+#13461 := [monotonicity #13455 #13458]: #13460
+#13466 := [trans #13461 #13464]: #13465
+#13469 := [monotonicity #13444 #13466]: #13468
+#13474 := [trans #13469 #13472]: #13473
+#12896 := (iff #12654 #12895)
+#12893 := (iff #12012 #12892)
+#12890 := (iff #12009 #12889)
+#12884 := (iff #3074 #12885)
+#12888 := [rewrite]: #12884
+#12882 := (iff #12008 #12881)
+#12879 := (iff #3071 #12876)
+#12873 := (and #4362 #12870)
+#12877 := (iff #12873 #12876)
+#12878 := [rewrite]: #12877
+#12874 := (iff #3071 #12873)
+#12871 := (iff #3070 #12870)
+#12872 := [rewrite]: #12871
+#12875 := [monotonicity #4364 #12872]: #12874
+#12880 := [trans #12875 #12878]: #12879
+#12883 := [monotonicity #12880]: #12882
+#12891 := [monotonicity #12883 #12888]: #12890
+#12894 := [quant-intro #12891]: #12893
+#12897 := [monotonicity #12894]: #12896
+#13477 := [monotonicity #12897 #13474]: #13476
+#13480 := [monotonicity #12894 #13477]: #13479
+#12867 := (iff #12666 #12866)
+#12864 := (iff #3069 #12863)
+#12865 := [rewrite]: #12864
+#12868 := [monotonicity #12865]: #12867
+#13483 := [monotonicity #12868 #13480]: #13482
+#13486 := [monotonicity #12865 #13483]: #13485
+#12861 := (iff #12678 #12860)
+#12858 := (iff #12005 #12855)
+#12293 := (and true true)
+#12837 := (and true #12293)
+#12840 := (and #3058 #12837)
+#12843 := (and #3057 #12840)
+#12846 := (and #3056 #12843)
+#12849 := (and #3055 #12846)
+#12852 := (and #11982 #12849)
+#12856 := (iff #12852 #12855)
+#12857 := [rewrite]: #12856
+#12853 := (iff #12005 #12852)
+#12850 := (iff #12002 #12849)
+#12847 := (iff #11999 #12846)
+#12844 := (iff #11996 #12843)
+#12841 := (iff #11993 #12840)
+#12838 := (iff #11990 #12837)
+#12835 := (iff #11987 #12293)
+#12833 := (iff #3060 true)
+#12834 := [rewrite]: #12833
+#12831 := (iff #3059 true)
+#12832 := [rewrite]: #12831
+#12836 := [monotonicity #12832 #12834]: #12835
+#12839 := [monotonicity #12832 #12836]: #12838
+#12842 := [monotonicity #12839]: #12841
+#12845 := [monotonicity #12842]: #12844
+#12848 := [monotonicity #12845]: #12847
+#12851 := [monotonicity #12848]: #12850
+#12854 := [monotonicity #12851]: #12853
+#12859 := [trans #12854 #12857]: #12858
+#12862 := [monotonicity #12859]: #12861
+#13489 := [monotonicity #12862 #13486]: #13488
+#13492 := [monotonicity #13489]: #13491
+#13497 := [trans #13492 #13495]: #13496
+#13500 := [monotonicity #13497]: #13499
+#13503 := [monotonicity #13500]: #13502
+#13508 := [trans #13503 #13506]: #13507
+#13511 := [monotonicity #13508]: #13510
+#13514 := [monotonicity #13511]: #13513
+#12829 := (iff #12714 #12828)
+#12826 := (iff #11964 #12823)
+#12772 := (and #12724 #12726)
+#12775 := (and #12772 #11920)
+#12778 := (and #3020 #12775)
+#12781 := (and #3017 #12778)
+#12784 := (and #3016 #12781)
+#12787 := (and #12737 #12784)
+#12790 := (and #11906 #12787)
+#12793 := (and #11897 #12790)
+#12796 := (and #11894 #12793)
+#12799 := (and #12741 #12796)
+#12802 := (and #12745 #12799)
+#12805 := (and #12749 #12751)
+#12808 := (and #12805 #12802)
+#12811 := (and #12757 #12759)
+#12814 := (and #12811 #12808)
+#12817 := (and #12765 #12767)
+#12820 := (and #12817 #12814)
+#12824 := (iff #12820 #12823)
+#12825 := [rewrite]: #12824
+#12821 := (iff #11964 #12820)
+#12815 := (iff #11961 #12814)
+#12809 := (iff #11958 #12808)
+#12803 := (iff #11955 #12802)
+#12800 := (iff #11952 #12799)
+#12797 := (iff #11949 #12796)
+#12794 := (iff #11946 #12793)
+#12791 := (iff #11943 #12790)
+#12788 := (iff #11940 #12787)
+#12785 := (iff #11937 #12784)
+#12782 := (iff #11934 #12781)
+#12779 := (iff #11931 #12778)
+#12776 := (iff #11926 #12775)
+#12773 := (iff #3028 #12772)
+#12729 := (iff #3027 #12726)
+#12730 := [rewrite]: #12729
+#12723 := (iff #3026 #12724)
+#12725 := [rewrite]: #12723
+#12774 := [monotonicity #12725 #12730]: #12773
+#12777 := [monotonicity #12774]: #12776
+#12780 := [monotonicity #12777]: #12779
+#12783 := [monotonicity #12780]: #12782
+#12786 := [monotonicity #12783]: #12785
+#12738 := (iff #3012 #12737)
+#12735 := (iff #3011 #12731)
+#12736 := [rewrite]: #12735
+#12739 := [quant-intro #12736]: #12738
+#12789 := [monotonicity #12739 #12786]: #12788
+#12792 := [monotonicity #12789]: #12791
+#12795 := [monotonicity #12792]: #12794
+#12798 := [monotonicity #12795]: #12797
+#12801 := [monotonicity #12743 #12798]: #12800
+#12746 := (iff #2974 #12745)
+#12747 := [rewrite]: #12746
+#12804 := [monotonicity #12747 #12801]: #12803
+#12806 := (iff #2972 #12805)
+#12754 := (iff #2971 #12751)
+#12755 := [rewrite]: #12754
+#12748 := (iff #2970 #12749)
+#12750 := [rewrite]: #12748
+#12807 := [monotonicity #12750 #12755]: #12806
+#12810 := [monotonicity #12807 #12804]: #12809
+#12812 := (iff #2968 #12811)
+#12762 := (iff #2967 #12759)
+#12763 := [rewrite]: #12762
+#12756 := (iff #2966 #12757)
+#12758 := [rewrite]: #12756
+#12813 := [monotonicity #12758 #12763]: #12812
+#12816 := [monotonicity #12813 #12810]: #12815
+#12818 := (iff #2964 #12817)
+#12770 := (iff #2963 #12767)
+#12771 := [rewrite]: #12770
+#12764 := (iff #2962 #12765)
+#12766 := [rewrite]: #12764
+#12819 := [monotonicity #12766 #12771]: #12818
+#12822 := [monotonicity #12819 #12816]: #12821
+#12827 := [trans #12822 #12825]: #12826
+#12830 := [monotonicity #12827]: #12829
+#13517 := [monotonicity #12830 #13514]: #13516
+#13520 := [monotonicity #13517]: #13519
+#12721 := (iff #3325 #12720)
+#12718 := (iff #3324 #12715)
+#12711 := (implies #11964 #12708)
+#12716 := (iff #12711 #12715)
+#12717 := [rewrite]: #12716
+#12712 := (iff #3324 #12711)
+#12709 := (iff #3323 #12708)
+#12706 := (iff #3322 #12703)
+#12699 := (implies #11967 #12696)
+#12704 := (iff #12699 #12703)
+#12705 := [rewrite]: #12704
+#12700 := (iff #3322 #12699)
+#12697 := (iff #3321 #12696)
+#12694 := (iff #3320 #12691)
+#12687 := (implies #11976 #12684)
+#12692 := (iff #12687 #12691)
+#12693 := [rewrite]: #12692
+#12688 := (iff #3320 #12687)
+#12685 := (iff #3319 #12684)
+#12682 := (iff #3318 #12679)
+#12675 := (implies #12005 #12672)
+#12680 := (iff #12675 #12679)
+#12681 := [rewrite]: #12680
+#12676 := (iff #3318 #12675)
+#12673 := (iff #3317 #12672)
+#12670 := (iff #3316 #12667)
+#12663 := (implies #3069 #12660)
+#12668 := (iff #12663 #12667)
+#12669 := [rewrite]: #12668
+#12664 := (iff #3316 #12663)
+#12661 := (iff #3315 #12660)
+#12658 := (iff #3314 #12655)
+#12651 := (implies #12012 #12648)
+#12656 := (iff #12651 #12655)
+#12657 := [rewrite]: #12656
+#12652 := (iff #3314 #12651)
+#12649 := (iff #3313 #12648)
+#12646 := (iff #3312 #12643)
+#12639 := (implies #12062 #12627)
+#12644 := (iff #12639 #12643)
+#12645 := [rewrite]: #12644
+#12640 := (iff #3312 #12639)
+#12637 := (iff #3311 #12627)
+#12632 := (and true #12627)
+#12635 := (iff #12632 #12627)
+#12636 := [rewrite]: #12635
+#12633 := (iff #3311 #12632)
+#12630 := (iff #3310 #12627)
+#12623 := (implies #12337 #12620)
+#12628 := (iff #12623 #12627)
+#12629 := [rewrite]: #12628
+#12624 := (iff #3310 #12623)
+#12621 := (iff #3309 #12620)
+#12618 := (iff #3308 #12615)
+#12611 := (implies #12608 #12174)
+#12616 := (iff #12611 #12615)
+#12617 := [rewrite]: #12616
+#12612 := (iff #3308 #12611)
+#12175 := (iff #3150 #12174)
+#12172 := (iff #3149 #12169)
+#12165 := (implies #12143 #12162)
+#12170 := (iff #12165 #12169)
+#12171 := [rewrite]: #12170
+#12166 := (iff #3149 #12165)
+#12163 := (iff #3148 #12162)
+#12160 := (iff #3147 #12157)
+#12153 := (implies #12150 #3146)
+#12158 := (iff #12153 #12157)
+#12159 := [rewrite]: #12158
+#12154 := (iff #3147 #12153)
+#12151 := (iff #3141 #12150)
+#12148 := (iff #3140 #12147)
+#12149 := [rewrite]: #12148
+#12152 := [quant-intro #12149]: #12151
+#12155 := [monotonicity #12152]: #12154
+#12161 := [trans #12155 #12159]: #12160
+#12164 := [monotonicity #12152 #12161]: #12163
+#12144 := (iff #3136 #12143)
+#12141 := (iff #3135 #12140)
+#12138 := (iff #3134 #12137)
+#12135 := (iff #3133 #12134)
+#12132 := (iff #3132 #12131)
+#12129 := (iff #3131 #12128)
+#12126 := (iff #3130 #12125)
+#12123 := (iff #3129 #12122)
+#12124 := [rewrite]: #12123
+#12120 := (iff #3127 #12119)
+#12121 := [rewrite]: #12120
+#12127 := [monotonicity #12121 #12124]: #12126
+#12117 := (iff #3125 #12116)
+#12118 := [rewrite]: #12117
+#12130 := [monotonicity #12118 #12127]: #12129
+#12114 := (iff #3123 #12113)
+#12115 := [rewrite]: #12114
+#12133 := [monotonicity #12115 #12130]: #12132
+#12021 := (iff #3092 #12020)
+#12022 := [rewrite]: #12021
+#12136 := [monotonicity #12022 #12133]: #12135
+#12139 := [monotonicity #12022 #12136]: #12138
+#12142 := [monotonicity #12022 #12139]: #12141
+#12145 := [monotonicity #12142]: #12144
+#12167 := [monotonicity #12145 #12164]: #12166
+#12173 := [trans #12167 #12171]: #12172
+#12176 := [monotonicity #12173]: #12175
+#12609 := (iff #3307 #12608)
+#12606 := (iff #3306 #12605)
+#12603 := (iff #3305 #12020)
+#12078 := (and #12020 #12020)
+#12081 := (iff #12078 #12020)
+#12082 := [rewrite]: #12081
+#12601 := (iff #3305 #12078)
+#12599 := (iff #3304 #12020)
+#12597 := (iff #3304 #12078)
+#12083 := (iff #3116 #12020)
+#12079 := (iff #3116 #12078)
+#12080 := [monotonicity #12022 #12022]: #12079
+#12084 := [trans #12080 #12082]: #12083
+#12598 := [monotonicity #12022 #12084]: #12597
+#12600 := [trans #12598 #12082]: #12599
+#12602 := [monotonicity #12022 #12600]: #12601
+#12604 := [trans #12602 #12082]: #12603
+#12607 := [monotonicity #12604]: #12606
+#12610 := [monotonicity #12022 #12607]: #12609
+#12613 := [monotonicity #12610 #12176]: #12612
+#12619 := [trans #12613 #12617]: #12618
+#12595 := (iff #3302 #12592)
+#12588 := (implies #12343 #12585)
+#12593 := (iff #12588 #12592)
+#12594 := [rewrite]: #12593
+#12589 := (iff #3302 #12588)
+#12586 := (iff #3301 #12585)
+#12583 := (iff #3300 #12580)
+#12577 := (implies #12352 #12574)
+#12581 := (iff #12577 #12580)
+#12582 := [rewrite]: #12581
+#12578 := (iff #3300 #12577)
+#12575 := (iff #3299 #12574)
+#12572 := (iff #3298 #12569)
+#12565 := (implies #12364 #12562)
+#12570 := (iff #12565 #12569)
+#12571 := [rewrite]: #12570
+#12566 := (iff #3298 #12565)
+#12563 := (iff #3297 #12562)
+#12560 := (iff #3296 #12557)
+#12553 := (implies #12550 #12485)
+#12558 := (iff #12553 #12557)
+#12559 := [rewrite]: #12558
+#12554 := (iff #3296 #12553)
+#12486 := (iff #3280 #12485)
+#12483 := (iff #3279 #12480)
+#12476 := (implies #12436 #12473)
+#12481 := (iff #12476 #12480)
+#12482 := [rewrite]: #12481
+#12477 := (iff #3279 #12476)
+#12474 := (iff #3278 #12473)
+#12471 := (iff #3277 #12468)
+#12464 := (implies #3264 #12461)
+#12469 := (iff #12464 #12468)
+#12470 := [rewrite]: #12469
+#12465 := (iff #3277 #12464)
+#12462 := (iff #3276 #12461)
+#12459 := (iff #3275 #12456)
+#12452 := (implies #12443 #12449)
+#12457 := (iff #12452 #12456)
+#12458 := [rewrite]: #12457
+#12453 := (iff #3275 #12452)
+#12450 := (iff #3274 #12449)
+#12447 := (iff #3273 #12446)
+#12448 := [rewrite]: #12447
+#12451 := [monotonicity #12448]: #12450
+#12444 := (iff #3269 #12443)
+#12441 := (iff #3268 #12440)
+#12442 := [rewrite]: #12441
+#12445 := [quant-intro #12442]: #12444
+#12454 := [monotonicity #12445 #12451]: #12453
+#12460 := [trans #12454 #12458]: #12459
+#12463 := [monotonicity #12445 #12460]: #12462
+#12466 := [monotonicity #12463]: #12465
+#12472 := [trans #12466 #12470]: #12471
+#12475 := [monotonicity #12472]: #12474
+#12437 := (iff #3263 #12436)
+#12434 := (iff #3262 #12433)
+#12431 := (iff #3261 #12430)
+#12428 := (iff #3260 #12427)
+#12429 := [rewrite]: #12428
+#12432 := [monotonicity #12429]: #12431
+#12425 := (iff #3257 #12424)
+#12413 := (= #3252 #12412)
+#12414 := [rewrite]: #12413
+#12426 := [monotonicity #12414]: #12425
+#12435 := [monotonicity #12426 #12432]: #12434
+#12422 := (iff #3255 #12421)
+#12419 := (iff #3254 #12418)
+#12420 := [monotonicity #12414]: #12419
+#12416 := (iff #3253 #12415)
+#12417 := [monotonicity #12414]: #12416
+#12423 := [monotonicity #12417 #12420]: #12422
+#12438 := [monotonicity #12423 #12435]: #12437
+#12478 := [monotonicity #12438 #12475]: #12477
+#12484 := [trans #12478 #12482]: #12483
+#12487 := [monotonicity #12423 #12484]: #12486
+#12551 := (iff #3295 #12550)
+#12548 := (iff #3294 #12547)
+#12545 := (iff #3293 #12544)
+#12542 := (iff #3292 #12541)
+#12539 := (iff #3291 #12538)
+#12536 := (iff #3290 #12535)
+#12533 := (iff #3289 #12530)
+#12527 := (and #12524 #3244)
+#12531 := (iff #12527 #12530)
+#12532 := [rewrite]: #12531
+#12528 := (iff #3289 #12527)
+#12525 := (iff #3288 #12524)
+#12526 := [rewrite]: #12525
+#12529 := [monotonicity #12526]: #12528
+#12534 := [trans #12529 #12532]: #12533
+#12522 := (iff #3287 #12521)
+#12523 := [rewrite]: #12522
+#12537 := [monotonicity #12523 #12534]: #12536
+#12540 := [monotonicity #12022 #12537]: #12539
+#12543 := [monotonicity #12022 #12540]: #12542
+#12546 := [monotonicity #12022 #12543]: #12545
+#12549 := [monotonicity #12546]: #12548
+#12552 := [monotonicity #12022 #12549]: #12551
+#12555 := [monotonicity #12552 #12487]: #12554
+#12561 := [trans #12555 #12559]: #12560
+#12519 := (iff #3285 #12516)
+#12512 := (implies #12372 #12509)
+#12517 := (iff #12512 #12516)
+#12518 := [rewrite]: #12517
+#12513 := (iff #3285 #12512)
+#12510 := (iff #3284 #12509)
+#12507 := (iff #3283 #12504)
+#12500 := (implies #12352 #12497)
+#12505 := (iff #12500 #12504)
+#12506 := [rewrite]: #12505
+#12501 := (iff #3283 #12500)
+#12498 := (iff #3282 #12497)
+#12495 := (iff #3281 #12492)
+#12488 := (implies #12409 #12485)
+#12493 := (iff #12488 #12492)
+#12494 := [rewrite]: #12493
+#12489 := (iff #3281 #12488)
+#12410 := (iff #3251 #12409)
+#12407 := (iff #3250 #12406)
+#12404 := (iff #3249 #12403)
+#12401 := (iff #3248 #12400)
+#12398 := (iff #3247 #12397)
+#12395 := (iff #3246 #12394)
+#12392 := (iff #3245 #12389)
+#12386 := (and #12383 #3244)
+#12390 := (iff #12386 #12389)
+#12391 := [rewrite]: #12390
+#12387 := (iff #3245 #12386)
+#12384 := (iff #3242 #12383)
+#12385 := [rewrite]: #12384
+#12388 := [monotonicity #12385]: #12387
+#12393 := [trans #12388 #12391]: #12392
+#12381 := (iff #3240 #12380)
+#12382 := [rewrite]: #12381
+#12396 := [monotonicity #12382 #12393]: #12395
+#12378 := (iff #3238 #3091)
+#12379 := [rewrite]: #12378
+#12399 := [monotonicity #12379 #12396]: #12398
+#12402 := [monotonicity #12399]: #12401
+#12405 := [monotonicity #12402]: #12404
+#12376 := (iff #3235 #12375)
+#12377 := [rewrite]: #12376
+#12408 := [monotonicity #12377 #12405]: #12407
+#12359 := (iff #3228 #12358)
+#12356 := (iff #3227 #12355)
+#12357 := [rewrite]: #12356
+#12347 := (iff #3222 #12346)
+#12348 := [rewrite]: #12347
+#12360 := [monotonicity #12348 #12357]: #12359
+#12411 := [monotonicity #12360 #12408]: #12410
+#12490 := [monotonicity #12411 #12487]: #12489
+#12496 := [trans #12490 #12494]: #12495
+#12499 := [monotonicity #12360 #12496]: #12498
+#12353 := (iff #3225 #12352)
+#12350 := (iff #3224 #12349)
+#12351 := [rewrite]: #12350
+#12354 := [monotonicity #12348 #12351]: #12353
+#12502 := [monotonicity #12354 #12499]: #12501
+#12508 := [trans #12502 #12506]: #12507
+#12511 := [monotonicity #12354 #12508]: #12510
+#12373 := (iff #3233 #12372)
+#12370 := (iff #3232 #12369)
+#12371 := [monotonicity #12084]: #12370
+#12374 := [monotonicity #12022 #12371]: #12373
+#12514 := [monotonicity #12374 #12511]: #12513
+#12520 := [trans #12514 #12518]: #12519
+#12564 := [monotonicity #12520 #12561]: #12563
+#12367 := (iff #3229 #12364)
+#12361 := (and #12358 #12020)
+#12365 := (iff #12361 #12364)
+#12366 := [rewrite]: #12365
+#12362 := (iff #3229 #12361)
+#12363 := [monotonicity #12360 #12022]: #12362
+#12368 := [trans #12363 #12366]: #12367
+#12567 := [monotonicity #12368 #12564]: #12566
+#12573 := [trans #12567 #12571]: #12572
+#12576 := [monotonicity #12360 #12573]: #12575
+#12579 := [monotonicity #12354 #12576]: #12578
+#12584 := [trans #12579 #12582]: #12583
+#12587 := [monotonicity #12354 #12584]: #12586
+#12344 := (iff #3219 #12343)
+#12341 := (iff #3218 #12340)
+#12342 := [monotonicity #12084]: #12341
+#12345 := [monotonicity #12022 #12342]: #12344
+#12590 := [monotonicity #12345 #12587]: #12589
+#12596 := [trans #12590 #12594]: #12595
+#12622 := [monotonicity #12596 #12619]: #12621
+#12338 := (iff #3216 #12337)
+#12335 := (iff #3215 #12334)
+#12332 := (iff #3214 #12331)
+#12329 := (iff #3213 #12328)
+#12326 := (iff #3212 #12325)
+#12323 := (iff #3211 #12322)
+#12320 := (iff #3210 #12319)
+#12317 := (iff #3209 #12316)
+#12314 := (iff #3208 #12313)
+#12311 := (iff #3207 #12310)
+#12308 := (iff #3206 #12307)
+#12305 := (iff #3205 #12020)
+#12300 := (and true #12020)
+#12303 := (iff #12300 #12020)
+#12304 := [rewrite]: #12303
+#12301 := (iff #3205 #12300)
+#12298 := (iff #3204 true)
+#12296 := (iff #12293 true)
+#12297 := [rewrite]: #12296
+#12294 := (iff #3204 #12293)
+#12291 := (iff #3203 true)
+#12292 := [rewrite]: #12291
+#12289 := (iff #3202 true)
+#12290 := [rewrite]: #12289
+#12295 := [monotonicity #12290 #12292]: #12294
+#12299 := [trans #12295 #12297]: #12298
+#12302 := [monotonicity #12299 #12022]: #12301
+#12306 := [trans #12302 #12304]: #12305
+#12309 := [monotonicity #12306]: #12308
+#12312 := [monotonicity #12309]: #12311
+#12315 := [monotonicity #12312]: #12314
+#12318 := [monotonicity #12315]: #12317
+#12321 := [monotonicity #12318]: #12320
+#12076 := (iff #3115 #12073)
+#12070 := (and #12067 #11903)
+#12074 := (iff #12070 #12073)
+#12075 := [rewrite]: #12074
+#12071 := (iff #3115 #12070)
+#11904 := (iff #3005 #11903)
+#11905 := [rewrite]: #11904
+#12068 := (iff #3114 #12067)
+#12069 := [rewrite]: #12068
+#12072 := [monotonicity #12069 #11905]: #12071
+#12077 := [trans #12072 #12075]: #12076
+#12324 := [monotonicity #12077 #12321]: #12323
+#12268 := (iff #3190 #12267)
+#12265 := (iff #3189 #12264)
+#12262 := (iff #3188 #12261)
+#12263 := [rewrite]: #12262
+#12266 := [monotonicity #12263]: #12265
+#12269 := [monotonicity #12266]: #12268
+#12327 := [monotonicity #12269 #12324]: #12326
+#12287 := (iff #3194 #12267)
+#12270 := (and true #12267)
+#12273 := (iff #12270 #12267)
+#12274 := [rewrite]: #12273
+#12285 := (iff #3194 #12270)
+#12283 := (iff #3193 #12267)
+#12281 := (iff #3193 #12270)
+#12279 := (iff #3192 #12267)
+#12277 := (iff #3192 #12270)
+#12275 := (iff #3191 #12267)
+#12271 := (iff #3191 #12270)
+#12259 := (iff #3181 true)
+#12254 := (forall (vars (?x782 T5)) (:pat #3177) true)
+#12257 := (iff #12254 true)
+#12258 := [elim-unused]: #12257
+#12255 := (iff #3181 #12254)
+#12252 := (iff #3180 true)
+#12204 := (= uf_9 #3165)
+#12216 := (implies #12204 #12204)
+#12219 := (iff #12216 true)
+#12220 := [rewrite]: #12219
+#12250 := (iff #3180 #12216)
+#12248 := (iff #3179 #12204)
+#12209 := (and true #12204)
+#12212 := (iff #12209 #12204)
+#12213 := [rewrite]: #12212
+#12246 := (iff #3179 #12209)
+#12205 := (iff #3166 #12204)
+#12206 := [rewrite]: #12205
+#12244 := (iff #3178 true)
+#12245 := [rewrite]: #12244
+#12247 := [monotonicity #12245 #12206]: #12246
+#12249 := [trans #12247 #12213]: #12248
+#12251 := [monotonicity #12206 #12249]: #12250
+#12253 := [trans #12251 #12220]: #12252
+#12256 := [quant-intro #12253]: #12255
+#12260 := [trans #12256 #12258]: #12259
+#12272 := [monotonicity #12260 #12269]: #12271
+#12276 := [trans #12272 #12274]: #12275
+#12242 := (iff #3174 true)
+#12197 := (forall (vars (?x779 T5)) (:pat #3154) true)
+#12200 := (iff #12197 true)
+#12201 := [elim-unused]: #12200
+#12240 := (iff #3174 #12197)
+#12238 := (iff #3173 true)
+#12236 := (iff #3173 #12216)
+#12234 := (iff #3172 #12204)
+#12232 := (iff #3172 #12209)
+#12230 := (iff #3171 true)
+#12231 := [rewrite]: #12230
+#12233 := [monotonicity #12231 #12206]: #12232
+#12235 := [trans #12233 #12213]: #12234
+#12237 := [monotonicity #12206 #12235]: #12236
+#12239 := [trans #12237 #12220]: #12238
+#12241 := [quant-intro #12239]: #12240
+#12243 := [trans #12241 #12201]: #12242
+#12278 := [monotonicity #12243 #12276]: #12277
+#12280 := [trans #12278 #12274]: #12279
+#12228 := (iff #3170 true)
+#12223 := (forall (vars (?x780 T5)) (:pat #3164) true)
+#12226 := (iff #12223 true)
+#12227 := [elim-unused]: #12226
+#12224 := (iff #3170 #12223)
+#12221 := (iff #3169 true)
+#12217 := (iff #3169 #12216)
+#12214 := (iff #3168 #12204)
+#12210 := (iff #3168 #12209)
+#12207 := (iff #3167 true)
+#12208 := [rewrite]: #12207
+#12211 := [monotonicity #12208 #12206]: #12210
+#12215 := [trans #12211 #12213]: #12214
+#12218 := [monotonicity #12206 #12215]: #12217
+#12222 := [trans #12218 #12220]: #12221
+#12225 := [quant-intro #12222]: #12224
+#12229 := [trans #12225 #12227]: #12228
+#12282 := [monotonicity #12229 #12280]: #12281
+#12284 := [trans #12282 #12274]: #12283
+#12202 := (iff #3161 true)
+#12198 := (iff #3161 #12197)
+#12195 := (iff #3160 true)
+#12184 := (= uf_261 #3157)
+#12187 := (not #12184)
+#12190 := (implies #12187 #12187)
+#12193 := (iff #12190 true)
+#12194 := [rewrite]: #12193
+#12191 := (iff #3160 #12190)
+#12188 := (iff #3159 #12187)
+#12185 := (iff #3158 #12184)
+#12186 := [rewrite]: #12185
+#12189 := [monotonicity #12186]: #12188
+#12192 := [monotonicity #12189 #12189]: #12191
+#12196 := [trans #12192 #12194]: #12195
+#12199 := [quant-intro #12196]: #12198
+#12203 := [trans #12199 #12201]: #12202
+#12286 := [monotonicity #12203 #12284]: #12285
+#12288 := [trans #12286 #12274]: #12287
+#12330 := [monotonicity #12288 #12327]: #12329
+#12333 := [monotonicity #12022 #12330]: #12332
+#12336 := [monotonicity #12022 #12333]: #12335
+#12339 := [monotonicity #12022 #12336]: #12338
+#12625 := [monotonicity #12339 #12622]: #12624
+#12631 := [trans #12625 #12629]: #12630
+#12182 := (iff #3151 true)
+#12177 := (implies false #12174)
+#12180 := (iff #12177 true)
+#12181 := [rewrite]: #12180
+#12178 := (iff #3151 #12177)
+#12111 := (iff #3121 false)
+#12106 := (and #12020 false)
+#12109 := (iff #12106 false)
+#12110 := [rewrite]: #12109
+#12107 := (iff #3121 #12106)
+#12104 := (iff #3120 false)
+#12088 := (and #12020 #12073)
+#12093 := (and #12020 #12088)
+#12096 := (and #12020 #12093)
+#12099 := (and false #12096)
+#12102 := (iff #12099 false)
+#12103 := [rewrite]: #12102
+#12100 := (iff #3120 #12099)
+#12097 := (iff #3119 #12096)
+#12094 := (iff #3118 #12093)
+#12091 := (iff #3117 #12088)
+#12085 := (and #12073 #12020)
+#12089 := (iff #12085 #12088)
+#12090 := [rewrite]: #12089
+#12086 := (iff #3117 #12085)
+#12087 := [monotonicity #12077 #12084]: #12086
+#12092 := [trans #12087 #12090]: #12091
+#12095 := [monotonicity #12022 #12092]: #12094
+#12098 := [monotonicity #12022 #12095]: #12097
+#12065 := (iff #3112 false)
+#12066 := [rewrite]: #12065
+#12101 := [monotonicity #12066 #12098]: #12100
+#12105 := [trans #12101 #12103]: #12104
+#12108 := [monotonicity #12022 #12105]: #12107
+#12112 := [trans #12108 #12110]: #12111
+#12179 := [monotonicity #12112 #12176]: #12178
+#12183 := [trans #12179 #12181]: #12182
+#12634 := [monotonicity #12183 #12631]: #12633
+#12638 := [trans #12634 #12636]: #12637
+#12063 := (iff #3111 #12062)
+#12060 := (iff #3110 #12059)
+#12057 := (iff #3109 #12056)
+#12054 := (iff #3108 #12053)
+#12051 := (iff #3107 #12050)
+#12048 := (iff #3106 #12047)
+#12045 := (iff #3105 #12044)
+#12042 := (iff #3104 #12039)
+#12036 := (and #12033 #12020)
+#12040 := (iff #12036 #12039)
+#12041 := [rewrite]: #12040
+#12037 := (iff #3104 #12036)
+#12034 := (iff #3103 #12033)
+#12031 := (iff #3102 #12030)
+#12032 := [rewrite]: #12031
+#12035 := [monotonicity #12032]: #12034
+#12038 := [monotonicity #12035 #12022]: #12037
+#12043 := [trans #12038 #12041]: #12042
+#12028 := (iff #3098 #12027)
+#12025 := (iff #3097 #12024)
+#12026 := [rewrite]: #12025
+#12029 := [quant-intro #12026]: #12028
+#12046 := [monotonicity #12029 #12043]: #12045
+#12049 := [monotonicity #12046]: #12048
+#12052 := [monotonicity #12022 #12049]: #12051
+#12055 := [monotonicity #12052]: #12054
+#12058 := [monotonicity #12055]: #12057
+#12061 := [monotonicity #12058]: #12060
+#12018 := (iff #3078 #12017)
+#12015 := (iff #3077 #3055)
+#12016 := [rewrite]: #12015
+#12019 := [monotonicity #12016]: #12018
+#12064 := [monotonicity #12019 #12061]: #12063
+#12641 := [monotonicity #12064 #12638]: #12640
+#12647 := [trans #12641 #12645]: #12646
+#12650 := [monotonicity #12019 #12647]: #12649
+#12013 := (iff #3076 #12012)
+#12010 := (iff #3075 #12009)
+#12011 := [rewrite]: #12010
+#12014 := [quant-intro #12011]: #12013
+#12653 := [monotonicity #12014 #12650]: #12652
+#12659 := [trans #12653 #12657]: #12658
+#12662 := [monotonicity #12014 #12659]: #12661
+#12665 := [monotonicity #12662]: #12664
+#12671 := [trans #12665 #12669]: #12670
+#12674 := [monotonicity #12671]: #12673
+#12006 := (iff #3068 #12005)
+#12003 := (iff #3067 #12002)
+#12000 := (iff #3066 #11999)
+#11997 := (iff #3065 #11996)
+#11994 := (iff #3064 #11993)
+#11991 := (iff #3063 #11990)
+#11988 := (iff #3062 #11987)
+#11985 := (iff #3061 #3060)
+#11986 := [rewrite]: #11985
+#11989 := [monotonicity #11986]: #11988
+#11992 := [monotonicity #11989]: #11991
+#11995 := [monotonicity #11992]: #11994
+#11998 := [monotonicity #11995]: #11997
+#12001 := [monotonicity #11998]: #12000
+#12004 := [monotonicity #12001]: #12003
+#11983 := (iff #3052 #11982)
+#11980 := (iff #3051 #11979)
+#11981 := [rewrite]: #11980
+#11971 := (iff #3046 #11970)
+#11972 := [rewrite]: #11971
+#11984 := [monotonicity #11972 #11981]: #11983
+#12007 := [monotonicity #11984 #12004]: #12006
+#12677 := [monotonicity #12007 #12674]: #12676
+#12683 := [trans #12677 #12681]: #12682
+#12686 := [monotonicity #11984 #12683]: #12685
+#11977 := (iff #3049 #11976)
+#11974 := (iff #3048 #11973)
+#11975 := [rewrite]: #11974
+#11978 := [monotonicity #11972 #11975]: #11977
+#12689 := [monotonicity #11978 #12686]: #12688
+#12695 := [trans #12689 #12693]: #12694
+#12698 := [monotonicity #11978 #12695]: #12697
+#11968 := (iff #3043 #11967)
+#11969 := [rewrite]: #11968
+#12701 := [monotonicity #11969 #12698]: #12700
+#12707 := [trans #12701 #12705]: #12706
+#12710 := [monotonicity #11969 #12707]: #12709
+#11965 := (iff #3041 #11964)
+#11962 := (iff #3040 #11961)
+#11959 := (iff #3039 #11958)
+#11956 := (iff #3038 #11955)
+#11953 := (iff #3037 #11952)
+#11950 := (iff #3036 #11949)
+#11947 := (iff #3035 #11946)
+#11944 := (iff #3034 #11943)
+#11941 := (iff #3033 #11940)
+#11938 := (iff #3032 #11937)
+#11935 := (iff #3031 #11934)
+#11932 := (iff #3030 #11931)
+#11929 := (iff #3029 #11926)
+#11923 := (and #11920 #3028)
+#11927 := (iff #11923 #11926)
+#11928 := [rewrite]: #11927
+#11924 := (iff #3029 #11923)
+#11921 := (iff #3025 #11920)
+#11918 := (iff #3024 #11915)
+#11912 := (iff #11909 false)
+#11916 := (iff #11912 #11915)
+#11917 := [rewrite]: #11916
+#11913 := (iff #3024 #11912)
+#11910 := (iff #3023 #11909)
+#11911 := [rewrite]: #11910
+#11914 := [monotonicity #11911]: #11913
+#11919 := [trans #11914 #11917]: #11918
+#11922 := [quant-intro #11919]: #11921
+#11925 := [monotonicity #11922]: #11924
+#11930 := [trans #11925 #11928]: #11929
+#11933 := [monotonicity #11930]: #11932
+#11936 := [monotonicity #11933]: #11935
+#11939 := [monotonicity #11936]: #11938
+#11942 := [monotonicity #11939]: #11941
+#11907 := (iff #3006 #11906)
+#11901 := (iff #3003 #11900)
+#11902 := [rewrite]: #11901
+#11908 := [monotonicity #11902 #11905]: #11907
+#11945 := [monotonicity #11908 #11942]: #11944
+#11898 := (iff #3001 #11897)
+#11899 := [rewrite]: #11898
+#11948 := [monotonicity #11899 #11945]: #11947
+#11895 := (iff #2999 #11894)
+#11892 := (iff #2998 #11891)
+#11889 := (iff #2997 #11888)
+#11886 := (iff #2996 #11885)
+#11883 := (iff #2995 #11882)
+#11880 := (iff #2994 #11879)
+#11881 := [rewrite]: #11880
+#11877 := (iff #2992 #11876)
+#11874 := (iff #2991 #11873)
+#11875 := [rewrite]: #11874
+#11878 := [monotonicity #11875]: #11877
+#11884 := [monotonicity #11878 #11881]: #11883
+#11871 := (iff #2989 #11870)
+#11872 := [rewrite]: #11871
+#11887 := [monotonicity #11872 #11884]: #11886
+#11868 := (iff #2987 #11867)
+#11869 := [rewrite]: #11868
+#11890 := [monotonicity #11869 #11887]: #11889
+#11865 := (iff #2985 #11864)
+#11866 := [rewrite]: #11865
+#11893 := [monotonicity #11866 #11890]: #11892
+#11862 := (iff #2983 #11860)
+#11863 := [rewrite]: #11862
+#11896 := [monotonicity #11863 #11893]: #11895
+#11951 := [monotonicity #11896 #11948]: #11950
+#11954 := [monotonicity #11951]: #11953
+#11957 := [monotonicity #11954]: #11956
+#11960 := [monotonicity #11957]: #11959
+#11963 := [monotonicity #11960]: #11962
+#11966 := [monotonicity #11963]: #11965
+#12713 := [monotonicity #11966 #12710]: #12712
+#12719 := [trans #12713 #12717]: #12718
+#12722 := [monotonicity #12719]: #12721
+#13522 := [trans #12722 #13520]: #13521
+#11859 := [asserted]: #3325
+#13523 := [mp #11859 #13522]: #13518
+#13524 := [not-or-elim #13523]: #12823
+#13531 := [and-elim #13524]: #11867
+#11145 := (not #11138)
+#1259 := (uf_116 #15)
+#2669 := (uf_43 #233 #1259)
+#2670 := (= #15 #2669)
+#11146 := (or #2670 #11145)
+#11151 := (forall (vars (?x710 T5) (?x711 T3)) (:pat #2667) #11146)
+#17463 := (~ #11151 #11151)
+#17461 := (~ #11146 #11146)
+#17462 := [refl]: #17461
+#17464 := [nnf-pos #17462]: #17463
+#2671 := (implies #2668 #2670)
+#2672 := (forall (vars (?x710 T5) (?x711 T3)) (:pat #2667) #2671)
+#11152 := (iff #2672 #11151)
+#11149 := (iff #2671 #11146)
+#11142 := (implies #11138 #2670)
+#11147 := (iff #11142 #11146)
+#11148 := [rewrite]: #11147
+#11143 := (iff #2671 #11142)
+#11144 := [monotonicity #11141]: #11143
+#11150 := [trans #11144 #11148]: #11149
+#11153 := [quant-intro #11150]: #11152
+#11137 := [asserted]: #2672
+#11156 := [mp #11137 #11153]: #11151
+#17465 := [mp~ #11156 #17464]: #11151
+#25398 := (not #11867)
+#25411 := (not #11151)
+#25412 := (or #25411 #25398 #25406)
+#25407 := (or #25406 #25398)
+#25413 := (or #25411 #25407)
+#25420 := (iff #25413 #25412)
+#25408 := (or #25398 #25406)
+#25415 := (or #25411 #25408)
+#25418 := (iff #25415 #25412)
+#25419 := [rewrite]: #25418
+#25416 := (iff #25413 #25415)
+#25409 := (iff #25407 #25408)
+#25410 := [rewrite]: #25409
+#25417 := [monotonicity #25410]: #25416
+#25421 := [trans #25417 #25419]: #25420
+#25414 := [quant-inst]: #25413
+#25422 := [mp #25414 #25421]: #25412
+#27875 := [unit-resolution #25422 #17465 #13531]: #25406
+#27877 := [symm #27875]: #27876
+#26168 := [monotonicity #27877]: #26162
+#26170 := [trans #26168 #27921]: #26169
+#26172 := [monotonicity #26170]: #26171
+#26174 := [monotonicity #26172]: #26195
+#26177 := [symm #26174]: #26176
+#26161 := [monotonicity #26177]: #26179
+#13533 := [and-elim #13524]: #11876
+#26175 := [mp #13533 #26161]: #26178
+decl uf_196 :: (-> T4 T5 T5 T2)
+#25944 := (uf_196 uf_287 #25399 #25399)
+#25945 := (= uf_9 #25944)
+#25966 := (not #25945)
+#25946 := (uf_200 uf_287 #25399 #25399 uf_282)
+#25947 := (= uf_9 #25946)
+#25949 := (iff #25945 #25947)
+#2245 := (:var 0 T16)
+#13 := (:var 3 T4)
+#2256 := (uf_200 #13 #21 #15 #2245)
+#2257 := (pattern #2256)
+#2259 := (uf_196 #13 #21 #15)
+#10101 := (= uf_9 #2259)
+#10097 := (= uf_9 #2256)
+#10104 := (iff #10097 #10101)
+#10107 := (forall (vars (?x586 T4) (?x587 T5) (?x588 T5) (?x589 T16)) (:pat #2257) #10104)
+#17105 := (~ #10107 #10107)
+#17103 := (~ #10104 #10104)
+#17104 := [refl]: #17103
+#17106 := [nnf-pos #17104]: #17105
+#2260 := (= #2259 uf_9)
+#2258 := (= #2256 uf_9)
+#2261 := (iff #2258 #2260)
+#2262 := (forall (vars (?x586 T4) (?x587 T5) (?x588 T5) (?x589 T16)) (:pat #2257) #2261)
+#10108 := (iff #2262 #10107)
+#10105 := (iff #2261 #10104)
+#10102 := (iff #2260 #10101)
+#10103 := [rewrite]: #10102
+#10099 := (iff #2258 #10097)
+#10100 := [rewrite]: #10099
+#10106 := [monotonicity #10100 #10103]: #10105
+#10109 := [quant-intro #10106]: #10108
+#10096 := [asserted]: #2262
+#10112 := [mp #10096 #10109]: #10107
+#17107 := [mp~ #10112 #17106]: #10107
+#25955 := (not #10107)
+#25953 := (or #25955 #25949)
+#25948 := (iff #25947 #25945)
+#25956 := (or #25955 #25948)
+#25958 := (iff #25956 #25953)
+#25960 := (iff #25953 #25953)
+#25961 := [rewrite]: #25960
+#25950 := (iff #25948 #25949)
+#25951 := [rewrite]: #25950
+#25959 := [monotonicity #25951]: #25958
+#25978 := [trans #25959 #25961]: #25958
+#25957 := [quant-inst]: #25956
+#25964 := [mp #25957 #25978]: #25953
+#26163 := [unit-resolution #25964 #17107]: #25949
+#25963 := (not #25947)
+#26189 := (iff #12702 #25963)
+#26188 := (iff #11967 #25947)
+#26167 := (iff #25947 #11967)
+#26165 := (= #25946 #3042)
+#26166 := [monotonicity #27877 #27877]: #26165
+#26187 := [monotonicity #26166]: #26167
+#26186 := [symm #26187]: #26188
+#26190 := [monotonicity #26186]: #26189
+#26164 := [hypothesis]: #12702
+#26191 := [mp #26164 #26190]: #25963
+#25967 := (not #25949)
+#25969 := (or #25967 #25966 #25947)
+#25970 := [def-axiom]: #25969
+#26192 := [unit-resolution #25970 #26191 #26163]: #25966
+#26048 := (uf_24 uf_287 #25399)
+#26049 := (= uf_9 #26048)
+#26196 := (= #2988 #26048)
+#26193 := (= #26048 #2988)
+#26194 := [monotonicity #27877]: #26193
+#26198 := [symm #26194]: #26196
+#13532 := [and-elim #13524]: #11870
+#26199 := [trans #13532 #26198]: #26049
+#26051 := (uf_48 #25399 #25810)
+#26052 := (= uf_9 #26051)
+#26202 := (= #2986 #26051)
+#26197 := (= #26051 #2986)
+#26200 := [monotonicity #27877 #26170]: #26197
+#26203 := [symm #26200]: #26202
+#26204 := [trans #13531 #26203]: #26052
+#26053 := (not #26052)
+#26050 := (not #26049)
+#26225 := (or #25945 #26050 #26053 #26059)
+#25822 := (uf_25 uf_287 #25399)
+#26054 := (= uf_26 #25822)
+#26213 := (= #2984 #25822)
+#26185 := (= #25822 #2984)
+#26211 := [monotonicity #27877]: #26185
+#26212 := [symm #26211]: #26213
+#13530 := [and-elim #13524]: #11864
+#26214 := [trans #13530 #26212]: #26054
+#25848 := (uf_27 uf_287 #25399)
+#25849 := (= uf_9 #25848)
+#26218 := (= #2982 #25848)
+#26215 := (= #25848 #2982)
+#26216 := [monotonicity #27877]: #26215
+#26224 := [symm #26216]: #26218
+#13529 := [and-elim #13524]: #11860
+#26246 := [trans #13529 #26224]: #25849
+#25811 := (uf_23 #25810)
+#25818 := (= uf_9 #25811)
+#26250 := (= #2993 #25811)
+#26247 := (= #25811 #2993)
+#26248 := [monotonicity #26170]: #26247
+#26251 := [symm #26248]: #26250
+#13534 := [and-elim #13524]: #11879
+#26252 := [trans #13534 #26251]: #25818
+#13537 := [and-elim #13524]: #11903
#47 := (:var 1 T4)
+#2217 := (uf_196 #47 #23 #23)
+#2218 := (pattern #2217)
+#9982 := (= uf_9 #2217)
+#227 := (uf_55 #47)
+#3897 := (= uf_9 #227)
+#18563 := (not #3897)
+#146 := (uf_24 #47 #23)
+#3621 := (= uf_9 #146)
+#11066 := (not #3621)
+#26 := (uf_13 #23)
+#144 := (uf_48 #23 #26)
+#3618 := (= uf_9 #144)
+#18438 := (not #3618)
+#142 := (uf_25 #47 #23)
+#3615 := (= uf_26 #142)
+#18437 := (not #3615)
+#135 := (uf_27 #47 #23)
+#3600 := (= uf_9 #135)
+#10691 := (not #3600)
+#71 := (uf_23 #26)
+#3453 := (= uf_9 #71)
+#10746 := (not #3453)
+#27 := (uf_12 #26)
+#29 := (= #27 uf_14)
+#20996 := (or #29 #10746 #10691 #18437 #18438 #11066 #18563 #9982)
+#21001 := (forall (vars (?x572 T4) (?x573 T5)) (:pat #2218) #20996)
+#52 := (not #29)
+#9997 := (and #52 #3453 #3600 #3615 #3618 #3621 #3897)
+#10000 := (not #9997)
+#10006 := (or #9982 #10000)
+#10011 := (forall (vars (?x572 T4) (?x573 T5)) (:pat #2218) #10006)
+#21002 := (iff #10011 #21001)
+#20999 := (iff #10006 #20996)
+#20982 := (or #29 #10746 #10691 #18437 #18438 #11066 #18563)
+#20993 := (or #9982 #20982)
+#20997 := (iff #20993 #20996)
+#20998 := [rewrite]: #20997
+#20994 := (iff #10006 #20993)
+#20991 := (iff #10000 #20982)
+#20983 := (not #20982)
+#20986 := (not #20983)
+#20989 := (iff #20986 #20982)
+#20990 := [rewrite]: #20989
+#20987 := (iff #10000 #20986)
+#20984 := (iff #9997 #20983)
+#20985 := [rewrite]: #20984
+#20988 := [monotonicity #20985]: #20987
+#20992 := [trans #20988 #20990]: #20991
+#20995 := [monotonicity #20992]: #20994
+#21000 := [trans #20995 #20998]: #20999
+#21003 := [quant-intro #21000]: #21002
+#17073 := (~ #10011 #10011)
+#17071 := (~ #10006 #10006)
+#17072 := [refl]: #17071
+#17074 := [nnf-pos #17072]: #17073
+#2225 := (= #2217 uf_9)
+#72 := (= #71 uf_9)
+#2219 := (and #52 #72)
+#147 := (= #146 uf_9)
+#2220 := (and #147 #2219)
+#145 := (= #144 uf_9)
+#2221 := (and #145 #2220)
+#143 := (= #142 uf_26)
+#2222 := (and #143 #2221)
+#136 := (= #135 uf_9)
+#2223 := (and #136 #2222)
+#229 := (= #227 uf_9)
+#2224 := (and #229 #2223)
+#2226 := (implies #2224 #2225)
+#2227 := (forall (vars (?x572 T4) (?x573 T5)) (:pat #2218) #2226)
+#10014 := (iff #2227 #10011)
+#9963 := (and #52 #3453)
+#9967 := (and #3621 #9963)
+#9970 := (and #3618 #9967)
+#9973 := (and #3615 #9970)
+#9976 := (and #3600 #9973)
+#9979 := (and #3897 #9976)
+#9988 := (not #9979)
+#9989 := (or #9988 #9982)
+#9994 := (forall (vars (?x572 T4) (?x573 T5)) (:pat #2218) #9989)
+#10012 := (iff #9994 #10011)
+#10009 := (iff #9989 #10006)
+#10003 := (or #10000 #9982)
+#10007 := (iff #10003 #10006)
+#10008 := [rewrite]: #10007
+#10004 := (iff #9989 #10003)
+#10001 := (iff #9988 #10000)
+#9998 := (iff #9979 #9997)
+#9999 := [rewrite]: #9998
+#10002 := [monotonicity #9999]: #10001
+#10005 := [monotonicity #10002]: #10004
+#10010 := [trans #10005 #10008]: #10009
+#10013 := [quant-intro #10010]: #10012
+#9995 := (iff #2227 #9994)
+#9992 := (iff #2226 #9989)
+#9985 := (implies #9979 #9982)
+#9990 := (iff #9985 #9989)
+#9991 := [rewrite]: #9990
+#9986 := (iff #2226 #9985)
+#9983 := (iff #2225 #9982)
+#9984 := [rewrite]: #9983
+#9980 := (iff #2224 #9979)
+#9977 := (iff #2223 #9976)
+#9974 := (iff #2222 #9973)
+#9971 := (iff #2221 #9970)
+#9968 := (iff #2220 #9967)
+#9965 := (iff #2219 #9963)
+#3454 := (iff #72 #3453)
+#3455 := [rewrite]: #3454
+#9966 := [monotonicity #3455]: #9965
+#3622 := (iff #147 #3621)
+#3623 := [rewrite]: #3622
+#9969 := [monotonicity #3623 #9966]: #9968
+#3619 := (iff #145 #3618)
+#3620 := [rewrite]: #3619
+#9972 := [monotonicity #3620 #9969]: #9971
+#3616 := (iff #143 #3615)
+#3617 := [rewrite]: #3616
+#9975 := [monotonicity #3617 #9972]: #9974
+#3602 := (iff #136 #3600)
+#3603 := [rewrite]: #3602
+#9978 := [monotonicity #3603 #9975]: #9977
+#3899 := (iff #229 #3897)
+#3900 := [rewrite]: #3899
+#9981 := [monotonicity #3900 #9978]: #9980
+#9987 := [monotonicity #9981 #9984]: #9986
+#9993 := [trans #9987 #9991]: #9992
+#9996 := [quant-intro #9993]: #9995
+#10015 := [trans #9996 #10013]: #10014
+#9962 := [asserted]: #2227
+#10016 := [mp #9962 #10015]: #10011
+#17075 := [mp~ #10016 #17074]: #10011
+#21004 := [mp #17075 #21003]: #21001
+#26055 := (not #26054)
+#25875 := (not #25849)
+#25819 := (not #25818)
+#23948 := (not #11903)
+#26041 := (not #21001)
+#26043 := (or #26041 #23948 #25819 #25875 #25945 #26050 #26053 #26055 #26059)
+#26057 := (= #26056 uf_14)
+#26058 := (or #26057 #25819 #25875 #26055 #26053 #26050 #23948 #25945)
+#26044 := (or #26041 #26058)
+#26031 := (iff #26044 #26043)
+#26065 := (or #23948 #25819 #25875 #25945 #26050 #26053 #26055 #26059)
+#26046 := (or #26041 #26065)
+#26028 := (iff #26046 #26043)
+#26030 := [rewrite]: #26028
+#26047 := (iff #26044 #26046)
+#26068 := (iff #26058 #26065)
+#26062 := (or #26059 #25819 #25875 #26055 #26053 #26050 #23948 #25945)
+#26066 := (iff #26062 #26065)
+#26067 := [rewrite]: #26066
+#26063 := (iff #26058 #26062)
+#26060 := (iff #26057 #26059)
+#26061 := [rewrite]: #26060
+#26064 := [monotonicity #26061]: #26063
+#26069 := [trans #26064 #26067]: #26068
+#26042 := [monotonicity #26069]: #26047
+#26032 := [trans #26042 #26030]: #26031
+#26045 := [quant-inst]: #26044
+#26033 := [mp #26045 #26032]: #26043
+#26253 := [unit-resolution #26033 #21004 #13537 #26252 #26246 #26214]: #26225
+#26254 := [unit-resolution #26253 #26204 #26199 #26192 #26175]: false
+#26255 := [lemma #26254]: #11967
+#22707 := (or #12702 #22704)
+#21694 := (forall (vars (?x778 int)) #21683)
+#21701 := (not #21694)
+#21679 := (forall (vars (?x776 int)) #21674)
+#21700 := (not #21679)
+#21702 := (or #21700 #21701)
+#21703 := (not #21702)
+#21732 := (or #21703 #21729)
+#21742 := (not #21732)
+#21743 := (or #21738 #21739 #21740 #21741 #21602 #21575 #12968 #21742)
+#21744 := (not #21743)
+#21497 := (forall (vars (?x785 int)) #21492)
+#21503 := (not #21497)
+#21504 := (or #21468 #21503)
+#21505 := (not #21504)
+#21534 := (or #21505 #21531)
+#21540 := (not #21534)
+#21541 := (or #13093 #21540)
+#21542 := (not #21541)
+#21547 := (or #13093 #21542)
+#21557 := (not #21547)
+#21558 := (or #21553 #21554 #17742 #21555 #21556 #17745 #21557)
+#21559 := (not #21558)
+#21564 := (or #17742 #17745 #21559)
+#21576 := (not #21564)
+#21615 := (or #21613 #21614 #21602 #21575 #12996 #21554 #21576)
+#21616 := (not #21615)
+#21577 := (or #21570 #21571 #17721 #17730 #21572 #21573 #21574 #21575 #21554 #21576)
+#21578 := (not #21577)
+#21583 := (or #17721 #17730 #21578)
+#21589 := (not #21583)
+#21590 := (or #17721 #17724 #21589)
+#21591 := (not #21590)
+#21596 := (or #17721 #17724 #21591)
+#21603 := (not #21596)
+#21604 := (or #21602 #21575 #12997 #21603)
+#21605 := (not #21604)
+#21621 := (or #21605 #21616)
+#21627 := (not #21621)
+#21628 := (or #17721 #17730 #21602 #21575 #21627)
+#21629 := (not #21628)
+#21634 := (or #17721 #17730 #21629)
+#21640 := (not #21634)
+#21641 := (or #17721 #17724 #21640)
+#21642 := (not #21641)
+#21647 := (or #17721 #17724 #21642)
+#21653 := (not #21647)
+#21654 := (or #21602 #21575 #12967 #21653)
+#21655 := (not #21654)
+#21749 := (or #21655 #21744)
+#21771 := (not #21749)
+#21463 := (forall (vars (?x775 int)) #21458)
+#21767 := (not #21463)
+#21772 := (or #17676 #21755 #21756 #21757 #21758 #21759 #21760 #21761 #21762 #21763 #12740 #21602 #21575 #13358 #21764 #21765 #21766 #21768 #21769 #21770 #21767 #21771)
+#21773 := (not #21772)
+#21778 := (or #17676 #12740 #21773)
+#21785 := (not #21778)
+#21441 := (forall (vars (?x773 int)) #21436)
+#21784 := (not #21441)
+#21786 := (or #21784 #21785)
+#21787 := (not #21786)
+#21792 := (or #21419 #21787)
+#21798 := (not #21792)
+#21799 := (or #12866 #21798)
+#21800 := (not #21799)
+#21805 := (or #12866 #21800)
+#21814 := (not #21805)
+#21815 := (or #17676 #21811 #21812 #21813 #17631 #17640 #21814)
+#21816 := (not #21815)
+#21821 := (or #17631 #17640 #21816)
+#21827 := (not #21821)
+#21828 := (or #17631 #17634 #21827)
+#21829 := (not #21828)
+#21834 := (or #17631 #17634 #21829)
+#21840 := (not #21834)
+#21841 := (or #12702 #21840)
+#21842 := (not #21841)
+#21847 := (or #12702 #21842)
+#22708 := (iff #21847 #22707)
+#22705 := (iff #21842 #22704)
+#22702 := (iff #21841 #22701)
+#22699 := (iff #21840 #22698)
+#22696 := (iff #21834 #22695)
+#22693 := (iff #21829 #22692)
+#22690 := (iff #21828 #22689)
+#22687 := (iff #21827 #22686)
+#22684 := (iff #21821 #22683)
+#22681 := (iff #21816 #22680)
+#22678 := (iff #21815 #22677)
+#22675 := (iff #21814 #22674)
+#22672 := (iff #21805 #22671)
+#22669 := (iff #21800 #22668)
+#22666 := (iff #21799 #22665)
+#22663 := (iff #21798 #22662)
+#22660 := (iff #21792 #22659)
+#22657 := (iff #21787 #22656)
+#22654 := (iff #21786 #22653)
+#22651 := (iff #21785 #22650)
+#22648 := (iff #21778 #22647)
+#22645 := (iff #21773 #22644)
+#22642 := (iff #21772 #22641)
+#22639 := (iff #21771 #22638)
+#22636 := (iff #21749 #22635)
+#22633 := (iff #21744 #22632)
+#22630 := (iff #21743 #22629)
+#22627 := (iff #21742 #22626)
+#22624 := (iff #21732 #22623)
+#22621 := (iff #21703 #22620)
+#22618 := (iff #21702 #22617)
+#22615 := (iff #21701 #22614)
+#22612 := (iff #21694 #22609)
+#22610 := (iff #21683 #21683)
+#22611 := [refl]: #22610
+#22613 := [quant-intro #22611]: #22612
+#22616 := [monotonicity #22613]: #22615
+#22607 := (iff #21700 #22606)
+#22604 := (iff #21679 #22601)
+#22602 := (iff #21674 #21674)
+#22603 := [refl]: #22602
+#22605 := [quant-intro #22603]: #22604
+#22608 := [monotonicity #22605]: #22607
+#22619 := [monotonicity #22608 #22616]: #22618
+#22622 := [monotonicity #22619]: #22621
+#22625 := [monotonicity #22622]: #22624
+#22628 := [monotonicity #22625]: #22627
+#22631 := [monotonicity #22628]: #22630
+#22634 := [monotonicity #22631]: #22633
+#22599 := (iff #21655 #22598)
+#22596 := (iff #21654 #22595)
+#22593 := (iff #21653 #22592)
+#22590 := (iff #21647 #22589)
+#22587 := (iff #21642 #22586)
+#22584 := (iff #21641 #22583)
+#22581 := (iff #21640 #22580)
+#22578 := (iff #21634 #22577)
+#22575 := (iff #21629 #22574)
+#22572 := (iff #21628 #22571)
+#22569 := (iff #21627 #22568)
+#22566 := (iff #21621 #22565)
+#22563 := (iff #21616 #22562)
+#22560 := (iff #21615 #22559)
+#22527 := (iff #21576 #22526)
+#22524 := (iff #21564 #22523)
+#22521 := (iff #21559 #22520)
+#22518 := (iff #21558 #22517)
+#22515 := (iff #21557 #22514)
+#22512 := (iff #21547 #22511)
+#22509 := (iff #21542 #22508)
+#22506 := (iff #21541 #22505)
+#22503 := (iff #21540 #22502)
+#22500 := (iff #21534 #22499)
+#22497 := (iff #21505 #22496)
+#22494 := (iff #21504 #22493)
+#22491 := (iff #21503 #22490)
+#22488 := (iff #21497 #22485)
+#22486 := (iff #21492 #21492)
+#22487 := [refl]: #22486
+#22489 := [quant-intro #22487]: #22488
+#22492 := [monotonicity #22489]: #22491
+#22495 := [monotonicity #22492]: #22494
+#22498 := [monotonicity #22495]: #22497
+#22501 := [monotonicity #22498]: #22500
+#22504 := [monotonicity #22501]: #22503
+#22507 := [monotonicity #22504]: #22506
+#22510 := [monotonicity #22507]: #22509
+#22513 := [monotonicity #22510]: #22512
+#22516 := [monotonicity #22513]: #22515
+#22519 := [monotonicity #22516]: #22518
+#22522 := [monotonicity #22519]: #22521
+#22525 := [monotonicity #22522]: #22524
+#22528 := [monotonicity #22525]: #22527
+#22561 := [monotonicity #22528]: #22560
+#22564 := [monotonicity #22561]: #22563
+#22557 := (iff #21605 #22556)
+#22554 := (iff #21604 #22553)
+#22551 := (iff #21603 #22550)
+#22548 := (iff #21596 #22547)
+#22545 := (iff #21591 #22544)
+#22542 := (iff #21590 #22541)
+#22539 := (iff #21589 #22538)
+#22536 := (iff #21583 #22535)
+#22533 := (iff #21578 #22532)
+#22530 := (iff #21577 #22529)
+#22531 := [monotonicity #22528]: #22530
+#22534 := [monotonicity #22531]: #22533
+#22537 := [monotonicity #22534]: #22536
+#22540 := [monotonicity #22537]: #22539
+#22543 := [monotonicity #22540]: #22542
+#22546 := [monotonicity #22543]: #22545
+#22549 := [monotonicity #22546]: #22548
+#22552 := [monotonicity #22549]: #22551
+#22555 := [monotonicity #22552]: #22554
+#22558 := [monotonicity #22555]: #22557
+#22567 := [monotonicity #22558 #22564]: #22566
+#22570 := [monotonicity #22567]: #22569
+#22573 := [monotonicity #22570]: #22572
+#22576 := [monotonicity #22573]: #22575
+#22579 := [monotonicity #22576]: #22578
+#22582 := [monotonicity #22579]: #22581
+#22585 := [monotonicity #22582]: #22584
+#22588 := [monotonicity #22585]: #22587
+#22591 := [monotonicity #22588]: #22590
+#22594 := [monotonicity #22591]: #22593
+#22597 := [monotonicity #22594]: #22596
+#22600 := [monotonicity #22597]: #22599
+#22637 := [monotonicity #22600 #22634]: #22636
+#22640 := [monotonicity #22637]: #22639
+#22483 := (iff #21767 #22482)
+#22480 := (iff #21463 #22477)
+#22478 := (iff #21458 #21458)
+#22479 := [refl]: #22478
+#22481 := [quant-intro #22479]: #22480
+#22484 := [monotonicity #22481]: #22483
+#22643 := [monotonicity #22484 #22640]: #22642
+#22646 := [monotonicity #22643]: #22645
+#22649 := [monotonicity #22646]: #22648
+#22652 := [monotonicity #22649]: #22651
+#22475 := (iff #21784 #22474)
+#22472 := (iff #21441 #22469)
+#22470 := (iff #21436 #21436)
+#22471 := [refl]: #22470
+#22473 := [quant-intro #22471]: #22472
+#22476 := [monotonicity #22473]: #22475
+#22655 := [monotonicity #22476 #22652]: #22654
+#22658 := [monotonicity #22655]: #22657
+#22661 := [monotonicity #22658]: #22660
+#22664 := [monotonicity #22661]: #22663
+#22667 := [monotonicity #22664]: #22666
+#22670 := [monotonicity #22667]: #22669
+#22673 := [monotonicity #22670]: #22672
+#22676 := [monotonicity #22673]: #22675
+#22679 := [monotonicity #22676]: #22678
+#22682 := [monotonicity #22679]: #22681
+#22685 := [monotonicity #22682]: #22684
+#22688 := [monotonicity #22685]: #22687
+#22691 := [monotonicity #22688]: #22690
+#22694 := [monotonicity #22691]: #22693
+#22697 := [monotonicity #22694]: #22696
+#22700 := [monotonicity #22697]: #22699
+#22703 := [monotonicity #22700]: #22702
+#22706 := [monotonicity #22703]: #22705
+#22709 := [monotonicity #22706]: #22708
+#18175 := (not #18170)
+#18181 := (and #17866 #17870 #18175)
+#18186 := (not #18181)
+#18189 := (or #18157 #18186)
+#18192 := (not #18189)
+#15260 := (and #3142 #4041 #13272 #14378)
+#17890 := (not #15260)
+#17893 := (forall (vars (?x778 int)) #17890)
+#15235 := (and #4041 #13272 #14378)
+#15240 := (not #15235)
+#15246 := (or #13288 #15240)
+#15251 := (forall (vars (?x776 int)) #15246)
+#17897 := (and #15251 #17893)
+#18198 := (or #17897 #18192)
+#18206 := (and #12113 #12116 #12119 #12122 #12910 #12913 #12967 #18198)
+#18027 := (not #18022)
+#18033 := (and #17757 #17761 #18027)
+#18038 := (not #18033)
+#18041 := (or #18009 #18038)
+#18044 := (not #18041)
+#15146 := (and #4041 #13098 #14378)
+#15151 := (not #15146)
+#15157 := (or #13114 #15151)
+#15162 := (forall (vars (?x785 int)) #15157)
+#17774 := (not #13135)
+#17784 := (and #17774 #15162)
+#18050 := (or #17784 #18044)
+#18055 := (and #13089 #18050)
+#18058 := (or #13093 #18055)
+#18064 := (and #3258 #13017 #13051 #13059 #13061 #15127 #18058)
+#18069 := (or #17742 #17745 #18064)
+#18107 := (and #12521 #12524 #12910 #12913 #12997 #13017 #18069)
+#18075 := (and #3236 #3237 #12346 #12355 #12375 #12380 #12383 #12913 #13017 #18069)
+#18080 := (or #17721 #17730 #18075)
+#18086 := (and #12346 #12349 #18080)
+#18091 := (or #17721 #17724 #18086)
+#18097 := (and #12910 #12913 #12996 #18091)
+#18112 := (or #18097 #18107)
+#18118 := (and #12346 #12355 #12910 #12913 #18112)
+#18123 := (or #17721 #17730 #18118)
+#18129 := (and #12346 #12349 #18123)
+#18134 := (or #17721 #17724 #18129)
+#18140 := (and #12910 #12913 #12968 #18134)
+#18211 := (or #18140 #18206)
+#15060 := (and #4041 #13364 #14378)
+#15065 := (not #15060)
+#15071 := (or #13380 #15065)
+#15076 := (forall (vars (?x775 int)) #15071)
+#18217 := (and #3055 #3195 #3196 #3197 #3198 #3199 #3200 #12030 #12067 #12261 #12741 #12910 #12913 #13359 #13389 #13392 #13402 #15076 #15085 #15096 #15107 #18211)
+#18222 := (or #17676 #12740 #18217)
+#15035 := (and #4041 #12870 #14378)
+#15040 := (not #15035)
+#15046 := (or #12885 #15040)
+#15051 := (forall (vars (?x773 int)) #15046)
+#18225 := (and #15051 #18222)
+#17654 := (not #17653)
+#17971 := (and #17652 #17654 #17655)
+#17974 := (not #17971)
+#17977 := (or #17662 #17974)
+#17980 := (not #17977)
+#18228 := (or #17980 #18225)
+#18231 := (and #12863 #18228)
+#18234 := (or #12866 #18231)
+#18240 := (and #3055 #3056 #3057 #3058 #11970 #11979 #18234)
+#18245 := (or #17631 #17640 #18240)
+#18251 := (and #11970 #11973 #18245)
+#18256 := (or #17631 #17634 #18251)
+#18259 := (and #11967 #18256)
+#18262 := (or #12702 #18259)
+#21848 := (iff #18262 #21847)
+#21845 := (iff #18259 #21842)
+#21837 := (and #11967 #21834)
+#21843 := (iff #21837 #21842)
+#21844 := [rewrite]: #21843
+#21838 := (iff #18259 #21837)
+#21835 := (iff #18256 #21834)
+#21832 := (iff #18251 #21829)
+#21824 := (and #11970 #11973 #21821)
+#21830 := (iff #21824 #21829)
+#21831 := [rewrite]: #21830
+#21825 := (iff #18251 #21824)
+#21822 := (iff #18245 #21821)
+#21819 := (iff #18240 #21816)
+#21808 := (and #3055 #3056 #3057 #3058 #11970 #11979 #21805)
+#21817 := (iff #21808 #21816)
+#21818 := [rewrite]: #21817
+#21809 := (iff #18240 #21808)
+#21806 := (iff #18234 #21805)
+#21803 := (iff #18231 #21800)
+#21795 := (and #12863 #21792)
+#21801 := (iff #21795 #21800)
+#21802 := [rewrite]: #21801
+#21796 := (iff #18231 #21795)
+#21793 := (iff #18228 #21792)
+#21790 := (iff #18225 #21787)
+#21781 := (and #21441 #21778)
+#21788 := (iff #21781 #21787)
+#21789 := [rewrite]: #21788
+#21782 := (iff #18225 #21781)
+#21779 := (iff #18222 #21778)
+#21776 := (iff #18217 #21773)
+#21752 := (and #3055 #3195 #3196 #3197 #3198 #3199 #3200 #12030 #12067 #12261 #12741 #12910 #12913 #13359 #13389 #13392 #13402 #21463 #15085 #15096 #15107 #21749)
+#21774 := (iff #21752 #21773)
+#21775 := [rewrite]: #21774
+#21753 := (iff #18217 #21752)
+#21750 := (iff #18211 #21749)
+#21747 := (iff #18206 #21744)
+#21735 := (and #12113 #12116 #12119 #12122 #12910 #12913 #12967 #21732)
+#21745 := (iff #21735 #21744)
+#21746 := [rewrite]: #21745
+#21736 := (iff #18206 #21735)
+#21733 := (iff #18198 #21732)
+#21730 := (iff #18192 #21729)
+#21727 := (iff #18189 #21724)
+#21710 := (or #21708 #21709 #18170)
+#21721 := (or #18157 #21710)
+#21725 := (iff #21721 #21724)
+#21726 := [rewrite]: #21725
+#21722 := (iff #18189 #21721)
+#21719 := (iff #18186 #21710)
+#21711 := (not #21710)
+#21714 := (not #21711)
+#21717 := (iff #21714 #21710)
+#21718 := [rewrite]: #21717
+#21715 := (iff #18186 #21714)
+#21712 := (iff #18181 #21711)
+#21713 := [rewrite]: #21712
+#21716 := [monotonicity #21713]: #21715
+#21720 := [trans #21716 #21718]: #21719
+#21723 := [monotonicity #21720]: #21722
+#21728 := [trans #21723 #21726]: #21727
+#21731 := [monotonicity #21728]: #21730
+#21706 := (iff #17897 #21703)
+#21697 := (and #21679 #21694)
+#21704 := (iff #21697 #21703)
+#21705 := [rewrite]: #21704
+#21698 := (iff #17897 #21697)
+#21695 := (iff #17893 #21694)
+#21692 := (iff #17890 #21683)
+#21684 := (not #21683)
+#21687 := (not #21684)
+#21690 := (iff #21687 #21683)
+#21691 := [rewrite]: #21690
+#21688 := (iff #17890 #21687)
+#21685 := (iff #15260 #21684)
+#21686 := [rewrite]: #21685
+#21689 := [monotonicity #21686]: #21688
+#21693 := [trans #21689 #21691]: #21692
+#21696 := [quant-intro #21693]: #21695
+#21680 := (iff #15251 #21679)
+#21677 := (iff #15246 #21674)
+#21660 := (or #4963 #13270 #18695)
+#21671 := (or #13288 #21660)
+#21675 := (iff #21671 #21674)
+#21676 := [rewrite]: #21675
+#21672 := (iff #15246 #21671)
+#21669 := (iff #15240 #21660)
+#21661 := (not #21660)
+#21664 := (not #21661)
+#21667 := (iff #21664 #21660)
+#21668 := [rewrite]: #21667
+#21665 := (iff #15240 #21664)
+#21662 := (iff #15235 #21661)
+#21663 := [rewrite]: #21662
+#21666 := [monotonicity #21663]: #21665
+#21670 := [trans #21666 #21668]: #21669
+#21673 := [monotonicity #21670]: #21672
+#21678 := [trans #21673 #21676]: #21677
+#21681 := [quant-intro #21678]: #21680
+#21699 := [monotonicity #21681 #21696]: #21698
+#21707 := [trans #21699 #21705]: #21706
+#21734 := [monotonicity #21707 #21731]: #21733
+#21737 := [monotonicity #21734]: #21736
+#21748 := [trans #21737 #21746]: #21747
+#21658 := (iff #18140 #21655)
+#21650 := (and #12910 #12913 #12968 #21647)
+#21656 := (iff #21650 #21655)
+#21657 := [rewrite]: #21656
+#21651 := (iff #18140 #21650)
+#21648 := (iff #18134 #21647)
+#21645 := (iff #18129 #21642)
+#21637 := (and #12346 #12349 #21634)
+#21643 := (iff #21637 #21642)
+#21644 := [rewrite]: #21643
+#21638 := (iff #18129 #21637)
+#21635 := (iff #18123 #21634)
+#21632 := (iff #18118 #21629)
+#21624 := (and #12346 #12355 #12910 #12913 #21621)
+#21630 := (iff #21624 #21629)
+#21631 := [rewrite]: #21630
+#21625 := (iff #18118 #21624)
+#21622 := (iff #18112 #21621)
+#21619 := (iff #18107 #21616)
+#21610 := (and #12521 #12524 #12910 #12913 #12997 #13017 #21564)
+#21617 := (iff #21610 #21616)
+#21618 := [rewrite]: #21617
+#21611 := (iff #18107 #21610)
+#21565 := (iff #18069 #21564)
+#21562 := (iff #18064 #21559)
+#21550 := (and #3258 #13017 #13051 #13059 #13061 #15127 #21547)
+#21560 := (iff #21550 #21559)
+#21561 := [rewrite]: #21560
+#21551 := (iff #18064 #21550)
+#21548 := (iff #18058 #21547)
+#21545 := (iff #18055 #21542)
+#21537 := (and #13089 #21534)
+#21543 := (iff #21537 #21542)
+#21544 := [rewrite]: #21543
+#21538 := (iff #18055 #21537)
+#21535 := (iff #18050 #21534)
+#21532 := (iff #18044 #21531)
+#21529 := (iff #18041 #21526)
+#21512 := (or #21510 #21511 #18022)
+#21523 := (or #18009 #21512)
+#21527 := (iff #21523 #21526)
+#21528 := [rewrite]: #21527
+#21524 := (iff #18041 #21523)
+#21521 := (iff #18038 #21512)
+#21513 := (not #21512)
+#21516 := (not #21513)
+#21519 := (iff #21516 #21512)
+#21520 := [rewrite]: #21519
+#21517 := (iff #18038 #21516)
+#21514 := (iff #18033 #21513)
+#21515 := [rewrite]: #21514
+#21518 := [monotonicity #21515]: #21517
+#21522 := [trans #21518 #21520]: #21521
+#21525 := [monotonicity #21522]: #21524
+#21530 := [trans #21525 #21528]: #21529
+#21533 := [monotonicity #21530]: #21532
+#21508 := (iff #17784 #21505)
+#21500 := (and #21467 #21497)
+#21506 := (iff #21500 #21505)
+#21507 := [rewrite]: #21506
+#21501 := (iff #17784 #21500)
+#21498 := (iff #15162 #21497)
+#21495 := (iff #15157 #21492)
+#21478 := (or #4963 #13096 #18695)
+#21489 := (or #13114 #21478)
+#21493 := (iff #21489 #21492)
+#21494 := [rewrite]: #21493
+#21490 := (iff #15157 #21489)
+#21487 := (iff #15151 #21478)
+#21479 := (not #21478)
+#21482 := (not #21479)
+#21485 := (iff #21482 #21478)
+#21486 := [rewrite]: #21485
+#21483 := (iff #15151 #21482)
+#21480 := (iff #15146 #21479)
+#21481 := [rewrite]: #21480
+#21484 := [monotonicity #21481]: #21483
+#21488 := [trans #21484 #21486]: #21487
+#21491 := [monotonicity #21488]: #21490
+#21496 := [trans #21491 #21494]: #21495
+#21499 := [quant-intro #21496]: #21498
+#21476 := (iff #17774 #21467)
+#21471 := (not #21468)
+#21474 := (iff #21471 #21467)
+#21475 := [rewrite]: #21474
+#21472 := (iff #17774 #21471)
+#21469 := (iff #13135 #21468)
+#21470 := [rewrite]: #21469
+#21473 := [monotonicity #21470]: #21472
+#21477 := [trans #21473 #21475]: #21476
+#21502 := [monotonicity #21477 #21499]: #21501
+#21509 := [trans #21502 #21507]: #21508
+#21536 := [monotonicity #21509 #21533]: #21535
+#21539 := [monotonicity #21536]: #21538
+#21546 := [trans #21539 #21544]: #21545
+#21549 := [monotonicity #21546]: #21548
+#21552 := [monotonicity #21549]: #21551
+#21563 := [trans #21552 #21561]: #21562
+#21566 := [monotonicity #21563]: #21565
+#21612 := [monotonicity #21566]: #21611
+#21620 := [trans #21612 #21618]: #21619
+#21608 := (iff #18097 #21605)
+#21599 := (and #12910 #12913 #12996 #21596)
+#21606 := (iff #21599 #21605)
+#21607 := [rewrite]: #21606
+#21600 := (iff #18097 #21599)
+#21597 := (iff #18091 #21596)
+#21594 := (iff #18086 #21591)
+#21586 := (and #12346 #12349 #21583)
+#21592 := (iff #21586 #21591)
+#21593 := [rewrite]: #21592
+#21587 := (iff #18086 #21586)
+#21584 := (iff #18080 #21583)
+#21581 := (iff #18075 #21578)
+#21567 := (and #3236 #3237 #12346 #12355 #12375 #12380 #12383 #12913 #13017 #21564)
+#21579 := (iff #21567 #21578)
+#21580 := [rewrite]: #21579
+#21568 := (iff #18075 #21567)
+#21569 := [monotonicity #21566]: #21568
+#21582 := [trans #21569 #21580]: #21581
+#21585 := [monotonicity #21582]: #21584
+#21588 := [monotonicity #21585]: #21587
+#21595 := [trans #21588 #21593]: #21594
+#21598 := [monotonicity #21595]: #21597
+#21601 := [monotonicity #21598]: #21600
+#21609 := [trans #21601 #21607]: #21608
+#21623 := [monotonicity #21609 #21620]: #21622
+#21626 := [monotonicity #21623]: #21625
+#21633 := [trans #21626 #21631]: #21632
+#21636 := [monotonicity #21633]: #21635
+#21639 := [monotonicity #21636]: #21638
+#21646 := [trans #21639 #21644]: #21645
+#21649 := [monotonicity #21646]: #21648
+#21652 := [monotonicity #21649]: #21651
+#21659 := [trans #21652 #21657]: #21658
+#21751 := [monotonicity #21659 #21748]: #21750
+#21464 := (iff #15076 #21463)
+#21461 := (iff #15071 #21458)
+#21444 := (or #4963 #13362 #18695)
+#21455 := (or #13380 #21444)
+#21459 := (iff #21455 #21458)
+#21460 := [rewrite]: #21459
+#21456 := (iff #15071 #21455)
+#21453 := (iff #15065 #21444)
+#21445 := (not #21444)
+#21448 := (not #21445)
+#21451 := (iff #21448 #21444)
+#21452 := [rewrite]: #21451
+#21449 := (iff #15065 #21448)
+#21446 := (iff #15060 #21445)
+#21447 := [rewrite]: #21446
+#21450 := [monotonicity #21447]: #21449
+#21454 := [trans #21450 #21452]: #21453
+#21457 := [monotonicity #21454]: #21456
+#21462 := [trans #21457 #21460]: #21461
+#21465 := [quant-intro #21462]: #21464
+#21754 := [monotonicity #21465 #21751]: #21753
+#21777 := [trans #21754 #21775]: #21776
+#21780 := [monotonicity #21777]: #21779
+#21442 := (iff #15051 #21441)
+#21439 := (iff #15046 #21436)
+#21422 := (or #4963 #12869 #18695)
+#21433 := (or #12885 #21422)
+#21437 := (iff #21433 #21436)
+#21438 := [rewrite]: #21437
+#21434 := (iff #15046 #21433)
+#21431 := (iff #15040 #21422)
+#21423 := (not #21422)
+#21426 := (not #21423)
+#21429 := (iff #21426 #21422)
+#21430 := [rewrite]: #21429
+#21427 := (iff #15040 #21426)
+#21424 := (iff #15035 #21423)
+#21425 := [rewrite]: #21424
+#21428 := [monotonicity #21425]: #21427
+#21432 := [trans #21428 #21430]: #21431
+#21435 := [monotonicity #21432]: #21434
+#21440 := [trans #21435 #21438]: #21439
+#21443 := [quant-intro #21440]: #21442
+#21783 := [monotonicity #21443 #21780]: #21782
+#21791 := [trans #21783 #21789]: #21790
+#21420 := (iff #17980 #21419)
+#21417 := (iff #17977 #21414)
+#21400 := (or #21398 #17653 #21399)
+#21411 := (or #17662 #21400)
+#21415 := (iff #21411 #21414)
+#21416 := [rewrite]: #21415
+#21412 := (iff #17977 #21411)
+#21409 := (iff #17974 #21400)
+#21401 := (not #21400)
+#21404 := (not #21401)
+#21407 := (iff #21404 #21400)
+#21408 := [rewrite]: #21407
+#21405 := (iff #17974 #21404)
+#21402 := (iff #17971 #21401)
+#21403 := [rewrite]: #21402
+#21406 := [monotonicity #21403]: #21405
+#21410 := [trans #21406 #21408]: #21409
+#21413 := [monotonicity #21410]: #21412
+#21418 := [trans #21413 #21416]: #21417
+#21421 := [monotonicity #21418]: #21420
+#21794 := [monotonicity #21421 #21791]: #21793
+#21797 := [monotonicity #21794]: #21796
+#21804 := [trans #21797 #21802]: #21803
+#21807 := [monotonicity #21804]: #21806
+#21810 := [monotonicity #21807]: #21809
+#21820 := [trans #21810 #21818]: #21819
+#21823 := [monotonicity #21820]: #21822
+#21826 := [monotonicity #21823]: #21825
+#21833 := [trans #21826 #21831]: #21832
+#21836 := [monotonicity #21833]: #21835
+#21839 := [monotonicity #21836]: #21838
+#21846 := [trans #21839 #21844]: #21845
+#21849 := [monotonicity #21846]: #21848
+#17867 := (+ ?x776!15 #12727)
+#17868 := (>= #17867 0::int)
+#17869 := (not #17868)
+#17871 := (and #17870 #17869 #17866)
+#17872 := (not #17871)
+#17875 := (+ #17874 #13286)
+#17876 := (<= #17875 0::int)
+#17877 := (or #17876 #17872)
+#17878 := (not #17877)
+#17901 := (or #17878 #17897)
+#14493 := (and #12113 #12116 #12119 #12122 #12910 #12913)
+#14498 := (not #14493)
+#17862 := (not #14498)
+#17859 := (not #13347)
+#17905 := (and #17859 #17862 #17901)
+#17758 := (+ ?x785!14 #13062)
+#17759 := (>= #17758 0::int)
+#17760 := (not #17759)
+#17762 := (and #17761 #17760 #17757)
+#17763 := (not #17762)
+#17766 := (+ #17765 #13112)
+#17767 := (<= #17766 0::int)
+#17768 := (or #17767 #17763)
+#17769 := (not #17768)
+#17788 := (or #17769 #17784)
+#17753 := (not #13093)
+#17792 := (and #17753 #17788)
+#17796 := (or #13093 #17792)
+#15135 := (and #3258 #13017 #13051 #13059 #13061 #15127)
+#15140 := (not #15135)
+#17748 := (not #15140)
+#17800 := (and #17748 #17796)
+#17804 := (or #17742 #17745 #17800)
+#17828 := (not #13216)
+#17831 := (and #17828 #17804)
+#17739 := (not #13048)
+#17808 := (and #17739 #17804)
+#17812 := (or #17721 #17730 #17808)
+#17727 := (not #12503)
+#17816 := (and #17727 #17812)
+#17820 := (or #17721 #17724 #17816)
+#17736 := (not #13013)
+#17824 := (and #17736 #17820)
+#17835 := (or #17824 #17831)
+#17733 := (not #12993)
+#17839 := (and #17733 #17835)
+#17843 := (or #17721 #17730 #17839)
+#17847 := (and #17727 #17843)
+#17851 := (or #17721 #17724 #17847)
+#17718 := (not #12982)
+#17855 := (and #17718 #17851)
+#17909 := (or #17855 #17905)
+#15115 := (and #3055 #12030 #12741 #12910 #12913 #13359 #13389 #13392 #13402 #15076 #15085 #15096 #15107)
+#14482 := (and #3195 #3196 #3197 #3198 #3199 #3200 #12067 #12261 #12910 #12913)
+#14487 := (not #14482)
+#17682 := (not #14487)
+#17913 := (and #17682 #15115 #17909)
+#17679 := (not #12741)
+#17917 := (or #17676 #17679 #17913)
+#17921 := (and #15051 #17917)
+#17656 := (and #17655 #17654 #17652)
+#17657 := (not #17656)
+#17663 := (or #17662 #17657)
+#17664 := (not #17663)
+#17925 := (or #17664 #17921)
+#17648 := (not #12866)
+#17929 := (and #17648 #17925)
+#17933 := (or #12866 #17929)
+#17643 := (not #12860)
+#17937 := (and #17643 #17933)
+#17941 := (or #17631 #17640 #17937)
+#17637 := (not #12690)
+#17945 := (and #17637 #17941)
+#17949 := (or #17631 #17634 #17945)
+#17628 := (not #12702)
+#17953 := (and #17628 #17949)
+#17957 := (or #12702 #17953)
+#18263 := (iff #17957 #18262)
+#18260 := (iff #17953 #18259)
+#18257 := (iff #17949 #18256)
+#18254 := (iff #17945 #18251)
+#18248 := (and #11976 #18245)
+#18252 := (iff #18248 #18251)
+#18253 := [rewrite]: #18252
+#18249 := (iff #17945 #18248)
+#18246 := (iff #17941 #18245)
+#18243 := (iff #17937 #18240)
+#18237 := (and #12855 #18234)
+#18241 := (iff #18237 #18240)
+#18242 := [rewrite]: #18241
+#18238 := (iff #17937 #18237)
+#18235 := (iff #17933 #18234)
+#18232 := (iff #17929 #18231)
+#18229 := (iff #17925 #18228)
+#18226 := (iff #17921 #18225)
+#18223 := (iff #17917 #18222)
+#18220 := (iff #17913 #18217)
+#18214 := (and #14482 #15115 #18211)
+#18218 := (iff #18214 #18217)
+#18219 := [rewrite]: #18218
+#18215 := (iff #17913 #18214)
+#18212 := (iff #17909 #18211)
+#18209 := (iff #17905 #18206)
+#18203 := (and #13342 #14493 #18198)
+#18207 := (iff #18203 #18206)
+#18208 := [rewrite]: #18207
+#18204 := (iff #17905 #18203)
+#18201 := (iff #17901 #18198)
+#18195 := (or #18192 #17897)
+#18199 := (iff #18195 #18198)
+#18200 := [rewrite]: #18199
+#18196 := (iff #17901 #18195)
+#18193 := (iff #17878 #18192)
+#18190 := (iff #17877 #18189)
+#18187 := (iff #17872 #18186)
+#18184 := (iff #17871 #18181)
+#18178 := (and #17870 #18175 #17866)
+#18182 := (iff #18178 #18181)
+#18183 := [rewrite]: #18182
+#18179 := (iff #17871 #18178)
+#18176 := (iff #17869 #18175)
+#18173 := (iff #17868 #18170)
+#18162 := (+ #12727 ?x776!15)
+#18165 := (>= #18162 0::int)
+#18171 := (iff #18165 #18170)
+#18172 := [rewrite]: #18171
+#18166 := (iff #17868 #18165)
+#18163 := (= #17867 #18162)
+#18164 := [rewrite]: #18163
+#18167 := [monotonicity #18164]: #18166
+#18174 := [trans #18167 #18172]: #18173
+#18177 := [monotonicity #18174]: #18176
+#18180 := [monotonicity #18177]: #18179
+#18185 := [trans #18180 #18183]: #18184
+#18188 := [monotonicity #18185]: #18187
+#18160 := (iff #17876 #18157)
+#18149 := (+ #13286 #17874)
+#18152 := (<= #18149 0::int)
+#18158 := (iff #18152 #18157)
+#18159 := [rewrite]: #18158
+#18153 := (iff #17876 #18152)
+#18150 := (= #17875 #18149)
+#18151 := [rewrite]: #18150
+#18154 := [monotonicity #18151]: #18153
+#18161 := [trans #18154 #18159]: #18160
+#18191 := [monotonicity #18161 #18188]: #18190
+#18194 := [monotonicity #18191]: #18193
+#18197 := [monotonicity #18194]: #18196
+#18202 := [trans #18197 #18200]: #18201
+#18147 := (iff #17862 #14493)
+#18148 := [rewrite]: #18147
+#18145 := (iff #17859 #13342)
+#18146 := [rewrite]: #18145
+#18205 := [monotonicity #18146 #18148 #18202]: #18204
+#18210 := [trans #18205 #18208]: #18209
+#18143 := (iff #17855 #18140)
+#18137 := (and #12977 #18134)
+#18141 := (iff #18137 #18140)
+#18142 := [rewrite]: #18141
+#18138 := (iff #17855 #18137)
+#18135 := (iff #17851 #18134)
+#18132 := (iff #17847 #18129)
+#18126 := (and #12352 #18123)
+#18130 := (iff #18126 #18129)
+#18131 := [rewrite]: #18130
+#18127 := (iff #17847 #18126)
+#18124 := (iff #17843 #18123)
+#18121 := (iff #17839 #18118)
+#18115 := (and #12988 #18112)
+#18119 := (iff #18115 #18118)
+#18120 := [rewrite]: #18119
+#18116 := (iff #17839 #18115)
+#18113 := (iff #17835 #18112)
+#18110 := (iff #17831 #18107)
+#18104 := (and #13211 #18069)
+#18108 := (iff #18104 #18107)
+#18109 := [rewrite]: #18108
+#18105 := (iff #17831 #18104)
+#18070 := (iff #17804 #18069)
+#18067 := (iff #17800 #18064)
+#18061 := (and #15135 #18058)
+#18065 := (iff #18061 #18064)
+#18066 := [rewrite]: #18065
+#18062 := (iff #17800 #18061)
+#18059 := (iff #17796 #18058)
+#18056 := (iff #17792 #18055)
+#18053 := (iff #17788 #18050)
+#18047 := (or #18044 #17784)
+#18051 := (iff #18047 #18050)
+#18052 := [rewrite]: #18051
+#18048 := (iff #17788 #18047)
+#18045 := (iff #17769 #18044)
+#18042 := (iff #17768 #18041)
+#18039 := (iff #17763 #18038)
+#18036 := (iff #17762 #18033)
+#18030 := (and #17761 #18027 #17757)
+#18034 := (iff #18030 #18033)
+#18035 := [rewrite]: #18034
+#18031 := (iff #17762 #18030)
+#18028 := (iff #17760 #18027)
+#18025 := (iff #17759 #18022)
+#18014 := (+ #13062 ?x785!14)
+#18017 := (>= #18014 0::int)
+#18023 := (iff #18017 #18022)
+#18024 := [rewrite]: #18023
+#18018 := (iff #17759 #18017)
+#18015 := (= #17758 #18014)
+#18016 := [rewrite]: #18015
+#18019 := [monotonicity #18016]: #18018
+#18026 := [trans #18019 #18024]: #18025
+#18029 := [monotonicity #18026]: #18028
+#18032 := [monotonicity #18029]: #18031
+#18037 := [trans #18032 #18035]: #18036
+#18040 := [monotonicity #18037]: #18039
+#18012 := (iff #17767 #18009)
+#18001 := (+ #13112 #17765)
+#18004 := (<= #18001 0::int)
+#18010 := (iff #18004 #18009)
+#18011 := [rewrite]: #18010
+#18005 := (iff #17767 #18004)
+#18002 := (= #17766 #18001)
+#18003 := [rewrite]: #18002
+#18006 := [monotonicity #18003]: #18005
+#18013 := [trans #18006 #18011]: #18012
+#18043 := [monotonicity #18013 #18040]: #18042
+#18046 := [monotonicity #18043]: #18045
+#18049 := [monotonicity #18046]: #18048
+#18054 := [trans #18049 #18052]: #18053
+#17999 := (iff #17753 #13089)
+#18000 := [rewrite]: #17999
+#18057 := [monotonicity #18000 #18054]: #18056
+#18060 := [monotonicity #18057]: #18059
+#17997 := (iff #17748 #15135)
+#17998 := [rewrite]: #17997
+#18063 := [monotonicity #17998 #18060]: #18062
+#18068 := [trans #18063 #18066]: #18067
+#18071 := [monotonicity #18068]: #18070
+#18102 := (iff #17828 #13211)
+#18103 := [rewrite]: #18102
+#18106 := [monotonicity #18103 #18071]: #18105
+#18111 := [trans #18106 #18109]: #18110
+#18100 := (iff #17824 #18097)
+#18094 := (and #13008 #18091)
+#18098 := (iff #18094 #18097)
+#18099 := [rewrite]: #18098
+#18095 := (iff #17824 #18094)
+#18092 := (iff #17820 #18091)
+#18089 := (iff #17816 #18086)
+#18083 := (and #12352 #18080)
+#18087 := (iff #18083 #18086)
+#18088 := [rewrite]: #18087
+#18084 := (iff #17816 #18083)
+#18081 := (iff #17812 #18080)
+#18078 := (iff #17808 #18075)
+#18072 := (and #13043 #18069)
+#18076 := (iff #18072 #18075)
+#18077 := [rewrite]: #18076
+#18073 := (iff #17808 #18072)
+#17995 := (iff #17739 #13043)
+#17996 := [rewrite]: #17995
+#18074 := [monotonicity #17996 #18071]: #18073
+#18079 := [trans #18074 #18077]: #18078
+#18082 := [monotonicity #18079]: #18081
+#17989 := (iff #17727 #12352)
+#17990 := [rewrite]: #17989
+#18085 := [monotonicity #17990 #18082]: #18084
+#18090 := [trans #18085 #18088]: #18089
+#18093 := [monotonicity #18090]: #18092
+#17993 := (iff #17736 #13008)
+#17994 := [rewrite]: #17993
+#18096 := [monotonicity #17994 #18093]: #18095
+#18101 := [trans #18096 #18099]: #18100
+#18114 := [monotonicity #18101 #18111]: #18113
+#17991 := (iff #17733 #12988)
+#17992 := [rewrite]: #17991
+#18117 := [monotonicity #17992 #18114]: #18116
+#18122 := [trans #18117 #18120]: #18121
+#18125 := [monotonicity #18122]: #18124
+#18128 := [monotonicity #17990 #18125]: #18127
+#18133 := [trans #18128 #18131]: #18132
+#18136 := [monotonicity #18133]: #18135
+#17987 := (iff #17718 #12977)
+#17988 := [rewrite]: #17987
+#18139 := [monotonicity #17988 #18136]: #18138
+#18144 := [trans #18139 #18142]: #18143
+#18213 := [monotonicity #18144 #18210]: #18212
+#17985 := (iff #17682 #14482)
+#17986 := [rewrite]: #17985
+#18216 := [monotonicity #17986 #18213]: #18215
+#18221 := [trans #18216 #18219]: #18220
+#17983 := (iff #17679 #12740)
+#17984 := [rewrite]: #17983
+#18224 := [monotonicity #17984 #18221]: #18223
+#18227 := [monotonicity #18224]: #18226
+#17981 := (iff #17664 #17980)
+#17978 := (iff #17663 #17977)
+#17975 := (iff #17657 #17974)
+#17972 := (iff #17656 #17971)
+#17973 := [rewrite]: #17972
+#17976 := [monotonicity #17973]: #17975
+#17979 := [monotonicity #17976]: #17978
+#17982 := [monotonicity #17979]: #17981
+#18230 := [monotonicity #17982 #18227]: #18229
+#17969 := (iff #17648 #12863)
+#17970 := [rewrite]: #17969
+#18233 := [monotonicity #17970 #18230]: #18232
+#18236 := [monotonicity #18233]: #18235
+#17967 := (iff #17643 #12855)
+#17968 := [rewrite]: #17967
+#18239 := [monotonicity #17968 #18236]: #18238
+#18244 := [trans #18239 #18242]: #18243
+#18247 := [monotonicity #18244]: #18246
+#17965 := (iff #17637 #11976)
+#17966 := [rewrite]: #17965
+#18250 := [monotonicity #17966 #18247]: #18249
+#18255 := [trans #18250 #18253]: #18254
+#18258 := [monotonicity #18255]: #18257
+#17963 := (iff #17628 #11967)
+#17964 := [rewrite]: #17963
+#18261 := [monotonicity #17964 #18258]: #18260
+#18264 := [monotonicity #18261]: #18263
+#15265 := (exists (vars (?x778 int)) #15260)
+#15254 := (not #15251)
+#15268 := (or #15254 #15265)
+#15271 := (and #15251 #15268)
+#15277 := (or #13347 #14498 #15271)
+#15165 := (not #15162)
+#15171 := (or #13135 #15165)
+#15176 := (and #15162 #15171)
+#15179 := (or #13093 #15176)
+#15182 := (and #13089 #15179)
+#15185 := (or #15140 #15182)
+#15188 := (and #13051 #15127 #15185)
+#15209 := (or #13216 #15188)
+#15191 := (or #13048 #15188)
+#15194 := (and #12346 #12355 #15191)
+#15197 := (or #12503 #15194)
+#15200 := (and #12346 #12349 #15197)
+#15203 := (or #13013 #15200)
+#15214 := (and #15203 #15209)
+#15217 := (or #12993 #15214)
+#15220 := (and #12346 #12355 #15217)
+#15223 := (or #12503 #15220)
+#15226 := (and #12346 #12349 #15223)
+#15229 := (or #12982 #15226)
+#15282 := (and #15229 #15277)
+#15120 := (not #15115)
+#15288 := (or #14487 #15120 #15282)
+#15293 := (and #3055 #12741 #15288)
+#15054 := (not #15051)
+#15296 := (or #15054 #15293)
+#15299 := (and #15051 #15296)
+#15302 := (or #12866 #15299)
+#15305 := (and #12863 #15302)
+#15308 := (or #12860 #15305)
+#15311 := (and #11970 #11979 #15308)
+#15314 := (or #12690 #15311)
+#15317 := (and #11970 #11973 #15314)
+#15320 := (or #12702 #15317)
+#15323 := (and #11967 #15320)
+#15326 := (not #15323)
+#17958 := (~ #15326 #17957)
+#17954 := (not #15320)
+#17955 := (~ #17954 #17953)
+#17950 := (not #15317)
+#17951 := (~ #17950 #17949)
+#17946 := (not #15314)
+#17947 := (~ #17946 #17945)
+#17942 := (not #15311)
+#17943 := (~ #17942 #17941)
+#17938 := (not #15308)
+#17939 := (~ #17938 #17937)
+#17934 := (not #15305)
+#17935 := (~ #17934 #17933)
+#17930 := (not #15302)
+#17931 := (~ #17930 #17929)
+#17926 := (not #15299)
+#17927 := (~ #17926 #17925)
+#17922 := (not #15296)
+#17923 := (~ #17922 #17921)
+#17918 := (not #15293)
+#17919 := (~ #17918 #17917)
+#17914 := (not #15288)
+#17915 := (~ #17914 #17913)
+#17910 := (not #15282)
+#17911 := (~ #17910 #17909)
+#17906 := (not #15277)
+#17907 := (~ #17906 #17905)
+#17902 := (not #15271)
+#17903 := (~ #17902 #17901)
+#17898 := (not #15268)
+#17899 := (~ #17898 #17897)
+#17894 := (not #15265)
+#17895 := (~ #17894 #17893)
+#17891 := (~ #17890 #17890)
+#17892 := [refl]: #17891
+#17896 := [nnf-neg #17892]: #17895
+#17887 := (not #15254)
+#17888 := (~ #17887 #15251)
+#17885 := (~ #15251 #15251)
+#17883 := (~ #15246 #15246)
+#17884 := [refl]: #17883
+#17886 := [nnf-pos #17884]: #17885
+#17889 := [nnf-neg #17886]: #17888
+#17900 := [nnf-neg #17889 #17896]: #17899
+#17879 := (~ #15254 #17878)
+#17880 := [sk]: #17879
+#17904 := [nnf-neg #17880 #17900]: #17903
+#17863 := (~ #17862 #17862)
+#17864 := [refl]: #17863
+#17860 := (~ #17859 #17859)
+#17861 := [refl]: #17860
+#17908 := [nnf-neg #17861 #17864 #17904]: #17907
+#17856 := (not #15229)
+#17857 := (~ #17856 #17855)
+#17852 := (not #15226)
+#17853 := (~ #17852 #17851)
+#17848 := (not #15223)
+#17849 := (~ #17848 #17847)
+#17844 := (not #15220)
+#17845 := (~ #17844 #17843)
+#17840 := (not #15217)
+#17841 := (~ #17840 #17839)
+#17836 := (not #15214)
+#17837 := (~ #17836 #17835)
+#17832 := (not #15209)
+#17833 := (~ #17832 #17831)
+#17805 := (not #15188)
+#17806 := (~ #17805 #17804)
+#17801 := (not #15185)
+#17802 := (~ #17801 #17800)
+#17797 := (not #15182)
+#17798 := (~ #17797 #17796)
+#17793 := (not #15179)
+#17794 := (~ #17793 #17792)
+#17789 := (not #15176)
+#17790 := (~ #17789 #17788)
+#17785 := (not #15171)
+#17786 := (~ #17785 #17784)
+#17781 := (not #15165)
+#17782 := (~ #17781 #15162)
+#17779 := (~ #15162 #15162)
+#17777 := (~ #15157 #15157)
+#17778 := [refl]: #17777
+#17780 := [nnf-pos #17778]: #17779
+#17783 := [nnf-neg #17780]: #17782
+#17775 := (~ #17774 #17774)
+#17776 := [refl]: #17775
+#17787 := [nnf-neg #17776 #17783]: #17786
+#17770 := (~ #15165 #17769)
+#17771 := [sk]: #17770
+#17791 := [nnf-neg #17771 #17787]: #17790
+#17754 := (~ #17753 #17753)
+#17755 := [refl]: #17754
+#17795 := [nnf-neg #17755 #17791]: #17794
+#17751 := (~ #13093 #13093)
+#17752 := [refl]: #17751
+#17799 := [nnf-neg #17752 #17795]: #17798
+#17749 := (~ #17748 #17748)
+#17750 := [refl]: #17749
+#17803 := [nnf-neg #17750 #17799]: #17802
+#17746 := (~ #17745 #17745)
+#17747 := [refl]: #17746
+#17743 := (~ #17742 #17742)
+#17744 := [refl]: #17743
+#17807 := [nnf-neg #17744 #17747 #17803]: #17806
+#17829 := (~ #17828 #17828)
+#17830 := [refl]: #17829
+#17834 := [nnf-neg #17830 #17807]: #17833
+#17825 := (not #15203)
+#17826 := (~ #17825 #17824)
+#17821 := (not #15200)
+#17822 := (~ #17821 #17820)
+#17817 := (not #15197)
+#17818 := (~ #17817 #17816)
+#17813 := (not #15194)
+#17814 := (~ #17813 #17812)
+#17809 := (not #15191)
+#17810 := (~ #17809 #17808)
+#17740 := (~ #17739 #17739)
+#17741 := [refl]: #17740
+#17811 := [nnf-neg #17741 #17807]: #17810
+#17731 := (~ #17730 #17730)
+#17732 := [refl]: #17731
+#17722 := (~ #17721 #17721)
+#17723 := [refl]: #17722
+#17815 := [nnf-neg #17723 #17732 #17811]: #17814
+#17728 := (~ #17727 #17727)
+#17729 := [refl]: #17728
+#17819 := [nnf-neg #17729 #17815]: #17818
+#17725 := (~ #17724 #17724)
+#17726 := [refl]: #17725
+#17823 := [nnf-neg #17723 #17726 #17819]: #17822
+#17737 := (~ #17736 #17736)
+#17738 := [refl]: #17737
+#17827 := [nnf-neg #17738 #17823]: #17826
+#17838 := [nnf-neg #17827 #17834]: #17837
+#17734 := (~ #17733 #17733)
+#17735 := [refl]: #17734
+#17842 := [nnf-neg #17735 #17838]: #17841
+#17846 := [nnf-neg #17723 #17732 #17842]: #17845
+#17850 := [nnf-neg #17729 #17846]: #17849
+#17854 := [nnf-neg #17723 #17726 #17850]: #17853
+#17719 := (~ #17718 #17718)
+#17720 := [refl]: #17719
+#17858 := [nnf-neg #17720 #17854]: #17857
+#17912 := [nnf-neg #17858 #17908]: #17911
+#17715 := (not #15120)
+#17716 := (~ #17715 #15115)
+#17713 := (~ #15115 #15115)
+#17711 := (~ #15107 #15107)
+#17712 := [refl]: #17711
+#17709 := (~ #15096 #15096)
+#17710 := [refl]: #17709
+#17707 := (~ #15085 #15085)
+#17708 := [refl]: #17707
+#17705 := (~ #15076 #15076)
+#17703 := (~ #15071 #15071)
+#17704 := [refl]: #17703
+#17706 := [nnf-pos #17704]: #17705
+#17701 := (~ #13402 #13402)
+#17702 := [refl]: #17701
+#17699 := (~ #13392 #13392)
+#17700 := [refl]: #17699
+#17697 := (~ #13389 #13389)
+#17698 := [refl]: #17697
+#17695 := (~ #13359 #13359)
+#17696 := [refl]: #17695
+#17693 := (~ #12913 #12913)
+#17694 := [refl]: #17693
+#17691 := (~ #12910 #12910)
+#17692 := [refl]: #17691
+#17689 := (~ #12741 #12741)
+#17690 := [refl]: #17689
+#17687 := (~ #12030 #12030)
+#17688 := [refl]: #17687
+#17685 := (~ #3055 #3055)
+#17686 := [refl]: #17685
+#17714 := [monotonicity #17686 #17688 #17690 #17692 #17694 #17696 #17698 #17700 #17702 #17706 #17708 #17710 #17712]: #17713
+#17717 := [nnf-neg #17714]: #17716
+#17683 := (~ #17682 #17682)
+#17684 := [refl]: #17683
+#17916 := [nnf-neg #17684 #17717 #17912]: #17915
+#17680 := (~ #17679 #17679)
+#17681 := [refl]: #17680
+#17677 := (~ #17676 #17676)
+#17678 := [refl]: #17677
+#17920 := [nnf-neg #17678 #17681 #17916]: #17919
+#17673 := (not #15054)
+#17674 := (~ #17673 #15051)
+#17671 := (~ #15051 #15051)
+#17669 := (~ #15046 #15046)
+#17670 := [refl]: #17669
+#17672 := [nnf-pos #17670]: #17671
+#17675 := [nnf-neg #17672]: #17674
+#17924 := [nnf-neg #17675 #17920]: #17923
+#17665 := (~ #15054 #17664)
+#17666 := [sk]: #17665
+#17928 := [nnf-neg #17666 #17924]: #17927
+#17649 := (~ #17648 #17648)
+#17650 := [refl]: #17649
+#17932 := [nnf-neg #17650 #17928]: #17931
+#17646 := (~ #12866 #12866)
+#17647 := [refl]: #17646
+#17936 := [nnf-neg #17647 #17932]: #17935
+#17644 := (~ #17643 #17643)
+#17645 := [refl]: #17644
+#17940 := [nnf-neg #17645 #17936]: #17939
+#17641 := (~ #17640 #17640)
+#17642 := [refl]: #17641
+#17632 := (~ #17631 #17631)
+#17633 := [refl]: #17632
+#17944 := [nnf-neg #17633 #17642 #17940]: #17943
+#17638 := (~ #17637 #17637)
+#17639 := [refl]: #17638
+#17948 := [nnf-neg #17639 #17944]: #17947
+#17635 := (~ #17634 #17634)
+#17636 := [refl]: #17635
+#17952 := [nnf-neg #17633 #17636 #17948]: #17951
+#17629 := (~ #17628 #17628)
+#17630 := [refl]: #17629
+#17956 := [nnf-neg #17630 #17952]: #17955
+#17626 := (~ #12702 #12702)
+#17627 := [refl]: #17626
+#17959 := [nnf-neg #17627 #17956]: #17958
+#14519 := (or #13325 #13347 #14498)
+#14524 := (and #13247 #14519)
+#14530 := (or #13453 #14487 #14524)
+#14535 := (and #3055 #12741 #14530)
+#14538 := (or #12895 #14535)
+#14541 := (and #12892 #14538)
+#14544 := (or #12866 #14541)
+#14547 := (and #12863 #14544)
+#14550 := (or #12860 #14547)
+#14553 := (and #11970 #11979 #14550)
+#14556 := (or #12690 #14553)
+#14559 := (and #11970 #11973 #14556)
+#14562 := (or #12702 #14559)
+#14565 := (and #11967 #14562)
+#14568 := (not #14565)
+#15327 := (iff #14568 #15326)
+#15324 := (iff #14565 #15323)
+#15321 := (iff #14562 #15320)
+#15318 := (iff #14559 #15317)
+#15315 := (iff #14556 #15314)
+#15312 := (iff #14553 #15311)
+#15309 := (iff #14550 #15308)
+#15306 := (iff #14547 #15305)
+#15303 := (iff #14544 #15302)
+#15300 := (iff #14541 #15299)
+#15297 := (iff #14538 #15296)
+#15294 := (iff #14535 #15293)
+#15291 := (iff #14530 #15288)
+#15285 := (or #15120 #14487 #15282)
+#15289 := (iff #15285 #15288)
+#15290 := [rewrite]: #15289
+#15286 := (iff #14530 #15285)
+#15283 := (iff #14524 #15282)
+#15280 := (iff #14519 #15277)
+#15274 := (or #15271 #13347 #14498)
+#15278 := (iff #15274 #15277)
+#15279 := [rewrite]: #15278
+#15275 := (iff #14519 #15274)
+#15272 := (iff #13325 #15271)
+#15269 := (iff #13320 #15268)
+#15266 := (iff #13311 #15265)
+#15263 := (iff #13306 #15260)
+#15257 := (and #3142 #4041 #14378 #13272)
+#15261 := (iff #15257 #15260)
+#15262 := [rewrite]: #15261
+#15258 := (iff #13306 #15257)
+#14373 := (iff #4355 #14378)
+#14394 := -4294967295::int
+#14386 := (+ -4294967295::int #161)
+#14379 := (<= #14386 0::int)
+#14375 := (iff #14379 #14378)
+#14376 := [rewrite]: #14375
+#14380 := (iff #4355 #14379)
+#14381 := (= #4354 #14386)
+#14387 := (+ #161 -4294967295::int)
+#14383 := (= #14387 #14386)
+#14384 := [rewrite]: #14383
+#14388 := (= #4354 #14387)
+#14389 := (= #4353 -4294967295::int)
+#14395 := (* -1::int 4294967295::int)
+#14391 := (= #14395 -4294967295::int)
+#14392 := [rewrite]: #14391
+#14396 := (= #4353 #14395)
+#7476 := (= uf_76 4294967295::int)
+#947 := 65536::int
+#1322 := (* 65536::int 65536::int)
+#1327 := (- #1322 1::int)
+#1328 := (= uf_76 #1327)
+#7477 := (iff #1328 #7476)
+#7474 := (= #1327 4294967295::int)
+#1010 := 4294967296::int
+#7467 := (- 4294967296::int 1::int)
+#7472 := (= #7467 4294967295::int)
+#7473 := [rewrite]: #7472
+#7469 := (= #1327 #7467)
+#7438 := (= #1322 4294967296::int)
+#7439 := [rewrite]: #7438
+#7470 := [monotonicity #7439]: #7469
+#7475 := [trans #7470 #7473]: #7474
+#7478 := [monotonicity #7475]: #7477
+#7466 := [asserted]: #1328
+#7481 := [mp #7466 #7478]: #7476
+#14393 := [monotonicity #7481]: #14396
+#14390 := [trans #14393 #14392]: #14389
+#14385 := [monotonicity #14390]: #14388
+#14382 := [trans #14385 #14384]: #14381
+#14377 := [monotonicity #14382]: #14380
+#14374 := [trans #14377 #14376]: #14373
+#15259 := [monotonicity #14374]: #15258
+#15264 := [trans #15259 #15262]: #15263
+#15267 := [quant-intro #15264]: #15266
+#15255 := (iff #13314 #15254)
+#15252 := (iff #13294 #15251)
+#15249 := (iff #13291 #15246)
+#15243 := (or #15240 #13288)
+#15247 := (iff #15243 #15246)
+#15248 := [rewrite]: #15247
+#15244 := (iff #13291 #15243)
+#15241 := (iff #13283 #15240)
+#15238 := (iff #13278 #15235)
+#15232 := (and #4041 #14378 #13272)
+#15236 := (iff #15232 #15235)
+#15237 := [rewrite]: #15236
+#15233 := (iff #13278 #15232)
+#15234 := [monotonicity #14374]: #15233
+#15239 := [trans #15234 #15237]: #15238
+#15242 := [monotonicity #15239]: #15241
+#15245 := [monotonicity #15242]: #15244
+#15250 := [trans #15245 #15248]: #15249
+#15253 := [quant-intro #15250]: #15252
+#15256 := [monotonicity #15253]: #15255
+#15270 := [monotonicity #15256 #15267]: #15269
+#15273 := [monotonicity #15253 #15270]: #15272
+#15276 := [monotonicity #15273]: #15275
+#15281 := [trans #15276 #15279]: #15280
+#15230 := (iff #13247 #15229)
+#15227 := (iff #13242 #15226)
+#15224 := (iff #13236 #15223)
+#15221 := (iff #13231 #15220)
+#15218 := (iff #13225 #15217)
+#15215 := (iff #13222 #15214)
+#15212 := (iff #13219 #15209)
+#15206 := (or #15188 #13216)
+#15210 := (iff #15206 #15209)
+#15211 := [rewrite]: #15210
+#15207 := (iff #13219 #15206)
+#15189 := (iff #13158 #15188)
+#15186 := (iff #13152 #15185)
+#15183 := (iff #13149 #15182)
+#15180 := (iff #13146 #15179)
+#15177 := (iff #13143 #15176)
+#15174 := (iff #13140 #15171)
+#15168 := (or #15165 #13135)
+#15172 := (iff #15168 #15171)
+#15173 := [rewrite]: #15172
+#15169 := (iff #13140 #15168)
+#15166 := (iff #13123 #15165)
+#15163 := (iff #13120 #15162)
+#15160 := (iff #13117 #15157)
+#15154 := (or #15151 #13114)
+#15158 := (iff #15154 #15157)
+#15159 := [rewrite]: #15158
+#15155 := (iff #13117 #15154)
+#15152 := (iff #13109 #15151)
+#15149 := (iff #13104 #15146)
+#15143 := (and #4041 #14378 #13098)
+#15147 := (iff #15143 #15146)
+#15148 := [rewrite]: #15147
+#15144 := (iff #13104 #15143)
+#15145 := [monotonicity #14374]: #15144
+#15150 := [trans #15145 #15148]: #15149
+#15153 := [monotonicity #15150]: #15152
+#15156 := [monotonicity #15153]: #15155
+#15161 := [trans #15156 #15159]: #15160
+#15164 := [quant-intro #15161]: #15163
+#15167 := [monotonicity #15164]: #15166
+#15170 := [monotonicity #15167]: #15169
+#15175 := [trans #15170 #15173]: #15174
+#15178 := [monotonicity #15164 #15175]: #15177
+#15181 := [monotonicity #15178]: #15180
+#15184 := [monotonicity #15181]: #15183
+#15141 := (iff #13086 #15140)
+#15138 := (iff #13081 #15135)
+#15132 := (and #3258 #13017 #13051 #15127 #13059 #13061)
+#15136 := (iff #15132 #15135)
+#15137 := [rewrite]: #15136
+#15133 := (iff #13081 #15132)
+#15130 := (iff #13054 #15127)
+#15079 := (+ 4294967295::int #12965)
+#15123 := (>= #15079 1::int)
+#15128 := (iff #15123 #15127)
+#15129 := [rewrite]: #15128
+#15124 := (iff #13054 #15123)
+#15080 := (= #13055 #15079)
+#15081 := [monotonicity #7481]: #15080
+#15125 := [monotonicity #15081]: #15124
+#15131 := [trans #15125 #15129]: #15130
+#15134 := [monotonicity #15131]: #15133
+#15139 := [trans #15134 #15137]: #15138
+#15142 := [monotonicity #15139]: #15141
+#15187 := [monotonicity #15142 #15184]: #15186
+#15190 := [monotonicity #15131 #15187]: #15189
+#15208 := [monotonicity #15190]: #15207
+#15213 := [trans #15208 #15211]: #15212
+#15204 := (iff #13185 #15203)
+#15201 := (iff #13180 #15200)
+#15198 := (iff #13174 #15197)
+#15195 := (iff #13169 #15194)
+#15192 := (iff #13163 #15191)
+#15193 := [monotonicity #15190]: #15192
+#15196 := [monotonicity #15193]: #15195
+#15199 := [monotonicity #15196]: #15198
+#15202 := [monotonicity #15199]: #15201
+#15205 := [monotonicity #15202]: #15204
+#15216 := [monotonicity #15205 #15213]: #15215
+#15219 := [monotonicity #15216]: #15218
+#15222 := [monotonicity #15219]: #15221
+#15225 := [monotonicity #15222]: #15224
+#15228 := [monotonicity #15225]: #15227
+#15231 := [monotonicity #15228]: #15230
+#15284 := [monotonicity #15231 #15281]: #15283
+#15121 := (iff #13453 #15120)
+#15118 := (iff #13448 #15115)
+#15112 := (and #3055 #12030 #12741 #12910 #12913 #13359 #15076 #13389 #13392 #15085 #15096 #13402 #15107)
+#15116 := (iff #15112 #15115)
+#15117 := [rewrite]: #15116
+#15113 := (iff #13448 #15112)
+#15110 := (iff #13405 #15107)
+#15101 := (+ 255::int #13378)
+#15104 := (>= #15101 0::int)
+#15108 := (iff #15104 #15107)
+#15109 := [rewrite]: #15108
+#15105 := (iff #13405 #15104)
+#15102 := (= #13406 #15101)
+#1332 := (= uf_78 255::int)
+#7480 := [asserted]: #1332
+#15103 := [monotonicity #7480]: #15102
+#15106 := [monotonicity #15103]: #15105
+#15111 := [trans #15106 #15109]: #15110
+#15099 := (iff #13398 #15096)
+#15090 := (+ 4294967295::int #13356)
+#15093 := (>= #15090 0::int)
+#15097 := (iff #15093 #15096)
+#15098 := [rewrite]: #15097
+#15094 := (iff #13398 #15093)
+#15091 := (= #13399 #15090)
+#15092 := [monotonicity #7481]: #15091
+#15095 := [monotonicity #15092]: #15094
+#15100 := [trans #15095 #15098]: #15099
+#15088 := (iff #13395 #15085)
+#15082 := (>= #15079 0::int)
+#15086 := (iff #15082 #15085)
+#15087 := [rewrite]: #15086
+#15083 := (iff #13395 #15082)
+#15084 := [monotonicity #15081]: #15083
+#15089 := [trans #15084 #15087]: #15088
+#15077 := (iff #13386 #15076)
+#15074 := (iff #13383 #15071)
+#15068 := (or #15065 #13380)
+#15072 := (iff #15068 #15071)
+#15073 := [rewrite]: #15072
+#15069 := (iff #13383 #15068)
+#15066 := (iff #13375 #15065)
+#15063 := (iff #13370 #15060)
+#15057 := (and #4041 #14378 #13364)
+#15061 := (iff #15057 #15060)
+#15062 := [rewrite]: #15061
+#15058 := (iff #13370 #15057)
+#15059 := [monotonicity #14374]: #15058
+#15064 := [trans #15059 #15062]: #15063
+#15067 := [monotonicity #15064]: #15066
+#15070 := [monotonicity #15067]: #15069
+#15075 := [trans #15070 #15073]: #15074
+#15078 := [quant-intro #15075]: #15077
+#15114 := [monotonicity #15078 #15089 #15100 #15111]: #15113
+#15119 := [trans #15114 #15117]: #15118
+#15122 := [monotonicity #15119]: #15121
+#15287 := [monotonicity #15122 #15284]: #15286
+#15292 := [trans #15287 #15290]: #15291
+#15295 := [monotonicity #15292]: #15294
+#15055 := (iff #12895 #15054)
+#15052 := (iff #12892 #15051)
+#15049 := (iff #12889 #15046)
+#15043 := (or #15040 #12885)
+#15047 := (iff #15043 #15046)
+#15048 := [rewrite]: #15047
+#15044 := (iff #12889 #15043)
+#15041 := (iff #12881 #15040)
+#15038 := (iff #12876 #15035)
+#15032 := (and #4041 #14378 #12870)
+#15036 := (iff #15032 #15035)
+#15037 := [rewrite]: #15036
+#15033 := (iff #12876 #15032)
+#15034 := [monotonicity #14374]: #15033
+#15039 := [trans #15034 #15037]: #15038
+#15042 := [monotonicity #15039]: #15041
+#15045 := [monotonicity #15042]: #15044
+#15050 := [trans #15045 #15048]: #15049
+#15053 := [quant-intro #15050]: #15052
+#15056 := [monotonicity #15053]: #15055
+#15298 := [monotonicity #15056 #15295]: #15297
+#15301 := [monotonicity #15053 #15298]: #15300
+#15304 := [monotonicity #15301]: #15303
+#15307 := [monotonicity #15304]: #15306
+#15310 := [monotonicity #15307]: #15309
+#15313 := [monotonicity #15310]: #15312
+#15316 := [monotonicity #15313]: #15315
+#15319 := [monotonicity #15316]: #15318
+#15322 := [monotonicity #15319]: #15321
+#15325 := [monotonicity #15322]: #15324
+#15328 := [monotonicity #15325]: #15327
+#13550 := (not #13512)
+#14569 := (iff #13550 #14568)
+#14566 := (iff #13512 #14565)
+#14563 := (iff #13509 #14562)
+#14560 := (iff #13504 #14559)
+#14557 := (iff #13498 #14556)
+#14554 := (iff #13493 #14553)
+#14551 := (iff #13487 #14550)
+#14548 := (iff #13484 #14547)
+#14545 := (iff #13481 #14544)
+#14542 := (iff #13478 #14541)
+#14539 := (iff #13475 #14538)
+#14536 := (iff #13470 #14535)
+#14533 := (iff #13462 #14530)
+#14527 := (or #14487 #14524 #13453)
+#14531 := (iff #14527 #14530)
+#14532 := [rewrite]: #14531
+#14528 := (iff #13462 #14527)
+#14525 := (iff #13353 #14524)
+#14522 := (iff #13350 #14519)
+#14504 := (or #13325 #14498)
+#14516 := (or #14504 #13347)
+#14520 := (iff #14516 #14519)
+#14521 := [rewrite]: #14520
+#14517 := (iff #13350 #14516)
+#14514 := (iff #13331 #14504)
+#14509 := (and true #14504)
+#14512 := (iff #14509 #14504)
+#14513 := [rewrite]: #14512
+#14510 := (iff #13331 #14509)
+#14507 := (iff #13328 #14504)
+#14501 := (or #14498 #13325)
+#14505 := (iff #14501 #14504)
+#14506 := [rewrite]: #14505
+#14502 := (iff #13328 #14501)
+#14499 := (iff #13267 #14498)
+#14496 := (iff #13262 #14493)
+#14490 := (and true #12113 #12116 #12119 #12122 #12910 #12913)
+#14494 := (iff #14490 #14493)
+#14495 := [rewrite]: #14494
+#14491 := (iff #13262 #14490)
+#14199 := (iff up_216 true)
+#10740 := [asserted]: up_216
+#14200 := [iff-true #10740]: #14199
+#14492 := [monotonicity #14200]: #14491
+#14497 := [trans #14492 #14495]: #14496
+#14500 := [monotonicity #14497]: #14499
+#14503 := [monotonicity #14500]: #14502
+#14508 := [trans #14503 #14506]: #14507
+#14511 := [monotonicity #14200 #14508]: #14510
+#14515 := [trans #14511 #14513]: #14514
+#14518 := [monotonicity #14515]: #14517
+#14523 := [trans #14518 #14521]: #14522
+#14526 := [monotonicity #14523]: #14525
+#14488 := (iff #12962 #14487)
+#14485 := (iff #12957 #14482)
+#14479 := (and #3195 #3196 #3197 #3198 #3199 #3200 true #12067 #12261 #12910 #12913)
+#14483 := (iff #14479 #14482)
+#14484 := [rewrite]: #14483
+#14480 := (iff #12957 #14479)
+#14453 := (iff #11903 true)
+#14454 := [iff-true #13537]: #14453
+#14481 := [monotonicity #14454]: #14480
+#14486 := [trans #14481 #14484]: #14485
+#14489 := [monotonicity #14486]: #14488
+#14529 := [monotonicity #14489 #14526]: #14528
+#14534 := [trans #14529 #14532]: #14533
+#14537 := [monotonicity #14534]: #14536
+#14540 := [monotonicity #14537]: #14539
+#14543 := [monotonicity #14540]: #14542
+#14546 := [monotonicity #14543]: #14545
+#14549 := [monotonicity #14546]: #14548
+#14552 := [monotonicity #14549]: #14551
+#14555 := [monotonicity #14552]: #14554
+#14558 := [monotonicity #14555]: #14557
+#14561 := [monotonicity #14558]: #14560
+#14564 := [monotonicity #14561]: #14563
+#14567 := [monotonicity #14564]: #14566
+#14570 := [monotonicity #14567]: #14569
+#13551 := [not-or-elim #13523]: #13550
+#14571 := [mp #13551 #14570]: #14568
+#15329 := [mp #14571 #15328]: #15326
+#17960 := [mp~ #15329 #17959]: #17957
+#17961 := [mp #17960 #18264]: #18262
+#21850 := [mp #17961 #21849]: #21847
+#22710 := [mp #21850 #22709]: #22707
+#28083 := [unit-resolution #22710 #26255]: #22704
+#23931 := (or #22701 #22695)
+#23932 := [def-axiom]: #23931
+#28084 := [unit-resolution #23932 #28083]: #22695
+decl uf_15 :: (-> T5 T6 T2)
+decl uf_16 :: (-> T4 T5 T6)
+#26003 := (uf_16 uf_287 #25399)
+#26479 := (uf_15 #26392 #26003)
+#26480 := (= uf_9 #26479)
+#26004 := (uf_15 #25399 #26003)
+#28108 := (= #26004 #26479)
+#28104 := (= #26479 #26004)
+#27888 := (= #26392 #25399)
+#27862 := (= #26392 #2981)
+#27860 := (= #26391 #2977)
+#27858 := (= #24110 uf_286)
+#24111 := (= uf_286 #24110)
+#1602 := (uf_143 #1358)
+#8264 := (= #161 #1602)
+#8267 := (forall (vars (?x386 T3) (?x387 int)) (:pat #1592) #8264)
+#16571 := (~ #8267 #8267)
+#16569 := (~ #8264 #8264)
+#16570 := [refl]: #16569
+#16572 := [nnf-pos #16570]: #16571
+#1603 := (= #1602 #161)
+#1604 := (forall (vars (?x386 T3) (?x387 int)) (:pat #1592) #1603)
+#8268 := (iff #1604 #8267)
+#8265 := (iff #1603 #8264)
+#8266 := [rewrite]: #8265
+#8269 := [quant-intro #8266]: #8268
+#8263 := [asserted]: #1604
+#8272 := [mp #8263 #8269]: #8267
+#16573 := [mp~ #8272 #16572]: #8267
+#24117 := (not #8267)
+#24118 := (or #24117 #24111)
+#24119 := [quant-inst]: #24118
+#27676 := [unit-resolution #24119 #16573]: #24111
+#27859 := [symm #27676]: #27858
+#27861 := [monotonicity #27656 #27859]: #27860
+#27863 := [monotonicity #27861]: #27862
+#27889 := [trans #27863 #27875]: #27888
+#28105 := [monotonicity #27889]: #28104
+#28109 := [symm #28105]: #28108
+#26005 := (= uf_9 #26004)
+decl uf_53 :: (-> T4 T5 T6)
+#25994 := (uf_53 uf_287 #25399)
+#25995 := (uf_15 #23 #25994)
+#26000 := (pattern #25995)
+decl up_197 :: (-> T3 bool)
+#25998 := (up_197 #25810)
+#25996 := (= uf_9 #25995)
+#25997 := (not #25996)
+decl uf_147 :: (-> T5 T6 T2)
+decl uf_192 :: (-> T7 T6)
+decl uf_11 :: (-> T4 T5 T7)
+#25990 := (uf_11 uf_287 #25399)
+#25991 := (uf_192 #25990)
+#25992 := (uf_147 #23 #25991)
+#25993 := (= uf_9 #25992)
+#26010 := (or #25993 #25997 #25998)
+#26013 := (forall (vars (?x577 T5)) (:pat #26000) #26010)
+#26016 := (not #26013)
+#26006 := (not #26005)
+#26019 := (or #25875 #26006 #26016)
+#26022 := (not #26019)
+#28086 := (= #3042 #25946)
+#28089 := [symm #26166]: #28086
+#28090 := [trans #26255 #28089]: #25947
+#25968 := (or #25967 #25945 #25963)
+#25965 := [def-axiom]: #25968
+#28091 := [unit-resolution #25965 #28090 #26163]: #25945
+#26025 := (or #25966 #26022)
+#14 := (:var 2 T4)
+#2166 := (uf_196 #14 #15 #23)
+#2228 := (pattern #2166)
+#2229 := (uf_53 #13 #21)
+#2230 := (uf_15 #23 #2229)
+#2231 := (pattern #2230)
+#2158 := (uf_11 #13 #15)
+#2236 := (uf_192 #2158)
+#2237 := (uf_147 #23 #2236)
+#10024 := (= uf_9 #2237)
+#10018 := (= uf_9 #2230)
+#21005 := (not #10018)
+#1382 := (uf_13 #21)
+#2232 := (up_197 #1382)
+#21020 := (or #2232 #21005 #10024)
+#21025 := (forall (vars (?x577 T5)) (:pat #2231) #21020)
+#21031 := (not #21025)
+#2145 := (uf_16 #14 #23)
+#2146 := (uf_15 #15 #2145)
+#9724 := (= uf_9 #2146)
+#20840 := (not #9724)
+#180 := (uf_27 #14 #15)
+#3718 := (= uf_9 #180)
+#10361 := (not #3718)
+#21032 := (or #10361 #20840 #21031)
+#21033 := (not #21032)
+#9772 := (= uf_9 #2166)
+#10048 := (not #9772)
+#21038 := (or #10048 #21033)
+#21041 := (forall (vars (?x574 T4) (?x575 T5) (?x576 T5)) (:pat #2228) #21038)
+#2233 := (not #2232)
+#10021 := (and #2233 #10018)
+#10030 := (not #10021)
+#10031 := (or #10030 #10024)
+#10036 := (forall (vars (?x577 T5)) (:pat #2231) #10031)
+#10057 := (and #3718 #9724 #10036)
+#10060 := (or #10048 #10057)
+#10063 := (forall (vars (?x574 T4) (?x575 T5) (?x576 T5)) (:pat #2228) #10060)
+#21042 := (iff #10063 #21041)
+#21039 := (iff #10060 #21038)
+#21036 := (iff #10057 #21033)
+#21028 := (and #3718 #9724 #21025)
+#21034 := (iff #21028 #21033)
+#21035 := [rewrite]: #21034
+#21029 := (iff #10057 #21028)
+#21026 := (iff #10036 #21025)
+#21023 := (iff #10031 #21020)
+#21006 := (or #2232 #21005)
+#21017 := (or #21006 #10024)
+#21021 := (iff #21017 #21020)
+#21022 := [rewrite]: #21021
+#21018 := (iff #10031 #21017)
+#21015 := (iff #10030 #21006)
+#21007 := (not #21006)
+#21010 := (not #21007)
+#21013 := (iff #21010 #21006)
+#21014 := [rewrite]: #21013
+#21011 := (iff #10030 #21010)
+#21008 := (iff #10021 #21007)
+#21009 := [rewrite]: #21008
+#21012 := [monotonicity #21009]: #21011
+#21016 := [trans #21012 #21014]: #21015
+#21019 := [monotonicity #21016]: #21018
+#21024 := [trans #21019 #21022]: #21023
+#21027 := [quant-intro #21024]: #21026
+#21030 := [monotonicity #21027]: #21029
+#21037 := [trans #21030 #21035]: #21036
+#21040 := [monotonicity #21037]: #21039
+#21043 := [quant-intro #21040]: #21042
+#17090 := (~ #10063 #10063)
+#17088 := (~ #10060 #10060)
+#17086 := (~ #10057 #10057)
+#17084 := (~ #10036 #10036)
+#17082 := (~ #10031 #10031)
+#17083 := [refl]: #17082
+#17085 := [nnf-pos #17083]: #17084
+#17080 := (~ #9724 #9724)
+#17081 := [refl]: #17080
+#17078 := (~ #3718 #3718)
+#17079 := [refl]: #17078
+#17087 := [monotonicity #17079 #17081 #17085]: #17086
+#17076 := (~ #10048 #10048)
+#17077 := [refl]: #17076
+#17089 := [monotonicity #17077 #17087]: #17088
+#17091 := [nnf-pos #17089]: #17090
+#2238 := (= #2237 uf_9)
+#2234 := (= #2230 uf_9)
+#2235 := (and #2233 #2234)
+#2239 := (implies #2235 #2238)
+#2240 := (forall (vars (?x577 T5)) (:pat #2231) #2239)
+#184 := (= #180 uf_9)
+#2241 := (and #184 #2240)
+#2151 := (= #2146 uf_9)
+#2242 := (and #2151 #2241)
+#2167 := (= #2166 uf_9)
+#2243 := (implies #2167 #2242)
+#2244 := (forall (vars (?x574 T4) (?x575 T5) (?x576 T5)) (:pat #2228) #2243)
+#10066 := (iff #2244 #10063)
+#10039 := (and #3718 #10036)
+#10042 := (and #9724 #10039)
+#10049 := (or #10048 #10042)
+#10054 := (forall (vars (?x574 T4) (?x575 T5) (?x576 T5)) (:pat #2228) #10049)
+#10064 := (iff #10054 #10063)
+#10061 := (iff #10049 #10060)
+#10058 := (iff #10042 #10057)
+#10059 := [rewrite]: #10058
+#10062 := [monotonicity #10059]: #10061
+#10065 := [quant-intro #10062]: #10064
+#10055 := (iff #2244 #10054)
+#10052 := (iff #2243 #10049)
+#10045 := (implies #9772 #10042)
+#10050 := (iff #10045 #10049)
+#10051 := [rewrite]: #10050
+#10046 := (iff #2243 #10045)
+#10043 := (iff #2242 #10042)
+#10040 := (iff #2241 #10039)
+#10037 := (iff #2240 #10036)
+#10034 := (iff #2239 #10031)
+#10027 := (implies #10021 #10024)
+#10032 := (iff #10027 #10031)
+#10033 := [rewrite]: #10032
+#10028 := (iff #2239 #10027)
+#10025 := (iff #2238 #10024)
+#10026 := [rewrite]: #10025
+#10022 := (iff #2235 #10021)
+#10019 := (iff #2234 #10018)
+#10020 := [rewrite]: #10019
+#10023 := [monotonicity #10020]: #10022
+#10029 := [monotonicity #10023 #10026]: #10028
+#10035 := [trans #10029 #10033]: #10034
+#10038 := [quant-intro #10035]: #10037
+#3719 := (iff #184 #3718)
+#3720 := [rewrite]: #3719
+#10041 := [monotonicity #3720 #10038]: #10040
+#9725 := (iff #2151 #9724)
+#9726 := [rewrite]: #9725
+#10044 := [monotonicity #9726 #10041]: #10043
+#9773 := (iff #2167 #9772)
+#9774 := [rewrite]: #9773
+#10047 := [monotonicity #9774 #10044]: #10046
+#10053 := [trans #10047 #10051]: #10052
+#10056 := [quant-intro #10053]: #10055
+#10067 := [trans #10056 #10065]: #10066
+#10017 := [asserted]: #2244
+#10068 := [mp #10017 #10067]: #10063
+#17092 := [mp~ #10068 #17091]: #10063
+#21044 := [mp #17092 #21043]: #21041
+#25985 := (not #21041)
+#25981 := (or #25985 #25966 #26022)
+#25999 := (or #25998 #25997 #25993)
+#26001 := (forall (vars (?x577 T5)) (:pat #26000) #25999)
+#26002 := (not #26001)
+#26007 := (or #25875 #26006 #26002)
+#26008 := (not #26007)
+#26009 := (or #25966 #26008)
+#25982 := (or #25985 #26009)
+#26075 := (iff #25982 #25981)
+#25987 := (or #25985 #26025)
+#26040 := (iff #25987 #25981)
+#26072 := [rewrite]: #26040
+#25988 := (iff #25982 #25987)
+#26026 := (iff #26009 #26025)
+#26023 := (iff #26008 #26022)
+#26020 := (iff #26007 #26019)
+#26017 := (iff #26002 #26016)
+#26014 := (iff #26001 #26013)
+#26011 := (iff #25999 #26010)
+#26012 := [rewrite]: #26011
+#26015 := [quant-intro #26012]: #26014
+#26018 := [monotonicity #26015]: #26017
+#26021 := [monotonicity #26018]: #26020
+#26024 := [monotonicity #26021]: #26023
+#26027 := [monotonicity #26024]: #26026
+#25989 := [monotonicity #26027]: #25988
+#26076 := [trans #25989 #26072]: #26075
+#25986 := [quant-inst]: #25982
+#26071 := [mp #25986 #26076]: #25981
+#28092 := [unit-resolution #26071 #21044]: #26025
+#28093 := [unit-resolution #28092 #28091]: #26022
+#26074 := (or #26019 #26005)
+#26078 := [def-axiom]: #26074
+#28094 := [unit-resolution #26078 #28093]: #26005
+#28088 := [trans #28094 #28109]: #26480
+#26481 := (not #26480)
+#26516 := (or #11973 #26481)
+#26470 := (uf_66 #26469 0::int #24108)
+#26474 := (uf_24 uf_287 #26470)
+#26475 := (= uf_9 #26474)
+#26476 := (not #26475)
+#26451 := (iff #17634 #26476)
+#26449 := (iff #11973 #26475)
+#26446 := (iff #26475 #11973)
+#26467 := (= #26474 #3047)
+#26465 := (= #26470 #3044)
+#26466 := [monotonicity #27663 #27656]: #26465
+#26468 := [monotonicity #26466]: #26467
+#26448 := [monotonicity #26468]: #26446
+#26450 := [symm #26448]: #26449
+#26452 := [monotonicity #26450]: #26451
+#26464 := [hypothesis]: #17634
+#26447 := [mp #26464 #26452]: #26476
+#26471 := (uf_58 #3175 #26470)
+#26472 := (uf_136 #26471)
+#26473 := (= uf_9 #26472)
+#26486 := (or #26473 #26476)
+#26489 := (not #26486)
+decl uf_22 :: (-> T3 T2)
+#26482 := (uf_22 #24108)
+#26483 := (= uf_9 #26482)
+#2783 := (uf_22 uf_7)
+#27694 := (= #2783 #26482)
+#27691 := (= #26482 #2783)
+#27692 := [monotonicity #27656]: #27691
+#27695 := [symm #27692]: #27694
+#11384 := (= uf_9 #2783)
+#2784 := (= #2783 uf_9)
+#11386 := (iff #2784 #11384)
+#11387 := [rewrite]: #11386
+#11383 := [asserted]: #2784
+#11390 := [mp #11383 #11387]: #11384
+#27696 := [trans #11390 #27695]: #26483
+#26484 := (not #26483)
+#26512 := (or #26484 #26489)
+#27697 := [hypothesis]: #26480
+#26427 := (<= #24110 0::int)
+#26454 := (not #26427)
+#13542 := [and-elim #13524]: #12741
+#26410 := (* -1::int #24110)
+#26519 := (+ uf_286 #26410)
+#26520 := (<= #26519 0::int)
+#27677 := (not #24111)
+#27678 := (or #27677 #26520)
+#27679 := [th-lemma]: #27678
+#27680 := [unit-resolution #27679 #27676]: #26520
+#27686 := (not #26520)
+#26455 := (or #26454 #12740 #27686)
+#26456 := [th-lemma]: #26455
+#26457 := [unit-resolution #26456 #27680 #13542]: #26454
+#237 := (uf_22 #233)
+#762 := (:var 4 int)
+#2069 := (uf_43 #233 #762)
+#2070 := (uf_66 #2069 #247 #233)
+#1373 := (:var 5 T4)
+#2086 := (uf_25 #1373 #2070)
+#1365 := (:var 3 T5)
+#2067 := (uf_16 #1373 #1365)
+#268 := (:var 2 int)
+#2065 := (uf_124 #233 #268)
+#2066 := (uf_43 #2065 #762)
+#2068 := (uf_15 #2066 #2067)
+#2087 := (pattern #2068 #2086 #237)
+#1545 := (uf_59 #1373)
+#2084 := (uf_58 #1545 #2070)
+#2085 := (pattern #2068 #2084 #237)
+#2090 := (uf_136 #2084)
+#9532 := (= uf_9 #2090)
+#2088 := (uf_24 #1373 #2070)
+#9529 := (= uf_9 #2088)
+#20750 := (not #9529)
+#20751 := (or #20750 #9532)
+#20752 := (not #20751)
+#9473 := (= uf_9 #2068)
+#20726 := (not #9473)
+#2073 := (uf_55 #1373)
+#9470 := (= uf_9 #2073)
+#20725 := (not #9470)
+#4045 := (* -1::int #268)
+#6109 := (+ #247 #4045)
+#6706 := (>= #6109 0::int)
+#4307 := (>= #247 0::int)
+#18663 := (not #4307)
+#3926 := (= uf_9 #237)
+#10244 := (not #3926)
+#20758 := (or #10244 #18663 #6706 #20725 #20726 #20752)
+#20763 := (forall (vars (?x526 T4) (?x527 int) (?x528 T5) (?x529 int) (?x530 int) (?x531 T3)) (:pat #2085 #2087) #20758)
+#9535 := (not #9532)
+#9538 := (and #9529 #9535)
+#7773 := (not #6706)
+#9511 := (and #3926 #4307 #7773 #9470 #9473)
+#9516 := (not #9511)
+#9552 := (or #9516 #9538)
+#9555 := (forall (vars (?x526 T4) (?x527 int) (?x528 T5) (?x529 int) (?x530 int) (?x531 T3)) (:pat #2085 #2087) #9552)
+#20764 := (iff #9555 #20763)
+#20761 := (iff #9552 #20758)
+#20727 := (or #10244 #18663 #6706 #20725 #20726)
+#20755 := (or #20727 #20752)
+#20759 := (iff #20755 #20758)
+#20760 := [rewrite]: #20759
+#20756 := (iff #9552 #20755)
+#20753 := (iff #9538 #20752)
+#20754 := [rewrite]: #20753
+#20736 := (iff #9516 #20727)
+#20728 := (not #20727)
+#20731 := (not #20728)
+#20734 := (iff #20731 #20727)
+#20735 := [rewrite]: #20734
+#20732 := (iff #9516 #20731)
+#20729 := (iff #9511 #20728)
+#20730 := [rewrite]: #20729
+#20733 := [monotonicity #20730]: #20732
+#20737 := [trans #20733 #20735]: #20736
+#20757 := [monotonicity #20737 #20754]: #20756
+#20762 := [trans #20757 #20760]: #20761
+#20765 := [quant-intro #20762]: #20764
+#16956 := (~ #9555 #9555)
+#16954 := (~ #9552 #9552)
+#16955 := [refl]: #16954
+#16957 := [nnf-pos #16955]: #16956
+#2091 := (= #2090 uf_9)
+#2092 := (not #2091)
+#2089 := (= #2088 uf_9)
+#2093 := (and #2089 #2092)
+#1434 := (< #247 #268)
+#397 := (<= 0::int #247)
+#1435 := (and #397 #1434)
+#2075 := (= #2068 uf_9)
+#2076 := (and #2075 #1435)
+#238 := (= #237 uf_9)
+#2077 := (and #238 #2076)
+#2074 := (= #2073 uf_9)
+#2078 := (and #2074 #2077)
+#2094 := (implies #2078 #2093)
+#2095 := (forall (vars (?x526 T4) (?x527 int) (?x528 T5) (?x529 int) (?x530 int) (?x531 T3)) (:pat #2085 #2087) #2094)
+#9558 := (iff #2095 #9555)
+#9479 := (and #1435 #9473)
+#9484 := (and #3926 #9479)
+#9487 := (and #9470 #9484)
+#9493 := (not #9487)
+#9544 := (or #9493 #9538)
+#9549 := (forall (vars (?x526 T4) (?x527 int) (?x528 T5) (?x529 int) (?x530 int) (?x531 T3)) (:pat #2085 #2087) #9544)
+#9556 := (iff #9549 #9555)
+#9553 := (iff #9544 #9552)
+#9517 := (iff #9493 #9516)
+#9514 := (iff #9487 #9511)
+#7776 := (and #4307 #7773)
+#9502 := (and #7776 #9473)
+#9505 := (and #3926 #9502)
+#9508 := (and #9470 #9505)
+#9512 := (iff #9508 #9511)
+#9513 := [rewrite]: #9512
+#9509 := (iff #9487 #9508)
+#9506 := (iff #9484 #9505)
+#9503 := (iff #9479 #9502)
+#7777 := (iff #1435 #7776)
+#7774 := (iff #1434 #7773)
+#7775 := [rewrite]: #7774
+#4306 := (iff #397 #4307)
+#4308 := [rewrite]: #4306
+#7778 := [monotonicity #4308 #7775]: #7777
+#9504 := [monotonicity #7778]: #9503
+#9507 := [monotonicity #9504]: #9506
+#9510 := [monotonicity #9507]: #9509
+#9515 := [trans #9510 #9513]: #9514
+#9518 := [monotonicity #9515]: #9517
+#9554 := [monotonicity #9518]: #9553
+#9557 := [quant-intro #9554]: #9556
+#9550 := (iff #2095 #9549)
+#9547 := (iff #2094 #9544)
+#9541 := (implies #9487 #9538)
+#9545 := (iff #9541 #9544)
+#9546 := [rewrite]: #9545
+#9542 := (iff #2094 #9541)
+#9539 := (iff #2093 #9538)
+#9536 := (iff #2092 #9535)
+#9533 := (iff #2091 #9532)
+#9534 := [rewrite]: #9533
+#9537 := [monotonicity #9534]: #9536
+#9530 := (iff #2089 #9529)
+#9531 := [rewrite]: #9530
+#9540 := [monotonicity #9531 #9537]: #9539
+#9488 := (iff #2078 #9487)
+#9485 := (iff #2077 #9484)
+#9482 := (iff #2076 #9479)
+#9476 := (and #9473 #1435)
+#9480 := (iff #9476 #9479)
+#9481 := [rewrite]: #9480
+#9477 := (iff #2076 #9476)
+#9474 := (iff #2075 #9473)
+#9475 := [rewrite]: #9474
+#9478 := [monotonicity #9475]: #9477
+#9483 := [trans #9478 #9481]: #9482
+#3927 := (iff #238 #3926)
+#3928 := [rewrite]: #3927
+#9486 := [monotonicity #3928 #9483]: #9485
+#9471 := (iff #2074 #9470)
+#9472 := [rewrite]: #9471
+#9489 := [monotonicity #9472 #9486]: #9488
+#9543 := [monotonicity #9489 #9540]: #9542
+#9548 := [trans #9543 #9546]: #9547
+#9551 := [quant-intro #9548]: #9550
+#9559 := [trans #9551 #9557]: #9558
+#9528 := [asserted]: #2095
+#9560 := [mp #9528 #9559]: #9555
+#16958 := [mp~ #9560 #16957]: #9555
+#20766 := [mp #16958 #20765]: #20763
+#26500 := (not #20763)
+#26360 := (or #26500 #23948 #26427 #26481 #26484 #26489)
+#26477 := (or #26476 #26473)
+#26478 := (not #26477)
+#26411 := (+ 0::int #26410)
+#26412 := (>= #26411 0::int)
+#26413 := (>= 0::int 0::int)
+#26414 := (not #26413)
+#26485 := (or #26484 #26414 #26412 #23948 #26481 #26478)
+#26361 := (or #26500 #26485)
+#26382 := (iff #26361 #26360)
+#26495 := (or #23948 #26427 #26481 #26484 #26489)
+#26378 := (or #26500 #26495)
+#26380 := (iff #26378 #26360)
+#26381 := [rewrite]: #26380
+#26379 := (iff #26361 #26378)
+#26498 := (iff #26485 #26495)
+#26492 := (or #26484 false #26427 #23948 #26481 #26489)
+#26496 := (iff #26492 #26495)
+#26497 := [rewrite]: #26496
+#26493 := (iff #26485 #26492)
+#26490 := (iff #26478 #26489)
+#26487 := (iff #26477 #26486)
+#26488 := [rewrite]: #26487
+#26491 := [monotonicity #26488]: #26490
+#26430 := (iff #26412 #26427)
+#26424 := (>= #26410 0::int)
+#26428 := (iff #26424 #26427)
+#26429 := [rewrite]: #26428
+#26425 := (iff #26412 #26424)
+#26422 := (= #26411 #26410)
+#26423 := [rewrite]: #26422
+#26426 := [monotonicity #26423]: #26425
+#26431 := [trans #26426 #26429]: #26430
+#26420 := (iff #26414 false)
+#26418 := (iff #26414 #3112)
+#26416 := (iff #26413 true)
+#26417 := [rewrite]: #26416
+#26419 := [monotonicity #26417]: #26418
+#26421 := [trans #26419 #12066]: #26420
+#26494 := [monotonicity #26421 #26431 #26491]: #26493
+#26499 := [trans #26494 #26497]: #26498
+#26377 := [monotonicity #26499]: #26379
+#26383 := [trans #26377 #26381]: #26382
+#26362 := [quant-inst]: #26361
+#26384 := [mp #26362 #26383]: #26360
+#26513 := [unit-resolution #26384 #20766 #13537 #26457 #27697]: #26512
+#26514 := [unit-resolution #26513 #27696]: #26489
+#26460 := (or #26486 #26475)
+#26461 := [def-axiom]: #26460
+#26515 := [unit-resolution #26461 #26514 #26447]: false
+#26517 := [lemma #26515]: #26516
+#28095 := [unit-resolution #26517 #28088]: #11973
+#23927 := (or #22698 #17631 #17634 #22692)
+#23928 := [def-axiom]: #23927
+#28097 := [unit-resolution #23928 #28095 #28084]: #28096
+#28098 := [unit-resolution #28097 #26648]: #22692
+#23917 := (or #22689 #22683)
+#23918 := [def-axiom]: #23917
+#28087 := [unit-resolution #23918 #28098]: #22683
+#26562 := (not #26473)
+#28100 := [unit-resolution #26384 #20766 #13537 #26457 #28088]: #26512
+#28101 := [unit-resolution #28100 #27696]: #26489
+#27944 := (or #26486 #26562)
+#27945 := [def-axiom]: #27944
+#28102 := [unit-resolution #27945 #28101]: #26562
+#28103 := (or #11979 #26473)
+#27942 := (or #11979 #26473 #17634)
+#26742 := (uf_58 #3175 #23935)
+#26748 := (uf_135 #26742)
+#26751 := (uf_25 uf_287 #26748)
+#26752 := (= uf_26 #26751)
+#26749 := (uf_210 uf_287 #26748)
+#26750 := (= uf_9 #26749)
+#27416 := (or #26750 #26752)
+#27420 := (not #27416)
+#26775 := (uf_12 #26369)
+#26938 := (= uf_14 #26775)
+#27415 := (not #26938)
+#26755 := (uf_13 #26748)
+#26756 := (uf_12 #26755)
+#26757 := (= uf_14 #26756)
+#26758 := (uf_27 uf_287 #26748)
+#26753 := (= uf_9 #26758)
+#26743 := (not #26753)
+#26744 := (uf_136 #26742)
+#26745 := (= uf_9 #26744)
+#26746 := (not #26745)
+#26747 := (or #26746 #26743)
+#26774 := (not #26747)
+#27536 := (or #26774 #26757 #27415 #27420)
+#27005 := (not #27536)
+#26817 := (uf_210 uf_287 #23935)
+#26862 := (= uf_9 #26817)
+#26863 := (uf_25 uf_287 #23935)
+#26810 := (= uf_26 #26863)
+#26811 := (or #26810 #26862)
+#26809 := (not #26811)
+#26987 := (or #26809 #26938)
+#27412 := (not #26987)
+#27012 := (or #27412 #27005)
+#27030 := (not #27012)
+#26867 := (uf_24 uf_287 #23935)
+#26988 := (= uf_9 #26867)
+#27804 := (= #3047 #26867)
+#27801 := (= #26867 #3047)
+#27799 := (= #23935 #3044)
+#27793 := (= #26311 #3044)
+#27794 := [symm #27788]: #27793
+#27797 := (= #23935 #26311)
+#27791 := (= #26333 #26311)
+#27792 := [symm #27786]: #27791
+#27798 := [trans #27796 #27792]: #27797
+#27800 := [trans #27798 #27794]: #27799
+#27802 := [monotonicity #27800]: #27801
+#27805 := [symm #27802]: #27804
+#27782 := [hypothesis]: #11973
+#27806 := [trans #27782 #27805]: #26988
+#26989 := (not #26988)
+#27033 := (or #26989 #27030)
+#27035 := (not #27033)
+#26992 := (uf_68 uf_287 #23935)
+#26997 := (= uf_9 #26992)
+#27541 := (iff #26997 #27035)
#2381 := (uf_68 #47 #23)
#2382 := (pattern #2381)
#282 := (uf_59 #47)
#2384 := (uf_58 #282 #23)
#2388 := (uf_135 #2384)
#2399 := (uf_210 #47 #2388)
-#10507 := (= uf_9 #2399)
+#10478 := (= uf_9 #2399)
#2397 := (uf_25 #47 #2388)
-#10504 := (= uf_26 #2397)
-#10510 := (or #10504 #10507)
-#21963 := (not #10510)
+#10475 := (= uf_26 #2397)
+#10481 := (or #10475 #10478)
+#21152 := (not #10481)
#2393 := (uf_13 #2388)
#2394 := (uf_12 #2393)
-#10498 := (= uf_14 #2394)
+#10469 := (= uf_14 #2394)
#2389 := (uf_27 #47 #2388)
-#10489 := (= uf_9 #2389)
-#10492 := (not #10489)
+#10460 := (= uf_9 #2389)
+#10463 := (not #10460)
#2385 := (uf_136 #2384)
-#10483 := (= uf_9 #2385)
-#10486 := (not #10483)
-#10495 := (or #10486 #10492)
-#21962 := (not #10495)
-#26 := (uf_13 #23)
-#27 := (uf_12 #26)
-#29 := (= #27 uf_14)
-#52 := (not #29)
-#21964 := (or #52 #21962 #10498 #21963)
-#21965 := (not #21964)
+#10454 := (= uf_9 #2385)
+#10457 := (not #10454)
+#10466 := (or #10457 #10463)
+#21151 := (not #10466)
+#21153 := (or #52 #21151 #10469 #21152)
+#21154 := (not #21153)
#2405 := (uf_210 #47 #23)
-#10522 := (= uf_9 #2405)
-#142 := (uf_25 #47 #23)
-#3644 := (= uf_26 #142)
-#10525 := (or #3644 #10522)
-#21957 := (not #10525)
-#21958 := (or #29 #21957)
-#21959 := (not #21958)
-#21968 := (or #21959 #21965)
-#21974 := (not #21968)
-#146 := (uf_24 #47 #23)
-#3650 := (= uf_9 #146)
-#11095 := (not #3650)
-#21975 := (or #11095 #21974)
-#21976 := (not #21975)
-#10479 := (= uf_9 #2381)
-#21981 := (iff #10479 #21976)
-#21984 := (forall (vars (?x632 T4) (?x633 T5)) (:pat #2382) #21981)
-#10501 := (not #10498)
-#10543 := (and #29 #10495 #10501 #10510)
-#10528 := (and #52 #10525)
-#10549 := (or #10528 #10543)
-#10554 := (and #3650 #10549)
-#10557 := (iff #10479 #10554)
-#10560 := (forall (vars (?x632 T4) (?x633 T5)) (:pat #2382) #10557)
-#21985 := (iff #10560 #21984)
-#21982 := (iff #10557 #21981)
-#21979 := (iff #10554 #21976)
-#21971 := (and #3650 #21968)
-#21977 := (iff #21971 #21976)
-#21978 := [rewrite]: #21977
-#21972 := (iff #10554 #21971)
-#21969 := (iff #10549 #21968)
-#21966 := (iff #10543 #21965)
-#21967 := [rewrite]: #21966
-#21960 := (iff #10528 #21959)
-#21961 := [rewrite]: #21960
-#21970 := [monotonicity #21961 #21967]: #21969
-#21973 := [monotonicity #21970]: #21972
-#21980 := [trans #21973 #21978]: #21979
-#21983 := [monotonicity #21980]: #21982
-#21986 := [quant-intro #21983]: #21985
-#17907 := (~ #10560 #10560)
-#17905 := (~ #10557 #10557)
-#17906 := [refl]: #17905
-#17908 := [nnf-pos #17906]: #17907
+#10493 := (= uf_9 #2405)
+#10496 := (or #3615 #10493)
+#21146 := (not #10496)
+#21147 := (or #29 #21146)
+#21148 := (not #21147)
+#21157 := (or #21148 #21154)
+#21163 := (not #21157)
+#21164 := (or #11066 #21163)
+#21165 := (not #21164)
+#10450 := (= uf_9 #2381)
+#21170 := (iff #10450 #21165)
+#21173 := (forall (vars (?x632 T4) (?x633 T5)) (:pat #2382) #21170)
+#10472 := (not #10469)
+#10514 := (and #29 #10466 #10472 #10481)
+#10499 := (and #52 #10496)
+#10520 := (or #10499 #10514)
+#10525 := (and #3621 #10520)
+#10528 := (iff #10450 #10525)
+#10531 := (forall (vars (?x632 T4) (?x633 T5)) (:pat #2382) #10528)
+#21174 := (iff #10531 #21173)
+#21171 := (iff #10528 #21170)
+#21168 := (iff #10525 #21165)
+#21160 := (and #3621 #21157)
+#21166 := (iff #21160 #21165)
+#21167 := [rewrite]: #21166
+#21161 := (iff #10525 #21160)
+#21158 := (iff #10520 #21157)
+#21155 := (iff #10514 #21154)
+#21156 := [rewrite]: #21155
+#21149 := (iff #10499 #21148)
+#21150 := [rewrite]: #21149
+#21159 := [monotonicity #21150 #21156]: #21158
+#21162 := [monotonicity #21159]: #21161
+#21169 := [trans #21162 #21167]: #21168
+#21172 := [monotonicity #21169]: #21171
+#21175 := [quant-intro #21172]: #21174
+#17195 := (~ #10531 #10531)
+#17193 := (~ #10528 #10528)
+#17194 := [refl]: #17193
+#17196 := [nnf-pos #17194]: #17195
#2406 := (= #2405 uf_9)
-#143 := (= #142 uf_26)
#2407 := (or #143 #2406)
#2408 := (and #52 #2407)
#2400 := (= #2399 uf_9)
@@ -326,6060 +5406,253 @@
#2403 := (and #2392 #2402)
#2404 := (and #29 #2403)
#2409 := (or #2404 #2408)
-#147 := (= #146 uf_9)
#2410 := (and #147 #2409)
#2383 := (= #2381 uf_9)
#2411 := (iff #2383 #2410)
#2412 := (forall (vars (?x632 T4) (?x633 T5)) (:pat #2382) #2411)
-#10563 := (iff #2412 #10560)
-#10513 := (and #10501 #10510)
-#10516 := (and #10495 #10513)
-#10519 := (and #29 #10516)
-#10531 := (or #10519 #10528)
-#10534 := (and #3650 #10531)
-#10537 := (iff #10479 #10534)
-#10540 := (forall (vars (?x632 T4) (?x633 T5)) (:pat #2382) #10537)
-#10561 := (iff #10540 #10560)
-#10558 := (iff #10537 #10557)
-#10555 := (iff #10534 #10554)
-#10552 := (iff #10531 #10549)
-#10546 := (or #10543 #10528)
-#10550 := (iff #10546 #10549)
-#10551 := [rewrite]: #10550
-#10547 := (iff #10531 #10546)
-#10544 := (iff #10519 #10543)
-#10545 := [rewrite]: #10544
-#10548 := [monotonicity #10545]: #10547
-#10553 := [trans #10548 #10551]: #10552
-#10556 := [monotonicity #10553]: #10555
-#10559 := [monotonicity #10556]: #10558
-#10562 := [quant-intro #10559]: #10561
-#10541 := (iff #2412 #10540)
-#10538 := (iff #2411 #10537)
-#10535 := (iff #2410 #10534)
-#10532 := (iff #2409 #10531)
-#10529 := (iff #2408 #10528)
-#10526 := (iff #2407 #10525)
-#10523 := (iff #2406 #10522)
-#10524 := [rewrite]: #10523
-#3645 := (iff #143 #3644)
-#3646 := [rewrite]: #3645
-#10527 := [monotonicity #3646 #10524]: #10526
+#10534 := (iff #2412 #10531)
+#10484 := (and #10472 #10481)
+#10487 := (and #10466 #10484)
+#10490 := (and #29 #10487)
+#10502 := (or #10490 #10499)
+#10505 := (and #3621 #10502)
+#10508 := (iff #10450 #10505)
+#10511 := (forall (vars (?x632 T4) (?x633 T5)) (:pat #2382) #10508)
+#10532 := (iff #10511 #10531)
+#10529 := (iff #10508 #10528)
+#10526 := (iff #10505 #10525)
+#10523 := (iff #10502 #10520)
+#10517 := (or #10514 #10499)
+#10521 := (iff #10517 #10520)
+#10522 := [rewrite]: #10521
+#10518 := (iff #10502 #10517)
+#10515 := (iff #10490 #10514)
+#10516 := [rewrite]: #10515
+#10519 := [monotonicity #10516]: #10518
+#10524 := [trans #10519 #10522]: #10523
+#10527 := [monotonicity #10524]: #10526
#10530 := [monotonicity #10527]: #10529
-#10520 := (iff #2404 #10519)
-#10517 := (iff #2403 #10516)
-#10514 := (iff #2402 #10513)
-#10511 := (iff #2401 #10510)
-#10508 := (iff #2400 #10507)
-#10509 := [rewrite]: #10508
-#10505 := (iff #2398 #10504)
-#10506 := [rewrite]: #10505
-#10512 := [monotonicity #10506 #10509]: #10511
-#10502 := (iff #2396 #10501)
-#10499 := (iff #2395 #10498)
-#10500 := [rewrite]: #10499
-#10503 := [monotonicity #10500]: #10502
-#10515 := [monotonicity #10503 #10512]: #10514
-#10496 := (iff #2392 #10495)
-#10493 := (iff #2391 #10492)
-#10490 := (iff #2390 #10489)
-#10491 := [rewrite]: #10490
-#10494 := [monotonicity #10491]: #10493
-#10487 := (iff #2387 #10486)
-#10484 := (iff #2386 #10483)
-#10485 := [rewrite]: #10484
-#10488 := [monotonicity #10485]: #10487
-#10497 := [monotonicity #10488 #10494]: #10496
-#10518 := [monotonicity #10497 #10515]: #10517
-#10521 := [monotonicity #10518]: #10520
-#10533 := [monotonicity #10521 #10530]: #10532
-#3651 := (iff #147 #3650)
-#3652 := [rewrite]: #3651
-#10536 := [monotonicity #3652 #10533]: #10535
-#10481 := (iff #2383 #10479)
-#10482 := [rewrite]: #10481
-#10539 := [monotonicity #10482 #10536]: #10538
-#10542 := [quant-intro #10539]: #10541
-#10564 := [trans #10542 #10562]: #10563
-#10478 := [asserted]: #2412
-#10565 := [mp #10478 #10564]: #10560
-#17909 := [mp~ #10565 #17908]: #10560
-#21987 := [mp #17909 #21986]: #21984
-#27470 := (not #21984)
-#30070 := (or #27470 #29389)
-#29305 := (or #29304 #29302)
-#29306 := (not #29305)
-#29316 := (or #29315 #29312)
-#29317 := (not #29316)
-#29319 := (= #29318 uf_14)
-#29320 := (not #29319)
-#29321 := (or #29320 #29317 #29309 #29306)
-#29322 := (not #29321)
-#29327 := (or #29326 #29324)
-#29328 := (not #29327)
-#29329 := (or #29319 #29328)
-#29330 := (not #29329)
-#29331 := (or #29330 #29322)
-#29332 := (not #29331)
-#29333 := (or #18452 #29332)
-#29334 := (not #29333)
-#29335 := (iff #12378 #29334)
-#30071 := (or #27470 #29335)
-#30073 := (iff #30071 #30070)
-#29901 := (iff #30070 #30070)
-#29902 := [rewrite]: #29901
-#29390 := (iff #29335 #29389)
-#29387 := (iff #29334 #29386)
-#29384 := (iff #29333 #29383)
-#29381 := (iff #29332 #29380)
-#29378 := (iff #29331 #29377)
-#29375 := (iff #29322 #29374)
-#29372 := (iff #29321 #29369)
-#29366 := (or #29351 #29357 #29309 #29363)
-#29370 := (iff #29366 #29369)
-#29371 := [rewrite]: #29370
-#29367 := (iff #29321 #29366)
-#29364 := (iff #29306 #29363)
-#29361 := (iff #29305 #29360)
-#29362 := [rewrite]: #29361
-#29365 := [monotonicity #29362]: #29364
-#29358 := (iff #29317 #29357)
-#29355 := (iff #29316 #29354)
-#29356 := [rewrite]: #29355
-#29359 := [monotonicity #29356]: #29358
-#29352 := (iff #29320 #29351)
-#29337 := (iff #29319 #29336)
-#29338 := [rewrite]: #29337
-#29353 := [monotonicity #29338]: #29352
-#29368 := [monotonicity #29353 #29359 #29365]: #29367
-#29373 := [trans #29368 #29371]: #29372
-#29376 := [monotonicity #29373]: #29375
-#29349 := (iff #29330 #29348)
-#29346 := (iff #29329 #29345)
-#29343 := (iff #29328 #29342)
-#29340 := (iff #29327 #29339)
-#29341 := [rewrite]: #29340
-#29344 := [monotonicity #29341]: #29343
-#29347 := [monotonicity #29338 #29344]: #29346
-#29350 := [monotonicity #29347]: #29349
-#29379 := [monotonicity #29350 #29376]: #29378
-#29382 := [monotonicity #29379]: #29381
-#29385 := [monotonicity #29382]: #29384
-#29388 := [monotonicity #29385]: #29387
-#29391 := [monotonicity #29388]: #29390
-#29900 := [monotonicity #29391]: #30073
-#29885 := [trans #29900 #29902]: #30073
-#30072 := [quant-inst]: #30071
-#29886 := [mp #30072 #29885]: #30070
-#30190 := [unit-resolution #29886 #21987 #30156]: false
-#30191 := [lemma #30190]: #29389
-#30791 := (or #29961 #12378)
-#30058 := [hypothesis]: #29369
-decl uf_116 :: (-> T5 int)
-#2980 := (uf_116 #2979)
-decl uf_124 :: (-> T3 int T3)
-#2977 := (uf_124 uf_7 uf_286)
-#2981 := (uf_43 #2977 #2980)
-#2984 := (uf_25 uf_287 #2981)
-#30109 := (= #2984 #29303)
-#30110 := (= #29303 #2984)
-#30604 := (= #29300 #2981)
-decl uf_143 :: (-> T3 int)
-#24856 := (uf_143 #2977)
-decl uf_144 :: (-> T3 T3)
-#24854 := (uf_144 #2977)
-#27136 := (uf_124 #24854 #24856)
-#27137 := (uf_43 #27136 #2980)
-#28502 := (= #27137 #2981)
-#28500 := (= #27136 #2977)
-#28498 := (= #24856 uf_286)
-#24857 := (= uf_286 #24856)
-#326 := (:var 1 T3)
-#1358 := (uf_124 #326 #161)
-#1592 := (pattern #1358)
-#1602 := (uf_143 #1358)
-#8293 := (= #161 #1602)
-#8296 := (forall (vars (?x386 T3) (?x387 int)) (:pat #1592) #8293)
-#17283 := (~ #8296 #8296)
-#17281 := (~ #8293 #8293)
-#17282 := [refl]: #17281
-#17284 := [nnf-pos #17282]: #17283
-#1603 := (= #1602 #161)
-#1604 := (forall (vars (?x386 T3) (?x387 int)) (:pat #1592) #1603)
-#8297 := (iff #1604 #8296)
-#8294 := (iff #1603 #8293)
-#8295 := [rewrite]: #8294
-#8298 := [quant-intro #8295]: #8297
-#8292 := [asserted]: #1604
-#8301 := [mp #8292 #8298]: #8296
-#17285 := [mp~ #8301 #17284]: #8296
-#24863 := (not #8296)
-#24864 := (or #24863 #24857)
-#24865 := [quant-inst]: #24864
-#28421 := [unit-resolution #24865 #17285]: #24857
-#28499 := [symm #28421]: #28498
-#28400 := (= #24854 uf_7)
-#24855 := (= uf_7 #24854)
-#1605 := (uf_144 #1358)
-#8300 := (= #326 #1605)
-#8304 := (forall (vars (?x388 T3) (?x389 int)) (:pat #1592) #8300)
-#17288 := (~ #8304 #8304)
-#17286 := (~ #8300 #8300)
-#17287 := [refl]: #17286
-#17289 := [nnf-pos #17287]: #17288
-#1606 := (= #1605 #326)
-#1607 := (forall (vars (?x388 T3) (?x389 int)) (:pat #1592) #1606)
-#8305 := (iff #1607 #8304)
-#8302 := (iff #1606 #8300)
-#8303 := [rewrite]: #8302
-#8306 := [quant-intro #8303]: #8305
-#8299 := [asserted]: #1607
-#8309 := [mp #8299 #8306]: #8304
-#17290 := [mp~ #8309 #17289]: #8304
-#24858 := (not #8304)
-#24859 := (or #24858 #24855)
-#24860 := [quant-inst]: #24859
-#28399 := [unit-resolution #24860 #17290]: #24855
-#28401 := [symm #28399]: #28400
-#28501 := [monotonicity #28401 #28499]: #28500
-#28503 := [monotonicity #28501]: #28502
-#30602 := (= #29300 #27137)
-decl uf_125 :: (-> T5 T5 int)
-decl uf_28 :: (-> int T5)
-decl uf_29 :: (-> T5 int)
-#3013 := (uf_29 #2979)
-#24681 := (uf_28 #3013)
-#27114 := (uf_13 #24681)
-#29133 := (uf_66 #24681 uf_298 #27114)
-#29134 := (uf_125 #29133 #24681)
-#29158 := (uf_66 #27137 #29134 #24854)
-#29162 := (uf_58 #3157 #29158)
-#29165 := (uf_135 #29162)
-#30601 := (= #29165 #27137)
-#29166 := (= #27137 #29165)
-decl up_67 :: (-> T14 bool)
-#29168 := (up_67 #29162)
-#29169 := (not #29168)
-#29167 := (not #29166)
-#29163 := (uf_136 #29162)
-#29164 := (= uf_9 #29163)
-#29159 := (uf_24 uf_287 #29158)
-#29160 := (= uf_9 #29159)
-#29161 := (not #29160)
-#29190 := (or #29161 #29164 #29167 #29169)
-#29193 := (not #29190)
-#29183 := (* -1::int #29134)
-#29184 := (+ #24856 #29183)
-#29185 := (<= #29184 0::int)
-#30532 := (not #29185)
-#29293 := (+ uf_298 #29183)
-#29295 := (>= #29293 0::int)
-#29135 := (= uf_298 #29134)
-#15 := (:var 1 T5)
-#1390 := (uf_13 #15)
-#1391 := (uf_66 #15 #161 #1390)
-#1392 := (pattern #1391)
-#1393 := (uf_125 #1391 #15)
-#7692 := (= #161 #1393)
-#7696 := (forall (vars (?x319 T5) (?x320 int)) (:pat #1392) #7692)
-#17000 := (~ #7696 #7696)
-#16998 := (~ #7692 #7692)
-#16999 := [refl]: #16998
-#17001 := [nnf-pos #16999]: #17000
-#1394 := (= #1393 #161)
-#1395 := (forall (vars (?x319 T5) (?x320 int)) (:pat #1392) #1394)
-#7697 := (iff #1395 #7696)
-#7694 := (iff #1394 #7692)
-#7695 := [rewrite]: #7694
-#7698 := [quant-intro #7695]: #7697
-#7691 := [asserted]: #1395
-#7701 := [mp #7691 #7698]: #7696
-#17002 := [mp~ #7701 #17001]: #7696
-#27121 := (not #7696)
-#29138 := (or #27121 #29135)
-#29139 := [quant-inst]: #29138
-#29491 := [unit-resolution #29139 #17002]: #29135
-#29492 := (not #29135)
-#30531 := (or #29492 #29295)
-#30526 := [th-lemma]: #30531
-#30527 := [unit-resolution #30526 #29491]: #29295
-#13457 := (* -1::int uf_298)
-#13720 := (+ uf_286 #13457)
-#13721 := (<= #13720 0::int)
-#13722 := (not #13721)
-#23320 := (or #18449 #18458 #23317)
-#23323 := (not #23320)
-#23326 := (or #18449 #18452 #23323)
-#23329 := (not #23326)
-#23332 := (or #18449 #18452 #23329)
-#23335 := (not #23332)
-#23338 := (or #22372 #13576 #13721 #23335)
-#23341 := (not #23338)
-decl ?x776!15 :: int
-#18631 := ?x776!15
-#18636 := (uf_66 #2979 ?x776!15 uf_7)
-#18637 := (uf_110 uf_287 #18636)
-#18982 := (* -1::int #18637)
-decl uf_302 :: int
-#3096 := uf_302
-#18983 := (+ uf_302 #18982)
-#18984 := (>= #18983 0::int)
-#18969 := (* -1::int ?x776!15)
-#18970 := (+ uf_286 #18969)
-#18971 := (<= #18970 0::int)
-#18633 := (>= ?x776!15 0::int)
-#22466 := (not #18633)
-#18632 := (<= ?x776!15 4294967295::int)
-#22465 := (not #18632)
-#22481 := (or #22465 #22466 #18971 #18984)
-#22486 := (not #22481)
-#13362 := (* -1::int uf_286)
-#13750 := (+ #161 #13362)
-#13749 := (>= #13750 0::int)
-#3103 := (= #3054 uf_302)
-#22439 := (not #3103)
-#22440 := (or #22439 #4992 #13749 #19506)
-#23352 := (forall (vars (?x778 int)) (:pat #23211) #22440)
-#23357 := (not #23352)
-#13761 := (* -1::int uf_302)
-#13762 := (+ #3054 #13761)
-#13763 := (<= #13762 0::int)
-#22431 := (or #4992 #13749 #13763 #19506)
-#23344 := (forall (vars (?x776 int)) (:pat #23211) #22431)
-#23349 := (not #23344)
-#23360 := (or #23349 #23357)
-#23363 := (not #23360)
-#23366 := (or #23363 #22486)
-#23369 := (not #23366)
-#12044 := (= uf_296 uf_302)
-#12093 := (not #12044)
-decl uf_301 :: int
-#3094 := uf_301
-#12041 := (= uf_297 uf_301)
-#12102 := (not #12041)
-decl uf_300 :: int
-#3092 := uf_300
-#12038 := (= uf_298 uf_300)
-#12111 := (not #12038)
-decl uf_299 :: int
-#3090 := uf_299
-#12035 := (= uf_296 uf_299)
-#12120 := (not #12035)
-#23372 := (or #12120 #12111 #12102 #12093 #22372 #13576 #13722 #23369)
-#23375 := (not #23372)
-#23378 := (or #23341 #23375)
-#23381 := (not #23378)
-#13926 := (* -1::int uf_296)
-#13927 := (+ #3054 #13926)
-#13928 := (<= #13927 0::int)
-#13915 := (+ #161 #13457)
-#13914 := (>= #13915 0::int)
-#22247 := (or #4992 #13914 #13928 #19506)
-#23220 := (forall (vars (?x775 int)) (:pat #23211) #22247)
-#23225 := (not #23220)
-#1331 := 255::int
-#15805 := (<= uf_296 255::int)
-#22516 := (not #15805)
-#15788 := (<= uf_297 4294967295::int)
-#22515 := (not #15788)
-#15771 := (<= uf_298 4294967295::int)
-#22514 := (not #15771)
-#13976 := (>= uf_296 0::int)
-#22512 := (not #13976)
-#13954 := (>= uf_298 0::int)
-#22511 := (not #13954)
-#13948 := (>= #13720 0::int)
-#13951 := (not #13948)
-#13897 := (* -1::int uf_297)
-#13898 := (+ uf_286 #13897)
-#13899 := (<= #13898 0::int)
-#13428 := (<= uf_286 0::int)
-decl uf_178 :: (-> T4 T4 T2)
-#3169 := (uf_178 uf_287 uf_287)
-#12330 := (= uf_9 #3169)
-#13894 := (not #12330)
-decl uf_202 :: (-> T1 T4 T2)
-decl uf_272 :: T1
-#2945 := uf_272
-#3087 := (uf_202 uf_272 uf_287)
-#12024 := (= uf_9 #3087)
-#15207 := (not #12024)
-#3082 := (uf_66 #2979 uf_297 uf_7)
-#3083 := (uf_110 uf_287 #3082)
-#12016 := (= uf_296 #3083)
-#22510 := (not #12016)
-decl up_292 :: (-> T4 T1 T1 T5 T3 bool)
-decl uf_6 :: (-> T3 T3)
-#11 := (uf_6 uf_7)
-decl uf_280 :: T1
-#2953 := uf_280
-#3182 := (up_292 uf_287 uf_272 uf_280 #2979 #11)
-#22509 := (not #3182)
-#3181 := (up_291 uf_287 uf_272 uf_280 #3013 #11)
-#22508 := (not #3181)
-decl uf_279 :: T1
-#2952 := uf_279
-#3180 := (up_291 uf_287 uf_272 uf_279 uf_286 uf_4)
-#12911 := (not #3180)
-#3179 := (up_291 uf_287 uf_272 uf_277 uf_296 uf_7)
-#12920 := (not #3179)
-#3178 := (up_291 uf_287 uf_272 uf_275 uf_297 uf_4)
-#12929 := (not #3178)
-#3177 := (up_291 uf_287 uf_272 uf_273 uf_298 uf_4)
-#12938 := (not #3177)
-#3031 := (uf_66 #2979 0::int uf_7)
-#3041 := (uf_110 uf_287 #3031)
-decl uf_295 :: int
-#3040 := uf_295
-#3042 := (= uf_295 #3041)
-#13173 := (not #3042)
-#23384 := (or #13173 #12938 #12929 #12920 #12911 #22508 #22509 #22510 #15207 #13894 #13428 #22372 #13576 #13899 #13951 #22511 #22512 #22514 #22515 #22516 #23225 #23381)
-#23387 := (not #23384)
-#23390 := (or #13173 #13428 #23387)
-#23393 := (not #23390)
-#13408 := (* -1::int #3054)
-#13409 := (+ uf_295 #13408)
-#13407 := (>= #13409 0::int)
-#13395 := (>= #161 1::int)
-#22236 := (or #4992 #13395 #13407 #19506)
-#23212 := (forall (vars (?x773 int)) (:pat #23211) #22236)
-#23217 := (not #23212)
-#23396 := (or #23217 #23393)
-#23399 := (not #23396)
-decl ?x773!13 :: int
-#18370 := ?x773!13
-#18380 := (>= ?x773!13 1::int)
-#18375 := (uf_66 #2979 ?x773!13 uf_7)
-#18376 := (uf_110 uf_287 #18375)
-#18377 := (* -1::int #18376)
-#18378 := (+ uf_295 #18377)
-#18379 := (>= #18378 0::int)
-#18372 := (>= ?x773!13 0::int)
-#22210 := (not #18372)
-#18371 := (<= ?x773!13 4294967295::int)
-#22209 := (not #18371)
-#22225 := (or #22209 #22210 #18379 #18380)
-#22230 := (not #22225)
-#23402 := (or #22230 #23399)
-#23405 := (not #23402)
-#13389 := (>= uf_286 1::int)
-#13392 := (not #13389)
-#23408 := (or #13392 #23405)
-#23411 := (not #23408)
-#23414 := (or #13392 #23411)
-#23417 := (not #23414)
-#3037 := (uf_68 uf_287 #3031)
-#11964 := (= uf_9 #3037)
-#18347 := (not #11964)
-#3032 := (uf_48 #3031 uf_7)
-#11955 := (= uf_9 #3032)
-#18338 := (not #11955)
-decl uf_274 :: T1
-#2947 := uf_274
-#3045 := (up_291 uf_287 uf_274 uf_273 1::int uf_4)
-#13146 := (not #3045)
-decl uf_276 :: T1
-#2949 := uf_276
-#3044 := (up_291 uf_287 uf_276 uf_275 0::int uf_4)
-#13155 := (not #3044)
-decl uf_278 :: T1
-#2951 := uf_278
-#3043 := (up_291 uf_287 uf_278 uf_277 uf_295 uf_7)
-#13164 := (not #3043)
-#23420 := (or #13173 #13164 #13155 #13146 #18338 #18347 #23417)
-#23423 := (not #23420)
-#23426 := (or #18338 #18347 #23423)
-#23429 := (not #23426)
-#3034 := (uf_24 uf_287 #3031)
-#11958 := (= uf_9 #3034)
-#18341 := (not #11958)
-#23432 := (or #18338 #18341 #23429)
-#23435 := (not #23432)
-#23438 := (or #18338 #18341 #23435)
-#23441 := (not #23438)
-decl uf_200 :: (-> T4 T5 T5 T16 T2)
-decl uf_282 :: T16
-#2957 := uf_282
-#3029 := (uf_200 uf_287 #2981 #2981 uf_282)
-#11952 := (= uf_9 #3029)
-#13206 := (not #11952)
-#23444 := (or #13206 #23441)
-#23447 := (not #23444)
-#24756 := (uf_116 #2981)
-#26144 := (uf_43 #2977 #24756)
-#26691 := (uf_200 uf_287 #26144 #26144 uf_282)
-#26936 := (= #26691 #3029)
-#26939 := (= #3029 #26691)
-#26151 := (= #2981 #26144)
-#2986 := (uf_48 #2981 #2977)
-#11896 := (= uf_9 #2986)
-decl uf_23 :: (-> T3 T2)
-#2993 := (uf_23 #2977)
-#11908 := (= uf_9 #2993)
-#2990 := (uf_12 #2977)
-#11902 := (= uf_14 #2990)
-#11905 := (not #11902)
-#2988 := (uf_24 uf_287 #2981)
-#11899 := (= uf_9 #2988)
-#11893 := (= uf_26 #2984)
-#2982 := (uf_27 uf_287 #2981)
-#11889 := (= uf_9 #2982)
-#14148 := (and #11889 #11893 #11896 #11899 #11905 #11908)
-decl uf_283 :: int
-#2961 := uf_283
-#14204 := (* -1::int uf_283)
-decl uf_78 :: int
-#429 := uf_78
-#14205 := (+ uf_78 #14204)
-#14203 := (>= #14205 0::int)
-#14201 := (>= uf_283 0::int)
-#14208 := (and #14201 #14203)
-#14211 := (not #14208)
-decl uf_284 :: int
-#2965 := uf_284
-#14190 := (* -1::int uf_284)
-decl uf_76 :: int
-#409 := uf_76
-#14191 := (+ uf_76 #14190)
-#14189 := (>= #14191 0::int)
-#14187 := (>= uf_284 0::int)
-#14194 := (and #14187 #14189)
-#14197 := (not #14194)
-decl uf_285 :: int
-#2969 := uf_285
-#14176 := (* -1::int uf_285)
-#14177 := (+ uf_76 #14176)
-#14175 := (>= #14177 0::int)
-#14173 := (>= uf_285 0::int)
-#14180 := (and #14173 #14175)
-#14183 := (not #14180)
-#1042 := 1099511627776::int
-#14161 := (>= uf_286 1099511627776::int)
-#14151 := (not #14148)
-decl uf_289 :: (-> T19 int)
-#3007 := (:var 0 T19)
-#3008 := (uf_289 #3007)
-#3009 := (pattern #3008)
-decl uf_290 :: int
-#3010 := uf_290
-#14138 := (* -1::int uf_290)
-#14139 := (+ #3008 #14138)
-#14137 := (>= #14139 0::int)
-#14136 := (not #14137)
-#14142 := (forall (vars (?x771 T19)) (:pat #3009) #14136)
-#14145 := (not #14142)
-#13429 := (not #13428)
-#13992 := (and #3042 #13429)
-#13997 := (not #13992)
-#13980 := (+ uf_78 #13926)
-#13979 := (>= #13980 0::int)
-#13983 := (and #13976 #13979)
-#13986 := (not #13983)
-#13967 := (+ uf_76 #13897)
-#13966 := (>= #13967 0::int)
-#13970 := (and #13433 #13966)
-#13973 := (not #13970)
-#13458 := (+ uf_76 #13457)
-#13957 := (>= #13458 0::int)
-#13960 := (and #13954 #13957)
-#13963 := (not #13960)
-#4382 := (* -1::int uf_76)
-#4383 := (+ #161 #4382)
-#4384 := (<= #4383 0::int)
-#4391 := (and #4070 #4384)
-#5606 := (not #4391)
-#13937 := (or #5606 #13914 #13928)
-#13942 := (forall (vars (?x775 int)) #13937)
-#13945 := (not #13942)
-#13900 := (not #13899)
-#13906 := (and #12016 #13900)
-#13911 := (not #13906)
-#13751 := (not #13749)
-#13789 := (and #3103 #4070 #4384 #13751)
-#13794 := (exists (vars (?x778 int)) #13789)
-#13772 := (or #5606 #13749 #13763)
-#13777 := (forall (vars (?x776 int)) #13772)
-#13797 := (not #13777)
-#13803 := (or #13797 #13794)
-#13808 := (and #13777 #13803)
-#13438 := (and #13433 #13436)
-#13441 := (not #13438)
-decl up_216 :: bool
-#2482 := up_216
-#12168 := (not up_216)
-#13835 := (or #12168 #12120 #12111 #12102 #12093 #13441 #13808)
-#13840 := (and up_216 #13835)
-#13865 := (or #13441 #13722 #13840)
-#13456 := (>= #13458 1::int)
-#13545 := (and #13454 #13456)
-#13548 := (not #13545)
-#13515 := (not #13514)
-#13521 := (and #12428 #13515)
-#13501 := (or #5606 #13478 #13492)
-#13506 := (forall (vars (?x786 int)) #13501)
-#13509 := (not #13506)
-#13526 := (or #13509 #13521)
-#13529 := (and #13506 #13526)
-#13532 := (or #13475 #13529)
-#13535 := (and #13470 #13532)
-#13464 := (and #13445 #13462)
-#13467 := (not #13464)
-#13563 := (or #12493 #13467 #13535 #13542 #13548)
-#13571 := (and #13454 #13456 #13563)
-#13447 := (and #13436 #13445)
-#13450 := (not #13447)
-#13682 := (or #12686 #12677 #13441 #13450 #13571 #13627)
-#12381 := (and #12369 #12378)
-#12600 := (not #12381)
-#13603 := (or #12582 #12573 #12600 #12591 #12548 #12539 #13576 #13450 #13571)
-#13611 := (and #12369 #12378 #13603)
-#12375 := (and #12369 #12372)
-#12612 := (not #12375)
-#13616 := (or #12612 #13611)
-#13622 := (and #12369 #12372 #13616)
-#13652 := (or #13441 #13622 #13628)
-#13687 := (and #13652 #13682)
-#13696 := (or #12600 #13441 #13687)
-#13704 := (and #12369 #12378 #13696)
-#13709 := (or #12612 #13704)
-#13715 := (and #12369 #12372 #13709)
-#13744 := (or #13441 #13715 #13721)
-#13870 := (and #13744 #13865)
-decl uf_55 :: (-> T4 T2)
-#3004 := (uf_55 uf_287)
-#11932 := (= uf_9 #3004)
-#12030 := (and #11932 #12024)
-#12203 := (not #12030)
-#3183 := (and #3181 #3182)
-#12902 := (not #3183)
-#14063 := (or #12938 #12929 #12920 #12911 #12902 #12203 #13894 #13441 #13870 #13911 #13945 #13951 #13963 #13973 #13986 #13997)
-#14071 := (and #3042 #13429 #14063)
-#13417 := (or #5606 #13395 #13407)
-#13422 := (forall (vars (?x773 int)) #13417)
-#13425 := (not #13422)
-#14076 := (or #13425 #14071)
-#14079 := (and #13422 #14076)
-#14082 := (or #13392 #14079)
-#14085 := (and #13389 #14082)
-#11967 := (and #11955 #11964)
-#13182 := (not #11967)
-#14106 := (or #13173 #13164 #13155 #13146 #13182 #14085)
-#14114 := (and #11955 #11964 #14106)
-#11961 := (and #11955 #11958)
-#13194 := (not #11961)
-#14119 := (or #13194 #14114)
-#14125 := (and #11955 #11958 #14119)
-#14130 := (or #13206 #14125)
-#14133 := (and #11952 #14130)
-#13363 := (+ uf_76 #13362)
-#13361 := (>= #13363 0::int)
-#13359 := (>= uf_286 0::int)
-#13366 := (and #13359 #13361)
-#13369 := (not #13366)
-decl uf_294 :: (-> int T5 T2)
-decl uf_293 :: int
-#3018 := uf_293
-#3021 := (uf_294 uf_293 #23)
-#3022 := (pattern #3021)
-#11938 := (= uf_9 #3021)
-#11944 := (not #11938)
-#11949 := (forall (vars (?x772 T5)) (:pat #3022) #11944)
-#13227 := (not #11949)
-decl uf_281 :: T1
-#2954 := uf_281
-#3002 := (uf_202 uf_281 uf_287)
-#11929 := (= uf_9 #3002)
-#11935 := (and #11929 #11932)
-#13272 := (not #11935)
-decl uf_203 :: (-> T4 T2)
-#3000 := (uf_203 uf_287)
-#11926 := (= uf_9 #3000)
-#13281 := (not #11926)
-decl uf_173 :: (-> T4 int)
-#3019 := (uf_173 uf_287)
-#3020 := (= uf_293 #3019)
-#13236 := (not #3020)
-#3017 := (up_291 uf_287 uf_281 uf_279 uf_286 uf_4)
-#13245 := (not #3017)
-#3015 := (up_292 uf_287 uf_281 uf_280 #2979 #11)
-#3014 := (up_291 uf_287 uf_281 uf_280 #3013 #11)
-#3016 := (and #3014 #3015)
-#13254 := (not #3016)
-#14256 := (or #13254 #13245 #13236 #13281 #13272 #13227 #13369 #13428 #14133 #14145 #14151 #14161 #14183 #14197 #14211)
-#14261 := (not #14256)
-#1 := true
-#3098 := (< #161 uf_286)
-#3104 := (and #3098 #3103)
-#411 := (<= #161 uf_76)
-#3105 := (and #411 #3104)
-#285 := (<= 0::int #161)
-#3106 := (and #285 #3105)
-#3107 := (exists (vars (?x778 int)) #3106)
-#3108 := (implies #3107 true)
-#3109 := (and #3107 #3108)
-#3099 := (<= #3054 uf_302)
-#3100 := (implies #3098 #3099)
-#412 := (and #285 #411)
-#3101 := (implies #412 #3100)
-#3102 := (forall (vars (?x776 int)) #3101)
-#3110 := (implies #3102 #3109)
-#3111 := (and #3102 #3110)
-#3112 := (implies true #3111)
-#3097 := (= uf_302 uf_296)
-#3113 := (implies #3097 #3112)
-#3095 := (= uf_301 uf_297)
-#3114 := (implies #3095 #3113)
-#3093 := (= uf_300 uf_298)
-#3115 := (implies #3093 #3114)
-#3091 := (= uf_299 uf_296)
-#3116 := (implies #3091 #3115)
-#3117 := (implies true #3116)
-#3066 := (<= 0::int uf_297)
-#3073 := (<= 1::int uf_298)
-#3074 := (and #3073 #3066)
-#3118 := (implies #3074 #3117)
-#3119 := (implies #3074 #3118)
-#3120 := (implies true #3119)
-#3121 := (implies #3074 #3120)
-#3122 := (implies up_216 #3121)
-#3123 := (and up_216 #3122)
-#3124 := (implies #3074 #3123)
-#3125 := (implies true #3124)
-#3126 := (implies #3074 #3125)
-#3288 := (implies #3074 #3126)
-#3289 := (implies true #3288)
-#3290 := (implies #3074 #3289)
-#3287 := (<= uf_286 uf_298)
-#3291 := (implies #3287 #3290)
-#3292 := (implies #3074 #3291)
-#3293 := (implies true #3292)
-#3230 := (implies false true)
-#3228 := (= #3227 uf_304)
-#3225 := (< uf_305 uf_286)
-#3229 := (and #3225 #3228)
-#3231 := (implies #3229 #3230)
-#3232 := (and #3229 #3231)
-#3221 := (<= #3054 uf_304)
-#3220 := (< #161 uf_306)
-#3222 := (implies #3220 #3221)
-#3223 := (implies #412 #3222)
-#3224 := (forall (vars (?x786 int)) #3223)
-#3233 := (implies #3224 #3232)
-#3234 := (and #3224 #3233)
-#3219 := (<= uf_306 uf_286)
-#3235 := (implies #3219 #3234)
-#3236 := (and #3219 #3235)
-#3237 := (implies true #3236)
-#3208 := (<= 0::int uf_305)
-#3217 := (<= 2::int uf_306)
-#3218 := (and #3217 #3208)
-#3238 := (implies #3218 #3237)
-#3239 := (implies #3216 #3238)
-#3210 := (+ uf_298 1::int)
-#3215 := (= uf_306 #3210)
-#3240 := (implies #3215 #3239)
-#3212 := (<= #3210 uf_76)
-#3211 := (<= 0::int #3210)
-#3213 := (and #3211 #3212)
-#3241 := (implies #3213 #3240)
-#3242 := (and #3213 #3241)
-#3209 := (and #3073 #3208)
-#3243 := (implies #3209 #3242)
-#3244 := (implies true #3243)
-#3264 := (= uf_305 uf_297)
-#3265 := (implies #3264 #3244)
-#3263 := (= uf_304 uf_296)
-#3266 := (implies #3263 #3265)
-#3267 := (implies true #3266)
-#3268 := (implies #3074 #3267)
-#3269 := (implies #3074 #3268)
-#3270 := (implies true #3269)
-#3271 := (implies #3074 #3270)
-#3262 := (<= #3197 uf_296)
-#3272 := (implies #3262 #3271)
-#3273 := (implies #3074 #3272)
-#3274 := (implies true #3273)
-#3207 := (= uf_305 uf_298)
-#3245 := (implies #3207 #3244)
-#3205 := (= uf_304 uf_303)
-#3246 := (implies #3205 #3245)
-#3247 := (implies true #3246)
-#3203 := (and #3073 #3073)
-#3248 := (implies #3203 #3247)
-#3249 := (implies #3202 #3248)
-#3250 := (implies #3201 #3249)
-#3200 := (= uf_303 #3197)
-#3251 := (implies #3200 #3250)
-#3195 := (= #3194 uf_9)
-#3190 := (= #3189 uf_9)
-#3196 := (and #3190 #3195)
-#3252 := (implies #3196 #3251)
-#3253 := (and #3196 #3252)
-#3192 := (= #3191 uf_9)
-#3193 := (and #3190 #3192)
-#3254 := (implies #3193 #3253)
-#3255 := (and #3193 #3254)
-#3256 := (implies #3074 #3255)
-#3257 := (implies true #3256)
-#3258 := (implies #3074 #3257)
-#3198 := (< uf_296 #3197)
-#3259 := (implies #3198 #3258)
-#3260 := (implies #3074 #3259)
-#3261 := (implies true #3260)
-#3275 := (and #3261 #3274)
-#3276 := (implies #3074 #3275)
-#3277 := (implies #3196 #3276)
-#3278 := (and #3196 #3277)
-#3279 := (implies #3193 #3278)
-#3280 := (and #3193 #3279)
-#3281 := (implies #3074 #3280)
-#3282 := (implies true #3281)
-#3283 := (implies #3074 #3282)
-#3187 := (< uf_298 uf_286)
-#3284 := (implies #3187 #3283)
-#3285 := (implies #3074 #3284)
-#3286 := (implies true #3285)
-#3294 := (and #3286 #3293)
-#3295 := (implies #3074 #3294)
-decl uf_41 :: (-> T4 T12)
-#3134 := (uf_41 uf_287)
-#3185 := (= #3134 #3134)
-#3184 := (= #3157 #3157)
-#3186 := (and #3184 #3185)
-#3296 := (implies #3186 #3295)
-#3297 := (implies #3183 #3296)
-#3298 := (implies #3180 #3297)
-#3299 := (implies #3179 #3298)
-#3300 := (implies #3178 #3299)
-#3301 := (implies #3177 #3300)
-#3005 := (= #3004 uf_9)
-#3088 := (= #3087 uf_9)
-#3089 := (and #3088 #3005)
-#3302 := (implies #3089 #3301)
-#3170 := (= #3169 uf_9)
-decl uf_172 :: (-> T4 T5 int)
-#3165 := (uf_172 uf_287 #23)
-#3166 := (pattern #3165)
-#3167 := (<= #3165 #3165)
-#3168 := (forall (vars (?x784 T5)) (:pat #3166) #3167)
-#3171 := (and #3168 #3170)
-#3164 := (<= #3019 #3019)
-#3172 := (and #3164 #3171)
-#3303 := (implies #3172 #3302)
-#3158 := (uf_58 #3157 #23)
-#3159 := (pattern #3158)
-#3147 := (uf_68 uf_287 #23)
-#3148 := (= #3147 uf_9)
-#3160 := (= #3158 #3158)
-#3161 := (and #3160 #3148)
-#3162 := (implies #3148 #3161)
-#3163 := (forall (vars (?x783 T5)) (:pat #3159) #3162)
-#3173 := (and #3163 #3172)
-decl uf_40 :: (-> T12 T5 T11)
-#3135 := (uf_40 #3134 #23)
-#3136 := (pattern #3135)
-#3153 := (= #3135 #3135)
-#3154 := (and #3153 #3148)
-#3155 := (implies #3148 #3154)
-#3156 := (forall (vars (?x782 T5)) (:pat #3136) #3155)
-#3174 := (and #3156 #3173)
-decl uf_19 :: (-> T9 T5 int)
-decl uf_20 :: (-> T4 T9)
-#3144 := (uf_20 uf_287)
-#3145 := (uf_19 #3144 #23)
-#3146 := (pattern #3145)
-#3149 := (= #3145 #3145)
-#3150 := (and #3149 #3148)
-#3151 := (implies #3148 #3150)
-#3152 := (forall (vars (?x781 T5)) (:pat #3146) #3151)
-#3175 := (and #3152 #3174)
-decl uf_261 :: T8
-#2837 := uf_261
-#3137 := (uf_25 uf_287 #23)
-#3138 := (uf_13 #3137)
-#3139 := (uf_12 #3138)
-#3140 := (= #3139 uf_261)
-#3141 := (not #3140)
-#3142 := (implies #3141 #3141)
-#3143 := (forall (vars (?x780 T5)) (:pat #3136) #3142)
-#3176 := (and #3143 #3175)
-#3304 := (implies #3176 #3303)
-#3305 := (implies #3074 #3304)
-#3306 := (implies true #3305)
-#3307 := (implies #3074 #3306)
-#3308 := (implies true #3307)
-#3309 := (implies #3074 #3308)
-#3310 := (implies true #3309)
-#3127 := (implies #3089 #3126)
-#3128 := (implies #3074 #3127)
-#3129 := (implies true #3128)
-#3130 := (implies #3074 #3129)
-#3086 := (not true)
-#3131 := (implies #3086 #3130)
-#3132 := (implies #3074 #3131)
-#3133 := (implies true #3132)
-#3311 := (and #3133 #3310)
-#3312 := (implies #3074 #3311)
-#3084 := (= #3083 uf_296)
-#3081 := (< uf_297 uf_286)
-#3085 := (and #3081 #3084)
-#3313 := (implies #3085 #3312)
-#3077 := (<= #3054 uf_296)
-#3076 := (< #161 uf_298)
-#3078 := (implies #3076 #3077)
-#3079 := (implies #412 #3078)
-#3080 := (forall (vars (?x775 int)) #3079)
-#3314 := (implies #3080 #3313)
-#3075 := (<= uf_298 uf_286)
-#3315 := (implies #3075 #3314)
-#3316 := (implies #3074 #3315)
-#3071 := (<= uf_298 uf_76)
-#3070 := (<= 0::int uf_298)
-#3072 := (and #3070 #3071)
-#3317 := (implies #3072 #3316)
-#3067 := (<= uf_297 uf_76)
-#3068 := (and #3066 #3067)
-#3318 := (implies #3068 #3317)
-#3063 := (<= uf_296 uf_78)
-#3062 := (<= 0::int uf_296)
-#3064 := (and #3062 #3063)
-#3319 := (implies #3064 #3318)
-#3320 := (implies true #3319)
-#3059 := (= #3041 uf_295)
-#2975 := (< 0::int uf_286)
-#3060 := (and #2975 #3059)
-#3321 := (implies #3060 #3320)
-#3322 := (and #3060 #3321)
-#3055 := (<= #3054 uf_295)
-#3052 := (< #161 1::int)
-#3056 := (implies #3052 #3055)
-#3057 := (implies #412 #3056)
-#3058 := (forall (vars (?x773 int)) #3057)
-#3323 := (implies #3058 #3322)
-#3324 := (and #3058 #3323)
-#3051 := (<= 1::int uf_286)
-#3325 := (implies #3051 #3324)
-#3326 := (and #3051 #3325)
-#3047 := (<= 0::int 0::int)
-#3048 := (and #3047 #3047)
-#3046 := (<= 1::int 1::int)
-#3049 := (and #3046 #3048)
-#3050 := (and #3046 #3049)
-#3327 := (implies #3050 #3326)
-#3328 := (implies #3045 #3327)
-#3329 := (implies #3044 #3328)
-#3330 := (implies #3043 #3329)
-#3331 := (implies #3042 #3330)
-#3038 := (= #3037 uf_9)
-#3033 := (= #3032 uf_9)
-#3039 := (and #3033 #3038)
-#3332 := (implies #3039 #3331)
-#3333 := (and #3039 #3332)
-#3035 := (= #3034 uf_9)
-#3036 := (and #3033 #3035)
-#3334 := (implies #3036 #3333)
-#3335 := (and #3036 #3334)
-#3030 := (= #3029 uf_9)
-#3336 := (implies #3030 #3335)
-#3337 := (and #3030 #3336)
-#3027 := (<= uf_286 uf_76)
-#3026 := (<= 0::int uf_286)
-#3028 := (and #3026 #3027)
-#3338 := (implies #3028 #3337)
-#3023 := (= #3021 uf_9)
-#3024 := (iff #3023 false)
-#3025 := (forall (vars (?x772 T5)) (:pat #3022) #3024)
-#3339 := (implies #3025 #3338)
-#3340 := (implies #3020 #3339)
-#3341 := (implies #3017 #3340)
-#3342 := (implies #3016 #3341)
-#3011 := (< #3008 uf_290)
-#3012 := (forall (vars (?x771 T19)) (:pat #3009) #3011)
-#3343 := (implies #3012 #3342)
-#3003 := (= #3002 uf_9)
-#3006 := (and #3003 #3005)
-#3344 := (implies #3006 #3343)
-#3001 := (= #3000 uf_9)
-#3345 := (implies #3001 #3344)
-#3346 := (implies true #3345)
-#2994 := (= #2993 uf_9)
-#2991 := (= #2990 uf_14)
-#2992 := (not #2991)
-#2995 := (and #2992 #2994)
-#2989 := (= #2988 uf_9)
-#2996 := (and #2989 #2995)
-#2987 := (= #2986 uf_9)
-#2997 := (and #2987 #2996)
-#2985 := (= #2984 uf_26)
-#2998 := (and #2985 #2997)
-#2983 := (= #2982 uf_9)
-#2999 := (and #2983 #2998)
-#3347 := (implies #2999 #3346)
-#3348 := (implies #2975 #3347)
-#2974 := (< uf_286 1099511627776::int)
-#3349 := (implies #2974 #3348)
-#2971 := (<= uf_285 uf_76)
-#2970 := (<= 0::int uf_285)
-#2972 := (and #2970 #2971)
-#3350 := (implies #2972 #3349)
-#2967 := (<= uf_284 uf_76)
-#2966 := (<= 0::int uf_284)
-#2968 := (and #2966 #2967)
-#3351 := (implies #2968 #3350)
-#2963 := (<= uf_283 uf_78)
-#2962 := (<= 0::int uf_283)
-#2964 := (and #2962 #2963)
-#3352 := (implies #2964 #3351)
-#3353 := (implies true #3352)
-#3354 := (not #3353)
-#14264 := (iff #3354 #14261)
-#12047 := (not #3098)
-#12048 := (or #12047 #3099)
-#5597 := (not #412)
-#12054 := (or #5597 #12048)
-#12059 := (forall (vars (?x776 int)) #12054)
-#12074 := (not #12059)
-#12075 := (or #3107 #12074)
-#12080 := (and #12059 #12075)
-#12094 := (or #12093 #12080)
-#12103 := (or #12102 #12094)
-#12112 := (or #12111 #12103)
-#12121 := (or #12120 #12112)
-#11998 := (and #3066 #3073)
-#12136 := (not #11998)
-#12137 := (or #12136 #12121)
-#12145 := (or #12136 #12137)
-#12160 := (or #12136 #12145)
-#12169 := (or #12168 #12160)
-#12174 := (and up_216 #12169)
-#12180 := (or #12136 #12174)
-#12195 := (or #12136 #12180)
-#12837 := (or #12136 #12195)
-#12852 := (or #12136 #12837)
-#12860 := (not #3287)
-#12861 := (or #12860 #12852)
-#12869 := (or #12136 #12861)
-#12431 := (and #3225 #12428)
-#12413 := (not #3220)
-#12414 := (or #12413 #3221)
-#12420 := (or #5597 #12414)
-#12425 := (forall (vars (?x786 int)) #12420)
-#12453 := (not #12425)
-#12454 := (or #12453 #12431)
-#12459 := (and #12425 #12454)
-#12465 := (not #3219)
-#12466 := (or #12465 #12459)
-#12471 := (and #3219 #12466)
-#12410 := (and #3208 #3217)
-#12484 := (not #12410)
-#12485 := (or #12484 #12471)
-#12494 := (or #12493 #12485)
-#12395 := (+ 1::int uf_298)
-#12407 := (= uf_306 #12395)
-#12502 := (not #12407)
-#12503 := (or #12502 #12494)
-#12401 := (<= #12395 uf_76)
-#12398 := (<= 0::int #12395)
-#12404 := (and #12398 #12401)
-#12511 := (not #12404)
-#12512 := (or #12511 #12503)
-#12517 := (and #12404 #12512)
-#12523 := (not #3209)
-#12524 := (or #12523 #12517)
-#12678 := (or #12524 #12677)
-#12687 := (or #12686 #12678)
-#12702 := (or #12136 #12687)
-#12710 := (or #12136 #12702)
-#12725 := (or #12136 #12710)
-#12733 := (not #3262)
-#12734 := (or #12733 #12725)
-#12742 := (or #12136 #12734)
-#12540 := (or #12539 #12524)
-#12549 := (or #12548 #12540)
-#12564 := (not #3073)
-#12565 := (or #12564 #12549)
-#12574 := (or #12573 #12565)
-#12583 := (or #12582 #12574)
-#12592 := (or #12591 #12583)
-#12601 := (or #12600 #12592)
-#12606 := (and #12381 #12601)
-#12613 := (or #12612 #12606)
-#12618 := (and #12375 #12613)
-#12624 := (or #12136 #12618)
-#12639 := (or #12136 #12624)
-#12647 := (not #3198)
-#12648 := (or #12647 #12639)
-#12656 := (or #12136 #12648)
-#12754 := (and #12656 #12742)
-#12760 := (or #12136 #12754)
-#12768 := (or #12600 #12760)
-#12773 := (and #12381 #12768)
-#12779 := (or #12612 #12773)
-#12784 := (and #12375 #12779)
-#12790 := (or #12136 #12784)
-#12805 := (or #12136 #12790)
-#12813 := (not #3187)
-#12814 := (or #12813 #12805)
-#12822 := (or #12136 #12814)
-#12881 := (and #12822 #12869)
-#12887 := (or #12136 #12881)
-#12903 := (or #12902 #12887)
-#12912 := (or #12911 #12903)
-#12921 := (or #12920 #12912)
-#12930 := (or #12929 #12921)
-#12939 := (or #12938 #12930)
-#12947 := (or #12203 #12939)
-#12333 := (and #3168 #12330)
-#12336 := (and #3164 #12333)
-#12955 := (not #12336)
-#12956 := (or #12955 #12947)
-#12964 := (or #12955 #12956)
-#12972 := (or #12136 #12964)
-#12987 := (or #12136 #12972)
-#13002 := (or #12136 #12987)
-#13024 := (or #12136 #13002)
-#12019 := (and #3081 #12016)
-#13032 := (not #12019)
-#13033 := (or #13032 #13024)
-#12001 := (not #3076)
-#12002 := (or #12001 #3077)
-#12008 := (or #5597 #12002)
-#12013 := (forall (vars (?x775 int)) #12008)
-#13041 := (not #12013)
-#13042 := (or #13041 #13033)
-#13050 := (not #3075)
-#13051 := (or #13050 #13042)
-#13059 := (or #12136 #13051)
-#13067 := (not #3072)
-#13068 := (or #13067 #13059)
-#13076 := (not #3068)
-#13077 := (or #13076 #13068)
-#13085 := (not #3064)
-#13086 := (or #13085 #13077)
-#11995 := (and #2975 #3042)
-#13101 := (not #11995)
-#13102 := (or #13101 #13086)
-#13107 := (and #11995 #13102)
-#11978 := (not #3052)
-#11979 := (or #11978 #3055)
-#11985 := (or #5597 #11979)
-#11990 := (forall (vars (?x773 int)) #11985)
-#13113 := (not #11990)
-#13114 := (or #13113 #13107)
-#13119 := (and #11990 #13114)
-#13125 := (not #3051)
-#13126 := (or #13125 #13119)
-#13131 := (and #3051 #13126)
-#11972 := (and #3046 #3047)
-#11975 := (and #3046 #11972)
-#13137 := (not #11975)
-#13138 := (or #13137 #13131)
-#13147 := (or #13146 #13138)
-#13156 := (or #13155 #13147)
-#13165 := (or #13164 #13156)
-#13174 := (or #13173 #13165)
-#13183 := (or #13182 #13174)
-#13188 := (and #11967 #13183)
-#13195 := (or #13194 #13188)
-#13200 := (and #11961 #13195)
-#13207 := (or #13206 #13200)
-#13212 := (and #11952 #13207)
-#13218 := (not #3028)
-#13219 := (or #13218 #13212)
-#13228 := (or #13227 #13219)
-#13237 := (or #13236 #13228)
-#13246 := (or #13245 #13237)
-#13255 := (or #13254 #13246)
-#13263 := (not #3012)
-#13264 := (or #13263 #13255)
-#13273 := (or #13272 #13264)
-#13282 := (or #13281 #13273)
-#11911 := (and #11905 #11908)
-#11914 := (and #11899 #11911)
-#11917 := (and #11896 #11914)
-#11920 := (and #11893 #11917)
-#11923 := (and #11889 #11920)
-#13297 := (not #11923)
-#13298 := (or #13297 #13282)
-#13306 := (not #2975)
-#13307 := (or #13306 #13298)
-#13315 := (not #2974)
-#13316 := (or #13315 #13307)
-#13324 := (not #2972)
-#13325 := (or #13324 #13316)
-#13333 := (not #2968)
-#13334 := (or #13333 #13325)
-#13342 := (not #2964)
-#13343 := (or #13342 #13334)
-#13355 := (not #13343)
-#14262 := (iff #13355 #14261)
-#14259 := (iff #13343 #14256)
-#14214 := (or #13369 #14133)
-#14217 := (or #13227 #14214)
-#14220 := (or #13236 #14217)
-#14223 := (or #13245 #14220)
-#14226 := (or #13254 #14223)
-#14229 := (or #14145 #14226)
-#14232 := (or #13272 #14229)
-#14235 := (or #13281 #14232)
-#14238 := (or #14151 #14235)
-#14241 := (or #13428 #14238)
-#14244 := (or #14161 #14241)
-#14247 := (or #14183 #14244)
-#14250 := (or #14197 #14247)
-#14253 := (or #14211 #14250)
-#14257 := (iff #14253 #14256)
-#14258 := [rewrite]: #14257
-#14254 := (iff #13343 #14253)
-#14251 := (iff #13334 #14250)
-#14248 := (iff #13325 #14247)
-#14245 := (iff #13316 #14244)
-#14242 := (iff #13307 #14241)
-#14239 := (iff #13298 #14238)
-#14236 := (iff #13282 #14235)
-#14233 := (iff #13273 #14232)
-#14230 := (iff #13264 #14229)
-#14227 := (iff #13255 #14226)
-#14224 := (iff #13246 #14223)
-#14221 := (iff #13237 #14220)
-#14218 := (iff #13228 #14217)
-#14215 := (iff #13219 #14214)
-#14134 := (iff #13212 #14133)
-#14131 := (iff #13207 #14130)
-#14128 := (iff #13200 #14125)
-#14122 := (and #11961 #14119)
-#14126 := (iff #14122 #14125)
-#14127 := [rewrite]: #14126
-#14123 := (iff #13200 #14122)
-#14120 := (iff #13195 #14119)
-#14117 := (iff #13188 #14114)
-#14111 := (and #11967 #14106)
-#14115 := (iff #14111 #14114)
-#14116 := [rewrite]: #14115
-#14112 := (iff #13188 #14111)
-#14109 := (iff #13183 #14106)
-#14088 := (or false #14085)
-#14091 := (or #13146 #14088)
-#14094 := (or #13155 #14091)
-#14097 := (or #13164 #14094)
-#14100 := (or #13173 #14097)
-#14103 := (or #13182 #14100)
-#14107 := (iff #14103 #14106)
-#14108 := [rewrite]: #14107
-#14104 := (iff #13183 #14103)
-#14101 := (iff #13174 #14100)
-#14098 := (iff #13165 #14097)
-#14095 := (iff #13156 #14094)
-#14092 := (iff #13147 #14091)
-#14089 := (iff #13138 #14088)
-#14086 := (iff #13131 #14085)
-#14083 := (iff #13126 #14082)
-#14080 := (iff #13119 #14079)
-#14077 := (iff #13114 #14076)
-#14074 := (iff #13107 #14071)
-#13989 := (and #13429 #3042)
-#14068 := (and #13989 #14063)
-#14072 := (iff #14068 #14071)
-#14073 := [rewrite]: #14072
-#14069 := (iff #13107 #14068)
-#14066 := (iff #13102 #14063)
-#14000 := (or #13441 #13870)
-#14003 := (or #12902 #14000)
-#14006 := (or #12911 #14003)
-#14009 := (or #12920 #14006)
-#14012 := (or #12929 #14009)
-#14015 := (or #12938 #14012)
-#14018 := (or #12203 #14015)
-#14021 := (or #13894 #14018)
-#14024 := (or #13894 #14021)
-#14027 := (or #13441 #14024)
-#14030 := (or #13441 #14027)
-#14033 := (or #13441 #14030)
-#14036 := (or #13441 #14033)
-#14039 := (or #13911 #14036)
-#14042 := (or #13945 #14039)
-#14045 := (or #13951 #14042)
-#14048 := (or #13441 #14045)
-#14051 := (or #13963 #14048)
-#14054 := (or #13973 #14051)
-#14057 := (or #13986 #14054)
-#14060 := (or #13997 #14057)
-#14064 := (iff #14060 #14063)
-#14065 := [rewrite]: #14064
-#14061 := (iff #13102 #14060)
-#14058 := (iff #13086 #14057)
-#14055 := (iff #13077 #14054)
-#14052 := (iff #13068 #14051)
-#14049 := (iff #13059 #14048)
-#14046 := (iff #13051 #14045)
-#14043 := (iff #13042 #14042)
-#14040 := (iff #13033 #14039)
-#14037 := (iff #13024 #14036)
-#14034 := (iff #13002 #14033)
-#14031 := (iff #12987 #14030)
-#14028 := (iff #12972 #14027)
-#14025 := (iff #12964 #14024)
-#14022 := (iff #12956 #14021)
-#14019 := (iff #12947 #14018)
-#14016 := (iff #12939 #14015)
-#14013 := (iff #12930 #14012)
-#14010 := (iff #12921 #14009)
-#14007 := (iff #12912 #14006)
-#14004 := (iff #12903 #14003)
-#14001 := (iff #12887 #14000)
-#13871 := (iff #12881 #13870)
-#13868 := (iff #12869 #13865)
-#13847 := (or #13441 #13840)
-#13850 := (or #13441 #13847)
-#13853 := (or #13441 #13850)
-#13856 := (or #13441 #13853)
-#13859 := (or #13722 #13856)
-#13862 := (or #13441 #13859)
-#13866 := (iff #13862 #13865)
-#13867 := [rewrite]: #13866
-#13863 := (iff #12869 #13862)
-#13860 := (iff #12861 #13859)
-#13857 := (iff #12852 #13856)
-#13854 := (iff #12837 #13853)
-#13851 := (iff #12195 #13850)
-#13848 := (iff #12180 #13847)
-#13841 := (iff #12174 #13840)
-#13838 := (iff #12169 #13835)
-#13811 := (or #12093 #13808)
-#13814 := (or #12102 #13811)
-#13817 := (or #12111 #13814)
-#13820 := (or #12120 #13817)
-#13823 := (or #13441 #13820)
-#13826 := (or #13441 #13823)
-#13829 := (or #13441 #13826)
-#13832 := (or #12168 #13829)
-#13836 := (iff #13832 #13835)
-#13837 := [rewrite]: #13836
-#13833 := (iff #12169 #13832)
-#13830 := (iff #12160 #13829)
-#13827 := (iff #12145 #13826)
-#13824 := (iff #12137 #13823)
-#13821 := (iff #12121 #13820)
-#13818 := (iff #12112 #13817)
-#13815 := (iff #12103 #13814)
-#13812 := (iff #12094 #13811)
-#13809 := (iff #12080 #13808)
-#13806 := (iff #12075 #13803)
-#13800 := (or #13794 #13797)
-#13804 := (iff #13800 #13803)
-#13805 := [rewrite]: #13804
-#13801 := (iff #12075 #13800)
-#13798 := (iff #12074 #13797)
-#13778 := (iff #12059 #13777)
-#13775 := (iff #12054 #13772)
-#13766 := (or #13749 #13763)
-#13769 := (or #5606 #13766)
-#13773 := (iff #13769 #13772)
-#13774 := [rewrite]: #13773
-#13770 := (iff #12054 #13769)
-#13767 := (iff #12048 #13766)
-#13764 := (iff #3099 #13763)
-#13765 := [rewrite]: #13764
-#13759 := (iff #12047 #13749)
-#13754 := (not #13751)
-#13757 := (iff #13754 #13749)
-#13758 := [rewrite]: #13757
-#13755 := (iff #12047 #13754)
-#13752 := (iff #3098 #13751)
-#13753 := [rewrite]: #13752
-#13756 := [monotonicity #13753]: #13755
-#13760 := [trans #13756 #13758]: #13759
-#13768 := [monotonicity #13760 #13765]: #13767
-#5607 := (iff #5597 #5606)
-#4392 := (iff #412 #4391)
-#4385 := (iff #411 #4384)
-#4386 := [rewrite]: #4385
-#4068 := (iff #285 #4070)
-#4069 := [rewrite]: #4068
-#4393 := [monotonicity #4069 #4386]: #4392
-#5608 := [monotonicity #4393]: #5607
-#13771 := [monotonicity #5608 #13768]: #13770
-#13776 := [trans #13771 #13774]: #13775
-#13779 := [quant-intro #13776]: #13778
-#13799 := [monotonicity #13779]: #13798
-#13795 := (iff #3107 #13794)
-#13792 := (iff #3106 #13789)
-#13780 := (and #13751 #3103)
-#13783 := (and #4384 #13780)
-#13786 := (and #4070 #13783)
-#13790 := (iff #13786 #13789)
-#13791 := [rewrite]: #13790
-#13787 := (iff #3106 #13786)
-#13784 := (iff #3105 #13783)
-#13781 := (iff #3104 #13780)
-#13782 := [monotonicity #13753]: #13781
-#13785 := [monotonicity #4386 #13782]: #13784
-#13788 := [monotonicity #4069 #13785]: #13787
-#13793 := [trans #13788 #13791]: #13792
-#13796 := [quant-intro #13793]: #13795
-#13802 := [monotonicity #13796 #13799]: #13801
-#13807 := [trans #13802 #13805]: #13806
-#13810 := [monotonicity #13779 #13807]: #13809
-#13813 := [monotonicity #13810]: #13812
-#13816 := [monotonicity #13813]: #13815
-#13819 := [monotonicity #13816]: #13818
-#13822 := [monotonicity #13819]: #13821
-#13442 := (iff #12136 #13441)
-#13439 := (iff #11998 #13438)
-#13435 := (iff #3073 #13436)
-#13437 := [rewrite]: #13435
-#13432 := (iff #3066 #13433)
-#13434 := [rewrite]: #13432
-#13440 := [monotonicity #13434 #13437]: #13439
-#13443 := [monotonicity #13440]: #13442
-#13825 := [monotonicity #13443 #13822]: #13824
-#13828 := [monotonicity #13443 #13825]: #13827
-#13831 := [monotonicity #13443 #13828]: #13830
-#13834 := [monotonicity #13831]: #13833
-#13839 := [trans #13834 #13837]: #13838
-#13842 := [monotonicity #13839]: #13841
-#13849 := [monotonicity #13443 #13842]: #13848
-#13852 := [monotonicity #13443 #13849]: #13851
-#13855 := [monotonicity #13443 #13852]: #13854
-#13858 := [monotonicity #13443 #13855]: #13857
-#13845 := (iff #12860 #13722)
-#13843 := (iff #3287 #13721)
-#13844 := [rewrite]: #13843
-#13846 := [monotonicity #13844]: #13845
-#13861 := [monotonicity #13846 #13858]: #13860
-#13864 := [monotonicity #13443 #13861]: #13863
-#13869 := [trans #13864 #13867]: #13868
-#13747 := (iff #12822 #13744)
-#13732 := (or #13441 #13715)
-#13735 := (or #13441 #13732)
-#13738 := (or #13721 #13735)
-#13741 := (or #13441 #13738)
-#13745 := (iff #13741 #13744)
-#13746 := [rewrite]: #13745
-#13742 := (iff #12822 #13741)
-#13739 := (iff #12814 #13738)
-#13736 := (iff #12805 #13735)
-#13733 := (iff #12790 #13732)
-#13718 := (iff #12784 #13715)
-#13712 := (and #12375 #13709)
-#13716 := (iff #13712 #13715)
-#13717 := [rewrite]: #13716
-#13713 := (iff #12784 #13712)
-#13710 := (iff #12779 #13709)
-#13707 := (iff #12773 #13704)
-#13701 := (and #12381 #13696)
-#13705 := (iff #13701 #13704)
-#13706 := [rewrite]: #13705
-#13702 := (iff #12773 #13701)
-#13699 := (iff #12768 #13696)
-#13690 := (or #13441 #13687)
-#13693 := (or #12600 #13690)
-#13697 := (iff #13693 #13696)
-#13698 := [rewrite]: #13697
-#13694 := (iff #12768 #13693)
-#13691 := (iff #12760 #13690)
-#13688 := (iff #12754 #13687)
-#13685 := (iff #12742 #13682)
-#13579 := (or #13450 #13571)
-#13661 := (or #13579 #12677)
-#13664 := (or #12686 #13661)
-#13667 := (or #13441 #13664)
-#13670 := (or #13441 #13667)
-#13673 := (or #13441 #13670)
-#13676 := (or #13627 #13673)
-#13679 := (or #13441 #13676)
-#13683 := (iff #13679 #13682)
-#13684 := [rewrite]: #13683
-#13680 := (iff #12742 #13679)
-#13677 := (iff #12734 #13676)
-#13674 := (iff #12725 #13673)
-#13671 := (iff #12710 #13670)
-#13668 := (iff #12702 #13667)
-#13665 := (iff #12687 #13664)
-#13662 := (iff #12678 #13661)
-#13580 := (iff #12524 #13579)
-#13574 := (iff #12517 #13571)
-#13568 := (and #13545 #13563)
-#13572 := (iff #13568 #13571)
-#13573 := [rewrite]: #13572
-#13569 := (iff #12517 #13568)
-#13566 := (iff #12512 #13563)
-#13551 := (or #13467 #13535)
-#13554 := (or #12493 #13551)
-#13557 := (or #13542 #13554)
-#13560 := (or #13548 #13557)
-#13564 := (iff #13560 #13563)
-#13565 := [rewrite]: #13564
-#13561 := (iff #12512 #13560)
-#13558 := (iff #12503 #13557)
-#13555 := (iff #12494 #13554)
-#13552 := (iff #12485 #13551)
-#13536 := (iff #12471 #13535)
-#13533 := (iff #12466 #13532)
-#13530 := (iff #12459 #13529)
-#13527 := (iff #12454 #13526)
-#13524 := (iff #12431 #13521)
-#13518 := (and #13515 #12428)
-#13522 := (iff #13518 #13521)
-#13523 := [rewrite]: #13522
-#13519 := (iff #12431 #13518)
-#13516 := (iff #3225 #13515)
-#13517 := [rewrite]: #13516
-#13520 := [monotonicity #13517]: #13519
-#13525 := [trans #13520 #13523]: #13524
-#13510 := (iff #12453 #13509)
-#13507 := (iff #12425 #13506)
-#13504 := (iff #12420 #13501)
-#13495 := (or #13478 #13492)
-#13498 := (or #5606 #13495)
-#13502 := (iff #13498 #13501)
-#13503 := [rewrite]: #13502
-#13499 := (iff #12420 #13498)
-#13496 := (iff #12414 #13495)
-#13493 := (iff #3221 #13492)
-#13494 := [rewrite]: #13493
-#13488 := (iff #12413 #13478)
-#13480 := (not #13478)
-#13483 := (not #13480)
-#13486 := (iff #13483 #13478)
-#13487 := [rewrite]: #13486
-#13484 := (iff #12413 #13483)
-#13481 := (iff #3220 #13480)
-#13482 := [rewrite]: #13481
-#13485 := [monotonicity #13482]: #13484
-#13489 := [trans #13485 #13487]: #13488
-#13497 := [monotonicity #13489 #13494]: #13496
-#13500 := [monotonicity #5608 #13497]: #13499
-#13505 := [trans #13500 #13503]: #13504
-#13508 := [quant-intro #13505]: #13507
-#13511 := [monotonicity #13508]: #13510
-#13528 := [monotonicity #13511 #13525]: #13527
-#13531 := [monotonicity #13508 #13528]: #13530
-#13476 := (iff #12465 #13475)
-#13473 := (iff #3219 #13470)
-#13474 := [rewrite]: #13473
-#13477 := [monotonicity #13474]: #13476
-#13534 := [monotonicity #13477 #13531]: #13533
-#13537 := [monotonicity #13474 #13534]: #13536
-#13468 := (iff #12484 #13467)
-#13465 := (iff #12410 #13464)
-#13461 := (iff #3217 #13462)
-#13463 := [rewrite]: #13461
-#13444 := (iff #3208 #13445)
-#13446 := [rewrite]: #13444
-#13466 := [monotonicity #13446 #13463]: #13465
-#13469 := [monotonicity #13466]: #13468
-#13553 := [monotonicity #13469 #13537]: #13552
-#13556 := [monotonicity #13553]: #13555
-#13543 := (iff #12502 #13542)
-#13540 := (iff #12407 #13538)
-#13541 := [rewrite]: #13540
-#13544 := [monotonicity #13541]: #13543
-#13559 := [monotonicity #13544 #13556]: #13558
-#13549 := (iff #12511 #13548)
-#13546 := (iff #12404 #13545)
-#13459 := (iff #12401 #13456)
-#13460 := [rewrite]: #13459
-#13453 := (iff #12398 #13454)
-#13455 := [rewrite]: #13453
-#13547 := [monotonicity #13455 #13460]: #13546
-#13550 := [monotonicity #13547]: #13549
-#13562 := [monotonicity #13550 #13559]: #13561
-#13567 := [trans #13562 #13565]: #13566
-#13570 := [monotonicity #13547 #13567]: #13569
-#13575 := [trans #13570 #13573]: #13574
-#13451 := (iff #12523 #13450)
-#13448 := (iff #3209 #13447)
-#13449 := [monotonicity #13437 #13446]: #13448
-#13452 := [monotonicity #13449]: #13451
-#13581 := [monotonicity #13452 #13575]: #13580
-#13663 := [monotonicity #13581]: #13662
-#13666 := [monotonicity #13663]: #13665
-#13669 := [monotonicity #13443 #13666]: #13668
-#13672 := [monotonicity #13443 #13669]: #13671
-#13675 := [monotonicity #13443 #13672]: #13674
-#13659 := (iff #12733 #13627)
-#13657 := (iff #3262 #13628)
-#13658 := [rewrite]: #13657
-#13660 := [monotonicity #13658]: #13659
-#13678 := [monotonicity #13660 #13675]: #13677
-#13681 := [monotonicity #13443 #13678]: #13680
-#13686 := [trans #13681 #13684]: #13685
-#13655 := (iff #12656 #13652)
-#13640 := (or #13441 #13622)
-#13643 := (or #13441 #13640)
-#13646 := (or #13628 #13643)
-#13649 := (or #13441 #13646)
-#13653 := (iff #13649 #13652)
-#13654 := [rewrite]: #13653
-#13650 := (iff #12656 #13649)
-#13647 := (iff #12648 #13646)
-#13644 := (iff #12639 #13643)
-#13641 := (iff #12624 #13640)
-#13625 := (iff #12618 #13622)
-#13619 := (and #12375 #13616)
-#13623 := (iff #13619 #13622)
-#13624 := [rewrite]: #13623
-#13620 := (iff #12618 #13619)
-#13617 := (iff #12613 #13616)
-#13614 := (iff #12606 #13611)
-#13608 := (and #12381 #13603)
-#13612 := (iff #13608 #13611)
-#13613 := [rewrite]: #13612
-#13609 := (iff #12606 #13608)
-#13606 := (iff #12601 #13603)
-#13582 := (or #12539 #13579)
-#13585 := (or #12548 #13582)
-#13588 := (or #13576 #13585)
-#13591 := (or #12573 #13588)
-#13594 := (or #12582 #13591)
-#13597 := (or #12591 #13594)
-#13600 := (or #12600 #13597)
-#13604 := (iff #13600 #13603)
-#13605 := [rewrite]: #13604
-#13601 := (iff #12601 #13600)
-#13598 := (iff #12592 #13597)
-#13595 := (iff #12583 #13594)
-#13592 := (iff #12574 #13591)
-#13589 := (iff #12565 #13588)
-#13586 := (iff #12549 #13585)
-#13583 := (iff #12540 #13582)
-#13584 := [monotonicity #13581]: #13583
-#13587 := [monotonicity #13584]: #13586
-#13577 := (iff #12564 #13576)
-#13578 := [monotonicity #13437]: #13577
-#13590 := [monotonicity #13578 #13587]: #13589
-#13593 := [monotonicity #13590]: #13592
-#13596 := [monotonicity #13593]: #13595
-#13599 := [monotonicity #13596]: #13598
-#13602 := [monotonicity #13599]: #13601
-#13607 := [trans #13602 #13605]: #13606
-#13610 := [monotonicity #13607]: #13609
-#13615 := [trans #13610 #13613]: #13614
-#13618 := [monotonicity #13615]: #13617
-#13621 := [monotonicity #13618]: #13620
-#13626 := [trans #13621 #13624]: #13625
-#13642 := [monotonicity #13443 #13626]: #13641
-#13645 := [monotonicity #13443 #13642]: #13644
-#13638 := (iff #12647 #13628)
-#13633 := (not #13627)
-#13636 := (iff #13633 #13628)
-#13637 := [rewrite]: #13636
-#13634 := (iff #12647 #13633)
-#13631 := (iff #3198 #13627)
-#13632 := [rewrite]: #13631
-#13635 := [monotonicity #13632]: #13634
-#13639 := [trans #13635 #13637]: #13638
-#13648 := [monotonicity #13639 #13645]: #13647
-#13651 := [monotonicity #13443 #13648]: #13650
-#13656 := [trans #13651 #13654]: #13655
-#13689 := [monotonicity #13656 #13686]: #13688
-#13692 := [monotonicity #13443 #13689]: #13691
-#13695 := [monotonicity #13692]: #13694
-#13700 := [trans #13695 #13698]: #13699
-#13703 := [monotonicity #13700]: #13702
-#13708 := [trans #13703 #13706]: #13707
-#13711 := [monotonicity #13708]: #13710
-#13714 := [monotonicity #13711]: #13713
-#13719 := [trans #13714 #13717]: #13718
-#13734 := [monotonicity #13443 #13719]: #13733
-#13737 := [monotonicity #13443 #13734]: #13736
-#13730 := (iff #12813 #13721)
-#13725 := (not #13722)
-#13728 := (iff #13725 #13721)
-#13729 := [rewrite]: #13728
-#13726 := (iff #12813 #13725)
-#13723 := (iff #3187 #13722)
-#13724 := [rewrite]: #13723
-#13727 := [monotonicity #13724]: #13726
-#13731 := [trans #13727 #13729]: #13730
-#13740 := [monotonicity #13731 #13737]: #13739
-#13743 := [monotonicity #13443 #13740]: #13742
-#13748 := [trans #13743 #13746]: #13747
-#13872 := [monotonicity #13748 #13869]: #13871
-#14002 := [monotonicity #13443 #13872]: #14001
-#14005 := [monotonicity #14002]: #14004
-#14008 := [monotonicity #14005]: #14007
-#14011 := [monotonicity #14008]: #14010
-#14014 := [monotonicity #14011]: #14013
-#14017 := [monotonicity #14014]: #14016
-#14020 := [monotonicity #14017]: #14019
-#13895 := (iff #12955 #13894)
-#13892 := (iff #12336 #12330)
-#13884 := (and true #12330)
-#13887 := (and true #13884)
-#13890 := (iff #13887 #12330)
-#13891 := [rewrite]: #13890
-#13888 := (iff #12336 #13887)
-#13885 := (iff #12333 #13884)
-#13880 := (iff #3168 true)
-#13875 := (forall (vars (?x784 T5)) (:pat #3166) true)
-#13878 := (iff #13875 true)
-#13879 := [elim-unused]: #13878
-#13876 := (iff #3168 #13875)
-#13873 := (iff #3167 true)
-#13874 := [rewrite]: #13873
-#13877 := [quant-intro #13874]: #13876
-#13881 := [trans #13877 #13879]: #13880
-#13886 := [monotonicity #13881]: #13885
-#13882 := (iff #3164 true)
-#13883 := [rewrite]: #13882
-#13889 := [monotonicity #13883 #13886]: #13888
-#13893 := [trans #13889 #13891]: #13892
-#13896 := [monotonicity #13893]: #13895
-#14023 := [monotonicity #13896 #14020]: #14022
-#14026 := [monotonicity #13896 #14023]: #14025
-#14029 := [monotonicity #13443 #14026]: #14028
-#14032 := [monotonicity #13443 #14029]: #14031
-#14035 := [monotonicity #13443 #14032]: #14034
-#14038 := [monotonicity #13443 #14035]: #14037
-#13912 := (iff #13032 #13911)
-#13909 := (iff #12019 #13906)
-#13903 := (and #13900 #12016)
-#13907 := (iff #13903 #13906)
-#13908 := [rewrite]: #13907
-#13904 := (iff #12019 #13903)
-#13901 := (iff #3081 #13900)
-#13902 := [rewrite]: #13901
-#13905 := [monotonicity #13902]: #13904
-#13910 := [trans #13905 #13908]: #13909
-#13913 := [monotonicity #13910]: #13912
-#14041 := [monotonicity #13913 #14038]: #14040
-#13946 := (iff #13041 #13945)
-#13943 := (iff #12013 #13942)
-#13940 := (iff #12008 #13937)
-#13931 := (or #13914 #13928)
-#13934 := (or #5606 #13931)
-#13938 := (iff #13934 #13937)
-#13939 := [rewrite]: #13938
-#13935 := (iff #12008 #13934)
-#13932 := (iff #12002 #13931)
-#13929 := (iff #3077 #13928)
-#13930 := [rewrite]: #13929
-#13924 := (iff #12001 #13914)
-#13916 := (not #13914)
-#13919 := (not #13916)
-#13922 := (iff #13919 #13914)
-#13923 := [rewrite]: #13922
-#13920 := (iff #12001 #13919)
-#13917 := (iff #3076 #13916)
-#13918 := [rewrite]: #13917
-#13921 := [monotonicity #13918]: #13920
-#13925 := [trans #13921 #13923]: #13924
-#13933 := [monotonicity #13925 #13930]: #13932
-#13936 := [monotonicity #5608 #13933]: #13935
-#13941 := [trans #13936 #13939]: #13940
-#13944 := [quant-intro #13941]: #13943
-#13947 := [monotonicity #13944]: #13946
-#14044 := [monotonicity #13947 #14041]: #14043
-#13952 := (iff #13050 #13951)
-#13949 := (iff #3075 #13948)
-#13950 := [rewrite]: #13949
-#13953 := [monotonicity #13950]: #13952
-#14047 := [monotonicity #13953 #14044]: #14046
-#14050 := [monotonicity #13443 #14047]: #14049
-#13964 := (iff #13067 #13963)
-#13961 := (iff #3072 #13960)
-#13958 := (iff #3071 #13957)
-#13959 := [rewrite]: #13958
-#13955 := (iff #3070 #13954)
-#13956 := [rewrite]: #13955
-#13962 := [monotonicity #13956 #13959]: #13961
-#13965 := [monotonicity #13962]: #13964
-#14053 := [monotonicity #13965 #14050]: #14052
-#13974 := (iff #13076 #13973)
-#13971 := (iff #3068 #13970)
-#13968 := (iff #3067 #13966)
-#13969 := [rewrite]: #13968
-#13972 := [monotonicity #13434 #13969]: #13971
-#13975 := [monotonicity #13972]: #13974
-#14056 := [monotonicity #13975 #14053]: #14055
-#13987 := (iff #13085 #13986)
-#13984 := (iff #3064 #13983)
-#13981 := (iff #3063 #13979)
-#13982 := [rewrite]: #13981
-#13977 := (iff #3062 #13976)
-#13978 := [rewrite]: #13977
-#13985 := [monotonicity #13978 #13982]: #13984
-#13988 := [monotonicity #13985]: #13987
-#14059 := [monotonicity #13988 #14056]: #14058
-#13998 := (iff #13101 #13997)
-#13995 := (iff #11995 #13992)
-#13993 := (iff #13989 #13992)
-#13994 := [rewrite]: #13993
-#13990 := (iff #11995 #13989)
-#13430 := (iff #2975 #13429)
-#13431 := [rewrite]: #13430
-#13991 := [monotonicity #13431]: #13990
-#13996 := [trans #13991 #13994]: #13995
-#13999 := [monotonicity #13996]: #13998
-#14062 := [monotonicity #13999 #14059]: #14061
-#14067 := [trans #14062 #14065]: #14066
-#14070 := [monotonicity #13991 #14067]: #14069
-#14075 := [trans #14070 #14073]: #14074
-#13426 := (iff #13113 #13425)
-#13423 := (iff #11990 #13422)
-#13420 := (iff #11985 #13417)
-#13411 := (or #13395 #13407)
-#13414 := (or #5606 #13411)
-#13418 := (iff #13414 #13417)
-#13419 := [rewrite]: #13418
-#13415 := (iff #11985 #13414)
-#13412 := (iff #11979 #13411)
-#13406 := (iff #3055 #13407)
-#13410 := [rewrite]: #13406
-#13404 := (iff #11978 #13395)
-#13396 := (not #13395)
-#13399 := (not #13396)
-#13402 := (iff #13399 #13395)
-#13403 := [rewrite]: #13402
-#13400 := (iff #11978 #13399)
-#13397 := (iff #3052 #13396)
-#13398 := [rewrite]: #13397
-#13401 := [monotonicity #13398]: #13400
-#13405 := [trans #13401 #13403]: #13404
-#13413 := [monotonicity #13405 #13410]: #13412
-#13416 := [monotonicity #5608 #13413]: #13415
-#13421 := [trans #13416 #13419]: #13420
-#13424 := [quant-intro #13421]: #13423
-#13427 := [monotonicity #13424]: #13426
-#14078 := [monotonicity #13427 #14075]: #14077
-#14081 := [monotonicity #13424 #14078]: #14080
-#13393 := (iff #13125 #13392)
-#13390 := (iff #3051 #13389)
-#13391 := [rewrite]: #13390
-#13394 := [monotonicity #13391]: #13393
-#14084 := [monotonicity #13394 #14081]: #14083
-#14087 := [monotonicity #13391 #14084]: #14086
-#13387 := (iff #13137 false)
-#12022 := (iff #3086 false)
-#12023 := [rewrite]: #12022
-#13385 := (iff #13137 #3086)
-#13383 := (iff #11975 true)
-#12362 := (and true true)
-#13378 := (and true #12362)
-#13381 := (iff #13378 true)
-#13382 := [rewrite]: #13381
-#13379 := (iff #11975 #13378)
-#13376 := (iff #11972 #12362)
-#13374 := (iff #3047 true)
-#13375 := [rewrite]: #13374
-#13372 := (iff #3046 true)
-#13373 := [rewrite]: #13372
-#13377 := [monotonicity #13373 #13375]: #13376
-#13380 := [monotonicity #13373 #13377]: #13379
-#13384 := [trans #13380 #13382]: #13383
-#13386 := [monotonicity #13384]: #13385
-#13388 := [trans #13386 #12023]: #13387
-#14090 := [monotonicity #13388 #14087]: #14089
-#14093 := [monotonicity #14090]: #14092
-#14096 := [monotonicity #14093]: #14095
-#14099 := [monotonicity #14096]: #14098
-#14102 := [monotonicity #14099]: #14101
-#14105 := [monotonicity #14102]: #14104
-#14110 := [trans #14105 #14108]: #14109
-#14113 := [monotonicity #14110]: #14112
-#14118 := [trans #14113 #14116]: #14117
-#14121 := [monotonicity #14118]: #14120
-#14124 := [monotonicity #14121]: #14123
-#14129 := [trans #14124 #14127]: #14128
-#14132 := [monotonicity #14129]: #14131
-#14135 := [monotonicity #14132]: #14134
-#13370 := (iff #13218 #13369)
-#13367 := (iff #3028 #13366)
-#13364 := (iff #3027 #13361)
-#13365 := [rewrite]: #13364
-#13358 := (iff #3026 #13359)
-#13360 := [rewrite]: #13358
-#13368 := [monotonicity #13360 #13365]: #13367
-#13371 := [monotonicity #13368]: #13370
-#14216 := [monotonicity #13371 #14135]: #14215
-#14219 := [monotonicity #14216]: #14218
-#14222 := [monotonicity #14219]: #14221
-#14225 := [monotonicity #14222]: #14224
-#14228 := [monotonicity #14225]: #14227
-#14146 := (iff #13263 #14145)
-#14143 := (iff #3012 #14142)
-#14140 := (iff #3011 #14136)
-#14141 := [rewrite]: #14140
-#14144 := [quant-intro #14141]: #14143
-#14147 := [monotonicity #14144]: #14146
-#14231 := [monotonicity #14147 #14228]: #14230
-#14234 := [monotonicity #14231]: #14233
-#14237 := [monotonicity #14234]: #14236
-#14152 := (iff #13297 #14151)
-#14149 := (iff #11923 #14148)
-#14150 := [rewrite]: #14149
-#14153 := [monotonicity #14150]: #14152
-#14240 := [monotonicity #14153 #14237]: #14239
-#14159 := (iff #13306 #13428)
-#14154 := (not #13429)
-#14157 := (iff #14154 #13428)
-#14158 := [rewrite]: #14157
-#14155 := (iff #13306 #14154)
-#14156 := [monotonicity #13431]: #14155
-#14160 := [trans #14156 #14158]: #14159
-#14243 := [monotonicity #14160 #14240]: #14242
-#14170 := (iff #13315 #14161)
-#14162 := (not #14161)
-#14165 := (not #14162)
-#14168 := (iff #14165 #14161)
-#14169 := [rewrite]: #14168
-#14166 := (iff #13315 #14165)
-#14163 := (iff #2974 #14162)
-#14164 := [rewrite]: #14163
-#14167 := [monotonicity #14164]: #14166
-#14171 := [trans #14167 #14169]: #14170
-#14246 := [monotonicity #14171 #14243]: #14245
-#14184 := (iff #13324 #14183)
-#14181 := (iff #2972 #14180)
-#14178 := (iff #2971 #14175)
-#14179 := [rewrite]: #14178
-#14172 := (iff #2970 #14173)
-#14174 := [rewrite]: #14172
-#14182 := [monotonicity #14174 #14179]: #14181
-#14185 := [monotonicity #14182]: #14184
-#14249 := [monotonicity #14185 #14246]: #14248
-#14198 := (iff #13333 #14197)
-#14195 := (iff #2968 #14194)
-#14192 := (iff #2967 #14189)
-#14193 := [rewrite]: #14192
-#14186 := (iff #2966 #14187)
-#14188 := [rewrite]: #14186
-#14196 := [monotonicity #14188 #14193]: #14195
-#14199 := [monotonicity #14196]: #14198
-#14252 := [monotonicity #14199 #14249]: #14251
-#14212 := (iff #13342 #14211)
-#14209 := (iff #2964 #14208)
-#14206 := (iff #2963 #14203)
-#14207 := [rewrite]: #14206
-#14200 := (iff #2962 #14201)
-#14202 := [rewrite]: #14200
-#14210 := [monotonicity #14202 #14207]: #14209
-#14213 := [monotonicity #14210]: #14212
-#14255 := [monotonicity #14213 #14252]: #14254
-#14260 := [trans #14255 #14258]: #14259
-#14263 := [monotonicity #14260]: #14262
-#13356 := (iff #3354 #13355)
-#13353 := (iff #3353 #13343)
-#13348 := (implies true #13343)
-#13351 := (iff #13348 #13343)
-#13352 := [rewrite]: #13351
-#13349 := (iff #3353 #13348)
-#13346 := (iff #3352 #13343)
-#13339 := (implies #2964 #13334)
-#13344 := (iff #13339 #13343)
-#13345 := [rewrite]: #13344
-#13340 := (iff #3352 #13339)
-#13337 := (iff #3351 #13334)
-#13330 := (implies #2968 #13325)
-#13335 := (iff #13330 #13334)
-#13336 := [rewrite]: #13335
-#13331 := (iff #3351 #13330)
-#13328 := (iff #3350 #13325)
-#13321 := (implies #2972 #13316)
-#13326 := (iff #13321 #13325)
-#13327 := [rewrite]: #13326
-#13322 := (iff #3350 #13321)
-#13319 := (iff #3349 #13316)
-#13312 := (implies #2974 #13307)
-#13317 := (iff #13312 #13316)
-#13318 := [rewrite]: #13317
-#13313 := (iff #3349 #13312)
-#13310 := (iff #3348 #13307)
-#13303 := (implies #2975 #13298)
-#13308 := (iff #13303 #13307)
-#13309 := [rewrite]: #13308
-#13304 := (iff #3348 #13303)
-#13301 := (iff #3347 #13298)
-#13294 := (implies #11923 #13282)
-#13299 := (iff #13294 #13298)
-#13300 := [rewrite]: #13299
-#13295 := (iff #3347 #13294)
-#13292 := (iff #3346 #13282)
-#13287 := (implies true #13282)
-#13290 := (iff #13287 #13282)
-#13291 := [rewrite]: #13290
-#13288 := (iff #3346 #13287)
-#13285 := (iff #3345 #13282)
-#13278 := (implies #11926 #13273)
-#13283 := (iff #13278 #13282)
-#13284 := [rewrite]: #13283
-#13279 := (iff #3345 #13278)
-#13276 := (iff #3344 #13273)
-#13269 := (implies #11935 #13264)
-#13274 := (iff #13269 #13273)
-#13275 := [rewrite]: #13274
-#13270 := (iff #3344 #13269)
-#13267 := (iff #3343 #13264)
-#13260 := (implies #3012 #13255)
-#13265 := (iff #13260 #13264)
-#13266 := [rewrite]: #13265
-#13261 := (iff #3343 #13260)
-#13258 := (iff #3342 #13255)
-#13251 := (implies #3016 #13246)
-#13256 := (iff #13251 #13255)
-#13257 := [rewrite]: #13256
-#13252 := (iff #3342 #13251)
-#13249 := (iff #3341 #13246)
-#13242 := (implies #3017 #13237)
-#13247 := (iff #13242 #13246)
-#13248 := [rewrite]: #13247
-#13243 := (iff #3341 #13242)
-#13240 := (iff #3340 #13237)
-#13233 := (implies #3020 #13228)
-#13238 := (iff #13233 #13237)
-#13239 := [rewrite]: #13238
-#13234 := (iff #3340 #13233)
-#13231 := (iff #3339 #13228)
-#13224 := (implies #11949 #13219)
-#13229 := (iff #13224 #13228)
-#13230 := [rewrite]: #13229
-#13225 := (iff #3339 #13224)
-#13222 := (iff #3338 #13219)
-#13215 := (implies #3028 #13212)
-#13220 := (iff #13215 #13219)
-#13221 := [rewrite]: #13220
-#13216 := (iff #3338 #13215)
-#13213 := (iff #3337 #13212)
-#13210 := (iff #3336 #13207)
-#13203 := (implies #11952 #13200)
-#13208 := (iff #13203 #13207)
-#13209 := [rewrite]: #13208
-#13204 := (iff #3336 #13203)
-#13201 := (iff #3335 #13200)
-#13198 := (iff #3334 #13195)
-#13191 := (implies #11961 #13188)
-#13196 := (iff #13191 #13195)
-#13197 := [rewrite]: #13196
-#13192 := (iff #3334 #13191)
-#13189 := (iff #3333 #13188)
-#13186 := (iff #3332 #13183)
-#13179 := (implies #11967 #13174)
-#13184 := (iff #13179 #13183)
-#13185 := [rewrite]: #13184
-#13180 := (iff #3332 #13179)
-#13177 := (iff #3331 #13174)
-#13170 := (implies #3042 #13165)
-#13175 := (iff #13170 #13174)
-#13176 := [rewrite]: #13175
-#13171 := (iff #3331 #13170)
-#13168 := (iff #3330 #13165)
-#13161 := (implies #3043 #13156)
-#13166 := (iff #13161 #13165)
-#13167 := [rewrite]: #13166
-#13162 := (iff #3330 #13161)
-#13159 := (iff #3329 #13156)
-#13152 := (implies #3044 #13147)
-#13157 := (iff #13152 #13156)
-#13158 := [rewrite]: #13157
-#13153 := (iff #3329 #13152)
-#13150 := (iff #3328 #13147)
-#13143 := (implies #3045 #13138)
-#13148 := (iff #13143 #13147)
-#13149 := [rewrite]: #13148
-#13144 := (iff #3328 #13143)
-#13141 := (iff #3327 #13138)
-#13134 := (implies #11975 #13131)
-#13139 := (iff #13134 #13138)
-#13140 := [rewrite]: #13139
-#13135 := (iff #3327 #13134)
-#13132 := (iff #3326 #13131)
-#13129 := (iff #3325 #13126)
-#13122 := (implies #3051 #13119)
-#13127 := (iff #13122 #13126)
-#13128 := [rewrite]: #13127
-#13123 := (iff #3325 #13122)
-#13120 := (iff #3324 #13119)
-#13117 := (iff #3323 #13114)
-#13110 := (implies #11990 #13107)
-#13115 := (iff #13110 #13114)
-#13116 := [rewrite]: #13115
-#13111 := (iff #3323 #13110)
-#13108 := (iff #3322 #13107)
-#13105 := (iff #3321 #13102)
-#13098 := (implies #11995 #13086)
-#13103 := (iff #13098 #13102)
-#13104 := [rewrite]: #13103
-#13099 := (iff #3321 #13098)
-#13096 := (iff #3320 #13086)
-#13091 := (implies true #13086)
-#13094 := (iff #13091 #13086)
-#13095 := [rewrite]: #13094
-#13092 := (iff #3320 #13091)
-#13089 := (iff #3319 #13086)
-#13082 := (implies #3064 #13077)
-#13087 := (iff #13082 #13086)
-#13088 := [rewrite]: #13087
-#13083 := (iff #3319 #13082)
-#13080 := (iff #3318 #13077)
-#13073 := (implies #3068 #13068)
-#13078 := (iff #13073 #13077)
-#13079 := [rewrite]: #13078
-#13074 := (iff #3318 #13073)
-#13071 := (iff #3317 #13068)
-#13064 := (implies #3072 #13059)
-#13069 := (iff #13064 #13068)
-#13070 := [rewrite]: #13069
-#13065 := (iff #3317 #13064)
-#13062 := (iff #3316 #13059)
-#13056 := (implies #11998 #13051)
-#13060 := (iff #13056 #13059)
-#13061 := [rewrite]: #13060
-#13057 := (iff #3316 #13056)
-#13054 := (iff #3315 #13051)
-#13047 := (implies #3075 #13042)
-#13052 := (iff #13047 #13051)
-#13053 := [rewrite]: #13052
-#13048 := (iff #3315 #13047)
-#13045 := (iff #3314 #13042)
-#13038 := (implies #12013 #13033)
-#13043 := (iff #13038 #13042)
-#13044 := [rewrite]: #13043
-#13039 := (iff #3314 #13038)
-#13036 := (iff #3313 #13033)
-#13029 := (implies #12019 #13024)
-#13034 := (iff #13029 #13033)
-#13035 := [rewrite]: #13034
-#13030 := (iff #3313 #13029)
-#13027 := (iff #3312 #13024)
-#13021 := (implies #11998 #13002)
-#13025 := (iff #13021 #13024)
-#13026 := [rewrite]: #13025
-#13022 := (iff #3312 #13021)
-#13019 := (iff #3311 #13002)
-#13014 := (and true #13002)
-#13017 := (iff #13014 #13002)
-#13018 := [rewrite]: #13017
-#13015 := (iff #3311 #13014)
-#13012 := (iff #3310 #13002)
-#13007 := (implies true #13002)
-#13010 := (iff #13007 #13002)
-#13011 := [rewrite]: #13010
-#13008 := (iff #3310 #13007)
-#13005 := (iff #3309 #13002)
-#12999 := (implies #11998 #12987)
-#13003 := (iff #12999 #13002)
-#13004 := [rewrite]: #13003
-#13000 := (iff #3309 #12999)
-#12997 := (iff #3308 #12987)
-#12992 := (implies true #12987)
-#12995 := (iff #12992 #12987)
-#12996 := [rewrite]: #12995
-#12993 := (iff #3308 #12992)
-#12990 := (iff #3307 #12987)
-#12984 := (implies #11998 #12972)
-#12988 := (iff #12984 #12987)
-#12989 := [rewrite]: #12988
-#12985 := (iff #3307 #12984)
-#12982 := (iff #3306 #12972)
-#12977 := (implies true #12972)
-#12980 := (iff #12977 #12972)
-#12981 := [rewrite]: #12980
-#12978 := (iff #3306 #12977)
-#12975 := (iff #3305 #12972)
-#12969 := (implies #11998 #12964)
-#12973 := (iff #12969 #12972)
-#12974 := [rewrite]: #12973
-#12970 := (iff #3305 #12969)
-#12967 := (iff #3304 #12964)
-#12961 := (implies #12336 #12956)
-#12965 := (iff #12961 #12964)
-#12966 := [rewrite]: #12965
-#12962 := (iff #3304 #12961)
-#12959 := (iff #3303 #12956)
-#12952 := (implies #12336 #12947)
-#12957 := (iff #12952 #12956)
-#12958 := [rewrite]: #12957
-#12953 := (iff #3303 #12952)
-#12950 := (iff #3302 #12947)
-#12944 := (implies #12030 #12939)
-#12948 := (iff #12944 #12947)
-#12949 := [rewrite]: #12948
-#12945 := (iff #3302 #12944)
-#12942 := (iff #3301 #12939)
-#12935 := (implies #3177 #12930)
-#12940 := (iff #12935 #12939)
-#12941 := [rewrite]: #12940
-#12936 := (iff #3301 #12935)
-#12933 := (iff #3300 #12930)
-#12926 := (implies #3178 #12921)
-#12931 := (iff #12926 #12930)
-#12932 := [rewrite]: #12931
-#12927 := (iff #3300 #12926)
-#12924 := (iff #3299 #12921)
-#12917 := (implies #3179 #12912)
-#12922 := (iff #12917 #12921)
-#12923 := [rewrite]: #12922
-#12918 := (iff #3299 #12917)
-#12915 := (iff #3298 #12912)
-#12908 := (implies #3180 #12903)
-#12913 := (iff #12908 #12912)
-#12914 := [rewrite]: #12913
-#12909 := (iff #3298 #12908)
-#12906 := (iff #3297 #12903)
-#12899 := (implies #3183 #12887)
-#12904 := (iff #12899 #12903)
-#12905 := [rewrite]: #12904
-#12900 := (iff #3297 #12899)
-#12897 := (iff #3296 #12887)
-#12892 := (implies true #12887)
-#12895 := (iff #12892 #12887)
-#12896 := [rewrite]: #12895
-#12893 := (iff #3296 #12892)
-#12890 := (iff #3295 #12887)
-#12884 := (implies #11998 #12881)
-#12888 := (iff #12884 #12887)
-#12889 := [rewrite]: #12888
-#12885 := (iff #3295 #12884)
-#12882 := (iff #3294 #12881)
-#12879 := (iff #3293 #12869)
-#12874 := (implies true #12869)
-#12877 := (iff #12874 #12869)
-#12878 := [rewrite]: #12877
-#12875 := (iff #3293 #12874)
-#12872 := (iff #3292 #12869)
-#12866 := (implies #11998 #12861)
-#12870 := (iff #12866 #12869)
-#12871 := [rewrite]: #12870
-#12867 := (iff #3292 #12866)
-#12864 := (iff #3291 #12861)
-#12857 := (implies #3287 #12852)
-#12862 := (iff #12857 #12861)
-#12863 := [rewrite]: #12862
-#12858 := (iff #3291 #12857)
-#12855 := (iff #3290 #12852)
-#12849 := (implies #11998 #12837)
-#12853 := (iff #12849 #12852)
-#12854 := [rewrite]: #12853
-#12850 := (iff #3290 #12849)
-#12847 := (iff #3289 #12837)
-#12842 := (implies true #12837)
-#12845 := (iff #12842 #12837)
-#12846 := [rewrite]: #12845
-#12843 := (iff #3289 #12842)
-#12840 := (iff #3288 #12837)
-#12834 := (implies #11998 #12195)
-#12838 := (iff #12834 #12837)
-#12839 := [rewrite]: #12838
-#12835 := (iff #3288 #12834)
-#12198 := (iff #3126 #12195)
-#12192 := (implies #11998 #12180)
-#12196 := (iff #12192 #12195)
-#12197 := [rewrite]: #12196
-#12193 := (iff #3126 #12192)
-#12190 := (iff #3125 #12180)
-#12185 := (implies true #12180)
-#12188 := (iff #12185 #12180)
-#12189 := [rewrite]: #12188
-#12186 := (iff #3125 #12185)
-#12183 := (iff #3124 #12180)
-#12177 := (implies #11998 #12174)
-#12181 := (iff #12177 #12180)
-#12182 := [rewrite]: #12181
-#12178 := (iff #3124 #12177)
-#12175 := (iff #3123 #12174)
-#12172 := (iff #3122 #12169)
-#12165 := (implies up_216 #12160)
-#12170 := (iff #12165 #12169)
-#12171 := [rewrite]: #12170
-#12166 := (iff #3122 #12165)
-#12163 := (iff #3121 #12160)
-#12157 := (implies #11998 #12145)
-#12161 := (iff #12157 #12160)
-#12162 := [rewrite]: #12161
-#12158 := (iff #3121 #12157)
-#12155 := (iff #3120 #12145)
-#12150 := (implies true #12145)
-#12153 := (iff #12150 #12145)
-#12154 := [rewrite]: #12153
-#12151 := (iff #3120 #12150)
-#12148 := (iff #3119 #12145)
-#12142 := (implies #11998 #12137)
-#12146 := (iff #12142 #12145)
-#12147 := [rewrite]: #12146
-#12143 := (iff #3119 #12142)
-#12140 := (iff #3118 #12137)
-#12133 := (implies #11998 #12121)
-#12138 := (iff #12133 #12137)
-#12139 := [rewrite]: #12138
-#12134 := (iff #3118 #12133)
-#12131 := (iff #3117 #12121)
-#12126 := (implies true #12121)
-#12129 := (iff #12126 #12121)
-#12130 := [rewrite]: #12129
-#12127 := (iff #3117 #12126)
-#12124 := (iff #3116 #12121)
-#12117 := (implies #12035 #12112)
-#12122 := (iff #12117 #12121)
-#12123 := [rewrite]: #12122
-#12118 := (iff #3116 #12117)
-#12115 := (iff #3115 #12112)
-#12108 := (implies #12038 #12103)
-#12113 := (iff #12108 #12112)
-#12114 := [rewrite]: #12113
-#12109 := (iff #3115 #12108)
-#12106 := (iff #3114 #12103)
-#12099 := (implies #12041 #12094)
-#12104 := (iff #12099 #12103)
-#12105 := [rewrite]: #12104
-#12100 := (iff #3114 #12099)
-#12097 := (iff #3113 #12094)
-#12090 := (implies #12044 #12080)
-#12095 := (iff #12090 #12094)
-#12096 := [rewrite]: #12095
-#12091 := (iff #3113 #12090)
-#12088 := (iff #3112 #12080)
-#12083 := (implies true #12080)
-#12086 := (iff #12083 #12080)
-#12087 := [rewrite]: #12086
-#12084 := (iff #3112 #12083)
-#12081 := (iff #3111 #12080)
-#12078 := (iff #3110 #12075)
-#12071 := (implies #12059 #3107)
-#12076 := (iff #12071 #12075)
-#12077 := [rewrite]: #12076
-#12072 := (iff #3110 #12071)
-#12069 := (iff #3109 #3107)
-#12064 := (and #3107 true)
-#12067 := (iff #12064 #3107)
-#12068 := [rewrite]: #12067
-#12065 := (iff #3109 #12064)
-#12062 := (iff #3108 true)
-#12063 := [rewrite]: #12062
-#12066 := [monotonicity #12063]: #12065
-#12070 := [trans #12066 #12068]: #12069
-#12060 := (iff #3102 #12059)
-#12057 := (iff #3101 #12054)
-#12051 := (implies #412 #12048)
-#12055 := (iff #12051 #12054)
-#12056 := [rewrite]: #12055
-#12052 := (iff #3101 #12051)
-#12049 := (iff #3100 #12048)
-#12050 := [rewrite]: #12049
-#12053 := [monotonicity #12050]: #12052
-#12058 := [trans #12053 #12056]: #12057
-#12061 := [quant-intro #12058]: #12060
-#12073 := [monotonicity #12061 #12070]: #12072
-#12079 := [trans #12073 #12077]: #12078
-#12082 := [monotonicity #12061 #12079]: #12081
-#12085 := [monotonicity #12082]: #12084
-#12089 := [trans #12085 #12087]: #12088
-#12045 := (iff #3097 #12044)
-#12046 := [rewrite]: #12045
-#12092 := [monotonicity #12046 #12089]: #12091
-#12098 := [trans #12092 #12096]: #12097
-#12042 := (iff #3095 #12041)
-#12043 := [rewrite]: #12042
-#12101 := [monotonicity #12043 #12098]: #12100
-#12107 := [trans #12101 #12105]: #12106
-#12039 := (iff #3093 #12038)
-#12040 := [rewrite]: #12039
-#12110 := [monotonicity #12040 #12107]: #12109
-#12116 := [trans #12110 #12114]: #12115
-#12036 := (iff #3091 #12035)
-#12037 := [rewrite]: #12036
-#12119 := [monotonicity #12037 #12116]: #12118
-#12125 := [trans #12119 #12123]: #12124
-#12128 := [monotonicity #12125]: #12127
-#12132 := [trans #12128 #12130]: #12131
-#11999 := (iff #3074 #11998)
-#12000 := [rewrite]: #11999
-#12135 := [monotonicity #12000 #12132]: #12134
-#12141 := [trans #12135 #12139]: #12140
-#12144 := [monotonicity #12000 #12141]: #12143
-#12149 := [trans #12144 #12147]: #12148
-#12152 := [monotonicity #12149]: #12151
-#12156 := [trans #12152 #12154]: #12155
-#12159 := [monotonicity #12000 #12156]: #12158
-#12164 := [trans #12159 #12162]: #12163
-#12167 := [monotonicity #12164]: #12166
-#12173 := [trans #12167 #12171]: #12172
-#12176 := [monotonicity #12173]: #12175
-#12179 := [monotonicity #12000 #12176]: #12178
-#12184 := [trans #12179 #12182]: #12183
-#12187 := [monotonicity #12184]: #12186
-#12191 := [trans #12187 #12189]: #12190
-#12194 := [monotonicity #12000 #12191]: #12193
-#12199 := [trans #12194 #12197]: #12198
-#12836 := [monotonicity #12000 #12199]: #12835
-#12841 := [trans #12836 #12839]: #12840
-#12844 := [monotonicity #12841]: #12843
-#12848 := [trans #12844 #12846]: #12847
-#12851 := [monotonicity #12000 #12848]: #12850
-#12856 := [trans #12851 #12854]: #12855
-#12859 := [monotonicity #12856]: #12858
-#12865 := [trans #12859 #12863]: #12864
-#12868 := [monotonicity #12000 #12865]: #12867
-#12873 := [trans #12868 #12871]: #12872
-#12876 := [monotonicity #12873]: #12875
-#12880 := [trans #12876 #12878]: #12879
-#12832 := (iff #3286 #12822)
-#12827 := (implies true #12822)
-#12830 := (iff #12827 #12822)
-#12831 := [rewrite]: #12830
-#12828 := (iff #3286 #12827)
-#12825 := (iff #3285 #12822)
-#12819 := (implies #11998 #12814)
-#12823 := (iff #12819 #12822)
-#12824 := [rewrite]: #12823
-#12820 := (iff #3285 #12819)
-#12817 := (iff #3284 #12814)
-#12810 := (implies #3187 #12805)
-#12815 := (iff #12810 #12814)
-#12816 := [rewrite]: #12815
-#12811 := (iff #3284 #12810)
-#12808 := (iff #3283 #12805)
-#12802 := (implies #11998 #12790)
-#12806 := (iff #12802 #12805)
-#12807 := [rewrite]: #12806
-#12803 := (iff #3283 #12802)
-#12800 := (iff #3282 #12790)
-#12795 := (implies true #12790)
-#12798 := (iff #12795 #12790)
-#12799 := [rewrite]: #12798
-#12796 := (iff #3282 #12795)
-#12793 := (iff #3281 #12790)
-#12787 := (implies #11998 #12784)
-#12791 := (iff #12787 #12790)
-#12792 := [rewrite]: #12791
-#12788 := (iff #3281 #12787)
-#12785 := (iff #3280 #12784)
-#12782 := (iff #3279 #12779)
-#12776 := (implies #12375 #12773)
-#12780 := (iff #12776 #12779)
-#12781 := [rewrite]: #12780
-#12777 := (iff #3279 #12776)
-#12774 := (iff #3278 #12773)
-#12771 := (iff #3277 #12768)
-#12765 := (implies #12381 #12760)
-#12769 := (iff #12765 #12768)
-#12770 := [rewrite]: #12769
-#12766 := (iff #3277 #12765)
-#12763 := (iff #3276 #12760)
-#12757 := (implies #11998 #12754)
-#12761 := (iff #12757 #12760)
-#12762 := [rewrite]: #12761
-#12758 := (iff #3276 #12757)
-#12755 := (iff #3275 #12754)
-#12752 := (iff #3274 #12742)
-#12747 := (implies true #12742)
-#12750 := (iff #12747 #12742)
-#12751 := [rewrite]: #12750
-#12748 := (iff #3274 #12747)
-#12745 := (iff #3273 #12742)
-#12739 := (implies #11998 #12734)
-#12743 := (iff #12739 #12742)
-#12744 := [rewrite]: #12743
-#12740 := (iff #3273 #12739)
-#12737 := (iff #3272 #12734)
-#12730 := (implies #3262 #12725)
-#12735 := (iff #12730 #12734)
-#12736 := [rewrite]: #12735
-#12731 := (iff #3272 #12730)
-#12728 := (iff #3271 #12725)
-#12722 := (implies #11998 #12710)
-#12726 := (iff #12722 #12725)
-#12727 := [rewrite]: #12726
-#12723 := (iff #3271 #12722)
-#12720 := (iff #3270 #12710)
-#12715 := (implies true #12710)
-#12718 := (iff #12715 #12710)
-#12719 := [rewrite]: #12718
-#12716 := (iff #3270 #12715)
-#12713 := (iff #3269 #12710)
-#12707 := (implies #11998 #12702)
-#12711 := (iff #12707 #12710)
-#12712 := [rewrite]: #12711
-#12708 := (iff #3269 #12707)
-#12705 := (iff #3268 #12702)
-#12699 := (implies #11998 #12687)
-#12703 := (iff #12699 #12702)
-#12704 := [rewrite]: #12703
-#12700 := (iff #3268 #12699)
-#12697 := (iff #3267 #12687)
-#12692 := (implies true #12687)
-#12695 := (iff #12692 #12687)
-#12696 := [rewrite]: #12695
-#12693 := (iff #3267 #12692)
-#12690 := (iff #3266 #12687)
-#12683 := (implies #12668 #12678)
-#12688 := (iff #12683 #12687)
-#12689 := [rewrite]: #12688
-#12684 := (iff #3266 #12683)
-#12681 := (iff #3265 #12678)
-#12674 := (implies #12671 #12524)
-#12679 := (iff #12674 #12678)
-#12680 := [rewrite]: #12679
-#12675 := (iff #3265 #12674)
-#12534 := (iff #3244 #12524)
-#12529 := (implies true #12524)
-#12532 := (iff #12529 #12524)
-#12533 := [rewrite]: #12532
-#12530 := (iff #3244 #12529)
-#12527 := (iff #3243 #12524)
-#12520 := (implies #3209 #12517)
-#12525 := (iff #12520 #12524)
-#12526 := [rewrite]: #12525
-#12521 := (iff #3243 #12520)
-#12518 := (iff #3242 #12517)
-#12515 := (iff #3241 #12512)
-#12508 := (implies #12404 #12503)
-#12513 := (iff #12508 #12512)
-#12514 := [rewrite]: #12513
-#12509 := (iff #3241 #12508)
-#12506 := (iff #3240 #12503)
-#12499 := (implies #12407 #12494)
-#12504 := (iff #12499 #12503)
-#12505 := [rewrite]: #12504
-#12500 := (iff #3240 #12499)
-#12497 := (iff #3239 #12494)
-#12490 := (implies #3216 #12485)
-#12495 := (iff #12490 #12494)
-#12496 := [rewrite]: #12495
-#12491 := (iff #3239 #12490)
-#12488 := (iff #3238 #12485)
-#12481 := (implies #12410 #12471)
-#12486 := (iff #12481 #12485)
-#12487 := [rewrite]: #12486
-#12482 := (iff #3238 #12481)
-#12479 := (iff #3237 #12471)
-#12474 := (implies true #12471)
-#12477 := (iff #12474 #12471)
-#12478 := [rewrite]: #12477
-#12475 := (iff #3237 #12474)
-#12472 := (iff #3236 #12471)
-#12469 := (iff #3235 #12466)
-#12462 := (implies #3219 #12459)
-#12467 := (iff #12462 #12466)
-#12468 := [rewrite]: #12467
-#12463 := (iff #3235 #12462)
-#12460 := (iff #3234 #12459)
-#12457 := (iff #3233 #12454)
-#12450 := (implies #12425 #12431)
-#12455 := (iff #12450 #12454)
-#12456 := [rewrite]: #12455
-#12451 := (iff #3233 #12450)
-#12448 := (iff #3232 #12431)
-#12443 := (and #12431 true)
-#12446 := (iff #12443 #12431)
-#12447 := [rewrite]: #12446
-#12444 := (iff #3232 #12443)
-#12441 := (iff #3231 true)
-#12436 := (implies #12431 true)
-#12439 := (iff #12436 true)
-#12440 := [rewrite]: #12439
-#12437 := (iff #3231 #12436)
-#12434 := (iff #3230 true)
-#12435 := [rewrite]: #12434
-#12432 := (iff #3229 #12431)
-#12429 := (iff #3228 #12428)
-#12430 := [rewrite]: #12429
-#12433 := [monotonicity #12430]: #12432
-#12438 := [monotonicity #12433 #12435]: #12437
-#12442 := [trans #12438 #12440]: #12441
-#12445 := [monotonicity #12433 #12442]: #12444
-#12449 := [trans #12445 #12447]: #12448
-#12426 := (iff #3224 #12425)
-#12423 := (iff #3223 #12420)
-#12417 := (implies #412 #12414)
-#12421 := (iff #12417 #12420)
-#12422 := [rewrite]: #12421
-#12418 := (iff #3223 #12417)
-#12415 := (iff #3222 #12414)
-#12416 := [rewrite]: #12415
-#12419 := [monotonicity #12416]: #12418
-#12424 := [trans #12419 #12422]: #12423
-#12427 := [quant-intro #12424]: #12426
-#12452 := [monotonicity #12427 #12449]: #12451
-#12458 := [trans #12452 #12456]: #12457
-#12461 := [monotonicity #12427 #12458]: #12460
-#12464 := [monotonicity #12461]: #12463
-#12470 := [trans #12464 #12468]: #12469
-#12473 := [monotonicity #12470]: #12472
-#12476 := [monotonicity #12473]: #12475
-#12480 := [trans #12476 #12478]: #12479
-#12411 := (iff #3218 #12410)
-#12412 := [rewrite]: #12411
-#12483 := [monotonicity #12412 #12480]: #12482
-#12489 := [trans #12483 #12487]: #12488
-#12492 := [monotonicity #12489]: #12491
-#12498 := [trans #12492 #12496]: #12497
-#12408 := (iff #3215 #12407)
-#12396 := (= #3210 #12395)
-#12397 := [rewrite]: #12396
-#12409 := [monotonicity #12397]: #12408
-#12501 := [monotonicity #12409 #12498]: #12500
-#12507 := [trans #12501 #12505]: #12506
-#12405 := (iff #3213 #12404)
-#12402 := (iff #3212 #12401)
-#12403 := [monotonicity #12397]: #12402
-#12399 := (iff #3211 #12398)
-#12400 := [monotonicity #12397]: #12399
-#12406 := [monotonicity #12400 #12403]: #12405
-#12510 := [monotonicity #12406 #12507]: #12509
-#12516 := [trans #12510 #12514]: #12515
-#12519 := [monotonicity #12406 #12516]: #12518
-#12522 := [monotonicity #12519]: #12521
-#12528 := [trans #12522 #12526]: #12527
-#12531 := [monotonicity #12528]: #12530
-#12535 := [trans #12531 #12533]: #12534
-#12672 := (iff #3264 #12671)
-#12673 := [rewrite]: #12672
-#12676 := [monotonicity #12673 #12535]: #12675
-#12682 := [trans #12676 #12680]: #12681
-#12669 := (iff #3263 #12668)
-#12670 := [rewrite]: #12669
-#12685 := [monotonicity #12670 #12682]: #12684
-#12691 := [trans #12685 #12689]: #12690
-#12694 := [monotonicity #12691]: #12693
-#12698 := [trans #12694 #12696]: #12697
-#12701 := [monotonicity #12000 #12698]: #12700
-#12706 := [trans #12701 #12704]: #12705
-#12709 := [monotonicity #12000 #12706]: #12708
-#12714 := [trans #12709 #12712]: #12713
-#12717 := [monotonicity #12714]: #12716
-#12721 := [trans #12717 #12719]: #12720
-#12724 := [monotonicity #12000 #12721]: #12723
-#12729 := [trans #12724 #12727]: #12728
-#12732 := [monotonicity #12729]: #12731
-#12738 := [trans #12732 #12736]: #12737
-#12741 := [monotonicity #12000 #12738]: #12740
-#12746 := [trans #12741 #12744]: #12745
-#12749 := [monotonicity #12746]: #12748
-#12753 := [trans #12749 #12751]: #12752
-#12666 := (iff #3261 #12656)
-#12661 := (implies true #12656)
-#12664 := (iff #12661 #12656)
-#12665 := [rewrite]: #12664
-#12662 := (iff #3261 #12661)
-#12659 := (iff #3260 #12656)
-#12653 := (implies #11998 #12648)
-#12657 := (iff #12653 #12656)
-#12658 := [rewrite]: #12657
-#12654 := (iff #3260 #12653)
-#12651 := (iff #3259 #12648)
-#12644 := (implies #3198 #12639)
-#12649 := (iff #12644 #12648)
-#12650 := [rewrite]: #12649
-#12645 := (iff #3259 #12644)
-#12642 := (iff #3258 #12639)
-#12636 := (implies #11998 #12624)
-#12640 := (iff #12636 #12639)
-#12641 := [rewrite]: #12640
-#12637 := (iff #3258 #12636)
-#12634 := (iff #3257 #12624)
-#12629 := (implies true #12624)
-#12632 := (iff #12629 #12624)
-#12633 := [rewrite]: #12632
-#12630 := (iff #3257 #12629)
-#12627 := (iff #3256 #12624)
-#12621 := (implies #11998 #12618)
-#12625 := (iff #12621 #12624)
-#12626 := [rewrite]: #12625
-#12622 := (iff #3256 #12621)
-#12619 := (iff #3255 #12618)
-#12616 := (iff #3254 #12613)
-#12609 := (implies #12375 #12606)
-#12614 := (iff #12609 #12613)
-#12615 := [rewrite]: #12614
-#12610 := (iff #3254 #12609)
-#12607 := (iff #3253 #12606)
-#12604 := (iff #3252 #12601)
-#12597 := (implies #12381 #12592)
-#12602 := (iff #12597 #12601)
-#12603 := [rewrite]: #12602
-#12598 := (iff #3252 #12597)
-#12595 := (iff #3251 #12592)
-#12588 := (implies #12384 #12583)
-#12593 := (iff #12588 #12592)
-#12594 := [rewrite]: #12593
-#12589 := (iff #3251 #12588)
-#12586 := (iff #3250 #12583)
-#12579 := (implies #3201 #12574)
-#12584 := (iff #12579 #12583)
-#12585 := [rewrite]: #12584
-#12580 := (iff #3250 #12579)
-#12577 := (iff #3249 #12574)
-#12570 := (implies #3202 #12565)
-#12575 := (iff #12570 #12574)
-#12576 := [rewrite]: #12575
-#12571 := (iff #3249 #12570)
-#12568 := (iff #3248 #12565)
-#12561 := (implies #3073 #12549)
-#12566 := (iff #12561 #12565)
-#12567 := [rewrite]: #12566
-#12562 := (iff #3248 #12561)
-#12559 := (iff #3247 #12549)
-#12554 := (implies true #12549)
-#12557 := (iff #12554 #12549)
-#12558 := [rewrite]: #12557
-#12555 := (iff #3247 #12554)
-#12552 := (iff #3246 #12549)
-#12545 := (implies #12389 #12540)
-#12550 := (iff #12545 #12549)
-#12551 := [rewrite]: #12550
-#12546 := (iff #3246 #12545)
-#12543 := (iff #3245 #12540)
-#12536 := (implies #12392 #12524)
-#12541 := (iff #12536 #12540)
-#12542 := [rewrite]: #12541
-#12537 := (iff #3245 #12536)
-#12393 := (iff #3207 #12392)
-#12394 := [rewrite]: #12393
-#12538 := [monotonicity #12394 #12535]: #12537
-#12544 := [trans #12538 #12542]: #12543
-#12390 := (iff #3205 #12389)
-#12391 := [rewrite]: #12390
-#12547 := [monotonicity #12391 #12544]: #12546
-#12553 := [trans #12547 #12551]: #12552
-#12556 := [monotonicity #12553]: #12555
-#12560 := [trans #12556 #12558]: #12559
-#12387 := (iff #3203 #3073)
-#12388 := [rewrite]: #12387
-#12563 := [monotonicity #12388 #12560]: #12562
-#12569 := [trans #12563 #12567]: #12568
-#12572 := [monotonicity #12569]: #12571
-#12578 := [trans #12572 #12576]: #12577
-#12581 := [monotonicity #12578]: #12580
-#12587 := [trans #12581 #12585]: #12586
-#12385 := (iff #3200 #12384)
-#12386 := [rewrite]: #12385
-#12590 := [monotonicity #12386 #12587]: #12589
-#12596 := [trans #12590 #12594]: #12595
-#12382 := (iff #3196 #12381)
-#12379 := (iff #3195 #12378)
-#12380 := [rewrite]: #12379
-#12370 := (iff #3190 #12369)
-#12371 := [rewrite]: #12370
-#12383 := [monotonicity #12371 #12380]: #12382
-#12599 := [monotonicity #12383 #12596]: #12598
-#12605 := [trans #12599 #12603]: #12604
-#12608 := [monotonicity #12383 #12605]: #12607
-#12376 := (iff #3193 #12375)
-#12373 := (iff #3192 #12372)
-#12374 := [rewrite]: #12373
-#12377 := [monotonicity #12371 #12374]: #12376
-#12611 := [monotonicity #12377 #12608]: #12610
-#12617 := [trans #12611 #12615]: #12616
-#12620 := [monotonicity #12377 #12617]: #12619
-#12623 := [monotonicity #12000 #12620]: #12622
-#12628 := [trans #12623 #12626]: #12627
-#12631 := [monotonicity #12628]: #12630
-#12635 := [trans #12631 #12633]: #12634
-#12638 := [monotonicity #12000 #12635]: #12637
-#12643 := [trans #12638 #12641]: #12642
-#12646 := [monotonicity #12643]: #12645
-#12652 := [trans #12646 #12650]: #12651
-#12655 := [monotonicity #12000 #12652]: #12654
-#12660 := [trans #12655 #12658]: #12659
-#12663 := [monotonicity #12660]: #12662
-#12667 := [trans #12663 #12665]: #12666
-#12756 := [monotonicity #12667 #12753]: #12755
-#12759 := [monotonicity #12000 #12756]: #12758
-#12764 := [trans #12759 #12762]: #12763
-#12767 := [monotonicity #12383 #12764]: #12766
-#12772 := [trans #12767 #12770]: #12771
-#12775 := [monotonicity #12383 #12772]: #12774
-#12778 := [monotonicity #12377 #12775]: #12777
-#12783 := [trans #12778 #12781]: #12782
-#12786 := [monotonicity #12377 #12783]: #12785
-#12789 := [monotonicity #12000 #12786]: #12788
-#12794 := [trans #12789 #12792]: #12793
-#12797 := [monotonicity #12794]: #12796
-#12801 := [trans #12797 #12799]: #12800
-#12804 := [monotonicity #12000 #12801]: #12803
-#12809 := [trans #12804 #12807]: #12808
-#12812 := [monotonicity #12809]: #12811
-#12818 := [trans #12812 #12816]: #12817
-#12821 := [monotonicity #12000 #12818]: #12820
-#12826 := [trans #12821 #12824]: #12825
-#12829 := [monotonicity #12826]: #12828
-#12833 := [trans #12829 #12831]: #12832
-#12883 := [monotonicity #12833 #12880]: #12882
-#12886 := [monotonicity #12000 #12883]: #12885
-#12891 := [trans #12886 #12889]: #12890
-#12367 := (iff #3186 true)
-#12365 := (iff #12362 true)
-#12366 := [rewrite]: #12365
-#12363 := (iff #3186 #12362)
-#12360 := (iff #3185 true)
-#12361 := [rewrite]: #12360
-#12358 := (iff #3184 true)
-#12359 := [rewrite]: #12358
-#12364 := [monotonicity #12359 #12361]: #12363
-#12368 := [trans #12364 #12366]: #12367
-#12894 := [monotonicity #12368 #12891]: #12893
-#12898 := [trans #12894 #12896]: #12897
-#12901 := [monotonicity #12898]: #12900
-#12907 := [trans #12901 #12905]: #12906
-#12910 := [monotonicity #12907]: #12909
-#12916 := [trans #12910 #12914]: #12915
-#12919 := [monotonicity #12916]: #12918
-#12925 := [trans #12919 #12923]: #12924
-#12928 := [monotonicity #12925]: #12927
-#12934 := [trans #12928 #12932]: #12933
-#12937 := [monotonicity #12934]: #12936
-#12943 := [trans #12937 #12941]: #12942
-#12033 := (iff #3089 #12030)
-#12027 := (and #12024 #11932)
-#12031 := (iff #12027 #12030)
-#12032 := [rewrite]: #12031
-#12028 := (iff #3089 #12027)
-#11933 := (iff #3005 #11932)
-#11934 := [rewrite]: #11933
-#12025 := (iff #3088 #12024)
-#12026 := [rewrite]: #12025
-#12029 := [monotonicity #12026 #11934]: #12028
-#12034 := [trans #12029 #12032]: #12033
-#12946 := [monotonicity #12034 #12943]: #12945
-#12951 := [trans #12946 #12949]: #12950
-#12337 := (iff #3172 #12336)
-#12334 := (iff #3171 #12333)
-#12331 := (iff #3170 #12330)
-#12332 := [rewrite]: #12331
-#12335 := [monotonicity #12332]: #12334
-#12338 := [monotonicity #12335]: #12337
-#12954 := [monotonicity #12338 #12951]: #12953
-#12960 := [trans #12954 #12958]: #12959
-#12356 := (iff #3176 #12336)
-#12339 := (and true #12336)
-#12342 := (iff #12339 #12336)
-#12343 := [rewrite]: #12342
-#12354 := (iff #3176 #12339)
-#12352 := (iff #3175 #12336)
-#12350 := (iff #3175 #12339)
-#12348 := (iff #3174 #12336)
-#12346 := (iff #3174 #12339)
-#12344 := (iff #3173 #12336)
-#12340 := (iff #3173 #12339)
-#12328 := (iff #3163 true)
-#12323 := (forall (vars (?x783 T5)) (:pat #3159) true)
-#12326 := (iff #12323 true)
-#12327 := [elim-unused]: #12326
-#12324 := (iff #3163 #12323)
-#12321 := (iff #3162 true)
-#12273 := (= uf_9 #3147)
-#12285 := (implies #12273 #12273)
-#12288 := (iff #12285 true)
-#12289 := [rewrite]: #12288
-#12319 := (iff #3162 #12285)
-#12317 := (iff #3161 #12273)
-#12278 := (and true #12273)
-#12281 := (iff #12278 #12273)
-#12282 := [rewrite]: #12281
-#12315 := (iff #3161 #12278)
-#12274 := (iff #3148 #12273)
-#12275 := [rewrite]: #12274
-#12313 := (iff #3160 true)
-#12314 := [rewrite]: #12313
-#12316 := [monotonicity #12314 #12275]: #12315
-#12318 := [trans #12316 #12282]: #12317
-#12320 := [monotonicity #12275 #12318]: #12319
-#12322 := [trans #12320 #12289]: #12321
-#12325 := [quant-intro #12322]: #12324
-#12329 := [trans #12325 #12327]: #12328
-#12341 := [monotonicity #12329 #12338]: #12340
-#12345 := [trans #12341 #12343]: #12344
-#12311 := (iff #3156 true)
-#12266 := (forall (vars (?x780 T5)) (:pat #3136) true)
-#12269 := (iff #12266 true)
-#12270 := [elim-unused]: #12269
-#12309 := (iff #3156 #12266)
-#12307 := (iff #3155 true)
-#12305 := (iff #3155 #12285)
-#12303 := (iff #3154 #12273)
-#12301 := (iff #3154 #12278)
-#12299 := (iff #3153 true)
-#12300 := [rewrite]: #12299
-#12302 := [monotonicity #12300 #12275]: #12301
-#12304 := [trans #12302 #12282]: #12303
-#12306 := [monotonicity #12275 #12304]: #12305
-#12308 := [trans #12306 #12289]: #12307
-#12310 := [quant-intro #12308]: #12309
-#12312 := [trans #12310 #12270]: #12311
-#12347 := [monotonicity #12312 #12345]: #12346
-#12349 := [trans #12347 #12343]: #12348
-#12297 := (iff #3152 true)
-#12292 := (forall (vars (?x781 T5)) (:pat #3146) true)
-#12295 := (iff #12292 true)
-#12296 := [elim-unused]: #12295
-#12293 := (iff #3152 #12292)
-#12290 := (iff #3151 true)
-#12286 := (iff #3151 #12285)
-#12283 := (iff #3150 #12273)
-#12279 := (iff #3150 #12278)
-#12276 := (iff #3149 true)
-#12277 := [rewrite]: #12276
-#12280 := [monotonicity #12277 #12275]: #12279
-#12284 := [trans #12280 #12282]: #12283
-#12287 := [monotonicity #12275 #12284]: #12286
-#12291 := [trans #12287 #12289]: #12290
-#12294 := [quant-intro #12291]: #12293
-#12298 := [trans #12294 #12296]: #12297
-#12351 := [monotonicity #12298 #12349]: #12350
-#12353 := [trans #12351 #12343]: #12352
-#12271 := (iff #3143 true)
-#12267 := (iff #3143 #12266)
-#12264 := (iff #3142 true)
-#12253 := (= uf_261 #3139)
-#12256 := (not #12253)
-#12259 := (implies #12256 #12256)
-#12262 := (iff #12259 true)
-#12263 := [rewrite]: #12262
-#12260 := (iff #3142 #12259)
-#12257 := (iff #3141 #12256)
-#12254 := (iff #3140 #12253)
-#12255 := [rewrite]: #12254
-#12258 := [monotonicity #12255]: #12257
-#12261 := [monotonicity #12258 #12258]: #12260
-#12265 := [trans #12261 #12263]: #12264
-#12268 := [quant-intro #12265]: #12267
-#12272 := [trans #12268 #12270]: #12271
-#12355 := [monotonicity #12272 #12353]: #12354
-#12357 := [trans #12355 #12343]: #12356
-#12963 := [monotonicity #12357 #12960]: #12962
-#12968 := [trans #12963 #12966]: #12967
-#12971 := [monotonicity #12000 #12968]: #12970
-#12976 := [trans #12971 #12974]: #12975
-#12979 := [monotonicity #12976]: #12978
-#12983 := [trans #12979 #12981]: #12982
-#12986 := [monotonicity #12000 #12983]: #12985
-#12991 := [trans #12986 #12989]: #12990
-#12994 := [monotonicity #12991]: #12993
-#12998 := [trans #12994 #12996]: #12997
-#13001 := [monotonicity #12000 #12998]: #13000
-#13006 := [trans #13001 #13004]: #13005
-#13009 := [monotonicity #13006]: #13008
-#13013 := [trans #13009 #13011]: #13012
-#12251 := (iff #3133 true)
-#12246 := (implies true true)
-#12249 := (iff #12246 true)
-#12250 := [rewrite]: #12249
-#12247 := (iff #3133 #12246)
-#12244 := (iff #3132 true)
-#12239 := (implies #11998 true)
-#12242 := (iff #12239 true)
-#12243 := [rewrite]: #12242
-#12240 := (iff #3132 #12239)
-#12237 := (iff #3131 true)
-#12204 := (or #12203 #12195)
-#12212 := (or #12136 #12204)
-#12227 := (or #12136 #12212)
-#12232 := (implies false #12227)
-#12235 := (iff #12232 true)
-#12236 := [rewrite]: #12235
-#12233 := (iff #3131 #12232)
-#12230 := (iff #3130 #12227)
-#12224 := (implies #11998 #12212)
-#12228 := (iff #12224 #12227)
-#12229 := [rewrite]: #12228
-#12225 := (iff #3130 #12224)
-#12222 := (iff #3129 #12212)
-#12217 := (implies true #12212)
-#12220 := (iff #12217 #12212)
-#12221 := [rewrite]: #12220
-#12218 := (iff #3129 #12217)
-#12215 := (iff #3128 #12212)
-#12209 := (implies #11998 #12204)
-#12213 := (iff #12209 #12212)
-#12214 := [rewrite]: #12213
-#12210 := (iff #3128 #12209)
-#12207 := (iff #3127 #12204)
-#12200 := (implies #12030 #12195)
-#12205 := (iff #12200 #12204)
-#12206 := [rewrite]: #12205
-#12201 := (iff #3127 #12200)
-#12202 := [monotonicity #12034 #12199]: #12201
-#12208 := [trans #12202 #12206]: #12207
-#12211 := [monotonicity #12000 #12208]: #12210
-#12216 := [trans #12211 #12214]: #12215
-#12219 := [monotonicity #12216]: #12218
-#12223 := [trans #12219 #12221]: #12222
-#12226 := [monotonicity #12000 #12223]: #12225
-#12231 := [trans #12226 #12229]: #12230
-#12234 := [monotonicity #12023 #12231]: #12233
-#12238 := [trans #12234 #12236]: #12237
-#12241 := [monotonicity #12000 #12238]: #12240
-#12245 := [trans #12241 #12243]: #12244
-#12248 := [monotonicity #12245]: #12247
-#12252 := [trans #12248 #12250]: #12251
-#13016 := [monotonicity #12252 #13013]: #13015
-#13020 := [trans #13016 #13018]: #13019
-#13023 := [monotonicity #12000 #13020]: #13022
-#13028 := [trans #13023 #13026]: #13027
-#12020 := (iff #3085 #12019)
-#12017 := (iff #3084 #12016)
-#12018 := [rewrite]: #12017
-#12021 := [monotonicity #12018]: #12020
-#13031 := [monotonicity #12021 #13028]: #13030
-#13037 := [trans #13031 #13035]: #13036
-#12014 := (iff #3080 #12013)
-#12011 := (iff #3079 #12008)
-#12005 := (implies #412 #12002)
-#12009 := (iff #12005 #12008)
-#12010 := [rewrite]: #12009
-#12006 := (iff #3079 #12005)
-#12003 := (iff #3078 #12002)
-#12004 := [rewrite]: #12003
-#12007 := [monotonicity #12004]: #12006
-#12012 := [trans #12007 #12010]: #12011
-#12015 := [quant-intro #12012]: #12014
-#13040 := [monotonicity #12015 #13037]: #13039
-#13046 := [trans #13040 #13044]: #13045
-#13049 := [monotonicity #13046]: #13048
-#13055 := [trans #13049 #13053]: #13054
-#13058 := [monotonicity #12000 #13055]: #13057
-#13063 := [trans #13058 #13061]: #13062
-#13066 := [monotonicity #13063]: #13065
-#13072 := [trans #13066 #13070]: #13071
-#13075 := [monotonicity #13072]: #13074
-#13081 := [trans #13075 #13079]: #13080
-#13084 := [monotonicity #13081]: #13083
-#13090 := [trans #13084 #13088]: #13089
-#13093 := [monotonicity #13090]: #13092
-#13097 := [trans #13093 #13095]: #13096
-#11996 := (iff #3060 #11995)
-#11993 := (iff #3059 #3042)
-#11994 := [rewrite]: #11993
-#11997 := [monotonicity #11994]: #11996
-#13100 := [monotonicity #11997 #13097]: #13099
-#13106 := [trans #13100 #13104]: #13105
-#13109 := [monotonicity #11997 #13106]: #13108
-#11991 := (iff #3058 #11990)
-#11988 := (iff #3057 #11985)
-#11982 := (implies #412 #11979)
-#11986 := (iff #11982 #11985)
-#11987 := [rewrite]: #11986
-#11983 := (iff #3057 #11982)
-#11980 := (iff #3056 #11979)
-#11981 := [rewrite]: #11980
-#11984 := [monotonicity #11981]: #11983
-#11989 := [trans #11984 #11987]: #11988
-#11992 := [quant-intro #11989]: #11991
-#13112 := [monotonicity #11992 #13109]: #13111
-#13118 := [trans #13112 #13116]: #13117
-#13121 := [monotonicity #11992 #13118]: #13120
-#13124 := [monotonicity #13121]: #13123
-#13130 := [trans #13124 #13128]: #13129
-#13133 := [monotonicity #13130]: #13132
-#11976 := (iff #3050 #11975)
-#11973 := (iff #3049 #11972)
-#11970 := (iff #3048 #3047)
-#11971 := [rewrite]: #11970
-#11974 := [monotonicity #11971]: #11973
-#11977 := [monotonicity #11974]: #11976
-#13136 := [monotonicity #11977 #13133]: #13135
-#13142 := [trans #13136 #13140]: #13141
-#13145 := [monotonicity #13142]: #13144
-#13151 := [trans #13145 #13149]: #13150
-#13154 := [monotonicity #13151]: #13153
-#13160 := [trans #13154 #13158]: #13159
-#13163 := [monotonicity #13160]: #13162
-#13169 := [trans #13163 #13167]: #13168
-#13172 := [monotonicity #13169]: #13171
-#13178 := [trans #13172 #13176]: #13177
-#11968 := (iff #3039 #11967)
-#11965 := (iff #3038 #11964)
-#11966 := [rewrite]: #11965
-#11956 := (iff #3033 #11955)
-#11957 := [rewrite]: #11956
-#11969 := [monotonicity #11957 #11966]: #11968
-#13181 := [monotonicity #11969 #13178]: #13180
-#13187 := [trans #13181 #13185]: #13186
-#13190 := [monotonicity #11969 #13187]: #13189
-#11962 := (iff #3036 #11961)
-#11959 := (iff #3035 #11958)
-#11960 := [rewrite]: #11959
-#11963 := [monotonicity #11957 #11960]: #11962
-#13193 := [monotonicity #11963 #13190]: #13192
-#13199 := [trans #13193 #13197]: #13198
-#13202 := [monotonicity #11963 #13199]: #13201
-#11953 := (iff #3030 #11952)
-#11954 := [rewrite]: #11953
-#13205 := [monotonicity #11954 #13202]: #13204
-#13211 := [trans #13205 #13209]: #13210
-#13214 := [monotonicity #11954 #13211]: #13213
-#13217 := [monotonicity #13214]: #13216
-#13223 := [trans #13217 #13221]: #13222
-#11950 := (iff #3025 #11949)
-#11947 := (iff #3024 #11944)
-#11941 := (iff #11938 false)
-#11945 := (iff #11941 #11944)
-#11946 := [rewrite]: #11945
-#11942 := (iff #3024 #11941)
-#11939 := (iff #3023 #11938)
-#11940 := [rewrite]: #11939
-#11943 := [monotonicity #11940]: #11942
-#11948 := [trans #11943 #11946]: #11947
-#11951 := [quant-intro #11948]: #11950
-#13226 := [monotonicity #11951 #13223]: #13225
-#13232 := [trans #13226 #13230]: #13231
-#13235 := [monotonicity #13232]: #13234
-#13241 := [trans #13235 #13239]: #13240
-#13244 := [monotonicity #13241]: #13243
-#13250 := [trans #13244 #13248]: #13249
-#13253 := [monotonicity #13250]: #13252
-#13259 := [trans #13253 #13257]: #13258
-#13262 := [monotonicity #13259]: #13261
-#13268 := [trans #13262 #13266]: #13267
-#11936 := (iff #3006 #11935)
-#11930 := (iff #3003 #11929)
-#11931 := [rewrite]: #11930
-#11937 := [monotonicity #11931 #11934]: #11936
-#13271 := [monotonicity #11937 #13268]: #13270
-#13277 := [trans #13271 #13275]: #13276
-#11927 := (iff #3001 #11926)
-#11928 := [rewrite]: #11927
-#13280 := [monotonicity #11928 #13277]: #13279
-#13286 := [trans #13280 #13284]: #13285
-#13289 := [monotonicity #13286]: #13288
-#13293 := [trans #13289 #13291]: #13292
-#11924 := (iff #2999 #11923)
-#11921 := (iff #2998 #11920)
-#11918 := (iff #2997 #11917)
-#11915 := (iff #2996 #11914)
-#11912 := (iff #2995 #11911)
-#11909 := (iff #2994 #11908)
-#11910 := [rewrite]: #11909
-#11906 := (iff #2992 #11905)
-#11903 := (iff #2991 #11902)
-#11904 := [rewrite]: #11903
-#11907 := [monotonicity #11904]: #11906
-#11913 := [monotonicity #11907 #11910]: #11912
-#11900 := (iff #2989 #11899)
-#11901 := [rewrite]: #11900
-#11916 := [monotonicity #11901 #11913]: #11915
-#11897 := (iff #2987 #11896)
-#11898 := [rewrite]: #11897
-#11919 := [monotonicity #11898 #11916]: #11918
-#11894 := (iff #2985 #11893)
-#11895 := [rewrite]: #11894
-#11922 := [monotonicity #11895 #11919]: #11921
-#11891 := (iff #2983 #11889)
-#11892 := [rewrite]: #11891
-#11925 := [monotonicity #11892 #11922]: #11924
-#13296 := [monotonicity #11925 #13293]: #13295
-#13302 := [trans #13296 #13300]: #13301
-#13305 := [monotonicity #13302]: #13304
-#13311 := [trans #13305 #13309]: #13310
-#13314 := [monotonicity #13311]: #13313
-#13320 := [trans #13314 #13318]: #13319
-#13323 := [monotonicity #13320]: #13322
-#13329 := [trans #13323 #13327]: #13328
-#13332 := [monotonicity #13329]: #13331
-#13338 := [trans #13332 #13336]: #13337
-#13341 := [monotonicity #13338]: #13340
-#13347 := [trans #13341 #13345]: #13346
-#13350 := [monotonicity #13347]: #13349
-#13354 := [trans #13350 #13352]: #13353
-#13357 := [monotonicity #13354]: #13356
-#14265 := [trans #13357 #14263]: #14264
-#11888 := [asserted]: #3354
-#14266 := [mp #11888 #14265]: #14261
-#14284 := [not-or-elim #14266]: #14148
-#14287 := [and-elim #14284]: #11896
-#233 := (:var 0 T3)
-#2666 := (uf_48 #15 #233)
-#2667 := (pattern #2666)
-#11167 := (= uf_9 #2666)
-#11174 := (not #11167)
-#1259 := (uf_116 #15)
-#2669 := (uf_43 #233 #1259)
-#2670 := (= #15 #2669)
-#11175 := (or #2670 #11174)
-#11180 := (forall (vars (?x710 T5) (?x711 T3)) (:pat #2667) #11175)
-#18175 := (~ #11180 #11180)
-#18173 := (~ #11175 #11175)
-#18174 := [refl]: #18173
-#18176 := [nnf-pos #18174]: #18175
-#2668 := (= #2666 uf_9)
-#2671 := (implies #2668 #2670)
-#2672 := (forall (vars (?x710 T5) (?x711 T3)) (:pat #2667) #2671)
-#11181 := (iff #2672 #11180)
-#11178 := (iff #2671 #11175)
-#11171 := (implies #11167 #2670)
-#11176 := (iff #11171 #11175)
-#11177 := [rewrite]: #11176
-#11172 := (iff #2671 #11171)
-#11169 := (iff #2668 #11167)
-#11170 := [rewrite]: #11169
-#11173 := [monotonicity #11170]: #11172
-#11179 := [trans #11173 #11177]: #11178
-#11182 := [quant-intro #11179]: #11181
-#11166 := [asserted]: #2672
-#11185 := [mp #11166 #11182]: #11180
-#18177 := [mp~ #11185 #18176]: #11180
-#26143 := (not #11896)
-#26156 := (not #11180)
-#26157 := (or #26156 #26143 #26151)
-#26152 := (or #26151 #26143)
-#26158 := (or #26156 #26152)
-#26165 := (iff #26158 #26157)
-#26153 := (or #26143 #26151)
-#26160 := (or #26156 #26153)
-#26163 := (iff #26160 #26157)
-#26164 := [rewrite]: #26163
-#26161 := (iff #26158 #26160)
-#26154 := (iff #26152 #26153)
-#26155 := [rewrite]: #26154
-#26162 := [monotonicity #26155]: #26161
-#26166 := [trans #26162 #26164]: #26165
-#26159 := [quant-inst]: #26158
-#26167 := [mp #26159 #26166]: #26157
-#28515 := [unit-resolution #26167 #18177 #14287]: #26151
-#26726 := [monotonicity #28515 #28515]: #26939
-#26788 := [symm #26726]: #26936
-#26692 := (= uf_9 #26691)
-decl uf_196 :: (-> T4 T5 T5 T2)
-#26689 := (uf_196 uf_287 #26144 #26144)
-#26690 := (= uf_9 #26689)
-#26694 := (iff #26690 #26692)
-#2245 := (:var 0 T16)
-#21 := (:var 2 T5)
-#13 := (:var 3 T4)
-#2256 := (uf_200 #13 #21 #15 #2245)
-#2257 := (pattern #2256)
-#2259 := (uf_196 #13 #21 #15)
-#10130 := (= uf_9 #2259)
-#10126 := (= uf_9 #2256)
-#10133 := (iff #10126 #10130)
-#10136 := (forall (vars (?x586 T4) (?x587 T5) (?x588 T5) (?x589 T16)) (:pat #2257) #10133)
-#17817 := (~ #10136 #10136)
-#17815 := (~ #10133 #10133)
-#17816 := [refl]: #17815
-#17818 := [nnf-pos #17816]: #17817
-#2260 := (= #2259 uf_9)
-#2258 := (= #2256 uf_9)
-#2261 := (iff #2258 #2260)
-#2262 := (forall (vars (?x586 T4) (?x587 T5) (?x588 T5) (?x589 T16)) (:pat #2257) #2261)
-#10137 := (iff #2262 #10136)
-#10134 := (iff #2261 #10133)
-#10131 := (iff #2260 #10130)
-#10132 := [rewrite]: #10131
-#10128 := (iff #2258 #10126)
-#10129 := [rewrite]: #10128
-#10135 := [monotonicity #10129 #10132]: #10134
-#10138 := [quant-intro #10135]: #10137
-#10125 := [asserted]: #2262
-#10141 := [mp #10125 #10138]: #10136
-#17819 := [mp~ #10141 #17818]: #10136
-#26712 := (not #10136)
-#26713 := (or #26712 #26694)
-#26693 := (iff #26692 #26690)
-#26710 := (or #26712 #26693)
-#26715 := (iff #26710 #26713)
-#26717 := (iff #26713 #26713)
-#26697 := [rewrite]: #26717
-#26695 := (iff #26693 #26694)
-#26696 := [rewrite]: #26695
-#26716 := [monotonicity #26696]: #26715
-#26699 := [trans #26716 #26697]: #26715
-#26714 := [quant-inst]: #26710
-#26700 := [mp #26714 #26699]: #26713
-#26910 := [unit-resolution #26700 #17819]: #26694
-#26701 := (not #26694)
-#26729 := (or #26701 #26692)
-#26555 := (uf_13 #26144)
-#26801 := (uf_12 #26555)
-#26804 := (= uf_14 #26801)
-#26923 := (not #26804)
-#26924 := (iff #11905 #26923)
-#26921 := (iff #11902 #26804)
-#26940 := (iff #26804 #11902)
-#26916 := (= #26801 #2990)
-#26914 := (= #26555 #2977)
-#24974 := (uf_13 #2981)
-#28563 := (= #24974 #2977)
-#24977 := (= #2977 #24974)
-#2697 := (uf_43 #326 #161)
-#23148 := (pattern #2697)
-#2701 := (uf_13 #2697)
-#11245 := (= #326 #2701)
-#23155 := (forall (vars (?x720 T3) (?x721 int)) (:pat #23148) #11245)
-#11249 := (forall (vars (?x720 T3) (?x721 int)) #11245)
-#23158 := (iff #11249 #23155)
-#23156 := (iff #11245 #11245)
-#23157 := [refl]: #23156
-#23159 := [quant-intro #23157]: #23158
-#18200 := (~ #11249 #11249)
-#18198 := (~ #11245 #11245)
-#18199 := [refl]: #18198
-#18201 := [nnf-pos #18199]: #18200
-#2702 := (= #2701 #326)
-#2703 := (forall (vars (?x720 T3) (?x721 int)) #2702)
-#11250 := (iff #2703 #11249)
-#11247 := (iff #2702 #11245)
-#11248 := [rewrite]: #11247
-#11251 := [quant-intro #11248]: #11250
-#11244 := [asserted]: #2703
-#11254 := [mp #11244 #11251]: #11249
-#18202 := [mp~ #11254 #18201]: #11249
-#23160 := [mp #18202 #23159]: #23155
-#24921 := (not #23155)
-#24982 := (or #24921 #24977)
-#24983 := [quant-inst]: #24982
-#28497 := [unit-resolution #24983 #23160]: #24977
-#28564 := [symm #28497]: #28563
-#26907 := (= #26555 #24974)
-#28516 := (= #26144 #2981)
-#28517 := [symm #28515]: #28516
-#26913 := [monotonicity #28517]: #26907
-#26915 := [trans #26913 #28564]: #26914
-#26917 := [monotonicity #26915]: #26916
-#26919 := [monotonicity #26917]: #26940
-#26922 := [symm #26919]: #26921
-#26938 := [monotonicity #26922]: #26924
-#14289 := [and-elim #14284]: #11905
-#26948 := [mp #14289 #26938]: #26923
-#26793 := (uf_24 uf_287 #26144)
-#26794 := (= uf_9 #26793)
-#26955 := (= #2988 #26793)
-#26949 := (= #26793 #2988)
-#26930 := [monotonicity #28517]: #26949
-#26958 := [symm #26930]: #26955
-#14288 := [and-elim #14284]: #11899
-#26957 := [trans #14288 #26958]: #26794
-#26796 := (uf_48 #26144 #26555)
-#26797 := (= uf_9 #26796)
-#26962 := (= #2986 #26796)
-#26959 := (= #26796 #2986)
-#26960 := [monotonicity #28517 #26915]: #26959
-#26963 := [symm #26960]: #26962
-#26969 := [trans #14287 #26963]: #26797
-#26798 := (not #26797)
-#26795 := (not #26794)
-#27030 := (or #26795 #26798 #26804)
-#26567 := (uf_25 uf_287 #26144)
-#26799 := (= uf_26 #26567)
-#26995 := (= #2984 #26567)
-#26991 := (= #26567 #2984)
-#26992 := [monotonicity #28517]: #26991
-#26996 := [symm #26992]: #26995
-#14286 := [and-elim #14284]: #11893
-#26990 := [trans #14286 #26996]: #26799
-#26711 := (not #26690)
-#26997 := [hypothesis]: #26711
-#26593 := (uf_27 uf_287 #26144)
-#26594 := (= uf_9 #26593)
-#27000 := (= #2982 #26593)
-#26970 := (= #26593 #2982)
-#26998 := [monotonicity #28517]: #26970
-#27021 := [symm #26998]: #27000
-#14285 := [and-elim #14284]: #11889
-#27022 := [trans #14285 #27021]: #26594
-#26556 := (uf_23 #26555)
-#26563 := (= uf_9 #26556)
-#27046 := (= #2993 #26556)
-#27023 := (= #26556 #2993)
-#27024 := [monotonicity #26915]: #27023
-#27011 := [symm #27024]: #27046
-#14290 := [and-elim #14284]: #11908
-#27047 := [trans #14290 #27011]: #26563
-#14273 := [not-or-elim #14266]: #11935
-#14275 := [and-elim #14273]: #11932
-#2217 := (uf_196 #47 #23 #23)
-#2218 := (pattern #2217)
-#10011 := (= uf_9 #2217)
-#227 := (uf_55 #47)
-#3926 := (= uf_9 #227)
-#19374 := (not #3926)
-#144 := (uf_48 #23 #26)
-#3647 := (= uf_9 #144)
-#19249 := (not #3647)
-#19248 := (not #3644)
-#135 := (uf_27 #47 #23)
-#3629 := (= uf_9 #135)
-#10720 := (not #3629)
-#71 := (uf_23 #26)
-#3482 := (= uf_9 #71)
-#10775 := (not #3482)
-#21807 := (or #29 #10775 #10720 #19248 #19249 #11095 #19374 #10011)
-#21812 := (forall (vars (?x572 T4) (?x573 T5)) (:pat #2218) #21807)
-#10026 := (and #52 #3482 #3629 #3644 #3647 #3650 #3926)
-#10029 := (not #10026)
-#10035 := (or #10011 #10029)
-#10040 := (forall (vars (?x572 T4) (?x573 T5)) (:pat #2218) #10035)
-#21813 := (iff #10040 #21812)
-#21810 := (iff #10035 #21807)
-#21793 := (or #29 #10775 #10720 #19248 #19249 #11095 #19374)
-#21804 := (or #10011 #21793)
-#21808 := (iff #21804 #21807)
-#21809 := [rewrite]: #21808
-#21805 := (iff #10035 #21804)
-#21802 := (iff #10029 #21793)
-#21794 := (not #21793)
-#21797 := (not #21794)
-#21800 := (iff #21797 #21793)
-#21801 := [rewrite]: #21800
-#21798 := (iff #10029 #21797)
-#21795 := (iff #10026 #21794)
-#21796 := [rewrite]: #21795
-#21799 := [monotonicity #21796]: #21798
-#21803 := [trans #21799 #21801]: #21802
-#21806 := [monotonicity #21803]: #21805
-#21811 := [trans #21806 #21809]: #21810
-#21814 := [quant-intro #21811]: #21813
-#17785 := (~ #10040 #10040)
-#17783 := (~ #10035 #10035)
-#17784 := [refl]: #17783
-#17786 := [nnf-pos #17784]: #17785
-#2225 := (= #2217 uf_9)
-#72 := (= #71 uf_9)
-#2219 := (and #52 #72)
-#2220 := (and #147 #2219)
-#145 := (= #144 uf_9)
-#2221 := (and #145 #2220)
-#2222 := (and #143 #2221)
-#136 := (= #135 uf_9)
-#2223 := (and #136 #2222)
-#229 := (= #227 uf_9)
-#2224 := (and #229 #2223)
-#2226 := (implies #2224 #2225)
-#2227 := (forall (vars (?x572 T4) (?x573 T5)) (:pat #2218) #2226)
-#10043 := (iff #2227 #10040)
-#9992 := (and #52 #3482)
-#9996 := (and #3650 #9992)
-#9999 := (and #3647 #9996)
-#10002 := (and #3644 #9999)
-#10005 := (and #3629 #10002)
-#10008 := (and #3926 #10005)
-#10017 := (not #10008)
-#10018 := (or #10017 #10011)
-#10023 := (forall (vars (?x572 T4) (?x573 T5)) (:pat #2218) #10018)
-#10041 := (iff #10023 #10040)
-#10038 := (iff #10018 #10035)
-#10032 := (or #10029 #10011)
-#10036 := (iff #10032 #10035)
-#10037 := [rewrite]: #10036
-#10033 := (iff #10018 #10032)
-#10030 := (iff #10017 #10029)
-#10027 := (iff #10008 #10026)
-#10028 := [rewrite]: #10027
-#10031 := [monotonicity #10028]: #10030
-#10034 := [monotonicity #10031]: #10033
-#10039 := [trans #10034 #10037]: #10038
-#10042 := [quant-intro #10039]: #10041
-#10024 := (iff #2227 #10023)
-#10021 := (iff #2226 #10018)
-#10014 := (implies #10008 #10011)
-#10019 := (iff #10014 #10018)
-#10020 := [rewrite]: #10019
-#10015 := (iff #2226 #10014)
-#10012 := (iff #2225 #10011)
-#10013 := [rewrite]: #10012
-#10009 := (iff #2224 #10008)
-#10006 := (iff #2223 #10005)
-#10003 := (iff #2222 #10002)
-#10000 := (iff #2221 #9999)
-#9997 := (iff #2220 #9996)
-#9994 := (iff #2219 #9992)
-#3483 := (iff #72 #3482)
-#3484 := [rewrite]: #3483
-#9995 := [monotonicity #3484]: #9994
-#9998 := [monotonicity #3652 #9995]: #9997
-#3648 := (iff #145 #3647)
-#3649 := [rewrite]: #3648
-#10001 := [monotonicity #3649 #9998]: #10000
-#10004 := [monotonicity #3646 #10001]: #10003
-#3631 := (iff #136 #3629)
-#3632 := [rewrite]: #3631
-#10007 := [monotonicity #3632 #10004]: #10006
-#3928 := (iff #229 #3926)
-#3929 := [rewrite]: #3928
-#10010 := [monotonicity #3929 #10007]: #10009
-#10016 := [monotonicity #10010 #10013]: #10015
-#10022 := [trans #10016 #10020]: #10021
-#10025 := [quant-intro #10022]: #10024
-#10044 := [trans #10025 #10042]: #10043
-#9991 := [asserted]: #2227
-#10045 := [mp #9991 #10044]: #10040
-#17787 := [mp~ #10045 #17786]: #10040
-#21815 := [mp #17787 #21814]: #21812
-#26800 := (not #26799)
-#26620 := (not #26594)
-#26564 := (not #26563)
-#24694 := (not #11932)
-#26728 := (not #21812)
-#26731 := (or #26728 #24694 #26564 #26620 #26690 #26795 #26798 #26800 #26804)
-#26802 := (= #26801 uf_14)
-#26803 := (or #26802 #26564 #26620 #26800 #26798 #26795 #24694 #26690)
-#26732 := (or #26728 #26803)
-#26783 := (iff #26732 #26731)
-#26810 := (or #24694 #26564 #26620 #26690 #26795 #26798 #26800 #26804)
-#26734 := (or #26728 #26810)
-#26781 := (iff #26734 #26731)
-#26782 := [rewrite]: #26781
-#26785 := (iff #26732 #26734)
-#26813 := (iff #26803 #26810)
-#26807 := (or #26804 #26564 #26620 #26800 #26798 #26795 #24694 #26690)
-#26811 := (iff #26807 #26810)
-#26812 := [rewrite]: #26811
-#26808 := (iff #26803 #26807)
-#26805 := (iff #26802 #26804)
-#26806 := [rewrite]: #26805
-#26809 := [monotonicity #26806]: #26808
-#26814 := [trans #26809 #26812]: #26813
-#26780 := [monotonicity #26814]: #26785
-#26779 := [trans #26780 #26782]: #26783
-#26733 := [quant-inst]: #26732
-#26784 := [mp #26733 #26779]: #26731
-#27048 := [unit-resolution #26784 #21815 #14275 #27047 #27022 #26997 #26990]: #27030
-#27049 := [unit-resolution #27048 #26969 #26957 #26948]: false
-#27110 := [lemma #27049]: #26690
-#26703 := (or #26701 #26711 #26692)
-#26704 := [def-axiom]: #26703
-#26724 := [unit-resolution #26704 #27110]: #26729
-#26730 := [unit-resolution #26724 #26910]: #26692
-#26789 := [trans #26730 #26788]: #11952
-#26725 := [hypothesis]: #13206
-#26791 := [unit-resolution #26725 #26789]: false
-#26792 := [lemma #26791]: #11952
-#23450 := (or #13206 #23447)
-#22451 := (forall (vars (?x778 int)) #22440)
-#22458 := (not #22451)
-#22436 := (forall (vars (?x776 int)) #22431)
-#22457 := (not #22436)
-#22459 := (or #22457 #22458)
-#22460 := (not #22459)
-#22489 := (or #22460 #22486)
-#22495 := (not #22489)
-#22496 := (or #12120 #12111 #12102 #12093 #22372 #13576 #13722 #22495)
-#22497 := (not #22496)
-#22275 := (forall (vars (?x786 int)) #22270)
-#22281 := (not #22275)
-#22282 := (or #22257 #22281)
-#22283 := (not #22282)
-#22312 := (or #22283 #22309)
-#22318 := (not #22312)
-#22319 := (or #13475 #22318)
-#22320 := (not #22319)
-#22325 := (or #13475 #22320)
-#22333 := (not #22325)
-#22334 := (or #12493 #22331 #18487 #22332 #13542 #18490 #22333)
-#22335 := (not #22334)
-#22340 := (or #18487 #18490 #22335)
-#22346 := (not #22340)
-#22383 := (or #12686 #12677 #22372 #13576 #22331 #13627 #22346)
-#22384 := (not #22383)
-#22347 := (or #12582 #12573 #18449 #18458 #12591 #12548 #12539 #13576 #22331 #22346)
-#22348 := (not #22347)
-#22353 := (or #18449 #18458 #22348)
-#22359 := (not #22353)
-#22360 := (or #18449 #18452 #22359)
-#22361 := (not #22360)
-#22366 := (or #18449 #18452 #22361)
-#22373 := (not #22366)
-#22374 := (or #22372 #13576 #13628 #22373)
-#22375 := (not #22374)
-#22389 := (or #22375 #22384)
-#22395 := (not #22389)
-#22396 := (or #18449 #18458 #22372 #13576 #22395)
-#22397 := (not #22396)
-#22402 := (or #18449 #18458 #22397)
-#22408 := (not #22402)
-#22409 := (or #18449 #18452 #22408)
-#22410 := (not #22409)
-#22415 := (or #18449 #18452 #22410)
-#22421 := (not #22415)
-#22422 := (or #22372 #13576 #13721 #22421)
-#22423 := (not #22422)
-#22502 := (or #22423 #22497)
-#22517 := (not #22502)
-#22252 := (forall (vars (?x775 int)) #22247)
-#22513 := (not #22252)
-#22518 := (or #13173 #12938 #12929 #12920 #12911 #22508 #22509 #22510 #15207 #13894 #13428 #22372 #13576 #13899 #13951 #22511 #22512 #22514 #22515 #22516 #22513 #22517)
-#22519 := (not #22518)
-#22524 := (or #13173 #13428 #22519)
-#22531 := (not #22524)
-#22241 := (forall (vars (?x773 int)) #22236)
-#22530 := (not #22241)
-#22532 := (or #22530 #22531)
-#22533 := (not #22532)
-#22538 := (or #22230 #22533)
-#22544 := (not #22538)
-#22545 := (or #13392 #22544)
-#22546 := (not #22545)
-#22551 := (or #13392 #22546)
-#22557 := (not #22551)
-#22558 := (or #13173 #13164 #13155 #13146 #18338 #18347 #22557)
-#22559 := (not #22558)
-#22564 := (or #18338 #18347 #22559)
-#22570 := (not #22564)
-#22571 := (or #18338 #18341 #22570)
-#22572 := (not #22571)
-#22577 := (or #18338 #18341 #22572)
-#22583 := (not #22577)
-#22584 := (or #13206 #22583)
-#22585 := (not #22584)
-#22590 := (or #13206 #22585)
-#23451 := (iff #22590 #23450)
-#23448 := (iff #22585 #23447)
-#23445 := (iff #22584 #23444)
-#23442 := (iff #22583 #23441)
-#23439 := (iff #22577 #23438)
-#23436 := (iff #22572 #23435)
-#23433 := (iff #22571 #23432)
-#23430 := (iff #22570 #23429)
-#23427 := (iff #22564 #23426)
-#23424 := (iff #22559 #23423)
-#23421 := (iff #22558 #23420)
-#23418 := (iff #22557 #23417)
-#23415 := (iff #22551 #23414)
-#23412 := (iff #22546 #23411)
-#23409 := (iff #22545 #23408)
-#23406 := (iff #22544 #23405)
-#23403 := (iff #22538 #23402)
-#23400 := (iff #22533 #23399)
-#23397 := (iff #22532 #23396)
-#23394 := (iff #22531 #23393)
-#23391 := (iff #22524 #23390)
-#23388 := (iff #22519 #23387)
-#23385 := (iff #22518 #23384)
-#23382 := (iff #22517 #23381)
-#23379 := (iff #22502 #23378)
-#23376 := (iff #22497 #23375)
-#23373 := (iff #22496 #23372)
-#23370 := (iff #22495 #23369)
-#23367 := (iff #22489 #23366)
-#23364 := (iff #22460 #23363)
-#23361 := (iff #22459 #23360)
-#23358 := (iff #22458 #23357)
-#23355 := (iff #22451 #23352)
-#23353 := (iff #22440 #22440)
-#23354 := [refl]: #23353
-#23356 := [quant-intro #23354]: #23355
-#23359 := [monotonicity #23356]: #23358
-#23350 := (iff #22457 #23349)
-#23347 := (iff #22436 #23344)
-#23345 := (iff #22431 #22431)
-#23346 := [refl]: #23345
-#23348 := [quant-intro #23346]: #23347
-#23351 := [monotonicity #23348]: #23350
-#23362 := [monotonicity #23351 #23359]: #23361
-#23365 := [monotonicity #23362]: #23364
-#23368 := [monotonicity #23365]: #23367
-#23371 := [monotonicity #23368]: #23370
-#23374 := [monotonicity #23371]: #23373
-#23377 := [monotonicity #23374]: #23376
-#23342 := (iff #22423 #23341)
-#23339 := (iff #22422 #23338)
-#23336 := (iff #22421 #23335)
-#23333 := (iff #22415 #23332)
-#23330 := (iff #22410 #23329)
-#23327 := (iff #22409 #23326)
-#23324 := (iff #22408 #23323)
-#23321 := (iff #22402 #23320)
-#23318 := (iff #22397 #23317)
-#23315 := (iff #22396 #23314)
-#23312 := (iff #22395 #23311)
-#23309 := (iff #22389 #23308)
-#23306 := (iff #22384 #23305)
-#23303 := (iff #22383 #23302)
-#23270 := (iff #22346 #23269)
-#23267 := (iff #22340 #23266)
-#23264 := (iff #22335 #23263)
-#23261 := (iff #22334 #23260)
-#23258 := (iff #22333 #23257)
-#23255 := (iff #22325 #23254)
-#23252 := (iff #22320 #23251)
-#23249 := (iff #22319 #23248)
-#23246 := (iff #22318 #23245)
-#23243 := (iff #22312 #23242)
-#23240 := (iff #22283 #23239)
-#23237 := (iff #22282 #23236)
-#23234 := (iff #22281 #23233)
-#23231 := (iff #22275 #23228)
-#23229 := (iff #22270 #22270)
-#23230 := [refl]: #23229
-#23232 := [quant-intro #23230]: #23231
-#23235 := [monotonicity #23232]: #23234
-#23238 := [monotonicity #23235]: #23237
-#23241 := [monotonicity #23238]: #23240
-#23244 := [monotonicity #23241]: #23243
-#23247 := [monotonicity #23244]: #23246
-#23250 := [monotonicity #23247]: #23249
-#23253 := [monotonicity #23250]: #23252
-#23256 := [monotonicity #23253]: #23255
-#23259 := [monotonicity #23256]: #23258
-#23262 := [monotonicity #23259]: #23261
-#23265 := [monotonicity #23262]: #23264
-#23268 := [monotonicity #23265]: #23267
-#23271 := [monotonicity #23268]: #23270
-#23304 := [monotonicity #23271]: #23303
-#23307 := [monotonicity #23304]: #23306
-#23300 := (iff #22375 #23299)
-#23297 := (iff #22374 #23296)
-#23294 := (iff #22373 #23293)
-#23291 := (iff #22366 #23290)
-#23288 := (iff #22361 #23287)
-#23285 := (iff #22360 #23284)
-#23282 := (iff #22359 #23281)
-#23279 := (iff #22353 #23278)
-#23276 := (iff #22348 #23275)
-#23273 := (iff #22347 #23272)
-#23274 := [monotonicity #23271]: #23273
-#23277 := [monotonicity #23274]: #23276
-#23280 := [monotonicity #23277]: #23279
-#23283 := [monotonicity #23280]: #23282
-#23286 := [monotonicity #23283]: #23285
-#23289 := [monotonicity #23286]: #23288
-#23292 := [monotonicity #23289]: #23291
-#23295 := [monotonicity #23292]: #23294
-#23298 := [monotonicity #23295]: #23297
-#23301 := [monotonicity #23298]: #23300
-#23310 := [monotonicity #23301 #23307]: #23309
-#23313 := [monotonicity #23310]: #23312
-#23316 := [monotonicity #23313]: #23315
-#23319 := [monotonicity #23316]: #23318
-#23322 := [monotonicity #23319]: #23321
-#23325 := [monotonicity #23322]: #23324
-#23328 := [monotonicity #23325]: #23327
-#23331 := [monotonicity #23328]: #23330
-#23334 := [monotonicity #23331]: #23333
-#23337 := [monotonicity #23334]: #23336
-#23340 := [monotonicity #23337]: #23339
-#23343 := [monotonicity #23340]: #23342
-#23380 := [monotonicity #23343 #23377]: #23379
-#23383 := [monotonicity #23380]: #23382
-#23226 := (iff #22513 #23225)
-#23223 := (iff #22252 #23220)
-#23221 := (iff #22247 #22247)
-#23222 := [refl]: #23221
-#23224 := [quant-intro #23222]: #23223
-#23227 := [monotonicity #23224]: #23226
-#23386 := [monotonicity #23227 #23383]: #23385
-#23389 := [monotonicity #23386]: #23388
-#23392 := [monotonicity #23389]: #23391
-#23395 := [monotonicity #23392]: #23394
-#23218 := (iff #22530 #23217)
-#23215 := (iff #22241 #23212)
-#23213 := (iff #22236 #22236)
-#23214 := [refl]: #23213
-#23216 := [quant-intro #23214]: #23215
-#23219 := [monotonicity #23216]: #23218
-#23398 := [monotonicity #23219 #23395]: #23397
-#23401 := [monotonicity #23398]: #23400
-#23404 := [monotonicity #23401]: #23403
-#23407 := [monotonicity #23404]: #23406
-#23410 := [monotonicity #23407]: #23409
-#23413 := [monotonicity #23410]: #23412
-#23416 := [monotonicity #23413]: #23415
-#23419 := [monotonicity #23416]: #23418
-#23422 := [monotonicity #23419]: #23421
-#23425 := [monotonicity #23422]: #23424
-#23428 := [monotonicity #23425]: #23427
-#23431 := [monotonicity #23428]: #23430
-#23434 := [monotonicity #23431]: #23433
-#23437 := [monotonicity #23434]: #23436
-#23440 := [monotonicity #23437]: #23439
-#23443 := [monotonicity #23440]: #23442
-#23446 := [monotonicity #23443]: #23445
-#23449 := [monotonicity #23446]: #23448
-#23452 := [monotonicity #23449]: #23451
-#18989 := (and #18632 #18633)
-#18992 := (not #18989)
-#18995 := (or #18971 #18984 #18992)
-#18998 := (not #18995)
-#15941 := (and #3103 #4070 #13751 #15127)
-#18655 := (not #15941)
-#18658 := (forall (vars (?x778 int)) #18655)
-#14370 := (and #4070 #15127)
-#14369 := (not #14370)
-#15927 := (or #13749 #13763 #14369)
-#15932 := (forall (vars (?x776 int)) #15927)
-#18662 := (and #15932 #18658)
-#19004 := (or #18662 #18998)
-#19012 := (and #12035 #12038 #12041 #12044 #13433 #13436 #13721 #19004)
-#18841 := (and #18508 #18509)
-#18844 := (not #18841)
-#18847 := (or #18823 #18836 #18844)
-#18850 := (not #18847)
-#15828 := (or #13478 #13492 #14369)
-#15833 := (forall (vars (?x786 int)) #15828)
-#18524 := (not #13521)
-#18534 := (and #18524 #15833)
-#18856 := (or #18534 #18850)
-#18861 := (and #13470 #18856)
-#18864 := (or #13475 #18861)
-#18872 := (and #3216 #13445 #13454 #13462 #13538 #15820 #18864)
-#18877 := (or #18487 #18490 #18872)
-#18917 := (and #12668 #12671 #13433 #13436 #13445 #13628 #18877)
-#18883 := (and #3201 #3202 #12369 #12378 #12384 #12389 #12392 #13436 #13445 #18877)
-#18888 := (or #18449 #18458 #18883)
-#18894 := (and #12369 #12372 #18888)
-#18899 := (or #18449 #18452 #18894)
-#18905 := (and #13433 #13436 #13627 #18899)
-#18922 := (or #18905 #18917)
-#18928 := (and #12369 #12378 #13433 #13436 #18922)
-#18933 := (or #18449 #18458 #18928)
-#18939 := (and #12369 #12372 #18933)
-#18944 := (or #18449 #18452 #18939)
-#18950 := (and #13433 #13436 #13722 #18944)
-#19017 := (or #18950 #19012)
-#15754 := (or #13914 #13928 #14369)
-#15759 := (forall (vars (?x775 int)) #15754)
-#19023 := (and #3042 #3177 #3178 #3179 #3180 #3181 #3182 #12016 #12024 #12330 #13429 #13433 #13436 #13900 #13948 #13954 #13976 #15759 #15771 #15788 #15805 #19017)
-#19028 := (or #13173 #13428 #19023)
-#15740 := (or #13395 #13407 #14369)
-#15745 := (forall (vars (?x773 int)) #15740)
-#19031 := (and #15745 #19028)
-#18744 := (and #18371 #18372)
-#18747 := (not #18744)
-#18753 := (or #18379 #18380 #18747)
-#18758 := (not #18753)
-#19034 := (or #18758 #19031)
-#19037 := (and #13389 #19034)
-#19040 := (or #13392 #19037)
-#19046 := (and #3042 #3043 #3044 #3045 #11955 #11964 #19040)
-#19051 := (or #18338 #18347 #19046)
-#19057 := (and #11955 #11958 #19051)
-#19062 := (or #18338 #18341 #19057)
-#19065 := (and #11952 #19062)
-#19068 := (or #13206 #19065)
-#22591 := (iff #19068 #22590)
-#22588 := (iff #19065 #22585)
-#22580 := (and #11952 #22577)
-#22586 := (iff #22580 #22585)
-#22587 := [rewrite]: #22586
-#22581 := (iff #19065 #22580)
-#22578 := (iff #19062 #22577)
-#22575 := (iff #19057 #22572)
-#22567 := (and #11955 #11958 #22564)
-#22573 := (iff #22567 #22572)
-#22574 := [rewrite]: #22573
-#22568 := (iff #19057 #22567)
-#22565 := (iff #19051 #22564)
-#22562 := (iff #19046 #22559)
-#22554 := (and #3042 #3043 #3044 #3045 #11955 #11964 #22551)
-#22560 := (iff #22554 #22559)
-#22561 := [rewrite]: #22560
-#22555 := (iff #19046 #22554)
-#22552 := (iff #19040 #22551)
-#22549 := (iff #19037 #22546)
-#22541 := (and #13389 #22538)
-#22547 := (iff #22541 #22546)
-#22548 := [rewrite]: #22547
-#22542 := (iff #19037 #22541)
-#22539 := (iff #19034 #22538)
-#22536 := (iff #19031 #22533)
-#22527 := (and #22241 #22524)
-#22534 := (iff #22527 #22533)
-#22535 := [rewrite]: #22534
-#22528 := (iff #19031 #22527)
-#22525 := (iff #19028 #22524)
-#22522 := (iff #19023 #22519)
-#22505 := (and #3042 #3177 #3178 #3179 #3180 #3181 #3182 #12016 #12024 #12330 #13429 #13433 #13436 #13900 #13948 #13954 #13976 #22252 #15771 #15788 #15805 #22502)
-#22520 := (iff #22505 #22519)
-#22521 := [rewrite]: #22520
-#22506 := (iff #19023 #22505)
-#22503 := (iff #19017 #22502)
-#22500 := (iff #19012 #22497)
-#22492 := (and #12035 #12038 #12041 #12044 #13433 #13436 #13721 #22489)
-#22498 := (iff #22492 #22497)
-#22499 := [rewrite]: #22498
-#22493 := (iff #19012 #22492)
-#22490 := (iff #19004 #22489)
-#22487 := (iff #18998 #22486)
-#22484 := (iff #18995 #22481)
-#22467 := (or #22465 #22466)
-#22478 := (or #18971 #18984 #22467)
-#22482 := (iff #22478 #22481)
-#22483 := [rewrite]: #22482
-#22479 := (iff #18995 #22478)
-#22476 := (iff #18992 #22467)
-#22468 := (not #22467)
-#22471 := (not #22468)
-#22474 := (iff #22471 #22467)
-#22475 := [rewrite]: #22474
-#22472 := (iff #18992 #22471)
-#22469 := (iff #18989 #22468)
-#22470 := [rewrite]: #22469
-#22473 := [monotonicity #22470]: #22472
-#22477 := [trans #22473 #22475]: #22476
-#22480 := [monotonicity #22477]: #22479
-#22485 := [trans #22480 #22483]: #22484
-#22488 := [monotonicity #22485]: #22487
-#22463 := (iff #18662 #22460)
-#22454 := (and #22436 #22451)
-#22461 := (iff #22454 #22460)
-#22462 := [rewrite]: #22461
-#22455 := (iff #18662 #22454)
-#22452 := (iff #18658 #22451)
-#22449 := (iff #18655 #22440)
-#22441 := (not #22440)
-#22444 := (not #22441)
-#22447 := (iff #22444 #22440)
-#22448 := [rewrite]: #22447
-#22445 := (iff #18655 #22444)
-#22442 := (iff #15941 #22441)
-#22443 := [rewrite]: #22442
-#22446 := [monotonicity #22443]: #22445
-#22450 := [trans #22446 #22448]: #22449
-#22453 := [quant-intro #22450]: #22452
-#22437 := (iff #15932 #22436)
-#22434 := (iff #15927 #22431)
-#20144 := (or #4992 #19506)
-#22428 := (or #13749 #13763 #20144)
-#22432 := (iff #22428 #22431)
-#22433 := [rewrite]: #22432
-#22429 := (iff #15927 #22428)
-#20153 := (iff #14369 #20144)
-#20145 := (not #20144)
-#20148 := (not #20145)
-#20151 := (iff #20148 #20144)
-#20152 := [rewrite]: #20151
-#20149 := (iff #14369 #20148)
-#20146 := (iff #14370 #20145)
-#20147 := [rewrite]: #20146
-#20150 := [monotonicity #20147]: #20149
-#20154 := [trans #20150 #20152]: #20153
-#22430 := [monotonicity #20154]: #22429
-#22435 := [trans #22430 #22433]: #22434
-#22438 := [quant-intro #22435]: #22437
-#22456 := [monotonicity #22438 #22453]: #22455
-#22464 := [trans #22456 #22462]: #22463
-#22491 := [monotonicity #22464 #22488]: #22490
-#22494 := [monotonicity #22491]: #22493
-#22501 := [trans #22494 #22499]: #22500
-#22426 := (iff #18950 #22423)
-#22418 := (and #13433 #13436 #13722 #22415)
-#22424 := (iff #22418 #22423)
-#22425 := [rewrite]: #22424
-#22419 := (iff #18950 #22418)
-#22416 := (iff #18944 #22415)
-#22413 := (iff #18939 #22410)
-#22405 := (and #12369 #12372 #22402)
-#22411 := (iff #22405 #22410)
-#22412 := [rewrite]: #22411
-#22406 := (iff #18939 #22405)
-#22403 := (iff #18933 #22402)
-#22400 := (iff #18928 #22397)
-#22392 := (and #12369 #12378 #13433 #13436 #22389)
-#22398 := (iff #22392 #22397)
-#22399 := [rewrite]: #22398
-#22393 := (iff #18928 #22392)
-#22390 := (iff #18922 #22389)
-#22387 := (iff #18917 #22384)
-#22380 := (and #12668 #12671 #13433 #13436 #13445 #13628 #22340)
-#22385 := (iff #22380 #22384)
-#22386 := [rewrite]: #22385
-#22381 := (iff #18917 #22380)
-#22341 := (iff #18877 #22340)
-#22338 := (iff #18872 #22335)
-#22328 := (and #3216 #13445 #13454 #13462 #13538 #15820 #22325)
-#22336 := (iff #22328 #22335)
-#22337 := [rewrite]: #22336
-#22329 := (iff #18872 #22328)
-#22326 := (iff #18864 #22325)
-#22323 := (iff #18861 #22320)
-#22315 := (and #13470 #22312)
-#22321 := (iff #22315 #22320)
-#22322 := [rewrite]: #22321
-#22316 := (iff #18861 #22315)
-#22313 := (iff #18856 #22312)
-#22310 := (iff #18850 #22309)
-#22307 := (iff #18847 #22304)
-#22290 := (or #22288 #22289)
-#22301 := (or #18823 #18836 #22290)
-#22305 := (iff #22301 #22304)
-#22306 := [rewrite]: #22305
-#22302 := (iff #18847 #22301)
-#22299 := (iff #18844 #22290)
-#22291 := (not #22290)
-#22294 := (not #22291)
-#22297 := (iff #22294 #22290)
-#22298 := [rewrite]: #22297
-#22295 := (iff #18844 #22294)
-#22292 := (iff #18841 #22291)
-#22293 := [rewrite]: #22292
-#22296 := [monotonicity #22293]: #22295
-#22300 := [trans #22296 #22298]: #22299
-#22303 := [monotonicity #22300]: #22302
-#22308 := [trans #22303 #22306]: #22307
-#22311 := [monotonicity #22308]: #22310
-#22286 := (iff #18534 #22283)
-#22278 := (and #22256 #22275)
-#22284 := (iff #22278 #22283)
-#22285 := [rewrite]: #22284
-#22279 := (iff #18534 #22278)
-#22276 := (iff #15833 #22275)
-#22273 := (iff #15828 #22270)
-#22267 := (or #13478 #13492 #20144)
-#22271 := (iff #22267 #22270)
-#22272 := [rewrite]: #22271
-#22268 := (iff #15828 #22267)
-#22269 := [monotonicity #20154]: #22268
-#22274 := [trans #22269 #22272]: #22273
-#22277 := [quant-intro #22274]: #22276
-#22265 := (iff #18524 #22256)
-#22260 := (not #22257)
-#22263 := (iff #22260 #22256)
-#22264 := [rewrite]: #22263
-#22261 := (iff #18524 #22260)
-#22258 := (iff #13521 #22257)
-#22259 := [rewrite]: #22258
-#22262 := [monotonicity #22259]: #22261
-#22266 := [trans #22262 #22264]: #22265
-#22280 := [monotonicity #22266 #22277]: #22279
-#22287 := [trans #22280 #22285]: #22286
-#22314 := [monotonicity #22287 #22311]: #22313
-#22317 := [monotonicity #22314]: #22316
-#22324 := [trans #22317 #22322]: #22323
-#22327 := [monotonicity #22324]: #22326
-#22330 := [monotonicity #22327]: #22329
-#22339 := [trans #22330 #22337]: #22338
-#22342 := [monotonicity #22339]: #22341
-#22382 := [monotonicity #22342]: #22381
-#22388 := [trans #22382 #22386]: #22387
-#22378 := (iff #18905 #22375)
-#22369 := (and #13433 #13436 #13627 #22366)
-#22376 := (iff #22369 #22375)
-#22377 := [rewrite]: #22376
-#22370 := (iff #18905 #22369)
-#22367 := (iff #18899 #22366)
-#22364 := (iff #18894 #22361)
-#22356 := (and #12369 #12372 #22353)
-#22362 := (iff #22356 #22361)
-#22363 := [rewrite]: #22362
-#22357 := (iff #18894 #22356)
-#22354 := (iff #18888 #22353)
-#22351 := (iff #18883 #22348)
-#22343 := (and #3201 #3202 #12369 #12378 #12384 #12389 #12392 #13436 #13445 #22340)
-#22349 := (iff #22343 #22348)
-#22350 := [rewrite]: #22349
-#22344 := (iff #18883 #22343)
-#22345 := [monotonicity #22342]: #22344
-#22352 := [trans #22345 #22350]: #22351
-#22355 := [monotonicity #22352]: #22354
-#22358 := [monotonicity #22355]: #22357
-#22365 := [trans #22358 #22363]: #22364
-#22368 := [monotonicity #22365]: #22367
-#22371 := [monotonicity #22368]: #22370
-#22379 := [trans #22371 #22377]: #22378
-#22391 := [monotonicity #22379 #22388]: #22390
-#22394 := [monotonicity #22391]: #22393
-#22401 := [trans #22394 #22399]: #22400
-#22404 := [monotonicity #22401]: #22403
-#22407 := [monotonicity #22404]: #22406
-#22414 := [trans #22407 #22412]: #22413
-#22417 := [monotonicity #22414]: #22416
-#22420 := [monotonicity #22417]: #22419
-#22427 := [trans #22420 #22425]: #22426
-#22504 := [monotonicity #22427 #22501]: #22503
-#22253 := (iff #15759 #22252)
-#22250 := (iff #15754 #22247)
-#22244 := (or #13914 #13928 #20144)
-#22248 := (iff #22244 #22247)
-#22249 := [rewrite]: #22248
-#22245 := (iff #15754 #22244)
-#22246 := [monotonicity #20154]: #22245
-#22251 := [trans #22246 #22249]: #22250
-#22254 := [quant-intro #22251]: #22253
-#22507 := [monotonicity #22254 #22504]: #22506
-#22523 := [trans #22507 #22521]: #22522
-#22526 := [monotonicity #22523]: #22525
-#22242 := (iff #15745 #22241)
-#22239 := (iff #15740 #22236)
-#22233 := (or #13395 #13407 #20144)
-#22237 := (iff #22233 #22236)
-#22238 := [rewrite]: #22237
-#22234 := (iff #15740 #22233)
-#22235 := [monotonicity #20154]: #22234
-#22240 := [trans #22235 #22238]: #22239
-#22243 := [quant-intro #22240]: #22242
-#22529 := [monotonicity #22243 #22526]: #22528
-#22537 := [trans #22529 #22535]: #22536
-#22231 := (iff #18758 #22230)
-#22228 := (iff #18753 #22225)
-#22211 := (or #22209 #22210)
-#22222 := (or #18379 #18380 #22211)
-#22226 := (iff #22222 #22225)
-#22227 := [rewrite]: #22226
-#22223 := (iff #18753 #22222)
-#22220 := (iff #18747 #22211)
-#22212 := (not #22211)
-#22215 := (not #22212)
-#22218 := (iff #22215 #22211)
-#22219 := [rewrite]: #22218
-#22216 := (iff #18747 #22215)
-#22213 := (iff #18744 #22212)
-#22214 := [rewrite]: #22213
-#22217 := [monotonicity #22214]: #22216
-#22221 := [trans #22217 #22219]: #22220
-#22224 := [monotonicity #22221]: #22223
-#22229 := [trans #22224 #22227]: #22228
-#22232 := [monotonicity #22229]: #22231
-#22540 := [monotonicity #22232 #22537]: #22539
-#22543 := [monotonicity #22540]: #22542
-#22550 := [trans #22543 #22548]: #22549
-#22553 := [monotonicity #22550]: #22552
-#22556 := [monotonicity #22553]: #22555
-#22563 := [trans #22556 #22561]: #22562
-#22566 := [monotonicity #22563]: #22565
-#22569 := [monotonicity #22566]: #22568
-#22576 := [trans #22569 #22574]: #22575
-#22579 := [monotonicity #22576]: #22578
-#22582 := [monotonicity #22579]: #22581
-#22589 := [trans #22582 #22587]: #22588
-#22592 := [monotonicity #22589]: #22591
-#18634 := (and #18633 #18632)
-#18635 := (not #18634)
-#18638 := (+ #18637 #13761)
-#18639 := (<= #18638 0::int)
-#18640 := (+ ?x776!15 #13362)
-#18641 := (>= #18640 0::int)
-#18642 := (or #18641 #18639 #18635)
-#18643 := (not #18642)
-#18666 := (or #18643 #18662)
-#18419 := (not #13441)
-#18626 := (not #12093)
-#18623 := (not #12102)
-#18620 := (not #12111)
-#18617 := (not #12120)
-#18670 := (and #18617 #18620 #18623 #18626 #18419 #13725 #18666)
-#15856 := (and #13454 #15820)
-#15859 := (not #15856)
-#18550 := (not #15859)
-#18510 := (and #18509 #18508)
-#18511 := (not #18510)
-#18514 := (+ #18513 #13490)
-#18515 := (<= #18514 0::int)
-#18516 := (+ ?x786!14 #13471)
-#18517 := (>= #18516 0::int)
-#18518 := (or #18517 #18515 #18511)
-#18519 := (not #18518)
-#18538 := (or #18519 #18534)
-#18504 := (not #13475)
-#18542 := (and #18504 #18538)
-#18546 := (or #13475 #18542)
-#18499 := (not #13542)
-#18496 := (not #13467)
-#18493 := (not #12493)
-#18553 := (and #18493 #18496 #18499 #18546 #18550)
-#18557 := (or #18487 #18490 #18553)
-#18484 := (not #13450)
-#18584 := (not #12677)
-#18581 := (not #12686)
-#18589 := (and #18581 #18584 #18419 #18484 #13633 #18557)
-#18481 := (not #13576)
-#18478 := (not #12539)
-#18475 := (not #12548)
-#18472 := (not #12591)
-#18461 := (not #12600)
-#18469 := (not #12573)
-#18466 := (not #12582)
-#18561 := (and #18466 #18469 #18461 #18472 #18475 #18478 #18481 #18484 #18557)
-#18565 := (or #18449 #18458 #18561)
-#18455 := (not #12612)
-#18569 := (and #18455 #18565)
-#18573 := (or #18449 #18452 #18569)
-#18577 := (and #18419 #13627 #18573)
-#18593 := (or #18577 #18589)
-#18597 := (and #18461 #18419 #18593)
-#18601 := (or #18449 #18458 #18597)
-#18605 := (and #18455 #18601)
-#18609 := (or #18449 #18452 #18605)
-#18613 := (and #18419 #13722 #18609)
-#18674 := (or #18613 #18670)
-#15810 := (and #13976 #15805)
-#15813 := (not #15810)
-#18444 := (not #15813)
-#15793 := (and #13433 #15788)
-#15796 := (not #15793)
-#18441 := (not #15796)
-#15776 := (and #13954 #15771)
-#15779 := (not #15776)
-#18438 := (not #15779)
-#18428 := (not #13997)
-#18425 := (not #13951)
-#18422 := (not #13911)
-#18416 := (not #13894)
-#18413 := (not #15207)
-#18410 := (not #12902)
-#18407 := (not #12911)
-#18404 := (not #12920)
-#18401 := (not #12929)
-#18398 := (not #12938)
-#18678 := (and #18398 #18401 #18404 #18407 #18410 #18413 #18416 #18419 #18422 #18425 #18428 #15759 #18438 #18441 #18444 #18674)
-#18682 := (or #13173 #14154 #18678)
-#18686 := (and #15745 #18682)
-#18373 := (and #18372 #18371)
-#18374 := (not #18373)
-#18381 := (or #18380 #18379 #18374)
-#18382 := (not #18381)
-#18690 := (or #18382 #18686)
-#18367 := (not #13392)
-#18694 := (and #18367 #18690)
-#18698 := (or #13392 #18694)
-#18362 := (not #13182)
-#18359 := (not #13146)
-#18356 := (not #13155)
-#18353 := (not #13164)
-#18350 := (not #13173)
-#18702 := (and #18350 #18353 #18356 #18359 #18362 #18698)
-#18706 := (or #18338 #18347 #18702)
-#18344 := (not #13194)
-#18710 := (and #18344 #18706)
-#18714 := (or #18338 #18341 #18710)
-#18335 := (not #13206)
-#18718 := (and #18335 #18714)
-#18722 := (or #13206 #18718)
-#19069 := (iff #18722 #19068)
-#19066 := (iff #18718 #19065)
-#19063 := (iff #18714 #19062)
-#19060 := (iff #18710 #19057)
-#19054 := (and #11961 #19051)
-#19058 := (iff #19054 #19057)
-#19059 := [rewrite]: #19058
-#19055 := (iff #18710 #19054)
-#19052 := (iff #18706 #19051)
-#19049 := (iff #18702 #19046)
-#19043 := (and #3042 #3043 #3044 #3045 #11967 #19040)
-#19047 := (iff #19043 #19046)
-#19048 := [rewrite]: #19047
-#19044 := (iff #18702 #19043)
-#19041 := (iff #18698 #19040)
-#19038 := (iff #18694 #19037)
-#19035 := (iff #18690 #19034)
-#19032 := (iff #18686 #19031)
-#19029 := (iff #18682 #19028)
-#19026 := (iff #18678 #19023)
-#19020 := (and #3177 #3178 #3179 #3180 #3183 #12024 #12330 #13438 #13906 #13948 #13992 #15759 #15776 #15793 #15810 #19017)
-#19024 := (iff #19020 #19023)
-#19025 := [rewrite]: #19024
-#19021 := (iff #18678 #19020)
-#19018 := (iff #18674 #19017)
-#19015 := (iff #18670 #19012)
-#19009 := (and #12035 #12038 #12041 #12044 #13438 #13721 #19004)
-#19013 := (iff #19009 #19012)
-#19014 := [rewrite]: #19013
-#19010 := (iff #18670 #19009)
-#19007 := (iff #18666 #19004)
-#19001 := (or #18998 #18662)
-#19005 := (iff #19001 #19004)
-#19006 := [rewrite]: #19005
-#19002 := (iff #18666 #19001)
-#18999 := (iff #18643 #18998)
-#18996 := (iff #18642 #18995)
-#18993 := (iff #18635 #18992)
-#18990 := (iff #18634 #18989)
-#18991 := [rewrite]: #18990
-#18994 := [monotonicity #18991]: #18993
-#18987 := (iff #18639 #18984)
-#18976 := (+ #13761 #18637)
-#18979 := (<= #18976 0::int)
-#18985 := (iff #18979 #18984)
-#18986 := [rewrite]: #18985
-#18980 := (iff #18639 #18979)
-#18977 := (= #18638 #18976)
-#18978 := [rewrite]: #18977
-#18981 := [monotonicity #18978]: #18980
-#18988 := [trans #18981 #18986]: #18987
-#18974 := (iff #18641 #18971)
-#18963 := (+ #13362 ?x776!15)
-#18966 := (>= #18963 0::int)
-#18972 := (iff #18966 #18971)
-#18973 := [rewrite]: #18972
-#18967 := (iff #18641 #18966)
-#18964 := (= #18640 #18963)
-#18965 := [rewrite]: #18964
-#18968 := [monotonicity #18965]: #18967
-#18975 := [trans #18968 #18973]: #18974
-#18997 := [monotonicity #18975 #18988 #18994]: #18996
-#19000 := [monotonicity #18997]: #18999
-#19003 := [monotonicity #19000]: #19002
-#19008 := [trans #19003 #19006]: #19007
-#18775 := (iff #18419 #13438)
-#18776 := [rewrite]: #18775
-#18961 := (iff #18626 #12044)
-#18962 := [rewrite]: #18961
-#18959 := (iff #18623 #12041)
-#18960 := [rewrite]: #18959
-#18957 := (iff #18620 #12038)
-#18958 := [rewrite]: #18957
-#18955 := (iff #18617 #12035)
-#18956 := [rewrite]: #18955
-#19011 := [monotonicity #18956 #18958 #18960 #18962 #18776 #13729 #19008]: #19010
-#19016 := [trans #19011 #19014]: #19015
-#18953 := (iff #18613 #18950)
-#18947 := (and #13438 #13722 #18944)
-#18951 := (iff #18947 #18950)
-#18952 := [rewrite]: #18951
-#18948 := (iff #18613 #18947)
-#18945 := (iff #18609 #18944)
-#18942 := (iff #18605 #18939)
-#18936 := (and #12375 #18933)
-#18940 := (iff #18936 #18939)
-#18941 := [rewrite]: #18940
-#18937 := (iff #18605 #18936)
-#18934 := (iff #18601 #18933)
-#18931 := (iff #18597 #18928)
-#18925 := (and #12381 #13438 #18922)
-#18929 := (iff #18925 #18928)
-#18930 := [rewrite]: #18929
-#18926 := (iff #18597 #18925)
-#18923 := (iff #18593 #18922)
-#18920 := (iff #18589 #18917)
-#18914 := (and #12668 #12671 #13438 #13447 #13628 #18877)
-#18918 := (iff #18914 #18917)
-#18919 := [rewrite]: #18918
-#18915 := (iff #18589 #18914)
-#18878 := (iff #18557 #18877)
-#18875 := (iff #18553 #18872)
-#18869 := (and #3216 #13464 #13538 #18864 #15856)
-#18873 := (iff #18869 #18872)
-#18874 := [rewrite]: #18873
-#18870 := (iff #18553 #18869)
-#18867 := (iff #18550 #15856)
-#18868 := [rewrite]: #18867
-#18865 := (iff #18546 #18864)
-#18862 := (iff #18542 #18861)
-#18859 := (iff #18538 #18856)
-#18853 := (or #18850 #18534)
-#18857 := (iff #18853 #18856)
-#18858 := [rewrite]: #18857
-#18854 := (iff #18538 #18853)
-#18851 := (iff #18519 #18850)
-#18848 := (iff #18518 #18847)
-#18845 := (iff #18511 #18844)
-#18842 := (iff #18510 #18841)
-#18843 := [rewrite]: #18842
-#18846 := [monotonicity #18843]: #18845
-#18839 := (iff #18515 #18836)
-#18828 := (+ #13490 #18513)
-#18831 := (<= #18828 0::int)
-#18837 := (iff #18831 #18836)
-#18838 := [rewrite]: #18837
-#18832 := (iff #18515 #18831)
-#18829 := (= #18514 #18828)
-#18830 := [rewrite]: #18829
-#18833 := [monotonicity #18830]: #18832
-#18840 := [trans #18833 #18838]: #18839
-#18826 := (iff #18517 #18823)
-#18815 := (+ #13471 ?x786!14)
-#18818 := (>= #18815 0::int)
-#18824 := (iff #18818 #18823)
-#18825 := [rewrite]: #18824
-#18819 := (iff #18517 #18818)
-#18816 := (= #18516 #18815)
-#18817 := [rewrite]: #18816
-#18820 := [monotonicity #18817]: #18819
-#18827 := [trans #18820 #18825]: #18826
-#18849 := [monotonicity #18827 #18840 #18846]: #18848
-#18852 := [monotonicity #18849]: #18851
-#18855 := [monotonicity #18852]: #18854
-#18860 := [trans #18855 #18858]: #18859
-#18813 := (iff #18504 #13470)
-#18814 := [rewrite]: #18813
-#18863 := [monotonicity #18814 #18860]: #18862
-#18866 := [monotonicity #18863]: #18865
-#18811 := (iff #18499 #13538)
-#18812 := [rewrite]: #18811
-#18809 := (iff #18496 #13464)
-#18810 := [rewrite]: #18809
-#18807 := (iff #18493 #3216)
-#18808 := [rewrite]: #18807
-#18871 := [monotonicity #18808 #18810 #18812 #18866 #18868]: #18870
-#18876 := [trans #18871 #18874]: #18875
-#18879 := [monotonicity #18876]: #18878
-#18805 := (iff #18484 #13447)
-#18806 := [rewrite]: #18805
-#18912 := (iff #18584 #12671)
-#18913 := [rewrite]: #18912
-#18910 := (iff #18581 #12668)
-#18911 := [rewrite]: #18910
-#18916 := [monotonicity #18911 #18913 #18776 #18806 #13637 #18879]: #18915
-#18921 := [trans #18916 #18919]: #18920
-#18908 := (iff #18577 #18905)
-#18902 := (and #13438 #13627 #18899)
-#18906 := (iff #18902 #18905)
-#18907 := [rewrite]: #18906
-#18903 := (iff #18577 #18902)
-#18900 := (iff #18573 #18899)
-#18897 := (iff #18569 #18894)
-#18891 := (and #12375 #18888)
-#18895 := (iff #18891 #18894)
-#18896 := [rewrite]: #18895
-#18892 := (iff #18569 #18891)
-#18889 := (iff #18565 #18888)
-#18886 := (iff #18561 #18883)
-#18880 := (and #3201 #3202 #12381 #12384 #12389 #12392 #13436 #13447 #18877)
-#18884 := (iff #18880 #18883)
-#18885 := [rewrite]: #18884
-#18881 := (iff #18561 #18880)
-#18803 := (iff #18481 #13436)
-#18804 := [rewrite]: #18803
-#18801 := (iff #18478 #12392)
-#18802 := [rewrite]: #18801
-#18799 := (iff #18475 #12389)
-#18800 := [rewrite]: #18799
-#18797 := (iff #18472 #12384)
-#18798 := [rewrite]: #18797
-#18791 := (iff #18461 #12381)
-#18792 := [rewrite]: #18791
-#18795 := (iff #18469 #3202)
-#18796 := [rewrite]: #18795
-#18793 := (iff #18466 #3201)
-#18794 := [rewrite]: #18793
-#18882 := [monotonicity #18794 #18796 #18792 #18798 #18800 #18802 #18804 #18806 #18879]: #18881
-#18887 := [trans #18882 #18885]: #18886
-#18890 := [monotonicity #18887]: #18889
-#18789 := (iff #18455 #12375)
-#18790 := [rewrite]: #18789
-#18893 := [monotonicity #18790 #18890]: #18892
-#18898 := [trans #18893 #18896]: #18897
-#18901 := [monotonicity #18898]: #18900
-#18904 := [monotonicity #18776 #18901]: #18903
-#18909 := [trans #18904 #18907]: #18908
-#18924 := [monotonicity #18909 #18921]: #18923
-#18927 := [monotonicity #18792 #18776 #18924]: #18926
-#18932 := [trans #18927 #18930]: #18931
-#18935 := [monotonicity #18932]: #18934
-#18938 := [monotonicity #18790 #18935]: #18937
-#18943 := [trans #18938 #18941]: #18942
-#18946 := [monotonicity #18943]: #18945
-#18949 := [monotonicity #18776 #18946]: #18948
-#18954 := [trans #18949 #18952]: #18953
-#19019 := [monotonicity #18954 #19016]: #19018
-#18787 := (iff #18444 #15810)
-#18788 := [rewrite]: #18787
-#18785 := (iff #18441 #15793)
-#18786 := [rewrite]: #18785
-#18783 := (iff #18438 #15776)
-#18784 := [rewrite]: #18783
-#18781 := (iff #18428 #13992)
-#18782 := [rewrite]: #18781
-#18779 := (iff #18425 #13948)
-#18780 := [rewrite]: #18779
-#18777 := (iff #18422 #13906)
-#18778 := [rewrite]: #18777
-#18773 := (iff #18416 #12330)
-#18774 := [rewrite]: #18773
-#18771 := (iff #18413 #12024)
-#18772 := [rewrite]: #18771
-#18769 := (iff #18410 #3183)
-#18770 := [rewrite]: #18769
-#18767 := (iff #18407 #3180)
-#18768 := [rewrite]: #18767
-#18765 := (iff #18404 #3179)
-#18766 := [rewrite]: #18765
-#18763 := (iff #18401 #3178)
-#18764 := [rewrite]: #18763
-#18761 := (iff #18398 #3177)
-#18762 := [rewrite]: #18761
-#19022 := [monotonicity #18762 #18764 #18766 #18768 #18770 #18772 #18774 #18776 #18778 #18780 #18782 #18784 #18786 #18788 #19019]: #19021
-#19027 := [trans #19022 #19025]: #19026
-#19030 := [monotonicity #14158 #19027]: #19029
-#19033 := [monotonicity #19030]: #19032
-#18759 := (iff #18382 #18758)
-#18756 := (iff #18381 #18753)
-#18750 := (or #18380 #18379 #18747)
-#18754 := (iff #18750 #18753)
-#18755 := [rewrite]: #18754
-#18751 := (iff #18381 #18750)
-#18748 := (iff #18374 #18747)
-#18745 := (iff #18373 #18744)
-#18746 := [rewrite]: #18745
-#18749 := [monotonicity #18746]: #18748
-#18752 := [monotonicity #18749]: #18751
-#18757 := [trans #18752 #18755]: #18756
-#18760 := [monotonicity #18757]: #18759
-#19036 := [monotonicity #18760 #19033]: #19035
-#18742 := (iff #18367 #13389)
-#18743 := [rewrite]: #18742
-#19039 := [monotonicity #18743 #19036]: #19038
-#19042 := [monotonicity #19039]: #19041
-#18740 := (iff #18362 #11967)
-#18741 := [rewrite]: #18740
-#18738 := (iff #18359 #3045)
-#18739 := [rewrite]: #18738
-#18736 := (iff #18356 #3044)
-#18737 := [rewrite]: #18736
-#18734 := (iff #18353 #3043)
-#18735 := [rewrite]: #18734
-#18732 := (iff #18350 #3042)
-#18733 := [rewrite]: #18732
-#19045 := [monotonicity #18733 #18735 #18737 #18739 #18741 #19042]: #19044
-#19050 := [trans #19045 #19048]: #19049
-#19053 := [monotonicity #19050]: #19052
-#18730 := (iff #18344 #11961)
-#18731 := [rewrite]: #18730
-#19056 := [monotonicity #18731 #19053]: #19055
-#19061 := [trans #19056 #19059]: #19060
-#19064 := [monotonicity #19061]: #19063
-#18728 := (iff #18335 #11952)
-#18729 := [rewrite]: #18728
-#19067 := [monotonicity #18729 #19064]: #19066
-#19070 := [monotonicity #19067]: #19069
-#15946 := (exists (vars (?x778 int)) #15941)
-#15935 := (not #15932)
-#15949 := (or #15935 #15946)
-#15952 := (and #15932 #15949)
-#15955 := (or #12120 #12111 #12102 #12093 #13441 #13722 #15952)
-#15836 := (not #15833)
-#15842 := (or #13521 #15836)
-#15847 := (and #15833 #15842)
-#15850 := (or #13475 #15847)
-#15853 := (and #13470 #15850)
-#15865 := (or #12493 #13467 #13542 #15853 #15859)
-#15870 := (and #13454 #15820 #15865)
-#15896 := (or #12686 #12677 #13441 #13450 #13627 #15870)
-#15873 := (or #12582 #12573 #12600 #12591 #12548 #12539 #13576 #13450 #15870)
-#15876 := (and #12369 #12378 #15873)
-#15879 := (or #12612 #15876)
-#15882 := (and #12369 #12372 #15879)
-#15888 := (or #13441 #13628 #15882)
-#15901 := (and #15888 #15896)
-#15904 := (or #12600 #13441 #15901)
-#15907 := (and #12369 #12378 #15904)
-#15910 := (or #12612 #15907)
-#15913 := (and #12369 #12372 #15910)
-#15919 := (or #13441 #13721 #15913)
-#15958 := (and #15919 #15955)
-#15762 := (not #15759)
-#15964 := (or #12938 #12929 #12920 #12911 #12902 #15207 #13894 #13441 #13911 #13951 #13997 #15762 #15779 #15796 #15813 #15958)
-#15969 := (and #3042 #13429 #15964)
-#15748 := (not #15745)
-#15972 := (or #15748 #15969)
-#15975 := (and #15745 #15972)
-#15978 := (or #13392 #15975)
-#15981 := (and #13389 #15978)
-#15984 := (or #13173 #13164 #13155 #13146 #13182 #15981)
-#15987 := (and #11955 #11964 #15984)
-#15990 := (or #13194 #15987)
-#15993 := (and #11955 #11958 #15990)
-#15996 := (or #13206 #15993)
-#15999 := (and #11952 #15996)
-#16002 := (not #15999)
-#18723 := (~ #16002 #18722)
-#18719 := (not #15996)
-#18720 := (~ #18719 #18718)
-#18715 := (not #15993)
-#18716 := (~ #18715 #18714)
-#18711 := (not #15990)
-#18712 := (~ #18711 #18710)
-#18707 := (not #15987)
-#18708 := (~ #18707 #18706)
-#18703 := (not #15984)
-#18704 := (~ #18703 #18702)
-#18699 := (not #15981)
-#18700 := (~ #18699 #18698)
-#18695 := (not #15978)
-#18696 := (~ #18695 #18694)
-#18691 := (not #15975)
-#18692 := (~ #18691 #18690)
-#18687 := (not #15972)
-#18688 := (~ #18687 #18686)
-#18683 := (not #15969)
-#18684 := (~ #18683 #18682)
-#18679 := (not #15964)
-#18680 := (~ #18679 #18678)
-#18675 := (not #15958)
-#18676 := (~ #18675 #18674)
-#18671 := (not #15955)
-#18672 := (~ #18671 #18670)
-#18667 := (not #15952)
-#18668 := (~ #18667 #18666)
-#18663 := (not #15949)
-#18664 := (~ #18663 #18662)
-#18659 := (not #15946)
-#18660 := (~ #18659 #18658)
-#18656 := (~ #18655 #18655)
-#18657 := [refl]: #18656
-#18661 := [nnf-neg #18657]: #18660
-#18652 := (not #15935)
-#18653 := (~ #18652 #15932)
-#18650 := (~ #15932 #15932)
-#18648 := (~ #15927 #15927)
-#18649 := [refl]: #18648
-#18651 := [nnf-pos #18649]: #18650
-#18654 := [nnf-neg #18651]: #18653
-#18665 := [nnf-neg #18654 #18661]: #18664
-#18644 := (~ #15935 #18643)
-#18645 := [sk]: #18644
-#18669 := [nnf-neg #18645 #18665]: #18668
-#18629 := (~ #13725 #13725)
-#18630 := [refl]: #18629
-#18420 := (~ #18419 #18419)
-#18421 := [refl]: #18420
-#18627 := (~ #18626 #18626)
-#18628 := [refl]: #18627
-#18624 := (~ #18623 #18623)
-#18625 := [refl]: #18624
-#18621 := (~ #18620 #18620)
-#18622 := [refl]: #18621
-#18618 := (~ #18617 #18617)
-#18619 := [refl]: #18618
-#18673 := [nnf-neg #18619 #18622 #18625 #18628 #18421 #18630 #18669]: #18672
-#18614 := (not #15919)
-#18615 := (~ #18614 #18613)
-#18610 := (not #15913)
-#18611 := (~ #18610 #18609)
-#18606 := (not #15910)
-#18607 := (~ #18606 #18605)
-#18602 := (not #15907)
-#18603 := (~ #18602 #18601)
-#18598 := (not #15904)
-#18599 := (~ #18598 #18597)
-#18594 := (not #15901)
-#18595 := (~ #18594 #18593)
-#18590 := (not #15896)
-#18591 := (~ #18590 #18589)
-#18558 := (not #15870)
-#18559 := (~ #18558 #18557)
-#18554 := (not #15865)
-#18555 := (~ #18554 #18553)
-#18551 := (~ #18550 #18550)
-#18552 := [refl]: #18551
-#18547 := (not #15853)
-#18548 := (~ #18547 #18546)
-#18543 := (not #15850)
-#18544 := (~ #18543 #18542)
-#18539 := (not #15847)
-#18540 := (~ #18539 #18538)
-#18535 := (not #15842)
-#18536 := (~ #18535 #18534)
-#18531 := (not #15836)
-#18532 := (~ #18531 #15833)
-#18529 := (~ #15833 #15833)
-#18527 := (~ #15828 #15828)
-#18528 := [refl]: #18527
-#18530 := [nnf-pos #18528]: #18529
-#18533 := [nnf-neg #18530]: #18532
-#18525 := (~ #18524 #18524)
-#18526 := [refl]: #18525
-#18537 := [nnf-neg #18526 #18533]: #18536
-#18520 := (~ #15836 #18519)
-#18521 := [sk]: #18520
-#18541 := [nnf-neg #18521 #18537]: #18540
-#18505 := (~ #18504 #18504)
-#18506 := [refl]: #18505
-#18545 := [nnf-neg #18506 #18541]: #18544
-#18502 := (~ #13475 #13475)
-#18503 := [refl]: #18502
-#18549 := [nnf-neg #18503 #18545]: #18548
-#18500 := (~ #18499 #18499)
-#18501 := [refl]: #18500
-#18497 := (~ #18496 #18496)
-#18498 := [refl]: #18497
-#18494 := (~ #18493 #18493)
-#18495 := [refl]: #18494
-#18556 := [nnf-neg #18495 #18498 #18501 #18549 #18552]: #18555
-#18491 := (~ #18490 #18490)
-#18492 := [refl]: #18491
-#18488 := (~ #18487 #18487)
-#18489 := [refl]: #18488
-#18560 := [nnf-neg #18489 #18492 #18556]: #18559
-#18587 := (~ #13633 #13633)
-#18588 := [refl]: #18587
-#18485 := (~ #18484 #18484)
-#18486 := [refl]: #18485
-#18585 := (~ #18584 #18584)
-#18586 := [refl]: #18585
-#18582 := (~ #18581 #18581)
-#18583 := [refl]: #18582
-#18592 := [nnf-neg #18583 #18586 #18421 #18486 #18588 #18560]: #18591
-#18578 := (not #15888)
-#18579 := (~ #18578 #18577)
-#18574 := (not #15882)
-#18575 := (~ #18574 #18573)
-#18570 := (not #15879)
-#18571 := (~ #18570 #18569)
-#18566 := (not #15876)
-#18567 := (~ #18566 #18565)
-#18562 := (not #15873)
-#18563 := (~ #18562 #18561)
-#18482 := (~ #18481 #18481)
-#18483 := [refl]: #18482
-#18479 := (~ #18478 #18478)
-#18480 := [refl]: #18479
-#18476 := (~ #18475 #18475)
-#18477 := [refl]: #18476
-#18473 := (~ #18472 #18472)
-#18474 := [refl]: #18473
-#18462 := (~ #18461 #18461)
-#18463 := [refl]: #18462
-#18470 := (~ #18469 #18469)
-#18471 := [refl]: #18470
-#18467 := (~ #18466 #18466)
-#18468 := [refl]: #18467
-#18564 := [nnf-neg #18468 #18471 #18463 #18474 #18477 #18480 #18483 #18486 #18560]: #18563
-#18459 := (~ #18458 #18458)
-#18460 := [refl]: #18459
-#18450 := (~ #18449 #18449)
-#18451 := [refl]: #18450
-#18568 := [nnf-neg #18451 #18460 #18564]: #18567
-#18456 := (~ #18455 #18455)
-#18457 := [refl]: #18456
-#18572 := [nnf-neg #18457 #18568]: #18571
-#18453 := (~ #18452 #18452)
-#18454 := [refl]: #18453
-#18576 := [nnf-neg #18451 #18454 #18572]: #18575
-#18464 := (~ #13627 #13627)
-#18465 := [refl]: #18464
-#18580 := [nnf-neg #18421 #18465 #18576]: #18579
-#18596 := [nnf-neg #18580 #18592]: #18595
-#18600 := [nnf-neg #18463 #18421 #18596]: #18599
-#18604 := [nnf-neg #18451 #18460 #18600]: #18603
-#18608 := [nnf-neg #18457 #18604]: #18607
-#18612 := [nnf-neg #18451 #18454 #18608]: #18611
-#18447 := (~ #13722 #13722)
-#18448 := [refl]: #18447
-#18616 := [nnf-neg #18421 #18448 #18612]: #18615
-#18677 := [nnf-neg #18616 #18673]: #18676
-#18445 := (~ #18444 #18444)
-#18446 := [refl]: #18445
-#18442 := (~ #18441 #18441)
-#18443 := [refl]: #18442
-#18439 := (~ #18438 #18438)
-#18440 := [refl]: #18439
-#18435 := (not #15762)
-#18436 := (~ #18435 #15759)
-#18433 := (~ #15759 #15759)
-#18431 := (~ #15754 #15754)
-#18432 := [refl]: #18431
-#18434 := [nnf-pos #18432]: #18433
-#18437 := [nnf-neg #18434]: #18436
-#18429 := (~ #18428 #18428)
-#18430 := [refl]: #18429
-#18426 := (~ #18425 #18425)
-#18427 := [refl]: #18426
-#18423 := (~ #18422 #18422)
-#18424 := [refl]: #18423
-#18417 := (~ #18416 #18416)
-#18418 := [refl]: #18417
-#18414 := (~ #18413 #18413)
-#18415 := [refl]: #18414
-#18411 := (~ #18410 #18410)
-#18412 := [refl]: #18411
-#18408 := (~ #18407 #18407)
-#18409 := [refl]: #18408
-#18405 := (~ #18404 #18404)
-#18406 := [refl]: #18405
-#18402 := (~ #18401 #18401)
-#18403 := [refl]: #18402
-#18399 := (~ #18398 #18398)
-#18400 := [refl]: #18399
-#18681 := [nnf-neg #18400 #18403 #18406 #18409 #18412 #18415 #18418 #18421 #18424 #18427 #18430 #18437 #18440 #18443 #18446 #18677]: #18680
-#18396 := (~ #14154 #14154)
-#18397 := [refl]: #18396
-#18394 := (~ #13173 #13173)
-#18395 := [refl]: #18394
-#18685 := [nnf-neg #18395 #18397 #18681]: #18684
-#18391 := (not #15748)
-#18392 := (~ #18391 #15745)
-#18389 := (~ #15745 #15745)
-#18387 := (~ #15740 #15740)
-#18388 := [refl]: #18387
-#18390 := [nnf-pos #18388]: #18389
-#18393 := [nnf-neg #18390]: #18392
-#18689 := [nnf-neg #18393 #18685]: #18688
-#18383 := (~ #15748 #18382)
-#18384 := [sk]: #18383
-#18693 := [nnf-neg #18384 #18689]: #18692
-#18368 := (~ #18367 #18367)
-#18369 := [refl]: #18368
-#18697 := [nnf-neg #18369 #18693]: #18696
-#18365 := (~ #13392 #13392)
-#18366 := [refl]: #18365
-#18701 := [nnf-neg #18366 #18697]: #18700
-#18363 := (~ #18362 #18362)
-#18364 := [refl]: #18363
-#18360 := (~ #18359 #18359)
-#18361 := [refl]: #18360
-#18357 := (~ #18356 #18356)
-#18358 := [refl]: #18357
-#18354 := (~ #18353 #18353)
-#18355 := [refl]: #18354
-#18351 := (~ #18350 #18350)
-#18352 := [refl]: #18351
-#18705 := [nnf-neg #18352 #18355 #18358 #18361 #18364 #18701]: #18704
-#18348 := (~ #18347 #18347)
-#18349 := [refl]: #18348
-#18339 := (~ #18338 #18338)
-#18340 := [refl]: #18339
-#18709 := [nnf-neg #18340 #18349 #18705]: #18708
-#18345 := (~ #18344 #18344)
-#18346 := [refl]: #18345
-#18713 := [nnf-neg #18346 #18709]: #18712
-#18342 := (~ #18341 #18341)
-#18343 := [refl]: #18342
-#18717 := [nnf-neg #18340 #18343 #18713]: #18716
-#18336 := (~ #18335 #18335)
-#18337 := [refl]: #18336
-#18721 := [nnf-neg #18337 #18717]: #18720
-#18333 := (~ #13206 #13206)
-#18334 := [refl]: #18333
-#18724 := [nnf-neg #18334 #18721]: #18723
-#15232 := (or #12120 #12111 #12102 #12093 #13441 #13722 #13808)
-#15237 := (and #13744 #15232)
-#15243 := (or #12938 #12929 #12920 #12911 #12902 #15207 #13894 #13441 #13911 #13945 #13951 #13963 #13973 #13986 #13997 #15237)
-#15248 := (and #3042 #13429 #15243)
-#15251 := (or #13425 #15248)
-#15254 := (and #13422 #15251)
-#15257 := (or #13392 #15254)
-#15260 := (and #13389 #15257)
-#15263 := (or #13173 #13164 #13155 #13146 #13182 #15260)
-#15266 := (and #11955 #11964 #15263)
-#15269 := (or #13194 #15266)
-#15272 := (and #11955 #11958 #15269)
-#15275 := (or #13206 #15272)
-#15278 := (and #11952 #15275)
-#15281 := (not #15278)
-#16003 := (iff #15281 #16002)
-#16000 := (iff #15278 #15999)
-#15997 := (iff #15275 #15996)
-#15994 := (iff #15272 #15993)
-#15991 := (iff #15269 #15990)
-#15988 := (iff #15266 #15987)
-#15985 := (iff #15263 #15984)
-#15982 := (iff #15260 #15981)
-#15979 := (iff #15257 #15978)
-#15976 := (iff #15254 #15975)
-#15973 := (iff #15251 #15972)
-#15970 := (iff #15248 #15969)
-#15967 := (iff #15243 #15964)
-#15961 := (or #12938 #12929 #12920 #12911 #12902 #15207 #13894 #13441 #13911 #15762 #13951 #15779 #15796 #15813 #13997 #15958)
-#15965 := (iff #15961 #15964)
-#15966 := [rewrite]: #15965
-#15962 := (iff #15243 #15961)
-#15959 := (iff #15237 #15958)
-#15956 := (iff #15232 #15955)
-#15953 := (iff #13808 #15952)
-#15950 := (iff #13803 #15949)
-#15947 := (iff #13794 #15946)
-#15944 := (iff #13789 #15941)
-#15938 := (and #3103 #4070 #15127 #13751)
-#15942 := (iff #15938 #15941)
-#15943 := [rewrite]: #15942
-#15939 := (iff #13789 #15938)
-#15122 := (iff #4384 #15127)
-#15143 := -4294967295::int
-#15135 := (+ -4294967295::int #161)
-#15128 := (<= #15135 0::int)
-#15124 := (iff #15128 #15127)
-#15125 := [rewrite]: #15124
-#15129 := (iff #4384 #15128)
-#15130 := (= #4383 #15135)
-#15136 := (+ #161 -4294967295::int)
-#15132 := (= #15136 #15135)
-#15133 := [rewrite]: #15132
-#15137 := (= #4383 #15136)
-#15138 := (= #4382 -4294967295::int)
-#15144 := (* -1::int 4294967295::int)
-#15140 := (= #15144 -4294967295::int)
-#15141 := [rewrite]: #15140
-#15145 := (= #4382 #15144)
-#7505 := (= uf_76 4294967295::int)
-#947 := 65536::int
-#1322 := (* 65536::int 65536::int)
-#1327 := (- #1322 1::int)
-#1328 := (= uf_76 #1327)
-#7506 := (iff #1328 #7505)
-#7503 := (= #1327 4294967295::int)
-#1010 := 4294967296::int
-#7496 := (- 4294967296::int 1::int)
-#7501 := (= #7496 4294967295::int)
-#7502 := [rewrite]: #7501
-#7498 := (= #1327 #7496)
-#7467 := (= #1322 4294967296::int)
-#7468 := [rewrite]: #7467
-#7499 := [monotonicity #7468]: #7498
-#7504 := [trans #7499 #7502]: #7503
-#7507 := [monotonicity #7504]: #7506
-#7495 := [asserted]: #1328
-#7510 := [mp #7495 #7507]: #7505
-#15142 := [monotonicity #7510]: #15145
-#15139 := [trans #15142 #15141]: #15138
-#15134 := [monotonicity #15139]: #15137
-#15131 := [trans #15134 #15133]: #15130
-#15126 := [monotonicity #15131]: #15129
-#15123 := [trans #15126 #15125]: #15122
-#15940 := [monotonicity #15123]: #15939
-#15945 := [trans #15940 #15943]: #15944
-#15948 := [quant-intro #15945]: #15947
-#15936 := (iff #13797 #15935)
-#15933 := (iff #13777 #15932)
-#15930 := (iff #13772 #15927)
-#15924 := (or #14369 #13749 #13763)
-#15928 := (iff #15924 #15927)
-#15929 := [rewrite]: #15928
-#15925 := (iff #13772 #15924)
-#14366 := (iff #5606 #14369)
-#14371 := (iff #4391 #14370)
-#14368 := [monotonicity #15123]: #14371
-#14367 := [monotonicity #14368]: #14366
-#15926 := [monotonicity #14367]: #15925
-#15931 := [trans #15926 #15929]: #15930
-#15934 := [quant-intro #15931]: #15933
-#15937 := [monotonicity #15934]: #15936
-#15951 := [monotonicity #15937 #15948]: #15950
-#15954 := [monotonicity #15934 #15951]: #15953
-#15957 := [monotonicity #15954]: #15956
-#15922 := (iff #13744 #15919)
-#15916 := (or #13441 #15913 #13721)
-#15920 := (iff #15916 #15919)
-#15921 := [rewrite]: #15920
-#15917 := (iff #13744 #15916)
-#15914 := (iff #13715 #15913)
-#15911 := (iff #13709 #15910)
-#15908 := (iff #13704 #15907)
-#15905 := (iff #13696 #15904)
-#15902 := (iff #13687 #15901)
-#15899 := (iff #13682 #15896)
-#15893 := (or #12686 #12677 #13441 #13450 #15870 #13627)
-#15897 := (iff #15893 #15896)
-#15898 := [rewrite]: #15897
-#15894 := (iff #13682 #15893)
-#15871 := (iff #13571 #15870)
-#15868 := (iff #13563 #15865)
-#15862 := (or #12493 #13467 #15853 #13542 #15859)
-#15866 := (iff #15862 #15865)
-#15867 := [rewrite]: #15866
-#15863 := (iff #13563 #15862)
-#15860 := (iff #13548 #15859)
-#15857 := (iff #13545 #15856)
-#15823 := (iff #13456 #15820)
-#15765 := (+ 4294967295::int #13457)
-#15816 := (>= #15765 1::int)
-#15821 := (iff #15816 #15820)
-#15822 := [rewrite]: #15821
-#15817 := (iff #13456 #15816)
-#15766 := (= #13458 #15765)
-#15767 := [monotonicity #7510]: #15766
-#15818 := [monotonicity #15767]: #15817
-#15824 := [trans #15818 #15822]: #15823
-#15858 := [monotonicity #15824]: #15857
-#15861 := [monotonicity #15858]: #15860
-#15854 := (iff #13535 #15853)
-#15851 := (iff #13532 #15850)
-#15848 := (iff #13529 #15847)
-#15845 := (iff #13526 #15842)
-#15839 := (or #15836 #13521)
-#15843 := (iff #15839 #15842)
-#15844 := [rewrite]: #15843
-#15840 := (iff #13526 #15839)
-#15837 := (iff #13509 #15836)
-#15834 := (iff #13506 #15833)
-#15831 := (iff #13501 #15828)
-#15825 := (or #14369 #13478 #13492)
-#15829 := (iff #15825 #15828)
-#15830 := [rewrite]: #15829
-#15826 := (iff #13501 #15825)
-#15827 := [monotonicity #14367]: #15826
-#15832 := [trans #15827 #15830]: #15831
-#15835 := [quant-intro #15832]: #15834
-#15838 := [monotonicity #15835]: #15837
-#15841 := [monotonicity #15838]: #15840
-#15846 := [trans #15841 #15844]: #15845
-#15849 := [monotonicity #15835 #15846]: #15848
-#15852 := [monotonicity #15849]: #15851
-#15855 := [monotonicity #15852]: #15854
-#15864 := [monotonicity #15855 #15861]: #15863
-#15869 := [trans #15864 #15867]: #15868
-#15872 := [monotonicity #15824 #15869]: #15871
-#15895 := [monotonicity #15872]: #15894
-#15900 := [trans #15895 #15898]: #15899
-#15891 := (iff #13652 #15888)
-#15885 := (or #13441 #15882 #13628)
-#15889 := (iff #15885 #15888)
-#15890 := [rewrite]: #15889
-#15886 := (iff #13652 #15885)
-#15883 := (iff #13622 #15882)
-#15880 := (iff #13616 #15879)
-#15877 := (iff #13611 #15876)
-#15874 := (iff #13603 #15873)
-#15875 := [monotonicity #15872]: #15874
-#15878 := [monotonicity #15875]: #15877
-#15881 := [monotonicity #15878]: #15880
-#15884 := [monotonicity #15881]: #15883
-#15887 := [monotonicity #15884]: #15886
-#15892 := [trans #15887 #15890]: #15891
-#15903 := [monotonicity #15892 #15900]: #15902
-#15906 := [monotonicity #15903]: #15905
-#15909 := [monotonicity #15906]: #15908
-#15912 := [monotonicity #15909]: #15911
-#15915 := [monotonicity #15912]: #15914
-#15918 := [monotonicity #15915]: #15917
-#15923 := [trans #15918 #15921]: #15922
-#15960 := [monotonicity #15923 #15957]: #15959
-#15814 := (iff #13986 #15813)
-#15811 := (iff #13983 #15810)
-#15808 := (iff #13979 #15805)
-#15799 := (+ 255::int #13926)
-#15802 := (>= #15799 0::int)
-#15806 := (iff #15802 #15805)
-#15807 := [rewrite]: #15806
-#15803 := (iff #13979 #15802)
-#15800 := (= #13980 #15799)
-#1332 := (= uf_78 255::int)
-#7509 := [asserted]: #1332
-#15801 := [monotonicity #7509]: #15800
-#15804 := [monotonicity #15801]: #15803
-#15809 := [trans #15804 #15807]: #15808
-#15812 := [monotonicity #15809]: #15811
-#15815 := [monotonicity #15812]: #15814
-#15797 := (iff #13973 #15796)
-#15794 := (iff #13970 #15793)
-#15791 := (iff #13966 #15788)
-#15782 := (+ 4294967295::int #13897)
-#15785 := (>= #15782 0::int)
-#15789 := (iff #15785 #15788)
-#15790 := [rewrite]: #15789
-#15786 := (iff #13966 #15785)
-#15783 := (= #13967 #15782)
-#15784 := [monotonicity #7510]: #15783
-#15787 := [monotonicity #15784]: #15786
-#15792 := [trans #15787 #15790]: #15791
-#15795 := [monotonicity #15792]: #15794
-#15798 := [monotonicity #15795]: #15797
-#15780 := (iff #13963 #15779)
-#15777 := (iff #13960 #15776)
-#15774 := (iff #13957 #15771)
-#15768 := (>= #15765 0::int)
-#15772 := (iff #15768 #15771)
-#15773 := [rewrite]: #15772
-#15769 := (iff #13957 #15768)
-#15770 := [monotonicity #15767]: #15769
-#15775 := [trans #15770 #15773]: #15774
-#15778 := [monotonicity #15775]: #15777
-#15781 := [monotonicity #15778]: #15780
-#15763 := (iff #13945 #15762)
-#15760 := (iff #13942 #15759)
-#15757 := (iff #13937 #15754)
-#15751 := (or #14369 #13914 #13928)
-#15755 := (iff #15751 #15754)
-#15756 := [rewrite]: #15755
-#15752 := (iff #13937 #15751)
-#15753 := [monotonicity #14367]: #15752
-#15758 := [trans #15753 #15756]: #15757
-#15761 := [quant-intro #15758]: #15760
-#15764 := [monotonicity #15761]: #15763
-#15963 := [monotonicity #15764 #15781 #15798 #15815 #15960]: #15962
-#15968 := [trans #15963 #15966]: #15967
-#15971 := [monotonicity #15968]: #15970
-#15749 := (iff #13425 #15748)
-#15746 := (iff #13422 #15745)
-#15743 := (iff #13417 #15740)
-#15737 := (or #14369 #13395 #13407)
-#15741 := (iff #15737 #15740)
-#15742 := [rewrite]: #15741
-#15738 := (iff #13417 #15737)
-#15739 := [monotonicity #14367]: #15738
-#15744 := [trans #15739 #15742]: #15743
-#15747 := [quant-intro #15744]: #15746
-#15750 := [monotonicity #15747]: #15749
-#15974 := [monotonicity #15750 #15971]: #15973
-#15977 := [monotonicity #15747 #15974]: #15976
-#15980 := [monotonicity #15977]: #15979
-#15983 := [monotonicity #15980]: #15982
-#15986 := [monotonicity #15983]: #15985
-#15989 := [monotonicity #15986]: #15988
-#15992 := [monotonicity #15989]: #15991
-#15995 := [monotonicity #15992]: #15994
-#15998 := [monotonicity #15995]: #15997
-#16001 := [monotonicity #15998]: #16000
-#16004 := [monotonicity #16001]: #16003
-#14281 := (not #14133)
-#15282 := (iff #14281 #15281)
-#15279 := (iff #14133 #15278)
-#15276 := (iff #14130 #15275)
-#15273 := (iff #14125 #15272)
-#15270 := (iff #14119 #15269)
-#15267 := (iff #14114 #15266)
-#15264 := (iff #14106 #15263)
-#15261 := (iff #14085 #15260)
-#15258 := (iff #14082 #15257)
-#15255 := (iff #14079 #15254)
-#15252 := (iff #14076 #15251)
-#15249 := (iff #14071 #15248)
-#15246 := (iff #14063 #15243)
-#15240 := (or #12938 #12929 #12920 #12911 #12902 #15207 #13894 #13441 #15237 #13911 #13945 #13951 #13963 #13973 #13986 #13997)
-#15244 := (iff #15240 #15243)
-#15245 := [rewrite]: #15244
-#15241 := (iff #14063 #15240)
-#15238 := (iff #13870 #15237)
-#15235 := (iff #13865 #15232)
-#15217 := (or #12120 #12111 #12102 #12093 #13441 #13808)
-#15229 := (or #13441 #13722 #15217)
-#15233 := (iff #15229 #15232)
-#15234 := [rewrite]: #15233
-#15230 := (iff #13865 #15229)
-#15227 := (iff #13840 #15217)
-#15222 := (and true #15217)
-#15225 := (iff #15222 #15217)
-#15226 := [rewrite]: #15225
-#15223 := (iff #13840 #15222)
-#15220 := (iff #13835 #15217)
-#15214 := (or false #12120 #12111 #12102 #12093 #13441 #13808)
-#15218 := (iff #15214 #15217)
-#15219 := [rewrite]: #15218
-#15215 := (iff #13835 #15214)
-#15212 := (iff #12168 false)
-#15210 := (iff #12168 #3086)
-#14948 := (iff up_216 true)
-#10769 := [asserted]: up_216
-#14949 := [iff-true #10769]: #14948
-#15211 := [monotonicity #14949]: #15210
-#15213 := [trans #15211 #12023]: #15212
-#15216 := [monotonicity #15213]: #15215
-#15221 := [trans #15216 #15219]: #15220
-#15224 := [monotonicity #14949 #15221]: #15223
-#15228 := [trans #15224 #15226]: #15227
-#15231 := [monotonicity #15228]: #15230
-#15236 := [trans #15231 #15234]: #15235
-#15239 := [monotonicity #15236]: #15238
-#15208 := (iff #12203 #15207)
-#15205 := (iff #12030 #12024)
-#15200 := (and true #12024)
-#15203 := (iff #15200 #12024)
-#15204 := [rewrite]: #15203
-#15201 := (iff #12030 #15200)
-#15190 := (iff #11932 true)
-#15191 := [iff-true #14275]: #15190
-#15202 := [monotonicity #15191]: #15201
-#15206 := [trans #15202 #15204]: #15205
-#15209 := [monotonicity #15206]: #15208
-#15242 := [monotonicity #15209 #15239]: #15241
-#15247 := [trans #15242 #15245]: #15246
-#15250 := [monotonicity #15247]: #15249
-#15253 := [monotonicity #15250]: #15252
-#15256 := [monotonicity #15253]: #15255
-#15259 := [monotonicity #15256]: #15258
-#15262 := [monotonicity #15259]: #15261
-#15265 := [monotonicity #15262]: #15264
-#15268 := [monotonicity #15265]: #15267
-#15271 := [monotonicity #15268]: #15270
-#15274 := [monotonicity #15271]: #15273
-#15277 := [monotonicity #15274]: #15276
-#15280 := [monotonicity #15277]: #15279
-#15283 := [monotonicity #15280]: #15282
-#14282 := [not-or-elim #14266]: #14281
-#15284 := [mp #14282 #15283]: #15281
-#16005 := [mp #15284 #16004]: #16002
-#18725 := [mp~ #16005 #18724]: #18722
-#18726 := [mp #18725 #19070]: #19068
-#22593 := [mp #18726 #22592]: #22590
-#23453 := [mp #22593 #23452]: #23450
-#28844 := [unit-resolution #23453 #26792]: #23447
-#24677 := (or #23444 #23438)
-#24678 := [def-axiom]: #24677
-#28845 := [unit-resolution #24678 #28844]: #23438
-decl uf_15 :: (-> T5 T6 T2)
-decl uf_16 :: (-> T4 T5 T6)
-#26748 := (uf_16 uf_287 #26144)
-#27224 := (uf_15 #27137 #26748)
-#27225 := (= uf_9 #27224)
-#26749 := (uf_15 #26144 #26748)
-#26750 := (= uf_9 #26749)
-#26946 := (or #13206 #26750)
-#26937 := [monotonicity #28517 #28517]: #26936
-#26954 := [symm #26937]: #26939
-#26911 := [hypothesis]: #11952
-#26941 := [trans #26911 #26954]: #26692
-decl uf_53 :: (-> T4 T5 T6)
-#26739 := (uf_53 uf_287 #26144)
-#26740 := (uf_15 #23 #26739)
-#26745 := (pattern #26740)
-decl up_197 :: (-> T3 bool)
-#26743 := (up_197 #26555)
-#26741 := (= uf_9 #26740)
-#26742 := (not #26741)
-decl uf_147 :: (-> T5 T6 T2)
-decl uf_192 :: (-> T7 T6)
-decl uf_11 :: (-> T4 T5 T7)
-#26735 := (uf_11 uf_287 #26144)
-#26736 := (uf_192 #26735)
-#26737 := (uf_147 #23 #26736)
-#26738 := (= uf_9 #26737)
-#26755 := (or #26738 #26742 #26743)
-#26758 := (forall (vars (?x577 T5)) (:pat #26745) #26755)
-#26761 := (not #26758)
-#26751 := (not #26750)
-#26764 := (or #26620 #26751 #26761)
-#26943 := [hypothesis]: #26751
-#26837 := (or #26764 #26750)
-#26841 := [def-axiom]: #26837
-#26944 := [unit-resolution #26841 #26943]: #26764
-#14 := (:var 2 T4)
-#2166 := (uf_196 #14 #15 #23)
-#2228 := (pattern #2166)
-#2229 := (uf_53 #13 #21)
-#2230 := (uf_15 #23 #2229)
-#2231 := (pattern #2230)
-#2158 := (uf_11 #13 #15)
-#2236 := (uf_192 #2158)
-#2237 := (uf_147 #23 #2236)
-#10053 := (= uf_9 #2237)
-#10047 := (= uf_9 #2230)
-#21816 := (not #10047)
-#1382 := (uf_13 #21)
-#2232 := (up_197 #1382)
-#21831 := (or #2232 #21816 #10053)
-#21836 := (forall (vars (?x577 T5)) (:pat #2231) #21831)
-#21842 := (not #21836)
-#2145 := (uf_16 #14 #23)
-#2146 := (uf_15 #15 #2145)
-#9753 := (= uf_9 #2146)
-#21651 := (not #9753)
-#180 := (uf_27 #14 #15)
-#3747 := (= uf_9 #180)
-#10390 := (not #3747)
-#21843 := (or #10390 #21651 #21842)
-#21844 := (not #21843)
-#9801 := (= uf_9 #2166)
-#10077 := (not #9801)
-#21849 := (or #10077 #21844)
-#21852 := (forall (vars (?x574 T4) (?x575 T5) (?x576 T5)) (:pat #2228) #21849)
-#2233 := (not #2232)
-#10050 := (and #2233 #10047)
-#10059 := (not #10050)
-#10060 := (or #10059 #10053)
-#10065 := (forall (vars (?x577 T5)) (:pat #2231) #10060)
-#10086 := (and #3747 #9753 #10065)
-#10089 := (or #10077 #10086)
-#10092 := (forall (vars (?x574 T4) (?x575 T5) (?x576 T5)) (:pat #2228) #10089)
-#21853 := (iff #10092 #21852)
-#21850 := (iff #10089 #21849)
-#21847 := (iff #10086 #21844)
-#21839 := (and #3747 #9753 #21836)
-#21845 := (iff #21839 #21844)
-#21846 := [rewrite]: #21845
-#21840 := (iff #10086 #21839)
-#21837 := (iff #10065 #21836)
-#21834 := (iff #10060 #21831)
-#21817 := (or #2232 #21816)
-#21828 := (or #21817 #10053)
-#21832 := (iff #21828 #21831)
-#21833 := [rewrite]: #21832
-#21829 := (iff #10060 #21828)
-#21826 := (iff #10059 #21817)
-#21818 := (not #21817)
-#21821 := (not #21818)
-#21824 := (iff #21821 #21817)
-#21825 := [rewrite]: #21824
-#21822 := (iff #10059 #21821)
-#21819 := (iff #10050 #21818)
-#21820 := [rewrite]: #21819
-#21823 := [monotonicity #21820]: #21822
-#21827 := [trans #21823 #21825]: #21826
-#21830 := [monotonicity #21827]: #21829
-#21835 := [trans #21830 #21833]: #21834
-#21838 := [quant-intro #21835]: #21837
-#21841 := [monotonicity #21838]: #21840
-#21848 := [trans #21841 #21846]: #21847
-#21851 := [monotonicity #21848]: #21850
-#21854 := [quant-intro #21851]: #21853
-#17802 := (~ #10092 #10092)
-#17800 := (~ #10089 #10089)
-#17798 := (~ #10086 #10086)
-#17796 := (~ #10065 #10065)
-#17794 := (~ #10060 #10060)
-#17795 := [refl]: #17794
-#17797 := [nnf-pos #17795]: #17796
-#17792 := (~ #9753 #9753)
-#17793 := [refl]: #17792
-#17790 := (~ #3747 #3747)
-#17791 := [refl]: #17790
-#17799 := [monotonicity #17791 #17793 #17797]: #17798
-#17788 := (~ #10077 #10077)
-#17789 := [refl]: #17788
-#17801 := [monotonicity #17789 #17799]: #17800
-#17803 := [nnf-pos #17801]: #17802
-#2238 := (= #2237 uf_9)
-#2234 := (= #2230 uf_9)
-#2235 := (and #2233 #2234)
-#2239 := (implies #2235 #2238)
-#2240 := (forall (vars (?x577 T5)) (:pat #2231) #2239)
-#184 := (= #180 uf_9)
-#2241 := (and #184 #2240)
-#2151 := (= #2146 uf_9)
-#2242 := (and #2151 #2241)
-#2167 := (= #2166 uf_9)
-#2243 := (implies #2167 #2242)
-#2244 := (forall (vars (?x574 T4) (?x575 T5) (?x576 T5)) (:pat #2228) #2243)
-#10095 := (iff #2244 #10092)
-#10068 := (and #3747 #10065)
-#10071 := (and #9753 #10068)
-#10078 := (or #10077 #10071)
-#10083 := (forall (vars (?x574 T4) (?x575 T5) (?x576 T5)) (:pat #2228) #10078)
-#10093 := (iff #10083 #10092)
-#10090 := (iff #10078 #10089)
-#10087 := (iff #10071 #10086)
-#10088 := [rewrite]: #10087
-#10091 := [monotonicity #10088]: #10090
-#10094 := [quant-intro #10091]: #10093
-#10084 := (iff #2244 #10083)
-#10081 := (iff #2243 #10078)
-#10074 := (implies #9801 #10071)
-#10079 := (iff #10074 #10078)
-#10080 := [rewrite]: #10079
-#10075 := (iff #2243 #10074)
-#10072 := (iff #2242 #10071)
-#10069 := (iff #2241 #10068)
-#10066 := (iff #2240 #10065)
-#10063 := (iff #2239 #10060)
-#10056 := (implies #10050 #10053)
-#10061 := (iff #10056 #10060)
-#10062 := [rewrite]: #10061
-#10057 := (iff #2239 #10056)
-#10054 := (iff #2238 #10053)
-#10055 := [rewrite]: #10054
-#10051 := (iff #2235 #10050)
-#10048 := (iff #2234 #10047)
-#10049 := [rewrite]: #10048
-#10052 := [monotonicity #10049]: #10051
-#10058 := [monotonicity #10052 #10055]: #10057
-#10064 := [trans #10058 #10062]: #10063
-#10067 := [quant-intro #10064]: #10066
-#3748 := (iff #184 #3747)
-#3749 := [rewrite]: #3748
-#10070 := [monotonicity #3749 #10067]: #10069
-#9754 := (iff #2151 #9753)
-#9755 := [rewrite]: #9754
-#10073 := [monotonicity #9755 #10070]: #10072
-#9802 := (iff #2167 #9801)
-#9803 := [rewrite]: #9802
-#10076 := [monotonicity #9803 #10073]: #10075
-#10082 := [trans #10076 #10080]: #10081
-#10085 := [quant-intro #10082]: #10084
-#10096 := [trans #10085 #10094]: #10095
-#10046 := [asserted]: #2244
-#10097 := [mp #10046 #10096]: #10092
-#17804 := [mp~ #10097 #17803]: #10092
-#21855 := [mp #17804 #21854]: #21852
-#26767 := (not #26764)
-#26859 := (not #21852)
-#26860 := (or #26859 #26711 #26767)
-#26744 := (or #26743 #26742 #26738)
-#26746 := (forall (vars (?x577 T5)) (:pat #26745) #26744)
-#26747 := (not #26746)
-#26752 := (or #26620 #26751 #26747)
-#26753 := (not #26752)
-#26754 := (or #26711 #26753)
-#26848 := (or #26859 #26754)
-#26832 := (iff #26848 #26860)
-#26770 := (or #26711 #26767)
-#26850 := (or #26859 #26770)
-#26888 := (iff #26850 #26860)
-#26836 := [rewrite]: #26888
-#26886 := (iff #26848 #26850)
-#26771 := (iff #26754 #26770)
-#26768 := (iff #26753 #26767)
-#26765 := (iff #26752 #26764)
-#26762 := (iff #26747 #26761)
-#26759 := (iff #26746 #26758)
-#26756 := (iff #26744 #26755)
-#26757 := [rewrite]: #26756
-#26760 := [quant-intro #26757]: #26759
-#26763 := [monotonicity #26760]: #26762
-#26766 := [monotonicity #26763]: #26765
-#26769 := [monotonicity #26766]: #26768
-#26772 := [monotonicity #26769]: #26771
-#26887 := [monotonicity #26772]: #26886
-#26838 := [trans #26887 #26836]: #26832
-#26849 := [quant-inst]: #26848
-#26834 := [mp #26849 #26838]: #26860
-#26942 := [unit-resolution #26834 #21855 #26944]: #26711
-#26708 := (not #26692)
-#26702 := (or #26701 #26690 #26708)
-#26698 := [def-axiom]: #26702
-#26945 := [unit-resolution #26698 #26942 #26941 #26910]: false
-#26947 := [lemma #26945]: #26946
-#28846 := [unit-resolution #26947 #26792]: #26750
-#26934 := (or #26751 #27225)
-#26912 := (= #26749 #27224)
-#26908 := (= #27224 #26749)
-#28528 := (= #27137 #26144)
-#28529 := [trans #28503 #28515]: #28528
-#26909 := [monotonicity #28529]: #26908
-#26932 := [symm #26909]: #26912
-#26920 := [hypothesis]: #26750
-#26933 := [trans #26920 #26932]: #27225
-#27226 := (not #27225)
-#26906 := [hypothesis]: #27226
-#26931 := [unit-resolution #26906 #26933]: false
-#26935 := [lemma #26931]: #26934
-#28847 := [unit-resolution #26935 #28846]: #27225
-#27261 := (or #11958 #27226)
-#27214 := (uf_43 #24854 #2980)
-#27215 := (uf_66 #27214 0::int #24854)
-#27219 := (uf_24 uf_287 #27215)
-#27220 := (= uf_9 #27219)
-#27221 := (not #27220)
-#27196 := (iff #18341 #27221)
-#27194 := (iff #11958 #27220)
-#27191 := (iff #27220 #11958)
-#27212 := (= #27219 #3034)
-#27210 := (= #27215 #3031)
-#28407 := (= #27214 #2979)
-#28405 := (= #2980 uf_288)
-#24973 := (= uf_288 #2980)
-#2698 := (uf_116 #2697)
-#11237 := (= #161 #2698)
-#23149 := (forall (vars (?x718 T3) (?x719 int)) (:pat #23148) #11237)
-#11241 := (forall (vars (?x718 T3) (?x719 int)) #11237)
-#23152 := (iff #11241 #23149)
-#23150 := (iff #11237 #11237)
-#23151 := [refl]: #23150
-#23153 := [quant-intro #23151]: #23152
-#18195 := (~ #11241 #11241)
-#18193 := (~ #11237 #11237)
-#18194 := [refl]: #18193
-#18196 := [nnf-pos #18194]: #18195
-#2699 := (= #2698 #161)
-#2700 := (forall (vars (?x718 T3) (?x719 int)) #2699)
-#11242 := (iff #2700 #11241)
-#11239 := (iff #2699 #11237)
-#11240 := [rewrite]: #11239
-#11243 := [quant-intro #11240]: #11242
-#11236 := [asserted]: #2700
-#11246 := [mp #11236 #11243]: #11241
-#18197 := [mp~ #11246 #18196]: #11241
-#23154 := [mp #18197 #23153]: #23149
-#24927 := (not #23149)
-#24978 := (or #24927 #24973)
-#24979 := [quant-inst]: #24978
-#28404 := [unit-resolution #24979 #23154]: #24973
-#28406 := [symm #28404]: #28405
-#28408 := [monotonicity #28401 #28406]: #28407
-#27211 := [monotonicity #28408 #28401]: #27210
-#27213 := [monotonicity #27211]: #27212
-#27193 := [monotonicity #27213]: #27191
-#27195 := [symm #27193]: #27194
-#27197 := [monotonicity #27195]: #27196
-#27209 := [hypothesis]: #18341
-#27192 := [mp #27209 #27197]: #27221
-#27216 := (uf_58 #3157 #27215)
-#27217 := (uf_136 #27216)
-#27218 := (= uf_9 #27217)
-#27231 := (or #27218 #27221)
-#27234 := (not #27231)
-decl uf_22 :: (-> T3 T2)
-#27227 := (uf_22 #24854)
-#27228 := (= uf_9 #27227)
-#2783 := (uf_22 uf_7)
-#28439 := (= #2783 #27227)
-#28436 := (= #27227 #2783)
-#28437 := [monotonicity #28401]: #28436
-#28440 := [symm #28437]: #28439
-#11413 := (= uf_9 #2783)
-#2784 := (= #2783 uf_9)
-#11415 := (iff #2784 #11413)
-#11416 := [rewrite]: #11415
-#11412 := [asserted]: #2784
-#11419 := [mp #11412 #11416]: #11413
-#28441 := [trans #11419 #28440]: #27228
-#27229 := (not #27228)
-#27257 := (or #27229 #27234)
-#28442 := [hypothesis]: #27225
-#27172 := (<= #24856 0::int)
-#27199 := (not #27172)
-#14280 := [not-or-elim #14266]: #13429
-#27155 := (* -1::int #24856)
-#27264 := (+ uf_286 #27155)
-#27265 := (<= #27264 0::int)
-#28422 := (not #24857)
-#28423 := (or #28422 #27265)
-#28424 := [th-lemma]: #28423
-#28425 := [unit-resolution #28424 #28421]: #27265
-#28431 := (not #27265)
-#27200 := (or #27199 #13428 #28431)
-#27201 := [th-lemma]: #27200
-#27202 := [unit-resolution #27201 #28425 #14280]: #27199
-#237 := (uf_22 #233)
-#247 := (:var 1 int)
-#762 := (:var 4 int)
-#2069 := (uf_43 #233 #762)
-#2070 := (uf_66 #2069 #247 #233)
-#1373 := (:var 5 T4)
-#2086 := (uf_25 #1373 #2070)
-#1365 := (:var 3 T5)
-#2067 := (uf_16 #1373 #1365)
-#268 := (:var 2 int)
-#2065 := (uf_124 #233 #268)
-#2066 := (uf_43 #2065 #762)
-#2068 := (uf_15 #2066 #2067)
-#2087 := (pattern #2068 #2086 #237)
-#1545 := (uf_59 #1373)
-#2084 := (uf_58 #1545 #2070)
-#2085 := (pattern #2068 #2084 #237)
-#2090 := (uf_136 #2084)
-#9561 := (= uf_9 #2090)
-#2088 := (uf_24 #1373 #2070)
-#9558 := (= uf_9 #2088)
-#21561 := (not #9558)
-#21562 := (or #21561 #9561)
-#21563 := (not #21562)
-#9502 := (= uf_9 #2068)
-#21537 := (not #9502)
-#2073 := (uf_55 #1373)
-#9499 := (= uf_9 #2073)
-#21536 := (not #9499)
-#4074 := (* -1::int #268)
-#6138 := (+ #247 #4074)
-#6735 := (>= #6138 0::int)
-#4336 := (>= #247 0::int)
-#19474 := (not #4336)
-#3955 := (= uf_9 #237)
-#10273 := (not #3955)
-#21569 := (or #10273 #19474 #6735 #21536 #21537 #21563)
-#21574 := (forall (vars (?x526 T4) (?x527 int) (?x528 T5) (?x529 int) (?x530 int) (?x531 T3)) (:pat #2085 #2087) #21569)
-#9564 := (not #9561)
-#9567 := (and #9558 #9564)
-#7802 := (not #6735)
-#9540 := (and #3955 #4336 #7802 #9499 #9502)
-#9545 := (not #9540)
-#9581 := (or #9545 #9567)
-#9584 := (forall (vars (?x526 T4) (?x527 int) (?x528 T5) (?x529 int) (?x530 int) (?x531 T3)) (:pat #2085 #2087) #9581)
-#21575 := (iff #9584 #21574)
-#21572 := (iff #9581 #21569)
-#21538 := (or #10273 #19474 #6735 #21536 #21537)
-#21566 := (or #21538 #21563)
-#21570 := (iff #21566 #21569)
-#21571 := [rewrite]: #21570
-#21567 := (iff #9581 #21566)
-#21564 := (iff #9567 #21563)
-#21565 := [rewrite]: #21564
-#21547 := (iff #9545 #21538)
-#21539 := (not #21538)
-#21542 := (not #21539)
-#21545 := (iff #21542 #21538)
-#21546 := [rewrite]: #21545
-#21543 := (iff #9545 #21542)
-#21540 := (iff #9540 #21539)
-#21541 := [rewrite]: #21540
-#21544 := [monotonicity #21541]: #21543
-#21548 := [trans #21544 #21546]: #21547
-#21568 := [monotonicity #21548 #21565]: #21567
-#21573 := [trans #21568 #21571]: #21572
-#21576 := [quant-intro #21573]: #21575
-#17668 := (~ #9584 #9584)
-#17666 := (~ #9581 #9581)
-#17667 := [refl]: #17666
-#17669 := [nnf-pos #17667]: #17668
-#2091 := (= #2090 uf_9)
-#2092 := (not #2091)
-#2089 := (= #2088 uf_9)
-#2093 := (and #2089 #2092)
-#1434 := (< #247 #268)
-#397 := (<= 0::int #247)
-#1435 := (and #397 #1434)
-#2075 := (= #2068 uf_9)
-#2076 := (and #2075 #1435)
-#238 := (= #237 uf_9)
-#2077 := (and #238 #2076)
-#2074 := (= #2073 uf_9)
-#2078 := (and #2074 #2077)
-#2094 := (implies #2078 #2093)
-#2095 := (forall (vars (?x526 T4) (?x527 int) (?x528 T5) (?x529 int) (?x530 int) (?x531 T3)) (:pat #2085 #2087) #2094)
-#9587 := (iff #2095 #9584)
-#9508 := (and #1435 #9502)
-#9513 := (and #3955 #9508)
-#9516 := (and #9499 #9513)
-#9522 := (not #9516)
-#9573 := (or #9522 #9567)
-#9578 := (forall (vars (?x526 T4) (?x527 int) (?x528 T5) (?x529 int) (?x530 int) (?x531 T3)) (:pat #2085 #2087) #9573)
-#9585 := (iff #9578 #9584)
-#9582 := (iff #9573 #9581)
-#9546 := (iff #9522 #9545)
-#9543 := (iff #9516 #9540)
-#7805 := (and #4336 #7802)
-#9531 := (and #7805 #9502)
-#9534 := (and #3955 #9531)
-#9537 := (and #9499 #9534)
-#9541 := (iff #9537 #9540)
-#9542 := [rewrite]: #9541
-#9538 := (iff #9516 #9537)
-#9535 := (iff #9513 #9534)
-#9532 := (iff #9508 #9531)
-#7806 := (iff #1435 #7805)
-#7803 := (iff #1434 #7802)
-#7804 := [rewrite]: #7803
-#4335 := (iff #397 #4336)
-#4337 := [rewrite]: #4335
-#7807 := [monotonicity #4337 #7804]: #7806
-#9533 := [monotonicity #7807]: #9532
-#9536 := [monotonicity #9533]: #9535
-#9539 := [monotonicity #9536]: #9538
-#9544 := [trans #9539 #9542]: #9543
-#9547 := [monotonicity #9544]: #9546
-#9583 := [monotonicity #9547]: #9582
-#9586 := [quant-intro #9583]: #9585
-#9579 := (iff #2095 #9578)
-#9576 := (iff #2094 #9573)
-#9570 := (implies #9516 #9567)
-#9574 := (iff #9570 #9573)
-#9575 := [rewrite]: #9574
-#9571 := (iff #2094 #9570)
-#9568 := (iff #2093 #9567)
-#9565 := (iff #2092 #9564)
-#9562 := (iff #2091 #9561)
-#9563 := [rewrite]: #9562
-#9566 := [monotonicity #9563]: #9565
-#9559 := (iff #2089 #9558)
-#9560 := [rewrite]: #9559
-#9569 := [monotonicity #9560 #9566]: #9568
-#9517 := (iff #2078 #9516)
-#9514 := (iff #2077 #9513)
-#9511 := (iff #2076 #9508)
-#9505 := (and #9502 #1435)
-#9509 := (iff #9505 #9508)
-#9510 := [rewrite]: #9509
-#9506 := (iff #2076 #9505)
-#9503 := (iff #2075 #9502)
-#9504 := [rewrite]: #9503
-#9507 := [monotonicity #9504]: #9506
-#9512 := [trans #9507 #9510]: #9511
-#3956 := (iff #238 #3955)
-#3957 := [rewrite]: #3956
-#9515 := [monotonicity #3957 #9512]: #9514
-#9500 := (iff #2074 #9499)
-#9501 := [rewrite]: #9500
-#9518 := [monotonicity #9501 #9515]: #9517
-#9572 := [monotonicity #9518 #9569]: #9571
-#9577 := [trans #9572 #9575]: #9576
-#9580 := [quant-intro #9577]: #9579
-#9588 := [trans #9580 #9586]: #9587
-#9557 := [asserted]: #2095
-#9589 := [mp #9557 #9588]: #9584
-#17670 := [mp~ #9589 #17669]: #9584
-#21577 := [mp #17670 #21576]: #21574
-#27245 := (not #21574)
-#27105 := (or #27245 #24694 #27172 #27226 #27229 #27234)
-#27222 := (or #27221 #27218)
-#27223 := (not #27222)
-#27156 := (+ 0::int #27155)
-#27157 := (>= #27156 0::int)
-#27158 := (>= 0::int 0::int)
-#27159 := (not #27158)
-#27230 := (or #27229 #27159 #27157 #24694 #27226 #27223)
-#27106 := (or #27245 #27230)
-#27127 := (iff #27106 #27105)
-#27240 := (or #24694 #27172 #27226 #27229 #27234)
-#27123 := (or #27245 #27240)
-#27125 := (iff #27123 #27105)
-#27126 := [rewrite]: #27125
-#27124 := (iff #27106 #27123)
-#27243 := (iff #27230 #27240)
-#27237 := (or #27229 false #27172 #24694 #27226 #27234)
-#27241 := (iff #27237 #27240)
-#27242 := [rewrite]: #27241
-#27238 := (iff #27230 #27237)
-#27235 := (iff #27223 #27234)
-#27232 := (iff #27222 #27231)
-#27233 := [rewrite]: #27232
-#27236 := [monotonicity #27233]: #27235
-#27175 := (iff #27157 #27172)
-#27169 := (>= #27155 0::int)
-#27173 := (iff #27169 #27172)
-#27174 := [rewrite]: #27173
-#27170 := (iff #27157 #27169)
-#27167 := (= #27156 #27155)
-#27168 := [rewrite]: #27167
-#27171 := [monotonicity #27168]: #27170
-#27176 := [trans #27171 #27174]: #27175
-#27165 := (iff #27159 false)
-#27163 := (iff #27159 #3086)
-#27161 := (iff #27158 true)
-#27162 := [rewrite]: #27161
-#27164 := [monotonicity #27162]: #27163
-#27166 := [trans #27164 #12023]: #27165
-#27239 := [monotonicity #27166 #27176 #27236]: #27238
-#27244 := [trans #27239 #27242]: #27243
-#27122 := [monotonicity #27244]: #27124
-#27128 := [trans #27122 #27126]: #27127
-#27107 := [quant-inst]: #27106
-#27129 := [mp #27107 #27128]: #27105
-#27258 := [unit-resolution #27129 #21577 #14275 #27202 #28442]: #27257
-#27259 := [unit-resolution #27258 #28441]: #27234
-#27205 := (or #27231 #27220)
-#27206 := [def-axiom]: #27205
-#27260 := [unit-resolution #27206 #27259 #27192]: false
-#27262 := [lemma #27260]: #27261
-#28848 := [unit-resolution #27262 #28847]: #11958
-#26988 := (or #23426 #18341 #23441)
-#26983 := [hypothesis]: #23438
-#26984 := [hypothesis]: #11958
-#26985 := [hypothesis]: #23429
-#24651 := (or #23426 #11955)
-#24652 := [def-axiom]: #24651
-#26980 := [unit-resolution #24652 #26985]: #11955
-#24663 := (or #23432 #23426)
-#24664 := [def-axiom]: #24663
-#26986 := [unit-resolution #24664 #26985]: #23432
-#24673 := (or #23441 #18338 #18341 #23435)
-#24674 := [def-axiom]: #24673
-#26987 := [unit-resolution #24674 #26986 #26980 #26984 #26983]: false
-#26989 := [lemma #26987]: #26988
-#28849 := [unit-resolution #26989 #28848 #28845]: #23426
-#28584 := (or #11964 #27226)
-#26967 := (uf_58 #3157 #3031)
-#27290 := (uf_135 #26967)
-#27293 := (uf_25 uf_287 #27290)
-#27294 := (= uf_26 #27293)
-#27291 := (uf_210 uf_287 #27290)
-#27292 := (= uf_9 #27291)
-#27400 := (or #27292 #27294)
-#27413 := (not #27400)
-#27282 := (uf_136 #26967)
-#27283 := (= uf_9 #27282)
-#27284 := (not #27283)
-#27280 := (uf_27 uf_287 #27290)
-#27281 := (= uf_9 #27280)
-#27276 := (not #27281)
-#27395 := (or #27276 #27284)
-#27397 := (not #27395)
-#27031 := (uf_13 #3031)
-#27305 := (uf_12 #27031)
-#27355 := (= uf_14 #27305)
-#27392 := (not #27355)
-#27277 := (uf_13 #27290)
-#27278 := (uf_12 #27277)
-#27279 := (= uf_14 #27278)
-#27438 := (or #27279 #27392 #27397 #27413)
-#27442 := (not #27438)
-#27311 := (uf_25 uf_287 #3031)
-#27312 := (= uf_26 #27311)
-#27304 := (uf_210 uf_287 #3031)
-#27310 := (= uf_9 #27304)
-#27357 := (or #27310 #27312)
-#27360 := (not #27357)
-#27403 := (or #27355 #27360)
-#27406 := (not #27403)
-#27450 := (or #27406 #27442)
-#27454 := (not #27450)
-#27451 := (or #18341 #27454)
-#27459 := (not #27451)
-#27466 := (iff #11964 #27459)
-#27471 := (or #27470 #27466)
-#27295 := (or #27294 #27292)
-#27296 := (not #27295)
-#27285 := (or #27284 #27276)
-#27286 := (not #27285)
-#27306 := (= #27305 uf_14)
-#27307 := (not #27306)
-#27308 := (or #27307 #27286 #27279 #27296)
-#27309 := (not #27308)
-#27313 := (or #27312 #27310)
-#27314 := (not #27313)
-#27362 := (or #27306 #27314)
-#27363 := (not #27362)
-#27364 := (or #27363 #27309)
-#27365 := (not #27364)
-#27366 := (or #18341 #27365)
-#27367 := (not #27366)
-#27354 := (iff #11964 #27367)
-#27472 := (or #27470 #27354)
-#27499 := (iff #27472 #27471)
-#27501 := (iff #27471 #27471)
-#27502 := [rewrite]: #27501
-#27467 := (iff #27354 #27466)
-#27460 := (iff #27367 #27459)
-#27457 := (iff #27366 #27451)
-#27455 := (iff #27365 #27454)
-#27452 := (iff #27364 #27450)
-#27443 := (iff #27309 #27442)
-#27440 := (iff #27308 #27438)
-#27435 := (or #27392 #27397 #27279 #27413)
-#27439 := (iff #27435 #27438)
-#27434 := [rewrite]: #27439
-#27436 := (iff #27308 #27435)
-#27414 := (iff #27296 #27413)
-#27401 := (iff #27295 #27400)
-#27412 := [rewrite]: #27401
-#27433 := [monotonicity #27412]: #27414
-#27398 := (iff #27286 #27397)
-#27396 := (iff #27285 #27395)
-#27391 := [rewrite]: #27396
-#27399 := [monotonicity #27391]: #27398
-#27393 := (iff #27307 #27392)
-#27353 := (iff #27306 #27355)
-#27356 := [rewrite]: #27353
-#27394 := [monotonicity #27356]: #27393
-#27437 := [monotonicity #27394 #27399 #27433]: #27436
-#27441 := [trans #27437 #27434]: #27440
-#27444 := [monotonicity #27441]: #27443
-#27407 := (iff #27363 #27406)
-#27404 := (iff #27362 #27403)
-#27361 := (iff #27314 #27360)
-#27358 := (iff #27313 #27357)
-#27359 := [rewrite]: #27358
-#27402 := [monotonicity #27359]: #27361
-#27405 := [monotonicity #27356 #27402]: #27404
-#27390 := [monotonicity #27405]: #27407
-#27453 := [monotonicity #27390 #27444]: #27452
-#27456 := [monotonicity #27453]: #27455
-#27458 := [monotonicity #27456]: #27457
-#27461 := [monotonicity #27458]: #27460
-#27468 := [monotonicity #27461]: #27467
-#27500 := [monotonicity #27468]: #27499
-#27503 := [trans #27500 #27502]: #27499
-#27498 := [quant-inst]: #27472
-#27488 := [mp #27498 #27503]: #27471
-#28379 := [unit-resolution #27488 #21987]: #27466
-#27641 := (not #27466)
-#28380 := (or #27641 #27451)
-#28374 := [hypothesis]: #18347
-#27644 := (or #27641 #11964 #27451)
-#27645 := [def-axiom]: #27644
-#27131 := [unit-resolution #27645 #28374]: #28380
-#27132 := [unit-resolution #27131 #28379]: #27451
-#27134 := (or #27459 #27454)
-#27133 := [unit-resolution #27262 #28442]: #11958
-#27642 := (or #27459 #18341 #27454)
-#27643 := [def-axiom]: #27642
-#27135 := [unit-resolution #27643 #27133]: #27134
-#27266 := [unit-resolution #27135 #27132]: #27454
-#27600 := (or #27450 #27438)
-#27598 := [def-axiom]: #27600
-#27368 := [unit-resolution #27598 #27266]: #27438
-#27756 := (not #27218)
-#27417 := (iff #27756 #27284)
-#27415 := (iff #27218 #27283)
-#27410 := (= #27217 #27282)
-#27371 := (= #27216 #26967)
-#27372 := [monotonicity #27211]: #27371
-#27411 := [monotonicity #27372]: #27410
-#27416 := [monotonicity #27411]: #27415
-#27686 := [monotonicity #27416]: #27417
-#27757 := (or #27231 #27756)
-#27758 := [def-axiom]: #27757
-#27370 := [unit-resolution #27758 #27259]: #27756
-#27687 := [mp #27370 #27686]: #27284
-#27521 := (or #27395 #27283)
-#27516 := [def-axiom]: #27521
-#27688 := [unit-resolution #27516 #27687]: #27395
-#25393 := (uf_12 uf_7)
-#28494 := (= #25393 #27305)
-#28490 := (= #27305 #25393)
-#28488 := (= #27031 uf_7)
-#24969 := (uf_13 #2979)
-#28486 := (= #24969 uf_7)
-#24970 := (= uf_7 #24969)
-#24975 := (or #24921 #24970)
-#24976 := [quant-inst]: #24975
-#27693 := [unit-resolution #24976 #23160]: #24970
-#28487 := [symm #27693]: #28486
-#28484 := (= #27031 #24969)
-#28467 := (= #3031 #2979)
-#27062 := (uf_116 #24681)
-#27078 := (uf_43 #24854 #27062)
-#28463 := (= #27078 #2979)
-#27759 := (= #27062 uf_288)
-#28455 := (= #27062 #2980)
-#28453 := (= #24681 #2979)
-#24682 := (= #2979 #24681)
-#93 := (uf_29 #23)
-#23059 := (pattern #93)
-#94 := (uf_28 #93)
-#3569 := (= #23 #94)
-#23060 := (forall (vars (?x14 T5)) (:pat #23059) #3569)
-#3572 := (forall (vars (?x14 T5)) #3569)
-#23061 := (iff #3572 #23060)
-#23063 := (iff #23060 #23060)
-#23064 := [rewrite]: #23063
-#23062 := [rewrite]: #23061
-#23065 := [trans #23062 #23064]: #23061
-#16237 := (~ #3572 #3572)
-#16227 := (~ #3569 #3569)
-#16228 := [refl]: #16227
-#16293 := [nnf-pos #16228]: #16237
-#95 := (= #94 #23)
-#96 := (forall (vars (?x14 T5)) #95)
-#3573 := (iff #96 #3572)
-#3570 := (iff #95 #3569)
-#3571 := [rewrite]: #3570
-#3574 := [quant-intro #3571]: #3573
-#3568 := [asserted]: #96
-#3577 := [mp #3568 #3574]: #3572
-#16294 := [mp~ #3577 #16293]: #3572
-#23066 := [mp #16294 #23065]: #23060
-#24685 := (not #23060)
-#24686 := (or #24685 #24682)
-#24687 := [quant-inst]: #24686
-#28452 := [unit-resolution #24687 #23066]: #24682
-#28454 := [symm #28452]: #28453
-#28456 := [monotonicity #28454]: #28455
-#27760 := [trans #28456 #28406]: #27759
-#28337 := [monotonicity #28401 #27760]: #28463
-#28477 := (= #3031 #27078)
-#27056 := (uf_66 #24681 0::int #24854)
-#27081 := (= #27056 #27078)
-#27084 := (not #27081)
-decl uf_138 :: (-> T5 T5 T2)
-#27057 := (uf_138 #27056 #24681)
-#27058 := (= uf_9 #27057)
-#27059 := (not #27058)
-#27090 := (or #27059 #27084)
-#27095 := (not #27090)
-#1576 := (uf_66 #21 #247 #233)
-#1577 := (pattern #1576)
-#1578 := (uf_138 #1576 #21)
-#8249 := (= uf_9 #1578)
-#21125 := (not #8249)
-decl uf_139 :: (-> T3 int)
-#1581 := (uf_139 #233)
-#1582 := (* #247 #1581)
-#1580 := (uf_116 #21)
-#1583 := (+ #1580 #1582)
-#1584 := (uf_43 #233 #1583)
-#1585 := (= #1576 #1584)
-#21124 := (not #1585)
-#21126 := (or #21124 #21125)
-#21127 := (not #21126)
-#21130 := (forall (vars (?x375 T5) (?x376 int) (?x377 T3)) (:pat #1577) #21127)
-#8255 := (and #1585 #8249)
-#8260 := (forall (vars (?x375 T5) (?x376 int) (?x377 T3)) (:pat #1577) #8255)
-#21131 := (iff #8260 #21130)
-#21128 := (iff #8255 #21127)
-#21129 := [rewrite]: #21128
-#21132 := [quant-intro #21129]: #21131
-#17258 := (~ #8260 #8260)
-#17256 := (~ #8255 #8255)
-#17257 := [refl]: #17256
-#17259 := [nnf-pos #17257]: #17258
-#1579 := (= #1578 uf_9)
-#1586 := (and #1579 #1585)
-#1587 := (forall (vars (?x375 T5) (?x376 int) (?x377 T3)) (:pat #1577) #1586)
-#8261 := (iff #1587 #8260)
-#8258 := (iff #1586 #8255)
-#8252 := (and #8249 #1585)
-#8256 := (iff #8252 #8255)
-#8257 := [rewrite]: #8256
-#8253 := (iff #1586 #8252)
-#8250 := (iff #1579 #8249)
-#8251 := [rewrite]: #8250
-#8254 := [monotonicity #8251]: #8253
-#8259 := [trans #8254 #8257]: #8258
-#8262 := [quant-intro #8259]: #8261
-#8248 := [asserted]: #1587
-#8265 := [mp #8248 #8262]: #8260
-#17260 := [mp~ #8265 #17259]: #8260
-#21133 := [mp #17260 #21132]: #21130
-#27098 := (not #21130)
-#27044 := (or #27098 #27095)
-#27060 := (uf_139 #24854)
-#27061 := (* 0::int #27060)
-#27063 := (+ #27062 #27061)
-#27064 := (uf_43 #24854 #27063)
-#27065 := (= #27056 #27064)
-#27066 := (not #27065)
-#27067 := (or #27066 #27059)
-#27068 := (not #27067)
-#27045 := (or #27098 #27068)
-#27130 := (iff #27045 #27044)
-#27204 := (iff #27044 #27044)
-#27207 := [rewrite]: #27204
-#27096 := (iff #27068 #27095)
-#27093 := (iff #27067 #27090)
-#27087 := (or #27084 #27059)
-#27091 := (iff #27087 #27090)
-#27092 := [rewrite]: #27091
-#27088 := (iff #27067 #27087)
-#27085 := (iff #27066 #27084)
-#27082 := (iff #27065 #27081)
-#27079 := (= #27064 #27078)
-#27076 := (= #27063 #27062)
-#27071 := (+ #27062 0::int)
-#27074 := (= #27071 #27062)
-#27075 := [rewrite]: #27074
-#27072 := (= #27063 #27071)
-#27069 := (= #27061 0::int)
-#27070 := [rewrite]: #27069
-#27073 := [monotonicity #27070]: #27072
-#27077 := [trans #27073 #27075]: #27076
-#27080 := [monotonicity #27077]: #27079
-#27083 := [monotonicity #27080]: #27082
-#27086 := [monotonicity #27083]: #27085
-#27089 := [monotonicity #27086]: #27088
-#27094 := [trans #27089 #27092]: #27093
-#27097 := [monotonicity #27094]: #27096
-#27203 := [monotonicity #27097]: #27130
-#27208 := [trans #27203 #27207]: #27130
-#27108 := [quant-inst]: #27045
-#27198 := [mp #27108 #27208]: #27044
-#27790 := [unit-resolution #27198 #21133]: #27095
-#27248 := (or #27090 #27081)
-#27249 := [def-axiom]: #27248
-#27845 := [unit-resolution #27249 #27790]: #27081
-#28296 := (= #3031 #27056)
-#28336 := [monotonicity #28452 #28399]: #28296
-#28478 := [trans #28336 #27845]: #28477
-#28468 := [trans #28478 #28337]: #28467
-#28485 := [monotonicity #28468]: #28484
-#28489 := [trans #28485 #28487]: #28488
-#28491 := [monotonicity #28489]: #28490
-#28495 := [symm #28491]: #28494
-#25394 := (= uf_14 #25393)
-#25401 := (iff #11413 #25394)
+#10533 := [quant-intro #10530]: #10532
+#10512 := (iff #2412 #10511)
+#10509 := (iff #2411 #10508)
+#10506 := (iff #2410 #10505)
+#10503 := (iff #2409 #10502)
+#10500 := (iff #2408 #10499)
+#10497 := (iff #2407 #10496)
+#10494 := (iff #2406 #10493)
+#10495 := [rewrite]: #10494
+#10498 := [monotonicity #3617 #10495]: #10497
+#10501 := [monotonicity #10498]: #10500
+#10491 := (iff #2404 #10490)
+#10488 := (iff #2403 #10487)
+#10485 := (iff #2402 #10484)
+#10482 := (iff #2401 #10481)
+#10479 := (iff #2400 #10478)
+#10480 := [rewrite]: #10479
+#10476 := (iff #2398 #10475)
+#10477 := [rewrite]: #10476
+#10483 := [monotonicity #10477 #10480]: #10482
+#10473 := (iff #2396 #10472)
+#10470 := (iff #2395 #10469)
+#10471 := [rewrite]: #10470
+#10474 := [monotonicity #10471]: #10473
+#10486 := [monotonicity #10474 #10483]: #10485
+#10467 := (iff #2392 #10466)
+#10464 := (iff #2391 #10463)
+#10461 := (iff #2390 #10460)
+#10462 := [rewrite]: #10461
+#10465 := [monotonicity #10462]: #10464
+#10458 := (iff #2387 #10457)
+#10455 := (iff #2386 #10454)
+#10456 := [rewrite]: #10455
+#10459 := [monotonicity #10456]: #10458
+#10468 := [monotonicity #10459 #10465]: #10467
+#10489 := [monotonicity #10468 #10486]: #10488
+#10492 := [monotonicity #10489]: #10491
+#10504 := [monotonicity #10492 #10501]: #10503
+#10507 := [monotonicity #3623 #10504]: #10506
+#10452 := (iff #2383 #10450)
+#10453 := [rewrite]: #10452
+#10510 := [monotonicity #10453 #10507]: #10509
+#10513 := [quant-intro #10510]: #10512
+#10535 := [trans #10513 #10533]: #10534
+#10449 := [asserted]: #2412
+#10536 := [mp #10449 #10535]: #10531
+#17197 := [mp~ #10536 #17196]: #10531
+#21176 := [mp #17197 #21175]: #21173
+#27590 := (not #21173)
+#27615 := (or #27590 #27541)
+#26727 := (or #26752 #26750)
+#26754 := (not #26727)
+#26776 := (= #26775 uf_14)
+#26814 := (not #26776)
+#26815 := (or #26814 #26774 #26757 #26754)
+#26816 := (not #26815)
+#26812 := (or #26776 #26809)
+#26807 := (not #26812)
+#26813 := (or #26807 #26816)
+#26866 := (not #26813)
+#26990 := (or #26989 #26866)
+#26991 := (not #26990)
+#26937 := (iff #26997 #26991)
+#26626 := (or #27590 #26937)
+#26665 := (iff #26626 #27615)
+#26670 := (iff #27615 #27615)
+#26671 := [rewrite]: #26670
+#27542 := (iff #26937 #27541)
+#27038 := (iff #26991 #27035)
+#27034 := (iff #26990 #27033)
+#27031 := (iff #26866 #27030)
+#27013 := (iff #26813 #27012)
+#27008 := (iff #26816 #27005)
+#27554 := (iff #26815 #27536)
+#27550 := (or #27415 #26774 #26757 #27420)
+#27534 := (iff #27550 #27536)
+#27533 := [rewrite]: #27534
+#27546 := (iff #26815 #27550)
+#27475 := (iff #26754 #27420)
+#27418 := (iff #26727 #27416)
+#27419 := [rewrite]: #27418
+#27549 := [monotonicity #27419]: #27475
+#27410 := (iff #26814 #27415)
+#26936 := (iff #26776 #26938)
+#26983 := [rewrite]: #26936
+#27048 := [monotonicity #26983]: #27410
+#27535 := [monotonicity #27048 #27549]: #27546
+#27539 := [trans #27535 #27533]: #27554
+#27011 := [monotonicity #27539]: #27008
+#27413 := (iff #26807 #27412)
+#27010 := (iff #26812 #26987)
+#26984 := (or #26938 #26809)
+#26982 := (iff #26984 #26987)
+#27009 := [rewrite]: #26982
+#26985 := (iff #26812 #26984)
+#26986 := [monotonicity #26983]: #26985
+#27047 := [trans #26986 #27009]: #27010
+#27414 := [monotonicity #27047]: #27413
+#27028 := [monotonicity #27414 #27011]: #27013
+#27032 := [monotonicity #27028]: #27031
+#27029 := [monotonicity #27032]: #27034
+#27540 := [monotonicity #27029]: #27038
+#27543 := [monotonicity #27540]: #27542
+#26666 := [monotonicity #27543]: #26665
+#26672 := [trans #26666 #26671]: #26665
+#26627 := [quant-inst]: #26626
+#26941 := [mp #26627 #26672]: #27615
+#27807 := [unit-resolution #26941 #21176]: #27541
+#27775 := (not #26997)
+#27815 := (iff #17640 #27775)
+#27813 := (iff #11979 #26997)
+#27811 := (iff #26997 #11979)
+#27809 := (= #26992 #3050)
+#27810 := [monotonicity #27800]: #27809
+#27812 := [monotonicity #27810]: #27811
+#27814 := [symm #27812]: #27813
+#27816 := [monotonicity #27814]: #27815
+#27808 := [hypothesis]: #17640
+#27817 := [mp #27808 #27816]: #27775
+#27772 := (not #27541)
+#27773 := (or #27772 #26997 #27033)
+#27774 := [def-axiom]: #27773
+#27818 := [unit-resolution #27774 #27817 #27807]: #27033
+#27770 := (or #27035 #26989 #27030)
+#27771 := [def-axiom]: #27770
+#27819 := [unit-resolution #27771 #27818 #27806]: #27030
+#27835 := (iff #26562 #26746)
+#27833 := (iff #26473 #26745)
+#27831 := (iff #26745 #26473)
+#27829 := (= #26744 #26472)
+#27827 := (= #26742 #26471)
+#26222 := (uf_58 #3175 #3044)
+#27825 := (= #26222 #26471)
+#27823 := (= #26471 #26222)
+#27824 := [monotonicity #26466]: #27823
+#27826 := [symm #27824]: #27825
+#27821 := (= #26742 #26222)
+#27822 := [monotonicity #27800]: #27821
+#27828 := [trans #27822 #27826]: #27827
+#27830 := [monotonicity #27828]: #27829
+#27832 := [monotonicity #27830]: #27831
+#27834 := [symm #27832]: #27833
+#27836 := [monotonicity #27834]: #27835
+#27820 := [hypothesis]: #26562
+#27837 := [mp #27820 #27836]: #26746
+#27732 := (or #26747 #26745)
+#27733 := [def-axiom]: #27732
+#27838 := [unit-resolution #27733 #27837]: #26747
+#24648 := (uf_12 uf_7)
+#27854 := (= #24648 #26775)
+#27850 := (= #26775 #24648)
+#27851 := [monotonicity #27849]: #27850
+#27855 := [symm #27851]: #27854
+#24649 := (= uf_14 #24648)
+#24656 := (iff #11384 #24649)
#2308 := (pattern #237)
#2836 := (uf_12 #233)
-#11586 := (= uf_14 #2836)
-#11590 := (iff #3955 #11586)
-#11593 := (forall (vars (?x761 T3)) (:pat #2308) #11590)
-#18295 := (~ #11593 #11593)
-#18293 := (~ #11590 #11590)
-#18294 := [refl]: #18293
-#18296 := [nnf-pos #18294]: #18295
+#11557 := (= uf_14 #2836)
+#11561 := (iff #3926 #11557)
+#11564 := (forall (vars (?x761 T3)) (:pat #2308) #11561)
+#17583 := (~ #11564 #11564)
+#17581 := (~ #11561 #11561)
+#17582 := [refl]: #17581
+#17584 := [nnf-pos #17582]: #17583
#2849 := (= #2836 uf_14)
#2850 := (iff #238 #2849)
#2851 := (forall (vars (?x761 T3)) (:pat #2308) #2850)
-#11594 := (iff #2851 #11593)
-#11591 := (iff #2850 #11590)
-#11588 := (iff #2849 #11586)
-#11589 := [rewrite]: #11588
-#11592 := [monotonicity #3957 #11589]: #11591
-#11595 := [quant-intro #11592]: #11594
-#11585 := [asserted]: #2851
-#11598 := [mp #11585 #11595]: #11593
-#18297 := [mp~ #11598 #18296]: #11593
-#25025 := (not #11593)
-#25404 := (or #25025 #25401)
-#25405 := [quant-inst]: #25404
-#27689 := [unit-resolution #25405 #18297]: #25401
-#25406 := (not #25401)
-#27690 := (or #25406 #25394)
-#25410 := (not #11413)
-#25411 := (or #25406 #25410 #25394)
-#25412 := [def-axiom]: #25411
-#27691 := [unit-resolution #25412 #11419]: #27690
-#27692 := [unit-resolution #27691 #27689]: #25394
-#28496 := [trans #27692 #28495]: #27355
-#27552 := (not #27279)
-#28573 := (iff #11905 #27552)
-#28571 := (iff #11902 #27279)
-#28569 := (iff #27279 #11902)
-#28567 := (= #27278 #2990)
-#28565 := (= #27277 #2977)
-#28561 := (= #27277 #24974)
-#28559 := (= #27290 #2981)
-#28557 := (= #27290 #27137)
-#27138 := (uf_66 #27137 0::int #24854)
-#27142 := (uf_58 #3157 #27138)
-#27145 := (uf_135 #27142)
-#28555 := (= #27145 #27137)
-#27146 := (= #27137 #27145)
-#27148 := (up_67 #27142)
-#27149 := (not #27148)
-#27147 := (not #27146)
-#27143 := (uf_136 #27142)
-#27144 := (= uf_9 #27143)
-#27139 := (uf_24 uf_287 #27138)
-#27140 := (= uf_9 #27139)
-#27141 := (not #27140)
-#27177 := (or #27141 #27144 #27147 #27149)
-#27180 := (not #27177)
-#27152 := (uf_24 uf_287 #27137)
-#27153 := (= uf_9 #27152)
-#28507 := (= #2988 #27152)
-#28504 := (= #27152 #2988)
-#28505 := [monotonicity #28503]: #28504
-#28508 := [symm #28505]: #28507
-#28509 := [trans #14288 #28508]: #27153
-#27154 := (not #27153)
-#28510 := (or #27154 #27180)
+#11565 := (iff #2851 #11564)
+#11562 := (iff #2850 #11561)
+#11559 := (iff #2849 #11557)
+#11560 := [rewrite]: #11559
+#11563 := [monotonicity #3928 #11560]: #11562
+#11566 := [quant-intro #11563]: #11565
+#11556 := [asserted]: #2851
+#11569 := [mp #11556 #11566]: #11564
+#17585 := [mp~ #11569 #17584]: #11564
+#24280 := (not #11564)
+#24659 := (or #24280 #24656)
+#24660 := [quant-inst]: #24659
+#27839 := [unit-resolution #24660 #17585]: #24656
+#24661 := (not #24656)
+#27840 := (or #24661 #24649)
+#24665 := (not #11384)
+#24666 := (or #24661 #24665 #24649)
+#24667 := [def-axiom]: #24666
+#27841 := [unit-resolution #24667 #11390]: #27840
+#27842 := [unit-resolution #27841 #27839]: #24649
+#27856 := [trans #27842 #27855]: #26938
+#27751 := (not #26757)
+#27930 := (iff #11876 #27751)
+#27928 := (iff #11873 #26757)
+#27926 := (iff #26757 #11873)
+#27924 := (= #26756 #2990)
+#27922 := (= #26755 #2977)
+#27918 := (= #26755 #24228)
+#27916 := (= #26748 #2981)
+#27914 := (= #26748 #26392)
+#26393 := (uf_66 #26392 0::int #24108)
+#26397 := (uf_58 #3175 #26393)
+#26400 := (uf_135 #26397)
+#27912 := (= #26400 #26392)
+#26401 := (= #26392 #26400)
+#26403 := (up_67 #26397)
+#26404 := (not #26403)
+#26402 := (not #26401)
+#26398 := (uf_136 #26397)
+#26399 := (= uf_9 #26398)
+#26394 := (uf_24 uf_287 #26393)
+#26395 := (= uf_9 #26394)
+#26396 := (not #26395)
+#26432 := (or #26396 #26399 #26402 #26404)
+#26435 := (not #26432)
+#26407 := (uf_24 uf_287 #26392)
+#26408 := (= uf_9 #26407)
+#27867 := (= #2988 #26407)
+#27864 := (= #26407 #2988)
+#27865 := [monotonicity #27863]: #27864
+#27868 := [symm #27865]: #27867
+#27869 := [trans #13532 #27868]: #26408
+#26409 := (not #26408)
+#27870 := (or #26409 #26435)
#277 := (:var 3 int)
#310 := (:var 2 T3)
#1470 := (uf_124 #310 #247)
@@ -6395,59 +5668,59 @@
#1472 := (pattern #1469 #1471)
#1478 := (uf_66 #1471 #161 #310)
#1486 := (uf_24 #35 #1478)
-#7960 := (= uf_9 #1486)
-#20901 := (not #7960)
+#7931 := (= uf_9 #1486)
+#20090 := (not #7931)
#1479 := (uf_58 #1473 #1478)
#1482 := (uf_136 #1479)
-#7954 := (= uf_9 #1482)
+#7925 := (= uf_9 #1482)
#1480 := (uf_135 #1479)
-#7951 := (= #1471 #1480)
-#20900 := (not #7951)
+#7922 := (= #1471 #1480)
+#20089 := (not #7922)
#1485 := (up_67 #1479)
-#20899 := (not #1485)
-#20902 := (or #20899 #20900 #7954 #20901)
-#20903 := (not #20902)
+#20088 := (not #1485)
+#20091 := (or #20088 #20089 #7925 #20090)
+#20092 := (not #20091)
#1476 := (uf_24 #35 #1471)
-#7948 := (= uf_9 #1476)
-#7983 := (not #7948)
-#5263 := (* -1::int #247)
-#6143 := (+ #161 #5263)
-#6144 := (>= #6143 0::int)
-#20909 := (or #4992 #6144 #7983 #20903)
-#20914 := (forall (vars (?x346 T4) (?x347 int) (?x348 T3) (?x349 int) (?x350 int)) (:pat #1472 #1475) #20909)
-#7957 := (not #7954)
-#7992 := (and #1485 #7951 #7957 #7960)
-#7647 := (not #6144)
-#7650 := (and #4070 #7647)
-#7653 := (not #7650)
-#8001 := (or #7653 #7983 #7992)
-#8006 := (forall (vars (?x346 T4) (?x347 int) (?x348 T3) (?x349 int) (?x350 int)) (:pat #1472 #1475) #8001)
-#20915 := (iff #8006 #20914)
-#20912 := (iff #8001 #20909)
-#20784 := (or #4992 #6144)
-#20906 := (or #20784 #7983 #20903)
-#20910 := (iff #20906 #20909)
-#20911 := [rewrite]: #20910
-#20907 := (iff #8001 #20906)
-#20904 := (iff #7992 #20903)
-#20905 := [rewrite]: #20904
-#20793 := (iff #7653 #20784)
-#20785 := (not #20784)
-#20788 := (not #20785)
-#20791 := (iff #20788 #20784)
-#20792 := [rewrite]: #20791
-#20789 := (iff #7653 #20788)
-#20786 := (iff #7650 #20785)
-#20787 := [rewrite]: #20786
-#20790 := [monotonicity #20787]: #20789
-#20794 := [trans #20790 #20792]: #20793
-#20908 := [monotonicity #20794 #20905]: #20907
-#20913 := [trans #20908 #20911]: #20912
-#20916 := [quant-intro #20913]: #20915
-#17035 := (~ #8006 #8006)
-#17033 := (~ #8001 #8001)
-#17034 := [refl]: #17033
-#17036 := [nnf-pos #17034]: #17035
+#7919 := (= uf_9 #1476)
+#7954 := (not #7919)
+#5234 := (* -1::int #247)
+#6114 := (+ #161 #5234)
+#6115 := (>= #6114 0::int)
+#20098 := (or #4963 #6115 #7954 #20092)
+#20103 := (forall (vars (?x346 T4) (?x347 int) (?x348 T3) (?x349 int) (?x350 int)) (:pat #1472 #1475) #20098)
+#7928 := (not #7925)
+#7963 := (and #1485 #7922 #7928 #7931)
+#7618 := (not #6115)
+#7621 := (and #4041 #7618)
+#7624 := (not #7621)
+#7972 := (or #7624 #7954 #7963)
+#7977 := (forall (vars (?x346 T4) (?x347 int) (?x348 T3) (?x349 int) (?x350 int)) (:pat #1472 #1475) #7972)
+#20104 := (iff #7977 #20103)
+#20101 := (iff #7972 #20098)
+#19973 := (or #4963 #6115)
+#20095 := (or #19973 #7954 #20092)
+#20099 := (iff #20095 #20098)
+#20100 := [rewrite]: #20099
+#20096 := (iff #7972 #20095)
+#20093 := (iff #7963 #20092)
+#20094 := [rewrite]: #20093
+#19982 := (iff #7624 #19973)
+#19974 := (not #19973)
+#19977 := (not #19974)
+#19980 := (iff #19977 #19973)
+#19981 := [rewrite]: #19980
+#19978 := (iff #7624 #19977)
+#19975 := (iff #7621 #19974)
+#19976 := [rewrite]: #19975
+#19979 := [monotonicity #19976]: #19978
+#19983 := [trans #19979 #19981]: #19982
+#20097 := [monotonicity #19983 #20094]: #20096
+#20102 := [trans #20097 #20100]: #20101
+#20105 := [quant-intro #20102]: #20104
+#16323 := (~ #7977 #7977)
+#16321 := (~ #7972 #7972)
+#16322 := [refl]: #16321
+#16324 := [nnf-pos #16322]: #16323
#1487 := (= #1486 uf_9)
#1488 := (and #1485 #1487)
#1483 := (= #1482 uf_9)
@@ -6461,1541 +5734,1585 @@
#1477 := (= #1476 uf_9)
#1492 := (implies #1477 #1491)
#1493 := (forall (vars (?x346 T4) (?x347 int) (?x348 T3) (?x349 int) (?x350 int)) (:pat #1472 #1475) #1492)
-#8009 := (iff #1493 #8006)
-#7963 := (and #1485 #7960)
-#7966 := (and #7957 #7963)
-#7969 := (and #7951 #7966)
-#7617 := (not #1372)
-#7975 := (or #7617 #7969)
-#7984 := (or #7983 #7975)
-#7989 := (forall (vars (?x346 T4) (?x347 int) (?x348 T3) (?x349 int) (?x350 int)) (:pat #1472 #1475) #7984)
-#8007 := (iff #7989 #8006)
-#8004 := (iff #7984 #8001)
-#7995 := (or #7653 #7992)
-#7998 := (or #7983 #7995)
-#8002 := (iff #7998 #8001)
-#8003 := [rewrite]: #8002
-#7999 := (iff #7984 #7998)
-#7996 := (iff #7975 #7995)
-#7993 := (iff #7969 #7992)
-#7994 := [rewrite]: #7993
-#7654 := (iff #7617 #7653)
-#7651 := (iff #1372 #7650)
-#7648 := (iff #1371 #7647)
-#7649 := [rewrite]: #7648
-#7652 := [monotonicity #4069 #7649]: #7651
-#7655 := [monotonicity #7652]: #7654
-#7997 := [monotonicity #7655 #7994]: #7996
-#8000 := [monotonicity #7997]: #7999
-#8005 := [trans #8000 #8003]: #8004
-#8008 := [quant-intro #8005]: #8007
-#7990 := (iff #1493 #7989)
-#7987 := (iff #1492 #7984)
-#7980 := (implies #7948 #7975)
-#7985 := (iff #7980 #7984)
-#7986 := [rewrite]: #7985
-#7981 := (iff #1492 #7980)
-#7978 := (iff #1491 #7975)
-#7972 := (implies #1372 #7969)
-#7976 := (iff #7972 #7975)
-#7977 := [rewrite]: #7976
-#7973 := (iff #1491 #7972)
-#7970 := (iff #1490 #7969)
-#7967 := (iff #1489 #7966)
-#7964 := (iff #1488 #7963)
-#7961 := (iff #1487 #7960)
-#7962 := [rewrite]: #7961
-#7965 := [monotonicity #7962]: #7964
-#7958 := (iff #1484 #7957)
-#7955 := (iff #1483 #7954)
-#7956 := [rewrite]: #7955
-#7959 := [monotonicity #7956]: #7958
-#7968 := [monotonicity #7959 #7965]: #7967
-#7952 := (iff #1481 #7951)
-#7953 := [rewrite]: #7952
-#7971 := [monotonicity #7953 #7968]: #7970
-#7974 := [monotonicity #7971]: #7973
-#7979 := [trans #7974 #7977]: #7978
-#7949 := (iff #1477 #7948)
-#7950 := [rewrite]: #7949
-#7982 := [monotonicity #7950 #7979]: #7981
-#7988 := [trans #7982 #7986]: #7987
-#7991 := [quant-intro #7988]: #7990
-#8010 := [trans #7991 #8008]: #8009
-#7947 := [asserted]: #1493
-#8011 := [mp #7947 #8010]: #8006
-#17037 := [mp~ #8011 #17036]: #8006
-#20917 := [mp #17037 #20916]: #20914
-#27680 := (not #20914)
-#27681 := (or #27680 #27154 #27172 #27180)
-#27150 := (or #27149 #27147 #27144 #27141)
-#27151 := (not #27150)
-#27160 := (or #27159 #27157 #27154 #27151)
-#27682 := (or #27680 #27160)
-#27727 := (iff #27682 #27681)
-#27186 := (or #27154 #27172 #27180)
-#27728 := (or #27680 #27186)
-#27731 := (iff #27728 #27681)
-#27732 := [rewrite]: #27731
-#27729 := (iff #27682 #27728)
-#27189 := (iff #27160 #27186)
-#27183 := (or false #27172 #27154 #27180)
-#27187 := (iff #27183 #27186)
-#27188 := [rewrite]: #27187
-#27184 := (iff #27160 #27183)
-#27181 := (iff #27151 #27180)
-#27178 := (iff #27150 #27177)
-#27179 := [rewrite]: #27178
-#27182 := [monotonicity #27179]: #27181
-#27185 := [monotonicity #27166 #27176 #27182]: #27184
-#27190 := [trans #27185 #27188]: #27189
-#27730 := [monotonicity #27190]: #27729
-#27733 := [trans #27730 #27732]: #27727
-#27683 := [quant-inst]: #27682
-#27734 := [mp #27683 #27733]: #27681
-#28511 := [unit-resolution #27734 #20917 #27202]: #28510
-#28512 := [unit-resolution #28511 #28509]: #27180
-#27751 := (or #27177 #27146)
-#27752 := [def-axiom]: #27751
-#28513 := [unit-resolution #27752 #28512]: #27146
-#28556 := [symm #28513]: #28555
-#28553 := (= #27290 #27145)
-#28551 := (= #26967 #27142)
-#28549 := (= #27142 #26967)
-#28547 := (= #27138 #3031)
-#28537 := (= #27056 #3031)
-#28538 := [symm #28336]: #28537
-#28545 := (= #27138 #27056)
-#28535 := (= #27078 #27056)
-#28536 := [symm #27845]: #28535
-#28543 := (= #27138 #27078)
-#28461 := (= #2979 #27078)
-#28534 := [symm #28337]: #28461
-#28541 := (= #27138 #2979)
-#27319 := (uf_116 #26144)
-#27333 := (uf_43 #24854 #27319)
-#28532 := (= #27333 #2979)
-#28524 := (= #27319 uf_288)
-#28522 := (= #27319 #2980)
-#28520 := (= #24756 #2980)
-#24980 := (= #2980 #24756)
-#24985 := (or #24927 #24980)
-#24986 := [quant-inst]: #24985
-#28514 := [unit-resolution #24986 #23154]: #24980
-#28521 := [symm #28514]: #28520
-#28518 := (= #27319 #24756)
-#28519 := [monotonicity #28517]: #28518
-#28523 := [trans #28519 #28521]: #28522
-#28525 := [trans #28523 #28406]: #28524
-#28533 := [monotonicity #28401 #28525]: #28532
-#28539 := (= #27138 #27333)
-#27315 := (uf_66 #26144 0::int #24854)
-#27336 := (= #27315 #27333)
-#27339 := (not #27336)
-#27316 := (uf_138 #27315 #26144)
-#27317 := (= uf_9 #27316)
-#27318 := (not #27317)
-#27345 := (or #27318 #27339)
-#27350 := (not #27345)
-#28295 := (or #27098 #27350)
-#27320 := (+ #27319 #27061)
-#27321 := (uf_43 #24854 #27320)
-#27322 := (= #27315 #27321)
-#27323 := (not #27322)
-#27324 := (or #27323 #27318)
-#27325 := (not #27324)
-#28280 := (or #27098 #27325)
-#28279 := (iff #28280 #28295)
-#28299 := (iff #28295 #28295)
-#28284 := [rewrite]: #28299
-#27351 := (iff #27325 #27350)
-#27348 := (iff #27324 #27345)
-#27342 := (or #27339 #27318)
-#27346 := (iff #27342 #27345)
-#27347 := [rewrite]: #27346
-#27343 := (iff #27324 #27342)
-#27340 := (iff #27323 #27339)
-#27337 := (iff #27322 #27336)
-#27334 := (= #27321 #27333)
-#27331 := (= #27320 #27319)
-#27326 := (+ #27319 0::int)
-#27329 := (= #27326 #27319)
-#27330 := [rewrite]: #27329
-#27327 := (= #27320 #27326)
-#27328 := [monotonicity #27070]: #27327
-#27332 := [trans #27328 #27330]: #27331
-#27335 := [monotonicity #27332]: #27334
-#27338 := [monotonicity #27335]: #27337
-#27341 := [monotonicity #27338]: #27340
-#27344 := [monotonicity #27341]: #27343
-#27349 := [trans #27344 #27347]: #27348
-#27352 := [monotonicity #27349]: #27351
-#28278 := [monotonicity #27352]: #28279
-#28285 := [trans #28278 #28284]: #28279
-#28281 := [quant-inst]: #28280
-#28286 := [mp #28281 #28285]: #28295
-#28526 := [unit-resolution #28286 #21133]: #27350
-#28335 := (or #27345 #27336)
-#28360 := [def-axiom]: #28335
-#28527 := [unit-resolution #28360 #28526]: #27336
-#28530 := (= #27138 #27315)
-#28531 := [monotonicity #28529]: #28530
-#28540 := [trans #28531 #28527]: #28539
-#28542 := [trans #28540 #28533]: #28541
-#28544 := [trans #28542 #28534]: #28543
-#28546 := [trans #28544 #28536]: #28545
-#28548 := [trans #28546 #28538]: #28547
-#28550 := [monotonicity #28548]: #28549
-#28552 := [symm #28550]: #28551
-#28554 := [monotonicity #28552]: #28553
-#28558 := [trans #28554 #28556]: #28557
-#28560 := [trans #28558 #28503]: #28559
-#28562 := [monotonicity #28560]: #28561
-#28566 := [trans #28562 #28564]: #28565
-#28568 := [monotonicity #28566]: #28567
-#28570 := [monotonicity #28568]: #28569
-#28572 := [symm #28570]: #28571
-#28574 := [monotonicity #28572]: #28573
-#28575 := [mp #14289 #28574]: #27552
-#28579 := (= #2984 #27293)
-#28576 := (= #27293 #2984)
-#28577 := [monotonicity #28560]: #28576
-#28580 := [symm #28577]: #28579
-#28581 := [trans #14286 #28580]: #27294
-#27553 := (not #27294)
-#27554 := (or #27400 #27553)
-#27555 := [def-axiom]: #27554
-#28582 := [unit-resolution #27555 #28581]: #27400
-#27610 := (or #27442 #27279 #27392 #27397 #27413)
-#27611 := [def-axiom]: #27610
-#28583 := [unit-resolution #27611 #28582 #28575 #28496 #27688 #27368]: false
-#28585 := [lemma #28583]: #28584
-#28850 := [unit-resolution #28585 #28847]: #11964
-#28851 := (or #23429 #18347 #23423)
-#27032 := (= #24854 #27031)
-#27005 := (uf_48 #3031 #24854)
-#27006 := (= uf_9 #27005)
-#27034 := (iff #27006 #27032)
-#9068 := (= #233 #1390)
-#11188 := (iff #9068 #11167)
-#23142 := (forall (vars (?x712 T5) (?x713 T3)) (:pat #2667) #11188)
-#11193 := (forall (vars (?x712 T5) (?x713 T3)) #11188)
-#23145 := (iff #11193 #23142)
-#23143 := (iff #11188 #11188)
-#23144 := [refl]: #23143
-#23146 := [quant-intro #23144]: #23145
-#18180 := (~ #11193 #11193)
-#18178 := (~ #11188 #11188)
-#18179 := [refl]: #18178
-#18181 := [nnf-pos #18179]: #18180
-#1890 := (= #1390 #233)
-#2673 := (iff #2668 #1890)
-#2674 := (forall (vars (?x712 T5) (?x713 T3)) #2673)
-#11194 := (iff #2674 #11193)
-#11191 := (iff #2673 #11188)
-#11184 := (iff #11167 #9068)
-#11189 := (iff #11184 #11188)
-#11190 := [rewrite]: #11189
-#11186 := (iff #2673 #11184)
-#9069 := (iff #1890 #9068)
-#9070 := [rewrite]: #9069
-#11187 := [monotonicity #11170 #9070]: #11186
-#11192 := [trans #11187 #11190]: #11191
-#11195 := [quant-intro #11192]: #11194
-#11183 := [asserted]: #2674
-#11198 := [mp #11183 #11195]: #11193
-#18182 := [mp~ #11198 #18181]: #11193
-#23147 := [mp #18182 #23146]: #23142
-#26172 := (not #23142)
-#26981 := (or #26172 #27034)
-#27033 := (iff #27032 #27006)
-#26982 := (or #26172 #27033)
-#27020 := (iff #26982 #26981)
-#27027 := (iff #26981 #26981)
-#27028 := [rewrite]: #27027
-#27035 := (iff #27033 #27034)
-#27036 := [rewrite]: #27035
-#27026 := [monotonicity #27036]: #27020
-#27029 := [trans #27026 #27028]: #27020
-#27025 := [quant-inst]: #26982
-#27007 := [mp #27025 #27029]: #26981
-#27009 := [unit-resolution #27007 #23147]: #27034
-#27013 := (not #27006)
-#27038 := (iff #18338 #27013)
-#27104 := (iff #11955 #27006)
-#27053 := (iff #27006 #11955)
-#27051 := (= #27005 #3032)
-#27052 := [monotonicity #28401]: #27051
-#27109 := [monotonicity #27052]: #27053
-#27054 := [symm #27109]: #27104
-#27039 := [monotonicity #27054]: #27038
-#27050 := [hypothesis]: #18338
-#27037 := [mp #27050 #27039]: #27013
-#27040 := (= #24969 #27031)
-#27041 := [symm #28485]: #27040
-#27055 := (= #24854 #24969)
-#27042 := [trans #28401 #27693]: #27055
-#27043 := [trans #27042 #27041]: #27032
-#27008 := (not #27032)
-#27010 := (not #27034)
-#26971 := (or #27010 #27006 #27008)
-#26994 := [def-axiom]: #26971
-#27111 := [unit-resolution #26994 #27043 #27037 #27009]: false
-#27112 := [lemma #27111]: #11955
-#24657 := (or #23429 #18338 #18347 #23423)
-#24658 := [def-axiom]: #24657
-#28852 := [unit-resolution #24658 #27112]: #28851
-#28853 := [unit-resolution #28852 #28850 #28849]: #23423
-#24635 := (or #23420 #3042)
-#24636 := [def-axiom]: #24635
-#28854 := [unit-resolution #24636 #28853]: #3042
-#24647 := (or #23420 #23414)
-#24648 := [def-axiom]: #24647
-#29533 := [unit-resolution #24648 #28853]: #23414
-#28732 := [hypothesis]: #13392
-#28733 := [th-lemma #14280 #28732]: false
-#28734 := [lemma #28733]: #13389
-#24633 := (or #23417 #13392 #23411)
-#24634 := [def-axiom]: #24633
-#29534 := [unit-resolution #24634 #28734 #29533]: #23411
-#24625 := (or #23408 #23402)
-#24626 := [def-axiom]: #24625
-#29543 := [unit-resolution #24626 #29534]: #23402
-#27621 := (* -1::int #3041)
-#27622 := (+ uf_295 #27621)
-#27623 := (>= #27622 0::int)
-#28855 := (or #13173 #27623)
-#28856 := [th-lemma]: #28855
-#28857 := [unit-resolution #28856 #28854]: #27623
-#24303 := (not #18379)
-#28858 := [hypothesis]: #22230
-#24304 := (or #22225 #24303)
-#24305 := [def-axiom]: #24304
-#28859 := [unit-resolution #24305 #28858]: #24303
-#28818 := (+ #3041 #18377)
-#28820 := (>= #28818 0::int)
-#28817 := (= #3041 #18376)
-#28866 := (= #18376 #3041)
-#28864 := (= #18375 #3031)
-#28862 := (= ?x773!13 0::int)
-#24306 := (not #18380)
-#24307 := (or #22225 #24306)
-#24308 := [def-axiom]: #24307
-#28860 := [unit-resolution #24308 #28858]: #24306
-#24301 := (or #22225 #18372)
-#24302 := [def-axiom]: #24301
-#28861 := [unit-resolution #24302 #28858]: #18372
-#28863 := [th-lemma #28861 #28860]: #28862
-#28865 := [monotonicity #28863]: #28864
-#28867 := [monotonicity #28865]: #28866
-#28868 := [symm #28867]: #28817
-#28869 := (not #28817)
-#28870 := (or #28869 #28820)
-#28871 := [th-lemma]: #28870
-#28872 := [unit-resolution #28871 #28868]: #28820
-#28873 := [th-lemma #28872 #28859 #28857]: false
-#28874 := [lemma #28873]: #22225
-#24621 := (or #23405 #22230 #23399)
-#24622 := [def-axiom]: #24621
-#29544 := [unit-resolution #24622 #28874 #29543]: #23399
-#24613 := (or #23396 #23390)
-#24614 := [def-axiom]: #24613
-#29545 := [unit-resolution #24614 #29544]: #23390
-#29546 := (or #23393 #13173 #23387)
-#24609 := (or #23393 #13173 #13428 #23387)
-#24610 := [def-axiom]: #24609
-#29547 := [unit-resolution #24610 #14280]: #29546
-#29548 := [unit-resolution #29547 #29545 #28854]: #23387
-#24599 := (or #23384 #23378)
-#24600 := [def-axiom]: #24599
-#29549 := [unit-resolution #24600 #29548]: #23378
-#24597 := (or #23384 #23220)
-#24598 := [def-axiom]: #24597
-#29550 := [unit-resolution #24598 #29548]: #23220
-#24571 := (or #23384 #12016)
-#24572 := [def-axiom]: #24571
-#29551 := [unit-resolution #24572 #29548]: #12016
-#24593 := (or #23384 #15788)
-#24594 := [def-axiom]: #24593
-#29552 := [unit-resolution #24594 #29548]: #15788
-#24583 := (or #23384 #13900)
-#24584 := [def-axiom]: #24583
-#29553 := [unit-resolution #24584 #29548]: #13900
-#27613 := (uf_13 #27078)
-#27614 := (uf_66 #27078 uf_297 #27613)
-#27615 := (uf_125 #27614 #27078)
-#27696 := (>= #27615 0::int)
-#24579 := (or #23384 #13433)
-#24580 := [def-axiom]: #24579
-#29554 := [unit-resolution #24580 #29548]: #13433
-#27712 := (* -1::int #27615)
-#27785 := (+ uf_297 #27712)
-#27786 := (<= #27785 0::int)
-#27616 := (= uf_297 #27615)
-#28905 := (uf_66 #24681 uf_297 #27114)
-#28906 := (uf_125 #28905 #24681)
-#29576 := (= #28906 #27615)
-#29574 := (= #27615 #28906)
-#29555 := (= #27078 #24681)
-#29556 := [trans #28337 #28452]: #29555
-#29571 := (= #27614 #28905)
-#29569 := (= #3082 #28905)
-#29567 := (= #28905 #3082)
-#29559 := (= #27114 uf_7)
-#29557 := (= #27114 #24969)
-#29558 := [monotonicity #28454]: #29557
-#29560 := [trans #29558 #28487]: #29559
-#29568 := [monotonicity #28454 #29560]: #29567
-#29570 := [symm #29568]: #29569
-#29565 := (= #27614 #3082)
-#29563 := (= #27613 uf_7)
-#29561 := (= #27613 #24969)
-#29562 := [monotonicity #28337]: #29561
-#29564 := [trans #29562 #28487]: #29563
-#29566 := [monotonicity #28337 #29564]: #29565
-#29572 := [trans #29566 #29570]: #29571
-#29575 := [monotonicity #29572 #29556]: #29574
-#29577 := [symm #29575]: #29576
-#28907 := (= uf_297 #28906)
-#28910 := (or #27121 #28907)
-#28911 := [quant-inst]: #28910
-#29573 := [unit-resolution #28911 #17002]: #28907
-#29578 := [trans #29573 #29577]: #27616
-#28124 := (not #27616)
-#29579 := (or #28124 #27786)
-#29580 := [th-lemma]: #29579
-#29581 := [unit-resolution #29580 #29578]: #27786
-#29582 := (not #27786)
-#29583 := (or #27696 #22372 #29582)
-#29584 := [th-lemma]: #29583
-#29585 := [unit-resolution #29584 #29581 #29554]: #27696
-#27697 := (not #27696)
-#28149 := (or #23372 #27697 #13899 #22515 #22510 #23225)
-#27994 := (uf_66 #2979 #27615 uf_7)
-#27995 := (uf_110 uf_287 #27994)
-#27998 := (= uf_302 #27995)
-#28104 := (= #3083 #27995)
-#28117 := (= #27995 #3083)
-#28111 := (= #27994 #3082)
-#28109 := (= #27615 uf_297)
-#27619 := (or #27121 #27616)
-#27620 := [quant-inst]: #27619
-#28108 := [unit-resolution #27620 #17002]: #27616
-#28110 := [symm #28108]: #28109
-#28112 := [monotonicity #28110]: #28111
-#28118 := [monotonicity #28112]: #28117
-#28119 := [symm #28118]: #28104
-#28120 := (= uf_302 #3083)
-#28113 := [hypothesis]: #12016
-#28114 := [hypothesis]: #23375
-#24539 := (or #23372 #12044)
-#24540 := [def-axiom]: #24539
-#28115 := [unit-resolution #24540 #28114]: #12044
-#28116 := [symm #28115]: #3097
-#28121 := [trans #28116 #28113]: #28120
-#28122 := [trans #28121 #28119]: #27998
-#27979 := (<= #27615 4294967295::int)
-#28123 := [hypothesis]: #15788
-#27787 := (>= #27785 0::int)
-#28125 := (or #28124 #27787)
-#28126 := [th-lemma]: #28125
-#28127 := [unit-resolution #28126 #28108]: #27787
-#28128 := (not #27787)
-#28129 := (or #27979 #22515 #28128)
-#28130 := [th-lemma]: #28129
-#28131 := [unit-resolution #28130 #28127 #28123]: #27979
-#28021 := (+ uf_286 #27712)
-#28022 := (<= #28021 0::int)
-#28133 := (not #28022)
-#28132 := [hypothesis]: #13900
-#28134 := (or #28133 #13899 #28128)
-#28135 := [th-lemma]: #28134
-#28136 := [unit-resolution #28135 #28127 #28132]: #28133
-#28001 := (not #27998)
-#27980 := (not #27979)
-#28146 := (or #27980 #28001 #28022)
-#28137 := [hypothesis]: #27696
-#24547 := (or #23372 #23366)
-#24548 := [def-axiom]: #24547
-#28138 := [unit-resolution #24548 #28114]: #23366
-#27791 := (+ uf_296 #13761)
-#27794 := (<= #27791 0::int)
-#28139 := (or #12093 #27794)
-#28140 := [th-lemma]: #28139
-#28141 := [unit-resolution #28140 #28115]: #27794
-#28045 := [hypothesis]: #23220
-#24545 := (or #23372 #13721)
-#24546 := [def-axiom]: #24545
-#28142 := [unit-resolution #24546 #28114]: #13721
-#28036 := (not #27794)
-#28051 := (or #22481 #13722 #23225 #28036)
-#28039 := [hypothesis]: #13721
-#27862 := (+ uf_298 #18969)
-#27863 := (<= #27862 0::int)
-#27874 := (+ uf_296 #18982)
-#27875 := (>= #27874 0::int)
-#28035 := (not #27875)
-#28029 := [hypothesis]: #27794
-#24522 := (not #18984)
-#28040 := [hypothesis]: #22486
-#24523 := (or #22481 #24522)
-#24524 := [def-axiom]: #24523
-#28041 := [unit-resolution #24524 #28040]: #24522
-#28037 := (or #28035 #18984 #28036)
-#28030 := [hypothesis]: #24522
-#28033 := [hypothesis]: #27875
-#28034 := [th-lemma #28033 #28030 #28029]: false
-#28038 := [lemma #28034]: #28037
-#28042 := [unit-resolution #28038 #28041 #28029]: #28035
-#28046 := (or #27863 #27875)
-#24517 := (or #22481 #18633)
-#24518 := [def-axiom]: #24517
-#28043 := [unit-resolution #24518 #28040]: #18633
-#24515 := (or #22481 #18632)
-#24516 := [def-axiom]: #24515
-#28044 := [unit-resolution #24516 #28040]: #18632
-#27888 := (or #23225 #22465 #22466 #27863 #27875)
-#27851 := (+ #18637 #13926)
-#27852 := (<= #27851 0::int)
-#27853 := (+ ?x776!15 #13457)
-#27854 := (>= #27853 0::int)
-#27855 := (or #22466 #27854 #27852 #22465)
-#27889 := (or #23225 #27855)
-#27896 := (iff #27889 #27888)
-#27883 := (or #22465 #22466 #27863 #27875)
-#27891 := (or #23225 #27883)
-#27894 := (iff #27891 #27888)
-#27895 := [rewrite]: #27894
-#27892 := (iff #27889 #27891)
-#27886 := (iff #27855 #27883)
-#27880 := (or #22466 #27863 #27875 #22465)
-#27884 := (iff #27880 #27883)
-#27885 := [rewrite]: #27884
-#27881 := (iff #27855 #27880)
-#27878 := (iff #27852 #27875)
-#27868 := (+ #13926 #18637)
-#27871 := (<= #27868 0::int)
-#27876 := (iff #27871 #27875)
-#27877 := [rewrite]: #27876
-#27872 := (iff #27852 #27871)
-#27869 := (= #27851 #27868)
-#27870 := [rewrite]: #27869
-#27873 := [monotonicity #27870]: #27872
-#27879 := [trans #27873 #27877]: #27878
-#27866 := (iff #27854 #27863)
-#27856 := (+ #13457 ?x776!15)
-#27859 := (>= #27856 0::int)
-#27864 := (iff #27859 #27863)
-#27865 := [rewrite]: #27864
-#27860 := (iff #27854 #27859)
-#27857 := (= #27853 #27856)
-#27858 := [rewrite]: #27857
-#27861 := [monotonicity #27858]: #27860
-#27867 := [trans #27861 #27865]: #27866
-#27882 := [monotonicity #27867 #27879]: #27881
-#27887 := [trans #27882 #27885]: #27886
-#27893 := [monotonicity #27887]: #27892
-#27897 := [trans #27893 #27895]: #27896
-#27890 := [quant-inst]: #27889
-#27898 := [mp #27890 #27897]: #27888
-#28047 := [unit-resolution #27898 #28045 #28044 #28043]: #28046
-#28048 := [unit-resolution #28047 #28042]: #27863
-#24519 := (not #18971)
-#24520 := (or #22481 #24519)
-#24521 := [def-axiom]: #24520
-#28049 := [unit-resolution #24521 #28040]: #24519
-#28050 := [th-lemma #28049 #28048 #28039]: false
-#28052 := [lemma #28050]: #28051
-#28143 := [unit-resolution #28052 #28142 #28045 #28141]: #22481
-#24531 := (or #23369 #23363 #22486)
-#24532 := [def-axiom]: #24531
-#28144 := [unit-resolution #24532 #28143 #28138]: #23363
-#24511 := (or #23360 #23352)
-#24512 := [def-axiom]: #24511
-#28145 := [unit-resolution #24512 #28144]: #23352
-#28058 := (or #23357 #27697 #27980 #28001 #28022)
-#27985 := (+ #27615 #13362)
-#27986 := (>= #27985 0::int)
-#27993 := (= #27995 uf_302)
-#27996 := (not #27993)
-#27997 := (or #27996 #27697 #27986 #27980)
-#28059 := (or #23357 #27997)
-#28066 := (iff #28059 #28058)
-#28053 := (or #27697 #27980 #28001 #28022)
-#28061 := (or #23357 #28053)
-#28064 := (iff #28061 #28058)
-#28065 := [rewrite]: #28064
-#28062 := (iff #28059 #28061)
-#28056 := (iff #27997 #28053)
-#28002 := (or #28001 #27697 #28022 #27980)
-#28054 := (iff #28002 #28053)
-#28055 := [rewrite]: #28054
-#28003 := (iff #27997 #28002)
-#28032 := (iff #27986 #28022)
-#28016 := (+ #13362 #27615)
-#28013 := (>= #28016 0::int)
-#28023 := (iff #28013 #28022)
-#28031 := [rewrite]: #28023
-#28019 := (iff #27986 #28013)
-#28017 := (= #27985 #28016)
-#28018 := [rewrite]: #28017
-#28020 := [monotonicity #28018]: #28019
-#27934 := [trans #28020 #28031]: #28032
-#28014 := (iff #27996 #28001)
-#27999 := (iff #27993 #27998)
-#28000 := [rewrite]: #27999
-#28015 := [monotonicity #28000]: #28014
-#28026 := [monotonicity #28015 #27934]: #28003
-#28057 := [trans #28026 #28055]: #28056
-#28063 := [monotonicity #28057]: #28062
-#28067 := [trans #28063 #28065]: #28066
-#28060 := [quant-inst]: #28059
-#28068 := [mp #28060 #28067]: #28058
-#28147 := [unit-resolution #28068 #28145 #28137]: #28146
-#28148 := [unit-resolution #28147 #28136 #28131 #28122]: false
-#28150 := [lemma #28148]: #28149
-#29586 := [unit-resolution #28150 #29585 #29553 #29552 #29551 #29550]: #23372
-#24555 := (or #23381 #23341 #23375)
-#24556 := [def-axiom]: #24555
-#29587 := [unit-resolution #24556 #29586 #29549]: #23341
-#24503 := (or #23338 #13722)
-#24504 := [def-axiom]: #24503
-#29588 := [unit-resolution #24504 #29587]: #13722
-#30529 := (not #29295)
-#30533 := (or #30532 #28431 #30529 #13721)
-#30534 := [th-lemma]: #30533
-#30535 := [unit-resolution #30534 #28425 #29588 #30527]: #30532
-#29174 := (>= #29134 0::int)
-#29175 := (not #29174)
-#29489 := [hypothesis]: #29175
-#24587 := (or #23384 #13954)
-#24588 := [def-axiom]: #24587
-#29490 := [unit-resolution #24588 #29548]: #13954
-#29294 := (<= #29293 0::int)
-#29493 := (or #29492 #29294)
-#29478 := [th-lemma]: #29493
-#29479 := [unit-resolution #29478 #29491]: #29294
-#29480 := [th-lemma #29479 #29490 #29489]: false
-#29481 := [lemma #29480]: #29174
-#30548 := (or #29175 #29185 #29193)
-#29204 := (or #27680 #27154 #29175 #29185 #29193)
-#29170 := (or #29169 #29167 #29164 #29161)
-#29171 := (not #29170)
-#29172 := (+ #29134 #27155)
-#29173 := (>= #29172 0::int)
-#29176 := (or #29175 #29173 #27154 #29171)
-#29205 := (or #27680 #29176)
-#29212 := (iff #29205 #29204)
-#29199 := (or #27154 #29175 #29185 #29193)
-#29207 := (or #27680 #29199)
-#29210 := (iff #29207 #29204)
-#29211 := [rewrite]: #29210
-#29208 := (iff #29205 #29207)
-#29202 := (iff #29176 #29199)
-#29196 := (or #29175 #29185 #27154 #29193)
-#29200 := (iff #29196 #29199)
-#29201 := [rewrite]: #29200
-#29197 := (iff #29176 #29196)
-#29194 := (iff #29171 #29193)
-#29191 := (iff #29170 #29190)
-#29192 := [rewrite]: #29191
-#29195 := [monotonicity #29192]: #29194
-#29188 := (iff #29173 #29185)
-#29177 := (+ #27155 #29134)
-#29180 := (>= #29177 0::int)
-#29186 := (iff #29180 #29185)
-#29187 := [rewrite]: #29186
-#29181 := (iff #29173 #29180)
-#29178 := (= #29172 #29177)
-#29179 := [rewrite]: #29178
-#29182 := [monotonicity #29179]: #29181
-#29189 := [trans #29182 #29187]: #29188
-#29198 := [monotonicity #29189 #29195]: #29197
-#29203 := [trans #29198 #29201]: #29202
-#29209 := [monotonicity #29203]: #29208
-#29213 := [trans #29209 #29211]: #29212
-#29206 := [quant-inst]: #29205
-#29214 := [mp #29206 #29213]: #29204
-#30553 := [unit-resolution #29214 #20917 #28509]: #30548
-#30542 := [unit-resolution #30553 #29481 #30535]: #29193
-#29220 := (or #29190 #29166)
-#29221 := [def-axiom]: #29220
-#30543 := [unit-resolution #29221 #30542]: #29166
-#30596 := [symm #30543]: #30601
-#30599 := (= #29300 #29165)
-#30564 := (= #27840 #29162)
-#30562 := (= #29162 #27840)
-#30560 := (= #29158 #3188)
-#29058 := (uf_116 #3188)
-#29062 := (uf_43 #24854 #29058)
-#30573 := (= #29062 #3188)
-#29063 := (= #3188 #29062)
-#28094 := (uf_48 #3188 #24854)
-#28095 := (= uf_9 #28094)
-#30567 := (= #3189 #28094)
-#28383 := (= #28094 #3189)
-#28384 := [monotonicity #28401]: #28383
-#30546 := [symm #28384]: #30567
-#28097 := (= #24854 #28096)
-#28198 := (* uf_298 #27060)
-#27568 := (uf_116 #27078)
-#28199 := (+ #27568 #28198)
-#28200 := (uf_43 #24854 #28199)
-#28342 := (uf_13 #28200)
-#28479 := (= #28342 #28096)
-#28475 := (= #28096 #28342)
-#28473 := (= #3188 #28200)
-#28194 := (uf_66 #27078 uf_298 #24854)
-#28201 := (= #28194 #28200)
-#28202 := (not #28201)
-#28195 := (uf_138 #28194 #27078)
-#28196 := (= uf_9 #28195)
-#28197 := (not #28196)
-#28205 := (or #28197 #28202)
-#28208 := (not #28205)
-#28213 := (or #27098 #28208)
-#28203 := (or #28202 #28197)
-#28204 := (not #28203)
-#28211 := (or #27098 #28204)
-#28215 := (iff #28211 #28213)
-#28217 := (iff #28213 #28213)
-#28218 := [rewrite]: #28217
-#28209 := (iff #28204 #28208)
-#28206 := (iff #28203 #28205)
-#28207 := [rewrite]: #28206
-#28210 := [monotonicity #28207]: #28209
-#28216 := [monotonicity #28210]: #28215
-#28219 := [trans #28216 #28218]: #28215
-#28214 := [quant-inst]: #28211
-#28289 := [mp #28214 #28219]: #28213
-#28465 := [unit-resolution #28289 #21133]: #28208
-#28292 := (or #28205 #28201)
-#28293 := [def-axiom]: #28292
-#28466 := [unit-resolution #28293 #28465]: #28201
-#28471 := (= #3188 #28194)
-#28469 := (= #28194 #3188)
-#28459 := (= uf_288 #27062)
-#28457 := (= #2980 #27062)
-#28458 := [symm #28456]: #28457
-#28460 := [trans #28404 #28458]: #28459
-#28462 := [monotonicity #28399 #28460]: #28461
-#28464 := [symm #28462]: #28463
-#28470 := [monotonicity #28464 #28401]: #28469
-#28472 := [symm #28470]: #28471
-#28474 := [trans #28472 #28466]: #28473
-#28476 := [monotonicity #28474]: #28475
-#28480 := [symm #28476]: #28479
-#28343 := (= #24854 #28342)
-#28282 := (or #24921 #28343)
-#28283 := [quant-inst]: #28282
-#28451 := [unit-resolution #28283 #23160]: #28343
-#28481 := [trans #28451 #28480]: #28097
-#27976 := (not #28097)
-#28093 := (iff #28095 #28097)
-#28101 := (or #26172 #28093)
-#28098 := (iff #28097 #28095)
-#28102 := (or #26172 #28098)
-#27972 := (iff #28102 #28101)
-#27974 := (iff #28101 #28101)
-#27936 := [rewrite]: #27974
-#28099 := (iff #28098 #28093)
-#28100 := [rewrite]: #28099
-#27973 := [monotonicity #28100]: #27972
-#27937 := [trans #27973 #27936]: #27972
-#27971 := [quant-inst]: #28102
-#27975 := [mp #27971 #27937]: #28101
-#28381 := [unit-resolution #27975 #23147]: #28093
-#28156 := (not #28095)
-#28364 := (iff #18449 #28156)
-#28229 := (iff #12369 #28095)
-#28392 := (iff #28095 #12369)
-#28393 := [monotonicity #28384]: #28392
-#28363 := [symm #28393]: #28229
-#28438 := [monotonicity #28363]: #28364
-#28382 := [hypothesis]: #18449
-#28449 := [mp #28382 #28438]: #28156
-#27970 := (not #28093)
-#27977 := (or #27970 #28095 #27976)
-#27978 := [def-axiom]: #27977
-#28450 := [unit-resolution #27978 #28449 #28381]: #27976
-#28482 := [unit-resolution #28450 #28481]: false
-#28483 := [lemma #28482]: #12369
-#30547 := [trans #28483 #30546]: #28095
-#29071 := (or #28156 #29063)
-#29074 := (or #26156 #28156 #29063)
-#29070 := (or #29063 #28156)
-#29075 := (or #26156 #29070)
-#29082 := (iff #29075 #29074)
-#29077 := (or #26156 #29071)
-#29080 := (iff #29077 #29074)
-#29081 := [rewrite]: #29080
-#29078 := (iff #29075 #29077)
-#29072 := (iff #29070 #29071)
-#29073 := [rewrite]: #29072
-#29079 := [monotonicity #29073]: #29078
-#29083 := [trans #29079 #29081]: #29082
-#29076 := [quant-inst]: #29075
-#29084 := [mp #29076 #29083]: #29074
-#30575 := [unit-resolution #29084 #18177]: #29071
-#30558 := [unit-resolution #30575 #30547]: #29063
-#30574 := [symm #30558]: #30573
-#30557 := (= #29158 #29062)
-#29400 := (* #27060 #29134)
-#29404 := (+ #27319 #29400)
-#29406 := (uf_43 #24854 #29404)
-#30571 := (= #29406 #29062)
-#30551 := (= #29404 #29058)
-#30515 := (= #29058 #29404)
-#30516 := (* -1::int #29404)
-#30517 := (+ #29058 #30516)
-#30518 := (<= #30517 0::int)
-#28931 := (* -1::int #27062)
-#28932 := (+ #24756 #28931)
-#28934 := (>= #28932 0::int)
-#28930 := (= #24756 #27062)
-#30513 := [trans #28521 #28458]: #28930
-#30589 := (not #28930)
-#30566 := (or #30589 #28934)
-#30536 := [th-lemma]: #30566
-#30595 := [unit-resolution #30536 #30513]: #28934
-#29598 := (* -1::int #27319)
-#29599 := (+ #24756 #29598)
-#29600 := (<= #29599 0::int)
-#29597 := (= #24756 #27319)
-#30597 := [symm #28519]: #29597
-#30598 := (not #29597)
-#30616 := (or #30598 #29600)
-#30617 := [th-lemma]: #30616
-#30612 := [unit-resolution #30617 #30597]: #29600
-#29106 := (+ #27062 #28198)
-#29107 := (uf_43 #24854 #29106)
-#29272 := (uf_116 #29107)
-#29276 := (* -1::int #29272)
-#29297 := (+ #29058 #29276)
-#29298 := (<= #29297 0::int)
-#29296 := (= #29058 #29272)
-#30658 := (= #29272 #29058)
-#30622 := (= #29107 #3188)
-#29102 := (uf_66 #24681 uf_298 #24854)
-#30620 := (= #29102 #3188)
-#30621 := [monotonicity #28454 #28401]: #30620
-#30615 := (= #29107 #29102)
-#29108 := (= #29102 #29107)
-#29109 := (not #29108)
-#29103 := (uf_138 #29102 #24681)
-#29104 := (= uf_9 #29103)
-#29105 := (not #29104)
-#29112 := (or #29105 #29109)
-#29115 := (not #29112)
-#29118 := (or #27098 #29115)
-#29110 := (or #29109 #29105)
-#29111 := (not #29110)
-#29119 := (or #27098 #29111)
-#29121 := (iff #29119 #29118)
-#29123 := (iff #29118 #29118)
-#29124 := [rewrite]: #29123
-#29116 := (iff #29111 #29115)
-#29113 := (iff #29110 #29112)
-#29114 := [rewrite]: #29113
-#29117 := [monotonicity #29114]: #29116
-#29122 := [monotonicity #29117]: #29121
-#29125 := [trans #29122 #29124]: #29121
-#29120 := [quant-inst]: #29119
-#29126 := [mp #29120 #29125]: #29118
-#30613 := [unit-resolution #29126 #21133]: #29115
-#29129 := (or #29112 #29108)
-#29130 := [def-axiom]: #29129
-#30618 := [unit-resolution #29130 #30613]: #29108
-#30619 := [symm #30618]: #30615
-#30632 := [trans #30619 #30621]: #30622
-#30659 := [monotonicity #30632]: #30658
-#30660 := [symm #30659]: #29296
-#30661 := (not #29296)
-#30656 := (or #30661 #29298)
-#30662 := [th-lemma]: #30656
-#30628 := [unit-resolution #30662 #30660]: #29298
-#29277 := (+ #28198 #29276)
-#29278 := (+ #27062 #29277)
-#29292 := (>= #29278 0::int)
-#29279 := (= #29278 0::int)
-#29282 := (or #24927 #29279)
-#29273 := (= #29106 #29272)
-#29283 := (or #24927 #29273)
-#29285 := (iff #29283 #29282)
-#29287 := (iff #29282 #29282)
-#29288 := [rewrite]: #29287
-#29280 := (iff #29273 #29279)
-#29281 := [rewrite]: #29280
-#29286 := [monotonicity #29281]: #29285
-#29289 := [trans #29286 #29288]: #29285
-#29284 := [quant-inst]: #29283
-#29290 := [mp #29284 #29289]: #29282
-#30663 := [unit-resolution #29290 #23154]: #29279
-#30664 := (not #29279)
-#30657 := (or #30664 #29292)
-#30665 := [th-lemma]: #30657
-#30666 := [unit-resolution #30665 #30663]: #29292
-#28937 := (>= #27060 1::int)
-#28935 := (= #27060 1::int)
+#7980 := (iff #1493 #7977)
+#7934 := (and #1485 #7931)
+#7937 := (and #7928 #7934)
+#7940 := (and #7922 #7937)
+#7588 := (not #1372)
+#7946 := (or #7588 #7940)
+#7955 := (or #7954 #7946)
+#7960 := (forall (vars (?x346 T4) (?x347 int) (?x348 T3) (?x349 int) (?x350 int)) (:pat #1472 #1475) #7955)
+#7978 := (iff #7960 #7977)
+#7975 := (iff #7955 #7972)
+#7966 := (or #7624 #7963)
+#7969 := (or #7954 #7966)
+#7973 := (iff #7969 #7972)
+#7974 := [rewrite]: #7973
+#7970 := (iff #7955 #7969)
+#7967 := (iff #7946 #7966)
+#7964 := (iff #7940 #7963)
+#7965 := [rewrite]: #7964
+#7625 := (iff #7588 #7624)
+#7622 := (iff #1372 #7621)
+#7619 := (iff #1371 #7618)
+#7620 := [rewrite]: #7619
+#7623 := [monotonicity #4040 #7620]: #7622
+#7626 := [monotonicity #7623]: #7625
+#7968 := [monotonicity #7626 #7965]: #7967
+#7971 := [monotonicity #7968]: #7970
+#7976 := [trans #7971 #7974]: #7975
+#7979 := [quant-intro #7976]: #7978
+#7961 := (iff #1493 #7960)
+#7958 := (iff #1492 #7955)
+#7951 := (implies #7919 #7946)
+#7956 := (iff #7951 #7955)
+#7957 := [rewrite]: #7956
+#7952 := (iff #1492 #7951)
+#7949 := (iff #1491 #7946)
+#7943 := (implies #1372 #7940)
+#7947 := (iff #7943 #7946)
+#7948 := [rewrite]: #7947
+#7944 := (iff #1491 #7943)
+#7941 := (iff #1490 #7940)
+#7938 := (iff #1489 #7937)
+#7935 := (iff #1488 #7934)
+#7932 := (iff #1487 #7931)
+#7933 := [rewrite]: #7932
+#7936 := [monotonicity #7933]: #7935
+#7929 := (iff #1484 #7928)
+#7926 := (iff #1483 #7925)
+#7927 := [rewrite]: #7926
+#7930 := [monotonicity #7927]: #7929
+#7939 := [monotonicity #7930 #7936]: #7938
+#7923 := (iff #1481 #7922)
+#7924 := [rewrite]: #7923
+#7942 := [monotonicity #7924 #7939]: #7941
+#7945 := [monotonicity #7942]: #7944
+#7950 := [trans #7945 #7948]: #7949
+#7920 := (iff #1477 #7919)
+#7921 := [rewrite]: #7920
+#7953 := [monotonicity #7921 #7950]: #7952
+#7959 := [trans #7953 #7957]: #7958
+#7962 := [quant-intro #7959]: #7961
+#7981 := [trans #7962 #7979]: #7980
+#7918 := [asserted]: #1493
+#7982 := [mp #7918 #7981]: #7977
+#16325 := [mp~ #7982 #16324]: #7977
+#20106 := [mp #16325 #20105]: #20103
+#26542 := (not #20103)
+#26543 := (or #26542 #26409 #26427 #26435)
+#26405 := (or #26404 #26402 #26399 #26396)
+#26406 := (not #26405)
+#26415 := (or #26414 #26412 #26409 #26406)
+#26544 := (or #26542 #26415)
+#26551 := (iff #26544 #26543)
+#26441 := (or #26409 #26427 #26435)
+#26546 := (or #26542 #26441)
+#26549 := (iff #26546 #26543)
+#26550 := [rewrite]: #26549
+#26547 := (iff #26544 #26546)
+#26444 := (iff #26415 #26441)
+#26438 := (or false #26427 #26409 #26435)
+#26442 := (iff #26438 #26441)
+#26443 := [rewrite]: #26442
+#26439 := (iff #26415 #26438)
+#26436 := (iff #26406 #26435)
+#26433 := (iff #26405 #26432)
+#26434 := [rewrite]: #26433
+#26437 := [monotonicity #26434]: #26436
+#26440 := [monotonicity #26421 #26431 #26437]: #26439
+#26445 := [trans #26440 #26443]: #26444
+#26548 := [monotonicity #26445]: #26547
+#26532 := [trans #26548 #26550]: #26551
+#26545 := [quant-inst]: #26544
+#26533 := [mp #26545 #26532]: #26543
+#27871 := [unit-resolution #26533 #20106 #26457]: #27870
+#27872 := [unit-resolution #27871 #27869]: #26435
+#26538 := (or #26432 #26401)
+#26539 := [def-axiom]: #26538
+#27873 := [unit-resolution #26539 #27872]: #26401
+#27913 := [symm #27873]: #27912
+#27910 := (= #26748 #26400)
+#27908 := (= #26742 #26397)
+#27906 := (= #26222 #26397)
+#27904 := (= #26397 #26222)
+#27902 := (= #26393 #3044)
+#27900 := (= #26393 #26311)
+#27898 := (= #26393 #26333)
+#27896 := (= #26393 #2979)
+#26574 := (uf_116 #25399)
+#26588 := (uf_43 #24108 #26574)
+#27892 := (= #26588 #2979)
+#27884 := (= #26574 uf_288)
+#27882 := (= #26574 #2980)
+#27880 := (= #24010 #2980)
+#24234 := (= #2980 #24010)
+#24239 := (or #24181 #24234)
+#24240 := [quant-inst]: #24239
+#27874 := [unit-resolution #24240 #22411]: #24234
+#27881 := [symm #27874]: #27880
+#27878 := (= #26574 #24010)
+#27879 := [monotonicity #27877]: #27878
+#27883 := [trans #27879 #27881]: #27882
+#27885 := [trans #27883 #27661]: #27884
+#27893 := [monotonicity #27656 #27885]: #27892
+#27894 := (= #26393 #26588)
+#26570 := (uf_66 #25399 0::int #24108)
+#26591 := (= #26570 #26588)
+#26594 := (not #26591)
+#26571 := (uf_138 #26570 #25399)
+#26572 := (= uf_9 #26571)
+#26573 := (not #26572)
+#26600 := (or #26573 #26594)
+#26605 := (not #26600)
+#26652 := (or #26353 #26605)
+#26575 := (+ #26574 #26316)
+#26576 := (uf_43 #24108 #26575)
+#26577 := (= #26570 #26576)
+#26578 := (not #26577)
+#26579 := (or #26578 #26573)
+#26580 := (not #26579)
+#26653 := (or #26353 #26580)
+#26655 := (iff #26653 #26652)
+#26667 := (iff #26652 #26652)
+#26668 := [rewrite]: #26667
+#26606 := (iff #26580 #26605)
+#26603 := (iff #26579 #26600)
+#26597 := (or #26594 #26573)
+#26601 := (iff #26597 #26600)
+#26602 := [rewrite]: #26601
+#26598 := (iff #26579 #26597)
+#26595 := (iff #26578 #26594)
+#26592 := (iff #26577 #26591)
+#26589 := (= #26576 #26588)
+#26586 := (= #26575 #26574)
+#26581 := (+ #26574 0::int)
+#26584 := (= #26581 #26574)
+#26585 := [rewrite]: #26584
+#26582 := (= #26575 #26581)
+#26583 := [monotonicity #26325]: #26582
+#26587 := [trans #26583 #26585]: #26586
+#26590 := [monotonicity #26587]: #26589
+#26593 := [monotonicity #26590]: #26592
+#26596 := [monotonicity #26593]: #26595
+#26599 := [monotonicity #26596]: #26598
+#26604 := [trans #26599 #26602]: #26603
+#26607 := [monotonicity #26604]: #26606
+#26656 := [monotonicity #26607]: #26655
+#26669 := [trans #26656 #26668]: #26655
+#26654 := [quant-inst]: #26653
+#26688 := [mp #26654 #26669]: #26652
+#27886 := [unit-resolution #26688 #20322]: #26605
+#26692 := (or #26600 #26591)
+#26693 := [def-axiom]: #26692
+#27887 := [unit-resolution #26693 #27886]: #26591
+#27890 := (= #26393 #26570)
+#27891 := [monotonicity #27889]: #27890
+#27895 := [trans #27891 #27887]: #27894
+#27897 := [trans #27895 #27893]: #27896
+#27899 := [trans #27897 #27790]: #27898
+#27901 := [trans #27899 #27792]: #27900
+#27903 := [trans #27901 #27794]: #27902
+#27905 := [monotonicity #27903]: #27904
+#27907 := [symm #27905]: #27906
+#27909 := [trans #27822 #27907]: #27908
+#27911 := [monotonicity #27909]: #27910
+#27915 := [trans #27911 #27913]: #27914
+#27917 := [trans #27915 #27863]: #27916
+#27919 := [monotonicity #27917]: #27918
+#27923 := [trans #27919 #27921]: #27922
+#27925 := [monotonicity #27923]: #27924
+#27927 := [monotonicity #27925]: #27926
+#27929 := [symm #27927]: #27928
+#27931 := [monotonicity #27929]: #27930
+#27932 := [mp #13533 #27931]: #27751
+#27936 := (= #2984 #26751)
+#27933 := (= #26751 #2984)
+#27934 := [monotonicity #27917]: #27933
+#27937 := [symm #27934]: #27936
+#27938 := [trans #13530 #27937]: #26752
+#27744 := (not #26752)
+#27745 := (or #27416 #27744)
+#27746 := [def-axiom]: #27745
+#27939 := [unit-resolution #27746 #27938]: #27416
+#27758 := (or #27005 #26774 #26757 #27415 #27420)
+#27759 := [def-axiom]: #27758
+#27940 := [unit-resolution #27759 #27939 #27932 #27856 #27838]: #27005
+#27762 := (or #27012 #27536)
+#27763 := [def-axiom]: #27762
+#27941 := [unit-resolution #27763 #27940 #27819]: false
+#27943 := [lemma #27941]: #27942
+#28112 := [unit-resolution #27943 #28095]: #28103
+#28113 := [unit-resolution #28112 #28102]: #11979
+#23911 := (or #22686 #17631 #17640 #22680)
+#23912 := [def-axiom]: #23911
+#28111 := [unit-resolution #23912 #28113 #26648 #28087]: #22680
+#23901 := (or #22677 #22671)
+#23902 := [def-axiom]: #23901
+#28114 := [unit-resolution #23902 #28111]: #22671
+#28072 := [hypothesis]: #12866
+#28073 := [th-lemma #13542 #28072]: false
+#28074 := [lemma #28073]: #12863
+#23887 := (or #22674 #12866 #22668)
+#23888 := [def-axiom]: #23887
+#28115 := [unit-resolution #23888 #28074 #28114]: #22668
+#23879 := (or #22665 #22659)
+#23880 := [def-axiom]: #23879
+#28116 := [unit-resolution #23880 #28115]: #22659
+#23889 := (or #22677 #3055)
+#23890 := [def-axiom]: #23889
+#28117 := [unit-resolution #23890 #28111]: #3055
+#26943 := [hypothesis]: #22641
+#26945 := (or #22650 #17676 #22644)
+#23863 := (or #22650 #17676 #12740 #22644)
+#23864 := [def-axiom]: #23863
+#26946 := [unit-resolution #23864 #13542]: #26945
+#28110 := [unit-resolution #26946 #26943 #28117]: #22650
+#23867 := (or #22653 #22647)
+#23868 := [def-axiom]: #23867
+#28118 := [unit-resolution #23868 #28110]: #22653
+#27722 := (or #17676 #17653 #22662 #22644)
+#26876 := (* -1::int #3054)
+#26877 := (+ uf_295 #26876)
+#26878 := (>= #26877 0::int)
+#26623 := (not #26878)
+#27630 := [hypothesis]: #17654
+#23562 := (not #17662)
+#26942 := [hypothesis]: #22659
+#26944 := [hypothesis]: #3055
+#26947 := [unit-resolution #26946 #26944 #26943]: #22650
+#26948 := [unit-resolution #23868 #26947]: #22653
+#23875 := (or #22662 #21419 #22656)
+#23876 := [def-axiom]: #23875
+#27014 := [unit-resolution #23876 #26948 #26942]: #21419
+#23563 := (or #21414 #23562)
+#23564 := [def-axiom]: #23563
+#27015 := [unit-resolution #23564 #27014]: #23562
+#23560 := (or #21414 #17655)
+#23561 := [def-axiom]: #23560
+#27045 := [unit-resolution #23561 #27014]: #17655
+#26624 := (or #26623 #17662 #21399 #17653)
+#27616 := [hypothesis]: #26878
+#27617 := [hypothesis]: #23562
+#27037 := (+ #3054 #17660)
+#27039 := (>= #27037 0::int)
+#27036 := (= #3054 #17659)
+#27629 := (= #17659 #3054)
+#27633 := (= #17658 #3044)
+#27614 := (= ?x773!13 0::int)
+#27631 := [hypothesis]: #17655
+#27632 := [th-lemma #27631 #27630]: #27614
+#27634 := [monotonicity #27632]: #27633
+#27635 := [monotonicity #27634]: #27629
+#26386 := [symm #27635]: #27036
+#26387 := (not #27036)
+#26388 := (or #26387 #27039)
+#26389 := [th-lemma]: #26388
+#26390 := [unit-resolution #26389 #26386]: #27039
+#26521 := [th-lemma #26390 #27617 #27616]: false
+#26625 := [lemma #26521]: #26624
+#27100 := [unit-resolution #26625 #27045 #27015 #27630]: #26623
+#27551 := (or #17676 #26878)
+#27591 := [th-lemma]: #27551
+#27592 := [unit-resolution #27591 #26944 #27100]: false
+#27723 := [lemma #27592]: #27722
+#28119 := [unit-resolution #27723 #26943 #28117 #28116]: #17653
+#23558 := (or #21414 #17654)
+#23559 := [def-axiom]: #23558
+#28120 := [unit-resolution #23559 #28119]: #21414
+#28121 := [unit-resolution #23876 #28120 #28118 #28116]: false
+#28122 := [lemma #28121]: #22644
+#23841 := (or #22641 #13392)
+#23842 := [def-axiom]: #23841
+#29445 := [unit-resolution #23842 #28122]: #13392
+#27504 := (not #27503)
+#29446 := [hypothesis]: #27504
+#28547 := (+ uf_298 #28457)
+#28548 := (<= #28547 0::int)
+#28874 := (not #28214)
+#28877 := (or #28874 #28548)
+#28878 := [th-lemma]: #28877
+#28879 := [unit-resolution #28878 #28873]: #28548
+#29447 := [th-lemma #28879 #29446 #29445 #29444]: false
+#29448 := [lemma #29447]: #27503
+#27512 := (* -1::int #27482)
+#27513 := (+ #24110 #27512)
+#27514 := (<= #27513 0::int)
+#27685 := (not #27514)
+#27640 := (+ uf_298 #27512)
+#27642 := (>= #27640 0::int)
+#28549 := (>= #28547 0::int)
+#29966 := (or #28874 #28549)
+#29967 := [th-lemma]: #29966
+#29968 := [unit-resolution #29967 #28873]: #28549
+#29981 := (not #28549)
+#29823 := (or #27642 #29981)
+#28076 := (<= #28081 0::int)
+#29972 := (or #29442 #28076)
+#29973 := [th-lemma]: #29972
+#29974 := [unit-resolution #29973 #29441]: #28076
+#29980 := (not #28076)
+#29822 := (or #27642 #29980 #29981)
+#29806 := [th-lemma]: #29822
+#29824 := [unit-resolution #29806 #29974]: #29823
+#29825 := [unit-resolution #29824 #29968]: #27642
+#26869 := (uf_66 #26333 uf_297 #26868)
+#26870 := (uf_125 #26869 #26333)
+#26951 := (>= #26870 0::int)
+#26967 := (* -1::int #26870)
+#27040 := (+ uf_297 #26967)
+#27041 := (<= #27040 0::int)
+#28165 := (uf_66 #23935 uf_297 #26369)
+#28166 := (uf_125 #28165 #23935)
+#28318 := (* -1::int #28166)
+#28542 := (+ uf_297 #28318)
+#28543 := (<= #28542 0::int)
+#28167 := (= uf_297 #28166)
+#28235 := (or #26376 #28167)
+#28236 := [quant-inst]: #28235
+#29826 := [unit-resolution #28236 #16290]: #28167
+#29683 := (not #28167)
+#29774 := (or #29683 #28543)
+#29726 := [th-lemma]: #29774
+#29779 := [unit-resolution #29726 #29826]: #28543
+#28480 := (not #28543)
+#28695 := (or #27041 #28480)
+#28077 := (+ #26870 #28318)
+#28079 := (>= #28077 0::int)
+#28075 := (= #26870 #28166)
+#29841 := (= #28166 #26870)
+#29845 := (= #28165 #26869)
+#29730 := (= #3100 #26869)
+#29728 := (= #26869 #3100)
+#29729 := [monotonicity #27789 #29424]: #29728
+#29731 := [symm #29729]: #29730
+#29679 := (= #28165 #3100)
+#29727 := [monotonicity #27709 #27849]: #29679
+#28496 := [trans #29727 #29731]: #29845
+#28499 := [monotonicity #28496 #27796]: #29841
+#28479 := [symm #28499]: #28075
+#28498 := (not #28075)
+#29842 := (or #28498 #28079)
+#29843 := [th-lemma]: #29842
+#29844 := [unit-resolution #29843 #28479]: #28079
+#29828 := (not #28079)
+#28495 := (or #27041 #29828 #28480)
+#28481 := [th-lemma]: #28495
+#29827 := [unit-resolution #28481 #29844]: #28695
+#28696 := [unit-resolution #29827 #29779]: #27041
+#28864 := (not #27041)
+#28989 := (or #26951 #28864)
+#23833 := (or #22641 #12910)
+#23834 := [def-axiom]: #23833
+#28862 := [unit-resolution #23834 #28122]: #12910
+#28497 := (or #26951 #21602 #28864)
+#28963 := [th-lemma]: #28497
+#28984 := [unit-resolution #28963 #28862]: #28989
+#28691 := [unit-resolution #28984 #28696]: #26951
+#26952 := (not #26951)
+#28840 := (or #22629 #26952)
+#23851 := (or #22641 #22477)
+#23852 := [def-axiom]: #23851
+#28540 := [unit-resolution #23852 #28122]: #22477
+#23825 := (or #22641 #12030)
+#23826 := [def-axiom]: #23825
+#28875 := [unit-resolution #23826 #28122]: #12030
+#23847 := (or #22641 #15096)
+#23848 := [def-axiom]: #23847
+#28849 := [unit-resolution #23848 #28122]: #15096
+#23837 := (or #22641 #13359)
+#23838 := [def-axiom]: #23837
+#28859 := [unit-resolution #23838 #28122]: #13359
+#27404 := (or #22629 #26952 #13358 #21769 #21761 #22482)
+#27249 := (uf_66 #2979 #26870 uf_7)
+#27250 := (uf_110 uf_287 #27249)
+#27253 := (= uf_302 #27250)
+#27359 := (= #3101 #27250)
+#27372 := (= #27250 #3101)
+#27366 := (= #27249 #3100)
+#27364 := (= #26870 uf_297)
+#26871 := (= uf_297 #26870)
+#26874 := (or #26376 #26871)
+#26875 := [quant-inst]: #26874
+#27363 := [unit-resolution #26875 #16290]: #26871
+#27365 := [symm #27363]: #27364
+#27367 := [monotonicity #27365]: #27366
+#27373 := [monotonicity #27367]: #27372
+#27374 := [symm #27373]: #27359
+#27375 := (= uf_302 #3101)
+#27368 := [hypothesis]: #12030
+#27369 := [hypothesis]: #22632
+#23793 := (or #22629 #12122)
+#23794 := [def-axiom]: #23793
+#27370 := [unit-resolution #23794 #27369]: #12122
+#27371 := [symm #27370]: #3129
+#27376 := [trans #27371 #27368]: #27375
+#27377 := [trans #27376 #27374]: #27253
+#27234 := (<= #26870 4294967295::int)
+#27378 := [hypothesis]: #15096
+#27042 := (>= #27040 0::int)
+#27379 := (not #26871)
+#27380 := (or #27379 #27042)
+#27381 := [th-lemma]: #27380
+#27382 := [unit-resolution #27381 #27363]: #27042
+#27383 := (not #27042)
+#27384 := (or #27234 #21769 #27383)
+#27385 := [th-lemma]: #27384
+#27386 := [unit-resolution #27385 #27382 #27378]: #27234
+#27276 := (+ uf_286 #26967)
+#27277 := (<= #27276 0::int)
+#27388 := (not #27277)
+#27387 := [hypothesis]: #13359
+#27389 := (or #27388 #13358 #27383)
+#27390 := [th-lemma]: #27389
+#27391 := [unit-resolution #27390 #27382 #27387]: #27388
+#27256 := (not #27253)
+#27235 := (not #27234)
+#27401 := (or #27235 #27256 #27277)
+#27392 := [hypothesis]: #26951
+#23801 := (or #22629 #22623)
+#23802 := [def-axiom]: #23801
+#27393 := [unit-resolution #23802 #27369]: #22623
+#27046 := (+ uf_296 #13286)
+#27049 := (<= #27046 0::int)
+#27394 := (or #21741 #27049)
+#27395 := [th-lemma]: #27394
+#27396 := [unit-resolution #27395 #27370]: #27049
+#27300 := [hypothesis]: #22477
+#23799 := (or #22629 #12967)
+#23800 := [def-axiom]: #23799
+#27397 := [unit-resolution #23800 #27369]: #12967
+#27291 := (not #27049)
+#27306 := (or #21724 #12968 #22482 #27291)
+#27294 := [hypothesis]: #12967
+#27117 := (+ uf_298 #18168)
+#27118 := (<= #27117 0::int)
+#27129 := (+ uf_296 #18155)
+#27130 := (>= #27129 0::int)
+#27290 := (not #27130)
+#27284 := [hypothesis]: #27049
+#23774 := (not #18157)
+#27295 := [hypothesis]: #21729
+#23775 := (or #21724 #23774)
+#23776 := [def-axiom]: #23775
+#27296 := [unit-resolution #23776 #27295]: #23774
+#27292 := (or #27290 #18157 #27291)
+#27285 := [hypothesis]: #23774
+#27288 := [hypothesis]: #27130
+#27289 := [th-lemma #27288 #27285 #27284]: false
+#27293 := [lemma #27289]: #27292
+#27297 := [unit-resolution #27293 #27296 #27284]: #27290
+#27301 := (or #27118 #27130)
+#23772 := (or #21724 #17870)
+#23773 := [def-axiom]: #23772
+#27298 := [unit-resolution #23773 #27295]: #17870
+#23770 := (or #21724 #17866)
+#23771 := [def-axiom]: #23770
+#27299 := [unit-resolution #23771 #27295]: #17866
+#27143 := (or #22482 #21708 #21709 #27118 #27130)
+#27106 := (+ #17874 #13378)
+#27107 := (<= #27106 0::int)
+#27108 := (+ ?x776!15 #12965)
+#27109 := (>= #27108 0::int)
+#27110 := (or #21709 #27109 #27107 #21708)
+#27144 := (or #22482 #27110)
+#27151 := (iff #27144 #27143)
+#27138 := (or #21708 #21709 #27118 #27130)
+#27146 := (or #22482 #27138)
+#27149 := (iff #27146 #27143)
+#27150 := [rewrite]: #27149
+#27147 := (iff #27144 #27146)
+#27141 := (iff #27110 #27138)
+#27135 := (or #21709 #27118 #27130 #21708)
+#27139 := (iff #27135 #27138)
+#27140 := [rewrite]: #27139
+#27136 := (iff #27110 #27135)
+#27133 := (iff #27107 #27130)
+#27123 := (+ #13378 #17874)
+#27126 := (<= #27123 0::int)
+#27131 := (iff #27126 #27130)
+#27132 := [rewrite]: #27131
+#27127 := (iff #27107 #27126)
+#27124 := (= #27106 #27123)
+#27125 := [rewrite]: #27124
+#27128 := [monotonicity #27125]: #27127
+#27134 := [trans #27128 #27132]: #27133
+#27121 := (iff #27109 #27118)
+#27111 := (+ #12965 ?x776!15)
+#27114 := (>= #27111 0::int)
+#27119 := (iff #27114 #27118)
+#27120 := [rewrite]: #27119
+#27115 := (iff #27109 #27114)
+#27112 := (= #27108 #27111)
+#27113 := [rewrite]: #27112
+#27116 := [monotonicity #27113]: #27115
+#27122 := [trans #27116 #27120]: #27121
+#27137 := [monotonicity #27122 #27134]: #27136
+#27142 := [trans #27137 #27140]: #27141
+#27148 := [monotonicity #27142]: #27147
+#27152 := [trans #27148 #27150]: #27151
+#27145 := [quant-inst]: #27144
+#27153 := [mp #27145 #27152]: #27143
+#27302 := [unit-resolution #27153 #27300 #27299 #27298]: #27301
+#27303 := [unit-resolution #27302 #27297]: #27118
+#23777 := (or #21724 #18175)
+#23778 := [def-axiom]: #23777
+#27304 := [unit-resolution #23778 #27295]: #18175
+#27305 := [th-lemma #27304 #27303 #27294]: false
+#27307 := [lemma #27305]: #27306
+#27398 := [unit-resolution #27307 #27397 #27300 #27396]: #21724
+#23785 := (or #22626 #22620 #21729)
+#23786 := [def-axiom]: #23785
+#27399 := [unit-resolution #23786 #27398 #27393]: #22620
+#23766 := (or #22617 #22609)
+#23767 := [def-axiom]: #23766
+#27400 := [unit-resolution #23767 #27399]: #22609
+#27313 := (or #22614 #26952 #27235 #27256 #27277)
+#27240 := (+ #26870 #12727)
+#27241 := (>= #27240 0::int)
+#27248 := (= #27250 uf_302)
+#27251 := (not #27248)
+#27252 := (or #27251 #26952 #27241 #27235)
+#27314 := (or #22614 #27252)
+#27321 := (iff #27314 #27313)
+#27308 := (or #26952 #27235 #27256 #27277)
+#27316 := (or #22614 #27308)
+#27319 := (iff #27316 #27313)
+#27320 := [rewrite]: #27319
+#27317 := (iff #27314 #27316)
+#27311 := (iff #27252 #27308)
+#27257 := (or #27256 #26952 #27277 #27235)
+#27309 := (iff #27257 #27308)
+#27310 := [rewrite]: #27309
+#27258 := (iff #27252 #27257)
+#27287 := (iff #27241 #27277)
+#27271 := (+ #12727 #26870)
+#27268 := (>= #27271 0::int)
+#27278 := (iff #27268 #27277)
+#27286 := [rewrite]: #27278
+#27274 := (iff #27241 #27268)
+#27272 := (= #27240 #27271)
+#27273 := [rewrite]: #27272
+#27275 := [monotonicity #27273]: #27274
+#27189 := [trans #27275 #27286]: #27287
+#27269 := (iff #27251 #27256)
+#27254 := (iff #27248 #27253)
+#27255 := [rewrite]: #27254
+#27270 := [monotonicity #27255]: #27269
+#27281 := [monotonicity #27270 #27189]: #27258
+#27312 := [trans #27281 #27310]: #27311
+#27318 := [monotonicity #27312]: #27317
+#27322 := [trans #27318 #27320]: #27321
+#27315 := [quant-inst]: #27314
+#27323 := [mp #27315 #27322]: #27313
+#27402 := [unit-resolution #27323 #27400 #27392]: #27401
+#27403 := [unit-resolution #27402 #27391 #27386 #27377]: false
+#27405 := [lemma #27403]: #27404
+#28978 := [unit-resolution #27405 #28859 #28849 #28875 #28540]: #28840
+#28985 := [unit-resolution #28978 #28691]: #22629
+#23853 := (or #22641 #22635)
+#23854 := [def-axiom]: #23853
+#29133 := [unit-resolution #23854 #28122]: #22635
+#23809 := (or #22638 #22598 #22632)
+#23810 := [def-axiom]: #23809
+#29134 := [unit-resolution #23810 #29133]: #22635
+#28988 := [unit-resolution #29134 #28985]: #22598
+#23758 := (or #22595 #12968)
+#23759 := [def-axiom]: #23758
+#28538 := [unit-resolution #23759 #28988]: #12968
+#27687 := (not #27642)
+#28488 := (or #27685 #27687 #12967)
+#27688 := (or #27685 #27686 #27687 #12967)
+#27689 := [th-lemma]: #27688
+#28539 := [unit-resolution #27689 #27680]: #28488
+#28525 := [unit-resolution #28539 #28538 #29825]: #27685
+#29715 := (or #26542 #26409 #27504 #27514 #27522)
+#27499 := (or #27498 #27496 #27493 #27490)
+#27500 := (not #27499)
+#27501 := (+ #27482 #26410)
+#27502 := (>= #27501 0::int)
+#27505 := (or #27504 #27502 #26409 #27500)
+#29716 := (or #26542 #27505)
+#29738 := (iff #29716 #29715)
+#27528 := (or #26409 #27504 #27514 #27522)
+#29718 := (or #26542 #27528)
+#29737 := (iff #29718 #29715)
+#29721 := [rewrite]: #29737
+#29735 := (iff #29716 #29718)
+#27531 := (iff #27505 #27528)
+#27525 := (or #27504 #27514 #26409 #27522)
+#27529 := (iff #27525 #27528)
+#27530 := [rewrite]: #27529
+#27526 := (iff #27505 #27525)
+#27523 := (iff #27500 #27522)
+#27520 := (iff #27499 #27519)
+#27521 := [rewrite]: #27520
+#27524 := [monotonicity #27521]: #27523
+#27517 := (iff #27502 #27514)
+#27506 := (+ #26410 #27482)
+#27509 := (>= #27506 0::int)
+#27515 := (iff #27509 #27514)
+#27516 := [rewrite]: #27515
+#27510 := (iff #27502 #27509)
+#27507 := (= #27501 #27506)
+#27508 := [rewrite]: #27507
+#27511 := [monotonicity #27508]: #27510
+#27518 := [trans #27511 #27516]: #27517
+#27527 := [monotonicity #27518 #27524]: #27526
+#27532 := [trans #27527 #27530]: #27531
+#29736 := [monotonicity #27532]: #29735
+#29739 := [trans #29736 #29721]: #29738
+#29717 := [quant-inst]: #29716
+#29740 := [mp #29717 #29739]: #29715
+#29511 := [unit-resolution #29740 #20106 #28525 #27869 #29448 #30005]: false
+#29510 := [lemma #29511]: #27522
+#28981 := (or #27519 #28652)
+#29570 := [hypothesis]: #28647
+#29836 := [hypothesis]: #27522
+#29772 := (or #27519 #27495)
+#29773 := [def-axiom]: #29772
+#29837 := [unit-resolution #29773 #29836]: #27495
+#27556 := (uf_58 #3175 #27555)
+#27557 := (uf_136 #27556)
+#27558 := (= uf_9 #27557)
+#29038 := (not #27558)
+#27559 := (uf_24 uf_287 #27555)
+#27560 := (= uf_9 #27559)
+#27561 := (not #27560)
+#27565 := (or #27558 #27561)
+#27568 := (not #27565)
+#27579 := (or #26500 #23948 #26481 #26484 #27504 #27514 #27568)
+#27562 := (or #27561 #27558)
+#27563 := (not #27562)
+#27564 := (or #26484 #27504 #27502 #23948 #26481 #27563)
+#27580 := (or #26500 #27564)
+#27587 := (iff #27580 #27579)
+#27574 := (or #23948 #26481 #26484 #27504 #27514 #27568)
+#27582 := (or #26500 #27574)
+#27585 := (iff #27582 #27579)
+#27586 := [rewrite]: #27585
+#27583 := (iff #27580 #27582)
+#27577 := (iff #27564 #27574)
+#27571 := (or #26484 #27504 #27514 #23948 #26481 #27568)
+#27575 := (iff #27571 #27574)
+#27576 := [rewrite]: #27575
+#27572 := (iff #27564 #27571)
+#27569 := (iff #27563 #27568)
+#27566 := (iff #27562 #27565)
+#27567 := [rewrite]: #27566
+#27570 := [monotonicity #27567]: #27569
+#27573 := [monotonicity #27518 #27570]: #27572
+#27578 := [trans #27573 #27576]: #27577
+#27584 := [monotonicity #27578]: #27583
+#27588 := [trans #27584 #27586]: #27587
+#27581 := [quant-inst]: #27580
+#27589 := [mp #27581 #27588]: #27579
+#29312 := [unit-resolution #27589 #20766 #13537 #28088 #28525 #27696 #29448]: #27568
+#29777 := (or #27565 #29038)
+#29778 := [def-axiom]: #29777
+#28973 := [unit-resolution #29778 #29312]: #29038
+#29704 := (or #28652 #27496 #27558)
+#29567 := (= #2984 #28581)
+#29558 := (= #28581 #2984)
+#29931 := (= #28578 #2981)
+#29929 := (= #28578 #26392)
+#29927 := (= #27494 #26392)
+#29874 := [hypothesis]: #27495
+#29928 := [symm #29874]: #29927
+#29925 := (= #28578 #27494)
+#29923 := (= #27095 #27491)
+#29921 := (= #27491 #27095)
+#29919 := (= #27485 #3220)
+#28238 := (uf_116 #3220)
+#28242 := (uf_43 #24108 #28238)
+#29913 := (= #28242 #3220)
+#28243 := (= #3220 #28242)
+#27349 := (uf_48 #3220 #24108)
+#27350 := (= uf_9 #27349)
+#29876 := (= #3221 #27349)
+#27638 := (= #27349 #3221)
+#27639 := [monotonicity #27656]: #27638
+#29877 := [symm #27639]: #29876
+#27352 := (= #24108 #27351)
+#27453 := (* uf_298 #26315)
+#26823 := (uf_116 #26333)
+#27454 := (+ #26823 #27453)
+#27455 := (uf_43 #24108 #27454)
+#27597 := (uf_13 #27455)
+#27734 := (= #27597 #27351)
+#27730 := (= #27351 #27597)
+#27728 := (= #3220 #27455)
+#27449 := (uf_66 #26333 uf_298 #24108)
+#27456 := (= #27449 #27455)
+#27457 := (not #27456)
+#27450 := (uf_138 #27449 #26333)
+#27451 := (= uf_9 #27450)
+#27452 := (not #27451)
+#27460 := (or #27452 #27457)
+#27463 := (not #27460)
+#27468 := (or #26353 #27463)
+#27458 := (or #27457 #27452)
+#27459 := (not #27458)
+#27466 := (or #26353 #27459)
+#27470 := (iff #27466 #27468)
+#27472 := (iff #27468 #27468)
+#27473 := [rewrite]: #27472
+#27464 := (iff #27459 #27463)
+#27461 := (iff #27458 #27460)
+#27462 := [rewrite]: #27461
+#27465 := [monotonicity #27462]: #27464
+#27471 := [monotonicity #27465]: #27470
+#27474 := [trans #27471 #27473]: #27470
+#27469 := [quant-inst]: #27466
+#27544 := [mp #27469 #27474]: #27468
+#27720 := [unit-resolution #27544 #20322]: #27463
+#27547 := (or #27460 #27456)
+#27548 := [def-axiom]: #27547
+#27721 := [unit-resolution #27548 #27720]: #27456
+#27726 := (= #3220 #27449)
+#27724 := (= #27449 #3220)
+#27714 := (= uf_288 #26317)
+#27712 := (= #2980 #26317)
+#27713 := [symm #27711]: #27712
+#27715 := [trans #27659 #27713]: #27714
+#27717 := [monotonicity #27654 #27715]: #27716
+#27719 := [symm #27717]: #27718
+#27725 := [monotonicity #27719 #27656]: #27724
+#27727 := [symm #27725]: #27726
+#27729 := [trans #27727 #27721]: #27728
+#27731 := [monotonicity #27729]: #27730
+#27735 := [symm #27731]: #27734
+#27598 := (= #24108 #27597)
+#27537 := (or #24175 #27598)
+#27538 := [quant-inst]: #27537
+#27706 := [unit-resolution #27538 #22417]: #27598
+#27736 := [trans #27706 #27735]: #27352
+#27231 := (not #27352)
+#27348 := (iff #27350 #27352)
+#27356 := (or #25427 #27348)
+#27353 := (iff #27352 #27350)
+#27357 := (or #25427 #27353)
+#27227 := (iff #27357 #27356)
+#27229 := (iff #27356 #27356)
+#27191 := [rewrite]: #27229
+#27354 := (iff #27353 #27348)
+#27355 := [rewrite]: #27354
+#27228 := [monotonicity #27355]: #27227
+#27192 := [trans #27228 #27191]: #27227
+#27226 := [quant-inst]: #27357
+#27230 := [mp #27226 #27192]: #27356
+#27636 := [unit-resolution #27230 #22404]: #27348
+#27411 := (not #27350)
+#27619 := (iff #17721 #27411)
+#27484 := (iff #12346 #27350)
+#27647 := (iff #27350 #12346)
+#27648 := [monotonicity #27639]: #27647
+#27618 := [symm #27648]: #27484
+#27693 := [monotonicity #27618]: #27619
+#27637 := [hypothesis]: #17721
+#27704 := [mp #27637 #27693]: #27411
+#27225 := (not #27348)
+#27232 := (or #27225 #27350 #27231)
+#27233 := [def-axiom]: #27232
+#27705 := [unit-resolution #27233 #27704 #27636]: #27231
+#27737 := [unit-resolution #27705 #27736]: false
+#27738 := [lemma #27737]: #12346
+#29878 := [trans #27738 #29877]: #27350
+#28251 := (or #27411 #28243)
+#28339 := (or #25411 #27411 #28243)
+#28250 := (or #28243 #27411)
+#28357 := (or #25411 #28250)
+#28372 := (iff #28357 #28339)
+#28362 := (or #25411 #28251)
+#28364 := (iff #28362 #28339)
+#28371 := [rewrite]: #28364
+#28363 := (iff #28357 #28362)
+#28252 := (iff #28250 #28251)
+#28253 := [rewrite]: #28252
+#28349 := [monotonicity #28253]: #28363
+#28380 := [trans #28349 #28371]: #28372
+#28361 := [quant-inst]: #28357
+#28381 := [mp #28361 #28380]: #28339
+#29879 := [unit-resolution #28381 #17465]: #28251
+#29880 := [unit-resolution #29879 #29878]: #28243
+#29914 := [symm #29880]: #29913
+#29917 := (= #27485 #28242)
+#29001 := (* #26315 #27482)
+#29004 := (+ #26574 #29001)
+#29007 := (uf_43 #24108 #29004)
+#29911 := (= #29007 #28242)
+#29905 := (= #29004 #28238)
+#28185 := (+ #26317 #27453)
+#28186 := (uf_43 #24108 #28185)
+#28369 := (uf_116 #28186)
+#29901 := (= #28369 #28238)
+#29887 := (= #28186 #3220)
+#28181 := (uf_66 #23935 uf_298 #24108)
+#29885 := (= #28181 #3220)
+#29886 := [monotonicity #27709 #27656]: #29885
+#29883 := (= #28186 #28181)
+#28187 := (= #28181 #28186)
+#28188 := (not #28187)
+#28182 := (uf_138 #28181 #23935)
+#28183 := (= uf_9 #28182)
+#28184 := (not #28183)
+#28191 := (or #28184 #28188)
+#28194 := (not #28191)
+#28286 := (or #26353 #28194)
+#28189 := (or #28188 #28184)
+#28190 := (not #28189)
+#28287 := (or #26353 #28190)
+#28288 := (iff #28287 #28286)
+#28290 := (iff #28286 #28286)
+#28262 := [rewrite]: #28290
+#28195 := (iff #28190 #28194)
+#28192 := (iff #28189 #28191)
+#28193 := [rewrite]: #28192
+#28196 := [monotonicity #28193]: #28195
+#28289 := [monotonicity #28196]: #28288
+#28263 := [trans #28289 #28262]: #28288
+#28283 := [quant-inst]: #28287
+#28254 := [mp #28283 #28263]: #28286
+#29881 := [unit-resolution #28254 #20322]: #28194
+#28267 := (or #28191 #28187)
+#28265 := [def-axiom]: #28267
+#29882 := [unit-resolution #28265 #29881]: #28187
+#29884 := [symm #29882]: #29883
+#29888 := [trans #29884 #29886]: #29887
+#29902 := [monotonicity #29888]: #29901
+#29903 := (= #29004 #28369)
+#28370 := (= #28185 #28369)
+#28373 := (* -1::int #28369)
+#28374 := (+ #27453 #28373)
+#28375 := (+ #26317 #28374)
+#28843 := (<= #28375 0::int)
+#28376 := (= #28375 0::int)
+#28838 := (or #24181 #28376)
+#28842 := (or #24181 #28370)
+#28912 := (iff #28842 #28838)
+#28907 := (iff #28838 #28838)
+#28908 := [rewrite]: #28907
+#28377 := (iff #28370 #28376)
+#28378 := [rewrite]: #28377
+#28841 := [monotonicity #28378]: #28912
+#28865 := [trans #28841 #28908]: #28912
+#28911 := [quant-inst]: #28842
+#28839 := [mp #28911 #28865]: #28838
+#29889 := [unit-resolution #28839 #22411]: #28376
+#29890 := (not #28376)
+#29891 := (or #29890 #28843)
+#29892 := [th-lemma]: #29891
+#29893 := [unit-resolution #29892 #29889]: #28843
+#28765 := (>= #28375 0::int)
+#29894 := (or #29890 #28765)
+#29895 := [th-lemma]: #29894
+#29896 := [unit-resolution #29895 #29889]: #28765
+#29897 := [th-lemma #29896 #29893]: #28370
+#29899 := (= #29004 #28185)
+#29868 := (= #28185 #29004)
+#29869 := (* -1::int #29004)
+#29870 := (+ #28185 #29869)
+#29871 := (<= #29870 0::int)
+#29267 := (* -1::int #26574)
+#29276 := (+ #24010 #29267)
+#29277 := (<= #29276 0::int)
+#29271 := (= #24010 #26574)
+#29860 := [symm #27879]: #29271
+#29861 := (not #29271)
+#29862 := (or #29861 #29277)
+#29863 := [th-lemma]: #29862
+#29935 := [unit-resolution #29863 #29860]: #29277
+#28718 := (>= #26315 1::int)
+#28559 := (= #26315 1::int)
#2932 := (uf_139 uf_7)
#2933 := (= #2932 1::int)
-#11840 := [asserted]: #2933
-#30673 := (= #27060 #2932)
-#30674 := [monotonicity #28401]: #30673
-#30675 := [trans #30674 #11840]: #28935
-#30676 := (not #28935)
-#30677 := (or #30676 #28937)
-#30672 := [th-lemma]: #30677
-#30627 := [unit-resolution #30672 #30675]: #28937
-#28936 := (<= #27060 1::int)
-#30678 := (or #30676 #28936)
-#30633 := [th-lemma]: #30678
-#30679 := [unit-resolution #30633 #30675]: #28936
-#30693 := (not #29600)
-#30698 := (not #29294)
-#30697 := (not #28936)
-#30695 := (not #28937)
-#30682 := (not #28934)
-#30681 := (not #29292)
-#30680 := (not #29298)
-#30694 := (or #30518 #30680 #30681 #30682 #30695 #30697 #30695 #30697 #30698 #30693)
-#30699 := [th-lemma]: #30694
-#30696 := [unit-resolution #30699 #30679 #30627 #30666 #29479 #30628 #30612 #30595]: #30518
-#30528 := (>= #30517 0::int)
-#28933 := (<= #28932 0::int)
-#30700 := (or #30589 #28933)
-#30701 := [th-lemma]: #30700
-#30702 := [unit-resolution #30701 #30513]: #28933
-#29601 := (>= #29599 0::int)
-#30703 := (or #30598 #29601)
-#30704 := [th-lemma]: #30703
-#30705 := [unit-resolution #30704 #30597]: #29601
-#29299 := (>= #29297 0::int)
-#30706 := (or #30661 #29299)
-#30709 := [th-lemma]: #30706
-#30741 := [unit-resolution #30709 #30660]: #29299
-#29291 := (<= #29278 0::int)
-#30742 := (or #30664 #29291)
-#30743 := [th-lemma]: #30742
-#30744 := [unit-resolution #30743 #30663]: #29291
-#30748 := (not #29601)
-#30747 := (not #28933)
-#30746 := (not #29291)
-#30745 := (not #29299)
-#30749 := (or #30528 #30745 #30746 #30747 #30695 #30697 #30695 #30697 #30529 #30748)
-#30750 := [th-lemma]: #30749
-#30751 := [unit-resolution #30750 #30679 #30627 #30744 #30527 #30741 #30705 #30702]: #30528
-#30753 := (not #30528)
-#30752 := (not #30518)
-#30754 := (or #30515 #30752 #30753)
-#30755 := [th-lemma]: #30754
-#30756 := [unit-resolution #30755 #30751 #30696]: #30515
-#30061 := [symm #30756]: #30551
-#30059 := [monotonicity #30061]: #30571
-#30587 := (= #29158 #29406)
-#29458 := (uf_66 #26144 #29134 #24854)
-#29409 := (= #29458 #29406)
-#29477 := (= #29406 #29458)
-#29487 := (not #29477)
-#29459 := (uf_138 #29458 #26144)
-#29460 := (= uf_9 #29459)
-#29393 := (not #29460)
-#29517 := (or #29393 #29487)
-#29521 := (not #29517)
-#29528 := (or #27098 #29521)
-#29394 := (* #29134 #27060)
-#29392 := (+ #27319 #29394)
-#29395 := (uf_43 #24854 #29392)
-#29396 := (= #29458 #29395)
-#29397 := (not #29396)
-#29398 := (or #29397 #29393)
-#29399 := (not #29398)
-#29529 := (or #27098 #29399)
-#29538 := (iff #29529 #29528)
-#29540 := (iff #29528 #29528)
-#29541 := [rewrite]: #29540
-#29522 := (iff #29399 #29521)
-#29519 := (iff #29398 #29517)
-#29514 := (or #29487 #29393)
-#29518 := (iff #29514 #29517)
-#29513 := [rewrite]: #29518
-#29515 := (iff #29398 #29514)
-#29498 := (iff #29397 #29487)
-#29485 := (iff #29396 #29477)
-#29483 := (iff #29409 #29477)
-#29484 := [rewrite]: #29483
-#29410 := (iff #29396 #29409)
-#29407 := (= #29395 #29406)
-#29402 := (= #29392 #29404)
-#29488 := (= #29394 #29400)
-#29403 := [rewrite]: #29488
-#29405 := [monotonicity #29403]: #29402
-#29408 := [monotonicity #29405]: #29407
-#29482 := [monotonicity #29408]: #29410
-#29486 := [trans #29482 #29484]: #29485
-#29499 := [monotonicity #29486]: #29498
-#29516 := [monotonicity #29499]: #29515
-#29520 := [trans #29516 #29513]: #29519
-#29523 := [monotonicity #29520]: #29522
-#29539 := [monotonicity #29523]: #29538
-#29542 := [trans #29539 #29541]: #29538
-#29537 := [quant-inst]: #29529
-#29535 := [mp #29537 #29542]: #29528
-#30545 := [unit-resolution #29535 #21133]: #29521
-#29593 := (or #29517 #29477)
-#29594 := [def-axiom]: #29593
-#30568 := [unit-resolution #29594 #30545]: #29477
-#30570 := [symm #30568]: #29409
-#30544 := (= #29158 #29458)
-#30569 := [monotonicity #28529]: #30544
-#30550 := [trans #30569 #30570]: #30587
-#30062 := [trans #30550 #30059]: #30557
-#30063 := [trans #30062 #30574]: #30560
-#30104 := [monotonicity #30063]: #30562
-#30064 := [symm #30104]: #30564
-#30065 := [monotonicity #30064]: #30599
-#30066 := [trans #30065 #30596]: #30602
-#30067 := [trans #30066 #28503]: #30604
-#30113 := [monotonicity #30067]: #30110
-#30121 := [symm #30113]: #30109
-#30077 := [trans #14286 #30121]: #29304
-#30074 := (not #29304)
-#30075 := (or #29360 #30074)
-#30076 := [def-axiom]: #30075
-#30079 := [unit-resolution #30076 #30077]: #29360
-#30192 := (= #25393 #29318)
-#30150 := (= #29318 #25393)
-#30108 := (= #28096 uf_7)
-#30078 := (= #28096 #24854)
-#29085 := (or #27970 #28156 #28097)
-#29086 := [def-axiom]: #29085
-#30080 := [unit-resolution #29086 #30547 #28381]: #28097
-#30115 := [symm #30080]: #30078
-#30149 := [trans #30115 #28401]: #30108
-#30189 := [monotonicity #30149]: #30150
-#30158 := [symm #30189]: #30192
-#30159 := [trans #27692 #30158]: #29336
-#29899 := (not #29309)
-#30092 := (iff #11905 #29899)
-#30090 := (iff #11902 #29309)
-#30088 := (iff #29309 #11902)
-#30103 := (= #29308 #2990)
-#30100 := (= #29307 #2977)
-#29898 := (= #29307 #24974)
-#30160 := [monotonicity #30067]: #29898
-#30102 := [trans #30160 #28564]: #30100
-#30101 := [monotonicity #30102]: #30103
-#30089 := [monotonicity #30101]: #30088
-#30091 := [symm #30089]: #30090
-#30087 := [monotonicity #30091]: #30092
-#30155 := [mp #14289 #30087]: #29899
-#29226 := (uf_66 #27214 #29134 #24854)
-#29227 := (uf_58 #3157 #29226)
-#29228 := (uf_136 #29227)
-#29229 := (= uf_9 #29228)
-#29261 := (not #29229)
-#30308 := (iff #29261 #29315)
-#30318 := (iff #29229 #29314)
-#30316 := (iff #29314 #29229)
-#30195 := (= #29313 #29228)
-#30206 := (= #27840 #29227)
-#30204 := (= #29227 #27840)
-#30161 := (= #29226 #3188)
-#30194 := (= #29134 uf_298)
-#30163 := [symm #29491]: #30194
-#30162 := [monotonicity #28408 #30163 #28401]: #30161
-#30205 := [monotonicity #30162]: #30204
-#29918 := [symm #30205]: #30206
-#30207 := [monotonicity #29918]: #30195
-#30317 := [monotonicity #30207]: #30316
-#30321 := [symm #30317]: #30318
-#30327 := [monotonicity #30321]: #30308
-#29230 := (uf_24 uf_287 #29226)
-#29231 := (= uf_9 #29230)
-#29232 := (not #29231)
-#29236 := (or #29229 #29232)
-#29239 := (not #29236)
-#30188 := (or #29175 #29185 #29239)
-#29250 := (or #27245 #24694 #27226 #27229 #29175 #29185 #29239)
-#29233 := (or #29232 #29229)
-#29234 := (not #29233)
-#29235 := (or #27229 #29175 #29173 #24694 #27226 #29234)
-#29251 := (or #27245 #29235)
-#29258 := (iff #29251 #29250)
-#29245 := (or #24694 #27226 #27229 #29175 #29185 #29239)
-#29253 := (or #27245 #29245)
-#29256 := (iff #29253 #29250)
-#29257 := [rewrite]: #29256
-#29254 := (iff #29251 #29253)
-#29248 := (iff #29235 #29245)
-#29242 := (or #27229 #29175 #29185 #24694 #27226 #29239)
-#29246 := (iff #29242 #29245)
-#29247 := [rewrite]: #29246
-#29243 := (iff #29235 #29242)
-#29240 := (iff #29234 #29239)
-#29237 := (iff #29233 #29236)
-#29238 := [rewrite]: #29237
-#29241 := [monotonicity #29238]: #29240
-#29244 := [monotonicity #29189 #29241]: #29243
-#29249 := [trans #29244 #29247]: #29248
-#29255 := [monotonicity #29249]: #29254
-#29259 := [trans #29255 #29257]: #29258
-#29252 := [quant-inst]: #29251
-#29260 := [mp #29252 #29259]: #29250
-#30196 := [unit-resolution #29260 #21577 #14275 #28847 #28441]: #30188
-#30197 := [unit-resolution #30196 #29481 #30535]: #29239
-#29262 := (or #29236 #29261)
-#29263 := [def-axiom]: #29262
-#30154 := [unit-resolution #29263 #30197]: #29261
-#30328 := [mp #30154 #30327]: #29315
-#30068 := (or #29354 #29314)
-#30069 := [def-axiom]: #30068
-#30362 := [unit-resolution #30069 #30328]: #29354
-#29889 := (or #29374 #29309 #29351 #29357 #29363)
-#29911 := [def-axiom]: #29889
-#30363 := [unit-resolution #29911 #30362 #30155 #30159 #30079 #30058]: false
-#30364 := [lemma #30363]: #29374
-#29947 := (or #29377 #29369)
-#29948 := [def-axiom]: #29947
-#30786 := [unit-resolution #29948 #30364]: #29377
-#24581 := (or #23384 #13436)
-#24582 := [def-axiom]: #24581
-#29645 := [unit-resolution #24582 #29548]: #13436
-#28447 := (or #12372 #13576 #27226 #13721)
-#28390 := [hypothesis]: #13436
-#28226 := (uf_66 #27078 uf_298 #27613)
-#28227 := (uf_125 #28226 #27078)
-#28257 := (* -1::int #28227)
-#28385 := (+ uf_298 #28257)
-#28386 := (<= #28385 0::int)
-#28228 := (= uf_298 #28227)
-#28231 := (or #27121 #28228)
-#28232 := [quant-inst]: #28231
-#28391 := [unit-resolution #28232 #17002]: #28228
-#28394 := (not #28228)
-#28395 := (or #28394 #28386)
-#28396 := [th-lemma]: #28395
-#28397 := [unit-resolution #28396 #28391]: #28386
-#28248 := (>= #28227 0::int)
-#28249 := (not #28248)
-#28300 := (uf_66 #27214 #28227 #24854)
-#28304 := (uf_24 uf_287 #28300)
-#28305 := (= uf_9 #28304)
-#28306 := (not #28305)
-#28301 := (uf_58 #3157 #28300)
-#28302 := (uf_136 #28301)
-#28303 := (= uf_9 #28302)
-#28310 := (or #28303 #28306)
-#28417 := (iff #18452 #28306)
-#28415 := (iff #12372 #28305)
-#28413 := (iff #28305 #12372)
-#28411 := (= #28304 #3191)
-#28409 := (= #28300 #3188)
-#28402 := (= #28227 uf_298)
-#28403 := [symm #28391]: #28402
-#28410 := [monotonicity #28408 #28403 #28401]: #28409
-#28412 := [monotonicity #28410]: #28411
-#28414 := [monotonicity #28412]: #28413
-#28416 := [symm #28414]: #28415
-#28418 := [monotonicity #28416]: #28417
-#28398 := [hypothesis]: #18452
-#28419 := [mp #28398 #28418]: #28306
-#28338 := (or #28310 #28305)
-#28339 := [def-axiom]: #28338
-#28420 := [unit-resolution #28339 #28419]: #28310
-#28258 := (+ #24856 #28257)
-#28259 := (<= #28258 0::int)
-#28430 := (not #28259)
-#28426 := [hypothesis]: #13722
-#28387 := (>= #28385 0::int)
-#28427 := (or #28394 #28387)
-#28428 := [th-lemma]: #28427
-#28429 := [unit-resolution #28428 #28391]: #28387
-#28432 := (not #28387)
-#28433 := (or #28430 #28431 #28432 #13721)
-#28434 := [th-lemma]: #28433
-#28435 := [unit-resolution #28434 #28429 #28426 #28425]: #28430
-#28313 := (not #28310)
-#28443 := (or #28249 #28259 #28313)
-#28324 := (or #27245 #24694 #27226 #27229 #28249 #28259 #28313)
-#28307 := (or #28306 #28303)
-#28308 := (not #28307)
-#28246 := (+ #28227 #27155)
-#28247 := (>= #28246 0::int)
-#28309 := (or #27229 #28249 #28247 #24694 #27226 #28308)
-#28325 := (or #27245 #28309)
-#28332 := (iff #28325 #28324)
-#28319 := (or #24694 #27226 #27229 #28249 #28259 #28313)
-#28327 := (or #27245 #28319)
-#28330 := (iff #28327 #28324)
-#28331 := [rewrite]: #28330
-#28328 := (iff #28325 #28327)
-#28322 := (iff #28309 #28319)
-#28316 := (or #27229 #28249 #28259 #24694 #27226 #28313)
-#28320 := (iff #28316 #28319)
-#28321 := [rewrite]: #28320
-#28317 := (iff #28309 #28316)
-#28314 := (iff #28308 #28313)
-#28311 := (iff #28307 #28310)
-#28312 := [rewrite]: #28311
-#28315 := [monotonicity #28312]: #28314
-#28262 := (iff #28247 #28259)
-#28251 := (+ #27155 #28227)
-#28254 := (>= #28251 0::int)
-#28260 := (iff #28254 #28259)
-#28261 := [rewrite]: #28260
-#28255 := (iff #28247 #28254)
-#28252 := (= #28246 #28251)
-#28253 := [rewrite]: #28252
-#28256 := [monotonicity #28253]: #28255
-#28263 := [trans #28256 #28261]: #28262
-#28318 := [monotonicity #28263 #28315]: #28317
-#28323 := [trans #28318 #28321]: #28322
-#28329 := [monotonicity #28323]: #28328
-#28333 := [trans #28329 #28331]: #28332
-#28326 := [quant-inst]: #28325
-#28334 := [mp #28326 #28333]: #28324
-#28444 := [unit-resolution #28334 #21577 #14275 #28442 #28441]: #28443
-#28445 := [unit-resolution #28444 #28435 #28420]: #28249
-#28446 := [th-lemma #28445 #28397 #28390]: false
-#28448 := [lemma #28446]: #28447
-#29646 := [unit-resolution #28448 #29588 #28847 #29645]: #12372
-#30096 := (or #29386 #18452 #29380)
-#30097 := [def-axiom]: #30096
-#30787 := [unit-resolution #30097 #29646 #30786]: #29386
-#29962 := (or #29961 #12378 #29383)
-#29960 := [def-axiom]: #29962
-#30792 := [unit-resolution #29960 #30787]: #30791
-#30793 := [unit-resolution #30792 #30191]: #12378
-#30801 := (or #18458 #23317)
-#24505 := (or #23338 #23332)
-#24506 := [def-axiom]: #24505
-#30794 := [unit-resolution #24506 #29587]: #23332
-#30795 := (or #23335 #18452 #23329)
-#24497 := (or #23335 #18449 #18452 #23329)
-#24498 := [def-axiom]: #24497
-#30796 := [unit-resolution #24498 #28483]: #30795
-#30797 := [unit-resolution #30796 #29646 #30794]: #23329
-#24487 := (or #23326 #23320)
-#24488 := [def-axiom]: #24487
-#30798 := [unit-resolution #24488 #30797]: #23320
-#30799 := (or #23323 #18458 #23317)
-#24481 := (or #23323 #18449 #18458 #23317)
-#24482 := [def-axiom]: #24481
-#30800 := [unit-resolution #24482 #28483]: #30799
-#30802 := [unit-resolution #30800 #30798]: #30801
-#30803 := [unit-resolution #30802 #30793]: #23317
-#30804 := (or #23314 #23275)
-#24329 := (not #18823)
-#30463 := (= #3083 #3227)
-#30461 := (= #3227 #3083)
-#30457 := (= #3226 #3082)
-#30454 := [hypothesis]: #23305
-#24443 := (or #23302 #12671)
-#24444 := [def-axiom]: #24443
-#30455 := [unit-resolution #24444 #30454]: #12671
-#30456 := [symm #30455]: #3264
-#30458 := [monotonicity #30456]: #30457
-#30462 := [monotonicity #30458]: #30461
-#30464 := [symm #30462]: #30463
-#30465 := (= uf_304 #3083)
-#24441 := (or #23302 #12668)
-#24442 := [def-axiom]: #24441
-#30459 := [unit-resolution #24442 #30454]: #12668
-#30460 := [symm #30459]: #3263
-#30466 := [trans #30460 #29551]: #30465
-#30467 := [trans #30466 #30464]: #12428
-#24438 := (+ uf_297 #13512)
-#24440 := (>= #24438 0::int)
-#30468 := (or #12677 #24440)
-#30469 := [th-lemma]: #30468
-#30420 := [unit-resolution #30469 #30455]: #24440
-#30082 := (not #24440)
-#30083 := (or #13515 #30082)
-#30111 := [hypothesis]: #24440
-#30112 := [hypothesis]: #13514
-#30081 := [th-lemma #30112 #30111 #29553]: false
-#30084 := [lemma #30081]: #30083
-#30424 := [unit-resolution #30084 #30420]: #13515
-#24317 := (or #22257 #22255 #13514)
-#24318 := [def-axiom]: #24317
-#30425 := [unit-resolution #24318 #30424 #30467]: #22257
-#24319 := (or #23236 #22256)
-#24320 := [def-axiom]: #24319
-#30421 := [unit-resolution #24320 #30425]: #23236
-#24453 := (or #23302 #23266)
-#24454 := [def-axiom]: #24453
-#30426 := [unit-resolution #24454 #30454]: #23266
-#30430 := (or #23269 #23263)
-#15731 := (<= uf_286 4294967295::int)
-#15734 := (iff #13361 #15731)
-#15725 := (+ 4294967295::int #13362)
-#15728 := (>= #15725 0::int)
-#15732 := (iff #15728 #15731)
-#15733 := [rewrite]: #15732
-#15729 := (iff #13361 #15728)
-#15726 := (= #13363 #15725)
-#15727 := [monotonicity #7510]: #15726
-#15730 := [monotonicity #15727]: #15729
-#15735 := [trans #15730 #15733]: #15734
-#14277 := [not-or-elim #14266]: #13366
-#14279 := [and-elim #14277]: #13361
-#15736 := [mp #14279 #15735]: #15731
-#29589 := [hypothesis]: #18490
-#29590 := [th-lemma #29589 #29588 #15736]: false
-#29591 := [lemma #29590]: #15820
-#30427 := (or #13576 #13454)
-#30428 := [th-lemma]: #30427
-#30429 := [unit-resolution #30428 #29645]: #13454
-#24377 := (or #23269 #18487 #18490 #23263)
-#24378 := [def-axiom]: #24377
-#30431 := [unit-resolution #24378 #30429 #29591]: #30430
-#30432 := [unit-resolution #30431 #30426]: #23263
-#24367 := (or #23260 #23254)
-#24368 := [def-axiom]: #24367
-#30434 := [unit-resolution #24368 #30432]: #23254
-#24312 := (>= #13539 -1::int)
-#24363 := (or #23260 #13538)
-#24364 := [def-axiom]: #24363
-#30435 := [unit-resolution #24364 #30432]: #13538
-#30433 := (or #13542 #24312)
-#30436 := [th-lemma]: #30433
-#30437 := [unit-resolution #30436 #30435]: #24312
-#30126 := (not #24312)
-#30127 := (or #13470 #30126)
-#30085 := [hypothesis]: #24312
-#30086 := [hypothesis]: #13475
-#30120 := [th-lemma #30086 #29588 #30085]: false
-#30128 := [lemma #30120]: #30127
-#30438 := [unit-resolution #30128 #30437]: #13470
-#24353 := (or #23257 #13475 #23251)
-#24354 := [def-axiom]: #24353
-#30439 := [unit-resolution #24354 #30438 #30434]: #23251
-#24345 := (or #23248 #23242)
-#24346 := [def-axiom]: #24345
-#30440 := [unit-resolution #24346 #30439]: #23242
-#24341 := (or #23245 #23239 #22309)
-#24342 := [def-axiom]: #24341
-#30441 := [unit-resolution #24342 #30440 #30421]: #22309
-#24330 := (or #22304 #24329)
-#24331 := [def-axiom]: #24330
-#30482 := [unit-resolution #24331 #30441]: #24329
-#30226 := (+ uf_298 #18821)
-#30453 := (>= #30226 0::int)
-#30481 := (not #30453)
-#30707 := (= uf_298 ?x786!14)
-#30738 := (not #30707)
-#24451 := (or #23302 #13628)
-#24452 := [def-axiom]: #24451
-#30443 := [unit-resolution #24452 #30454]: #13628
-#24435 := (+ uf_296 #13490)
-#24436 := (<= #24435 0::int)
-#30445 := (or #12686 #24436)
-#30446 := [th-lemma]: #30445
-#30444 := [unit-resolution #30446 #30459]: #24436
-#24332 := (not #18836)
-#24333 := (or #22304 #24332)
-#24334 := [def-axiom]: #24333
-#30447 := [unit-resolution #24334 #30441]: #24332
-#30725 := (not #24436)
-#30726 := (or #30724 #18836 #30725 #13627)
-#30719 := [hypothesis]: #13628
-#30720 := [hypothesis]: #24436
-#30721 := [hypothesis]: #24332
-#30722 := [hypothesis]: #30714
-#30723 := [th-lemma #30722 #30721 #30720 #30719]: false
-#30727 := [lemma #30723]: #30726
-#30448 := [unit-resolution #30727 #30447 #30444 #30443]: #30724
-#30715 := (or #30713 #30714)
-#30716 := [th-lemma]: #30715
-#30449 := [unit-resolution #30716 #30448]: #30713
-#30739 := (or #30738 #30708)
-#30734 := (= #18513 #3197)
-#30732 := (= #18512 #3188)
-#30730 := (= ?x786!14 uf_298)
-#30729 := [hypothesis]: #30707
-#30731 := [symm #30729]: #30730
-#30733 := [monotonicity #30731]: #30732
-#30735 := [monotonicity #30733]: #30734
-#30736 := [symm #30735]: #30708
-#30728 := [hypothesis]: #30713
-#30737 := [unit-resolution #30728 #30736]: false
-#30740 := [lemma #30737]: #30739
-#30450 := [unit-resolution #30740 #30449]: #30738
-#30485 := (or #30707 #30481)
-#30224 := (<= #30226 0::int)
-#30262 := (+ uf_296 #18834)
-#30263 := (>= #30262 0::int)
-#30451 := (not #30263)
-#30452 := (or #30451 #18836 #30725)
-#30494 := [th-lemma]: #30452
-#30495 := [unit-resolution #30494 #30444 #30447]: #30451
-#30497 := (or #30224 #30263)
-#24327 := (or #22304 #18509)
-#24328 := [def-axiom]: #24327
-#30496 := [unit-resolution #24328 #30441]: #18509
-#24325 := (or #22304 #18508)
-#24326 := [def-axiom]: #24325
-#30491 := [unit-resolution #24326 #30441]: #18508
-#30279 := (or #23225 #22288 #22289 #30224 #30263)
-#30208 := (+ #18513 #13926)
-#30203 := (<= #30208 0::int)
-#30209 := (+ ?x786!14 #13457)
-#30210 := (>= #30209 0::int)
-#30212 := (or #22289 #30210 #30203 #22288)
-#30280 := (or #23225 #30212)
-#30287 := (iff #30280 #30279)
-#30274 := (or #22288 #22289 #30224 #30263)
-#30282 := (or #23225 #30274)
-#30285 := (iff #30282 #30279)
-#30286 := [rewrite]: #30285
-#30283 := (iff #30280 #30282)
-#30277 := (iff #30212 #30274)
-#30271 := (or #22289 #30224 #30263 #22288)
-#30275 := (iff #30271 #30274)
-#30276 := [rewrite]: #30275
-#30272 := (iff #30212 #30271)
-#30269 := (iff #30203 #30263)
-#30257 := (+ #13926 #18513)
-#30259 := (<= #30257 0::int)
-#30267 := (iff #30259 #30263)
-#30268 := [rewrite]: #30267
-#30260 := (iff #30203 #30259)
-#30258 := (= #30208 #30257)
-#30253 := [rewrite]: #30258
-#30261 := [monotonicity #30253]: #30260
-#30270 := [trans #30261 #30268]: #30269
-#30241 := (iff #30210 #30224)
-#30213 := (+ #13457 ?x786!14)
-#30223 := (>= #30213 0::int)
-#30227 := (iff #30223 #30224)
-#30228 := [rewrite]: #30227
-#30211 := (iff #30210 #30223)
-#30218 := (= #30209 #30213)
-#30219 := [rewrite]: #30218
-#30225 := [monotonicity #30219]: #30211
-#30242 := [trans #30225 #30228]: #30241
-#30273 := [monotonicity #30242 #30270]: #30272
-#30278 := [trans #30273 #30276]: #30277
-#30284 := [monotonicity #30278]: #30283
-#30288 := [trans #30284 #30286]: #30287
-#30281 := [quant-inst]: #30280
-#30289 := [mp #30281 #30288]: #30279
-#30498 := [unit-resolution #30289 #29550 #30491 #30496]: #30497
-#30499 := [unit-resolution #30498 #30495]: #30224
-#30500 := (not #30224)
-#30510 := (or #30707 #30500 #30481)
-#30484 := [th-lemma]: #30510
-#30480 := [unit-resolution #30484 #30499]: #30485
-#30486 := [unit-resolution #30480 #30450]: #30481
-#30487 := [th-lemma #30486 #30437 #30482]: false
-#30483 := [lemma #30487]: #23302
-#29652 := (or #23314 #23305 #23275)
-#29632 := [hypothesis]: #23272
-#29633 := [hypothesis]: #23317
-#24465 := (or #23314 #12378)
-#24466 := [def-axiom]: #24465
-#29636 := [unit-resolution #24466 #29633]: #12378
-#29637 := (or #23281 #18458 #23275)
-#24407 := (or #23281 #18449 #18458 #23275)
-#24408 := [def-axiom]: #24407
-#29638 := [unit-resolution #24408 #28483]: #29637
-#29639 := [unit-resolution #29638 #29636 #29632]: #23281
-#24413 := (or #23284 #23278)
-#24414 := [def-axiom]: #24413
-#29640 := [unit-resolution #24414 #29639]: #23284
-#29641 := [hypothesis]: #23302
-#24471 := (or #23314 #23308)
-#24472 := [def-axiom]: #24471
-#29642 := [unit-resolution #24472 #29633]: #23308
-#24461 := (or #23311 #23299 #23305)
-#24462 := [def-axiom]: #24461
-#29643 := [unit-resolution #24462 #29642 #29641]: #23299
-#24431 := (or #23296 #23290)
-#24432 := [def-axiom]: #24431
-#29644 := [unit-resolution #24432 #29643]: #23290
-#29649 := (or #23293 #23287)
-#29647 := (or #23293 #18452 #23287)
-#24423 := (or #23293 #18449 #18452 #23287)
-#24424 := [def-axiom]: #24423
-#29648 := [unit-resolution #24424 #28483]: #29647
-#29650 := [unit-resolution #29648 #29646]: #29649
-#29651 := [unit-resolution #29650 #29644 #29640]: false
-#29653 := [lemma #29651]: #29652
-#30805 := [unit-resolution #29653 #30483]: #30804
-#30806 := [unit-resolution #30805 #30803]: #23275
-#24389 := (or #23272 #12389)
-#24390 := [def-axiom]: #24389
-#30807 := [unit-resolution #24390 #30806]: #12389
-#24387 := (or #23272 #12384)
-#24388 := [def-axiom]: #24387
-#30808 := [unit-resolution #24388 #30806]: #12384
-#30809 := [trans #30808 #30807]: #30365
-#30810 := (not #30365)
-#30811 := (or #30810 #30319)
-#30812 := [th-lemma]: #30811
-#30813 := [unit-resolution #30812 #30809]: #30319
-#24397 := (or #23272 #23266)
-#24398 := [def-axiom]: #24397
-#30814 := [unit-resolution #24398 #30806]: #23266
-#30815 := [unit-resolution #30431 #30814]: #23263
-#30816 := [unit-resolution #24368 #30815]: #23254
-#30817 := [unit-resolution #24364 #30815]: #13538
-#30818 := [unit-resolution #30436 #30817]: #24312
-#30819 := [unit-resolution #30128 #30818]: #13470
-#30820 := [unit-resolution #24354 #30819 #30816]: #23251
-#30821 := [unit-resolution #24346 #30820]: #23242
-#30830 := (= #3197 #3227)
-#30826 := (= #3227 #3197)
-#30824 := (= #3226 #3188)
-#24391 := (or #23272 #12392)
-#24392 := [def-axiom]: #24391
-#30822 := [unit-resolution #24392 #30806]: #12392
-#30823 := [symm #30822]: #3207
-#30825 := [monotonicity #30823]: #30824
-#30827 := [monotonicity #30825]: #30826
-#30831 := [symm #30827]: #30830
-#30832 := (= uf_304 #3197)
-#30829 := [symm #30808]: #3200
-#30828 := [symm #30807]: #3205
-#30833 := [trans #30828 #30829]: #30832
-#30834 := [trans #30833 #30831]: #12428
-#30114 := (+ uf_298 #13512)
-#30060 := (>= #30114 0::int)
-#30835 := (or #12539 #30060)
-#30836 := [th-lemma]: #30835
-#30837 := [unit-resolution #30836 #30822]: #30060
-#30838 := (not #30060)
-#30839 := (or #13515 #30838 #13721)
-#30840 := [th-lemma]: #30839
-#30841 := [unit-resolution #30840 #29588 #30837]: #13515
-#30842 := [unit-resolution #24318 #30841 #30834]: #22257
-#30843 := [unit-resolution #24320 #30842]: #23236
-#30844 := [unit-resolution #24342 #30843 #30821]: #22309
-#30845 := [unit-resolution #24334 #30844]: #24332
-#30846 := (not #30319)
-#30847 := (or #30724 #18836 #30846)
-#30848 := [th-lemma]: #30847
-#30849 := [unit-resolution #30848 #30845 #30813]: #30724
-#30850 := [unit-resolution #30716 #30849]: #30713
-#30851 := [unit-resolution #24331 #30844]: #24329
-#30852 := (or #30453 #18823 #30126)
-#30853 := [th-lemma]: #30852
-#30854 := [unit-resolution #30853 #30851 #30818]: #30453
-#30855 := [unit-resolution #24472 #30803]: #23308
-#30856 := (or #23311 #23299)
-#30857 := [unit-resolution #24462 #30483]: #30856
-#30858 := [unit-resolution #30857 #30855]: #23299
-#24429 := (or #23296 #13627)
-#24430 := [def-axiom]: #24429
-#30859 := [unit-resolution #24430 #30858]: #13627
-#30860 := (or #24436 #30846 #13628)
-#30861 := [th-lemma]: #30860
-#30862 := [unit-resolution #30861 #30859 #30813]: #24436
-#30863 := [unit-resolution #30494 #30845 #30862]: #30451
-#30864 := [unit-resolution #24328 #30844]: #18509
-#30865 := [unit-resolution #24326 #30844]: #18508
-#30866 := [unit-resolution #30289 #29550 #30865 #30864 #30863]: #30224
-#30867 := [unit-resolution #30484 #30866 #30854]: #30707
-[unit-resolution #30740 #30867 #30850]: false
+#11811 := [asserted]: #2933
+#29875 := (= #26315 #2932)
+#29943 := [monotonicity #27656]: #29875
+#29944 := [trans #29943 #11811]: #28559
+#29945 := (not #28559)
+#29946 := (or #29945 #28718)
+#29947 := [th-lemma]: #29946
+#29948 := [unit-resolution #29947 #29944]: #28718
+#28736 := (<= #26315 1::int)
+#29949 := (or #29945 #28736)
+#29950 := [th-lemma]: #29949
+#29951 := [unit-resolution #29950 #29944]: #28736
+#28558 := (* -1::int #26317)
+#28545 := (+ #24010 #28558)
+#28694 := (>= #28545 0::int)
+#28692 := (= #24010 #26317)
+#29952 := [trans #27881 #27713]: #28692
+#29953 := (not #28692)
+#29954 := (or #29953 #28694)
+#29955 := [th-lemma]: #29954
+#29956 := [unit-resolution #29955 #29952]: #28694
+#29962 := (not #28548)
+#29961 := (not #28082)
+#29960 := (not #29277)
+#29958 := (not #28736)
+#29957 := (not #28718)
+#29959 := (not #28694)
+#29963 := (or #29871 #29957 #29958 #29959 #29957 #29958 #29960 #29961 #29962)
+#29964 := [th-lemma]: #29963
+#29965 := [unit-resolution #29964 #28879 #29956 #29951 #29948 #29935 #29444]: #29871
+#29872 := (>= #29870 0::int)
+#29285 := (>= #29276 0::int)
+#29969 := (or #29861 #29285)
+#29970 := [th-lemma]: #29969
+#29971 := [unit-resolution #29970 #29860]: #29285
+#28693 := (<= #28545 0::int)
+#29975 := (or #29953 #28693)
+#29976 := [th-lemma]: #29975
+#29977 := [unit-resolution #29976 #29952]: #28693
+#29979 := (not #29285)
+#29978 := (not #28693)
+#29982 := (or #29872 #29957 #29958 #29978 #29957 #29958 #29979 #29980 #29981)
+#29983 := [th-lemma]: #29982
+#29984 := [unit-resolution #29983 #29977 #29974 #29951 #29948 #29971 #29968]: #29872
+#29986 := (not #29872)
+#29985 := (not #29871)
+#29987 := (or #29868 #29985 #29986)
+#29988 := [th-lemma]: #29987
+#29989 := [unit-resolution #29988 #29984 #29965]: #29868
+#29577 := [symm #29989]: #29899
+#29571 := [trans #29577 #29897]: #29903
+#29578 := [trans #29571 #29902]: #29905
+#29562 := [monotonicity #29578]: #29911
+#29915 := (= #27485 #29007)
+#28990 := (uf_66 #25399 #27482 #24108)
+#29010 := (= #28990 #29007)
+#29013 := (not #29010)
+#28991 := (uf_138 #28990 #25399)
+#28992 := (= uf_9 #28991)
+#28993 := (not #28992)
+#29019 := (or #28993 #29013)
+#29024 := (not #29019)
+#29041 := (or #26353 #29024)
+#28994 := (* #27482 #26315)
+#28995 := (+ #26574 #28994)
+#28996 := (uf_43 #24108 #28995)
+#28997 := (= #28990 #28996)
+#28998 := (not #28997)
+#28999 := (or #28998 #28993)
+#29000 := (not #28999)
+#29040 := (or #26353 #29000)
+#29039 := (iff #29040 #29041)
+#29056 := (iff #29041 #29041)
+#29073 := [rewrite]: #29056
+#29025 := (iff #29000 #29024)
+#29022 := (iff #28999 #29019)
+#29016 := (or #29013 #28993)
+#29020 := (iff #29016 #29019)
+#29021 := [rewrite]: #29020
+#29017 := (iff #28999 #29016)
+#29014 := (iff #28998 #29013)
+#29011 := (iff #28997 #29010)
+#29008 := (= #28996 #29007)
+#29005 := (= #28995 #29004)
+#29002 := (= #28994 #29001)
+#29003 := [rewrite]: #29002
+#29006 := [monotonicity #29003]: #29005
+#29009 := [monotonicity #29006]: #29008
+#29012 := [monotonicity #29009]: #29011
+#29015 := [monotonicity #29012]: #29014
+#29018 := [monotonicity #29015]: #29017
+#29023 := [trans #29018 #29021]: #29022
+#29026 := [monotonicity #29023]: #29025
+#29072 := [monotonicity #29026]: #29039
+#29107 := [trans #29072 #29073]: #29039
+#29071 := [quant-inst]: #29040
+#29108 := [mp #29071 #29107]: #29041
+#29907 := [unit-resolution #29108 #20322]: #29024
+#29209 := (or #29019 #29010)
+#29207 := [def-axiom]: #29209
+#29908 := [unit-resolution #29207 #29907]: #29010
+#29909 := (= #27485 #28990)
+#29910 := [monotonicity #27889]: #29909
+#29916 := [trans #29910 #29908]: #29915
+#29561 := [trans #29916 #29562]: #29917
+#29502 := [trans #29561 #29914]: #29919
+#29563 := [monotonicity #29502]: #29921
+#29555 := [symm #29563]: #29923
+#29392 := [monotonicity #29555]: #29925
+#29556 := [trans #29392 #29928]: #29929
+#29557 := [trans #29556 #27863]: #29931
+#29559 := [monotonicity #29557]: #29558
+#29568 := [symm #29559]: #29567
+#29565 := [trans #13530 #29568]: #28582
+#29532 := (not #28582)
+#29533 := (or #28638 #29532)
+#29534 := [def-axiom]: #29533
+#29569 := [unit-resolution #29534 #29565]: #28638
+#29695 := (= #24648 #28596)
+#29691 := (= #28596 #24648)
+#29574 := (= #27351 uf_7)
+#29573 := (= #27351 #24108)
+#28379 := (or #27225 #27411 #27352)
+#28382 := [def-axiom]: #28379
+#29572 := [unit-resolution #28382 #29878 #27636]: #27352
+#29566 := [symm #29572]: #29573
+#29575 := [trans #29566 #27656]: #29574
+#29692 := [monotonicity #29575]: #29691
+#29696 := [symm #29692]: #29695
+#29582 := [trans #27842 #29696]: #28614
+#29507 := (not #28587)
+#29688 := (iff #11876 #29507)
+#29586 := (iff #11873 #28587)
+#28726 := (iff #28587 #11873)
+#29579 := (= #28586 #2990)
+#29592 := (= #28585 #2977)
+#29583 := (= #28585 #24228)
+#29581 := [monotonicity #29557]: #29583
+#29593 := [trans #29581 #27921]: #29592
+#29584 := [monotonicity #29593]: #29579
+#29585 := [monotonicity #29584]: #28726
+#29687 := [symm #29585]: #29586
+#29689 := [monotonicity #29687]: #29688
+#29690 := [mp #13533 #29689]: #29507
+#29698 := (iff #29038 #28593)
+#29699 := (iff #27558 #28592)
+#29686 := (iff #28592 #27558)
+#29677 := (= #28591 #27557)
+#29676 := (= #27095 #27556)
+#29673 := (= #27556 #27095)
+#29675 := [monotonicity #29712]: #29673
+#29674 := [symm #29675]: #29676
+#29684 := [monotonicity #29674]: #29677
+#29697 := [monotonicity #29684]: #29686
+#29700 := [symm #29697]: #29699
+#29685 := [monotonicity #29700]: #29698
+#29706 := [hypothesis]: #29038
+#29701 := [mp #29706 #29685]: #28593
+#29506 := (or #28632 #28592)
+#29471 := [def-axiom]: #29506
+#29702 := [unit-resolution #29471 #29701]: #28632
+#29554 := (or #28652 #28587 #28629 #28635 #28641)
+#29541 := [def-axiom]: #29554
+#29703 := [unit-resolution #29541 #29702 #29690 #29582 #29569 #29570]: false
+#29705 := [lemma #29703]: #29704
+#29110 := [unit-resolution #29705 #28973 #29837 #29570]: false
+#28980 := [lemma #29110]: #28981
+#30221 := [unit-resolution #28980 #29510]: #28652
+#28160 := (or #28655 #28647)
+#28159 := [def-axiom]: #28160
+#30222 := [unit-resolution #28159 #30221]: #28655
+#29319 := (or #28664 #28658)
+#29325 := (or #12349 #12967)
+#23835 := (or #22641 #12913)
+#23836 := [def-axiom]: #23835
+#29317 := [unit-resolution #23836 #28122]: #12913
+#29318 := (or #12349 #21575 #12967)
+#27702 := (or #12349 #21575 #26481 #12967)
+#27645 := [hypothesis]: #12913
+#27641 := (<= #27640 0::int)
+#27483 := (= uf_298 #27482)
+#27486 := (or #26376 #27483)
+#27487 := [quant-inst]: #27486
+#27646 := [unit-resolution #27487 #16290]: #27483
+#27649 := (not #27483)
+#27650 := (or #27649 #27641)
+#27651 := [th-lemma]: #27650
+#27652 := [unit-resolution #27651 #27646]: #27641
+#27672 := (iff #17724 #27561)
+#27670 := (iff #12349 #27560)
+#27668 := (iff #27560 #12349)
+#27666 := (= #27559 #3223)
+#27658 := [symm #27646]: #27657
+#27665 := [monotonicity #27663 #27658 #27656]: #27664
+#27667 := [monotonicity #27665]: #27666
+#27669 := [monotonicity #27667]: #27668
+#27671 := [symm #27669]: #27670
+#27673 := [monotonicity #27671]: #27672
+#27653 := [hypothesis]: #17724
+#27674 := [mp #27653 #27673]: #27561
+#27593 := (or #27565 #27560)
+#27594 := [def-axiom]: #27593
+#27675 := [unit-resolution #27594 #27674]: #27565
+#27681 := [hypothesis]: #12968
+#27682 := (or #27649 #27642)
+#27683 := [th-lemma]: #27682
+#27684 := [unit-resolution #27683 #27646]: #27642
+#27690 := [unit-resolution #27689 #27684 #27681 #27680]: #27685
+#27698 := (or #27504 #27514 #27568)
+#27699 := [unit-resolution #27589 #20766 #13537 #27697 #27696]: #27698
+#27700 := [unit-resolution #27699 #27690 #27675]: #27504
+#27701 := [th-lemma #27700 #27652 #27645]: false
+#27703 := [lemma #27701]: #27702
+#29322 := [unit-resolution #27703 #28088]: #29318
+#29326 := [unit-resolution #29322 #29317]: #29325
+#29327 := [unit-resolution #29326 #28538]: #12349
+#28201 := (or #28664 #17724 #28658)
+#28206 := [def-axiom]: #28201
+#29320 := [unit-resolution #28206 #29327]: #29319
+#30225 := [unit-resolution #29320 #30222]: #28664
+#29321 := (or #12355 #28661)
+#28667 := (iff #12355 #28664)
+#29514 := (not #28667)
+#29693 := [hypothesis]: #29514
+#29473 := (or #27590 #28667)
+#28583 := (or #28582 #28580)
+#28584 := (not #28583)
+#28594 := (or #28593 #28590)
+#28595 := (not #28594)
+#28597 := (= #28596 uf_14)
+#28598 := (not #28597)
+#28599 := (or #28598 #28595 #28587 #28584)
+#28600 := (not #28599)
+#28605 := (or #28604 #28602)
+#28606 := (not #28605)
+#28607 := (or #28597 #28606)
+#28608 := (not #28607)
+#28609 := (or #28608 #28600)
+#28610 := (not #28609)
+#28611 := (or #17724 #28610)
+#28612 := (not #28611)
+#28613 := (iff #12355 #28612)
+#29474 := (or #27590 #28613)
+#29531 := (iff #29474 #29473)
+#29536 := (iff #29473 #29473)
+#29509 := [rewrite]: #29536
+#28668 := (iff #28613 #28667)
+#28665 := (iff #28612 #28664)
+#28662 := (iff #28611 #28661)
+#28659 := (iff #28610 #28658)
+#28656 := (iff #28609 #28655)
+#28653 := (iff #28600 #28652)
+#28650 := (iff #28599 #28647)
+#28644 := (or #28629 #28635 #28587 #28641)
+#28648 := (iff #28644 #28647)
+#28649 := [rewrite]: #28648
+#28645 := (iff #28599 #28644)
+#28642 := (iff #28584 #28641)
+#28639 := (iff #28583 #28638)
+#28640 := [rewrite]: #28639
+#28643 := [monotonicity #28640]: #28642
+#28636 := (iff #28595 #28635)
+#28633 := (iff #28594 #28632)
+#28634 := [rewrite]: #28633
+#28637 := [monotonicity #28634]: #28636
+#28630 := (iff #28598 #28629)
+#28615 := (iff #28597 #28614)
+#28616 := [rewrite]: #28615
+#28631 := [monotonicity #28616]: #28630
+#28646 := [monotonicity #28631 #28637 #28643]: #28645
+#28651 := [trans #28646 #28649]: #28650
+#28654 := [monotonicity #28651]: #28653
+#28627 := (iff #28608 #28626)
+#28624 := (iff #28607 #28623)
+#28621 := (iff #28606 #28620)
+#28618 := (iff #28605 #28617)
+#28619 := [rewrite]: #28618
+#28622 := [monotonicity #28619]: #28621
+#28625 := [monotonicity #28616 #28622]: #28624
+#28628 := [monotonicity #28625]: #28627
+#28657 := [monotonicity #28628 #28654]: #28656
+#28660 := [monotonicity #28657]: #28659
+#28663 := [monotonicity #28660]: #28662
+#28666 := [monotonicity #28663]: #28665
+#28669 := [monotonicity #28666]: #28668
+#29535 := [monotonicity #28669]: #29531
+#29520 := [trans #29535 #29509]: #29531
+#29472 := [quant-inst]: #29474
+#29508 := [mp #29472 #29520]: #29473
+#29694 := [unit-resolution #29508 #21176 #29693]: false
+#29580 := [lemma #29694]: #28667
+#28208 := (or #29514 #12355 #28661)
+#28211 := [def-axiom]: #28208
+#29345 := [unit-resolution #28211 #29580]: #29321
+#30226 := [unit-resolution #29345 #30225]: #12355
+#29853 := (or #17730 #22574)
+#23760 := (or #22595 #22589)
+#23761 := [def-axiom]: #23760
+#29386 := [unit-resolution #23761 #28988]: #22589
+#29389 := (or #22592 #17724 #22586)
+#23752 := (or #22592 #17721 #17724 #22586)
+#23753 := [def-axiom]: #23752
+#29181 := [unit-resolution #23753 #27738]: #29389
+#29213 := [unit-resolution #29181 #29386 #29327]: #22586
+#23742 := (or #22583 #22577)
+#23743 := [def-axiom]: #23742
+#29850 := [unit-resolution #23743 #29213]: #22577
+#29851 := (or #22580 #17730 #22574)
+#23736 := (or #22580 #17721 #17730 #22574)
+#23737 := [def-axiom]: #23736
+#29852 := [unit-resolution #23737 #27738]: #29851
+#29854 := [unit-resolution #29852 #29850]: #29853
+#30229 := [unit-resolution #29854 #30226]: #22574
+#30232 := (or #22571 #22532)
+#28690 := [unit-resolution #29740 #20106 #28525 #27869 #29448]: #27522
+#28541 := [unit-resolution #29773 #28690]: #27495
+#29313 := [unit-resolution #29705 #28973 #28541]: #28652
+#29314 := [unit-resolution #28159 #29313]: #28655
+#29323 := [unit-resolution #29320 #29314]: #28664
+#29346 := [unit-resolution #29345 #29323]: #12355
+#29264 := [unit-resolution #29854 #29346]: #22574
+#29855 := [hypothesis]: #22526
+#23652 := (or #22529 #22523)
+#23653 := [def-axiom]: #23652
+#29856 := [unit-resolution #23653 #29855]: #22529
+#23708 := (or #22559 #22523)
+#23709 := [def-axiom]: #23708
+#29857 := [unit-resolution #23709 #29855]: #22559
+#29858 := (or #22571 #22562 #22532)
+#28575 := (or #22571 #17724 #22562 #22532)
+#28560 := [hypothesis]: #12349
+#28561 := [hypothesis]: #22529
+#28562 := [hypothesis]: #22574
+#23720 := (or #22571 #12355)
+#23721 := [def-axiom]: #23720
+#28563 := [unit-resolution #23721 #28562]: #12355
+#28564 := (or #22538 #17730 #22532)
+#23662 := (or #22538 #17721 #17730 #22532)
+#23663 := [def-axiom]: #23662
+#28565 := [unit-resolution #23663 #27738]: #28564
+#28566 := [unit-resolution #28565 #28563 #28561]: #22538
+#23668 := (or #22541 #22535)
+#23669 := [def-axiom]: #23668
+#28567 := [unit-resolution #23669 #28566]: #22541
+#28568 := [hypothesis]: #22559
+#23726 := (or #22571 #22565)
+#23727 := [def-axiom]: #23726
+#28569 := [unit-resolution #23727 #28562]: #22565
+#23716 := (or #22568 #22556 #22562)
+#23717 := [def-axiom]: #23716
+#28570 := [unit-resolution #23717 #28569 #28568]: #22556
+#23686 := (or #22553 #22547)
+#23687 := [def-axiom]: #23686
+#28571 := [unit-resolution #23687 #28570]: #22547
+#28572 := (or #22550 #17724 #22544)
+#23678 := (or #22550 #17721 #17724 #22544)
+#23679 := [def-axiom]: #23678
+#28573 := [unit-resolution #23679 #27738]: #28572
+#28574 := [unit-resolution #28573 #28571 #28567 #28560]: false
+#28576 := [lemma #28574]: #28575
+#29859 := [unit-resolution #28576 #29327]: #29858
+#29993 := [unit-resolution #29859 #29857 #29856 #29264]: false
+#29994 := [lemma #29993]: #22523
+#30143 := (or #22526 #22520)
+#28967 := (or #15127 #12967)
+#14990 := (<= uf_286 4294967295::int)
+#14993 := (iff #12726 #14990)
+#14984 := (+ 4294967295::int #12727)
+#14987 := (>= #14984 0::int)
+#14991 := (iff #14987 #14990)
+#14992 := [rewrite]: #14991
+#14988 := (iff #12726 #14987)
+#14985 := (= #12728 #14984)
+#14986 := [monotonicity #7481]: #14985
+#14989 := [monotonicity #14986]: #14988
+#14994 := [trans #14989 #14992]: #14993
+#13540 := [and-elim #13524]: #12726
+#14995 := [mp #13540 #14994]: #14990
+#28956 := [hypothesis]: #17745
+#28957 := [th-lemma #27681 #28956 #14995]: false
+#28968 := [lemma #28957]: #28967
+#30138 := [unit-resolution #28968 #28538]: #15127
+#30141 := (or #22526 #17745 #22520)
+#30136 := (or #21575 #13051)
+#30139 := [th-lemma]: #30136
+#30140 := [unit-resolution #30139 #29317]: #13051
+#23632 := (or #22526 #17742 #17745 #22520)
+#23633 := [def-axiom]: #23632
+#30142 := [unit-resolution #23633 #30140]: #30141
+#30144 := [unit-resolution #30142 #30138]: #30143
+#30150 := [unit-resolution #30144 #29994]: #22520
+#23622 := (or #22517 #22511)
+#23623 := [def-axiom]: #23622
+#30127 := [unit-resolution #23623 #30150]: #22511
+#23568 := (>= #13063 -1::int)
+#23618 := (or #22517 #13061)
+#23619 := [def-axiom]: #23618
+#30124 := [unit-resolution #23619 #30150]: #13061
+#30129 := (or #21556 #23568)
+#30128 := [th-lemma]: #30129
+#30130 := [unit-resolution #30128 #30124]: #23568
+#29095 := (not #23568)
+#30131 := (or #13089 #29095)
+#29096 := (or #13089 #12967 #29095)
+#29081 := [hypothesis]: #23568
+#29082 := [hypothesis]: #13093
+#29094 := [th-lemma #29082 #27681 #29081]: false
+#29097 := [lemma #29094]: #29096
+#30132 := [unit-resolution #29097 #28538]: #30131
+#30133 := [unit-resolution #30132 #30130]: #13089
+#23608 := (or #22514 #13093 #22508)
+#23609 := [def-axiom]: #23608
+#30134 := [unit-resolution #23609 #30133 #30127]: #22508
+#23600 := (or #22505 #22499)
+#23601 := [def-axiom]: #23600
+#30135 := [unit-resolution #23601 #30134]: #22499
+#23693 := (+ uf_297 #13126)
+#23695 := (>= #23693 0::int)
+#30148 := [hypothesis]: #22562
+#23698 := (or #22559 #12524)
+#23699 := [def-axiom]: #23698
+#30154 := [unit-resolution #23699 #30148]: #12524
+#30155 := (or #21614 #23695)
+#30153 := [th-lemma]: #30155
+#30156 := [unit-resolution #30153 #30154]: #23695
+#30164 := (not #23695)
+#30174 := (or #13129 #30164)
+#30165 := (or #13129 #13358 #30164)
+#30173 := [th-lemma]: #30165
+#30146 := [unit-resolution #30173 #28859]: #30174
+#30163 := [unit-resolution #30146 #30156]: #13129
+#30161 := (= #3101 #3272)
+#30159 := (= #3272 #3101)
+#30157 := (= #3271 #3100)
+#30149 := (= #28166 uf_297)
+#30166 := [symm #29826]: #30149
+#30147 := [trans #28479 #30166]: #27364
+#30152 := [monotonicity #30147]: #27366
+#30172 := (= #3271 #27249)
+#30170 := (= uf_305 #26870)
+#30168 := (= uf_305 #28166)
+#30167 := [symm #30154]: #3288
+#30169 := [trans #30167 #29826]: #30168
+#30171 := [trans #30169 #28499]: #30170
+#30180 := [monotonicity #30171]: #30172
+#30151 := [trans #30180 #30152]: #30157
+#30160 := [monotonicity #30151]: #30159
+#30162 := [symm #30160]: #30161
+#30194 := (= uf_304 #3101)
+#23696 := (or #22559 #12521)
+#23697 := [def-axiom]: #23696
+#30145 := [unit-resolution #23697 #30148]: #12521
+#30158 := [symm #30145]: #3287
+#30195 := [trans #30158 #28875]: #30194
+#30190 := [trans #30195 #30162]: #12446
+#23573 := (or #21468 #21466 #13128)
+#23574 := [def-axiom]: #23573
+#30196 := [unit-resolution #23574 #30190 #30163]: #21468
+#23575 := (or #22493 #21467)
+#23576 := [def-axiom]: #23575
+#30197 := [unit-resolution #23576 #30196]: #22493
+#23596 := (or #22502 #22496 #21531)
+#23597 := [def-axiom]: #23596
+#30176 := [unit-resolution #23597 #30197 #30135]: #21531
+#23588 := (or #21526 #18027)
+#23589 := [def-axiom]: #23588
+#30178 := [unit-resolution #23589 #30176]: #18027
+#29153 := (+ uf_298 #18020)
+#30137 := (>= #29153 0::int)
+#30208 := (not #30137)
+#30288 := (not #30181)
+#30175 := (= #3230 #17765)
+#30256 := (not #30175)
+#30251 := (+ #3230 #18007)
+#30253 := (>= #30251 0::int)
+#30263 := (not #30253)
+#23704 := (or #22559 #12997)
+#23705 := [def-axiom]: #23704
+#30179 := [unit-resolution #23705 #30148]: #12997
+#23690 := (+ uf_296 #13112)
+#23691 := (<= #23690 0::int)
+#30199 := (or #21613 #23691)
+#30200 := [th-lemma]: #30199
+#30182 := [unit-resolution #30200 #30145]: #23691
+#23585 := (not #18009)
+#23586 := (or #21526 #23585)
+#23587 := [def-axiom]: #23586
+#30183 := [unit-resolution #23587 #30176]: #23585
+#30264 := (not #23691)
+#30265 := (or #30263 #18009 #30264 #12996)
+#30258 := [hypothesis]: #12997
+#30259 := [hypothesis]: #23691
+#30260 := [hypothesis]: #23585
+#30261 := [hypothesis]: #30253
+#30262 := [th-lemma #30261 #30260 #30259 #30258]: false
+#30266 := [lemma #30262]: #30265
+#30184 := [unit-resolution #30266 #30183 #30182 #30179]: #30263
+#30257 := (or #30256 #30253)
+#30267 := [th-lemma]: #30257
+#30185 := [unit-resolution #30267 #30184]: #30256
+#30289 := (or #30288 #30175)
+#30284 := (= #17765 #3230)
+#30282 := (= #17764 #3220)
+#30280 := (= #17764 #27555)
+#30276 := (= ?x785!14 #27482)
+#30274 := (= ?x785!14 #28213)
+#30272 := (= ?x785!14 uf_298)
+#30271 := [hypothesis]: #30181
+#30273 := [symm #30271]: #30272
+#30275 := [trans #30273 #28873]: #30274
+#30277 := [trans #30275 #29458]: #30276
+#30278 := (= #2979 #26469)
+#30279 := [symm #27663]: #30278
+#30281 := [monotonicity #30279 #30277 #27654]: #30280
+#30283 := [trans #30281 #29712]: #30282
+#30285 := [monotonicity #30283]: #30284
+#30286 := [symm #30285]: #30175
+#30250 := [hypothesis]: #30256
+#30287 := [unit-resolution #30250 #30286]: false
+#30290 := [lemma #30287]: #30289
+#30186 := [unit-resolution #30290 #30185]: #30288
+#30213 := (or #30181 #30208)
+#29154 := (<= #29153 0::int)
+#29165 := (+ uf_296 #18007)
+#29166 := (>= #29165 0::int)
+#30187 := (not #29166)
+#30177 := (or #30187 #18009 #30264)
+#30207 := [th-lemma]: #30177
+#30189 := [unit-resolution #30207 #30183 #30182]: #30187
+#23583 := (or #21526 #17761)
+#23584 := [def-axiom]: #23583
+#30191 := [unit-resolution #23584 #30176]: #17761
+#23581 := (or #21526 #17757)
+#23582 := [def-axiom]: #23581
+#30192 := [unit-resolution #23582 #30176]: #17757
+#30050 := (or #22482 #21510 #21511 #29154 #29166)
+#29142 := (+ #17765 #13378)
+#29143 := (<= #29142 0::int)
+#29144 := (+ ?x785!14 #12965)
+#29145 := (>= #29144 0::int)
+#29146 := (or #21511 #29145 #29143 #21510)
+#30051 := (or #22482 #29146)
+#30058 := (iff #30051 #30050)
+#29174 := (or #21510 #21511 #29154 #29166)
+#30053 := (or #22482 #29174)
+#30056 := (iff #30053 #30050)
+#30057 := [rewrite]: #30056
+#30054 := (iff #30051 #30053)
+#29177 := (iff #29146 #29174)
+#29171 := (or #21511 #29154 #29166 #21510)
+#29175 := (iff #29171 #29174)
+#29176 := [rewrite]: #29175
+#29172 := (iff #29146 #29171)
+#29169 := (iff #29143 #29166)
+#29159 := (+ #13378 #17765)
+#29162 := (<= #29159 0::int)
+#29167 := (iff #29162 #29166)
+#29168 := [rewrite]: #29167
+#29163 := (iff #29143 #29162)
+#29160 := (= #29142 #29159)
+#29161 := [rewrite]: #29160
+#29164 := [monotonicity #29161]: #29163
+#29170 := [trans #29164 #29168]: #29169
+#29157 := (iff #29145 #29154)
+#29147 := (+ #12965 ?x785!14)
+#29150 := (>= #29147 0::int)
+#29155 := (iff #29150 #29154)
+#29156 := [rewrite]: #29155
+#29151 := (iff #29145 #29150)
+#29148 := (= #29144 #29147)
+#29149 := [rewrite]: #29148
+#29152 := [monotonicity #29149]: #29151
+#29158 := [trans #29152 #29156]: #29157
+#29173 := [monotonicity #29158 #29170]: #29172
+#29178 := [trans #29173 #29176]: #29177
+#30055 := [monotonicity #29178]: #30054
+#30059 := [trans #30055 #30057]: #30058
+#30052 := [quant-inst]: #30051
+#30060 := [mp #30052 #30059]: #30050
+#30193 := [unit-resolution #30060 #28540 #30192 #30191 #30189]: #29154
+#30210 := (not #29154)
+#30211 := (or #30181 #30210 #30208)
+#30212 := [th-lemma]: #30211
+#30214 := [unit-resolution #30212 #30193]: #30213
+#30215 := [unit-resolution #30214 #30186]: #30208
+#30216 := [th-lemma #30130 #30215 #30178]: false
+#30219 := [lemma #30216]: #22559
+#30241 := [unit-resolution #29859 #30219]: #30232
+#30236 := [unit-resolution #30241 #30229]: #22532
+#23646 := (or #22529 #12383)
+#23647 := [def-axiom]: #23646
+#30293 := [unit-resolution #23647 #30236]: #12383
+#28986 := [symm #30293]: #3242
+#29080 := [trans #28986 #28873]: #28951
+#29269 := [trans #29080 #29458]: #29074
+#29334 := [monotonicity #30279 #29269 #27654]: #29208
+#29357 := [trans #29334 #29712]: #29112
+#29125 := [monotonicity #29357]: #28979
+#28987 := [symm #29125]: #28730
+#29079 := (= uf_304 #3230)
+#23642 := (or #22529 #12375)
+#23643 := [def-axiom]: #23642
+#28983 := [unit-resolution #23643 #30236]: #12375
+#29211 := [symm #28983]: #3235
+#23644 := (or #22529 #12380)
+#23645 := [def-axiom]: #23644
+#29034 := [unit-resolution #23645 #30236]: #12380
+#28946 := [symm #29034]: #3240
+#28982 := [trans #28946 #29211]: #29079
+#29212 := [trans #28982 #28987]: #12446
+#28731 := [hypothesis]: #21466
+#29304 := [unit-resolution #28731 #29212]: false
+#28976 := [lemma #29304]: #12446
+#30304 := (or #21466 #21468)
+#28528 := (+ uf_298 #13126)
+#28529 := (>= #28528 0::int)
+#30294 := (or #21574 #28529)
+#30295 := [th-lemma]: #30294
+#30296 := [unit-resolution #30295 #30293]: #28529
+#30297 := [hypothesis]: #21467
+#30298 := [hypothesis]: #12446
+#30299 := [unit-resolution #23574 #30298 #30297]: #13128
+#30300 := (not #28529)
+#30301 := (or #13129 #12967 #30300)
+#30302 := [th-lemma]: #30301
+#30303 := [unit-resolution #30302 #30299 #30296 #28538]: false
+#30305 := [lemma #30303]: #30304
+#29076 := [unit-resolution #30305 #28976]: #21468
+#29124 := [unit-resolution #23576 #29076]: #22493
+#28962 := [unit-resolution #23597 #30135]: #22499
+#29358 := [unit-resolution #28962 #29124]: #21531
+#23692 := (>= #23690 0::int)
+#30027 := (not #23692)
+#29111 := [unit-resolution #23727 #30229]: #22565
+#29289 := (or #22568 #22556)
+#29333 := [unit-resolution #23717 #30219]: #29289
+#29311 := [unit-resolution #29333 #29111]: #22556
+#23684 := (or #22553 #12996)
+#23685 := [def-axiom]: #23684
+#29113 := [unit-resolution #23685 #29311]: #12996
+#29780 := (+ #3230 #13112)
+#29809 := (<= #29780 0::int)
+#29719 := (= #3230 uf_304)
+#29564 := (or #22529 #29719)
+#29576 := [hypothesis]: #22532
+#29720 := [unit-resolution #23645 #29576]: #12380
+#29761 := [unit-resolution #23643 #29576]: #12375
+#29723 := [trans #29761 #29720]: #29719
+#30001 := (not #29719)
+#29829 := [hypothesis]: #30001
+#29786 := [unit-resolution #29829 #29723]: false
+#29787 := [lemma #29786]: #29564
+#30237 := [unit-resolution #29787 #30236]: #29719
+#30002 := (or #30001 #29809)
+#30003 := [th-lemma]: #30002
+#30238 := [unit-resolution #30003 #30237]: #29809
+#30026 := (not #29809)
+#30028 := (or #30026 #30027 #12997)
+#30023 := [hypothesis]: #29809
+#30024 := [hypothesis]: #12996
+#30025 := [hypothesis]: #23692
+#30020 := [th-lemma #30025 #30024 #30023]: false
+#30029 := [lemma #30020]: #30028
+#29339 := [unit-resolution #30029 #30238 #29113]: #30027
+#29268 := (or #23692 #23691)
+#29308 := [th-lemma]: #29268
+#29306 := [unit-resolution #29308 #29339]: #23691
+#29309 := [unit-resolution #23587 #29358]: #23585
+#29290 := [unit-resolution #30207 #29309 #29306]: #30187
+#30126 := (or #21526 #29166 #30181)
+#30089 := [hypothesis]: #30288
+#30114 := [hypothesis]: #21531
+#30115 := [unit-resolution #23589 #30114]: #18027
+#30110 := (or #30137 #18022 #29095)
+#30116 := [th-lemma]: #30110
+#30117 := [unit-resolution #30116 #30115 #30130]: #30137
+#30113 := [unit-resolution #30212 #30117 #30089]: #30210
+#30118 := [hypothesis]: #30187
+#30119 := [unit-resolution #23584 #30114]: #17761
+#30120 := [unit-resolution #23582 #30114]: #17757
+#30125 := [unit-resolution #30060 #28540 #30120 #30119 #30118 #30113]: false
+#30076 := [lemma #30125]: #30126
+#29305 := [unit-resolution #30076 #29290 #29358]: #30181
+#30233 := (or #30263 #18009)
+#30239 := [th-lemma #30261 #30260 #30238]: false
+#30234 := [lemma #30239]: #30233
+#28425 := [unit-resolution #30234 #29309]: #30263
+#29077 := [unit-resolution #30267 #28425]: #30256
+[unit-resolution #30290 #29077 #29305]: false
unsat
--- a/src/HOL/Boogie/Tools/boogie_commands.ML Fri Dec 11 15:06:12 2009 +0100
+++ b/src/HOL/Boogie/Tools/boogie_commands.ML Fri Dec 11 15:35:29 2009 +0100
@@ -12,7 +12,9 @@
structure Boogie_Commands: BOOGIE_COMMANDS =
struct
-fun boogie_load (verbose, base_name) thy =
+(* commands *)
+
+fun boogie_load (quiet, base_name) thy =
let
val path = Path.explode (base_name ^ ".b2i")
val _ = File.exists path orelse
@@ -20,36 +22,143 @@
val _ = Boogie_VCs.is_closed thy orelse
error ("Found the beginning of a new Boogie environment, " ^
"but another Boogie environment is still open.")
- in Boogie_Loader.load_b2i verbose path thy end
+ in Boogie_Loader.load_b2i (not quiet) path thy end
+
+
+datatype vc_opts = VC_full of bool | VC_only of string list
+
+fun get_vc thy vc_name =
+ (case Boogie_VCs.lookup thy vc_name of
+ SOME vc => vc
+ | NONE =>
+ (case AList.lookup (op =) (Boogie_VCs.state_of thy) vc_name of
+ SOME Boogie_VCs.Proved => error ("The verification condition " ^
+ quote vc_name ^ " has already been proved.")
+ | _ => error ("There is no verification condition " ^
+ quote vc_name ^ ".")))
+
+fun boogie_vc (vc_opts, vc_name) lthy =
+ let
+ val thy = ProofContext.theory_of lthy
+ val vc = get_vc thy vc_name
+
+ val (sks, ts) = split_list
+ (case vc_opts of
+ VC_full paths => [vc] |> paths ? maps Boogie_VCs.paths_of
+ | VC_only assertions => map_filter (Boogie_VCs.extract vc) assertions)
+
+ val discharge = fold (Boogie_VCs.discharge o pair vc_name)
+ fun after_qed [thms] = Local_Theory.theory (discharge (sks ~~ thms))
+ | after_qed _ = I
+ in
+ lthy
+ |> fold Variable.auto_fixes ts
+ |> Proof.theorem_i NONE after_qed [map (rpair []) ts]
+ end
+
+
+fun write_list head pp xs = Pretty.writeln (Pretty.big_list head (map pp xs))
fun boogie_status thy =
- let
- fun pretty_vc_name (name, _, proved) =
- let val s = if proved then "proved" else "not proved"
- in Pretty.str (name ^ " (" ^ s ^ ")") end
+ let
+ fun string_of_state Boogie_VCs.Proved = "proved"
+ | string_of_state Boogie_VCs.NotProved = "not proved"
+ | string_of_state Boogie_VCs.PartiallyProved = "partially proved"
+
+ fun pretty_vc_name (name, proved) =
+ Pretty.str (name ^ " (" ^ string_of_state proved ^ ")")
in
- Pretty.writeln (Pretty.big_list "Boogie verification conditions:"
- (map pretty_vc_name (Boogie_VCs.as_list thy)))
+ Boogie_VCs.state_of thy
+ |> write_list "Boogie verification conditions:" pretty_vc_name
+ end
+
+fun boogie_status_vc (full, vc_name) thy =
+ let
+ fun pretty tag s = s ^ " (" ^ tag ^ ")"
+
+ val (ps, us) = Boogie_VCs.state_of_vc thy vc_name
+ in
+ if full
+ then write_list ("Assertions of Boogie verification condition " ^
+ quote vc_name ^ ":") Pretty.str (sort fast_string_ord
+ (map (pretty "proved") ps @ map (pretty "not proved") us))
+ else write_list ("Unproved assertions of Boogie verification condition " ^
+ quote vc_name ^ ":") Pretty.str (sort fast_string_ord us)
end
-fun boogie_vc name lthy =
- (case Boogie_VCs.lookup (ProofContext.theory_of lthy) name of
- NONE => error ("There is no verification condition " ^ quote name ^ ".")
- | SOME vc =>
- let
- fun store thm = Boogie_VCs.discharge (name, thm)
- fun after_qed [[thm]] = Local_Theory.theory (store thm)
- | after_qed _ = I
- in
- lthy
- |> Variable.auto_fixes vc
- |> Proof.theorem_i NONE after_qed [[(vc, [])]]
- end)
+
+local
+ fun trying s = tracing ("Trying " ^ s ^ " ...")
+ fun success_on s = tracing ("Succeeded on " ^ s ^ ".")
+ fun failure_on s c = tracing ("Failed on " ^ s ^ c)
+
+ fun string_of_asserts goal = space_implode ", " (Boogie_VCs.names_of goal)
+
+ fun string_of_path (i, n) =
+ "path " ^ string_of_int i ^ " of " ^ string_of_int n
+
+ fun itemize_paths ps =
+ let val n = length ps
+ in fst (fold_map (fn p => fn i => (((i, n), p), i+1)) ps 1) end
+
+ fun par_map f = flat o Par_List.map f
+
+ fun divide f goal =
+ let val n = Boogie_VCs.size_of goal
+ in
+ if n = 1 then Boogie_VCs.names_of goal
+ else
+ let val (goal1, goal2) = Boogie_VCs.split_path (n div 2) goal
+ in par_map f [goal1, goal2] end
+ end
+
+ fun prove thy meth (_, goal) =
+ ProofContext.init thy
+ |> Proof.theorem_i NONE (K I) [[(goal, [])]]
+ |> Proof.apply meth
+ |> Seq.hd
+ |> Proof.global_done_proof
+in
+fun boogie_narrow_vc ((timeout, vc_name), meth) thy =
+ let
+ val tp = TimeLimit.timeLimit (Time.fromSeconds timeout) (prove thy meth)
+
+ fun try_goal (tag, split_tag) split goal = (trying tag;
+ (case try tp goal of
+ SOME _ => (success_on tag; [])
+ | NONE => (failure_on tag split_tag; split goal)))
+
+ fun group_goal goal =
+ try_goal (string_of_asserts goal,
+ if Boogie_VCs.size_of goal = 1 then "." else ", further splitting ...")
+ (divide group_goal) goal
+
+ fun path_goal p =
+ try_goal (string_of_path p, ", splitting into assertions ...")
+ (divide group_goal)
+
+ val full_goal =
+ try_goal ("full goal", ", splitting into paths ...")
+ (par_map (uncurry path_goal) o itemize_paths o Boogie_VCs.paths_of)
+
+ val unsolved = full_goal (get_vc thy vc_name)
+ in
+ if null unsolved
+ then writeln ("Completely solved Boogie verification condition " ^
+ quote vc_name ^ ".")
+ else write_list ("Unsolved assertions of Boogie verification condition " ^
+ quote vc_name ^ ":") Pretty.str (sort fast_string_ord unsolved)
+ end
+end
+
+
fun boogie_end thy =
let
- fun not_proved (name, _, proved) = if not proved then SOME name else NONE
- val unproved = map_filter not_proved (Boogie_VCs.as_list thy)
+ fun not_proved (_, Boogie_VCs.Proved) = NONE
+ | not_proved (name, _) = SOME name
+
+ val unproved = map_filter not_proved (Boogie_VCs.state_of thy)
in
if null unproved then Boogie_VCs.close thy
else error (Pretty.string_of (Pretty.big_list
@@ -57,26 +166,57 @@
(map Pretty.str unproved)))
end
+
+
+(* syntax and setup *)
+
+fun scan_val n f = Args.$$$ n -- Args.colon |-- f
+fun scan_arg f = Args.parens f
+fun scan_opt n = Scan.optional (scan_arg (Args.$$$ n) >> K true) false
+
+
val _ =
OuterSyntax.command "boogie_open"
"Open a new Boogie environment and load a Boogie-generated .b2i file."
OuterKeyword.thy_decl
- (Scan.optional (Args.parens (Args.$$$ "quiet") >> K false) true --
- OuterParse.name >> (Toplevel.theory o boogie_load))
+ (scan_opt "quiet" -- OuterParse.name >> (Toplevel.theory o boogie_load))
+
+
+val vc_name = OuterParse.name
+val vc_opts =
+ scan_arg (scan_val "only" (Scan.repeat1 OuterParse.name)) >> VC_only ||
+ scan_opt "paths" >> VC_full
+
+fun vc_cmd f = Toplevel.print o Toplevel.local_theory_to_proof NONE f
+
+val _ =
+ OuterSyntax.command "boogie_vc"
+ "Enter into proof mode for a specific verification condition."
+ OuterKeyword.thy_goal
+ (vc_opts -- vc_name >> (vc_cmd o boogie_vc))
+
+
+val default_timeout = 5
+
+val status_narrow_vc =
+ scan_arg (Args.$$$ "narrow" |--
+ Scan.optional (scan_val "timeout" OuterParse.nat) default_timeout) --
+ vc_name --
+ scan_arg (scan_val "try" Method.parse)
+
+val status_vc_opts = scan_opt "full"
+
+fun status_cmd f = Toplevel.no_timing o Toplevel.keep (fn state =>
+ f (Toplevel.theory_of state))
val _ =
OuterSyntax.improper_command "boogie_status"
"Show the name and state of all loaded verification conditions."
OuterKeyword.diag
- (Scan.succeed (Toplevel.no_timing o Toplevel.keep (fn state =>
- boogie_status (Toplevel.theory_of state))))
+ (status_narrow_vc >> (status_cmd o boogie_narrow_vc) ||
+ status_vc_opts -- vc_name >> (status_cmd o boogie_status_vc) ||
+ Scan.succeed (status_cmd boogie_status))
-val _ =
- OuterSyntax.command "boogie_vc"
- "Enter the proof mode for a specific verification condition."
- OuterKeyword.thy_goal
- (OuterParse.name >> (fn name => Toplevel.print o
- Toplevel.local_theory_to_proof NONE (boogie_vc name)))
val _ =
OuterSyntax.command "boogie_end"
@@ -84,6 +224,7 @@
OuterKeyword.thy_decl
(Scan.succeed (Toplevel.theory boogie_end))
+
val setup = Theory.at_end (fn thy =>
let
val _ = Boogie_VCs.is_closed thy
--- a/src/HOL/Boogie/Tools/boogie_loader.ML Fri Dec 11 15:06:12 2009 +0100
+++ b/src/HOL/Boogie/Tools/boogie_loader.ML Fri Dec 11 15:35:29 2009 +0100
@@ -33,7 +33,16 @@
val short_name =
translate_string (fn s => if Symbol.is_letdig s then s else "")
-fun label_name line col = "L_" ^ string_of_int line ^ "_" ^ string_of_int col
+(* these prefixes must be distinct: *)
+val var_prefix = "V_"
+val label_prefix = "L_"
+val position_prefix = "P_"
+
+val no_label_name = label_prefix ^ "unknown"
+fun label_name line col =
+ if line = 0 orelse col = 0 then no_label_name
+ else label_prefix ^ string_of_int line ^ "_" ^ string_of_int col
+
datatype attribute_value = StringValue of string | TermValue of term
@@ -99,33 +108,42 @@
fun maybe_builtin T =
let
fun const name = SOME (Const (name, T))
+ fun const2_abs name =
+ let
+ val U = Term.domain_type T
+ val T' = Term.range_type (Term.range_type T)
+ in
+ SOME (Abs (Name.uu, U, Abs (Name.uu, U,
+ Const (name, T) $ Bound 0 $ Bound 1)))
+ end
fun choose builtin =
(case builtin of
- "bvnot" => const @{const_name z3_bvnot}
- | "bvand" => const @{const_name z3_bvand}
- | "bvor" => const @{const_name z3_bvor}
- | "bvxor" => const @{const_name z3_bvxor}
- | "bvadd" => const @{const_name z3_bvadd}
- | "bvneg" => const @{const_name z3_bvneg}
- | "bvsub" => const @{const_name z3_bvsub}
- | "bvmul" => const @{const_name z3_bvmul}
- | "bvudiv" => const @{const_name z3_bvudiv}
- | "bvurem" => const @{const_name z3_bvurem}
- | "bvsdiv" => const @{const_name z3_bvsdiv}
- | "bvsrem" => const @{const_name z3_bvsrem}
- | "bvshl" => const @{const_name z3_bvshl}
- | "bvlshr" => const @{const_name z3_bvlshr}
- | "bvashr" => const @{const_name z3_bvashr}
- | "bvult" => const @{const_name z3_bvult}
- | "bvule" => const @{const_name z3_bvule}
- | "bvugt" => const @{const_name z3_bvugt}
- | "bvuge" => const @{const_name z3_bvuge}
- | "bvslt" => const @{const_name z3_bvslt}
- | "bvsle" => const @{const_name z3_bvsle}
- | "bvsgt" => const @{const_name z3_bvsgt}
- | "bvsge" => const @{const_name z3_bvsge}
- | "sign_extend" => const @{const_name z3_sign_extend}
+ "bvnot" => const @{const_name bitNOT}
+ | "bvand" => const @{const_name bitAND}
+ | "bvor" => const @{const_name bitOR}
+ | "bvxor" => const @{const_name bitXOR}
+ | "bvadd" => const @{const_name plus}
+ | "bvneg" => const @{const_name uminus}
+ | "bvsub" => const @{const_name minus}
+ | "bvmul" => const @{const_name times}
+ | "bvudiv" => const @{const_name div}
+ | "bvurem" => const @{const_name mod}
+ | "bvsdiv" => const @{const_name sdiv}
+ | "bvsrem" => const @{const_name srem}
+ | "bvshl" => const @{const_name bv_shl}
+ | "bvlshr" => const @{const_name bv_lshr}
+ | "bvashr" => const @{const_name bv_ashr}
+ | "bvult" => const @{const_name less}
+ | "bvule" => const @{const_name less_eq}
+ | "bvugt" => const2_abs @{const_name less}
+ | "bvuge" => const2_abs @{const_name less_eq}
+ | "bvslt" => const @{const_name word_sless}
+ | "bvsle" => const @{const_name word_sle}
+ | "bvsgt" => const2_abs @{const_name word_sless}
+ | "bvsge" => const2_abs @{const_name word_sle}
+ | "zero_extend" => const @{const_name ucast}
+ | "sign_extend" => const @{const_name scast}
| _ => NONE)
fun is_builtin att =
@@ -170,7 +188,7 @@
fun uniques fns fds =
let
- fun is_unique (name, (([], T), atts)) =
+ fun is_unique (name, (([], _), atts)) =
(case AList.lookup (op =) atts "unique" of
SOME _ => Symtab.lookup fds name
| NONE => NONE)
@@ -211,7 +229,7 @@
let
fun split (_, Used thm) (used, new) = (thm :: used, new)
| split (t, New name) (used, new) = (used, (name, t) :: new)
- | split (t, Unused thm) (used, new) =
+ | split (_, Unused thm) (used, new) =
(warning (Pretty.str_of
(Pretty.big_list "This background axiom has not been loaded:"
[Display.pretty_thm_global thy thm]));
@@ -233,8 +251,6 @@
|> log verbose "The following axioms were added:" (map fst new)
|> (fn thy' => log verbose "The following axioms already existed:"
(map (Display.string_of_thm_global thy') used) thy')
- |> Context.theory_map (fn context => fold Split_VC_SMT_Rules.add_thm
- (Boogie_Axioms.get (Context.proof_of context)) context)
end
end
@@ -260,11 +276,7 @@
|> indexed
in
fun add_vcs verbose vcs thy =
- let
- val vcs' =
- vcs
- |> burrow_fst rename
- |> map (apsnd HOLogic.mk_Trueprop)
+ let val vcs' = burrow_fst rename vcs
in
thy
|> Boogie_VCs.set vcs'
@@ -396,53 +408,70 @@
fun mk_trigger [] t = t
| mk_trigger ps t = @{term trigger} $ HOLogic.mk_list patternT ps $ t
- val label_kind =
- $$$ "pos" >> K @{term block_at} ||
- $$$ "neg" >> K @{term assert_at} ||
+ fun make_label (line, col) = Free (label_name line col, @{typ bool})
+ fun labelled_by kind pos t = kind $ make_label pos $ t
+ val label =
+ $$$ "pos" |-- num -- num >> (fn (pos as (line, col)) =>
+ if label_name line col = no_label_name then I
+ else labelled_by @{term block_at} pos) ||
+ $$$ "neg" |-- num -- num >> labelled_by @{term assert_at} ||
scan_fail "illegal label kind"
- fun mk_select (m, k) =
- let
- val mT = Term.fastype_of m and kT = Term.fastype_of k
- val vT = Term.range_type mT
- in Const (@{const_name boogie_select}, mT --> kT --> vT) $ m $ k end
-
fun mk_store ((m, k), v) =
let
val mT = Term.fastype_of m and kT = Term.fastype_of k
val vT = Term.fastype_of v
- in
- Const (@{const_name boogie_store}, mT --> kT --> vT --> mT) $ m $ k $ v
- end
+ in Const (@{const_name fun_upd}, mT --> kT --> vT --> mT) $ m $ k $ v end
fun mk_extract ((msb, lsb), t) =
let
val dT = Term.fastype_of t and rT = mk_wordT (msb - lsb)
val nT = @{typ nat}
val mk_nat_num = HOLogic.mk_number @{typ nat}
- in
- Const (@{const_name boogie_bv_extract}, [nT, nT, dT] ---> rT) $
- mk_nat_num msb $ mk_nat_num lsb $ t
- end
+ in Const (@{const_name slice}, [nT, dT] ---> rT) $ mk_nat_num lsb $ t end
fun mk_concat (t1, t2) =
let
val T1 = Term.fastype_of t1 and T2 = Term.fastype_of t2
val U = mk_wordT (dest_wordT T1 + dest_wordT T2)
- in Const (@{const_name boogie_bv_concat}, [T1, T2] ---> U) $ t1 $ t2 end
+ in Const (@{const_name word_cat}, [T1, T2] ---> U) $ t1 $ t2 end
+
+ fun unique_labels t =
+ let
+ fun names_of (@{term assert_at} $ Free (n, _) $ t) = cons n #> names_of t
+ | names_of (t $ u) = names_of t #> names_of u
+ | names_of (Abs (_, _, t)) = names_of t
+ | names_of _ = I
+ val nctxt = Name.make_context (duplicates (op =) (names_of t []))
- val var_name = str >> prefix "V"
- val dest_var_name = unprefix "V"
+ fun fresh (i, nctxt) = (position_prefix ^ string_of_int i, (i+1, nctxt))
+ fun renamed n (i, nctxt) =
+ yield_singleton Name.variants n nctxt ||> pair i
+ fun mk_label (name, t) = @{term assert_at} $ Free (name, @{typ bool}) $ t
+
+ fun unique t =
+ (case t of
+ @{term assert_at} $ Free (n, _) $ u =>
+ if n = no_label_name
+ then fresh ##>> unique u #>> mk_label
+ else renamed n ##>> unique u #>> mk_label
+ | u1 $ u2 => unique u1 ##>> unique u2 #>> (op $)
+ | Abs (n, T, u) => unique u #>> (fn u' => Abs (n, T, u'))
+ | _ => pair t)
+ in fst (unique t (1, nctxt)) end
+
+ val var_name = str >> prefix var_prefix
+ val dest_var_name = unprefix var_prefix
fun rename_variables t =
let
- fun make_context names = Name.make_context (duplicates (op =) names)
fun short_var_name n = short_name (dest_var_name n)
+ val all_names = Term.add_free_names t []
val (names, nctxt) =
- Term.add_free_names t []
+ all_names
|> map_filter (try (fn n => (n, short_var_name n)))
|> split_list
- ||>> (fn names => Name.variants names (make_context names))
+ ||>> (fn names => Name.variants names (Name.make_context all_names))
|>> Symtab.make o (op ~~)
fun rename_free n = the_default n (Symtab.lookup names n)
@@ -484,10 +513,7 @@
scan_count (attribute tds fds) i --
exp >> (fn (((vs, ps), _), t) =>
fold_rev (mk_quant q) vs (mk_trigger ps t))) ||
- scan_line "label" (label_kind -- num -- num) -- exp >>
- (fn (((l, line), col), t) =>
- if line = 0 orelse col = 0 then t
- else l $ Free (label_name line col, @{typ bool}) $ t) ||
+ scan_line "label" label -- exp >> (fn (mk, t) => mk t) ||
scan_line "int-num" num >> HOLogic.mk_number @{typ int} ||
binexp "<" (binop @{term "op < :: int => _"}) ||
binexp "<=" (binop @{term "op <= :: int => _"}) ||
@@ -499,8 +525,7 @@
binexp "/" (binop @{term boogie_div}) ||
binexp "%" (binop @{term boogie_mod}) ||
scan_line "select" num :|-- (fn arity =>
- exp -- (scan_count exp (arity - 1) >> HOLogic.mk_tuple) >>
- mk_select) ||
+ exp -- (scan_count exp (arity - 1) >> HOLogic.mk_tuple) >> (op $)) ||
scan_line "store" num :|-- (fn arity =>
exp -- (scan_count exp (arity - 2) >> HOLogic.mk_tuple) -- exp >>
mk_store) ||
@@ -509,7 +534,7 @@
scan_line "bv-extract" (num -- num) -- exp >> mk_extract ||
binexp "bv-concat" mk_concat ||
scan_fail "illegal expression") st
- in exp >> rename_variables end
+ in exp >> (rename_variables o unique_labels) end
and attribute tds fds =
let
@@ -548,8 +573,7 @@
fun vcs verbose tds fds = Scan.depend (fn thy =>
Scan.repeat (scan_line "vc" (str -- num) --
- (expr tds fds >> Sign.cert_term thy)) >>
- (fn vcs => ((), add_vcs verbose vcs thy)))
+ (expr tds fds)) >> (fn vcs => ((), add_vcs verbose vcs thy)))
fun parse verbose thy = Scan.pass thy
(type_decls verbose :|-- (fn tds =>
--- a/src/HOL/Boogie/Tools/boogie_split.ML Fri Dec 11 15:06:12 2009 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,174 +0,0 @@
-(* Title: HOL/Boogie/Tools/boogie_split.ML
- Author: Sascha Boehme, TU Muenchen
-
-Method to split Boogie-generate verification conditions and pass the splinters
-to user-supplied automated sub-tactics.
-*)
-
-signature BOOGIE_SPLIT =
-sig
- val add_sub_tactic: string * (Proof.context -> int -> tactic) -> theory ->
- theory
- val sub_tactic_names: theory -> string list
- val setup: theory -> theory
-end
-
-structure Boogie_Split: BOOGIE_SPLIT =
-struct
-
-(* sub-tactics store *)
-
-structure Sub_Tactics = Theory_Data
-(
- type T = (Proof.context -> int -> tactic) Symtab.table
- val empty = Symtab.empty
- val extend = I
- fun merge data = Symtab.merge (K true) data
-)
-
-fun add_sub_tactic tac = Sub_Tactics.map (Symtab.update_new tac)
-
-val sub_tactic_names = Symtab.keys o Sub_Tactics.get
-
-fun lookup_sub_tactic ctxt name =
- (case Symtab.lookup (Sub_Tactics.get (ProofContext.theory_of ctxt)) name of
- SOME tac => SOME (name, tac)
- | NONE => (warning ("Unknown sub-tactic: " ^ quote name); NONE))
-
-fun as_meta_eq eq = eq RS @{thm eq_reflection}
-
-fun full_implies_conv cv ct =
- (case try Thm.dest_implies ct of
- NONE => cv ct
- | SOME (cp, cc) => Drule.imp_cong_rule (cv cp) (full_implies_conv cv cc))
-
-fun sub_tactics_tac ctxt (verbose, sub_tac_names) case_name =
- let
- fun trace msg = if verbose then tracing msg else ()
- fun trying name = trace ("Trying tactic " ^ quote name ^ " on case " ^
- quote case_name ^ " ...")
- fun solved () = trace ("Case solved: " ^ quote case_name)
- fun failed () = trace ("Case remains unsolved: " ^ quote case_name)
-
- infix IF_UNSOLVED
- fun (tac1 IF_UNSOLVED tac2) i st = st |> (tac1 i THEN (fn st' =>
- let val j = i + Thm.nprems_of st' - Thm.nprems_of st
- in
- if i > j then (solved (); Tactical.all_tac st')
- else Seq.INTERVAL tac2 i j st'
- end))
-
- fun TRY_ALL [] _ st = (failed (); Tactical.no_tac st)
- | TRY_ALL ((name, tac) :: tacs) i st = (trying name;
- (TRY o tac ctxt IF_UNSOLVED TRY_ALL tacs) i st)
-
- val unfold_labels = Conv.try_conv (Conv.arg_conv
- (More_Conv.rewrs_conv (map as_meta_eq @{thms labels})))
- in
- CONVERSION (full_implies_conv unfold_labels)
- THEN' TRY_ALL (map_filter (lookup_sub_tactic ctxt) sub_tac_names)
- end
-
-
-(* case names *)
-
-fun case_name_of t =
- (case HOLogic.dest_Trueprop (Logic.strip_imp_concl t) of
- @{term assert_at} $ Free (n, _) $ _ => n
- | _ => raise TERM ("case_name_of", [t]))
-
-local
- fun make_case_name (i, t) =
- the_default ("B_" ^ string_of_int (i + 1)) (try case_name_of t)
-
- val assert_intro = Thm.symmetric (as_meta_eq @{thm assert_at_def})
- val l = Thm.dest_arg1 (Thm.rhs_of assert_intro)
- fun intro new =
- Conv.rewr_conv (Thm.instantiate ([], [(l, new)]) assert_intro)
-
- val assert_eq = @{lemma "assert_at x t == assert_at y t"
- by (simp add: assert_at_def)}
- val (x, y) = pairself Thm.dest_arg1 (Thm.dest_binop (Thm.cprop_of assert_eq))
- fun rename old new =
- Conv.rewr_conv (Thm.instantiate ([], [(x, old), (y, new)]) assert_eq)
-
- fun at_concl cv = Conv.concl_conv ~1 (Conv.arg_conv cv)
- fun make_label thy name = Thm.cterm_of thy (Free (name, @{typ bool}))
-
- fun rename_case thy new ct =
- (case try case_name_of (Thm.term_of ct) of
- NONE => at_concl (intro (make_label thy new))
- | SOME old => if new = old then Conv.all_conv
- else at_concl (rename (make_label thy old) (make_label thy new))) ct
-in
-fun unique_case_names thy st =
- let
- val names = map_index make_case_name (Thm.prems_of st)
- |> ` (duplicates (op =)) |-> Name.variant_list
- in ALLGOALS (fn i => CONVERSION (rename_case thy (nth names (i-1))) i) st end
-end
-
-
-(* splitting method *)
-
-local
- val intro_rules = [@{thm impI}, @{thm conjI}, @{thm TrueI}]
- val elim_rules = [@{thm conjE}, @{thm TrueE}]
- val split_tac = REPEAT_ALL_NEW
- (Tactic.resolve_tac intro_rules ORELSE' Tactic.eresolve_tac elim_rules)
-
- fun ALL_GOALS false tac = ALLGOALS tac
- | ALL_GOALS true tac = PARALLEL_GOALS (HEADGOAL tac)
-
- fun prep_tac ctxt ((parallel, verbose), subs) facts =
- HEADGOAL (Method.insert_tac facts)
- THEN HEADGOAL split_tac
- THEN unique_case_names (ProofContext.theory_of ctxt)
- THEN ALL_GOALS parallel (SUBGOAL (fn (t, i) =>
- TRY (sub_tactics_tac ctxt (verbose, subs) (case_name_of t) i)))
-
- val drop_assert_at = Conv.concl_conv ~1 (Conv.try_conv (Conv.arg_conv
- (Conv.rewr_conv (as_meta_eq @{thm assert_at_def}))))
-
- fun split_vc args ctxt = METHOD_CASES (fn facts =>
- prep_tac ctxt args facts #>
- Seq.maps (fn st =>
- st
- |> ALLGOALS (CONVERSION drop_assert_at)
- |> Seq.map (pair (map (rpair [] o case_name_of) (Thm.prems_of st)))) #>
- Seq.maps (fn (names, st) =>
- CASES
- (Rule_Cases.make_common false
- (ProofContext.theory_of ctxt, Thm.prop_of st) names)
- Tactical.all_tac st))
-
- val options =
- Args.parens (Args.$$$ "verbose") >> K (false, true) ||
- Args.parens (Args.$$$ "fast_verbose") >> K (true, true)
- val sub_tactics = Args.$$$ "try" -- Args.colon |-- Scan.repeat1 Args.name
-in
-val setup_split_vc = Method.setup @{binding split_vc}
- (Scan.lift (Scan.optional options (true, false) --
- Scan.optional sub_tactics []) >> split_vc)
- ("Splits a Boogie-generated verification conditions into smaller problems" ^
- " and tries to solve the splinters with the supplied sub-tactics.")
-end
-
-
-(* predefined sub-tactics *)
-
-fun fast_sub_tac ctxt = Classical.fast_tac (Classical.claset_of ctxt)
-fun simp_sub_tac ctxt =
- Simplifier.asm_full_simp_tac (Simplifier.simpset_of ctxt)
-fun smt_sub_tac ctxt = SMT_Solver.smt_tac ctxt (Split_VC_SMT_Rules.get ctxt)
-
-
-(* setup *)
-
-val setup =
- setup_split_vc #>
- add_sub_tactic ("fast", fast_sub_tac) #>
- add_sub_tactic ("simp", simp_sub_tac) #>
- add_sub_tactic ("smt", smt_sub_tac)
-
-end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Boogie/Tools/boogie_tactics.ML Fri Dec 11 15:35:29 2009 +0100
@@ -0,0 +1,86 @@
+(* Title: HOL/Boogie/Tools/boogie_tactics.ML
+ Author: Sascha Boehme, TU Muenchen
+
+Boogie tactics and Boogie methods.
+*)
+
+signature BOOGIE_TACTICS =
+sig
+ val boogie_tac: Proof.context -> thm list -> int -> tactic
+ val boogie_all_tac: Proof.context -> thm list -> tactic
+ val setup: theory -> theory
+end
+
+structure Boogie_Tactics: BOOGIE_TACTICS =
+struct
+
+fun as_meta_eq eq = eq RS @{thm eq_reflection}
+
+val assert_at_def = as_meta_eq @{thm assert_at_def}
+val block_at_def = as_meta_eq @{thm block_at_def}
+val label_eqs = [assert_at_def, block_at_def]
+
+fun unfold_labels_tac ctxt =
+ let val unfold = More_Conv.rewrs_conv label_eqs
+ in CONVERSION (More_Conv.top_sweep_conv (K unfold) ctxt) end
+
+fun boogie_tac ctxt rules =
+ unfold_labels_tac ctxt
+ THEN' SMT_Solver.smt_tac ctxt (Boogie_Axioms.get ctxt @ rules)
+
+fun boogie_all_tac ctxt rules =
+ PARALLEL_GOALS (HEADGOAL (boogie_tac ctxt rules))
+
+fun boogie_method all =
+ Scan.optional Attrib.thms [] >> (fn thms => fn ctxt => METHOD (fn facts =>
+ let val tac = if all then boogie_all_tac else HEADGOAL oo boogie_tac
+ in tac ctxt (thms @ facts) end))
+
+val setup_boogie = Method.setup @{binding boogie} (boogie_method false)
+ ("Applies an SMT solver to the current goal " ^
+ "using the current set of Boogie background axioms")
+
+val setup_boogie_all = Method.setup @{binding boogie_all} (boogie_method true)
+ ("Applies an SMT solver to all goals " ^
+ "using the current set of Boogie background axioms")
+
+
+local
+ fun prep_tac facts =
+ Method.insert_tac facts
+ THEN' REPEAT_ALL_NEW
+ (Tactic.resolve_tac [@{thm impI}, @{thm conjI}, @{thm TrueI}]
+ ORELSE' Tactic.eresolve_tac [@{thm conjE}, @{thm TrueE}])
+
+ val drop_assert_at = Conv.concl_conv ~1 (Conv.try_conv (Conv.arg_conv
+ (Conv.rewr_conv assert_at_def)))
+
+ fun case_name_of t =
+ (case HOLogic.dest_Trueprop (Logic.strip_imp_concl t) of
+ @{term assert_at} $ Free (n, _) $ _ => n
+ | _ => raise TERM ("case_name_of", [t]))
+
+ fun boogie_cases ctxt = METHOD_CASES (fn facts =>
+ ALLGOALS (prep_tac facts) #>
+ Seq.maps (fn st =>
+ st
+ |> ALLGOALS (CONVERSION drop_assert_at)
+ |> Seq.map (pair (map (rpair [] o case_name_of) (Thm.prems_of st)))) #>
+ Seq.maps (fn (names, st) =>
+ CASES
+ (Rule_Cases.make_common false
+ (ProofContext.theory_of ctxt, Thm.prop_of st) names)
+ Tactical.all_tac st))
+in
+val setup_boogie_cases = Method.setup @{binding boogie_cases}
+ (Scan.succeed boogie_cases)
+ "Prepares a set of Boogie assertions for case-based proofs"
+end
+
+
+val setup =
+ setup_boogie #>
+ setup_boogie_all #>
+ setup_boogie_cases
+
+end
--- a/src/HOL/Boogie/Tools/boogie_vcs.ML Fri Dec 11 15:06:12 2009 +0100
+++ b/src/HOL/Boogie/Tools/boogie_vcs.ML Fri Dec 11 15:35:29 2009 +0100
@@ -6,54 +6,378 @@
signature BOOGIE_VCS =
sig
+ datatype skel = Nil of string | Imp of skel | Con1 of skel | Con2 of skel |
+ Conj of skel * skel
+ val size_of: skel * 'a -> int
+ val names_of: skel * 'a -> string list
+ val paths_of: skel * term -> (skel * term) list
+ val split_path: int -> skel * term -> (skel * term) * (skel * term)
+ val extract: skel * term -> string -> (skel * term) option
+
val set: (string * term) list -> theory -> theory
- val lookup: theory -> string -> term option
- val discharge: string * thm -> theory -> theory
+ val lookup: theory -> string -> (skel * term) option
+ val discharge: string * (skel * thm) -> theory -> theory
val close: theory -> theory
val is_closed: theory -> bool
- val as_list: theory -> (string * term * bool) list
+
+ datatype state = Proved | NotProved | PartiallyProved
+ val state_of: theory -> (string * state) list
+ val state_of_vc: theory -> string -> string list * string list
end
structure Boogie_VCs: BOOGIE_VCS =
struct
+(* simplify VCs: drop True, fold premises into conjunctions *)
+
+local
+ val dest_eq = HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of
+
+ val true_rules = map dest_eq @{lemma
+ "(True & P) = P"
+ "(P & True) = P"
+ "(P --> True) = True"
+ "(True --> P) = P"
+ by simp_all}
+
+ val fold_prems =
+ dest_eq @{lemma "(P1 --> P2 --> Q) = ((P1 & P2) --> Q)" by fast}
+in
+fun simp_vc thy = Pattern.rewrite_term thy (true_rules @ [fold_prems]) []
+end
+
+
+(* VC skeleton: reflect the structure of implications and conjunctions *)
+
+datatype skel =
+ Nil of string | (* assertion with label name *)
+ Imp of skel | (* implication: only conclusion is relevant *)
+ Con1 of skel | (* only left conjunct *)
+ Con2 of skel | (* right left conjunct *)
+ Conj of skel * skel (* full conjunction *)
+
+val skel_of =
+ let
+ fun skel (@{term "op -->"} $ _ $ u) = Imp (skel u)
+ | skel (@{term "op &"} $ t $ u) = Conj (skel t, skel u)
+ | skel (@{term assert_at} $ Free (n, _) $ _) = Nil n
+ | skel t = raise TERM ("skel_of", [t])
+ in skel o HOLogic.dest_Trueprop end
+
+fun size (Nil _) = 1
+ | size (Imp sk) = size sk
+ | size (Con1 sk) = size sk
+ | size (Con2 sk) = size sk
+ | size (Conj (sk1, sk2)) = size sk1 + size sk2
+
+fun size_of (sk, _) = size sk
+
+fun add_names (Nil name) = cons name (* label names are unique! *)
+ | add_names (Imp sk) = add_names sk
+ | add_names (Con1 sk) = add_names sk
+ | add_names (Con2 sk) = add_names sk
+ | add_names (Conj (sk1, sk2)) = add_names sk1 #> add_names sk2
+
+fun names_of (sk, _) = add_names sk []
+
+
+fun make_app t u = t $ u
+fun dest_app (t $ u) = (t, u)
+ | dest_app t = raise TERM ("dest_app", [t])
+
+val make_prop = HOLogic.mk_Trueprop
+val dest_prop = HOLogic.dest_Trueprop
+
+val make_imp = curry HOLogic.mk_imp
+val dest_imp = HOLogic.dest_imp
+
+val make_conj = curry HOLogic.mk_conj
+fun dest_conj (@{term "op &"} $ t $ u) = (t, u)
+ | dest_conj t = raise TERM ("dest_conj", [t])
+
+
+fun app_both f g (x, y) = (f x, g y)
+
+fun under_imp f r (sk, t) =
+ let val (t1, t2) = dest_imp t
+ in f (sk, t2) |> r (app_both Imp (make_imp t1)) end
+
+fun split_conj f r1 r2 r (sk1, sk2, t) =
+ let val (t1, t2) = dest_conj t
+ in
+ f (sk2, t2)
+ |>> r1 (app_both (Conj o pair sk1) (make_conj t1))
+ ||> r2 (apfst Con2)
+ |> r
+ end
+
+val paths_of =
+ let
+ fun chop1 xs = (hd xs, tl xs)
+
+ fun split (sk, t) =
+ (case sk of
+ Nil _ => [(sk, t)]
+ | Imp sk' => under_imp split map (sk', t)
+ | Con1 sk1 => splt Con1 (sk1, t)
+ | Con2 sk2 => splt Con2 (sk2, t)
+ | Conj (sk1 as Nil _, sk2) =>
+ split_conj (chop1 o split) I map (op ::) (sk1, sk2, t)
+ | Conj (sk1, sk2) =>
+ let val (t1, t2) = dest_conj t
+ in splt Con1 (sk1, t1) @ splt Con2 (sk2, t2) end)
+ and splt c st = split st |> map (apfst c)
+
+ in map (apsnd make_prop) o split o apsnd dest_prop end
+
+fun split_path i =
+ let
+ fun split_at i (sk, t) =
+ (case sk of
+ Imp sk' => under_imp (split_at i) pairself (sk', t)
+ | Con1 sk1 => split_at i (sk1, t) |> pairself (apfst Con1)
+ | Con2 sk2 => split_at i (sk2, t) |> pairself (apfst Con2)
+ | Conj (sk1, sk2) =>
+ if i > 1
+ then split_conj (split_at (i-1)) I I I (sk1, sk2, t)
+ else app_both (pair (Con1 sk1)) (pair (Con2 sk2)) (dest_conj t)
+ | _ => raise TERM ("split_path: " ^ string_of_int i, [t]))
+ in pairself (apsnd make_prop) o split_at i o apsnd dest_prop end
+
+fun extract (sk, t) name =
+ let
+ fun is_in (Nil n) = (n = name, false)
+ | is_in (Imp sk) = is_in sk
+ | is_in (Con1 sk) = is_in sk
+ | is_in (Con2 sk) = is_in sk
+ | is_in (Conj (sk1, sk2)) =
+ if fst (is_in sk1) then (true, true) else is_in sk2 ||> K true
+
+ fun extr (sk, t) =
+ (case sk of
+ Nil _ => (sk, t)
+ | Imp sk' => under_imp extr I (sk', t)
+ | Con1 sk1 => extr (sk1, t) |> apfst Con1
+ | Con2 sk2 => extr (sk2, t) |> apfst Con2
+ | Conj (sk1, sk2) =>
+ (case is_in sk1 of
+ (true, true) => extr (sk1, fst (dest_conj t)) |> apfst Con1
+ | (true, false) => (Con1 sk1, fst (dest_conj t))
+ | (false, _) => extr (sk2, snd (dest_conj t)) |> apfst Con2))
+ in
+ (case is_in sk of
+ (true, true) => SOME ((sk, dest_prop t) |> extr |> apsnd make_prop)
+ | (true, false) => SOME (sk, t)
+ | (false, _) => NONE)
+ end
+
+
+fun cut thy =
+ let
+ fun opt_map_both f g = Option.map (app_both f g)
+
+ fun cut_err (u, t) = raise TERM ("cut: not matching", [u, t])
+
+ val matches = Pattern.matches thy
+
+ fun sub (usk, u) (sk, t) =
+ (case (usk, sk) of
+ (Nil _, Nil _) => if matches (u, t) then NONE else cut_err (u, t)
+ | (Imp usk', Imp sk') => sub_imp (usk', u) (sk', t)
+
+ | (Con1 usk1, Con1 sk1) =>
+ sub (usk1, u) (sk1, t) |> Option.map (apfst Con1)
+ | (Con2 usk2, Con2 sk2) =>
+ sub (usk2, u) (sk2, t) |> Option.map (apfst Con2)
+
+ | (Con1 _, Con2 _) => SOME (sk, t)
+ | (Con2 _, Con1 _) => SOME (sk, t)
+
+ | (Con1 usk1, Conj (sk1, sk2)) =>
+ let val (t1, t2) = dest_conj t
+ in
+ sub (usk1, u) (sk1, t1)
+ |> opt_map_both (Conj o rpair sk2) (fn t1' => make_conj t1' t2)
+ |> SOME o the_default (Con2 sk2, t2)
+ end
+ | (Con2 usk2, Conj (sk1, sk2)) =>
+ let val (t1, t2) = dest_conj t
+ in
+ sub (usk2, u) (sk2, t2)
+ |> opt_map_both (Conj o pair sk1) (make_conj t1)
+ |> SOME o the_default (Con1 sk1, t1)
+ end
+ | (Conj (usk1, _), Con1 sk1) =>
+ let val (u1, _) = dest_conj u
+ in sub (usk1, u1) (sk1, t) |> Option.map (apfst Con1) end
+ | (Conj (_, usk2), Con2 sk2) =>
+ let val (_, u2) = dest_conj u
+ in sub (usk2, u2) (sk2, t) |> Option.map (apfst Con2) end
+
+ | (Conj (usk1, usk2), Conj (sk1, sk2)) =>
+ sub_conj (usk1, usk2, u) (sk1, sk2, t)
+
+ | _ => cut_err (u, t))
+
+ and sub_imp (usk', u) (sk', t) =
+ let
+ val (u1, u2) = dest_imp u
+ val (t1, t2) = dest_imp t
+ in
+ if matches (u1, t1)
+ then sub (usk', u2) (sk', t2) |> opt_map_both Imp (make_imp t1)
+ else cut_err (u, t)
+ end
+
+ and sub_conj (usk1, usk2, u) (sk1, sk2, t) =
+ let
+ val (u1, u2) = dest_conj u
+ val (t1, t2) = dest_conj t
+ in
+ (case (sub (usk1, u1) (sk1, t1), sub (usk2, u2) (sk2, t2)) of
+ (NONE, NONE) => NONE
+ | (NONE, SOME (sk2', t2')) => SOME (Con2 sk2', t2')
+ | (SOME (sk1', t1'), NONE) => SOME (Con1 sk1', t1')
+ | (SOME (sk1', t1'), SOME (sk2', t2')) =>
+ SOME (Conj (sk1', sk2'), make_conj t1' t2'))
+ end
+ in
+ (fn su => fn st =>
+ (su, st)
+ |> pairself (apsnd dest_prop)
+ |> uncurry sub
+ |> Option.map (apsnd make_prop))
+ end
+
+
+local
+ fun imp f (lsk, lthm) (rsk, rthm) =
+ let
+ val mk_prop = Thm.capply @{cterm Trueprop}
+ val ct = Thm.cprop_of lthm |> Thm.dest_arg |> Thm.dest_arg1 |> mk_prop
+ val th = Thm.assume ct
+ fun imp_elim thm = @{thm mp} OF [thm, th]
+ fun imp_intr thm = Thm.implies_intr ct thm COMP_INCR @{thm impI}
+ val (sk, thm) = f (lsk, imp_elim lthm) (rsk, imp_elim rthm)
+ in (Imp sk, imp_intr thm) end
+
+ fun conj1 thm = @{thm conjunct1} OF [thm]
+ fun conj2 thm = @{thm conjunct2} OF [thm]
+in
+fun join (lsk, lthm) (rsk, rthm) =
+ (case (lsk, rsk) of
+ (Nil _, Nil _) => (lsk, lthm)
+ | (Imp lsk', Imp rsk') => imp join (lsk', lthm) (rsk', rthm)
+
+ | (Con1 lsk1, Con1 rsk1) => join (lsk1, lthm) (rsk1, rthm) |>> Con1
+ | (Con2 lsk2, Con2 rsk2) => join (lsk2, lthm) (rsk2, rthm) |>> Con2
+
+ | (Con1 lsk1, Con2 rsk2) => (Conj (lsk1, rsk2), @{thm conjI} OF [lthm, rthm])
+ | (Con2 lsk2, Con1 rsk1) => (Conj (rsk1, lsk2), @{thm conjI} OF [rthm, lthm])
+
+ | (Con1 lsk1, Conj (rsk1, rsk2)) =>
+ let val (sk1, thm) = join (lsk1, lthm) (rsk1, conj1 rthm)
+ in (Conj (sk1, rsk2), @{thm conjI} OF [thm, conj2 rthm]) end
+ | (Con2 lsk2, Conj (rsk1, rsk2)) =>
+ let val (sk2, thm) = join (lsk2, lthm) (rsk2, conj2 rthm)
+ in (Conj (rsk1, sk2), @{thm conjI} OF [conj1 rthm, thm]) end
+ | (Conj (lsk1, lsk2), Con1 rsk1) =>
+ let val (sk1, thm) = join (lsk1, conj1 lthm) (rsk1, rthm)
+ in (Conj (sk1, lsk2), @{thm conjI} OF [thm, conj2 lthm]) end
+ | (Conj (lsk1, lsk2), Con2 rsk2) =>
+ let val (sk2, thm) = join (lsk2, conj2 lthm) (rsk2, rthm)
+ in (Conj (lsk1, sk2), @{thm conjI} OF [conj1 lthm, thm]) end
+
+ | (Conj (lsk1, lsk2), Conj (rsk1, rsk2)) =>
+ let
+ val (sk1, th1) = join (lsk1, conj1 lthm) (rsk1, conj1 rthm)
+ val (sk2, th2) = join (lsk2, conj2 lthm) (rsk2, conj2 rthm)
+ in (Conj (sk1, sk2), @{thm conjI} OF [th1, th2]) end
+
+ | _ => raise THM ("join: different structure", 0, [lthm, rthm]))
+end
+
+
+(* the VC store *)
+
fun err_vcs () = error "undischarged Boogie verification conditions found"
+datatype vc =
+ VC_NotProved of skel * term |
+ VC_PartiallyProved of (skel * term) * (skel * thm) |
+ VC_Proved of skel * thm
+
+fun goal_of (VC_NotProved g) = SOME g
+ | goal_of (VC_PartiallyProved (g, _)) = SOME g
+ | goal_of (VC_Proved _) = NONE
+
+fun proof_of (VC_NotProved _) = NONE
+ | proof_of (VC_PartiallyProved (_, p)) = SOME p
+ | proof_of (VC_Proved p) = SOME p
+
structure VCs = Theory_Data
(
- type T = (term * bool) Symtab.table option
+ type T = vc Symtab.table option
val empty = NONE
val extend = I
fun merge (NONE, NONE) = NONE
| merge _ = err_vcs ()
)
-fun set vcs = VCs.map (fn
- NONE => SOME (Symtab.make (map (apsnd (rpair false)) vcs))
- | SOME _ => err_vcs ())
+fun prep thy t = VC_NotProved (` skel_of (simp_vc thy (HOLogic.mk_Trueprop t)))
+
+fun set vcs thy = VCs.map (fn
+ NONE => SOME (Symtab.make (map (apsnd (prep thy)) vcs))
+ | SOME _ => err_vcs ()) thy
-fun lookup thy name =
+fun lookup thy name =
(case VCs.get thy of
- SOME vcs => Option.map fst (Symtab.lookup vcs name)
+ SOME vcs =>
+ (case Symtab.lookup vcs name of
+ SOME vc => goal_of vc
+ | NONE => NONE)
| NONE => NONE)
-fun discharge (name, thm) thy = VCs.map (fn
- SOME vcs => SOME (Symtab.map_entry name (fn (t, proved) =>
- if proved then (t, proved)
- else (t, Pattern.matches thy (Thm.prop_of thm, t))) vcs)
- | NONE => NONE) thy
+fun discharge (name, prf) thy = thy |> VCs.map (Option.map
+ (Symtab.map_entry name (fn
+ VC_NotProved goal =>
+ (case cut thy (apsnd Thm.prop_of prf) goal of
+ SOME goal' => VC_PartiallyProved (goal', prf)
+ | NONE => VC_Proved prf)
+ | VC_PartiallyProved (goal, proof) =>
+ (case cut thy (apsnd Thm.prop_of prf) goal of
+ SOME goal' => VC_PartiallyProved (goal', join prf proof)
+ | NONE => VC_Proved (join prf proof))
+ | VC_Proved proof => VC_Proved proof)))
val close = VCs.map (fn
SOME vcs =>
- if Symtab.exists (fn (_, (_, proved)) => not proved) vcs then err_vcs ()
+ if Symtab.exists (is_some o goal_of o snd) vcs
+ then err_vcs ()
else NONE
| NONE => NONE)
val is_closed = is_none o VCs.get
-fun as_list thy =
+
+datatype state = Proved | NotProved | PartiallyProved
+
+fun state_of thy =
(case VCs.get thy of
- SOME vcs => map (fn (n, (t, p)) => (n, t, p)) (Symtab.dest vcs)
+ SOME vcs => Symtab.dest vcs |> map (apsnd (fn
+ VC_NotProved _ => NotProved
+ | VC_PartiallyProved _ => PartiallyProved
+ | VC_Proved _ => Proved))
| NONE => [])
+fun names_of' skx = these (Option.map names_of skx)
+
+fun state_of_vc thy name =
+ (case VCs.get thy of
+ SOME vcs =>
+ (case Symtab.lookup vcs name of
+ SOME vc => (names_of' (proof_of vc), names_of' (goal_of vc))
+ | NONE => ([], []))
+ | NONE => ([], []))
+
end