# HG changeset patch # User schirmer # Date 1258188027 -3600 # Node ID ac68c3ee4c2e5ebeb285130f94a75105e5064530 # Parent 3db22a5707f30fc69f9321f1598ab66c71a0caff# Parent cddea85bc87b8829ad32b2cff9c7fd9df2080779 merged diff -r 3db22a5707f3 -r ac68c3ee4c2e NEWS --- a/NEWS Sat Nov 14 09:31:54 2009 +0100 +++ b/NEWS Sat Nov 14 09:40:27 2009 +0100 @@ -241,6 +241,8 @@ sizechange -> size_change induct_scheme -> induction_schema +* Lemma name change: replaced "anti_sym" by "antisym" everywhere. + *** ML *** diff -r 3db22a5707f3 -r ac68c3ee4c2e doc-src/IsarImplementation/Thy/Local_Theory.thy --- a/doc-src/IsarImplementation/Thy/Local_Theory.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/doc-src/IsarImplementation/Thy/Local_Theory.thy Sat Nov 14 09:40:27 2009 +0100 @@ -97,13 +97,12 @@ text %mlref {* \begin{mldecls} @{index_ML_type local_theory: Proof.context} \\ - @{index_ML TheoryTarget.init: "string option -> theory -> local_theory"} \\[1ex] - @{index_ML LocalTheory.define: "string -> + @{index_ML Theory_Target.init: "string option -> theory -> local_theory"} \\[1ex] + @{index_ML Local_Theory.define: "string -> (binding * mixfix) * (Attrib.binding * term) -> local_theory -> (term * (string * thm)) * local_theory"} \\ - @{index_ML LocalTheory.note: "string -> - Attrib.binding * thm list -> local_theory -> - (string * thm list) * local_theory"} \\ + @{index_ML Local_Theory.note: "Attrib.binding * thm list -> + local_theory -> (string * thm list) * local_theory"} \\ \end{mldecls} \begin{description} @@ -116,7 +115,7 @@ with operations on expecting a regular @{text "ctxt:"}~@{ML_type Proof.context}. - \item @{ML TheoryTarget.init}~@{text "NONE thy"} initializes a + \item @{ML Theory_Target.init}~@{text "NONE thy"} initializes a trivial local theory from the given background theory. Alternatively, @{text "SOME name"} may be given to initialize a @{command locale} or @{command class} context (a fully-qualified @@ -124,7 +123,7 @@ --- normally the Isar toplevel already takes care to initialize the local theory context. - \item @{ML LocalTheory.define}~@{text "kind ((b, mx), (a, rhs)) + \item @{ML Local_Theory.define}~@{text "kind ((b, mx), (a, rhs)) lthy"} defines a local entity according to the specification that is given relatively to the current @{text "lthy"} context. In particular the term of the RHS may refer to earlier local entities @@ -145,13 +144,13 @@ @{attribute simplified} are better avoided. The @{text kind} determines the theorem kind tag of the resulting - fact. Typical examples are @{ML Thm.definitionK}, @{ML - Thm.theoremK}, or @{ML Thm.internalK}. + fact. Typical examples are @{ML Thm.definitionK} or @{ML + Thm.theoremK}. - \item @{ML LocalTheory.note}~@{text "kind (a, ths) lthy"} is - analogous to @{ML LocalTheory.define}, but defines facts instead of + \item @{ML Local_Theory.note}~@{text "(a, ths) lthy"} is + analogous to @{ML Local_Theory.define}, but defines facts instead of terms. There is also a slightly more general variant @{ML - LocalTheory.notes} that defines several facts (with attribute + Local_Theory.notes} that defines several facts (with attribute expressions) simultaneously. This is essentially the internal version of the @{command lemmas} diff -r 3db22a5707f3 -r ac68c3ee4c2e doc-src/IsarImplementation/Thy/document/Local_Theory.tex --- a/doc-src/IsarImplementation/Thy/document/Local_Theory.tex Sat Nov 14 09:31:54 2009 +0100 +++ b/doc-src/IsarImplementation/Thy/document/Local_Theory.tex Sat Nov 14 09:40:27 2009 +0100 @@ -123,13 +123,12 @@ \begin{isamarkuptext}% \begin{mldecls} \indexdef{}{ML type}{local\_theory}\verb|type local_theory = Proof.context| \\ - \indexdef{}{ML}{TheoryTarget.init}\verb|TheoryTarget.init: string option -> theory -> local_theory| \\[1ex] - \indexdef{}{ML}{LocalTheory.define}\verb|LocalTheory.define: string ->|\isasep\isanewline% + \indexdef{}{ML}{Theory\_Target.init}\verb|Theory_Target.init: string option -> theory -> local_theory| \\[1ex] + \indexdef{}{ML}{Local\_Theory.define}\verb|Local_Theory.define: string ->|\isasep\isanewline% \verb| (binding * mixfix) * (Attrib.binding * term) -> local_theory ->|\isasep\isanewline% \verb| (term * (string * thm)) * local_theory| \\ - \indexdef{}{ML}{LocalTheory.note}\verb|LocalTheory.note: string ->|\isasep\isanewline% -\verb| Attrib.binding * thm list -> local_theory ->|\isasep\isanewline% -\verb| (string * thm list) * local_theory| \\ + \indexdef{}{ML}{Local\_Theory.note}\verb|Local_Theory.note: Attrib.binding * thm list ->|\isasep\isanewline% +\verb| local_theory -> (string * thm list) * local_theory| \\ \end{mldecls} \begin{description} @@ -141,7 +140,7 @@ any value \isa{lthy{\isacharcolon}}~\verb|local_theory| can be also used with operations on expecting a regular \isa{ctxt{\isacharcolon}}~\verb|Proof.context|. - \item \verb|TheoryTarget.init|~\isa{NONE\ thy} initializes a + \item \verb|Theory_Target.init|~\isa{NONE\ thy} initializes a trivial local theory from the given background theory. Alternatively, \isa{SOME\ name} may be given to initialize a \hyperlink{command.locale}{\mbox{\isa{\isacommand{locale}}}} or \hyperlink{command.class}{\mbox{\isa{\isacommand{class}}}} context (a fully-qualified @@ -149,7 +148,7 @@ --- normally the Isar toplevel already takes care to initialize the local theory context. - \item \verb|LocalTheory.define|~\isa{kind\ {\isacharparenleft}{\isacharparenleft}b{\isacharcomma}\ mx{\isacharparenright}{\isacharcomma}\ {\isacharparenleft}a{\isacharcomma}\ rhs{\isacharparenright}{\isacharparenright}\ lthy} defines a local entity according to the specification that is + \item \verb|Local_Theory.define|~\isa{kind\ {\isacharparenleft}{\isacharparenleft}b{\isacharcomma}\ mx{\isacharparenright}{\isacharcomma}\ {\isacharparenleft}a{\isacharcomma}\ rhs{\isacharparenright}{\isacharparenright}\ lthy} defines a local entity according to the specification that is given relatively to the current \isa{lthy} context. In particular the term of the RHS may refer to earlier local entities from the auxiliary context, or hypothetical parameters from the @@ -169,11 +168,11 @@ \hyperlink{attribute.simplified}{\mbox{\isa{simplified}}} are better avoided. The \isa{kind} determines the theorem kind tag of the resulting - fact. Typical examples are \verb|Thm.definitionK|, \verb|Thm.theoremK|, or \verb|Thm.internalK|. + fact. Typical examples are \verb|Thm.definitionK| or \verb|Thm.theoremK|. - \item \verb|LocalTheory.note|~\isa{kind\ {\isacharparenleft}a{\isacharcomma}\ ths{\isacharparenright}\ lthy} is - analogous to \verb|LocalTheory.define|, but defines facts instead of - terms. There is also a slightly more general variant \verb|LocalTheory.notes| that defines several facts (with attribute + \item \verb|Local_Theory.note|~\isa{{\isacharparenleft}a{\isacharcomma}\ ths{\isacharparenright}\ lthy} is + analogous to \verb|Local_Theory.define|, but defines facts instead of + terms. There is also a slightly more general variant \verb|Local_Theory.notes| that defines several facts (with attribute expressions) simultaneously. This is essentially the internal version of the \hyperlink{command.lemmas}{\mbox{\isa{\isacommand{lemmas}}}} diff -r 3db22a5707f3 -r ac68c3ee4c2e src/FOL/ex/LocaleTest.thy --- a/src/FOL/ex/LocaleTest.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/FOL/ex/LocaleTest.thy Sat Nov 14 09:40:27 2009 +0100 @@ -195,7 +195,7 @@ begin thm lor_def -(* Can we get rid the the additional hypothesis, caused by LocalTheory.notes? *) +(* Can we get rid the the additional hypothesis, caused by Local_Theory.notes? *) lemma "x || y = --(-- x && --y)" by (unfold lor_def) (rule refl) diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Algebra/IntRing.thy --- a/src/HOL/Algebra/IntRing.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Algebra/IntRing.thy Sat Nov 14 09:40:27 2009 +0100 @@ -12,52 +12,11 @@ subsection {* Some properties of @{typ int} *} -lemma dvds_imp_abseq: - "\l dvd k; k dvd l\ \ abs l = abs (k::int)" -apply (subst abs_split, rule conjI) - apply (clarsimp, subst abs_split, rule conjI) - apply (clarsimp) - apply (cases "k=0", simp) - apply (cases "l=0", simp) - apply (simp add: zdvd_anti_sym) - apply clarsimp - apply (cases "k=0", simp) - apply (simp add: zdvd_anti_sym) -apply (clarsimp, subst abs_split, rule conjI) - apply (clarsimp) - apply (cases "l=0", simp) - apply (simp add: zdvd_anti_sym) -apply (clarsimp) -apply (subgoal_tac "-l = -k", simp) -apply (intro zdvd_anti_sym, simp+) -done - -lemma abseq_imp_dvd: - assumes a_lk: "abs l = abs (k::int)" - shows "l dvd k" -proof - - from a_lk - have "nat (abs l) = nat (abs k)" by simp - hence "nat (abs l) dvd nat (abs k)" by simp - hence "int (nat (abs l)) dvd k" by (subst int_dvd_iff) - hence "abs l dvd k" by simp - thus "l dvd k" - apply (unfold dvd_def, cases "l<0") - defer 1 apply clarsimp - proof (clarsimp) - fix k - assume l0: "l < 0" - have "- (l * k) = l * (-k)" by simp - thus "\ka. - (l * k) = l * ka" by fast - qed -qed - lemma dvds_eq_abseq: "(l dvd k \ k dvd l) = (abs l = abs (k::int))" apply rule - apply (simp add: dvds_imp_abseq) -apply (rule conjI) - apply (simp add: abseq_imp_dvd)+ + apply (simp add: zdvd_antisym_abs) +apply (simp add: dvd_if_abs_eq) done diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Algebra/Lattice.thy --- a/src/HOL/Algebra/Lattice.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Algebra/Lattice.thy Sat Nov 14 09:40:27 2009 +0100 @@ -18,7 +18,7 @@ locale weak_partial_order = equivalence L for L (structure) + assumes le_refl [intro, simp]: "x \ carrier L ==> x \ x" - and weak_le_anti_sym [intro]: + and weak_le_antisym [intro]: "[| x \ y; y \ x; x \ carrier L; y \ carrier L |] ==> x .= y" and le_trans [trans]: "[| x \ y; y \ z; x \ carrier L; y \ carrier L; z \ carrier L |] ==> x \ z" @@ -636,7 +636,7 @@ fix s assume sup: "least L s (Upper L {x, y, z})" show "x \ (y \ z) .= s" - proof (rule weak_le_anti_sym) + proof (rule weak_le_antisym) from sup L show "x \ (y \ z) \ s" by (fastsimp intro!: join_le elim: least_Upper_above) next @@ -877,7 +877,7 @@ fix i assume inf: "greatest L i (Lower L {x, y, z})" show "x \ (y \ z) .= i" - proof (rule weak_le_anti_sym) + proof (rule weak_le_antisym) from inf L show "i \ x \ (y \ z)" by (fastsimp intro!: meet_le elim: greatest_Lower_below) next @@ -1089,11 +1089,11 @@ assumes eq_is_equal: "op .= = op =" begin -declare weak_le_anti_sym [rule del] +declare weak_le_antisym [rule del] -lemma le_anti_sym [intro]: +lemma le_antisym [intro]: "[| x \ y; y \ x; x \ carrier L; y \ carrier L |] ==> x = y" - using weak_le_anti_sym unfolding eq_is_equal . + using weak_le_antisym unfolding eq_is_equal . lemma lless_eq: "x \ y \ x \ y & x \ y" diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Algebra/Sylow.thy --- a/src/HOL/Algebra/Sylow.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Algebra/Sylow.thy Sat Nov 14 09:40:27 2009 +0100 @@ -344,7 +344,7 @@ done lemma (in sylow_central) card_H_eq: "card(H) = p^a" -by (blast intro: le_anti_sym lemma_leq1 lemma_leq2) +by (blast intro: le_antisym lemma_leq1 lemma_leq2) lemma (in sylow) sylow_thm: "\H. subgroup H G & card(H) = p^a" apply (cut_tac lemma_A1, clarify) diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Algebra/UnivPoly.thy --- a/src/HOL/Algebra/UnivPoly.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Algebra/UnivPoly.thy Sat Nov 14 09:40:27 2009 +0100 @@ -811,7 +811,7 @@ lemma deg_eqI: "[| !!m. n < m ==> coeff P p m = \; !!n. n ~= 0 ==> coeff P p n ~= \; p \ carrier P |] ==> deg R p = n" -by (fast intro: le_anti_sym deg_aboveI deg_belowI) +by (fast intro: le_antisym deg_aboveI deg_belowI) text {* Degree and polynomial operations *} @@ -826,11 +826,11 @@ lemma deg_monom [simp]: "[| a ~= \; a \ carrier R |] ==> deg R (monom P a n) = n" - by (fastsimp intro: le_anti_sym deg_aboveI deg_belowI) + by (fastsimp intro: le_antisym deg_aboveI deg_belowI) lemma deg_const [simp]: assumes R: "a \ carrier R" shows "deg R (monom P a 0) = 0" -proof (rule le_anti_sym) +proof (rule le_antisym) show "deg R (monom P a 0) <= 0" by (rule deg_aboveI) (simp_all add: R) next show "0 <= deg R (monom P a 0)" by (rule deg_belowI) (simp_all add: R) @@ -838,7 +838,7 @@ lemma deg_zero [simp]: "deg R \\<^bsub>P\<^esub> = 0" -proof (rule le_anti_sym) +proof (rule le_antisym) show "deg R \\<^bsub>P\<^esub> <= 0" by (rule deg_aboveI) simp_all next show "0 <= deg R \\<^bsub>P\<^esub>" by (rule deg_belowI) simp_all @@ -846,7 +846,7 @@ lemma deg_one [simp]: "deg R \\<^bsub>P\<^esub> = 0" -proof (rule le_anti_sym) +proof (rule le_antisym) show "deg R \\<^bsub>P\<^esub> <= 0" by (rule deg_aboveI) simp_all next show "0 <= deg R \\<^bsub>P\<^esub>" by (rule deg_belowI) simp_all @@ -854,7 +854,7 @@ lemma deg_uminus [simp]: assumes R: "p \ carrier P" shows "deg R (\\<^bsub>P\<^esub> p) = deg R p" -proof (rule le_anti_sym) +proof (rule le_antisym) show "deg R (\\<^bsub>P\<^esub> p) <= deg R p" by (simp add: deg_aboveI deg_aboveD R) next show "deg R p <= deg R (\\<^bsub>P\<^esub> p)" @@ -878,7 +878,7 @@ lemma deg_smult [simp]: assumes R: "a \ carrier R" "p \ carrier P" shows "deg R (a \\<^bsub>P\<^esub> p) = (if a = \ then 0 else deg R p)" -proof (rule le_anti_sym) +proof (rule le_antisym) show "deg R (a \\<^bsub>P\<^esub> p) <= (if a = \ then 0 else deg R p)" using R by (rule deg_smult_ring) next @@ -920,7 +920,7 @@ lemma deg_mult [simp]: "[| p ~= \\<^bsub>P\<^esub>; q ~= \\<^bsub>P\<^esub>; p \ carrier P; q \ carrier P |] ==> deg R (p \\<^bsub>P\<^esub> q) = deg R p + deg R q" -proof (rule le_anti_sym) +proof (rule le_antisym) assume "p \ carrier P" " q \ carrier P" then show "deg R (p \\<^bsub>P\<^esub> q) <= deg R p + deg R q" by (rule deg_mult_ring) next diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Algebra/poly/UnivPoly2.thy --- a/src/HOL/Algebra/poly/UnivPoly2.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Algebra/poly/UnivPoly2.thy Sat Nov 14 09:40:27 2009 +0100 @@ -557,7 +557,7 @@ lemma deg_eqI: "[| !!m. n < m ==> coeff p m = 0; !!n. n ~= 0 ==> coeff p n ~= 0|] ==> deg p = n" -by (fast intro: le_anti_sym deg_aboveI deg_belowI) +by (fast intro: le_antisym deg_aboveI deg_belowI) (* Degree and polynomial operations *) @@ -571,11 +571,11 @@ lemma deg_monom [simp]: "a ~= 0 ==> deg (monom a n::'a::ring up) = n" -by (fastsimp intro: le_anti_sym deg_aboveI deg_belowI) +by (fastsimp intro: le_antisym deg_aboveI deg_belowI) lemma deg_const [simp]: "deg (monom (a::'a::ring) 0) = 0" -proof (rule le_anti_sym) +proof (rule le_antisym) show "deg (monom a 0) <= 0" by (rule deg_aboveI) simp next show "0 <= deg (monom a 0)" by (rule deg_belowI) simp @@ -583,7 +583,7 @@ lemma deg_zero [simp]: "deg 0 = 0" -proof (rule le_anti_sym) +proof (rule le_antisym) show "deg 0 <= 0" by (rule deg_aboveI) simp next show "0 <= deg 0" by (rule deg_belowI) simp @@ -591,7 +591,7 @@ lemma deg_one [simp]: "deg 1 = 0" -proof (rule le_anti_sym) +proof (rule le_antisym) show "deg 1 <= 0" by (rule deg_aboveI) simp next show "0 <= deg 1" by (rule deg_belowI) simp @@ -612,7 +612,7 @@ lemma deg_uminus [simp]: "deg (-p::('a::ring) up) = deg p" -proof (rule le_anti_sym) +proof (rule le_antisym) show "deg (- p) <= deg p" by (simp add: deg_aboveI deg_aboveD) next show "deg p <= deg (- p)" @@ -626,7 +626,7 @@ lemma deg_smult [simp]: "deg ((a::'a::domain) *s p) = (if a = 0 then 0 else deg p)" -proof (rule le_anti_sym) +proof (rule le_antisym) show "deg (a *s p) <= (if a = 0 then 0 else deg p)" by (rule deg_smult_ring) next show "(if a = 0 then 0 else deg p) <= deg (a *s p)" @@ -657,7 +657,7 @@ lemma deg_mult [simp]: "[| (p::'a::domain up) ~= 0; q ~= 0|] ==> deg (p * q) = deg p + deg q" -proof (rule le_anti_sym) +proof (rule le_antisym) show "deg (p * q) <= deg p + deg q" by (rule deg_mult_ring) next let ?s = "(%i. coeff p i * coeff q (deg p + deg q - i))" diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Boogie/Examples/Boogie_Max.thy --- a/src/HOL/Boogie/Examples/Boogie_Max.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Boogie/Examples/Boogie_Max.thy Sat Nov 14 09:40:27 2009 +0100 @@ -57,14 +57,20 @@ proof (split_vc (verbose) try: fast simp) print_cases case L_14_5c - thus ?case by (metis abs_of_nonneg zabs_less_one_iff zle_linear) + 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 linorder_not_le xt1(10) zle_linear - zless_add1_eq) + thus ?case by (metis less_le_not_le xt1(10) zle_linear zless_add1_eq) next case L_14_5a - thus ?case by (metis less_le_not_le zle_add1_eq_le zless_linear) + note max0 = `max0 = array 0` + { + fix i :: int + assume "0 \ i \ i < 1" + hence "i = 0" by simp + with max0 have "array i \ max0" by simp + } + thus ?case by simp qed boogie_end diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Boogie/Examples/cert/Boogie_Dijkstra --- a/src/HOL/Boogie/Examples/cert/Boogie_Dijkstra Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Boogie/Examples/cert/Boogie_Dijkstra Sat Nov 14 09:40:27 2009 +0100 @@ -12,18 +12,18 @@ (uf_10 T1 Int) (uf_11 T2) (uf_9 Int) - (uf_21 T2) - (uf_16 T2) - (uf_20 T2) + (uf_18 T2) + (uf_22 T2) + (uf_17 T2) (uf_12 T2 Int) (uf_14 T3) - (uf_18 T2 Int) - (uf_22 T3) - (uf_24 T3) - (uf_23 T3) + (uf_24 T2 Int) + (uf_19 T3) + (uf_21 T3) + (uf_20 T3) (uf_15 T4) - (uf_17 T4) - (uf_19 T4) + (uf_23 T4) + (uf_16 T4) ) :extrapreds ( (up_13 T2) @@ -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 (implies (= (uf_12 uf_11) 0) (and (implies (forall (?x27 T2) (<= 0 (uf_12 ?x27))) (and (implies (forall (?x28 T2) (?x29 T2) (implies (and (up_13 ?x28) (not (up_13 ?x29))) (<= (uf_12 ?x28) (uf_12 ?x29)))) (and (implies (forall (?x30 T2) (?x31 T2) (implies (and (< (uf_10 (uf_1 ?x31 ?x30)) uf_9) (up_13 ?x31)) (<= (uf_12 ?x30) (+ (uf_12 ?x31) (uf_10 (uf_1 ?x31 ?x30)))))) (and (implies (forall (?x32 T2) (implies (and (< (uf_12 ?x32) uf_9) (not (= ?x32 uf_11))) (exists (?x33 T2) (and (= (uf_12 ?x32) (+ (uf_12 ?x33) (uf_10 (uf_1 ?x33 ?x32)))) (and (up_13 ?x33) (< (uf_12 ?x33) (uf_12 ?x32))))))) (implies true (implies true (implies (= (uf_4 uf_14 uf_11) 0) (implies (forall (?x34 T2) (<= 0 (uf_4 uf_14 ?x34))) (implies (forall (?x35 T2) (?x36 T2) (implies (and (= (uf_6 uf_15 ?x35) uf_8) (not (= (uf_6 uf_15 ?x36) uf_8))) (<= (uf_4 uf_14 ?x35) (uf_4 uf_14 ?x36)))) (implies (forall (?x37 T2) (?x38 T2) (implies (and (< (uf_10 (uf_1 ?x38 ?x37)) uf_9) (= (uf_6 uf_15 ?x38) uf_8)) (<= (uf_4 uf_14 ?x37) (+ (uf_4 uf_14 ?x38) (uf_10 (uf_1 ?x38 ?x37)))))) (implies (forall (?x39 T2) (implies (and (< (uf_4 uf_14 ?x39) uf_9) (not (= ?x39 uf_11))) (exists (?x40 T2) (and (= (uf_4 uf_14 ?x39) (+ (uf_4 uf_14 ?x40) (uf_10 (uf_1 ?x40 ?x39)))) (and (= (uf_6 uf_15 ?x40) uf_8) (< (uf_4 uf_14 ?x40) (uf_4 uf_14 ?x39))))))) (implies true (and (implies true (implies true (implies (exists (?x41 T2) (and (< (uf_4 uf_14 ?x41) uf_9) (not (= (uf_6 uf_15 ?x41) uf_8)))) (implies (not (= (uf_6 uf_15 uf_16) uf_8)) (implies (< (uf_4 uf_14 uf_16) uf_9) (implies (forall (?x42 T2) (implies (not (= (uf_6 uf_15 ?x42) uf_8)) (<= (uf_4 uf_14 uf_16) (uf_4 uf_14 ?x42)))) (implies (= uf_17 (uf_7 uf_15 uf_16 uf_8)) (implies (forall (?x43 T2) (implies (and (< (+ (uf_4 uf_14 uf_16) (uf_10 (uf_1 uf_16 ?x43))) (uf_4 uf_14 ?x43)) (< (uf_10 (uf_1 uf_16 ?x43)) uf_9)) (= (uf_18 ?x43) (+ (uf_4 uf_14 uf_16) (uf_10 (uf_1 uf_16 ?x43)))))) (implies (forall (?x44 T2) (implies (not (and (< (+ (uf_4 uf_14 uf_16) (uf_10 (uf_1 uf_16 ?x44))) (uf_4 uf_14 ?x44)) (< (uf_10 (uf_1 uf_16 ?x44)) uf_9))) (= (uf_18 ?x44) (uf_4 uf_14 ?x44)))) (and (implies (forall (?x45 T2) (<= (uf_18 ?x45) (uf_4 uf_14 ?x45))) (and (implies (forall (?x46 T2) (implies (= (uf_6 uf_17 ?x46) uf_8) (= (uf_18 ?x46) (uf_4 uf_14 ?x46)))) (implies true (implies true (and (implies (= (uf_18 uf_11) 0) (and (implies (forall (?x47 T2) (<= 0 (uf_18 ?x47))) (and (implies (forall (?x48 T2) (?x49 T2) (implies (and (= (uf_6 uf_17 ?x48) uf_8) (not (= (uf_6 uf_17 ?x49) uf_8))) (<= (uf_18 ?x48) (uf_18 ?x49)))) (and (implies (forall (?x50 T2) (?x51 T2) (implies (and (< (uf_10 (uf_1 ?x51 ?x50)) uf_9) (= (uf_6 uf_17 ?x51) uf_8)) (<= (uf_18 ?x50) (+ (uf_18 ?x51) (uf_10 (uf_1 ?x51 ?x50)))))) (and (implies (forall (?x52 T2) (implies (and (< (uf_18 ?x52) uf_9) (not (= ?x52 uf_11))) (exists (?x53 T2) (and (= (uf_18 ?x52) (+ (uf_18 ?x53) (uf_10 (uf_1 ?x53 ?x52)))) (and (= (uf_6 uf_17 ?x53) uf_8) (< (uf_18 ?x53) (uf_18 ?x52))))))) (implies false true)) (forall (?x54 T2) (implies (and (< (uf_18 ?x54) uf_9) (not (= ?x54 uf_11))) (exists (?x55 T2) (and (= (uf_18 ?x54) (+ (uf_18 ?x55) (uf_10 (uf_1 ?x55 ?x54)))) (and (= (uf_6 uf_17 ?x55) uf_8) (< (uf_18 ?x55) (uf_18 ?x54))))))))) (forall (?x56 T2) (?x57 T2) (implies (and (< (uf_10 (uf_1 ?x57 ?x56)) uf_9) (= (uf_6 uf_17 ?x57) uf_8)) (<= (uf_18 ?x56) (+ (uf_18 ?x57) (uf_10 (uf_1 ?x57 ?x56)))))))) (forall (?x58 T2) (?x59 T2) (implies (and (= (uf_6 uf_17 ?x58) uf_8) (not (= (uf_6 uf_17 ?x59) uf_8))) (<= (uf_18 ?x58) (uf_18 ?x59)))))) (forall (?x60 T2) (<= 0 (uf_18 ?x60))))) (= (uf_18 uf_11) 0))))) (forall (?x61 T2) (implies (= (uf_6 uf_17 ?x61) uf_8) (= (uf_18 ?x61) (uf_4 uf_14 ?x61)))))) (forall (?x62 T2) (<= (uf_18 ?x62) (uf_4 uf_14 ?x62))))))))))))) (implies true (implies true (implies (not (exists (?x63 T2) (and (< (uf_4 uf_14 ?x63) uf_9) (not (= (uf_6 uf_15 ?x63) uf_8))))) (implies true (implies true (implies (= uf_19 uf_15) (implies (= uf_20 uf_21) (implies (= uf_22 uf_14) (implies (= uf_23 uf_24) (implies true (and (implies (forall (?x64 T2) (implies (and (< (uf_4 uf_22 ?x64) uf_9) (not (= ?x64 uf_11))) (exists (?x65 T2) (and (= (uf_4 uf_22 ?x64) (+ (uf_4 uf_22 ?x65) (uf_10 (uf_1 ?x65 ?x64)))) (< (uf_4 uf_22 ?x65) (uf_4 uf_22 ?x64)))))) (and (implies (forall (?x66 T2) (?x67 T2) (implies (and (< (uf_10 (uf_1 ?x67 ?x66)) uf_9) (< (uf_4 uf_22 ?x67) uf_9)) (<= (uf_4 uf_22 ?x66) (+ (uf_4 uf_22 ?x67) (uf_10 (uf_1 ?x67 ?x66)))))) (and (implies (= (uf_4 uf_22 uf_11) 0) true) (= (uf_4 uf_22 uf_11) 0))) (forall (?x68 T2) (?x69 T2) (implies (and (< (uf_10 (uf_1 ?x69 ?x68)) uf_9) (< (uf_4 uf_22 ?x69) uf_9)) (<= (uf_4 uf_22 ?x68) (+ (uf_4 uf_22 ?x69) (uf_10 (uf_1 ?x69 ?x68)))))))) (forall (?x70 T2) (implies (and (< (uf_4 uf_22 ?x70) uf_9) (not (= ?x70 uf_11))) (exists (?x71 T2) (and (= (uf_4 uf_22 ?x70) (+ (uf_4 uf_22 ?x71) (uf_10 (uf_1 ?x71 ?x70)))) (< (uf_4 uf_22 ?x71) (uf_4 uf_22 ?x70))))))))))))))))))))))))))) (forall (?x72 T2) (implies (and (< (uf_12 ?x72) uf_9) (not (= ?x72 uf_11))) (exists (?x73 T2) (and (= (uf_12 ?x72) (+ (uf_12 ?x73) (uf_10 (uf_1 ?x73 ?x72)))) (and (up_13 ?x73) (< (uf_12 ?x73) (uf_12 ?x72))))))))) (forall (?x74 T2) (?x75 T2) (implies (and (< (uf_10 (uf_1 ?x75 ?x74)) uf_9) (up_13 ?x75)) (<= (uf_12 ?x74) (+ (uf_12 ?x75) (uf_10 (uf_1 ?x75 ?x74)))))))) (forall (?x76 T2) (?x77 T2) (implies (and (up_13 ?x76) (not (up_13 ?x77))) (<= (uf_12 ?x76) (uf_12 ?x77)))))) (forall (?x78 T2) (<= 0 (uf_12 ?x78))))) (= (uf_12 uf_11) 0))))))))) +: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 (= (uf_6 uf_15 ?x43) uf_8)) (= (uf_6 uf_15 ?x42) uf_8)) (<= (uf_4 uf_14 ?x42) (uf_4 uf_14 ?x43)))) (implies (forall (?x44 T2) (?x45 T2) (implies (and (= (uf_6 uf_15 ?x45) uf_8) (< (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 (= (uf_6 uf_15 ?x47) uf_8) (= (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 (= (uf_6 uf_15 ?x48) uf_8)) (< (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 (= (uf_6 uf_15 ?x57) uf_8)) (< (uf_4 uf_14 ?x57) uf_9))) (implies (not (= (uf_6 uf_15 uf_22) uf_8)) (implies (< (uf_4 uf_14 uf_22) uf_9) (implies (forall (?x58 T2) (implies (not (= (uf_6 uf_15 ?x58) uf_8)) (<= (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 (= (uf_6 uf_23 ?x63) uf_8) (= (uf_24 ?x63) (uf_4 uf_14 ?x63)))) (implies (forall (?x64 T2) (implies (= (uf_6 uf_23 ?x64) uf_8) (= (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 (= (uf_6 uf_23 ?x68) uf_8)) (= (uf_6 uf_23 ?x67) uf_8)) (<= (uf_24 ?x67) (uf_24 ?x68)))) (implies (forall (?x69 T2) (?x70 T2) (implies (and (not (= (uf_6 uf_23 ?x70) uf_8)) (= (uf_6 uf_23 ?x69) uf_8)) (<= (uf_24 ?x69) (uf_24 ?x70)))) (and (forall (?x71 T2) (?x72 T2) (implies (and (= (uf_6 uf_23 ?x72) uf_8) (< (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 (= (uf_6 uf_23 ?x74) uf_8) (< (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 (= (uf_6 uf_23 ?x76) uf_8) (= (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 (= (uf_6 uf_23 ?x78) uf_8) (= (uf_24 ?x77) (+ (uf_24 ?x78) (uf_10 (uf_1 ?x78 ?x77))))))))) (implies false true)))))))))))))))))))))))))))))))))))))))))))))))))))) :formula true ) diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Boogie/Examples/cert/Boogie_Dijkstra.proof --- a/src/HOL/Boogie/Examples/cert/Boogie_Dijkstra.proof Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Boogie/Examples/cert/Boogie_Dijkstra.proof Sat Nov 14 09:40:27 2009 +0100 @@ -1,4237 +1,239 @@ #2 := false #55 := 0::int -decl uf_4 :: (-> T3 T2 int) -decl ?x40!7 :: (-> T2 T2) -decl ?x52!15 :: T2 -#2305 := ?x52!15 -#15992 := (?x40!7 ?x52!15) -decl uf_14 :: T3 -#107 := uf_14 -#15996 := (uf_4 uf_14 #15992) -#20405 := (>= #15996 0::int) -#11 := (:var 0 T2) -#110 := (uf_4 uf_14 #11) -#4403 := (pattern #110) -#1843 := (>= #110 0::int) -#4404 := (forall (vars (?x34 T2)) (:pat #4403) #1843) -decl uf_10 :: (-> T1 int) +decl uf_24 :: (-> T2 int) +decl ?x68!16 :: T2 +#2296 := ?x68!16 +#2300 := (uf_24 ?x68!16) +#1220 := -1::int +#2894 := (* -1::int #2300) +decl ?x67!17 :: T2 +#2297 := ?x67!17 +#2298 := (uf_24 ?x67!17) +#2895 := (+ #2298 #2894) +#2896 := (<= #2895 0::int) +#4133 := (not #2896) +decl uf_6 :: (-> T4 T2 T5) +decl uf_23 :: T4 +#193 := uf_23 +#2305 := (uf_6 uf_23 ?x68!16) +decl uf_8 :: T5 +#33 := uf_8 +#2306 := (= uf_8 #2305) +#2303 := (uf_6 uf_23 ?x67!17) +#2304 := (= uf_8 #2303) +#3433 := (not #2304) +#3448 := (or #3433 #2306 #2896) +#3453 := (not #3448) decl uf_1 :: (-> T2 T2 T1) -decl ?x66!20 :: T2 -#2511 := ?x66!20 -decl ?x67!19 :: T2 -#2510 := ?x67!19 -#2516 := (uf_1 ?x67!19 ?x66!20) -#2517 := (uf_10 #2516) -#1320 := -1::int -#2524 := (* -1::int #2517) -decl uf_22 :: T3 -#230 := uf_22 -#2514 := (uf_4 uf_22 ?x67!19) -#2520 := (* -1::int #2514) -#3094 := (+ #2520 #2524) -#2512 := (uf_4 uf_22 ?x66!20) -#3095 := (+ #2512 #3094) -#3096 := (<= #3095 0::int) -decl uf_9 :: int -#56 := uf_9 -#2525 := (+ uf_9 #2524) -#2526 := (<= #2525 0::int) -#2521 := (+ uf_9 #2520) -#2522 := (<= #2521 0::int) -#3693 := (or #2522 #2526 #3096) -#3698 := (not #3693) +decl ?x75!20 :: T2 +#2354 := ?x75!20 +#11 := (:var 0 T2) +#2358 := (uf_1 #11 ?x75!20) +#4486 := (pattern #2358) +#202 := (uf_24 #11) +#4426 := (pattern #202) +#212 := (uf_6 uf_23 #11) +#4452 := (pattern #212) +decl uf_10 :: (-> T1 int) +#2359 := (uf_10 #2358) +#2355 := (uf_24 ?x75!20) +#2356 := (* -1::int #2355) +#2958 := (+ #2356 #2359) +#2959 := (+ #202 #2958) +#2962 := (= #2959 0::int) +#3524 := (not #2962) +#2357 := (+ #202 #2356) +#2362 := (>= #2357 0::int) +#773 := (= uf_8 #212) +#779 := (not #773) +#3525 := (or #779 #2362 #3524) +#4487 := (forall (vars (?x76 T2)) (:pat #4452 #4426 #4486) #3525) +#4492 := (not #4487) #10 := (:var 1 T2) #90 := (uf_1 #11 #10) -#4379 := (pattern #90) -#238 := (uf_4 uf_22 #10) -#1720 := (* -1::int #238) -#235 := (uf_4 uf_22 #11) -#1721 := (+ #235 #1720) +#4281 := (pattern #90) +#224 := (uf_24 #10) +#1505 := (* -1::int #224) +#1506 := (+ #202 #1505) #91 := (uf_10 #90) -#1727 := (+ #91 #1721) -#1750 := (>= #1727 0::int) -#1707 := (* -1::int #235) -#1708 := (+ uf_9 #1707) -#1709 := (<= #1708 0::int) -#1343 := (* -1::int #91) -#1346 := (+ uf_9 #1343) -#1347 := (<= #1346 0::int) -#3661 := (or #1347 #1709 #1750) -#4633 := (forall (vars (?x66 T2) (?x67 T2)) (:pat #4379) #3661) -#4638 := (not #4633) -decl uf_11 :: T2 -#67 := uf_11 -#250 := (uf_4 uf_22 uf_11) -#251 := (= #250 0::int) -#4641 := (or #251 #4638) -#4644 := (not #4641) -#4647 := (or #4644 #3698) -#4650 := (not #4647) -#4609 := (pattern #235) -decl ?x65!18 :: (-> T2 T2) -#2487 := (?x65!18 #11) -#2490 := (uf_1 #2487 #11) -#2491 := (uf_10 #2490) -#3064 := (* -1::int #2491) -#2488 := (uf_4 uf_22 #2487) -#3047 := (* -1::int #2488) -#3065 := (+ #3047 #3064) -#3066 := (+ #235 #3065) -#3067 := (= #3066 0::int) -#3631 := (not #3067) -#3048 := (+ #235 #3047) -#3049 := (<= #3048 0::int) -#3632 := (or #3049 #3631) -#3633 := (not #3632) -#68 := (= #11 uf_11) -#3639 := (or #68 #1709 #3633) -#4625 := (forall (vars (?x64 T2)) (:pat #4609) #3639) -#4630 := (not #4625) -#4653 := (or #4630 #4650) -#4656 := (not #4653) -decl ?x64!17 :: T2 -#2447 := ?x64!17 -#2451 := (uf_1 #11 ?x64!17) -#4610 := (pattern #2451) -#2452 := (uf_10 #2451) -#2448 := (uf_4 uf_22 ?x64!17) -#2449 := (* -1::int #2448) -#3017 := (+ #2449 #2452) -#3018 := (+ #235 #3017) -#3021 := (= #3018 0::int) -#3595 := (not #3021) -#2450 := (+ #235 #2449) -#2455 := (>= #2450 0::int) -#3596 := (or #2455 #3595) -#4611 := (forall (vars (?x65 T2)) (:pat #4609 #4610) #3596) -#4616 := (not #4611) -#2993 := (= uf_11 ?x64!17) -#2459 := (+ uf_9 #2449) -#2460 := (<= #2459 0::int) -#4619 := (or #2460 #2993 #4616) -#4622 := (not #4619) -#4659 := (or #4622 #4656) -#4662 := (not #4659) -decl uf_6 :: (-> T4 T2 T5) -decl uf_15 :: T4 -#113 := uf_15 -#116 := (uf_6 uf_15 #11) -#4445 := (pattern #116) -#1404 := (* -1::int #110) -#1405 := (+ uf_9 #1404) -#1406 := (<= #1405 0::int) -decl uf_8 :: T5 -#33 := uf_8 -#505 := (= uf_8 #116) -#3581 := (or #505 #1406) -#4601 := (forall (vars (?x41 T2)) (:pat #4445 #4403) #3581) -#4606 := (not #4601) -#933 := (= uf_14 uf_22) -#1053 := (not #933) -decl uf_19 :: T4 -#225 := uf_19 -#930 := (= uf_15 uf_19) -#1071 := (not #930) -decl uf_24 :: T3 -#233 := uf_24 -decl uf_23 :: T3 -#232 := uf_23 -#234 := (= uf_23 uf_24) -#1044 := (not #234) -decl uf_21 :: T2 -#228 := uf_21 -decl uf_20 :: T2 -#227 := uf_20 -#229 := (= uf_20 uf_21) -#1062 := (not #229) -#4665 := (or #1062 #1044 #1071 #1053 #4606 #4662) -#4668 := (not #4665) -#2309 := (uf_1 #11 ?x52!15) -#4514 := (pattern #2309) -decl uf_18 :: (-> T2 int) -#158 := (uf_18 #11) -#4454 := (pattern #158) -decl uf_17 :: T4 -#149 := uf_17 -#168 := (uf_6 uf_17 #11) -#4480 := (pattern #168) -#2310 := (uf_10 #2309) -#2306 := (uf_18 ?x52!15) -#2307 := (* -1::int #2306) -#2917 := (+ #2307 #2310) -#2918 := (+ #158 #2917) -#2921 := (= #2918 0::int) -#3474 := (not #2921) -#2308 := (+ #158 #2307) -#2313 := (>= #2308 0::int) -#630 := (= uf_8 #168) -#636 := (not #630) -#3475 := (or #636 #2313 #3474) -#4515 := (forall (vars (?x53 T2)) (:pat #4480 #4454 #4514) #3475) -#4520 := (not #4515) -#180 := (uf_18 #10) -#1505 := (* -1::int #180) -#1506 := (+ #158 #1505) #1536 := (+ #91 #1506) #1534 := (>= #1536 0::int) -#3466 := (or #636 #1347 #1534) -#4506 := (forall (vars (?x50 T2) (?x51 T2)) (:pat #4379) #3466) -#4511 := (not #4506) -#2893 := (= uf_11 ?x52!15) -#2317 := (+ uf_9 #2307) -#2318 := (<= #2317 0::int) -#4523 := (or #2318 #2893 #4511 #4520) -#4526 := (not #4523) -decl ?x50!14 :: T2 -#2275 := ?x50!14 -decl ?x51!13 :: T2 -#2274 := ?x51!13 -#2280 := (uf_1 ?x51!13 ?x50!14) -#2281 := (uf_10 #2280) -#2284 := (* -1::int #2281) -#2278 := (uf_18 ?x51!13) -#2879 := (* -1::int #2278) -#2880 := (+ #2879 #2284) -#2276 := (uf_18 ?x50!14) -#2881 := (+ #2276 #2880) -#2882 := (<= #2881 0::int) -#2288 := (uf_6 uf_17 ?x51!13) -#2289 := (= uf_8 #2288) -#3429 := (not #2289) -#2285 := (+ uf_9 #2284) -#2286 := (<= #2285 0::int) -#3444 := (or #2286 #3429 #2882) -#3449 := (not #3444) -#4529 := (or #3449 #4526) -#4532 := (not #4529) -#4497 := (pattern #158 #180) +#1235 := (* -1::int #91) +decl uf_9 :: int +#56 := uf_9 +#1236 := (+ uf_9 #1235) +#1237 := (<= #1236 0::int) +#3516 := (or #779 #1237 #1534) +#4478 := (forall (vars (?x71 T2) (?x72 T2)) (:pat #4281) #3516) +#4483 := (not #4478) +decl uf_11 :: T2 +#67 := uf_11 +#2934 := (= uf_11 ?x75!20) +#2366 := (+ uf_9 #2356) +#2367 := (<= #2366 0::int) +#4495 := (or #2367 #2934 #4483 #4492) +#4498 := (not #4495) +decl ?x71!19 :: T2 +#2324 := ?x71!19 +decl ?x72!18 :: T2 +#2323 := ?x72!18 +#2329 := (uf_1 ?x72!18 ?x71!19) +#2330 := (uf_10 #2329) +#2333 := (* -1::int #2330) +#2327 := (uf_24 ?x72!18) +#2920 := (* -1::int #2327) +#2921 := (+ #2920 #2333) +#2325 := (uf_24 ?x71!19) +#2922 := (+ #2325 #2921) +#2923 := (<= #2922 0::int) +#2337 := (uf_6 uf_23 ?x72!18) +#2338 := (= uf_8 #2337) +#3479 := (not #2338) +#2334 := (+ uf_9 #2333) +#2335 := (<= #2334 0::int) +#3494 := (or #2335 #3479 #2923) +#3499 := (not #3494) +#4501 := (or #3499 #4498) +#4504 := (not #4501) +#4469 := (pattern #202 #224) #1504 := (>= #1506 0::int) -#176 := (uf_6 uf_17 #10) -#648 := (= uf_8 #176) -#3406 := (not #648) -#3421 := (or #630 #3406 #1504) -#4498 := (forall (vars (?x48 T2) (?x49 T2)) (:pat #4497) #3421) -#4503 := (not #4498) -#4535 := (or #4503 #4532) -#4538 := (not #4535) -decl ?x49!11 :: T2 -#2247 := ?x49!11 -#2251 := (uf_18 ?x49!11) -#2853 := (* -1::int #2251) -decl ?x48!12 :: T2 -#2248 := ?x48!12 -#2249 := (uf_18 ?x48!12) -#2854 := (+ #2249 #2853) -#2855 := (<= #2854 0::int) -#2256 := (uf_6 uf_17 ?x49!11) -#2257 := (= uf_8 #2256) -#2254 := (uf_6 uf_17 ?x48!12) -#2255 := (= uf_8 #2254) -#3383 := (not #2255) -#3398 := (or #3383 #2257 #2855) -#3403 := (not #3398) -#4541 := (or #3403 #4538) -#4544 := (not #4541) -#1495 := (>= #158 0::int) -#4489 := (forall (vars (?x47 T2)) (:pat #4454) #1495) -#4494 := (not #4489) -#4547 := (or #4494 #4544) -#4550 := (not #4547) -decl ?x47!10 :: T2 -#2232 := ?x47!10 -#2233 := (uf_18 ?x47!10) -#2234 := (>= #2233 0::int) -#2235 := (not #2234) -#4553 := (or #2235 #4550) -#4556 := (not #4553) -#172 := (uf_18 uf_11) -#173 := (= #172 0::int) -#1492 := (not #173) -#4559 := (or #1492 #4556) -#4562 := (not #4559) -#4565 := (or #1492 #4562) -#4568 := (not #4565) -#616 := (= #110 #158) -#637 := (or #616 #636) -#4481 := (forall (vars (?x46 T2)) (:pat #4403 #4454 #4480) #637) -#4486 := (not #4481) -#4571 := (or #4486 #4568) -#4574 := (not #4571) -decl ?x46!9 :: T2 -#2207 := ?x46!9 -#2212 := (uf_4 uf_14 ?x46!9) -#2211 := (uf_18 ?x46!9) -#2825 := (= #2211 #2212) -#2208 := (uf_6 uf_17 ?x46!9) -#2209 := (= uf_8 #2208) -#2210 := (not #2209) -#2831 := (or #2210 #2825) -#2836 := (not #2831) -#4577 := (or #2836 #4574) -#4580 := (not #4577) -#1480 := (* -1::int #158) -#1481 := (+ #110 #1480) -#1479 := (>= #1481 0::int) -#4472 := (forall (vars (?x45 T2)) (:pat #4403 #4454) #1479) -#4477 := (not #4472) -#4583 := (or #4477 #4580) -#4586 := (not #4583) -decl ?x45!8 :: T2 -#2189 := ?x45!8 -#2192 := (uf_4 uf_14 ?x45!8) -#2815 := (* -1::int #2192) -#2190 := (uf_18 ?x45!8) -#2816 := (+ #2190 #2815) -#2817 := (<= #2816 0::int) -#2822 := (not #2817) -#4589 := (or #2822 #4586) -#4592 := (not #4589) -decl uf_16 :: T2 -#140 := uf_16 -#152 := (uf_1 uf_16 #11) -#4455 := (pattern #152) -#153 := (uf_10 #152) -#1623 := (+ #153 #1480) -#144 := (uf_4 uf_14 uf_16) -#1624 := (+ #144 #1623) -#1625 := (= #1624 0::int) -#1450 := (* -1::int #153) -#1457 := (+ uf_9 #1450) -#1458 := (<= #1457 0::int) -#1449 := (* -1::int #144) -#1451 := (+ #1449 #1450) -#1452 := (+ #110 #1451) -#1453 := (<= #1452 0::int) -#3375 := (or #1453 #1458 #1625) -#4464 := (forall (vars (?x43 T2)) (:pat #4403 #4455 #4454) #3375) -#4469 := (not #4464) -#3355 := (or #1453 #1458) -#3356 := (not #3355) -#3359 := (or #616 #3356) -#4456 := (forall (vars (?x44 T2)) (:pat #4403 #4454 #4455) #3359) -#4461 := (not #4456) -decl ?x41!16 :: T2 -#2408 := ?x41!16 -#2414 := (uf_6 uf_15 ?x41!16) -#2415 := (= uf_8 #2414) -#2409 := (uf_4 uf_14 ?x41!16) -#2410 := (* -1::int #2409) -#2411 := (+ uf_9 #2410) -#2412 := (<= #2411 0::int) -#1655 := (+ uf_9 #1449) -#1656 := (<= #1655 0::int) -#1638 := (+ #110 #1449) -#1637 := (>= #1638 0::int) -#1644 := (or #505 #1637) -#4446 := (forall (vars (?x42 T2)) (:pat #4445 #4403) #1644) -#4451 := (not #4446) -#141 := (uf_6 uf_15 uf_16) -#585 := (= uf_8 #141) +#221 := (uf_6 uf_23 #10) +#793 := (= uf_8 #221) +#3456 := (not #793) +#3471 := (or #773 #3456 #1504) +#4470 := (forall (vars (?x67 T2) (?x68 T2)) (:pat #4469) #3471) +#4475 := (not #4470) +#4507 := (or #4475 #4504) +#7658 := [hypothesis]: #3499 +#2336 := (not #2335) +#4131 := (or #3494 #2336) +#4137 := [def-axiom]: #4131 +#17052 := [unit-resolution #4137 #7658]: #2336 +#17077 := (or #3494 #2335) +decl uf_4 :: (-> T3 T2 int) +decl uf_14 :: T3 +#107 := uf_14 +#110 := (uf_4 uf_14 #11) +#4305 := (pattern #110) +#759 := (= #110 #202) +#780 := (or #759 #779) +#4453 := (forall (vars (?x63 T2)) (:pat #4305 #4426 #4452) #780) +#4510 := (not #4507) +#4513 := (or #3453 #4510) +#4516 := (not #4513) +#1495 := (>= #202 0::int) +#4461 := (forall (vars (?x65 T2)) (:pat #4426) #1495) +#4466 := (not #4461) +#4519 := (or #4466 #4516) +#4522 := (not #4519) +decl ?x65!15 :: T2 +#2281 := ?x65!15 +#2282 := (uf_24 ?x65!15) +#2283 := (>= #2282 0::int) +#2284 := (not #2283) +#4525 := (or #2284 #4522) +#4528 := (not #4525) +#216 := (uf_24 uf_11) +#217 := (= #216 0::int) +#1492 := (not #217) +#4531 := (or #1492 #4528) +#4534 := (not #4531) +#4537 := (or #1492 #4534) +#4540 := (not #4537) +#4458 := (not #4453) +#4543 := (or #4458 #4540) +#4546 := (not #4543) +decl ?x63!14 :: T2 +#2256 := ?x63!14 +#2261 := (uf_4 uf_14 ?x63!14) +#2260 := (uf_24 ?x63!14) +#2866 := (= #2260 #2261) +#2257 := (uf_6 uf_23 ?x63!14) +#2258 := (= uf_8 #2257) +#2259 := (not #2258) +#2872 := (or #2259 #2866) +#2877 := (not #2872) +#10222 := [hypothesis]: #2877 +#4144 := (or #2872 #2258) +#4145 := [def-axiom]: #4144 +#10559 := [unit-resolution #4145 #10222]: #2258 +#4140 := (not #2866) +#4141 := (or #2872 #4140) +#4146 := [def-axiom]: #4141 +#10294 := [unit-resolution #4146 #10222]: #4140 +decl uf_3 :: (-> T1 T2) +decl uf_22 :: T2 +#184 := uf_22 +#4728 := (uf_1 uf_22 uf_22) +#9695 := (uf_3 #4728) +#10367 := (uf_1 #9695 ?x63!14) +#10448 := (uf_3 #10367) +#11132 := (uf_4 uf_14 #10448) +#13212 := (= #11132 #2261) +#12385 := (= #2261 #11132) +#10449 := (= ?x63!14 #10448) +#12 := (uf_1 #10 #11) +#4196 := (pattern #12) +#13 := (uf_3 #12) +#317 := (= #11 #13) +#4197 := (forall (vars (?x2 T2) (?x3 T2)) (:pat #4196) #317) +#321 := (forall (vars (?x2 T2) (?x3 T2)) #317) +#4200 := (iff #321 #4197) +#4198 := (iff #317 #317) +#4199 := [refl]: #4198 +#4201 := [quant-intro #4199]: #4200 +#1843 := (~ #321 #321) +#1875 := (~ #317 #317) +#1876 := [refl]: #1875 +#1841 := [nnf-pos #1876]: #1843 +#14 := (= #13 #11) +#15 := (forall (vars (?x2 T2) (?x3 T2)) #14) +#322 := (iff #15 #321) +#319 := (iff #14 #317) +#320 := [rewrite]: #319 +#323 := [quant-intro #320]: #322 +#316 := [asserted]: #15 +#326 := [mp #316 #323]: #321 +#1877 := [mp~ #326 #1841]: #321 +#4202 := [mp #1877 #4201]: #4197 +#8139 := (not #4197) +#12947 := (or #8139 #10449) +#12948 := [quant-inst]: #12947 +#13195 := [unit-resolution #12948 #4202]: #10449 +#13203 := [monotonicity #13195]: #12385 +#13213 := [symm #13203]: #13212 +#13222 := (= #2260 #11132) +#188 := (uf_4 uf_14 uf_22) +#13623 := (= #188 #11132) +#13621 := (= #11132 #188) +#13610 := (= #10448 uf_22) +#10707 := (= #9695 uf_22) +#9696 := (= uf_22 #9695) +#9727 := (or #8139 #9696) +#9731 := [quant-inst]: #9727 +#10706 := [unit-resolution #9731 #4202]: #9696 +#10708 := [symm #10706]: #10707 +#13609 := (= #10448 #9695) +#10319 := (= ?x63!14 #9695) +decl uf_15 :: T4 +#113 := uf_15 +#9518 := (uf_6 uf_15 ?x63!14) +#9519 := (= uf_8 #9518) decl uf_7 :: (-> T4 T2 T5 T4) -#150 := (uf_7 uf_15 uf_16 uf_8) -#151 := (= uf_17 #150) -#876 := (not #151) -#4595 := (or #876 #585 #4451 #1656 #2412 #2415 #4461 #4469 #4592) -#4598 := (not #4595) -#4671 := (or #4598 #4668) -#4674 := (not #4671) -#2152 := (?x40!7 #11) -#2155 := (uf_1 #2152 #11) -#2156 := (uf_10 #2155) -#2790 := (* -1::int #2156) -#2153 := (uf_4 uf_14 #2152) -#2773 := (* -1::int #2153) -#2791 := (+ #2773 #2790) -#2792 := (+ #110 #2791) -#2793 := (= #2792 0::int) -#3339 := (not #2793) -#2774 := (+ #110 #2773) -#2775 := (<= #2774 0::int) -#2161 := (uf_6 uf_15 #2152) -#2162 := (= uf_8 #2161) -#3338 := (not #2162) -#3340 := (or #3338 #2775 #3339) -#3341 := (not #3340) -#3347 := (or #68 #1406 #3341) -#4437 := (forall (vars (?x39 T2)) (:pat #4403) #3347) -#4442 := (not #4437) -decl uf_12 :: (-> T2 int) -#69 := (uf_12 #11) -#4355 := (pattern #69) -decl ?x33!6 :: (-> T2 T2) -#2123 := (?x33!6 #11) -#2127 := (uf_12 #2123) -#2728 := (* -1::int #2127) -#2124 := (uf_1 #2123 #11) -#2125 := (uf_10 #2124) -#2745 := (* -1::int #2125) -#2746 := (+ #2745 #2728) -#2747 := (+ #69 #2746) -#2748 := (= #2747 0::int) -#3311 := (not #2748) -#2729 := (+ #69 #2728) -#2730 := (<= #2729 0::int) -decl up_13 :: (-> T2 bool) -#2133 := (up_13 #2123) -#3310 := (not #2133) -#3312 := (or #3310 #2730 #3311) -#3313 := (not #3312) -#1386 := (* -1::int #69) -#1387 := (+ uf_9 #1386) -#1388 := (<= #1387 0::int) -#3319 := (or #68 #1388 #3313) -#4429 := (forall (vars (?x32 T2)) (:pat #4355) #3319) -#4434 := (not #4429) -#114 := (uf_6 uf_15 #10) -#4420 := (pattern #114 #116) -#120 := (uf_4 uf_14 #10) -#1417 := (* -1::int #120) -#1418 := (+ #110 #1417) -#1416 := (>= #1418 0::int) -#502 := (= uf_8 #114) -#3276 := (not #502) -#3291 := (or #3276 #505 #1416) -#4421 := (forall (vars (?x35 T2) (?x36 T2)) (:pat #4420) #3291) -#4426 := (not #4421) -#1424 := (+ #91 #1418) -#1815 := (>= #1424 0::int) -#508 := (not #505) -#3268 := (or #508 #1347 #1815) -#4412 := (forall (vars (?x37 T2) (?x38 T2)) (:pat #4379) #3268) -#4417 := (not #4412) -#4409 := (not #4404) -#108 := (uf_4 uf_14 uf_11) -#109 := (= #108 0::int) -#1854 := (not #109) -#4677 := (or #1854 #4409 #4417 #4426 #4434 #4442 #4674) -#4680 := (not #4677) -decl ?x32!5 :: T2 -#2081 := ?x32!5 -#2091 := (uf_1 #11 ?x32!5) -#4388 := (pattern #2091) -#77 := (up_13 #11) -#4348 := (pattern #77) -#2082 := (uf_12 ?x32!5) -#2083 := (* -1::int #2082) -#2096 := (+ #69 #2083) -#2097 := (>= #2096 0::int) -#2092 := (uf_10 #2091) -#2093 := (+ #2083 #2092) -#2094 := (+ #69 #2093) -#2095 := (= #2094 0::int) -#3229 := (not #2095) -#78 := (not #77) -#3230 := (or #78 #3229 #2097) -#4389 := (forall (vars (?x33 T2)) (:pat #4348 #4355 #4388) #3230) -#4394 := (not #4389) -#2688 := (= uf_11 ?x32!5) -#2084 := (+ uf_9 #2083) -#2085 := (<= #2084 0::int) -#4397 := (or #2085 #2688 #4394) -#4400 := (not #4397) -#4683 := (or #4400 #4680) -#4686 := (not #4683) -#86 := (uf_12 #10) -#1323 := (* -1::int #86) -#1344 := (+ #1323 #91) -#1345 := (+ #69 #1344) -#1342 := (>= #1345 0::int) -#3221 := (or #78 #1342 #1347) -#4380 := (forall (vars (?x30 T2) (?x31 T2)) (:pat #4379) #3221) -#4385 := (not #4380) -#4689 := (or #4385 #4686) -#4692 := (not #4689) -decl ?x31!3 :: T2 -#2051 := ?x31!3 -#2065 := (uf_12 ?x31!3) -decl ?x30!4 :: T2 -#2052 := ?x30!4 -#2062 := (uf_12 ?x30!4) -#2063 := (* -1::int #2062) -#2660 := (+ #2063 #2065) -#2053 := (uf_1 ?x31!3 ?x30!4) -#2054 := (uf_10 #2053) -#2661 := (+ #2054 #2660) -#2664 := (>= #2661 0::int) -#2059 := (up_13 ?x31!3) -#3184 := (not #2059) -#2055 := (* -1::int #2054) -#2056 := (+ uf_9 #2055) -#2057 := (<= #2056 0::int) -#3199 := (or #2057 #3184 #2664) -#3204 := (not #3199) -#4695 := (or #3204 #4692) -#4698 := (not #4695) -#84 := (up_13 #10) -#4370 := (pattern #77 #84) -#1324 := (+ #69 #1323) -#1322 := (>= #1324 0::int) -#2632 := (not #84) -#3176 := (or #77 #2632 #1322) -#4371 := (forall (vars (?x28 T2) (?x29 T2)) (:pat #4370) #3176) -#4376 := (not #4371) -#4701 := (or #4376 #4698) -#4704 := (not #4701) -decl ?x29!1 :: T2 -#2026 := ?x29!1 -#2030 := (uf_12 ?x29!1) -#2647 := (* -1::int #2030) -decl ?x28!2 :: T2 -#2027 := ?x28!2 -#2028 := (uf_12 ?x28!2) -#2648 := (+ #2028 #2647) -#2649 := (<= #2648 0::int) -#2034 := (up_13 ?x29!1) -#2033 := (up_13 ?x28!2) -#2266 := (not #2033) -#2166 := (or #2266 #2034 #2649) -#6004 := [hypothesis]: #2033 -#4349 := (forall (vars (?x26 T2)) (:pat #4348) #78) -#79 := (forall (vars (?x26 T2)) #78) -#4352 := (iff #79 #4349) -#4350 := (iff #78 #78) -#4351 := [refl]: #4350 -#4353 := [quant-intro #4351]: #4352 -#1965 := (~ #79 #79) -#2002 := (~ #78 #78) -#2003 := [refl]: #2002 -#1966 := [nnf-pos #2003]: #1965 -#70 := (= #69 0::int) -#73 := (not #68) -#1912 := (or #73 #70) -#1915 := (forall (vars (?x24 T2)) #1912) -#1918 := (not #1915) -#1846 := (forall (vars (?x34 T2)) #1843) -#1849 := (not #1846) -#511 := (and #502 #508) -#517 := (not #511) -#1832 := (or #517 #1416) -#1837 := (forall (vars (?x35 T2) (?x36 T2)) #1832) -#1840 := (not #1837) -#1348 := (not #1347) -#1807 := (and #505 #1348) -#1812 := (not #1807) -#1818 := (or #1812 #1815) -#1821 := (forall (vars (?x37 T2) (?x38 T2)) #1818) -#1824 := (not #1821) -#1710 := (not #1709) -#1744 := (and #1348 #1710) -#1747 := (not #1744) -#1753 := (or #1747 #1750) -#1756 := (forall (vars (?x66 T2) (?x67 T2)) #1753) -#1759 := (not #1756) -#1767 := (or #251 #1759) -#1772 := (and #1756 #1767) -#1725 := (= #1727 0::int) -#1719 := (>= #1721 0::int) -#1722 := (not #1719) -#1729 := (and #1722 #1725) -#1732 := (exists (vars (?x65 T2)) #1729) -#1713 := (and #73 #1710) -#1716 := (not #1713) -#1735 := (or #1716 #1732) -#1738 := (forall (vars (?x64 T2)) #1735) -#1741 := (not #1738) -#1775 := (or #1741 #1772) -#1778 := (and #1738 #1775) -#1407 := (not #1406) -#1670 := (and #508 #1407) -#1675 := (exists (vars (?x41 T2)) #1670) -#1796 := (or #1062 #1044 #1071 #1053 #1675 #1778) -#1678 := (not #1675) -#1649 := (forall (vars (?x42 T2)) #1644) -#1652 := (not #1649) -#1459 := (not #1458) -#1454 := (not #1453) -#1462 := (and #1454 #1459) -#1620 := (not #1462) -#1628 := (or #1620 #1625) -#1631 := (forall (vars (?x43 T2)) #1628) -#1634 := (not #1631) -#1561 := (= #1536 0::int) -#1558 := (not #1504) -#1570 := (and #630 #1558 #1561) -#1575 := (exists (vars (?x53 T2)) #1570) -#1547 := (+ uf_9 #1480) -#1548 := (<= #1547 0::int) -#1549 := (not #1548) -#1552 := (and #73 #1549) -#1555 := (not #1552) -#1578 := (or #1555 #1575) -#1581 := (forall (vars (?x52 T2)) #1578) -#1526 := (and #630 #1348) -#1531 := (not #1526) -#1538 := (or #1531 #1534) -#1541 := (forall (vars (?x50 T2) (?x51 T2)) #1538) -#1544 := (not #1541) -#1584 := (or #1544 #1581) -#1587 := (and #1541 #1584) -#656 := (and #636 #648) -#664 := (not #656) -#1512 := (or #664 #1504) -#1517 := (forall (vars (?x48 T2) (?x49 T2)) #1512) -#1520 := (not #1517) -#1590 := (or #1520 #1587) -#1593 := (and #1517 #1590) -#1498 := (forall (vars (?x47 T2)) #1495) -#1501 := (not #1498) -#1596 := (or #1501 #1593) -#1599 := (and #1498 #1596) -#1602 := (or #1492 #1599) -#1605 := (and #173 #1602) -#642 := (forall (vars (?x46 T2)) #637) -#824 := (not #642) -#1608 := (or #824 #1605) -#1611 := (and #642 #1608) -#1484 := (forall (vars (?x45 T2)) #1479) -#1487 := (not #1484) -#1614 := (or #1487 #1611) -#1617 := (and #1484 #1614) -#1468 := (or #616 #1462) -#1473 := (forall (vars (?x44 T2)) #1468) -#1476 := (not #1473) -#1702 := (or #876 #585 #1476 #1617 #1634 #1652 #1656 #1678) -#1801 := (and #1702 #1796) -#1422 := (= #1424 0::int) -#1419 := (not #1416) -#1432 := (and #505 #1419 #1422) -#1437 := (exists (vars (?x40 T2)) #1432) -#1410 := (and #73 #1407) -#1413 := (not #1410) -#1440 := (or #1413 #1437) -#1443 := (forall (vars (?x39 T2)) #1440) -#1446 := (not #1443) -#1389 := (not #1388) -#1392 := (and #73 #1389) -#1395 := (not #1392) -#1370 := (= #1345 0::int) -#1366 := (not #1322) -#1378 := (and #77 #1366 #1370) -#1383 := (exists (vars (?x33 T2)) #1378) -#1398 := (or #1383 #1395) -#1401 := (forall (vars (?x32 T2)) #1398) -#1857 := (not #1401) -#1878 := (or #1854 #1857 #1446 #1801 #1824 #1840 #1849) -#1883 := (and #1401 #1878) -#1351 := (and #77 #1348) -#1354 := (not #1351) -#1357 := (or #1342 #1354) -#1360 := (forall (vars (?x30 T2) (?x31 T2)) #1357) -#1363 := (not #1360) -#1886 := (or #1363 #1883) -#1889 := (and #1360 #1886) -#454 := (and #78 #84) -#460 := (not #454) -#1329 := (or #460 #1322) -#1334 := (forall (vars (?x28 T2) (?x29 T2)) #1329) -#1337 := (not #1334) -#1892 := (or #1337 #1889) -#1895 := (and #1334 #1892) -#1313 := (>= #69 0::int) -#1314 := (forall (vars (?x27 T2)) #1313) -#1317 := (not #1314) -#1898 := (or #1317 #1895) -#1901 := (and #1314 #1898) -#80 := (uf_12 uf_11) -#81 := (= #80 0::int) -#1308 := (not #81) -#1904 := (or #1308 #1901) -#1907 := (and #81 #1904) -#437 := (= uf_9 #69) -#443 := (or #68 #437) -#448 := (forall (vars (?x25 T2)) #443) -#1277 := (not #448) -#1268 := (not #79) -#1930 := (or #1268 #1277 #1907 #1918) -#1935 := (not #1930) -#82 := (<= 0::int #69) -#83 := (forall (vars (?x27 T2)) #82) -#87 := (<= #86 #69) -#85 := (and #84 #78) -#88 := (implies #85 #87) -#89 := (forall (vars (?x28 T2) (?x29 T2)) #88) -#94 := (+ #69 #91) -#95 := (<= #86 #94) -#92 := (< #91 uf_9) -#93 := (and #92 #77) -#96 := (implies #93 #95) -#97 := (forall (vars (?x30 T2) (?x31 T2)) #96) -#101 := (< #69 #86) -#102 := (and #77 #101) -#100 := (= #86 #94) -#103 := (and #100 #102) -#104 := (exists (vars (?x33 T2)) #103) -#98 := (< #69 uf_9) -#99 := (and #98 #73) -#105 := (implies #99 #104) -#106 := (forall (vars (?x32 T2)) #105) -#241 := (< #235 #238) -#239 := (+ #235 #91) -#240 := (= #238 #239) -#242 := (and #240 #241) -#243 := (exists (vars (?x65 T2)) #242) -#236 := (< #235 uf_9) -#237 := (and #236 #73) -#244 := (implies #237 #243) -#245 := (forall (vars (?x64 T2)) #244) -#247 := (<= #238 #239) -#246 := (and #92 #236) -#248 := (implies #246 #247) -#249 := (forall (vars (?x66 T2) (?x67 T2)) #248) -#1 := true -#252 := (implies #251 true) -#253 := (and #252 #251) -#254 := (implies #249 #253) -#255 := (and #254 #249) -#256 := (implies #245 #255) -#257 := (and #256 #245) -#258 := (implies true #257) -#259 := (implies #234 #258) -#231 := (= uf_22 uf_14) -#260 := (implies #231 #259) -#261 := (implies #229 #260) -#226 := (= uf_19 uf_15) -#262 := (implies #226 #261) -#263 := (implies true #262) -#264 := (implies true #263) -#117 := (= #116 uf_8) -#118 := (not #117) -#129 := (< #110 uf_9) -#138 := (and #129 #118) -#139 := (exists (vars (?x41 T2)) #138) -#224 := (not #139) -#265 := (implies #224 #264) -#266 := (implies true #265) -#267 := (implies true #266) -#166 := (<= #158 #110) -#167 := (forall (vars (?x45 T2)) #166) -#163 := (= #158 #110) -#169 := (= #168 uf_8) -#170 := (implies #169 #163) -#171 := (forall (vars (?x46 T2)) #170) -#174 := (<= 0::int #158) -#175 := (forall (vars (?x47 T2)) #174) -#181 := (<= #180 #158) -#178 := (not #169) -#177 := (= #176 uf_8) -#179 := (and #177 #178) -#182 := (implies #179 #181) -#183 := (forall (vars (?x48 T2) (?x49 T2)) #182) -#185 := (+ #158 #91) -#186 := (<= #180 #185) -#184 := (and #92 #169) -#187 := (implies #184 #186) -#188 := (forall (vars (?x50 T2) (?x51 T2)) #187) -#192 := (< #158 #180) -#193 := (and #169 #192) -#191 := (= #180 #185) -#194 := (and #191 #193) -#195 := (exists (vars (?x53 T2)) #194) -#189 := (< #158 uf_9) -#190 := (and #189 #73) -#196 := (implies #190 #195) -#197 := (forall (vars (?x52 T2)) #196) -#198 := (implies false true) -#199 := (implies #197 #198) -#200 := (and #199 #197) -#201 := (implies #188 #200) -#202 := (and #201 #188) -#203 := (implies #183 #202) -#204 := (and #203 #183) -#205 := (implies #175 #204) -#206 := (and #205 #175) -#207 := (implies #173 #206) -#208 := (and #207 #173) -#209 := (implies true #208) -#210 := (implies true #209) -#211 := (implies #171 #210) -#212 := (and #211 #171) -#213 := (implies #167 #212) -#214 := (and #213 #167) -#156 := (< #153 uf_9) -#154 := (+ #144 #153) -#155 := (< #154 #110) -#157 := (and #155 #156) -#162 := (not #157) -#164 := (implies #162 #163) -#165 := (forall (vars (?x44 T2)) #164) -#215 := (implies #165 #214) -#159 := (= #158 #154) -#160 := (implies #157 #159) -#161 := (forall (vars (?x43 T2)) #160) -#216 := (implies #161 #215) -#217 := (implies #151 #216) -#146 := (<= #144 #110) -#147 := (implies #118 #146) -#148 := (forall (vars (?x42 T2)) #147) -#218 := (implies #148 #217) -#145 := (< #144 uf_9) -#219 := (implies #145 #218) -#142 := (= #141 uf_8) -#143 := (not #142) -#220 := (implies #143 #219) -#221 := (implies #139 #220) -#222 := (implies true #221) -#223 := (implies true #222) -#268 := (and #223 #267) -#269 := (implies true #268) -#132 := (< #110 #120) -#133 := (and #117 #132) -#125 := (+ #110 #91) -#131 := (= #120 #125) -#134 := (and #131 #133) -#135 := (exists (vars (?x40 T2)) #134) -#130 := (and #129 #73) -#136 := (implies #130 #135) -#137 := (forall (vars (?x39 T2)) #136) -#270 := (implies #137 #269) -#126 := (<= #120 #125) -#124 := (and #92 #117) -#127 := (implies #124 #126) -#128 := (forall (vars (?x37 T2) (?x38 T2)) #127) -#271 := (implies #128 #270) -#121 := (<= #120 #110) -#115 := (= #114 uf_8) -#119 := (and #115 #118) -#122 := (implies #119 #121) -#123 := (forall (vars (?x35 T2) (?x36 T2)) #122) -#272 := (implies #123 #271) -#111 := (<= 0::int #110) -#112 := (forall (vars (?x34 T2)) #111) -#273 := (implies #112 #272) -#274 := (implies #109 #273) -#275 := (implies true #274) -#276 := (implies true #275) -#277 := (implies #106 #276) -#278 := (and #277 #106) -#279 := (implies #97 #278) -#280 := (and #279 #97) -#281 := (implies #89 #280) -#282 := (and #281 #89) -#283 := (implies #83 #282) -#284 := (and #283 #83) -#285 := (implies #81 #284) -#286 := (and #285 #81) -#287 := (implies true #286) -#288 := (implies #79 #287) -#74 := (= #69 uf_9) -#75 := (implies #73 #74) -#76 := (forall (vars (?x25 T2)) #75) -#289 := (implies #76 #288) -#71 := (implies #68 #70) -#72 := (forall (vars (?x24 T2)) #71) -#290 := (implies #72 #289) -#291 := (implies true #290) -#292 := (implies true #291) -#293 := (not #292) -#1938 := (iff #293 #1935) -#983 := (= 0::int #250) -#939 := (+ #91 #235) -#968 := (<= #238 #939) -#974 := (not #246) -#975 := (or #974 #968) -#980 := (forall (vars (?x66 T2) (?x67 T2)) #975) -#1003 := (not #980) -#1004 := (or #1003 #983) -#1012 := (and #980 #1004) -#942 := (= #238 #939) -#948 := (and #241 #942) -#953 := (exists (vars (?x65 T2)) #948) -#936 := (and #73 #236) -#959 := (not #936) -#960 := (or #959 #953) -#965 := (forall (vars (?x64 T2)) #960) -#1020 := (not #965) -#1021 := (or #1020 #1012) -#1029 := (and #965 #1021) -#1045 := (or #1044 #1029) -#1054 := (or #1053 #1045) -#1063 := (or #1062 #1054) -#1072 := (or #1071 #1063) -#579 := (and #129 #508) -#582 := (exists (vars (?x41 T2)) #579) -#1091 := (or #582 #1072) -#703 := (and #192 #630) -#676 := (+ #91 #158) -#697 := (= #180 #676) -#708 := (and #697 #703) -#711 := (exists (vars (?x53 T2)) #708) -#694 := (and #73 #189) -#717 := (not #694) -#718 := (or #717 #711) -#723 := (forall (vars (?x52 T2)) #718) -#679 := (<= #180 #676) -#673 := (and #92 #630) -#685 := (not #673) -#686 := (or #685 #679) -#691 := (forall (vars (?x50 T2) (?x51 T2)) #686) -#745 := (not #691) -#746 := (or #745 #723) -#754 := (and #691 #746) -#665 := (or #181 #664) -#670 := (forall (vars (?x48 T2) (?x49 T2)) #665) -#762 := (not #670) -#763 := (or #762 #754) -#771 := (and #670 #763) -#779 := (not #175) -#780 := (or #779 #771) -#788 := (and #175 #780) -#645 := (= 0::int #172) -#796 := (not #645) -#797 := (or #796 #788) -#805 := (and #645 #797) -#825 := (or #824 #805) -#833 := (and #642 #825) -#841 := (not #167) -#842 := (or #841 #833) -#850 := (and #167 #842) -#622 := (or #157 #616) -#627 := (forall (vars (?x44 T2)) #622) -#858 := (not #627) -#859 := (or #858 #850) -#602 := (= #154 #158) -#608 := (or #162 #602) -#613 := (forall (vars (?x43 T2)) #608) -#867 := (not #613) -#868 := (or #867 #859) -#877 := (or #876 #868) -#594 := (or #146 #505) -#599 := (forall (vars (?x42 T2)) #594) -#885 := (not #599) -#886 := (or #885 #877) -#894 := (not #145) -#895 := (or #894 #886) -#903 := (or #585 #895) -#911 := (not #582) -#912 := (or #911 #903) -#1107 := (and #912 #1091) -#556 := (and #132 #505) -#529 := (+ #91 #110) -#550 := (= #120 #529) -#561 := (and #550 #556) -#564 := (exists (vars (?x40 T2)) #561) -#547 := (and #73 #129) -#570 := (not #547) -#571 := (or #570 #564) -#576 := (forall (vars (?x39 T2)) #571) -#1120 := (not #576) -#1121 := (or #1120 #1107) -#532 := (<= #120 #529) -#526 := (and #92 #505) -#538 := (not #526) -#539 := (or #538 #532) -#544 := (forall (vars (?x37 T2) (?x38 T2)) #539) -#1129 := (not #544) -#1130 := (or #1129 #1121) -#518 := (or #121 #517) -#523 := (forall (vars (?x35 T2) (?x36 T2)) #518) -#1138 := (not #523) -#1139 := (or #1138 #1130) -#1147 := (not #112) -#1148 := (or #1147 #1139) -#499 := (= 0::int #108) -#1156 := (not #499) -#1157 := (or #1156 #1148) -#484 := (and #73 #98) -#490 := (not #484) -#491 := (or #104 #490) -#496 := (forall (vars (?x32 T2)) #491) -#1176 := (not #496) -#1177 := (or #1176 #1157) -#1185 := (and #496 #1177) -#469 := (and #77 #92) -#475 := (not #469) -#476 := (or #95 #475) -#481 := (forall (vars (?x30 T2) (?x31 T2)) #476) -#1193 := (not #481) -#1194 := (or #1193 #1185) -#1202 := (and #481 #1194) -#461 := (or #87 #460) -#466 := (forall (vars (?x28 T2) (?x29 T2)) #461) -#1210 := (not #466) -#1211 := (or #1210 #1202) -#1219 := (and #466 #1211) -#1227 := (not #83) -#1228 := (or #1227 #1219) -#1236 := (and #83 #1228) -#451 := (= 0::int #80) -#1244 := (not #451) -#1245 := (or #1244 #1236) -#1253 := (and #451 #1245) -#1269 := (or #1268 #1253) -#1278 := (or #1277 #1269) -#423 := (= 0::int #69) -#429 := (or #73 #423) -#434 := (forall (vars (?x24 T2)) #429) -#1286 := (not #434) -#1287 := (or #1286 #1278) -#1303 := (not #1287) -#1936 := (iff #1303 #1935) -#1933 := (iff #1287 #1930) -#1921 := (or #1268 #1907) -#1924 := (or #1277 #1921) -#1927 := (or #1918 #1924) -#1931 := (iff #1927 #1930) -#1932 := [rewrite]: #1931 -#1928 := (iff #1287 #1927) -#1925 := (iff #1278 #1924) -#1922 := (iff #1269 #1921) -#1908 := (iff #1253 #1907) -#1905 := (iff #1245 #1904) -#1902 := (iff #1236 #1901) -#1899 := (iff #1228 #1898) -#1896 := (iff #1219 #1895) -#1893 := (iff #1211 #1892) -#1890 := (iff #1202 #1889) -#1887 := (iff #1194 #1886) -#1884 := (iff #1185 #1883) -#1881 := (iff #1177 #1878) -#1860 := (or #1446 #1801) -#1863 := (or #1824 #1860) -#1866 := (or #1840 #1863) -#1869 := (or #1849 #1866) -#1872 := (or #1854 #1869) -#1875 := (or #1857 #1872) -#1879 := (iff #1875 #1878) -#1880 := [rewrite]: #1879 -#1876 := (iff #1177 #1875) -#1873 := (iff #1157 #1872) -#1870 := (iff #1148 #1869) -#1867 := (iff #1139 #1866) -#1864 := (iff #1130 #1863) -#1861 := (iff #1121 #1860) -#1802 := (iff #1107 #1801) -#1799 := (iff #1091 #1796) -#1781 := (or #1044 #1778) -#1784 := (or #1053 #1781) -#1787 := (or #1062 #1784) -#1790 := (or #1071 #1787) -#1793 := (or #1675 #1790) -#1797 := (iff #1793 #1796) -#1798 := [rewrite]: #1797 -#1794 := (iff #1091 #1793) -#1791 := (iff #1072 #1790) -#1788 := (iff #1063 #1787) -#1785 := (iff #1054 #1784) -#1782 := (iff #1045 #1781) -#1779 := (iff #1029 #1778) -#1776 := (iff #1021 #1775) -#1773 := (iff #1012 #1772) -#1770 := (iff #1004 #1767) -#1764 := (or #1759 #251) -#1768 := (iff #1764 #1767) -#1769 := [rewrite]: #1768 -#1765 := (iff #1004 #1764) -#1762 := (iff #983 #251) -#1763 := [rewrite]: #1762 -#1760 := (iff #1003 #1759) -#1757 := (iff #980 #1756) -#1754 := (iff #975 #1753) -#1751 := (iff #968 #1750) -#1752 := [rewrite]: #1751 -#1748 := (iff #974 #1747) -#1745 := (iff #246 #1744) -#1711 := (iff #236 #1710) -#1712 := [rewrite]: #1711 -#1349 := (iff #92 #1348) -#1350 := [rewrite]: #1349 -#1746 := [monotonicity #1350 #1712]: #1745 -#1749 := [monotonicity #1746]: #1748 -#1755 := [monotonicity #1749 #1752]: #1754 -#1758 := [quant-intro #1755]: #1757 -#1761 := [monotonicity #1758]: #1760 -#1766 := [monotonicity #1761 #1763]: #1765 -#1771 := [trans #1766 #1769]: #1770 -#1774 := [monotonicity #1758 #1771]: #1773 -#1742 := (iff #1020 #1741) -#1739 := (iff #965 #1738) -#1736 := (iff #960 #1735) -#1733 := (iff #953 #1732) -#1730 := (iff #948 #1729) -#1726 := (iff #942 #1725) -#1728 := [rewrite]: #1726 -#1723 := (iff #241 #1722) -#1724 := [rewrite]: #1723 -#1731 := [monotonicity #1724 #1728]: #1730 -#1734 := [quant-intro #1731]: #1733 -#1717 := (iff #959 #1716) -#1714 := (iff #936 #1713) -#1715 := [monotonicity #1712]: #1714 -#1718 := [monotonicity #1715]: #1717 -#1737 := [monotonicity #1718 #1734]: #1736 -#1740 := [quant-intro #1737]: #1739 -#1743 := [monotonicity #1740]: #1742 -#1777 := [monotonicity #1743 #1774]: #1776 -#1780 := [monotonicity #1740 #1777]: #1779 -#1783 := [monotonicity #1780]: #1782 -#1786 := [monotonicity #1783]: #1785 -#1789 := [monotonicity #1786]: #1788 -#1792 := [monotonicity #1789]: #1791 -#1676 := (iff #582 #1675) -#1673 := (iff #579 #1670) -#1667 := (and #1407 #508) -#1671 := (iff #1667 #1670) -#1672 := [rewrite]: #1671 -#1668 := (iff #579 #1667) -#1408 := (iff #129 #1407) -#1409 := [rewrite]: #1408 -#1669 := [monotonicity #1409]: #1668 -#1674 := [trans #1669 #1672]: #1673 -#1677 := [quant-intro #1674]: #1676 -#1795 := [monotonicity #1677 #1792]: #1794 -#1800 := [trans #1795 #1798]: #1799 -#1705 := (iff #912 #1702) -#1681 := (or #1476 #1617) -#1684 := (or #1634 #1681) -#1687 := (or #876 #1684) -#1690 := (or #1652 #1687) -#1693 := (or #1656 #1690) -#1696 := (or #585 #1693) -#1699 := (or #1678 #1696) -#1703 := (iff #1699 #1702) -#1704 := [rewrite]: #1703 -#1700 := (iff #912 #1699) -#1697 := (iff #903 #1696) -#1694 := (iff #895 #1693) -#1691 := (iff #886 #1690) -#1688 := (iff #877 #1687) -#1685 := (iff #868 #1684) -#1682 := (iff #859 #1681) -#1618 := (iff #850 #1617) -#1615 := (iff #842 #1614) -#1612 := (iff #833 #1611) -#1609 := (iff #825 #1608) -#1606 := (iff #805 #1605) -#1603 := (iff #797 #1602) -#1600 := (iff #788 #1599) -#1597 := (iff #780 #1596) -#1594 := (iff #771 #1593) -#1591 := (iff #763 #1590) -#1588 := (iff #754 #1587) -#1585 := (iff #746 #1584) -#1582 := (iff #723 #1581) -#1579 := (iff #718 #1578) -#1576 := (iff #711 #1575) -#1573 := (iff #708 #1570) -#1564 := (and #1558 #630) -#1567 := (and #1561 #1564) -#1571 := (iff #1567 #1570) -#1572 := [rewrite]: #1571 -#1568 := (iff #708 #1567) -#1565 := (iff #703 #1564) -#1559 := (iff #192 #1558) -#1560 := [rewrite]: #1559 -#1566 := [monotonicity #1560]: #1565 -#1562 := (iff #697 #1561) -#1563 := [rewrite]: #1562 -#1569 := [monotonicity #1563 #1566]: #1568 -#1574 := [trans #1569 #1572]: #1573 -#1577 := [quant-intro #1574]: #1576 -#1556 := (iff #717 #1555) -#1553 := (iff #694 #1552) -#1550 := (iff #189 #1549) -#1551 := [rewrite]: #1550 -#1554 := [monotonicity #1551]: #1553 -#1557 := [monotonicity #1554]: #1556 -#1580 := [monotonicity #1557 #1577]: #1579 -#1583 := [quant-intro #1580]: #1582 -#1545 := (iff #745 #1544) -#1542 := (iff #691 #1541) -#1539 := (iff #686 #1538) -#1535 := (iff #679 #1534) -#1537 := [rewrite]: #1535 -#1532 := (iff #685 #1531) -#1529 := (iff #673 #1526) -#1523 := (and #1348 #630) -#1527 := (iff #1523 #1526) -#1528 := [rewrite]: #1527 -#1524 := (iff #673 #1523) -#1525 := [monotonicity #1350]: #1524 -#1530 := [trans #1525 #1528]: #1529 -#1533 := [monotonicity #1530]: #1532 -#1540 := [monotonicity #1533 #1537]: #1539 -#1543 := [quant-intro #1540]: #1542 -#1546 := [monotonicity #1543]: #1545 -#1586 := [monotonicity #1546 #1583]: #1585 -#1589 := [monotonicity #1543 #1586]: #1588 -#1521 := (iff #762 #1520) -#1518 := (iff #670 #1517) -#1515 := (iff #665 #1512) -#1509 := (or #1504 #664) -#1513 := (iff #1509 #1512) -#1514 := [rewrite]: #1513 -#1510 := (iff #665 #1509) -#1507 := (iff #181 #1504) -#1508 := [rewrite]: #1507 -#1511 := [monotonicity #1508]: #1510 -#1516 := [trans #1511 #1514]: #1515 -#1519 := [quant-intro #1516]: #1518 -#1522 := [monotonicity #1519]: #1521 -#1592 := [monotonicity #1522 #1589]: #1591 -#1595 := [monotonicity #1519 #1592]: #1594 -#1502 := (iff #779 #1501) -#1499 := (iff #175 #1498) -#1496 := (iff #174 #1495) -#1497 := [rewrite]: #1496 -#1500 := [quant-intro #1497]: #1499 -#1503 := [monotonicity #1500]: #1502 -#1598 := [monotonicity #1503 #1595]: #1597 -#1601 := [monotonicity #1500 #1598]: #1600 -#1493 := (iff #796 #1492) -#1490 := (iff #645 #173) -#1491 := [rewrite]: #1490 -#1494 := [monotonicity #1491]: #1493 -#1604 := [monotonicity #1494 #1601]: #1603 -#1607 := [monotonicity #1491 #1604]: #1606 -#1610 := [monotonicity #1607]: #1609 -#1613 := [monotonicity #1610]: #1612 -#1488 := (iff #841 #1487) -#1485 := (iff #167 #1484) -#1482 := (iff #166 #1479) -#1483 := [rewrite]: #1482 -#1486 := [quant-intro #1483]: #1485 -#1489 := [monotonicity #1486]: #1488 -#1616 := [monotonicity #1489 #1613]: #1615 -#1619 := [monotonicity #1486 #1616]: #1618 -#1477 := (iff #858 #1476) -#1474 := (iff #627 #1473) -#1471 := (iff #622 #1468) -#1465 := (or #1462 #616) -#1469 := (iff #1465 #1468) -#1470 := [rewrite]: #1469 -#1466 := (iff #622 #1465) -#1463 := (iff #157 #1462) -#1460 := (iff #156 #1459) -#1461 := [rewrite]: #1460 -#1455 := (iff #155 #1454) -#1456 := [rewrite]: #1455 -#1464 := [monotonicity #1456 #1461]: #1463 -#1467 := [monotonicity #1464]: #1466 -#1472 := [trans #1467 #1470]: #1471 -#1475 := [quant-intro #1472]: #1474 -#1478 := [monotonicity #1475]: #1477 -#1683 := [monotonicity #1478 #1619]: #1682 -#1635 := (iff #867 #1634) -#1632 := (iff #613 #1631) -#1629 := (iff #608 #1628) -#1626 := (iff #602 #1625) -#1627 := [rewrite]: #1626 -#1621 := (iff #162 #1620) -#1622 := [monotonicity #1464]: #1621 -#1630 := [monotonicity #1622 #1627]: #1629 -#1633 := [quant-intro #1630]: #1632 -#1636 := [monotonicity #1633]: #1635 -#1686 := [monotonicity #1636 #1683]: #1685 -#1689 := [monotonicity #1686]: #1688 -#1653 := (iff #885 #1652) -#1650 := (iff #599 #1649) -#1647 := (iff #594 #1644) -#1641 := (or #1637 #505) -#1645 := (iff #1641 #1644) -#1646 := [rewrite]: #1645 -#1642 := (iff #594 #1641) -#1639 := (iff #146 #1637) -#1640 := [rewrite]: #1639 -#1643 := [monotonicity #1640]: #1642 -#1648 := [trans #1643 #1646]: #1647 -#1651 := [quant-intro #1648]: #1650 -#1654 := [monotonicity #1651]: #1653 -#1692 := [monotonicity #1654 #1689]: #1691 -#1665 := (iff #894 #1656) -#1657 := (not #1656) -#1660 := (not #1657) -#1663 := (iff #1660 #1656) -#1664 := [rewrite]: #1663 -#1661 := (iff #894 #1660) -#1658 := (iff #145 #1657) -#1659 := [rewrite]: #1658 -#1662 := [monotonicity #1659]: #1661 -#1666 := [trans #1662 #1664]: #1665 -#1695 := [monotonicity #1666 #1692]: #1694 -#1698 := [monotonicity #1695]: #1697 -#1679 := (iff #911 #1678) -#1680 := [monotonicity #1677]: #1679 -#1701 := [monotonicity #1680 #1698]: #1700 -#1706 := [trans #1701 #1704]: #1705 -#1803 := [monotonicity #1706 #1800]: #1802 -#1447 := (iff #1120 #1446) -#1444 := (iff #576 #1443) -#1441 := (iff #571 #1440) -#1438 := (iff #564 #1437) -#1435 := (iff #561 #1432) -#1426 := (and #1419 #505) -#1429 := (and #1422 #1426) -#1433 := (iff #1429 #1432) -#1434 := [rewrite]: #1433 -#1430 := (iff #561 #1429) -#1427 := (iff #556 #1426) -#1420 := (iff #132 #1419) -#1421 := [rewrite]: #1420 -#1428 := [monotonicity #1421]: #1427 -#1423 := (iff #550 #1422) -#1425 := [rewrite]: #1423 -#1431 := [monotonicity #1425 #1428]: #1430 -#1436 := [trans #1431 #1434]: #1435 -#1439 := [quant-intro #1436]: #1438 -#1414 := (iff #570 #1413) -#1411 := (iff #547 #1410) -#1412 := [monotonicity #1409]: #1411 -#1415 := [monotonicity #1412]: #1414 -#1442 := [monotonicity #1415 #1439]: #1441 -#1445 := [quant-intro #1442]: #1444 -#1448 := [monotonicity #1445]: #1447 -#1862 := [monotonicity #1448 #1803]: #1861 -#1825 := (iff #1129 #1824) -#1822 := (iff #544 #1821) -#1819 := (iff #539 #1818) -#1816 := (iff #532 #1815) -#1817 := [rewrite]: #1816 -#1813 := (iff #538 #1812) -#1810 := (iff #526 #1807) -#1804 := (and #1348 #505) -#1808 := (iff #1804 #1807) -#1809 := [rewrite]: #1808 -#1805 := (iff #526 #1804) -#1806 := [monotonicity #1350]: #1805 -#1811 := [trans #1806 #1809]: #1810 -#1814 := [monotonicity #1811]: #1813 -#1820 := [monotonicity #1814 #1817]: #1819 -#1823 := [quant-intro #1820]: #1822 -#1826 := [monotonicity #1823]: #1825 -#1865 := [monotonicity #1826 #1862]: #1864 -#1841 := (iff #1138 #1840) -#1838 := (iff #523 #1837) -#1835 := (iff #518 #1832) -#1829 := (or #1416 #517) -#1833 := (iff #1829 #1832) -#1834 := [rewrite]: #1833 -#1830 := (iff #518 #1829) -#1827 := (iff #121 #1416) -#1828 := [rewrite]: #1827 -#1831 := [monotonicity #1828]: #1830 -#1836 := [trans #1831 #1834]: #1835 -#1839 := [quant-intro #1836]: #1838 -#1842 := [monotonicity #1839]: #1841 -#1868 := [monotonicity #1842 #1865]: #1867 -#1850 := (iff #1147 #1849) -#1847 := (iff #112 #1846) -#1844 := (iff #111 #1843) -#1845 := [rewrite]: #1844 -#1848 := [quant-intro #1845]: #1847 -#1851 := [monotonicity #1848]: #1850 -#1871 := [monotonicity #1851 #1868]: #1870 -#1855 := (iff #1156 #1854) -#1852 := (iff #499 #109) -#1853 := [rewrite]: #1852 -#1856 := [monotonicity #1853]: #1855 -#1874 := [monotonicity #1856 #1871]: #1873 -#1858 := (iff #1176 #1857) -#1402 := (iff #496 #1401) -#1399 := (iff #491 #1398) -#1396 := (iff #490 #1395) -#1393 := (iff #484 #1392) -#1390 := (iff #98 #1389) -#1391 := [rewrite]: #1390 -#1394 := [monotonicity #1391]: #1393 -#1397 := [monotonicity #1394]: #1396 -#1384 := (iff #104 #1383) -#1381 := (iff #103 #1378) -#1372 := (and #77 #1366) -#1375 := (and #1370 #1372) -#1379 := (iff #1375 #1378) -#1380 := [rewrite]: #1379 -#1376 := (iff #103 #1375) -#1373 := (iff #102 #1372) -#1367 := (iff #101 #1366) -#1368 := [rewrite]: #1367 -#1374 := [monotonicity #1368]: #1373 -#1369 := (iff #100 #1370) -#1371 := [rewrite]: #1369 -#1377 := [monotonicity #1371 #1374]: #1376 -#1382 := [trans #1377 #1380]: #1381 -#1385 := [quant-intro #1382]: #1384 -#1400 := [monotonicity #1385 #1397]: #1399 -#1403 := [quant-intro #1400]: #1402 -#1859 := [monotonicity #1403]: #1858 -#1877 := [monotonicity #1859 #1874]: #1876 -#1882 := [trans #1877 #1880]: #1881 -#1885 := [monotonicity #1403 #1882]: #1884 -#1364 := (iff #1193 #1363) -#1361 := (iff #481 #1360) -#1358 := (iff #476 #1357) -#1355 := (iff #475 #1354) -#1352 := (iff #469 #1351) -#1353 := [monotonicity #1350]: #1352 -#1356 := [monotonicity #1353]: #1355 -#1341 := (iff #95 #1342) -#1340 := [rewrite]: #1341 -#1359 := [monotonicity #1340 #1356]: #1358 -#1362 := [quant-intro #1359]: #1361 -#1365 := [monotonicity #1362]: #1364 -#1888 := [monotonicity #1365 #1885]: #1887 -#1891 := [monotonicity #1362 #1888]: #1890 -#1338 := (iff #1210 #1337) -#1335 := (iff #466 #1334) -#1332 := (iff #461 #1329) -#1326 := (or #1322 #460) -#1330 := (iff #1326 #1329) -#1331 := [rewrite]: #1330 -#1327 := (iff #461 #1326) -#1321 := (iff #87 #1322) -#1325 := [rewrite]: #1321 -#1328 := [monotonicity #1325]: #1327 -#1333 := [trans #1328 #1331]: #1332 -#1336 := [quant-intro #1333]: #1335 -#1339 := [monotonicity #1336]: #1338 -#1894 := [monotonicity #1339 #1891]: #1893 -#1897 := [monotonicity #1336 #1894]: #1896 -#1318 := (iff #1227 #1317) -#1315 := (iff #83 #1314) -#1311 := (iff #82 #1313) -#1312 := [rewrite]: #1311 -#1316 := [quant-intro #1312]: #1315 -#1319 := [monotonicity #1316]: #1318 -#1900 := [monotonicity #1319 #1897]: #1899 -#1903 := [monotonicity #1316 #1900]: #1902 -#1309 := (iff #1244 #1308) -#1306 := (iff #451 #81) -#1307 := [rewrite]: #1306 -#1310 := [monotonicity #1307]: #1309 -#1906 := [monotonicity #1310 #1903]: #1905 -#1909 := [monotonicity #1307 #1906]: #1908 -#1923 := [monotonicity #1909]: #1922 -#1926 := [monotonicity #1923]: #1925 -#1919 := (iff #1286 #1918) -#1916 := (iff #434 #1915) -#1913 := (iff #429 #1912) -#1910 := (iff #423 #70) -#1911 := [rewrite]: #1910 -#1914 := [monotonicity #1911]: #1913 -#1917 := [quant-intro #1914]: #1916 -#1920 := [monotonicity #1917]: #1919 -#1929 := [monotonicity #1920 #1926]: #1928 -#1934 := [trans #1929 #1932]: #1933 -#1937 := [monotonicity #1934]: #1936 -#1304 := (iff #293 #1303) -#1301 := (iff #292 #1287) -#1292 := (implies true #1287) -#1295 := (iff #1292 #1287) -#1296 := [rewrite]: #1295 -#1299 := (iff #292 #1292) -#1297 := (iff #291 #1287) -#1293 := (iff #291 #1292) -#1290 := (iff #290 #1287) -#1283 := (implies #434 #1278) -#1288 := (iff #1283 #1287) -#1289 := [rewrite]: #1288 -#1284 := (iff #290 #1283) -#1281 := (iff #289 #1278) -#1274 := (implies #448 #1269) -#1279 := (iff #1274 #1278) -#1280 := [rewrite]: #1279 -#1275 := (iff #289 #1274) -#1272 := (iff #288 #1269) -#1265 := (implies #79 #1253) -#1270 := (iff #1265 #1269) -#1271 := [rewrite]: #1270 -#1266 := (iff #288 #1265) -#1263 := (iff #287 #1253) -#1258 := (implies true #1253) -#1261 := (iff #1258 #1253) -#1262 := [rewrite]: #1261 -#1259 := (iff #287 #1258) -#1256 := (iff #286 #1253) -#1250 := (and #1245 #451) -#1254 := (iff #1250 #1253) -#1255 := [rewrite]: #1254 -#1251 := (iff #286 #1250) -#452 := (iff #81 #451) -#453 := [rewrite]: #452 -#1248 := (iff #285 #1245) -#1241 := (implies #451 #1236) -#1246 := (iff #1241 #1245) -#1247 := [rewrite]: #1246 -#1242 := (iff #285 #1241) -#1239 := (iff #284 #1236) -#1233 := (and #1228 #83) -#1237 := (iff #1233 #1236) -#1238 := [rewrite]: #1237 -#1234 := (iff #284 #1233) -#1231 := (iff #283 #1228) -#1224 := (implies #83 #1219) -#1229 := (iff #1224 #1228) -#1230 := [rewrite]: #1229 -#1225 := (iff #283 #1224) -#1222 := (iff #282 #1219) -#1216 := (and #1211 #466) -#1220 := (iff #1216 #1219) -#1221 := [rewrite]: #1220 -#1217 := (iff #282 #1216) -#467 := (iff #89 #466) -#464 := (iff #88 #461) -#457 := (implies #454 #87) -#462 := (iff #457 #461) -#463 := [rewrite]: #462 -#458 := (iff #88 #457) -#455 := (iff #85 #454) -#456 := [rewrite]: #455 -#459 := [monotonicity #456]: #458 -#465 := [trans #459 #463]: #464 -#468 := [quant-intro #465]: #467 -#1214 := (iff #281 #1211) -#1207 := (implies #466 #1202) -#1212 := (iff #1207 #1211) -#1213 := [rewrite]: #1212 -#1208 := (iff #281 #1207) -#1205 := (iff #280 #1202) -#1199 := (and #1194 #481) -#1203 := (iff #1199 #1202) -#1204 := [rewrite]: #1203 -#1200 := (iff #280 #1199) -#482 := (iff #97 #481) -#479 := (iff #96 #476) -#472 := (implies #469 #95) -#477 := (iff #472 #476) -#478 := [rewrite]: #477 -#473 := (iff #96 #472) -#470 := (iff #93 #469) -#471 := [rewrite]: #470 -#474 := [monotonicity #471]: #473 -#480 := [trans #474 #478]: #479 -#483 := [quant-intro #480]: #482 -#1197 := (iff #279 #1194) -#1190 := (implies #481 #1185) -#1195 := (iff #1190 #1194) -#1196 := [rewrite]: #1195 -#1191 := (iff #279 #1190) -#1188 := (iff #278 #1185) -#1182 := (and #1177 #496) -#1186 := (iff #1182 #1185) -#1187 := [rewrite]: #1186 -#1183 := (iff #278 #1182) -#497 := (iff #106 #496) -#494 := (iff #105 #491) -#487 := (implies #484 #104) -#492 := (iff #487 #491) -#493 := [rewrite]: #492 -#488 := (iff #105 #487) -#485 := (iff #99 #484) -#486 := [rewrite]: #485 -#489 := [monotonicity #486]: #488 -#495 := [trans #489 #493]: #494 -#498 := [quant-intro #495]: #497 -#1180 := (iff #277 #1177) -#1173 := (implies #496 #1157) -#1178 := (iff #1173 #1177) -#1179 := [rewrite]: #1178 -#1174 := (iff #277 #1173) -#1171 := (iff #276 #1157) -#1162 := (implies true #1157) -#1165 := (iff #1162 #1157) -#1166 := [rewrite]: #1165 -#1169 := (iff #276 #1162) -#1167 := (iff #275 #1157) -#1163 := (iff #275 #1162) -#1160 := (iff #274 #1157) -#1153 := (implies #499 #1148) -#1158 := (iff #1153 #1157) -#1159 := [rewrite]: #1158 -#1154 := (iff #274 #1153) -#1151 := (iff #273 #1148) -#1144 := (implies #112 #1139) -#1149 := (iff #1144 #1148) -#1150 := [rewrite]: #1149 -#1145 := (iff #273 #1144) -#1142 := (iff #272 #1139) -#1135 := (implies #523 #1130) -#1140 := (iff #1135 #1139) -#1141 := [rewrite]: #1140 -#1136 := (iff #272 #1135) -#1133 := (iff #271 #1130) -#1126 := (implies #544 #1121) -#1131 := (iff #1126 #1130) -#1132 := [rewrite]: #1131 -#1127 := (iff #271 #1126) -#1124 := (iff #270 #1121) -#1117 := (implies #576 #1107) -#1122 := (iff #1117 #1121) -#1123 := [rewrite]: #1122 -#1118 := (iff #270 #1117) -#1115 := (iff #269 #1107) -#1110 := (implies true #1107) -#1113 := (iff #1110 #1107) -#1114 := [rewrite]: #1113 -#1111 := (iff #269 #1110) -#1108 := (iff #268 #1107) -#1105 := (iff #267 #1091) -#1096 := (implies true #1091) -#1099 := (iff #1096 #1091) -#1100 := [rewrite]: #1099 -#1103 := (iff #267 #1096) -#1101 := (iff #266 #1091) -#1097 := (iff #266 #1096) -#1094 := (iff #265 #1091) -#1088 := (implies #911 #1072) -#1092 := (iff #1088 #1091) -#1093 := [rewrite]: #1092 -#1089 := (iff #265 #1088) -#1086 := (iff #264 #1072) -#1077 := (implies true #1072) -#1080 := (iff #1077 #1072) -#1081 := [rewrite]: #1080 -#1084 := (iff #264 #1077) -#1082 := (iff #263 #1072) -#1078 := (iff #263 #1077) -#1075 := (iff #262 #1072) -#1068 := (implies #930 #1063) -#1073 := (iff #1068 #1072) -#1074 := [rewrite]: #1073 -#1069 := (iff #262 #1068) -#1066 := (iff #261 #1063) -#1059 := (implies #229 #1054) -#1064 := (iff #1059 #1063) -#1065 := [rewrite]: #1064 -#1060 := (iff #261 #1059) -#1057 := (iff #260 #1054) -#1050 := (implies #933 #1045) -#1055 := (iff #1050 #1054) -#1056 := [rewrite]: #1055 -#1051 := (iff #260 #1050) -#1048 := (iff #259 #1045) -#1041 := (implies #234 #1029) -#1046 := (iff #1041 #1045) -#1047 := [rewrite]: #1046 -#1042 := (iff #259 #1041) -#1039 := (iff #258 #1029) -#1034 := (implies true #1029) -#1037 := (iff #1034 #1029) -#1038 := [rewrite]: #1037 -#1035 := (iff #258 #1034) -#1032 := (iff #257 #1029) -#1026 := (and #1021 #965) -#1030 := (iff #1026 #1029) -#1031 := [rewrite]: #1030 -#1027 := (iff #257 #1026) -#966 := (iff #245 #965) -#963 := (iff #244 #960) -#956 := (implies #936 #953) -#961 := (iff #956 #960) -#962 := [rewrite]: #961 -#957 := (iff #244 #956) -#954 := (iff #243 #953) -#951 := (iff #242 #948) -#945 := (and #942 #241) -#949 := (iff #945 #948) -#950 := [rewrite]: #949 -#946 := (iff #242 #945) -#943 := (iff #240 #942) -#940 := (= #239 #939) -#941 := [rewrite]: #940 -#944 := [monotonicity #941]: #943 -#947 := [monotonicity #944]: #946 -#952 := [trans #947 #950]: #951 -#955 := [quant-intro #952]: #954 -#937 := (iff #237 #936) -#938 := [rewrite]: #937 -#958 := [monotonicity #938 #955]: #957 -#964 := [trans #958 #962]: #963 -#967 := [quant-intro #964]: #966 -#1024 := (iff #256 #1021) -#1017 := (implies #965 #1012) -#1022 := (iff #1017 #1021) -#1023 := [rewrite]: #1022 -#1018 := (iff #256 #1017) -#1015 := (iff #255 #1012) -#1009 := (and #1004 #980) -#1013 := (iff #1009 #1012) -#1014 := [rewrite]: #1013 -#1010 := (iff #255 #1009) -#981 := (iff #249 #980) -#978 := (iff #248 #975) -#971 := (implies #246 #968) -#976 := (iff #971 #975) -#977 := [rewrite]: #976 -#972 := (iff #248 #971) -#969 := (iff #247 #968) -#970 := [monotonicity #941]: #969 -#973 := [monotonicity #970]: #972 -#979 := [trans #973 #977]: #978 -#982 := [quant-intro #979]: #981 -#1007 := (iff #254 #1004) -#1000 := (implies #980 #983) -#1005 := (iff #1000 #1004) -#1006 := [rewrite]: #1005 -#1001 := (iff #254 #1000) -#998 := (iff #253 #983) -#993 := (and true #983) -#996 := (iff #993 #983) -#997 := [rewrite]: #996 -#994 := (iff #253 #993) -#984 := (iff #251 #983) -#985 := [rewrite]: #984 -#991 := (iff #252 true) -#986 := (implies #983 true) -#989 := (iff #986 true) -#990 := [rewrite]: #989 -#987 := (iff #252 #986) -#988 := [monotonicity #985]: #987 -#992 := [trans #988 #990]: #991 -#995 := [monotonicity #992 #985]: #994 -#999 := [trans #995 #997]: #998 -#1002 := [monotonicity #982 #999]: #1001 -#1008 := [trans #1002 #1006]: #1007 -#1011 := [monotonicity #1008 #982]: #1010 -#1016 := [trans #1011 #1014]: #1015 -#1019 := [monotonicity #967 #1016]: #1018 -#1025 := [trans #1019 #1023]: #1024 -#1028 := [monotonicity #1025 #967]: #1027 -#1033 := [trans #1028 #1031]: #1032 -#1036 := [monotonicity #1033]: #1035 -#1040 := [trans #1036 #1038]: #1039 -#1043 := [monotonicity #1040]: #1042 -#1049 := [trans #1043 #1047]: #1048 -#934 := (iff #231 #933) -#935 := [rewrite]: #934 -#1052 := [monotonicity #935 #1049]: #1051 -#1058 := [trans #1052 #1056]: #1057 -#1061 := [monotonicity #1058]: #1060 -#1067 := [trans #1061 #1065]: #1066 -#931 := (iff #226 #930) -#932 := [rewrite]: #931 -#1070 := [monotonicity #932 #1067]: #1069 -#1076 := [trans #1070 #1074]: #1075 -#1079 := [monotonicity #1076]: #1078 -#1083 := [trans #1079 #1081]: #1082 -#1085 := [monotonicity #1083]: #1084 -#1087 := [trans #1085 #1081]: #1086 -#928 := (iff #224 #911) -#583 := (iff #139 #582) -#580 := (iff #138 #579) -#509 := (iff #118 #508) -#506 := (iff #117 #505) -#507 := [rewrite]: #506 -#510 := [monotonicity #507]: #509 -#581 := [monotonicity #510]: #580 -#584 := [quant-intro #581]: #583 -#929 := [monotonicity #584]: #928 -#1090 := [monotonicity #929 #1087]: #1089 -#1095 := [trans #1090 #1093]: #1094 -#1098 := [monotonicity #1095]: #1097 -#1102 := [trans #1098 #1100]: #1101 -#1104 := [monotonicity #1102]: #1103 -#1106 := [trans #1104 #1100]: #1105 -#926 := (iff #223 #912) -#917 := (implies true #912) -#920 := (iff #917 #912) -#921 := [rewrite]: #920 -#924 := (iff #223 #917) -#922 := (iff #222 #912) -#918 := (iff #222 #917) -#915 := (iff #221 #912) -#908 := (implies #582 #903) -#913 := (iff #908 #912) -#914 := [rewrite]: #913 -#909 := (iff #221 #908) -#906 := (iff #220 #903) -#588 := (not #585) -#900 := (implies #588 #895) -#904 := (iff #900 #903) -#905 := [rewrite]: #904 -#901 := (iff #220 #900) -#898 := (iff #219 #895) -#891 := (implies #145 #886) -#896 := (iff #891 #895) -#897 := [rewrite]: #896 -#892 := (iff #219 #891) -#889 := (iff #218 #886) -#882 := (implies #599 #877) -#887 := (iff #882 #886) -#888 := [rewrite]: #887 -#883 := (iff #218 #882) -#880 := (iff #217 #877) -#873 := (implies #151 #868) -#878 := (iff #873 #877) -#879 := [rewrite]: #878 -#874 := (iff #217 #873) -#871 := (iff #216 #868) -#864 := (implies #613 #859) -#869 := (iff #864 #868) -#870 := [rewrite]: #869 -#865 := (iff #216 #864) -#862 := (iff #215 #859) -#855 := (implies #627 #850) -#860 := (iff #855 #859) -#861 := [rewrite]: #860 -#856 := (iff #215 #855) -#853 := (iff #214 #850) -#847 := (and #842 #167) -#851 := (iff #847 #850) -#852 := [rewrite]: #851 -#848 := (iff #214 #847) -#845 := (iff #213 #842) -#838 := (implies #167 #833) -#843 := (iff #838 #842) -#844 := [rewrite]: #843 -#839 := (iff #213 #838) -#836 := (iff #212 #833) -#830 := (and #825 #642) -#834 := (iff #830 #833) -#835 := [rewrite]: #834 -#831 := (iff #212 #830) -#643 := (iff #171 #642) -#640 := (iff #170 #637) -#633 := (implies #630 #616) -#638 := (iff #633 #637) -#639 := [rewrite]: #638 -#634 := (iff #170 #633) -#617 := (iff #163 #616) -#618 := [rewrite]: #617 -#631 := (iff #169 #630) -#632 := [rewrite]: #631 -#635 := [monotonicity #632 #618]: #634 -#641 := [trans #635 #639]: #640 -#644 := [quant-intro #641]: #643 -#828 := (iff #211 #825) -#821 := (implies #642 #805) -#826 := (iff #821 #825) -#827 := [rewrite]: #826 -#822 := (iff #211 #821) -#819 := (iff #210 #805) -#810 := (implies true #805) -#813 := (iff #810 #805) -#814 := [rewrite]: #813 -#817 := (iff #210 #810) -#815 := (iff #209 #805) -#811 := (iff #209 #810) -#808 := (iff #208 #805) -#802 := (and #797 #645) -#806 := (iff #802 #805) -#807 := [rewrite]: #806 -#803 := (iff #208 #802) -#646 := (iff #173 #645) -#647 := [rewrite]: #646 -#800 := (iff #207 #797) -#793 := (implies #645 #788) -#798 := (iff #793 #797) -#799 := [rewrite]: #798 -#794 := (iff #207 #793) -#791 := (iff #206 #788) -#785 := (and #780 #175) -#789 := (iff #785 #788) -#790 := [rewrite]: #789 -#786 := (iff #206 #785) -#783 := (iff #205 #780) -#776 := (implies #175 #771) -#781 := (iff #776 #780) -#782 := [rewrite]: #781 -#777 := (iff #205 #776) -#774 := (iff #204 #771) -#768 := (and #763 #670) -#772 := (iff #768 #771) -#773 := [rewrite]: #772 -#769 := (iff #204 #768) -#671 := (iff #183 #670) -#668 := (iff #182 #665) -#661 := (implies #656 #181) -#666 := (iff #661 #665) -#667 := [rewrite]: #666 -#662 := (iff #182 #661) -#659 := (iff #179 #656) -#653 := (and #648 #636) -#657 := (iff #653 #656) -#658 := [rewrite]: #657 -#654 := (iff #179 #653) -#651 := (iff #178 #636) -#652 := [monotonicity #632]: #651 -#649 := (iff #177 #648) -#650 := [rewrite]: #649 -#655 := [monotonicity #650 #652]: #654 -#660 := [trans #655 #658]: #659 -#663 := [monotonicity #660]: #662 -#669 := [trans #663 #667]: #668 -#672 := [quant-intro #669]: #671 -#766 := (iff #203 #763) -#759 := (implies #670 #754) -#764 := (iff #759 #763) -#765 := [rewrite]: #764 -#760 := (iff #203 #759) -#757 := (iff #202 #754) -#751 := (and #746 #691) -#755 := (iff #751 #754) -#756 := [rewrite]: #755 -#752 := (iff #202 #751) -#692 := (iff #188 #691) -#689 := (iff #187 #686) -#682 := (implies #673 #679) -#687 := (iff #682 #686) -#688 := [rewrite]: #687 -#683 := (iff #187 #682) -#680 := (iff #186 #679) -#677 := (= #185 #676) -#678 := [rewrite]: #677 -#681 := [monotonicity #678]: #680 -#674 := (iff #184 #673) -#675 := [monotonicity #632]: #674 -#684 := [monotonicity #675 #681]: #683 -#690 := [trans #684 #688]: #689 -#693 := [quant-intro #690]: #692 -#749 := (iff #201 #746) -#742 := (implies #691 #723) -#747 := (iff #742 #746) -#748 := [rewrite]: #747 -#743 := (iff #201 #742) -#740 := (iff #200 #723) -#735 := (and true #723) -#738 := (iff #735 #723) -#739 := [rewrite]: #738 -#736 := (iff #200 #735) -#724 := (iff #197 #723) -#721 := (iff #196 #718) -#714 := (implies #694 #711) -#719 := (iff #714 #718) -#720 := [rewrite]: #719 -#715 := (iff #196 #714) -#712 := (iff #195 #711) -#709 := (iff #194 #708) -#706 := (iff #193 #703) -#700 := (and #630 #192) -#704 := (iff #700 #703) -#705 := [rewrite]: #704 -#701 := (iff #193 #700) -#702 := [monotonicity #632]: #701 -#707 := [trans #702 #705]: #706 -#698 := (iff #191 #697) -#699 := [monotonicity #678]: #698 -#710 := [monotonicity #699 #707]: #709 -#713 := [quant-intro #710]: #712 -#695 := (iff #190 #694) -#696 := [rewrite]: #695 -#716 := [monotonicity #696 #713]: #715 -#722 := [trans #716 #720]: #721 -#725 := [quant-intro #722]: #724 -#733 := (iff #199 true) -#728 := (implies #723 true) -#731 := (iff #728 true) -#732 := [rewrite]: #731 -#729 := (iff #199 #728) -#726 := (iff #198 true) -#727 := [rewrite]: #726 -#730 := [monotonicity #725 #727]: #729 -#734 := [trans #730 #732]: #733 -#737 := [monotonicity #734 #725]: #736 -#741 := [trans #737 #739]: #740 -#744 := [monotonicity #693 #741]: #743 -#750 := [trans #744 #748]: #749 -#753 := [monotonicity #750 #693]: #752 -#758 := [trans #753 #756]: #757 -#761 := [monotonicity #672 #758]: #760 -#767 := [trans #761 #765]: #766 -#770 := [monotonicity #767 #672]: #769 -#775 := [trans #770 #773]: #774 -#778 := [monotonicity #775]: #777 -#784 := [trans #778 #782]: #783 -#787 := [monotonicity #784]: #786 -#792 := [trans #787 #790]: #791 -#795 := [monotonicity #647 #792]: #794 -#801 := [trans #795 #799]: #800 -#804 := [monotonicity #801 #647]: #803 -#809 := [trans #804 #807]: #808 -#812 := [monotonicity #809]: #811 -#816 := [trans #812 #814]: #815 -#818 := [monotonicity #816]: #817 -#820 := [trans #818 #814]: #819 -#823 := [monotonicity #644 #820]: #822 -#829 := [trans #823 #827]: #828 -#832 := [monotonicity #829 #644]: #831 -#837 := [trans #832 #835]: #836 -#840 := [monotonicity #837]: #839 -#846 := [trans #840 #844]: #845 -#849 := [monotonicity #846]: #848 -#854 := [trans #849 #852]: #853 -#628 := (iff #165 #627) -#625 := (iff #164 #622) -#619 := (implies #162 #616) -#623 := (iff #619 #622) -#624 := [rewrite]: #623 -#620 := (iff #164 #619) -#621 := [monotonicity #618]: #620 -#626 := [trans #621 #624]: #625 -#629 := [quant-intro #626]: #628 -#857 := [monotonicity #629 #854]: #856 -#863 := [trans #857 #861]: #862 -#614 := (iff #161 #613) -#611 := (iff #160 #608) -#605 := (implies #157 #602) -#609 := (iff #605 #608) -#610 := [rewrite]: #609 -#606 := (iff #160 #605) -#603 := (iff #159 #602) -#604 := [rewrite]: #603 -#607 := [monotonicity #604]: #606 -#612 := [trans #607 #610]: #611 -#615 := [quant-intro #612]: #614 -#866 := [monotonicity #615 #863]: #865 -#872 := [trans #866 #870]: #871 -#875 := [monotonicity #872]: #874 -#881 := [trans #875 #879]: #880 -#600 := (iff #148 #599) -#597 := (iff #147 #594) -#591 := (implies #508 #146) -#595 := (iff #591 #594) -#596 := [rewrite]: #595 -#592 := (iff #147 #591) -#593 := [monotonicity #510]: #592 -#598 := [trans #593 #596]: #597 -#601 := [quant-intro #598]: #600 -#884 := [monotonicity #601 #881]: #883 -#890 := [trans #884 #888]: #889 -#893 := [monotonicity #890]: #892 -#899 := [trans #893 #897]: #898 -#589 := (iff #143 #588) -#586 := (iff #142 #585) -#587 := [rewrite]: #586 -#590 := [monotonicity #587]: #589 -#902 := [monotonicity #590 #899]: #901 -#907 := [trans #902 #905]: #906 -#910 := [monotonicity #584 #907]: #909 -#916 := [trans #910 #914]: #915 -#919 := [monotonicity #916]: #918 -#923 := [trans #919 #921]: #922 -#925 := [monotonicity #923]: #924 -#927 := [trans #925 #921]: #926 -#1109 := [monotonicity #927 #1106]: #1108 -#1112 := [monotonicity #1109]: #1111 -#1116 := [trans #1112 #1114]: #1115 -#577 := (iff #137 #576) -#574 := (iff #136 #571) -#567 := (implies #547 #564) -#572 := (iff #567 #571) -#573 := [rewrite]: #572 -#568 := (iff #136 #567) -#565 := (iff #135 #564) -#562 := (iff #134 #561) -#559 := (iff #133 #556) -#553 := (and #505 #132) -#557 := (iff #553 #556) -#558 := [rewrite]: #557 -#554 := (iff #133 #553) -#555 := [monotonicity #507]: #554 -#560 := [trans #555 #558]: #559 -#551 := (iff #131 #550) -#530 := (= #125 #529) -#531 := [rewrite]: #530 -#552 := [monotonicity #531]: #551 -#563 := [monotonicity #552 #560]: #562 -#566 := [quant-intro #563]: #565 -#548 := (iff #130 #547) -#549 := [rewrite]: #548 -#569 := [monotonicity #549 #566]: #568 -#575 := [trans #569 #573]: #574 -#578 := [quant-intro #575]: #577 -#1119 := [monotonicity #578 #1116]: #1118 -#1125 := [trans #1119 #1123]: #1124 -#545 := (iff #128 #544) -#542 := (iff #127 #539) -#535 := (implies #526 #532) -#540 := (iff #535 #539) -#541 := [rewrite]: #540 -#536 := (iff #127 #535) -#533 := (iff #126 #532) -#534 := [monotonicity #531]: #533 -#527 := (iff #124 #526) -#528 := [monotonicity #507]: #527 -#537 := [monotonicity #528 #534]: #536 -#543 := [trans #537 #541]: #542 -#546 := [quant-intro #543]: #545 -#1128 := [monotonicity #546 #1125]: #1127 -#1134 := [trans #1128 #1132]: #1133 -#524 := (iff #123 #523) -#521 := (iff #122 #518) -#514 := (implies #511 #121) -#519 := (iff #514 #518) -#520 := [rewrite]: #519 -#515 := (iff #122 #514) -#512 := (iff #119 #511) -#503 := (iff #115 #502) -#504 := [rewrite]: #503 -#513 := [monotonicity #504 #510]: #512 -#516 := [monotonicity #513]: #515 -#522 := [trans #516 #520]: #521 -#525 := [quant-intro #522]: #524 -#1137 := [monotonicity #525 #1134]: #1136 -#1143 := [trans #1137 #1141]: #1142 -#1146 := [monotonicity #1143]: #1145 -#1152 := [trans #1146 #1150]: #1151 -#500 := (iff #109 #499) -#501 := [rewrite]: #500 -#1155 := [monotonicity #501 #1152]: #1154 -#1161 := [trans #1155 #1159]: #1160 -#1164 := [monotonicity #1161]: #1163 -#1168 := [trans #1164 #1166]: #1167 -#1170 := [monotonicity #1168]: #1169 -#1172 := [trans #1170 #1166]: #1171 -#1175 := [monotonicity #498 #1172]: #1174 -#1181 := [trans #1175 #1179]: #1180 -#1184 := [monotonicity #1181 #498]: #1183 -#1189 := [trans #1184 #1187]: #1188 -#1192 := [monotonicity #483 #1189]: #1191 -#1198 := [trans #1192 #1196]: #1197 -#1201 := [monotonicity #1198 #483]: #1200 -#1206 := [trans #1201 #1204]: #1205 -#1209 := [monotonicity #468 #1206]: #1208 -#1215 := [trans #1209 #1213]: #1214 -#1218 := [monotonicity #1215 #468]: #1217 -#1223 := [trans #1218 #1221]: #1222 -#1226 := [monotonicity #1223]: #1225 -#1232 := [trans #1226 #1230]: #1231 -#1235 := [monotonicity #1232]: #1234 -#1240 := [trans #1235 #1238]: #1239 -#1243 := [monotonicity #453 #1240]: #1242 -#1249 := [trans #1243 #1247]: #1248 -#1252 := [monotonicity #1249 #453]: #1251 -#1257 := [trans #1252 #1255]: #1256 -#1260 := [monotonicity #1257]: #1259 -#1264 := [trans #1260 #1262]: #1263 -#1267 := [monotonicity #1264]: #1266 -#1273 := [trans #1267 #1271]: #1272 -#449 := (iff #76 #448) -#446 := (iff #75 #443) -#440 := (implies #73 #437) -#444 := (iff #440 #443) -#445 := [rewrite]: #444 -#441 := (iff #75 #440) -#438 := (iff #74 #437) -#439 := [rewrite]: #438 -#442 := [monotonicity #439]: #441 -#447 := [trans #442 #445]: #446 -#450 := [quant-intro #447]: #449 -#1276 := [monotonicity #450 #1273]: #1275 -#1282 := [trans #1276 #1280]: #1281 -#435 := (iff #72 #434) -#432 := (iff #71 #429) -#426 := (implies #68 #423) -#430 := (iff #426 #429) -#431 := [rewrite]: #430 -#427 := (iff #71 #426) -#424 := (iff #70 #423) -#425 := [rewrite]: #424 -#428 := [monotonicity #425]: #427 -#433 := [trans #428 #431]: #432 -#436 := [quant-intro #433]: #435 -#1285 := [monotonicity #436 #1282]: #1284 -#1291 := [trans #1285 #1289]: #1290 -#1294 := [monotonicity #1291]: #1293 -#1298 := [trans #1294 #1296]: #1297 -#1300 := [monotonicity #1298]: #1299 -#1302 := [trans #1300 #1296]: #1301 -#1305 := [monotonicity #1302]: #1304 -#1939 := [trans #1305 #1937]: #1938 -#422 := [asserted]: #293 -#1940 := [mp #422 #1939]: #1935 -#1941 := [not-or-elim #1940]: #79 -#2004 := [mp~ #1941 #1966]: #79 -#4354 := [mp #2004 #4353]: #4349 -#5709 := (not #4349) -#5715 := (or #5709 #2266) -#5470 := [quant-inst]: #5715 -#6005 := [unit-resolution #5470 #4354 #6004]: false -#6009 := [lemma #6005]: #2266 -#3870 := (or #2166 #2033) -#3957 := [def-axiom]: #3870 -#7294 := [unit-resolution #3957 #6009]: #2166 -#2421 := (not #2166) -#4707 := (or #2421 #4704) -#4710 := (not #4707) -#4362 := (forall (vars (?x27 T2)) (:pat #4355) #1313) -#4367 := (not #4362) -#4713 := (or #4367 #4710) -#4716 := (not #4713) -decl ?x27!0 :: T2 -#2011 := ?x27!0 -#2012 := (uf_12 ?x27!0) -#2013 := (>= #2012 0::int) -#2014 := (not #2013) -#4719 := (or #2014 #4716) -#4722 := (not #4719) -#4725 := (or #1308 #4722) -#4728 := (not #4725) -#4735 := (forall (vars (?x24 T2)) (:pat #4355) #1912) -#4738 := (iff #1915 #4735) -#4736 := (iff #1912 #1912) -#4737 := [refl]: #4736 -#4739 := [quant-intro #4737]: #4738 -#2238 := (~ #1915 #1915) -#2072 := (~ #1912 #1912) -#2073 := [refl]: #2072 -#2239 := [nnf-pos #2073]: #2238 -#1945 := [not-or-elim #1940]: #1915 -#2265 := [mp~ #1945 #2239]: #1915 -#4740 := [mp #2265 #4739]: #4735 -#5485 := [hypothesis]: #1308 -#5347 := (not #4735) -#5348 := (or #5347 #81) -#4030 := (= uf_11 uf_11) -#4033 := (not #4030) -#4034 := (or #4033 #81) -#5341 := (or #5347 #4034) -#5350 := (iff #5341 #5348) -#5361 := (iff #5348 #5348) -#5363 := [rewrite]: #5361 -#5108 := (iff #4034 #81) -#5046 := (or false #81) -#5106 := (iff #5046 #81) -#5107 := [rewrite]: #5106 -#5104 := (iff #4034 #5046) -#3956 := (iff #4033 false) -#6849 := (not true) -#6852 := (iff #6849 false) -#6853 := [rewrite]: #6852 -#5051 := (iff #4033 #6849) -#4038 := (iff #4030 true) -#5050 := [rewrite]: #4038 -#5052 := [monotonicity #5050]: #5051 -#3939 := [trans #5052 #6853]: #3956 -#5105 := [monotonicity #3939]: #5104 -#5346 := [trans #5105 #5107]: #5108 -#5351 := [monotonicity #5346]: #5350 -#5364 := [trans #5351 #5363]: #5350 -#5349 := [quant-inst]: #5341 -#5484 := [mp #5349 #5364]: #5348 -#5486 := [unit-resolution #5484 #5485 #4740]: false -#5487 := [lemma #5486]: #81 -#4731 := (or #1308 #4728) -#3666 := (forall (vars (?x66 T2) (?x67 T2)) #3661) -#3672 := (not #3666) -#3673 := (or #251 #3672) -#3674 := (not #3673) -#3701 := (or #3674 #3698) -#3708 := (not #3701) -#3644 := (forall (vars (?x64 T2)) #3639) -#3707 := (not #3644) -#3709 := (or #3707 #3708) -#3710 := (not #3709) -#3607 := (forall (vars (?x65 T2)) #3596) -#3613 := (not #3607) -#3614 := (or #2460 #2993 #3613) -#3615 := (not #3614) -#3715 := (or #3615 #3710) -#3722 := (not #3715) -#3592 := (forall (vars (?x41 T2)) #3581) -#3721 := (not #3592) -#3723 := (or #1062 #1044 #1071 #1053 #3721 #3722) -#3724 := (not #3723) -#3486 := (forall (vars (?x53 T2)) #3475) -#3493 := (not #3486) -#3471 := (forall (vars (?x50 T2) (?x51 T2)) #3466) -#3492 := (not #3471) -#3494 := (or #2318 #2893 #3492 #3493) -#3495 := (not #3494) -#3500 := (or #3449 #3495) -#3507 := (not #3500) -#3426 := (forall (vars (?x48 T2) (?x49 T2)) #3421) -#3506 := (not #3426) -#3508 := (or #3506 #3507) -#3509 := (not #3508) -#3514 := (or #3403 #3509) -#3520 := (not #3514) -#3521 := (or #1501 #3520) -#3522 := (not #3521) -#3527 := (or #2235 #3522) -#3533 := (not #3527) -#3534 := (or #1492 #3533) -#3535 := (not #3534) -#3540 := (or #1492 #3535) -#3546 := (not #3540) -#3547 := (or #824 #3546) -#3548 := (not #3547) -#3553 := (or #2836 #3548) -#3559 := (not #3553) -#3560 := (or #1487 #3559) -#3561 := (not #3560) -#3566 := (or #2822 #3561) -#3574 := (not #3566) -#3380 := (forall (vars (?x43 T2)) #3375) -#3573 := (not #3380) -#3362 := (forall (vars (?x44 T2)) #3359) -#3572 := (not #3362) -#3575 := (or #876 #585 #1652 #1656 #2412 #2415 #3572 #3573 #3574) -#3576 := (not #3575) -#3729 := (or #3576 #3724) -#3739 := (not #3729) -#3352 := (forall (vars (?x39 T2)) #3347) -#3738 := (not #3352) -#3324 := (forall (vars (?x32 T2)) #3319) -#3737 := (not #3324) -#3296 := (forall (vars (?x35 T2) (?x36 T2)) #3291) -#3736 := (not #3296) -#3273 := (forall (vars (?x37 T2) (?x38 T2)) #3268) -#3735 := (not #3273) -#3740 := (or #1854 #1849 #3735 #3736 #3737 #3738 #3739) -#3741 := (not #3740) -#3241 := (forall (vars (?x33 T2)) #3230) -#3247 := (not #3241) -#3248 := (or #2085 #2688 #3247) -#3249 := (not #3248) -#3746 := (or #3249 #3741) -#3753 := (not #3746) -#3226 := (forall (vars (?x30 T2) (?x31 T2)) #3221) -#3752 := (not #3226) -#3754 := (or #3752 #3753) -#3755 := (not #3754) -#3760 := (or #3204 #3755) -#3767 := (not #3760) -#3181 := (forall (vars (?x28 T2) (?x29 T2)) #3176) -#3766 := (not #3181) -#3768 := (or #3766 #3767) -#3769 := (not #3768) -#3774 := (or #2421 #3769) -#3780 := (not #3774) -#3781 := (or #1317 #3780) -#3782 := (not #3781) -#3787 := (or #2014 #3782) -#3793 := (not #3787) -#3794 := (or #1308 #3793) -#3795 := (not #3794) -#3800 := (or #1308 #3795) -#4732 := (iff #3800 #4731) -#4729 := (iff #3795 #4728) -#4726 := (iff #3794 #4725) -#4723 := (iff #3793 #4722) -#4720 := (iff #3787 #4719) -#4717 := (iff #3782 #4716) -#4714 := (iff #3781 #4713) -#4711 := (iff #3780 #4710) -#4708 := (iff #3774 #4707) -#4705 := (iff #3769 #4704) -#4702 := (iff #3768 #4701) -#4699 := (iff #3767 #4698) -#4696 := (iff #3760 #4695) -#4693 := (iff #3755 #4692) -#4690 := (iff #3754 #4689) -#4687 := (iff #3753 #4686) -#4684 := (iff #3746 #4683) -#4681 := (iff #3741 #4680) -#4678 := (iff #3740 #4677) -#4675 := (iff #3739 #4674) -#4672 := (iff #3729 #4671) -#4669 := (iff #3724 #4668) -#4666 := (iff #3723 #4665) -#4663 := (iff #3722 #4662) -#4660 := (iff #3715 #4659) -#4657 := (iff #3710 #4656) -#4654 := (iff #3709 #4653) -#4651 := (iff #3708 #4650) -#4648 := (iff #3701 #4647) -#4645 := (iff #3674 #4644) -#4642 := (iff #3673 #4641) -#4639 := (iff #3672 #4638) -#4636 := (iff #3666 #4633) -#4634 := (iff #3661 #3661) -#4635 := [refl]: #4634 -#4637 := [quant-intro #4635]: #4636 -#4640 := [monotonicity #4637]: #4639 -#4643 := [monotonicity #4640]: #4642 -#4646 := [monotonicity #4643]: #4645 -#4649 := [monotonicity #4646]: #4648 -#4652 := [monotonicity #4649]: #4651 -#4631 := (iff #3707 #4630) -#4628 := (iff #3644 #4625) -#4626 := (iff #3639 #3639) -#4627 := [refl]: #4626 -#4629 := [quant-intro #4627]: #4628 -#4632 := [monotonicity #4629]: #4631 -#4655 := [monotonicity #4632 #4652]: #4654 -#4658 := [monotonicity #4655]: #4657 -#4623 := (iff #3615 #4622) -#4620 := (iff #3614 #4619) -#4617 := (iff #3613 #4616) -#4614 := (iff #3607 #4611) -#4612 := (iff #3596 #3596) -#4613 := [refl]: #4612 -#4615 := [quant-intro #4613]: #4614 -#4618 := [monotonicity #4615]: #4617 -#4621 := [monotonicity #4618]: #4620 -#4624 := [monotonicity #4621]: #4623 -#4661 := [monotonicity #4624 #4658]: #4660 -#4664 := [monotonicity #4661]: #4663 -#4607 := (iff #3721 #4606) -#4604 := (iff #3592 #4601) -#4602 := (iff #3581 #3581) -#4603 := [refl]: #4602 -#4605 := [quant-intro #4603]: #4604 -#4608 := [monotonicity #4605]: #4607 -#4667 := [monotonicity #4608 #4664]: #4666 -#4670 := [monotonicity #4667]: #4669 -#4599 := (iff #3576 #4598) -#4596 := (iff #3575 #4595) -#4593 := (iff #3574 #4592) -#4590 := (iff #3566 #4589) -#4587 := (iff #3561 #4586) -#4584 := (iff #3560 #4583) -#4581 := (iff #3559 #4580) -#4578 := (iff #3553 #4577) -#4575 := (iff #3548 #4574) -#4572 := (iff #3547 #4571) -#4569 := (iff #3546 #4568) -#4566 := (iff #3540 #4565) -#4563 := (iff #3535 #4562) -#4560 := (iff #3534 #4559) -#4557 := (iff #3533 #4556) -#4554 := (iff #3527 #4553) -#4551 := (iff #3522 #4550) -#4548 := (iff #3521 #4547) -#4545 := (iff #3520 #4544) -#4542 := (iff #3514 #4541) -#4539 := (iff #3509 #4538) -#4536 := (iff #3508 #4535) -#4533 := (iff #3507 #4532) -#4530 := (iff #3500 #4529) -#4527 := (iff #3495 #4526) -#4524 := (iff #3494 #4523) -#4521 := (iff #3493 #4520) -#4518 := (iff #3486 #4515) -#4516 := (iff #3475 #3475) -#4517 := [refl]: #4516 -#4519 := [quant-intro #4517]: #4518 -#4522 := [monotonicity #4519]: #4521 -#4512 := (iff #3492 #4511) -#4509 := (iff #3471 #4506) -#4507 := (iff #3466 #3466) -#4508 := [refl]: #4507 -#4510 := [quant-intro #4508]: #4509 -#4513 := [monotonicity #4510]: #4512 -#4525 := [monotonicity #4513 #4522]: #4524 -#4528 := [monotonicity #4525]: #4527 -#4531 := [monotonicity #4528]: #4530 -#4534 := [monotonicity #4531]: #4533 -#4504 := (iff #3506 #4503) -#4501 := (iff #3426 #4498) -#4499 := (iff #3421 #3421) -#4500 := [refl]: #4499 -#4502 := [quant-intro #4500]: #4501 -#4505 := [monotonicity #4502]: #4504 -#4537 := [monotonicity #4505 #4534]: #4536 -#4540 := [monotonicity #4537]: #4539 -#4543 := [monotonicity #4540]: #4542 -#4546 := [monotonicity #4543]: #4545 -#4495 := (iff #1501 #4494) -#4492 := (iff #1498 #4489) -#4490 := (iff #1495 #1495) -#4491 := [refl]: #4490 -#4493 := [quant-intro #4491]: #4492 -#4496 := [monotonicity #4493]: #4495 -#4549 := [monotonicity #4496 #4546]: #4548 -#4552 := [monotonicity #4549]: #4551 -#4555 := [monotonicity #4552]: #4554 -#4558 := [monotonicity #4555]: #4557 -#4561 := [monotonicity #4558]: #4560 -#4564 := [monotonicity #4561]: #4563 -#4567 := [monotonicity #4564]: #4566 -#4570 := [monotonicity #4567]: #4569 -#4487 := (iff #824 #4486) -#4484 := (iff #642 #4481) -#4482 := (iff #637 #637) -#4483 := [refl]: #4482 -#4485 := [quant-intro #4483]: #4484 -#4488 := [monotonicity #4485]: #4487 -#4573 := [monotonicity #4488 #4570]: #4572 -#4576 := [monotonicity #4573]: #4575 -#4579 := [monotonicity #4576]: #4578 -#4582 := [monotonicity #4579]: #4581 -#4478 := (iff #1487 #4477) -#4475 := (iff #1484 #4472) -#4473 := (iff #1479 #1479) -#4474 := [refl]: #4473 -#4476 := [quant-intro #4474]: #4475 -#4479 := [monotonicity #4476]: #4478 -#4585 := [monotonicity #4479 #4582]: #4584 -#4588 := [monotonicity #4585]: #4587 -#4591 := [monotonicity #4588]: #4590 -#4594 := [monotonicity #4591]: #4593 -#4470 := (iff #3573 #4469) -#4467 := (iff #3380 #4464) -#4465 := (iff #3375 #3375) -#4466 := [refl]: #4465 -#4468 := [quant-intro #4466]: #4467 -#4471 := [monotonicity #4468]: #4470 -#4462 := (iff #3572 #4461) -#4459 := (iff #3362 #4456) -#4457 := (iff #3359 #3359) -#4458 := [refl]: #4457 -#4460 := [quant-intro #4458]: #4459 -#4463 := [monotonicity #4460]: #4462 -#4452 := (iff #1652 #4451) -#4449 := (iff #1649 #4446) -#4447 := (iff #1644 #1644) -#4448 := [refl]: #4447 -#4450 := [quant-intro #4448]: #4449 -#4453 := [monotonicity #4450]: #4452 -#4597 := [monotonicity #4453 #4463 #4471 #4594]: #4596 -#4600 := [monotonicity #4597]: #4599 -#4673 := [monotonicity #4600 #4670]: #4672 -#4676 := [monotonicity #4673]: #4675 -#4443 := (iff #3738 #4442) -#4440 := (iff #3352 #4437) -#4438 := (iff #3347 #3347) -#4439 := [refl]: #4438 -#4441 := [quant-intro #4439]: #4440 -#4444 := [monotonicity #4441]: #4443 -#4435 := (iff #3737 #4434) -#4432 := (iff #3324 #4429) -#4430 := (iff #3319 #3319) -#4431 := [refl]: #4430 -#4433 := [quant-intro #4431]: #4432 -#4436 := [monotonicity #4433]: #4435 -#4427 := (iff #3736 #4426) -#4424 := (iff #3296 #4421) -#4422 := (iff #3291 #3291) -#4423 := [refl]: #4422 -#4425 := [quant-intro #4423]: #4424 -#4428 := [monotonicity #4425]: #4427 -#4418 := (iff #3735 #4417) -#4415 := (iff #3273 #4412) -#4413 := (iff #3268 #3268) -#4414 := [refl]: #4413 -#4416 := [quant-intro #4414]: #4415 -#4419 := [monotonicity #4416]: #4418 -#4410 := (iff #1849 #4409) -#4407 := (iff #1846 #4404) -#4405 := (iff #1843 #1843) -#4406 := [refl]: #4405 -#4408 := [quant-intro #4406]: #4407 -#4411 := [monotonicity #4408]: #4410 -#4679 := [monotonicity #4411 #4419 #4428 #4436 #4444 #4676]: #4678 -#4682 := [monotonicity #4679]: #4681 -#4401 := (iff #3249 #4400) -#4398 := (iff #3248 #4397) -#4395 := (iff #3247 #4394) -#4392 := (iff #3241 #4389) -#4390 := (iff #3230 #3230) -#4391 := [refl]: #4390 -#4393 := [quant-intro #4391]: #4392 -#4396 := [monotonicity #4393]: #4395 -#4399 := [monotonicity #4396]: #4398 -#4402 := [monotonicity #4399]: #4401 -#4685 := [monotonicity #4402 #4682]: #4684 -#4688 := [monotonicity #4685]: #4687 -#4386 := (iff #3752 #4385) -#4383 := (iff #3226 #4380) -#4381 := (iff #3221 #3221) -#4382 := [refl]: #4381 -#4384 := [quant-intro #4382]: #4383 -#4387 := [monotonicity #4384]: #4386 -#4691 := [monotonicity #4387 #4688]: #4690 -#4694 := [monotonicity #4691]: #4693 -#4697 := [monotonicity #4694]: #4696 -#4700 := [monotonicity #4697]: #4699 -#4377 := (iff #3766 #4376) -#4374 := (iff #3181 #4371) -#4372 := (iff #3176 #3176) -#4373 := [refl]: #4372 -#4375 := [quant-intro #4373]: #4374 -#4378 := [monotonicity #4375]: #4377 -#4703 := [monotonicity #4378 #4700]: #4702 -#4706 := [monotonicity #4703]: #4705 -#4709 := [monotonicity #4706]: #4708 -#4712 := [monotonicity #4709]: #4711 -#4368 := (iff #1317 #4367) -#4365 := (iff #1314 #4362) -#4363 := (iff #1313 #1313) -#4364 := [refl]: #4363 -#4366 := [quant-intro #4364]: #4365 -#4369 := [monotonicity #4366]: #4368 -#4715 := [monotonicity #4369 #4712]: #4714 -#4718 := [monotonicity #4715]: #4717 -#4721 := [monotonicity #4718]: #4720 -#4724 := [monotonicity #4721]: #4723 -#4727 := [monotonicity #4724]: #4726 -#4730 := [monotonicity #4727]: #4729 -#4733 := [monotonicity #4730]: #4732 -#2527 := (not #2526) -#2523 := (not #2522) -#3081 := (and #2523 #2527) -#3084 := (not #3081) -#3101 := (or #3084 #3096) -#3104 := (not #3101) -#2536 := (not #251) -#2546 := (and #2536 #1756) -#3110 := (or #2546 #3104) -#3054 := (not #3049) -#3072 := (and #3054 #3067) -#3075 := (or #1716 #3072) -#3078 := (forall (vars (?x64 T2)) #3075) -#3115 := (and #3078 #3110) -#2456 := (not #2455) -#3024 := (and #2456 #3021) -#3027 := (not #3024) -#3030 := (forall (vars (?x65 T2)) #3027) -#2996 := (not #2993) -#2461 := (not #2460) -#3036 := (and #2461 #2996 #3030) -#3118 := (or #3036 #3115) -#2441 := (not #1670) -#2444 := (forall (vars (?x41 T2)) #2441) -#3121 := (and #229 #234 #930 #933 #2444 #3118) -#2314 := (not #2313) -#2924 := (and #630 #2314 #2921) -#2927 := (not #2924) -#2930 := (forall (vars (?x53 T2)) #2927) -#2896 := (not #2893) -#2319 := (not #2318) -#2939 := (and #1541 #2319 #2896 #2930) -#2287 := (not #2286) -#2866 := (and #2287 #2289) -#2869 := (not #2866) -#2887 := (or #2869 #2882) -#2890 := (not #2887) -#2944 := (or #2890 #2939) -#2947 := (and #1517 #2944) -#2258 := (not #2257) -#2841 := (and #2255 #2258) -#2844 := (not #2841) -#2860 := (or #2844 #2855) -#2863 := (not #2860) -#2950 := (or #2863 #2947) -#2953 := (and #1498 #2950) -#2956 := (or #2235 #2953) -#2959 := (and #173 #2956) -#2962 := (or #1492 #2959) -#2965 := (and #642 #2962) -#2968 := (or #2836 #2965) -#2971 := (and #1484 #2968) -#2974 := (or #2822 #2971) -#2416 := (not #2415) -#2413 := (not #2412) -#2980 := (and #151 #588 #1473 #1631 #1649 #1657 #2413 #2416 #2974) -#3124 := (or #2980 #3121) -#2780 := (not #2775) -#2798 := (and #2162 #2780 #2793) -#2801 := (or #1413 #2798) -#2804 := (forall (vars (?x39 T2)) #2801) -#2735 := (not #2730) -#2753 := (and #2133 #2735 #2748) -#2759 := (or #1395 #2753) -#2764 := (forall (vars (?x32 T2)) #2759) -#3130 := (and #109 #1821 #1837 #1846 #2764 #2804 #3124) -#2691 := (not #2688) -#2098 := (not #2097) -#2679 := (and #77 #2095 #2098) -#2682 := (not #2679) -#2685 := (forall (vars (?x33 T2)) #2682) -#2086 := (not #2085) -#2715 := (and #2086 #2685 #2691) -#3135 := (or #2715 #3130) -#3138 := (and #1360 #3135) -#2058 := (not #2057) -#2667 := (and #2058 #2059) -#2670 := (not #2667) -#2673 := (or #2664 #2670) -#2676 := (not #2673) -#3141 := (or #2676 #3138) -#3144 := (and #1334 #3141) -#2035 := (not #2034) -#2635 := (and #2033 #2035) -#2638 := (not #2635) -#2654 := (or #2638 #2649) -#2657 := (not #2654) -#3147 := (or #2657 #3144) -#3150 := (and #1314 #3147) -#3153 := (or #2014 #3150) -#3156 := (and #81 #3153) -#3159 := (or #1308 #3156) -#3801 := (iff #3159 #3800) -#3798 := (iff #3156 #3795) -#3790 := (and #81 #3787) -#3796 := (iff #3790 #3795) -#3797 := [rewrite]: #3796 -#3791 := (iff #3156 #3790) -#3788 := (iff #3153 #3787) -#3785 := (iff #3150 #3782) -#3777 := (and #1314 #3774) -#3783 := (iff #3777 #3782) -#3784 := [rewrite]: #3783 -#3778 := (iff #3150 #3777) -#3775 := (iff #3147 #3774) -#3772 := (iff #3144 #3769) -#3763 := (and #3181 #3760) -#3770 := (iff #3763 #3769) -#3771 := [rewrite]: #3770 -#3764 := (iff #3144 #3763) -#3761 := (iff #3141 #3760) -#3758 := (iff #3138 #3755) -#3749 := (and #3226 #3746) -#3756 := (iff #3749 #3755) -#3757 := [rewrite]: #3756 -#3750 := (iff #3138 #3749) -#3747 := (iff #3135 #3746) -#3744 := (iff #3130 #3741) -#3732 := (and #109 #3273 #3296 #1846 #3324 #3352 #3729) -#3742 := (iff #3732 #3741) -#3743 := [rewrite]: #3742 -#3733 := (iff #3130 #3732) -#3730 := (iff #3124 #3729) -#3727 := (iff #3121 #3724) -#3718 := (and #229 #234 #930 #933 #3592 #3715) -#3725 := (iff #3718 #3724) -#3726 := [rewrite]: #3725 -#3719 := (iff #3121 #3718) -#3716 := (iff #3118 #3715) -#3713 := (iff #3115 #3710) -#3704 := (and #3644 #3701) -#3711 := (iff #3704 #3710) -#3712 := [rewrite]: #3711 -#3705 := (iff #3115 #3704) -#3702 := (iff #3110 #3701) -#3699 := (iff #3104 #3698) -#3696 := (iff #3101 #3693) -#3679 := (or #2522 #2526) -#3690 := (or #3679 #3096) -#3694 := (iff #3690 #3693) -#3695 := [rewrite]: #3694 -#3691 := (iff #3101 #3690) -#3688 := (iff #3084 #3679) -#3680 := (not #3679) -#3683 := (not #3680) -#3686 := (iff #3683 #3679) -#3687 := [rewrite]: #3686 -#3684 := (iff #3084 #3683) -#3681 := (iff #3081 #3680) -#3682 := [rewrite]: #3681 -#3685 := [monotonicity #3682]: #3684 -#3689 := [trans #3685 #3687]: #3688 -#3692 := [monotonicity #3689]: #3691 -#3697 := [trans #3692 #3695]: #3696 -#3700 := [monotonicity #3697]: #3699 -#3677 := (iff #2546 #3674) -#3669 := (and #2536 #3666) -#3675 := (iff #3669 #3674) -#3676 := [rewrite]: #3675 -#3670 := (iff #2546 #3669) -#3667 := (iff #1756 #3666) -#3664 := (iff #1753 #3661) -#3647 := (or #1347 #1709) -#3658 := (or #3647 #1750) -#3662 := (iff #3658 #3661) -#3663 := [rewrite]: #3662 -#3659 := (iff #1753 #3658) -#3656 := (iff #1747 #3647) -#3648 := (not #3647) -#3651 := (not #3648) -#3654 := (iff #3651 #3647) -#3655 := [rewrite]: #3654 -#3652 := (iff #1747 #3651) -#3649 := (iff #1744 #3648) -#3650 := [rewrite]: #3649 -#3653 := [monotonicity #3650]: #3652 -#3657 := [trans #3653 #3655]: #3656 -#3660 := [monotonicity #3657]: #3659 -#3665 := [trans #3660 #3663]: #3664 -#3668 := [quant-intro #3665]: #3667 -#3671 := [monotonicity #3668]: #3670 -#3678 := [trans #3671 #3676]: #3677 -#3703 := [monotonicity #3678 #3700]: #3702 -#3645 := (iff #3078 #3644) -#3642 := (iff #3075 #3639) -#3620 := (or #68 #1709) -#3636 := (or #3620 #3633) -#3640 := (iff #3636 #3639) -#3641 := [rewrite]: #3640 -#3637 := (iff #3075 #3636) -#3634 := (iff #3072 #3633) -#3635 := [rewrite]: #3634 -#3629 := (iff #1716 #3620) -#3621 := (not #3620) -#3624 := (not #3621) -#3627 := (iff #3624 #3620) -#3628 := [rewrite]: #3627 -#3625 := (iff #1716 #3624) -#3622 := (iff #1713 #3621) -#3623 := [rewrite]: #3622 -#3626 := [monotonicity #3623]: #3625 -#3630 := [trans #3626 #3628]: #3629 -#3638 := [monotonicity #3630 #3635]: #3637 -#3643 := [trans #3638 #3641]: #3642 -#3646 := [quant-intro #3643]: #3645 -#3706 := [monotonicity #3646 #3703]: #3705 -#3714 := [trans #3706 #3712]: #3713 -#3618 := (iff #3036 #3615) -#3610 := (and #2461 #2996 #3607) -#3616 := (iff #3610 #3615) -#3617 := [rewrite]: #3616 -#3611 := (iff #3036 #3610) -#3608 := (iff #3030 #3607) -#3605 := (iff #3027 #3596) -#3597 := (not #3596) -#3600 := (not #3597) -#3603 := (iff #3600 #3596) -#3604 := [rewrite]: #3603 -#3601 := (iff #3027 #3600) -#3598 := (iff #3024 #3597) -#3599 := [rewrite]: #3598 -#3602 := [monotonicity #3599]: #3601 -#3606 := [trans #3602 #3604]: #3605 -#3609 := [quant-intro #3606]: #3608 -#3612 := [monotonicity #3609]: #3611 -#3619 := [trans #3612 #3617]: #3618 -#3717 := [monotonicity #3619 #3714]: #3716 -#3593 := (iff #2444 #3592) -#3590 := (iff #2441 #3581) -#3582 := (not #3581) -#3585 := (not #3582) -#3588 := (iff #3585 #3581) -#3589 := [rewrite]: #3588 -#3586 := (iff #2441 #3585) -#3583 := (iff #1670 #3582) -#3584 := [rewrite]: #3583 -#3587 := [monotonicity #3584]: #3586 -#3591 := [trans #3587 #3589]: #3590 -#3594 := [quant-intro #3591]: #3593 -#3720 := [monotonicity #3594 #3717]: #3719 -#3728 := [trans #3720 #3726]: #3727 -#3579 := (iff #2980 #3576) -#3569 := (and #151 #588 #3362 #3380 #1649 #1657 #2413 #2416 #3566) -#3577 := (iff #3569 #3576) -#3578 := [rewrite]: #3577 -#3570 := (iff #2980 #3569) -#3567 := (iff #2974 #3566) -#3564 := (iff #2971 #3561) -#3556 := (and #1484 #3553) -#3562 := (iff #3556 #3561) -#3563 := [rewrite]: #3562 -#3557 := (iff #2971 #3556) -#3554 := (iff #2968 #3553) -#3551 := (iff #2965 #3548) -#3543 := (and #642 #3540) -#3549 := (iff #3543 #3548) -#3550 := [rewrite]: #3549 -#3544 := (iff #2965 #3543) -#3541 := (iff #2962 #3540) -#3538 := (iff #2959 #3535) -#3530 := (and #173 #3527) -#3536 := (iff #3530 #3535) -#3537 := [rewrite]: #3536 -#3531 := (iff #2959 #3530) -#3528 := (iff #2956 #3527) -#3525 := (iff #2953 #3522) -#3517 := (and #1498 #3514) -#3523 := (iff #3517 #3522) -#3524 := [rewrite]: #3523 -#3518 := (iff #2953 #3517) -#3515 := (iff #2950 #3514) -#3512 := (iff #2947 #3509) -#3503 := (and #3426 #3500) -#3510 := (iff #3503 #3509) -#3511 := [rewrite]: #3510 -#3504 := (iff #2947 #3503) -#3501 := (iff #2944 #3500) -#3498 := (iff #2939 #3495) -#3489 := (and #3471 #2319 #2896 #3486) -#3496 := (iff #3489 #3495) -#3497 := [rewrite]: #3496 -#3490 := (iff #2939 #3489) -#3487 := (iff #2930 #3486) -#3484 := (iff #2927 #3475) -#3476 := (not #3475) -#3479 := (not #3476) -#3482 := (iff #3479 #3475) -#3483 := [rewrite]: #3482 -#3480 := (iff #2927 #3479) -#3477 := (iff #2924 #3476) -#3478 := [rewrite]: #3477 -#3481 := [monotonicity #3478]: #3480 -#3485 := [trans #3481 #3483]: #3484 -#3488 := [quant-intro #3485]: #3487 -#3472 := (iff #1541 #3471) -#3469 := (iff #1538 #3466) -#3452 := (or #636 #1347) -#3463 := (or #3452 #1534) -#3467 := (iff #3463 #3466) -#3468 := [rewrite]: #3467 -#3464 := (iff #1538 #3463) -#3461 := (iff #1531 #3452) -#3453 := (not #3452) -#3456 := (not #3453) -#3459 := (iff #3456 #3452) -#3460 := [rewrite]: #3459 -#3457 := (iff #1531 #3456) -#3454 := (iff #1526 #3453) -#3455 := [rewrite]: #3454 -#3458 := [monotonicity #3455]: #3457 -#3462 := [trans #3458 #3460]: #3461 -#3465 := [monotonicity #3462]: #3464 -#3470 := [trans #3465 #3468]: #3469 -#3473 := [quant-intro #3470]: #3472 -#3491 := [monotonicity #3473 #3488]: #3490 -#3499 := [trans #3491 #3497]: #3498 -#3450 := (iff #2890 #3449) -#3447 := (iff #2887 #3444) -#3430 := (or #2286 #3429) -#3441 := (or #3430 #2882) -#3445 := (iff #3441 #3444) -#3446 := [rewrite]: #3445 -#3442 := (iff #2887 #3441) -#3439 := (iff #2869 #3430) -#3431 := (not #3430) -#3434 := (not #3431) -#3437 := (iff #3434 #3430) -#3438 := [rewrite]: #3437 -#3435 := (iff #2869 #3434) -#3432 := (iff #2866 #3431) -#3433 := [rewrite]: #3432 -#3436 := [monotonicity #3433]: #3435 -#3440 := [trans #3436 #3438]: #3439 -#3443 := [monotonicity #3440]: #3442 -#3448 := [trans #3443 #3446]: #3447 -#3451 := [monotonicity #3448]: #3450 -#3502 := [monotonicity #3451 #3499]: #3501 -#3427 := (iff #1517 #3426) -#3424 := (iff #1512 #3421) -#3407 := (or #630 #3406) -#3418 := (or #3407 #1504) -#3422 := (iff #3418 #3421) -#3423 := [rewrite]: #3422 -#3419 := (iff #1512 #3418) -#3416 := (iff #664 #3407) -#3408 := (not #3407) -#3411 := (not #3408) -#3414 := (iff #3411 #3407) -#3415 := [rewrite]: #3414 -#3412 := (iff #664 #3411) -#3409 := (iff #656 #3408) -#3410 := [rewrite]: #3409 -#3413 := [monotonicity #3410]: #3412 -#3417 := [trans #3413 #3415]: #3416 -#3420 := [monotonicity #3417]: #3419 -#3425 := [trans #3420 #3423]: #3424 -#3428 := [quant-intro #3425]: #3427 -#3505 := [monotonicity #3428 #3502]: #3504 -#3513 := [trans #3505 #3511]: #3512 -#3404 := (iff #2863 #3403) -#3401 := (iff #2860 #3398) -#3384 := (or #3383 #2257) -#3395 := (or #3384 #2855) -#3399 := (iff #3395 #3398) -#3400 := [rewrite]: #3399 -#3396 := (iff #2860 #3395) -#3393 := (iff #2844 #3384) -#3385 := (not #3384) -#3388 := (not #3385) -#3391 := (iff #3388 #3384) -#3392 := [rewrite]: #3391 -#3389 := (iff #2844 #3388) -#3386 := (iff #2841 #3385) -#3387 := [rewrite]: #3386 -#3390 := [monotonicity #3387]: #3389 -#3394 := [trans #3390 #3392]: #3393 -#3397 := [monotonicity #3394]: #3396 -#3402 := [trans #3397 #3400]: #3401 -#3405 := [monotonicity #3402]: #3404 -#3516 := [monotonicity #3405 #3513]: #3515 -#3519 := [monotonicity #3516]: #3518 -#3526 := [trans #3519 #3524]: #3525 -#3529 := [monotonicity #3526]: #3528 -#3532 := [monotonicity #3529]: #3531 -#3539 := [trans #3532 #3537]: #3538 -#3542 := [monotonicity #3539]: #3541 -#3545 := [monotonicity #3542]: #3544 -#3552 := [trans #3545 #3550]: #3551 -#3555 := [monotonicity #3552]: #3554 -#3558 := [monotonicity #3555]: #3557 -#3565 := [trans #3558 #3563]: #3564 -#3568 := [monotonicity #3565]: #3567 -#3381 := (iff #1631 #3380) -#3378 := (iff #1628 #3375) -#3372 := (or #3355 #1625) -#3376 := (iff #3372 #3375) -#3377 := [rewrite]: #3376 -#3373 := (iff #1628 #3372) -#3370 := (iff #1620 #3355) -#3365 := (not #3356) -#3368 := (iff #3365 #3355) -#3369 := [rewrite]: #3368 -#3366 := (iff #1620 #3365) -#3357 := (iff #1462 #3356) -#3358 := [rewrite]: #3357 -#3367 := [monotonicity #3358]: #3366 -#3371 := [trans #3367 #3369]: #3370 -#3374 := [monotonicity #3371]: #3373 -#3379 := [trans #3374 #3377]: #3378 -#3382 := [quant-intro #3379]: #3381 -#3363 := (iff #1473 #3362) -#3360 := (iff #1468 #3359) -#3361 := [monotonicity #3358]: #3360 -#3364 := [quant-intro #3361]: #3363 -#3571 := [monotonicity #3364 #3382 #3568]: #3570 -#3580 := [trans #3571 #3578]: #3579 -#3731 := [monotonicity #3580 #3728]: #3730 -#3353 := (iff #2804 #3352) -#3350 := (iff #2801 #3347) -#3327 := (or #68 #1406) -#3344 := (or #3327 #3341) -#3348 := (iff #3344 #3347) -#3349 := [rewrite]: #3348 -#3345 := (iff #2801 #3344) -#3342 := (iff #2798 #3341) -#3343 := [rewrite]: #3342 -#3336 := (iff #1413 #3327) -#3328 := (not #3327) -#3331 := (not #3328) -#3334 := (iff #3331 #3327) -#3335 := [rewrite]: #3334 -#3332 := (iff #1413 #3331) -#3329 := (iff #1410 #3328) -#3330 := [rewrite]: #3329 -#3333 := [monotonicity #3330]: #3332 -#3337 := [trans #3333 #3335]: #3336 -#3346 := [monotonicity #3337 #3343]: #3345 -#3351 := [trans #3346 #3349]: #3350 -#3354 := [quant-intro #3351]: #3353 -#3325 := (iff #2764 #3324) -#3322 := (iff #2759 #3319) -#3299 := (or #68 #1388) -#3316 := (or #3299 #3313) -#3320 := (iff #3316 #3319) -#3321 := [rewrite]: #3320 -#3317 := (iff #2759 #3316) -#3314 := (iff #2753 #3313) -#3315 := [rewrite]: #3314 -#3308 := (iff #1395 #3299) -#3300 := (not #3299) -#3303 := (not #3300) -#3306 := (iff #3303 #3299) -#3307 := [rewrite]: #3306 -#3304 := (iff #1395 #3303) -#3301 := (iff #1392 #3300) -#3302 := [rewrite]: #3301 -#3305 := [monotonicity #3302]: #3304 -#3309 := [trans #3305 #3307]: #3308 -#3318 := [monotonicity #3309 #3315]: #3317 -#3323 := [trans #3318 #3321]: #3322 -#3326 := [quant-intro #3323]: #3325 -#3297 := (iff #1837 #3296) -#3294 := (iff #1832 #3291) -#3277 := (or #3276 #505) -#3288 := (or #3277 #1416) -#3292 := (iff #3288 #3291) -#3293 := [rewrite]: #3292 -#3289 := (iff #1832 #3288) -#3286 := (iff #517 #3277) -#3278 := (not #3277) -#3281 := (not #3278) -#3284 := (iff #3281 #3277) -#3285 := [rewrite]: #3284 -#3282 := (iff #517 #3281) -#3279 := (iff #511 #3278) -#3280 := [rewrite]: #3279 -#3283 := [monotonicity #3280]: #3282 -#3287 := [trans #3283 #3285]: #3286 -#3290 := [monotonicity #3287]: #3289 -#3295 := [trans #3290 #3293]: #3294 -#3298 := [quant-intro #3295]: #3297 -#3274 := (iff #1821 #3273) -#3271 := (iff #1818 #3268) -#3254 := (or #508 #1347) -#3265 := (or #3254 #1815) -#3269 := (iff #3265 #3268) -#3270 := [rewrite]: #3269 -#3266 := (iff #1818 #3265) -#3263 := (iff #1812 #3254) -#3255 := (not #3254) -#3258 := (not #3255) -#3261 := (iff #3258 #3254) -#3262 := [rewrite]: #3261 -#3259 := (iff #1812 #3258) -#3256 := (iff #1807 #3255) -#3257 := [rewrite]: #3256 -#3260 := [monotonicity #3257]: #3259 -#3264 := [trans #3260 #3262]: #3263 -#3267 := [monotonicity #3264]: #3266 -#3272 := [trans #3267 #3270]: #3271 -#3275 := [quant-intro #3272]: #3274 -#3734 := [monotonicity #3275 #3298 #3326 #3354 #3731]: #3733 -#3745 := [trans #3734 #3743]: #3744 -#3252 := (iff #2715 #3249) -#3244 := (and #2086 #3241 #2691) -#3250 := (iff #3244 #3249) -#3251 := [rewrite]: #3250 -#3245 := (iff #2715 #3244) -#3242 := (iff #2685 #3241) -#3239 := (iff #2682 #3230) -#3231 := (not #3230) -#3234 := (not #3231) -#3237 := (iff #3234 #3230) -#3238 := [rewrite]: #3237 -#3235 := (iff #2682 #3234) -#3232 := (iff #2679 #3231) -#3233 := [rewrite]: #3232 -#3236 := [monotonicity #3233]: #3235 -#3240 := [trans #3236 #3238]: #3239 -#3243 := [quant-intro #3240]: #3242 -#3246 := [monotonicity #3243]: #3245 -#3253 := [trans #3246 #3251]: #3252 -#3748 := [monotonicity #3253 #3745]: #3747 -#3227 := (iff #1360 #3226) -#3224 := (iff #1357 #3221) -#3207 := (or #78 #1347) -#3218 := (or #1342 #3207) -#3222 := (iff #3218 #3221) -#3223 := [rewrite]: #3222 -#3219 := (iff #1357 #3218) -#3216 := (iff #1354 #3207) -#3208 := (not #3207) -#3211 := (not #3208) -#3214 := (iff #3211 #3207) -#3215 := [rewrite]: #3214 -#3212 := (iff #1354 #3211) -#3209 := (iff #1351 #3208) -#3210 := [rewrite]: #3209 -#3213 := [monotonicity #3210]: #3212 -#3217 := [trans #3213 #3215]: #3216 -#3220 := [monotonicity #3217]: #3219 -#3225 := [trans #3220 #3223]: #3224 -#3228 := [quant-intro #3225]: #3227 -#3751 := [monotonicity #3228 #3748]: #3750 -#3759 := [trans #3751 #3757]: #3758 -#3205 := (iff #2676 #3204) -#3202 := (iff #2673 #3199) -#3185 := (or #2057 #3184) -#3196 := (or #2664 #3185) -#3200 := (iff #3196 #3199) -#3201 := [rewrite]: #3200 -#3197 := (iff #2673 #3196) -#3194 := (iff #2670 #3185) -#3186 := (not #3185) -#3189 := (not #3186) -#3192 := (iff #3189 #3185) -#3193 := [rewrite]: #3192 -#3190 := (iff #2670 #3189) -#3187 := (iff #2667 #3186) -#3188 := [rewrite]: #3187 -#3191 := [monotonicity #3188]: #3190 -#3195 := [trans #3191 #3193]: #3194 -#3198 := [monotonicity #3195]: #3197 -#3203 := [trans #3198 #3201]: #3202 -#3206 := [monotonicity #3203]: #3205 -#3762 := [monotonicity #3206 #3759]: #3761 -#3182 := (iff #1334 #3181) -#3179 := (iff #1329 #3176) -#3162 := (or #77 #2632) -#3173 := (or #3162 #1322) -#3177 := (iff #3173 #3176) -#3178 := [rewrite]: #3177 -#3174 := (iff #1329 #3173) -#3171 := (iff #460 #3162) -#3163 := (not #3162) -#3166 := (not #3163) -#3169 := (iff #3166 #3162) -#3170 := [rewrite]: #3169 -#3167 := (iff #460 #3166) -#3164 := (iff #454 #3163) -#3165 := [rewrite]: #3164 -#3168 := [monotonicity #3165]: #3167 -#3172 := [trans #3168 #3170]: #3171 -#3175 := [monotonicity #3172]: #3174 -#3180 := [trans #3175 #3178]: #3179 -#3183 := [quant-intro #3180]: #3182 -#3765 := [monotonicity #3183 #3762]: #3764 -#3773 := [trans #3765 #3771]: #3772 -#2534 := (iff #2657 #2421) -#2199 := (iff #2654 #2166) -#2042 := (or #2266 #2034) -#2297 := (or #2042 #2649) -#2167 := (iff #2297 #2166) -#2198 := [rewrite]: #2167 -#2499 := (iff #2654 #2297) -#2138 := (iff #2638 #2042) -#1973 := (not #2042) -#2218 := (not #1973) -#2018 := (iff #2218 #2042) -#2137 := [rewrite]: #2018 -#2219 := (iff #2638 #2218) -#1974 := (iff #2635 #1973) -#2043 := [rewrite]: #1974 -#2017 := [monotonicity #2043]: #2219 -#2296 := [trans #2017 #2137]: #2138 -#2500 := [monotonicity #2296]: #2499 -#2420 := [trans #2500 #2198]: #2199 -#2535 := [monotonicity #2420]: #2534 -#3776 := [monotonicity #2535 #3773]: #3775 -#3779 := [monotonicity #3776]: #3778 -#3786 := [trans #3779 #3784]: #3785 -#3789 := [monotonicity #3786]: #3788 -#3792 := [monotonicity #3789]: #3791 -#3799 := [trans #3792 #3797]: #3798 -#3802 := [monotonicity #3799]: #3801 -#2513 := (* -1::int #2512) -#2515 := (+ #2514 #2513) -#2518 := (+ #2517 #2515) -#2519 := (>= #2518 0::int) -#2528 := (and #2527 #2523) -#2529 := (not #2528) -#2530 := (or #2529 #2519) -#2531 := (not #2530) -#2550 := (or #2531 #2546) -#2489 := (+ #2488 #1707) -#2492 := (+ #2491 #2489) -#2493 := (= #2492 0::int) -#2494 := (>= #2489 0::int) -#2495 := (not #2494) -#2496 := (and #2495 #2493) -#2501 := (or #1716 #2496) -#2504 := (forall (vars (?x64 T2)) #2501) -#2554 := (and #2504 #2550) -#2453 := (+ #2452 #2450) -#2454 := (= #2453 0::int) -#2457 := (and #2456 #2454) -#2473 := (not #2457) -#2476 := (forall (vars (?x65 T2)) #2473) -#2462 := (= ?x64!17 uf_11) -#2463 := (not #2462) -#2464 := (and #2463 #2461) -#2465 := (not #2464) -#2470 := (not #2465) -#2480 := (and #2470 #2476) -#2558 := (or #2480 #2554) -#2438 := (not #1053) -#2435 := (not #1071) -#2432 := (not #1044) -#2429 := (not #1062) -#2562 := (and #2429 #2432 #2435 #2438 #2444 #2558) -#2417 := (and #2416 #2413) -#2311 := (+ #2310 #2308) -#2312 := (= #2311 0::int) -#2315 := (and #630 #2314 #2312) -#2332 := (not #2315) -#2335 := (forall (vars (?x53 T2)) #2332) -#2320 := (= ?x52!15 uf_11) -#2321 := (not #2320) -#2322 := (and #2321 #2319) -#2323 := (not #2322) -#2329 := (not #2323) -#2339 := (and #2329 #2335) -#2344 := (and #1541 #2339) -#2277 := (* -1::int #2276) -#2279 := (+ #2278 #2277) -#2282 := (+ #2281 #2279) -#2283 := (>= #2282 0::int) -#2290 := (and #2289 #2287) -#2291 := (not #2290) -#2292 := (or #2291 #2283) -#2293 := (not #2292) -#2348 := (or #2293 #2344) -#2352 := (and #1517 #2348) -#2250 := (* -1::int #2249) -#2252 := (+ #2251 #2250) -#2253 := (>= #2252 0::int) -#2259 := (and #2258 #2255) -#2260 := (not #2259) -#2261 := (or #2260 #2253) -#2262 := (not #2261) -#2356 := (or #2262 #2352) -#2360 := (and #1498 #2356) -#2364 := (or #2235 #2360) -#2229 := (not #1492) -#2368 := (and #2229 #2364) -#2372 := (or #1492 #2368) -#2376 := (and #642 #2372) -#2213 := (= #2212 #2211) -#2214 := (or #2213 #2210) -#2215 := (not #2214) -#2380 := (or #2215 #2376) -#2384 := (and #1484 #2380) -#2191 := (* -1::int #2190) -#2193 := (+ #2192 #2191) -#2194 := (>= #2193 0::int) -#2195 := (not #2194) -#2388 := (or #2195 #2384) -#2177 := (not #876) -#2425 := (and #2177 #588 #1473 #2388 #1631 #1649 #1657 #2417) -#2566 := (or #2425 #2562) -#2154 := (+ #2153 #1404) -#2157 := (+ #2156 #2154) -#2158 := (= #2157 0::int) -#2159 := (>= #2154 0::int) -#2160 := (not #2159) -#2163 := (and #2162 #2160 #2158) -#2168 := (or #1413 #2163) -#2171 := (forall (vars (?x39 T2)) #2168) -#2126 := (+ #1386 #2125) -#2128 := (+ #2127 #2126) -#2129 := (= #2128 0::int) -#2130 := (+ #2127 #1386) -#2131 := (>= #2130 0::int) -#2132 := (not #2131) -#2134 := (and #2133 #2132 #2129) -#2141 := (or #2134 #1395) -#2144 := (forall (vars (?x32 T2)) #2141) -#2120 := (not #1854) -#2591 := (and #2120 #2144 #2171 #2566 #1821 #1837 #1846) -#2087 := (= ?x32!5 uf_11) -#2088 := (not #2087) -#2089 := (and #2088 #2086) -#2090 := (not #2089) -#2112 := (not #2090) -#2099 := (and #77 #2098 #2095) -#2105 := (not #2099) -#2108 := (forall (vars (?x33 T2)) #2105) -#2115 := (and #2108 #2112) -#2595 := (or #2115 #2591) -#2599 := (and #1360 #2595) -#2060 := (and #2059 #2058) -#2061 := (not #2060) -#2064 := (+ #2063 #2054) -#2066 := (+ #2065 #2064) -#2067 := (>= #2066 0::int) -#2068 := (or #2067 #2061) -#2069 := (not #2068) -#2603 := (or #2069 #2599) -#2607 := (and #1334 #2603) -#2029 := (* -1::int #2028) -#2031 := (+ #2030 #2029) -#2032 := (>= #2031 0::int) -#2036 := (and #2035 #2033) -#2037 := (not #2036) -#2038 := (or #2037 #2032) -#2039 := (not #2038) -#2611 := (or #2039 #2607) -#2615 := (and #1314 #2611) -#2619 := (or #2014 #2615) -#1969 := (not #1308) -#2623 := (and #1969 #2619) -#2627 := (or #1308 #2623) -#3160 := (iff #2627 #3159) -#3157 := (iff #2623 #3156) -#3154 := (iff #2619 #3153) -#3151 := (iff #2615 #3150) -#3148 := (iff #2611 #3147) -#3145 := (iff #2607 #3144) -#3142 := (iff #2603 #3141) -#3139 := (iff #2599 #3138) -#3136 := (iff #2595 #3135) -#3133 := (iff #2591 #3130) -#3127 := (and #109 #2764 #2804 #3124 #1821 #1837 #1846) -#3131 := (iff #3127 #3130) -#3132 := [rewrite]: #3131 -#3128 := (iff #2591 #3127) -#3125 := (iff #2566 #3124) -#3122 := (iff #2562 #3121) -#3119 := (iff #2558 #3118) -#3116 := (iff #2554 #3115) -#3113 := (iff #2550 #3110) -#3107 := (or #3104 #2546) -#3111 := (iff #3107 #3110) -#3112 := [rewrite]: #3111 -#3108 := (iff #2550 #3107) -#3105 := (iff #2531 #3104) -#3102 := (iff #2530 #3101) -#3099 := (iff #2519 #3096) -#3087 := (+ #2514 #2517) -#3088 := (+ #2513 #3087) -#3091 := (>= #3088 0::int) -#3097 := (iff #3091 #3096) -#3098 := [rewrite]: #3097 -#3092 := (iff #2519 #3091) -#3089 := (= #2518 #3088) -#3090 := [rewrite]: #3089 -#3093 := [monotonicity #3090]: #3092 -#3100 := [trans #3093 #3098]: #3099 -#3085 := (iff #2529 #3084) -#3082 := (iff #2528 #3081) -#3083 := [rewrite]: #3082 -#3086 := [monotonicity #3083]: #3085 -#3103 := [monotonicity #3086 #3100]: #3102 -#3106 := [monotonicity #3103]: #3105 -#3109 := [monotonicity #3106]: #3108 -#3114 := [trans #3109 #3112]: #3113 -#3079 := (iff #2504 #3078) -#3076 := (iff #2501 #3075) -#3073 := (iff #2496 #3072) -#3070 := (iff #2493 #3067) -#3057 := (+ #2488 #2491) -#3058 := (+ #1707 #3057) -#3061 := (= #3058 0::int) -#3068 := (iff #3061 #3067) -#3069 := [rewrite]: #3068 -#3062 := (iff #2493 #3061) -#3059 := (= #2492 #3058) -#3060 := [rewrite]: #3059 -#3063 := [monotonicity #3060]: #3062 -#3071 := [trans #3063 #3069]: #3070 -#3055 := (iff #2495 #3054) -#3052 := (iff #2494 #3049) -#3041 := (+ #1707 #2488) -#3044 := (>= #3041 0::int) -#3050 := (iff #3044 #3049) -#3051 := [rewrite]: #3050 -#3045 := (iff #2494 #3044) -#3042 := (= #2489 #3041) -#3043 := [rewrite]: #3042 -#3046 := [monotonicity #3043]: #3045 -#3053 := [trans #3046 #3051]: #3052 -#3056 := [monotonicity #3053]: #3055 -#3074 := [monotonicity #3056 #3071]: #3073 -#3077 := [monotonicity #3074]: #3076 -#3080 := [quant-intro #3077]: #3079 -#3117 := [monotonicity #3080 #3114]: #3116 -#3039 := (iff #2480 #3036) -#3002 := (and #2461 #2996) -#3033 := (and #3002 #3030) -#3037 := (iff #3033 #3036) -#3038 := [rewrite]: #3037 -#3034 := (iff #2480 #3033) -#3031 := (iff #2476 #3030) -#3028 := (iff #2473 #3027) -#3025 := (iff #2457 #3024) -#3022 := (iff #2454 #3021) -#3019 := (= #2453 #3018) -#3020 := [rewrite]: #3019 -#3023 := [monotonicity #3020]: #3022 -#3026 := [monotonicity #3023]: #3025 -#3029 := [monotonicity #3026]: #3028 -#3032 := [quant-intro #3029]: #3031 -#3015 := (iff #2470 #3002) -#3007 := (not #3002) -#3010 := (not #3007) -#3013 := (iff #3010 #3002) -#3014 := [rewrite]: #3013 -#3011 := (iff #2470 #3010) -#3008 := (iff #2465 #3007) -#3005 := (iff #2464 #3002) -#2999 := (and #2996 #2461) -#3003 := (iff #2999 #3002) -#3004 := [rewrite]: #3003 -#3000 := (iff #2464 #2999) -#2997 := (iff #2463 #2996) -#2994 := (iff #2462 #2993) -#2995 := [rewrite]: #2994 -#2998 := [monotonicity #2995]: #2997 -#3001 := [monotonicity #2998]: #3000 -#3006 := [trans #3001 #3004]: #3005 -#3009 := [monotonicity #3006]: #3008 -#3012 := [monotonicity #3009]: #3011 -#3016 := [trans #3012 #3014]: #3015 -#3035 := [monotonicity #3016 #3032]: #3034 -#3040 := [trans #3035 #3038]: #3039 -#3120 := [monotonicity #3040 #3117]: #3119 -#2991 := (iff #2438 #933) -#2992 := [rewrite]: #2991 -#2989 := (iff #2435 #930) -#2990 := [rewrite]: #2989 -#2987 := (iff #2432 #234) -#2988 := [rewrite]: #2987 -#2985 := (iff #2429 #229) -#2986 := [rewrite]: #2985 -#3123 := [monotonicity #2986 #2988 #2990 #2992 #3120]: #3122 -#2983 := (iff #2425 #2980) -#2977 := (and #151 #588 #1473 #2974 #1631 #1649 #1657 #2417) -#2981 := (iff #2977 #2980) -#2982 := [rewrite]: #2981 -#2978 := (iff #2425 #2977) -#2975 := (iff #2388 #2974) -#2972 := (iff #2384 #2971) -#2969 := (iff #2380 #2968) -#2966 := (iff #2376 #2965) -#2963 := (iff #2372 #2962) -#2960 := (iff #2368 #2959) -#2957 := (iff #2364 #2956) -#2954 := (iff #2360 #2953) -#2951 := (iff #2356 #2950) -#2948 := (iff #2352 #2947) -#2945 := (iff #2348 #2944) -#2942 := (iff #2344 #2939) -#2902 := (and #2319 #2896) -#2933 := (and #2902 #2930) -#2936 := (and #1541 #2933) -#2940 := (iff #2936 #2939) -#2941 := [rewrite]: #2940 -#2937 := (iff #2344 #2936) -#2934 := (iff #2339 #2933) -#2931 := (iff #2335 #2930) -#2928 := (iff #2332 #2927) -#2925 := (iff #2315 #2924) -#2922 := (iff #2312 #2921) -#2919 := (= #2311 #2918) -#2920 := [rewrite]: #2919 -#2923 := [monotonicity #2920]: #2922 -#2926 := [monotonicity #2923]: #2925 -#2929 := [monotonicity #2926]: #2928 -#2932 := [quant-intro #2929]: #2931 -#2915 := (iff #2329 #2902) -#2907 := (not #2902) -#2910 := (not #2907) -#2913 := (iff #2910 #2902) -#2914 := [rewrite]: #2913 -#2911 := (iff #2329 #2910) -#2908 := (iff #2323 #2907) -#2905 := (iff #2322 #2902) -#2899 := (and #2896 #2319) -#2903 := (iff #2899 #2902) -#2904 := [rewrite]: #2903 -#2900 := (iff #2322 #2899) -#2897 := (iff #2321 #2896) -#2894 := (iff #2320 #2893) -#2895 := [rewrite]: #2894 -#2898 := [monotonicity #2895]: #2897 -#2901 := [monotonicity #2898]: #2900 -#2906 := [trans #2901 #2904]: #2905 -#2909 := [monotonicity #2906]: #2908 -#2912 := [monotonicity #2909]: #2911 -#2916 := [trans #2912 #2914]: #2915 -#2935 := [monotonicity #2916 #2932]: #2934 -#2938 := [monotonicity #2935]: #2937 -#2943 := [trans #2938 #2941]: #2942 -#2891 := (iff #2293 #2890) -#2888 := (iff #2292 #2887) -#2885 := (iff #2283 #2882) -#2872 := (+ #2278 #2281) -#2873 := (+ #2277 #2872) -#2876 := (>= #2873 0::int) -#2883 := (iff #2876 #2882) -#2884 := [rewrite]: #2883 -#2877 := (iff #2283 #2876) -#2874 := (= #2282 #2873) -#2875 := [rewrite]: #2874 -#2878 := [monotonicity #2875]: #2877 -#2886 := [trans #2878 #2884]: #2885 -#2870 := (iff #2291 #2869) -#2867 := (iff #2290 #2866) -#2868 := [rewrite]: #2867 -#2871 := [monotonicity #2868]: #2870 -#2889 := [monotonicity #2871 #2886]: #2888 -#2892 := [monotonicity #2889]: #2891 -#2946 := [monotonicity #2892 #2943]: #2945 -#2949 := [monotonicity #2946]: #2948 -#2864 := (iff #2262 #2863) -#2861 := (iff #2261 #2860) -#2858 := (iff #2253 #2855) -#2847 := (+ #2250 #2251) -#2850 := (>= #2847 0::int) -#2856 := (iff #2850 #2855) -#2857 := [rewrite]: #2856 -#2851 := (iff #2253 #2850) -#2848 := (= #2252 #2847) -#2849 := [rewrite]: #2848 -#2852 := [monotonicity #2849]: #2851 -#2859 := [trans #2852 #2857]: #2858 -#2845 := (iff #2260 #2844) -#2842 := (iff #2259 #2841) -#2843 := [rewrite]: #2842 -#2846 := [monotonicity #2843]: #2845 -#2862 := [monotonicity #2846 #2859]: #2861 -#2865 := [monotonicity #2862]: #2864 -#2952 := [monotonicity #2865 #2949]: #2951 -#2955 := [monotonicity #2952]: #2954 -#2958 := [monotonicity #2955]: #2957 -#2839 := (iff #2229 #173) -#2840 := [rewrite]: #2839 -#2961 := [monotonicity #2840 #2958]: #2960 -#2964 := [monotonicity #2961]: #2963 -#2967 := [monotonicity #2964]: #2966 -#2837 := (iff #2215 #2836) -#2834 := (iff #2214 #2831) -#2828 := (or #2825 #2210) -#2832 := (iff #2828 #2831) -#2833 := [rewrite]: #2832 -#2829 := (iff #2214 #2828) -#2826 := (iff #2213 #2825) -#2827 := [rewrite]: #2826 -#2830 := [monotonicity #2827]: #2829 -#2835 := [trans #2830 #2833]: #2834 -#2838 := [monotonicity #2835]: #2837 -#2970 := [monotonicity #2838 #2967]: #2969 -#2973 := [monotonicity #2970]: #2972 -#2823 := (iff #2195 #2822) -#2820 := (iff #2194 #2817) -#2809 := (+ #2191 #2192) -#2812 := (>= #2809 0::int) -#2818 := (iff #2812 #2817) -#2819 := [rewrite]: #2818 -#2813 := (iff #2194 #2812) -#2810 := (= #2193 #2809) -#2811 := [rewrite]: #2810 -#2814 := [monotonicity #2811]: #2813 -#2821 := [trans #2814 #2819]: #2820 -#2824 := [monotonicity #2821]: #2823 -#2976 := [monotonicity #2824 #2973]: #2975 -#2807 := (iff #2177 #151) -#2808 := [rewrite]: #2807 -#2979 := [monotonicity #2808 #2976]: #2978 -#2984 := [trans #2979 #2982]: #2983 -#3126 := [monotonicity #2984 #3123]: #3125 -#2805 := (iff #2171 #2804) -#2802 := (iff #2168 #2801) -#2799 := (iff #2163 #2798) -#2796 := (iff #2158 #2793) -#2783 := (+ #2153 #2156) -#2784 := (+ #1404 #2783) -#2787 := (= #2784 0::int) -#2794 := (iff #2787 #2793) -#2795 := [rewrite]: #2794 -#2788 := (iff #2158 #2787) -#2785 := (= #2157 #2784) -#2786 := [rewrite]: #2785 -#2789 := [monotonicity #2786]: #2788 -#2797 := [trans #2789 #2795]: #2796 -#2781 := (iff #2160 #2780) -#2778 := (iff #2159 #2775) -#2767 := (+ #1404 #2153) -#2770 := (>= #2767 0::int) -#2776 := (iff #2770 #2775) -#2777 := [rewrite]: #2776 -#2771 := (iff #2159 #2770) -#2768 := (= #2154 #2767) -#2769 := [rewrite]: #2768 -#2772 := [monotonicity #2769]: #2771 -#2779 := [trans #2772 #2777]: #2778 -#2782 := [monotonicity #2779]: #2781 -#2800 := [monotonicity #2782 #2797]: #2799 -#2803 := [monotonicity #2800]: #2802 -#2806 := [quant-intro #2803]: #2805 -#2765 := (iff #2144 #2764) -#2762 := (iff #2141 #2759) -#2756 := (or #2753 #1395) -#2760 := (iff #2756 #2759) -#2761 := [rewrite]: #2760 -#2757 := (iff #2141 #2756) -#2754 := (iff #2134 #2753) -#2751 := (iff #2129 #2748) -#2738 := (+ #2125 #2127) -#2739 := (+ #1386 #2738) -#2742 := (= #2739 0::int) -#2749 := (iff #2742 #2748) -#2750 := [rewrite]: #2749 -#2743 := (iff #2129 #2742) -#2740 := (= #2128 #2739) -#2741 := [rewrite]: #2740 -#2744 := [monotonicity #2741]: #2743 -#2752 := [trans #2744 #2750]: #2751 -#2736 := (iff #2132 #2735) -#2733 := (iff #2131 #2730) -#2722 := (+ #1386 #2127) -#2725 := (>= #2722 0::int) -#2731 := (iff #2725 #2730) -#2732 := [rewrite]: #2731 -#2726 := (iff #2131 #2725) -#2723 := (= #2130 #2722) -#2724 := [rewrite]: #2723 -#2727 := [monotonicity #2724]: #2726 -#2734 := [trans #2727 #2732]: #2733 -#2737 := [monotonicity #2734]: #2736 -#2755 := [monotonicity #2737 #2752]: #2754 -#2758 := [monotonicity #2755]: #2757 -#2763 := [trans #2758 #2761]: #2762 -#2766 := [quant-intro #2763]: #2765 -#2720 := (iff #2120 #109) -#2721 := [rewrite]: #2720 -#3129 := [monotonicity #2721 #2766 #2806 #3126]: #3128 -#3134 := [trans #3129 #3132]: #3133 -#2718 := (iff #2115 #2715) -#2697 := (and #2086 #2691) -#2712 := (and #2685 #2697) -#2716 := (iff #2712 #2715) -#2717 := [rewrite]: #2716 -#2713 := (iff #2115 #2712) -#2710 := (iff #2112 #2697) -#2702 := (not #2697) -#2705 := (not #2702) -#2708 := (iff #2705 #2697) -#2709 := [rewrite]: #2708 -#2706 := (iff #2112 #2705) -#2703 := (iff #2090 #2702) -#2700 := (iff #2089 #2697) -#2694 := (and #2691 #2086) -#2698 := (iff #2694 #2697) -#2699 := [rewrite]: #2698 -#2695 := (iff #2089 #2694) -#2692 := (iff #2088 #2691) -#2689 := (iff #2087 #2688) -#2690 := [rewrite]: #2689 -#2693 := [monotonicity #2690]: #2692 -#2696 := [monotonicity #2693]: #2695 -#2701 := [trans #2696 #2699]: #2700 -#2704 := [monotonicity #2701]: #2703 -#2707 := [monotonicity #2704]: #2706 -#2711 := [trans #2707 #2709]: #2710 -#2686 := (iff #2108 #2685) -#2683 := (iff #2105 #2682) -#2680 := (iff #2099 #2679) -#2681 := [rewrite]: #2680 -#2684 := [monotonicity #2681]: #2683 -#2687 := [quant-intro #2684]: #2686 -#2714 := [monotonicity #2687 #2711]: #2713 -#2719 := [trans #2714 #2717]: #2718 -#3137 := [monotonicity #2719 #3134]: #3136 -#3140 := [monotonicity #3137]: #3139 -#2677 := (iff #2069 #2676) -#2674 := (iff #2068 #2673) -#2671 := (iff #2061 #2670) -#2668 := (iff #2060 #2667) -#2669 := [rewrite]: #2668 -#2672 := [monotonicity #2669]: #2671 -#2665 := (iff #2067 #2664) -#2662 := (= #2066 #2661) -#2663 := [rewrite]: #2662 -#2666 := [monotonicity #2663]: #2665 -#2675 := [monotonicity #2666 #2672]: #2674 -#2678 := [monotonicity #2675]: #2677 -#3143 := [monotonicity #2678 #3140]: #3142 -#3146 := [monotonicity #3143]: #3145 -#2658 := (iff #2039 #2657) -#2655 := (iff #2038 #2654) -#2652 := (iff #2032 #2649) -#2641 := (+ #2029 #2030) -#2644 := (>= #2641 0::int) -#2650 := (iff #2644 #2649) -#2651 := [rewrite]: #2650 -#2645 := (iff #2032 #2644) -#2642 := (= #2031 #2641) -#2643 := [rewrite]: #2642 -#2646 := [monotonicity #2643]: #2645 -#2653 := [trans #2646 #2651]: #2652 -#2639 := (iff #2037 #2638) -#2636 := (iff #2036 #2635) -#2637 := [rewrite]: #2636 -#2640 := [monotonicity #2637]: #2639 -#2656 := [monotonicity #2640 #2653]: #2655 -#2659 := [monotonicity #2656]: #2658 -#3149 := [monotonicity #2659 #3146]: #3148 -#3152 := [monotonicity #3149]: #3151 -#3155 := [monotonicity #3152]: #3154 -#2633 := (iff #1969 #81) -#2634 := [rewrite]: #2633 -#3158 := [monotonicity #2634 #3155]: #3157 -#3161 := [monotonicity #3158]: #3160 -#1943 := (not #1907) -#2628 := (~ #1943 #2627) -#2624 := (not #1904) -#2625 := (~ #2624 #2623) -#2620 := (not #1901) -#2621 := (~ #2620 #2619) -#2616 := (not #1898) -#2617 := (~ #2616 #2615) -#2612 := (not #1895) -#2613 := (~ #2612 #2611) -#2608 := (not #1892) -#2609 := (~ #2608 #2607) -#2604 := (not #1889) -#2605 := (~ #2604 #2603) -#2600 := (not #1886) -#2601 := (~ #2600 #2599) -#2596 := (not #1883) -#2597 := (~ #2596 #2595) -#2592 := (not #1878) -#2593 := (~ #2592 #2591) -#2588 := (not #1849) -#2589 := (~ #2588 #1846) -#2586 := (~ #1846 #1846) -#2584 := (~ #1843 #1843) -#2585 := [refl]: #2584 -#2587 := [nnf-pos #2585]: #2586 -#2590 := [nnf-neg #2587]: #2589 -#2581 := (not #1840) -#2582 := (~ #2581 #1837) -#2579 := (~ #1837 #1837) -#2577 := (~ #1832 #1832) -#2578 := [refl]: #2577 -#2580 := [nnf-pos #2578]: #2579 -#2583 := [nnf-neg #2580]: #2582 -#2574 := (not #1824) -#2575 := (~ #2574 #1821) -#2572 := (~ #1821 #1821) -#2570 := (~ #1818 #1818) -#2571 := [refl]: #2570 -#2573 := [nnf-pos #2571]: #2572 -#2576 := [nnf-neg #2573]: #2575 -#2567 := (not #1801) -#2568 := (~ #2567 #2566) -#2563 := (not #1796) -#2564 := (~ #2563 #2562) -#2559 := (not #1778) -#2560 := (~ #2559 #2558) -#2555 := (not #1775) -#2556 := (~ #2555 #2554) -#2551 := (not #1772) -#2552 := (~ #2551 #2550) -#2547 := (not #1767) -#2548 := (~ #2547 #2546) -#2543 := (not #1759) -#2544 := (~ #2543 #1756) -#2541 := (~ #1756 #1756) -#2539 := (~ #1753 #1753) -#2540 := [refl]: #2539 -#2542 := [nnf-pos #2540]: #2541 -#2545 := [nnf-neg #2542]: #2544 -#2537 := (~ #2536 #2536) -#2538 := [refl]: #2537 -#2549 := [nnf-neg #2538 #2545]: #2548 -#2532 := (~ #1759 #2531) -#2533 := [sk]: #2532 -#2553 := [nnf-neg #2533 #2549]: #2552 -#2507 := (not #1741) -#2508 := (~ #2507 #2504) -#2505 := (~ #1738 #2504) -#2502 := (~ #1735 #2501) -#2497 := (~ #1732 #2496) -#2498 := [sk]: #2497 -#2485 := (~ #1716 #1716) -#2486 := [refl]: #2485 -#2503 := [monotonicity #2486 #2498]: #2502 -#2506 := [nnf-pos #2503]: #2505 -#2509 := [nnf-neg #2506]: #2508 -#2557 := [nnf-neg #2509 #2553]: #2556 -#2483 := (~ #1741 #2480) -#2458 := (exists (vars (?x65 T2)) #2457) -#2466 := (or #2465 #2458) -#2467 := (not #2466) -#2481 := (~ #2467 #2480) -#2477 := (not #2458) -#2478 := (~ #2477 #2476) -#2474 := (~ #2473 #2473) -#2475 := [refl]: #2474 -#2479 := [nnf-neg #2475]: #2478 -#2471 := (~ #2470 #2470) -#2472 := [refl]: #2471 -#2482 := [nnf-neg #2472 #2479]: #2481 -#2468 := (~ #1741 #2467) -#2469 := [sk]: #2468 -#2484 := [trans #2469 #2482]: #2483 -#2561 := [nnf-neg #2484 #2557]: #2560 -#2445 := (~ #1678 #2444) -#2442 := (~ #2441 #2441) -#2443 := [refl]: #2442 -#2446 := [nnf-neg #2443]: #2445 -#2439 := (~ #2438 #2438) -#2440 := [refl]: #2439 -#2436 := (~ #2435 #2435) -#2437 := [refl]: #2436 -#2433 := (~ #2432 #2432) -#2434 := [refl]: #2433 -#2430 := (~ #2429 #2429) -#2431 := [refl]: #2430 -#2565 := [nnf-neg #2431 #2434 #2437 #2440 #2446 #2561]: #2564 -#2426 := (not #1702) -#2427 := (~ #2426 #2425) -#2422 := (not #1678) -#2423 := (~ #2422 #2417) -#2418 := (~ #1675 #2417) -#2419 := [sk]: #2418 -#2424 := [nnf-neg #2419]: #2423 -#2406 := (~ #1657 #1657) -#2407 := [refl]: #2406 -#2403 := (not #1652) -#2404 := (~ #2403 #1649) -#2401 := (~ #1649 #1649) -#2399 := (~ #1644 #1644) -#2400 := [refl]: #2399 -#2402 := [nnf-pos #2400]: #2401 -#2405 := [nnf-neg #2402]: #2404 -#2396 := (not #1634) -#2397 := (~ #2396 #1631) -#2394 := (~ #1631 #1631) -#2392 := (~ #1628 #1628) -#2393 := [refl]: #2392 -#2395 := [nnf-pos #2393]: #2394 -#2398 := [nnf-neg #2395]: #2397 -#2389 := (not #1617) -#2390 := (~ #2389 #2388) -#2385 := (not #1614) -#2386 := (~ #2385 #2384) -#2381 := (not #1611) -#2382 := (~ #2381 #2380) -#2377 := (not #1608) -#2378 := (~ #2377 #2376) -#2373 := (not #1605) -#2374 := (~ #2373 #2372) -#2369 := (not #1602) -#2370 := (~ #2369 #2368) -#2365 := (not #1599) -#2366 := (~ #2365 #2364) -#2361 := (not #1596) -#2362 := (~ #2361 #2360) -#2357 := (not #1593) -#2358 := (~ #2357 #2356) -#2353 := (not #1590) -#2354 := (~ #2353 #2352) -#2349 := (not #1587) -#2350 := (~ #2349 #2348) -#2345 := (not #1584) -#2346 := (~ #2345 #2344) -#2326 := (not #1581) -#2342 := (~ #2326 #2339) -#2316 := (exists (vars (?x53 T2)) #2315) -#2324 := (or #2323 #2316) -#2325 := (not #2324) -#2340 := (~ #2325 #2339) -#2336 := (not #2316) -#2337 := (~ #2336 #2335) -#2333 := (~ #2332 #2332) -#2334 := [refl]: #2333 -#2338 := [nnf-neg #2334]: #2337 -#2330 := (~ #2329 #2329) -#2331 := [refl]: #2330 -#2341 := [nnf-neg #2331 #2338]: #2340 -#2327 := (~ #2326 #2325) -#2328 := [sk]: #2327 -#2343 := [trans #2328 #2341]: #2342 -#2302 := (not #1544) -#2303 := (~ #2302 #1541) -#2300 := (~ #1541 #1541) -#2298 := (~ #1538 #1538) -#2299 := [refl]: #2298 -#2301 := [nnf-pos #2299]: #2300 -#2304 := [nnf-neg #2301]: #2303 -#2347 := [nnf-neg #2304 #2343]: #2346 -#2294 := (~ #1544 #2293) -#2295 := [sk]: #2294 -#2351 := [nnf-neg #2295 #2347]: #2350 -#2271 := (not #1520) -#2272 := (~ #2271 #1517) -#2269 := (~ #1517 #1517) -#2267 := (~ #1512 #1512) -#2268 := [refl]: #2267 -#2270 := [nnf-pos #2268]: #2269 -#2273 := [nnf-neg #2270]: #2272 -#2355 := [nnf-neg #2273 #2351]: #2354 -#2263 := (~ #1520 #2262) -#2264 := [sk]: #2263 -#2359 := [nnf-neg #2264 #2355]: #2358 -#2244 := (not #1501) -#2245 := (~ #2244 #1498) -#2242 := (~ #1498 #1498) -#2240 := (~ #1495 #1495) -#2241 := [refl]: #2240 -#2243 := [nnf-pos #2241]: #2242 -#2246 := [nnf-neg #2243]: #2245 -#2363 := [nnf-neg #2246 #2359]: #2362 -#2236 := (~ #1501 #2235) -#2237 := [sk]: #2236 -#2367 := [nnf-neg #2237 #2363]: #2366 -#2230 := (~ #2229 #2229) -#2231 := [refl]: #2230 -#2371 := [nnf-neg #2231 #2367]: #2370 -#2227 := (~ #1492 #1492) -#2228 := [refl]: #2227 -#2375 := [nnf-neg #2228 #2371]: #2374 -#2224 := (not #824) -#2225 := (~ #2224 #642) -#2222 := (~ #642 #642) -#2220 := (~ #637 #637) -#2221 := [refl]: #2220 -#2223 := [nnf-pos #2221]: #2222 -#2226 := [nnf-neg #2223]: #2225 -#2379 := [nnf-neg #2226 #2375]: #2378 -#2216 := (~ #824 #2215) -#2217 := [sk]: #2216 -#2383 := [nnf-neg #2217 #2379]: #2382 -#2204 := (not #1487) -#2205 := (~ #2204 #1484) -#2202 := (~ #1484 #1484) -#2200 := (~ #1479 #1479) -#2201 := [refl]: #2200 -#2203 := [nnf-pos #2201]: #2202 -#2206 := [nnf-neg #2203]: #2205 -#2387 := [nnf-neg #2206 #2383]: #2386 -#2196 := (~ #1487 #2195) -#2197 := [sk]: #2196 -#2391 := [nnf-neg #2197 #2387]: #2390 -#2186 := (not #1476) -#2187 := (~ #2186 #1473) -#2184 := (~ #1473 #1473) -#2182 := (~ #1468 #1468) -#2183 := [refl]: #2182 -#2185 := [nnf-pos #2183]: #2184 -#2188 := [nnf-neg #2185]: #2187 -#2180 := (~ #588 #588) -#2181 := [refl]: #2180 -#2178 := (~ #2177 #2177) -#2179 := [refl]: #2178 -#2428 := [nnf-neg #2179 #2181 #2188 #2391 #2398 #2405 #2407 #2424]: #2427 -#2569 := [nnf-neg #2428 #2565]: #2568 -#2174 := (not #1446) -#2175 := (~ #2174 #2171) -#2172 := (~ #1443 #2171) -#2169 := (~ #1440 #2168) -#2164 := (~ #1437 #2163) -#2165 := [sk]: #2164 -#2150 := (~ #1413 #1413) -#2151 := [refl]: #2150 -#2170 := [monotonicity #2151 #2165]: #2169 -#2173 := [nnf-pos #2170]: #2172 -#2176 := [nnf-neg #2173]: #2175 -#2147 := (not #1857) -#2148 := (~ #2147 #2144) -#2145 := (~ #1401 #2144) -#2142 := (~ #1398 #2141) -#2139 := (~ #1395 #1395) -#2140 := [refl]: #2139 -#2135 := (~ #1383 #2134) -#2136 := [sk]: #2135 -#2143 := [monotonicity #2136 #2140]: #2142 -#2146 := [nnf-pos #2143]: #2145 -#2149 := [nnf-neg #2146]: #2148 -#2121 := (~ #2120 #2120) -#2122 := [refl]: #2121 -#2594 := [nnf-neg #2122 #2149 #2176 #2569 #2576 #2583 #2590]: #2593 -#2118 := (~ #1857 #2115) -#2100 := (exists (vars (?x33 T2)) #2099) -#2101 := (or #2100 #2090) -#2102 := (not #2101) -#2116 := (~ #2102 #2115) -#2113 := (~ #2112 #2112) -#2114 := [refl]: #2113 -#2109 := (not #2100) -#2110 := (~ #2109 #2108) -#2106 := (~ #2105 #2105) -#2107 := [refl]: #2106 -#2111 := [nnf-neg #2107]: #2110 -#2117 := [nnf-neg #2111 #2114]: #2116 -#2103 := (~ #1857 #2102) -#2104 := [sk]: #2103 -#2119 := [trans #2104 #2117]: #2118 -#2598 := [nnf-neg #2119 #2594]: #2597 -#2078 := (not #1363) -#2079 := (~ #2078 #1360) -#2076 := (~ #1360 #1360) -#2074 := (~ #1357 #1357) -#2075 := [refl]: #2074 -#2077 := [nnf-pos #2075]: #2076 -#2080 := [nnf-neg #2077]: #2079 -#2602 := [nnf-neg #2080 #2598]: #2601 -#2070 := (~ #1363 #2069) -#2071 := [sk]: #2070 -#2606 := [nnf-neg #2071 #2602]: #2605 -#2048 := (not #1337) -#2049 := (~ #2048 #1334) -#2046 := (~ #1334 #1334) -#2044 := (~ #1329 #1329) -#2045 := [refl]: #2044 -#2047 := [nnf-pos #2045]: #2046 -#2050 := [nnf-neg #2047]: #2049 -#2610 := [nnf-neg #2050 #2606]: #2609 -#2040 := (~ #1337 #2039) -#2041 := [sk]: #2040 -#2614 := [nnf-neg #2041 #2610]: #2613 -#2023 := (not #1317) -#2024 := (~ #2023 #1314) -#2021 := (~ #1314 #1314) -#2019 := (~ #1313 #1313) -#2020 := [refl]: #2019 -#2022 := [nnf-pos #2020]: #2021 -#2025 := [nnf-neg #2022]: #2024 -#2618 := [nnf-neg #2025 #2614]: #2617 -#2015 := (~ #1317 #2014) -#2016 := [sk]: #2015 -#2622 := [nnf-neg #2016 #2618]: #2621 -#1970 := (~ #1969 #1969) -#2010 := [refl]: #1970 -#2626 := [nnf-neg #2010 #2622]: #2625 -#2008 := (~ #1308 #1308) -#2009 := [refl]: #2008 -#2629 := [nnf-neg #2009 #2626]: #2628 -#1944 := [not-or-elim #1940]: #1943 -#2630 := [mp~ #1944 #2629]: #2627 -#2631 := [mp #2630 #3161]: #3159 -#3803 := [mp #2631 #3802]: #3800 -#4734 := [mp #3803 #4733]: #4731 -#7295 := [unit-resolution #4734 #5487]: #4728 -#4058 := (or #4725 #4719) -#4059 := [def-axiom]: #4058 -#7296 := [unit-resolution #4059 #7295]: #4719 -#373 := (<= uf_9 0::int) -#374 := (not #373) -#57 := (< 0::int uf_9) -#375 := (iff #57 #374) -#376 := [rewrite]: #375 -#369 := [asserted]: #57 -#377 := [mp #369 #376]: #374 -#5901 := (* -1::int #2012) -#5891 := (+ uf_9 #5901) -#5892 := (<= #5891 0::int) -#5472 := (= uf_9 #2012) -#5745 := (= uf_11 ?x27!0) -#5918 := (not #5745) -#5916 := (= #2012 0::int) -#5836 := (not #5916) -#5835 := [hypothesis]: #2014 -#5837 := (or #5836 #2013) -#5896 := [th-lemma]: #5837 -#5922 := [unit-resolution #5896 #5835]: #5836 -#5981 := (or #5347 #5918 #5916) -#5477 := (= ?x27!0 uf_11) -#5917 := (not #5477) -#5890 := (or #5917 #5916) -#5982 := (or #5347 #5890) -#5987 := (iff #5982 #5981) -#5915 := (or #5918 #5916) -#5984 := (or #5347 #5915) -#5985 := (iff #5984 #5981) -#5986 := [rewrite]: #5985 -#6000 := (iff #5982 #5984) -#5921 := (iff #5890 #5915) -#5919 := (iff #5917 #5918) -#5743 := (iff #5477 #5745) -#5746 := [rewrite]: #5743 -#5920 := [monotonicity #5746]: #5919 -#5980 := [monotonicity #5920]: #5921 -#5979 := [monotonicity #5980]: #6000 -#5988 := [trans #5979 #5986]: #5987 -#5983 := [quant-inst]: #5982 -#6010 := [mp #5983 #5988]: #5981 -#5923 := [unit-resolution #6010 #4740 #5922]: #5918 -#5748 := (or #5472 #5745) -#4356 := (forall (vars (?x25 T2)) (:pat #4355) #443) -#4359 := (iff #448 #4356) -#4357 := (iff #443 #443) -#4358 := [refl]: #4357 -#4360 := [quant-intro #4358]: #4359 -#1967 := (~ #448 #448) -#2005 := (~ #443 #443) -#2006 := [refl]: #2005 -#1968 := [nnf-pos #2006]: #1967 -#1942 := [not-or-elim #1940]: #448 -#2007 := [mp~ #1942 #1968]: #448 -#4361 := [mp #2007 #4360]: #4356 -#5821 := (not #4356) -#5827 := (or #5821 #5472 #5745) -#5737 := (or #5477 #5472) -#5828 := (or #5821 #5737) -#5894 := (iff #5828 #5827) -#5830 := (or #5821 #5748) -#5846 := (iff #5830 #5827) -#5847 := [rewrite]: #5846 -#5826 := (iff #5828 #5830) -#5756 := (iff #5737 #5748) -#5736 := (or #5745 #5472) -#5752 := (iff #5736 #5748) -#5753 := [rewrite]: #5752 -#5747 := (iff #5737 #5736) -#5744 := [monotonicity #5746]: #5747 -#5834 := [trans #5744 #5753]: #5756 -#5831 := [monotonicity #5834]: #5826 -#5895 := [trans #5831 #5847]: #5894 -#5829 := [quant-inst]: #5828 -#5900 := [mp #5829 #5895]: #5827 -#5924 := [unit-resolution #5900 #4361]: #5748 -#5989 := [unit-resolution #5924 #5923]: #5472 -#6012 := (not #5472) -#6013 := (or #6012 #5892) -#6014 := [th-lemma]: #6013 -#6042 := [unit-resolution #6014 #5989]: #5892 -#6011 := (<= #2012 0::int) -#6043 := (or #6011 #2013) -#6044 := [th-lemma]: #6043 -#6045 := [unit-resolution #6044 #5835]: #6011 -#6046 := [th-lemma #6045 #6042 #377]: false -#6041 := [lemma #6046]: #2013 -#4053 := (or #4722 #2014 #4716) -#4054 := [def-axiom]: #4053 -#7297 := [unit-resolution #4054 #6041 #7296]: #4716 -#4077 := (or #4713 #4707) -#4078 := [def-axiom]: #4077 -#7298 := [unit-resolution #4078 #7297]: #4707 -#4071 := (or #4710 #2421 #4704) -#4073 := [def-axiom]: #4071 -#7299 := [unit-resolution #4073 #7298 #7294]: #4704 -#4098 := (or #4701 #4695) -#4099 := [def-axiom]: #4098 -#7300 := [unit-resolution #4099 #7299]: #4695 -#6817 := [hypothesis]: #2059 -#6051 := (or #5709 #3184) -#6081 := [quant-inst]: #6051 -#6818 := [unit-resolution #6081 #4354 #6817]: false -#6839 := [lemma #6818]: #3184 -#3960 := (or #3199 #2059) -#3964 := [def-axiom]: #3960 -#7301 := [unit-resolution #3964 #6839]: #3199 -#4094 := (or #4698 #3204 #4692) -#4095 := [def-axiom]: #4094 -#7302 := [unit-resolution #4095 #7301 #7300]: #4692 -#4108 := (or #4689 #4683) -#4129 := [def-axiom]: #4108 -#7303 := [unit-resolution #4129 #7302]: #4683 -#6633 := (= uf_9 #2082) -#6706 := (not #6633) -#6684 := [hypothesis]: #4400 -#4274 := (or #4397 #2086) -#3948 := [def-axiom]: #4274 -#6685 := [unit-resolution #3948 #6684]: #2086 -#6798 := (or #6706 #2085) -#6840 := [th-lemma]: #6798 -#6835 := [unit-resolution #6840 #6685]: #6706 -#3949 := (or #4397 #2691) -#4281 := [def-axiom]: #3949 -#6841 := [unit-resolution #4281 #6684]: #2691 -#6695 := (or #5821 #2688 #6633) -#6680 := (or #2087 #6633) -#6696 := (or #5821 #6680) -#6732 := (iff #6696 #6695) -#6548 := (or #2688 #6633) -#6694 := (or #5821 #6548) -#6699 := (iff #6694 #6695) -#6705 := [rewrite]: #6699 -#6697 := (iff #6696 #6694) -#6608 := (iff #6680 #6548) -#6665 := [monotonicity #2690]: #6608 -#6698 := [monotonicity #6665]: #6697 -#6728 := [trans #6698 #6705]: #6732 -#6547 := [quant-inst]: #6696 -#6734 := [mp #6547 #6728]: #6695 -#6842 := [unit-resolution #6734 #4361 #6841 #6835]: false -#6843 := [lemma #6842]: #4397 -#4116 := (or #4686 #4400 #4680) -#4117 := [def-axiom]: #4116 -#7304 := [unit-resolution #4117 #6843 #7303]: #4680 -#4149 := (or #4677 #4404) -#4145 := [def-axiom]: #4149 -#7305 := [unit-resolution #4145 #7304]: #4404 -#24124 := (or #4409 #20405) -#25438 := [quant-inst]: #24124 -#29563 := [unit-resolution #25438 #7305]: #20405 -#15997 := (* -1::int #15996) -#15993 := (uf_1 #15992 ?x52!15) -#15994 := (uf_10 #15993) -#15995 := (* -1::int #15994) -#16012 := (+ #15995 #15997) -#15362 := (uf_4 uf_14 ?x52!15) -#16013 := (+ #15362 #16012) -#22006 := (>= #16013 0::int) -#16016 := (= #16013 0::int) -#16019 := (not #16016) -#16004 := (uf_6 uf_15 #15992) -#16005 := (= uf_8 #16004) -#16006 := (not #16005) -#16002 := (+ #15362 #15997) -#16003 := (<= #16002 0::int) -#16025 := (or #16003 #16006 #16019) -#16030 := (not #16025) -#15397 := (* -1::int #15362) -#16009 := (+ uf_9 #15397) -#16010 := (<= #16009 0::int) -#34324 := (not #16010) -#15398 := (+ #2306 #15397) -#14218 := (>= #15398 0::int) -#15367 := (= #2306 #15362) -decl uf_3 :: (-> T1 T2) -#20604 := (uf_3 #15993) -#32578 := (uf_6 uf_15 #20604) -#32576 := (= uf_8 #32578) -#5319 := (uf_6 #150 uf_16) -#5314 := (= uf_8 #5319) -decl uf_2 :: (-> T1 T2) -#6008 := (uf_1 uf_16 uf_11) -#7128 := (uf_2 #6008) -#32602 := (= #7128 #20604) -#32591 := (ite #32602 #5314 #32576) -#7203 := (uf_7 uf_15 #7128 #5319) -#32554 := (uf_6 #7203 #20604) -#32538 := (= uf_8 #32554) -#32564 := (iff #32538 #32591) +#194 := (uf_7 uf_15 uf_22 uf_8) +#3894 := (uf_6 #194 uf_22) +#3895 := (= uf_8 #3894) +#10330 := (ite #10319 #3895 #9519) +#10323 := (uf_7 uf_15 #9695 #3894) +#10324 := (uf_6 #10323 ?x63!14) +#10327 := (= uf_8 #10324) +#10333 := (iff #10327 #10330) #30 := (:var 1 T5) #20 := (:var 2 T2) #29 := (:var 3 T4) #31 := (uf_7 #29 #20 #30) #32 := (uf_6 #31 #11) -#4314 := (pattern #32) +#4216 := (pattern #32) #36 := (uf_6 #29 #11) #335 := (= uf_8 #36) #35 := (= #30 uf_8) @@ -4239,16 +241,16 @@ #338 := (ite #24 #35 #335) #34 := (= #32 uf_8) #341 := (iff #34 #338) -#4315 := (forall (vars (?x10 T4) (?x11 T2) (?x12 T5) (?x13 T2)) (:pat #4314) #341) +#4217 := (forall (vars (?x10 T4) (?x11 T2) (?x12 T5) (?x13 T2)) (:pat #4216) #341) #344 := (forall (vars (?x10 T4) (?x11 T2) (?x12 T5) (?x13 T2)) #341) -#4318 := (iff #344 #4315) -#4316 := (iff #341 #341) -#4317 := [refl]: #4316 -#4319 := [quant-intro #4317]: #4318 -#1953 := (~ #344 #344) -#1989 := (~ #341 #341) -#1990 := [refl]: #1989 -#1954 := [nnf-pos #1990]: #1953 +#4220 := (iff #344 #4217) +#4218 := (iff #341 #341) +#4219 := [refl]: #4218 +#4221 := [quant-intro #4219]: #4220 +#1848 := (~ #344 #344) +#1884 := (~ #341 #341) +#1885 := [refl]: #1884 +#1849 := [nnf-pos #1885]: #1848 #37 := (= #36 uf_8) #38 := (ite #24 #35 #37) #39 := (iff #34 #38) @@ -4263,460 +265,4364 @@ #346 := [quant-intro #343]: #345 #333 := [asserted]: #40 #349 := [mp #333 #346]: #344 -#1991 := [mp~ #349 #1954]: #344 -#4320 := [mp #1991 #4319]: #4315 -#7026 := (not #4315) -#32561 := (or #7026 #32564) -#6089 := (= #5319 uf_8) -#32580 := (= #20604 #7128) -#32553 := (ite #32580 #6089 #32576) -#32556 := (= #32554 uf_8) -#32579 := (iff #32556 #32553) -#32603 := (or #7026 #32579) -#32548 := (iff #32603 #32561) -#32606 := (iff #32561 #32561) -#32631 := [rewrite]: #32606 -#32629 := (iff #32579 #32564) -#32535 := (iff #32553 #32591) -#6101 := (iff #6089 #5314) -#6102 := [rewrite]: #6101 -#32601 := (iff #32580 #32602) -#32563 := [rewrite]: #32601 -#32560 := [monotonicity #32563 #6102]: #32535 -#32537 := (iff #32556 #32538) -#32557 := [rewrite]: #32537 -#32565 := [monotonicity #32557 #32560]: #32629 -#32587 := [monotonicity #32565]: #32548 -#32604 := [trans #32587 #32631]: #32548 -#32558 := [quant-inst]: #32603 -#32623 := [mp #32558 #32604]: #32561 -#32954 := [unit-resolution #32623 #4320]: #32564 -#30313 := (not #32538) -#8488 := (uf_6 uf_17 ?x52!15) -#9319 := (= uf_8 #8488) -#9846 := (not #9319) -#32968 := (iff #9846 #30313) -#32955 := (iff #9319 #32538) -#32985 := (iff #32538 #9319) -#32980 := (= #32554 #8488) -#32983 := (= #20604 ?x52!15) -#20605 := (= ?x52!15 #20604) -#12 := (uf_1 #10 #11) -#4294 := (pattern #12) -#13 := (uf_3 #12) -#317 := (= #11 #13) -#4295 := (forall (vars (?x2 T2) (?x3 T2)) (:pat #4294) #317) -#321 := (forall (vars (?x2 T2) (?x3 T2)) #317) -#4298 := (iff #321 #4295) -#4296 := (iff #317 #317) -#4297 := [refl]: #4296 -#4299 := [quant-intro #4297]: #4298 -#1948 := (~ #321 #321) -#1980 := (~ #317 #317) -#1981 := [refl]: #1980 -#1946 := [nnf-pos #1981]: #1948 -#14 := (= #13 #11) -#15 := (forall (vars (?x2 T2) (?x3 T2)) #14) -#322 := (iff #15 #321) -#319 := (iff #14 #317) -#320 := [rewrite]: #319 -#323 := [quant-intro #320]: #322 -#316 := [asserted]: #15 -#326 := [mp #316 #323]: #321 -#1982 := [mp~ #326 #1946]: #321 -#4300 := [mp #1982 #4299]: #4295 -#5378 := (not #4295) -#27981 := (or #5378 #20605) -#27945 := [quant-inst]: #27981 -#32953 := [unit-resolution #27945 #4300]: #20605 -#32984 := [symm #32953]: #32983 -#8596 := (= #7203 uf_17) -#8594 := (= #150 uf_17) -#4147 := (or #4677 #109) -#4148 := [def-axiom]: #4147 -#7307 := [unit-resolution #4148 #7304]: #109 -#4150 := (or #4677 #4412) -#4130 := [def-axiom]: #4150 -#8027 := [unit-resolution #4130 #7304]: #4412 -#4137 := (or #4677 #4437) -#4132 := [def-axiom]: #4137 -#8030 := [unit-resolution #4132 #7304]: #4437 -#5284 := (or #4665 #4442 #4417 #1854) -#4776 := (uf_4 uf_14 ?x64!17) -#4777 := (* -1::int #4776) -#4778 := (+ uf_9 #4777) -#4779 := (<= #4778 0::int) -#4845 := (?x40!7 ?x64!17) -#4941 := (uf_6 uf_15 #4845) -#4942 := (= uf_8 #4941) -#4943 := (not #4942) -#4848 := (uf_4 uf_14 #4845) -#4849 := (* -1::int #4848) -#4939 := (+ #4776 #4849) -#4940 := (<= #4939 0::int) -#4846 := (uf_1 #4845 ?x64!17) -#4847 := (uf_10 #4846) -#4842 := (* -1::int #4847) -#4926 := (+ #4842 #4849) -#4927 := (+ #4776 #4926) -#4930 := (= #4927 0::int) -#4932 := (not #4930) -#5006 := (or #4932 #4940 #4943) -#5242 := [hypothesis]: #4668 -#4159 := (or #4665 #933) -#4154 := [def-axiom]: #4159 -#5243 := [unit-resolution #4154 #5242]: #933 -#4134 := (or #4665 #4659) -#4135 := [def-axiom]: #4134 -#5244 := [unit-resolution #4135 #5242]: #4659 -#5245 := [hypothesis]: #109 -#5247 := (= #250 #108) -#5246 := [symm #5243]: #231 -#5248 := [monotonicity #5246]: #5247 -#5249 := [trans #5248 #5245]: #251 -#4193 := (or #4641 #2536) -#4198 := [def-axiom]: #4193 -#5250 := [unit-resolution #4198 #5249]: #4641 -#5087 := [hypothesis]: #4412 -#4160 := (or #4665 #4601) -#4133 := [def-axiom]: #4160 -#5230 := [unit-resolution #4133 #5242]: #4601 -#5100 := (or #3693 #4417 #4606 #1053) -#4862 := (uf_4 uf_14 ?x67!19) -#3931 := (uf_4 uf_14 ?x66!20) -#3932 := (* -1::int #3931) -#4955 := (+ #3932 #4862) -#4956 := (+ #2517 #4955) -#4959 := (>= #4956 0::int) -#4866 := (uf_6 uf_15 ?x67!19) -#4867 := (= uf_8 #4866) -#4863 := (* -1::int #4862) -#4864 := (+ uf_9 #4863) -#4865 := (<= #4864 0::int) -#5068 := (not #4865) -#5072 := [hypothesis]: #3698 -#4189 := (or #3693 #2523) -#4186 := [def-axiom]: #4189 -#5073 := [unit-resolution #4186 #5072]: #2523 -#5061 := (+ #2514 #4863) -#5063 := (>= #5061 0::int) -#5060 := (= #2514 #4862) -#5075 := (= #4862 #2514) -#5074 := [hypothesis]: #933 -#5076 := [monotonicity #5074]: #5075 -#5077 := [symm #5076]: #5060 -#5078 := (not #5060) -#5079 := (or #5078 #5063) -#5080 := [th-lemma]: #5079 -#5081 := [unit-resolution #5080 #5077]: #5063 -#5069 := (not #5063) -#5070 := (or #5068 #5069 #2522) -#5064 := [hypothesis]: #2523 -#5065 := [hypothesis]: #4865 -#5066 := [hypothesis]: #5063 -#5067 := [th-lemma #5066 #5065 #5064]: false -#5071 := [lemma #5067]: #5070 -#5082 := [unit-resolution #5071 #5081 #5073]: #5068 -#4869 := (or #4865 #4867) -#5083 := [hypothesis]: #4601 -#4872 := (or #4606 #4865 #4867) -#4868 := (or #4867 #4865) -#4873 := (or #4606 #4868) -#4880 := (iff #4873 #4872) -#4875 := (or #4606 #4869) -#4878 := (iff #4875 #4872) -#4879 := [rewrite]: #4878 -#4876 := (iff #4873 #4875) -#4870 := (iff #4868 #4869) -#4871 := [rewrite]: #4870 -#4877 := [monotonicity #4871]: #4876 -#4881 := [trans #4877 #4879]: #4880 -#4874 := [quant-inst]: #4873 -#4882 := [mp #4874 #4881]: #4872 -#5084 := [unit-resolution #4882 #5083]: #4869 -#5085 := [unit-resolution #5084 #5082]: #4867 -#4953 := (not #4867) -#5088 := (or #4953 #4959) -#4190 := (or #3693 #2527) -#4170 := [def-axiom]: #4190 -#5086 := [unit-resolution #4170 #5072]: #2527 -#4970 := (or #4417 #2526 #4953 #4959) -#4948 := (+ #4862 #3932) -#4949 := (+ #2517 #4948) -#4952 := (>= #4949 0::int) -#4954 := (or #4953 #2526 #4952) -#4971 := (or #4417 #4954) -#4978 := (iff #4971 #4970) -#4965 := (or #2526 #4953 #4959) -#4973 := (or #4417 #4965) -#4976 := (iff #4973 #4970) -#4977 := [rewrite]: #4976 -#4974 := (iff #4971 #4973) -#4968 := (iff #4954 #4965) -#4962 := (or #4953 #2526 #4959) -#4966 := (iff #4962 #4965) -#4967 := [rewrite]: #4966 -#4963 := (iff #4954 #4962) -#4960 := (iff #4952 #4959) -#4957 := (= #4949 #4956) -#4958 := [rewrite]: #4957 -#4961 := [monotonicity #4958]: #4960 -#4964 := [monotonicity #4961]: #4963 -#4969 := [trans #4964 #4967]: #4968 -#4975 := [monotonicity #4969]: #4974 -#4979 := [trans #4975 #4977]: #4978 -#4972 := [quant-inst]: #4971 -#4980 := [mp #4972 #4979]: #4970 -#5089 := [unit-resolution #4980 #5087 #5086]: #5088 -#5090 := [unit-resolution #5089 #5085]: #4959 -#4171 := (not #3096) -#4173 := (or #3693 #4171) -#4174 := [def-axiom]: #4173 -#5091 := [unit-resolution #4174 #5072]: #4171 -#5054 := (+ #2512 #3932) -#5058 := (<= #5054 0::int) -#5053 := (= #2512 #3931) -#5092 := (= #3931 #2512) -#5093 := [monotonicity #5074]: #5092 -#5094 := [symm #5093]: #5053 -#5095 := (not #5053) -#5096 := (or #5095 #5058) -#5097 := [th-lemma]: #5096 -#5098 := [unit-resolution #5097 #5094]: #5058 -#5099 := [th-lemma #5098 #5091 #5081 #5090]: false -#5101 := [lemma #5099]: #5100 -#5231 := [unit-resolution #5101 #5230 #5087 #5243]: #3693 -#4181 := (or #4650 #4644 #3698) -#4182 := [def-axiom]: #4181 -#5232 := [unit-resolution #4182 #5231 #5250]: #4650 -#4161 := (or #4653 #4647) -#4162 := [def-axiom]: #4161 -#5233 := [unit-resolution #4162 #5232]: #4653 -#4169 := (or #4662 #4622 #4656) -#4155 := [def-axiom]: #4169 -#5234 := [unit-resolution #4155 #5233 #5244]: #4622 -#4194 := (or #4619 #4611) -#4195 := [def-axiom]: #4194 -#5229 := [unit-resolution #4195 #5234]: #4611 -#5713 := (or #5006 #4616 #1053) -#5123 := (uf_4 uf_22 #4845) -#5136 := (* -1::int #5123) -#5137 := (+ #2448 #5136) -#5138 := (<= #5137 0::int) -#5150 := (+ #4842 #5136) -#5151 := (+ #2448 #5150) -#5152 := (= #5151 0::int) -#5382 := (+ #4848 #5136) -#5387 := (>= #5382 0::int) -#5381 := (= #4848 #5123) -#5643 := (= #5123 #4848) -#5642 := [symm #5074]: #231 -#5644 := [monotonicity #5642]: #5643 -#5645 := [symm #5644]: #5381 -#5646 := (not #5381) -#5647 := (or #5646 #5387) -#5648 := [th-lemma]: #5647 -#5649 := [unit-resolution #5648 #5645]: #5387 -#5119 := (+ #2448 #4777) -#5121 := (>= #5119 0::int) -#5118 := (= #2448 #4776) -#5629 := (= #4776 #2448) -#5630 := [monotonicity #5074]: #5629 -#5631 := [symm #5630]: #5118 -#5632 := (not #5118) -#5633 := (or #5632 #5121) -#5628 := [th-lemma]: #5633 -#5634 := [unit-resolution #5628 #5631]: #5121 -#5040 := (>= #4927 0::int) -#5005 := (not #5006) -#5635 := [hypothesis]: #5005 -#5042 := (or #5006 #4930) -#5043 := [def-axiom]: #5042 -#5636 := [unit-resolution #5043 #5635]: #4930 -#5637 := (or #4932 #5040) -#5638 := [th-lemma]: #5637 -#5653 := [unit-resolution #5638 #5636]: #5040 -#5386 := (<= #5382 0::int) -#5654 := (or #5646 #5386) -#5675 := [th-lemma]: #5654 -#5676 := [unit-resolution #5675 #5645]: #5386 -#5120 := (<= #5119 0::int) -#5677 := (or #5632 #5120) -#5678 := [th-lemma]: #5677 -#5679 := [unit-resolution #5678 #5631]: #5120 -#5034 := (<= #4927 0::int) -#5674 := (or #4932 #5034) -#5680 := [th-lemma]: #5674 -#5681 := [unit-resolution #5680 #5636]: #5034 -#5562 := (not #5387) -#5567 := (not #5121) -#5566 := (not #5040) -#5779 := (not #5386) -#5778 := (not #5120) -#5777 := (not #5034) -#5572 := (or #5152 #5777 #5778 #5779 #5566 #5567 #5562) -#5774 := [hypothesis]: #5386 -#5775 := [hypothesis]: #5120 -#5776 := [hypothesis]: #5034 -#5157 := (not #5152) -#5772 := [hypothesis]: #5157 -#5175 := (>= #5151 0::int) -#5563 := [hypothesis]: #5387 -#5564 := [hypothesis]: #5121 -#5565 := [hypothesis]: #5040 -#5568 := (or #5175 #5566 #5567 #5562) -#5569 := [th-lemma]: #5568 -#5570 := [unit-resolution #5569 #5565 #5564 #5563]: #5175 -#5784 := (not #5175) -#5788 := (or #5784 #5152 #5777 #5778 #5779) -#5773 := [hypothesis]: #5175 -#5174 := (<= #5151 0::int) -#5780 := (or #5174 #5777 #5778 #5779) -#5781 := [th-lemma]: #5780 -#5782 := [unit-resolution #5781 #5776 #5775 #5774]: #5174 -#5783 := (not #5174) -#5785 := (or #5152 #5783 #5784) -#5786 := [th-lemma]: #5785 -#5787 := [unit-resolution #5786 #5782 #5773 #5772]: false -#5789 := [lemma #5787]: #5788 -#5571 := [unit-resolution #5789 #5570 #5772 #5776 #5775 #5774]: false -#5641 := [lemma #5571]: #5572 -#5682 := [unit-resolution #5641 #5681 #5679 #5676 #5653 #5634 #5649]: #5152 -#5160 := (or #5138 #5157) -#5683 := [hypothesis]: #4611 -#5163 := (or #4616 #5138 #5157) -#5122 := (+ #2449 #4847) -#5124 := (+ #5123 #5122) -#5125 := (= #5124 0::int) -#5126 := (not #5125) -#5127 := (+ #5123 #2449) -#5128 := (>= #5127 0::int) -#5129 := (or #5128 #5126) -#5164 := (or #4616 #5129) -#5171 := (iff #5164 #5163) -#5166 := (or #4616 #5160) -#5169 := (iff #5166 #5163) -#5170 := [rewrite]: #5169 -#5167 := (iff #5164 #5166) -#5161 := (iff #5129 #5160) -#5158 := (iff #5126 #5157) -#5155 := (iff #5125 #5152) -#5143 := (+ #4847 #5123) -#5144 := (+ #2449 #5143) -#5147 := (= #5144 0::int) -#5153 := (iff #5147 #5152) -#5154 := [rewrite]: #5153 -#5148 := (iff #5125 #5147) -#5145 := (= #5124 #5144) -#5146 := [rewrite]: #5145 -#5149 := [monotonicity #5146]: #5148 -#5156 := [trans #5149 #5154]: #5155 -#5159 := [monotonicity #5156]: #5158 -#5141 := (iff #5128 #5138) -#5130 := (+ #2449 #5123) -#5133 := (>= #5130 0::int) -#5139 := (iff #5133 #5138) -#5140 := [rewrite]: #5139 -#5134 := (iff #5128 #5133) -#5131 := (= #5127 #5130) -#5132 := [rewrite]: #5131 -#5135 := [monotonicity #5132]: #5134 -#5142 := [trans #5135 #5140]: #5141 -#5162 := [monotonicity #5142 #5159]: #5161 -#5168 := [monotonicity #5162]: #5167 -#5172 := [trans #5168 #5170]: #5171 -#5165 := [quant-inst]: #5164 -#5173 := [mp #5165 #5172]: #5163 -#5684 := [unit-resolution #5173 #5683]: #5160 -#5710 := [unit-resolution #5684 #5682]: #5138 -#5041 := (not #4940) -#5044 := (or #5006 #5041) -#5045 := [def-axiom]: #5044 -#5711 := [unit-resolution #5045 #5635]: #5041 -#5712 := [th-lemma #5634 #5711 #5649 #5710]: false -#5714 := [lemma #5712]: #5713 -#5235 := [unit-resolution #5714 #5229 #5243]: #5006 -#5238 := (or #4779 #5005) -#4191 := (or #4619 #2996) -#4192 := [def-axiom]: #4191 -#5236 := [unit-resolution #4192 #5234]: #2996 -#5237 := [hypothesis]: #4437 -#5023 := (or #4442 #2993 #4779 #5005) -#4850 := (+ #4849 #4842) -#4851 := (+ #4776 #4850) -#4852 := (= #4851 0::int) -#4938 := (not #4852) -#4944 := (or #4943 #4940 #4938) -#4945 := (not #4944) -#4946 := (or #2462 #4779 #4945) -#5025 := (or #4442 #4946) -#5031 := (iff #5025 #5023) -#5013 := (or #2993 #4779 #5005) -#5027 := (or #4442 #5013) -#5024 := (iff #5027 #5023) -#5030 := [rewrite]: #5024 -#5028 := (iff #5025 #5027) -#5014 := (iff #4946 #5013) -#5011 := (iff #4945 #5005) -#5009 := (iff #4944 #5006) -#4935 := (or #4943 #4940 #4932) -#5007 := (iff #4935 #5006) -#5008 := [rewrite]: #5007 -#4950 := (iff #4944 #4935) -#4933 := (iff #4938 #4932) -#4925 := (iff #4852 #4930) -#4928 := (= #4851 #4927) -#4929 := [rewrite]: #4928 -#4931 := [monotonicity #4929]: #4925 -#4934 := [monotonicity #4931]: #4933 -#4951 := [monotonicity #4934]: #4950 -#5010 := [trans #4951 #5008]: #5009 -#5012 := [monotonicity #5010]: #5011 -#5015 := [monotonicity #2995 #5012]: #5014 -#5029 := [monotonicity #5015]: #5028 -#5032 := [trans #5029 #5030]: #5031 -#5026 := [quant-inst]: #5025 -#5033 := [mp #5026 #5032]: #5023 -#5239 := [unit-resolution #5033 #5237 #5236]: #5238 -#5254 := [unit-resolution #5239 #5235]: #4779 -#4200 := (or #4619 #2461) -#4207 := [def-axiom]: #4200 -#5255 := [unit-resolution #4207 #5234]: #2461 -#5280 := [monotonicity #5243]: #5629 -#5281 := [symm #5280]: #5118 -#5282 := [unit-resolution #5628 #5281]: #5121 -#5283 := [th-lemma #5282 #5255 #5254]: false -#5279 := [lemma #5283]: #5284 -#8031 := [unit-resolution #5279 #8030 #8027 #7307]: #4665 -#4138 := (or #4677 #4671) -#4106 := [def-axiom]: #4138 -#7357 := [unit-resolution #4106 #7304]: #4671 -#4143 := (or #4674 #4598 #4668) -#4144 := [def-axiom]: #4143 -#7389 := [unit-resolution #4144 #7357 #8031]: #4598 -#4125 := (or #4595 #151) -#4126 := [def-axiom]: #4125 -#8578 := [unit-resolution #4126 #7389]: #151 -#8595 := [symm #8578]: #8594 -#8592 := (= #7203 #150) +#1886 := [mp~ #349 #1849]: #344 +#4222 := [mp #1886 #4221]: #4217 +#4987 := (not #4217) +#13028 := (or #4987 #10333) +#4958 := (= #3894 uf_8) +#10322 := (ite #10319 #4958 #9519) +#10325 := (= #10324 uf_8) +#10326 := (iff #10325 #10322) +#13031 := (or #4987 #10326) +#13033 := (iff #13031 #13028) +#13035 := (iff #13028 #13028) +#13036 := [rewrite]: #13035 +#10334 := (iff #10326 #10333) +#10331 := (iff #10322 #10330) +#4970 := (iff #4958 #3895) +#4971 := [rewrite]: #4970 +#10332 := [monotonicity #4971]: #10331 +#10328 := (iff #10325 #10327) +#10329 := [rewrite]: #10328 +#10335 := [monotonicity #10329 #10332]: #10334 +#13034 := [monotonicity #10335]: #13033 +#13037 := [trans #13034 #13036]: #13033 +#13032 := [quant-inst]: #13031 +#13051 := [mp #13032 #13037]: #13028 +#13579 := [unit-resolution #13051 #4222]: #10333 +#13595 := (= #2257 #10324) +#13584 := (= #10324 #2257) +#13582 := (= #10323 uf_23) +#7680 := (= #194 uf_23) +#195 := (= uf_23 #194) +#4549 := (or #2877 #4546) +#4552 := (not #4549) +#1480 := (* -1::int #202) +#1481 := (+ #110 #1480) +#1479 := (>= #1481 0::int) +#4444 := (forall (vars (?x61 T2)) (:pat #4305 #4426) #1479) +#4449 := (not #4444) +#4555 := (or #4449 #4552) +#4558 := (not #4555) +decl ?x61!13 :: T2 +#2238 := ?x61!13 +#2241 := (uf_4 uf_14 ?x61!13) +#2856 := (* -1::int #2241) +#2239 := (uf_24 ?x61!13) +#2857 := (+ #2239 #2856) +#2858 := (<= #2857 0::int) +#2863 := (not #2858) +#4561 := (or #2863 #4558) +#4564 := (not #4561) +#196 := (uf_1 uf_22 #11) +#4427 := (pattern #196) +#197 := (uf_10 #196) +#1623 := (+ #197 #1480) +#1624 := (+ #188 #1623) +#1625 := (= #1624 0::int) +#1449 := (* -1::int #197) +#1455 := (* -1::int #188) +#1456 := (+ #1455 #1449) +#1457 := (+ #110 #1456) +#1458 := (<= #1457 0::int) +#1450 := (+ uf_9 #1449) +#1451 := (<= #1450 0::int) +#3425 := (or #1451 #1458 #1625) +#4436 := (forall (vars (?x59 T2)) (:pat #4427 #4305 #4426) #3425) +#4441 := (not #4436) +#3405 := (or #1451 #1458) +#3406 := (not #3405) +#3409 := (or #759 #3406) +#4428 := (forall (vars (?x60 T2)) (:pat #4305 #4426 #4427) #3409) +#4433 := (not #4428) +decl ?x48!12 :: T2 +#2214 := ?x48!12 +#2220 := (uf_6 uf_15 ?x48!12) +#2221 := (= uf_8 #2220) +#2215 := (uf_4 uf_14 ?x48!12) +#2216 := (* -1::int #2215) +#2217 := (+ uf_9 #2216) +#2218 := (<= #2217 0::int) +#1655 := (+ uf_9 #1455) +#1656 := (<= #1655 0::int) +#114 := (uf_6 uf_15 #11) +#4347 := (pattern #114) +#1638 := (+ #110 #1455) +#1637 := (>= #1638 0::int) +#478 := (= uf_8 #114) +#1644 := (or #478 #1637) +#4418 := (forall (vars (?x58 T2)) (:pat #4347 #4305) #1644) +#4423 := (not #4418) +#185 := (uf_6 uf_15 uf_22) +#728 := (= uf_8 #185) +#981 := (not #195) +#4567 := (or #981 #728 #4423 #1656 #2218 #2221 #4433 #4441 #4564) +#4570 := (not #4567) +decl ?x53!11 :: T2 +#2148 := ?x53!11 +decl ?x54!10 :: T2 +#2147 := ?x54!10 +#2153 := (uf_1 ?x54!10 ?x53!11) +#2154 := (uf_10 #2153) +#2161 := (* -1::int #2154) +decl uf_19 :: T3 +#146 := uf_19 +#2151 := (uf_4 uf_19 ?x54!10) +#2157 := (* -1::int #2151) +#2813 := (+ #2157 #2161) +#2149 := (uf_4 uf_19 ?x53!11) +#2814 := (+ #2149 #2813) +#2815 := (<= #2814 0::int) +#2162 := (+ uf_9 #2161) +#2163 := (<= #2162 0::int) +#2158 := (+ uf_9 #2157) +#2159 := (<= #2158 0::int) +#3369 := (or #2159 #2163 #2815) +#3374 := (not #3369) +#154 := (uf_4 uf_19 #10) +#1357 := (* -1::int #154) +#151 := (uf_4 uf_19 #11) +#1358 := (+ #151 #1357) +#1364 := (+ #91 #1358) +#1387 := (>= #1364 0::int) +#1344 := (* -1::int #151) +#1345 := (+ uf_9 #1344) +#1346 := (<= #1345 0::int) +#3337 := (or #1237 #1346 #1387) +#4380 := (forall (vars (?x53 T2) (?x54 T2)) (:pat #4281) #3337) +#4385 := (not #4380) +#166 := (uf_4 uf_19 uf_11) +#167 := (= #166 0::int) +#4388 := (or #167 #4385) +#4391 := (not #4388) +#4394 := (or #4391 #3374) +#4397 := (not #4394) +#4356 := (pattern #151) +decl ?x50!9 :: (-> T2 T2) +#2124 := (?x50!9 #11) +#2127 := (uf_1 #2124 #11) +#2128 := (uf_10 #2127) +#2783 := (* -1::int #2128) +#2125 := (uf_4 uf_19 #2124) +#2766 := (* -1::int #2125) +#2784 := (+ #2766 #2783) +#2785 := (+ #151 #2784) +#2786 := (= #2785 0::int) +#3307 := (not #2786) +#2767 := (+ #151 #2766) +#2768 := (<= #2767 0::int) +#3308 := (or #2768 #3307) +#3309 := (not #3308) +#68 := (= #11 uf_11) +#3315 := (or #68 #1346 #3309) +#4372 := (forall (vars (?x49 T2)) (:pat #4356) #3315) +#4377 := (not #4372) +#4400 := (or #4377 #4397) +#4403 := (not #4400) +decl ?x49!8 :: T2 +#2084 := ?x49!8 +#2088 := (uf_1 #11 ?x49!8) +#4357 := (pattern #2088) +#2089 := (uf_10 #2088) +#2085 := (uf_4 uf_19 ?x49!8) +#2086 := (* -1::int #2085) +#2736 := (+ #2086 #2089) +#2737 := (+ #151 #2736) +#2740 := (= #2737 0::int) +#3271 := (not #2740) +#2087 := (+ #151 #2086) +#2092 := (>= #2087 0::int) +#3272 := (or #2092 #3271) +#4358 := (forall (vars (?x50 T2)) (:pat #4356 #4357) #3272) +#4363 := (not #4358) +#2712 := (= uf_11 ?x49!8) +#2096 := (+ uf_9 #2086) +#2097 := (<= #2096 0::int) +#4366 := (or #2097 #2712 #4363) +#4369 := (not #4366) +#4406 := (or #4369 #4403) +#4409 := (not #4406) +#1299 := (* -1::int #110) +#1300 := (+ uf_9 #1299) +#1301 := (<= #1300 0::int) +#3257 := (or #478 #1301) +#4348 := (forall (vars (?x48 T2)) (:pat #4347 #4305) #3257) +#4353 := (not #4348) +#569 := (= uf_14 uf_19) +#674 := (not #569) +decl uf_16 :: T4 +#141 := uf_16 +#566 := (= uf_15 uf_16) +#692 := (not #566) +decl uf_21 :: T3 +#149 := uf_21 +decl uf_20 :: T3 +#148 := uf_20 +#150 := (= uf_20 uf_21) +#665 := (not #150) +decl uf_18 :: T2 +#144 := uf_18 +decl uf_17 :: T2 +#143 := uf_17 +#145 := (= uf_17 uf_18) +#683 := (not #145) +#4412 := (or #683 #665 #692 #674 #4353 #4409) +#108 := (uf_4 uf_14 uf_11) +#109 := (= #108 0::int) +#4415 := (not #4412) +#4573 := (or #4415 #4570) +#4576 := (not #4573) +decl ?x47!7 :: (-> T2 T2) +#2047 := (?x47!7 #11) +#2048 := (uf_4 uf_14 #2047) +#2671 := (* -1::int #2048) +#2686 := (+ #110 #2671) +#2687 := (<= #2686 0::int) +#2052 := (uf_1 #2047 #11) +#2053 := (uf_10 #2052) +#2672 := (* -1::int #2053) +#2673 := (+ #2671 #2672) +#2674 := (+ #110 #2673) +#2675 := (= #2674 0::int) +#3241 := (not #2675) +#2056 := (uf_6 uf_15 #2047) +#2057 := (= uf_8 #2056) +#3240 := (not #2057) +#3242 := (or #3240 #3241 #2687) +#3243 := (not #3242) +#3249 := (or #68 #1301 #3243) +#4339 := (forall (vars (?x46 T2)) (:pat #4305) #3249) +#4344 := (not #4339) +decl uf_12 :: (-> T2 int) +#69 := (uf_12 #11) +#4257 := (pattern #69) +decl ?x38!6 :: (-> T2 T2) +#2020 := (?x38!6 #11) +#2024 := (uf_12 #2020) +#2630 := (* -1::int #2024) +#2021 := (uf_1 #2020 #11) +#2022 := (uf_10 #2021) +#2647 := (* -1::int #2022) +#2648 := (+ #2647 #2630) +#2649 := (+ #69 #2648) +#2650 := (= #2649 0::int) +#3213 := (not #2650) +#2631 := (+ #69 #2630) +#2632 := (<= #2631 0::int) +decl up_13 :: (-> T2 bool) +#2030 := (up_13 #2020) +#3212 := (not #2030) +#3214 := (or #3212 #2632 #3213) +#3215 := (not #3214) +#1261 := (* -1::int #69) +#1262 := (+ uf_9 #1261) +#1263 := (<= #1262 0::int) +#3221 := (or #68 #1263 #3215) +#4331 := (forall (vars (?x37 T2)) (:pat #4257) #3221) +#4336 := (not #4331) +#117 := (uf_6 uf_15 #10) +#4322 := (pattern #114 #117) +#120 := (uf_4 uf_14 #10) +#1313 := (* -1::int #120) +#1314 := (+ #110 #1313) +#1317 := (>= #1314 0::int) +#484 := (= uf_8 #117) +#3178 := (not #484) +#3193 := (or #478 #3178 #1317) +#4323 := (forall (vars (?x42 T2) (?x43 T2)) (:pat #4322) #3193) +#4328 := (not #4323) +#1315 := (+ #91 #1314) +#1710 := (>= #1315 0::int) +#481 := (not #478) +#3170 := (or #481 #1237 #1710) +#4314 := (forall (vars (?x44 T2) (?x45 T2)) (:pat #4281) #3170) +#4319 := (not #4314) +#1738 := (>= #110 0::int) +#4306 := (forall (vars (?x41 T2)) (:pat #4305) #1738) +#4311 := (not #4306) +#1749 := (not #109) +#4579 := (or #1749 #4311 #4319 #4328 #4336 #4344 #4576) +#4582 := (not #4579) +decl ?x37!5 :: T2 +#1976 := ?x37!5 +#1977 := (uf_1 #11 ?x37!5) +#4290 := (pattern #1977) +#77 := (up_13 #11) +#4250 := (pattern #77) +#1979 := (uf_12 ?x37!5) +#1980 := (* -1::int #1979) +#1978 := (uf_10 #1977) +#2598 := (+ #1978 #1980) +#2599 := (+ #69 #2598) +#2602 := (= #2599 0::int) +#3131 := (not #2602) +#1984 := (+ #69 #1980) +#1985 := (>= #1984 0::int) +#78 := (not #77) +#3132 := (or #78 #1985 #3131) +#4291 := (forall (vars (?x38 T2)) (:pat #4250 #4257 #4290) #3132) +#4296 := (not #4291) +#2574 := (= uf_11 ?x37!5) +#1989 := (+ uf_9 #1980) +#1990 := (<= #1989 0::int) +#4299 := (or #1990 #2574 #4296) +#5019 := (= uf_9 #1979) +#5185 := (not #5019) +#1991 := (not #1990) +#4302 := (not #4299) +#5183 := [hypothesis]: #4302 +#4176 := (or #4299 #1991) +#3850 := [def-axiom]: #4176 +#5184 := [unit-resolution #3850 #5183]: #1991 +#5186 := (or #5185 #1990) +#5193 := [th-lemma]: #5186 +#5194 := [unit-resolution #5193 #5184]: #5185 +#2577 := (not #2574) +#3851 := (or #4299 #2577) +#4183 := [def-axiom]: #3851 +#5192 := [unit-resolution #4183 #5183]: #2577 +#437 := (= uf_9 #69) +#443 := (or #68 #437) +#4258 := (forall (vars (?x25 T2)) (:pat #4257) #443) +#448 := (forall (vars (?x25 T2)) #443) +#4261 := (iff #448 #4258) +#4259 := (iff #443 #443) +#4260 := [refl]: #4259 +#4262 := [quant-intro #4260]: #4261 +#1862 := (~ #448 #448) +#1900 := (~ #443 #443) +#1901 := [refl]: #1900 +#1863 := [nnf-pos #1901]: #1862 +#70 := (= #69 0::int) +#73 := (not #68) +#1807 := (or #73 #70) +#1810 := (forall (vars (?x24 T2)) #1807) +#1813 := (not #1810) +#1741 := (forall (vars (?x41 T2)) #1738) +#1744 := (not #1741) +#487 := (and #481 #484) +#493 := (not #487) +#1727 := (or #493 #1317) +#1732 := (forall (vars (?x42 T2) (?x43 T2)) #1727) +#1735 := (not #1732) +#1238 := (not #1237) +#1702 := (and #478 #1238) +#1707 := (not #1702) +#1713 := (or #1707 #1710) +#1716 := (forall (vars (?x44 T2) (?x45 T2)) #1713) +#1719 := (not #1716) +#1649 := (forall (vars (?x58 T2)) #1644) +#1652 := (not #1649) +#1459 := (not #1458) +#1452 := (not #1451) +#1462 := (and #1452 #1459) +#1620 := (not #1462) +#1628 := (or #1620 #1625) +#1631 := (forall (vars (?x59 T2)) #1628) +#1634 := (not #1631) +#1558 := (= #1536 0::int) +#1561 := (not #1504) +#1570 := (and #773 #1561 #1558) +#1575 := (exists (vars (?x76 T2)) #1570) +#1547 := (+ uf_9 #1480) +#1548 := (<= #1547 0::int) +#1549 := (not #1548) +#1552 := (and #73 #1549) +#1555 := (not #1552) +#1578 := (or #1555 #1575) +#1581 := (forall (vars (?x75 T2)) #1578) +#1526 := (and #773 #1238) +#1531 := (not #1526) +#1538 := (or #1531 #1534) +#1541 := (forall (vars (?x71 T2) (?x72 T2)) #1538) +#1544 := (not #1541) +#1584 := (or #1544 #1581) +#1587 := (and #1541 #1584) +#796 := (and #779 #793) +#802 := (not #796) +#1512 := (or #802 #1504) +#1517 := (forall (vars (?x67 T2) (?x68 T2)) #1512) +#1520 := (not #1517) +#1590 := (or #1520 #1587) +#1593 := (and #1517 #1590) +#1498 := (forall (vars (?x65 T2)) #1495) +#1501 := (not #1498) +#1596 := (or #1501 #1593) +#1599 := (and #1498 #1596) +#1602 := (or #1492 #1599) +#1605 := (and #217 #1602) +#785 := (forall (vars (?x63 T2)) #780) +#939 := (not #785) +#1608 := (or #939 #1605) +#1611 := (and #785 #1608) +#1484 := (forall (vars (?x61 T2)) #1479) +#1487 := (not #1484) +#1614 := (or #1487 #1611) +#1617 := (and #1484 #1614) +#1468 := (or #759 #1462) +#1473 := (forall (vars (?x60 T2)) #1468) +#1476 := (not #1473) +#1302 := (not #1301) +#1421 := (and #481 #1302) +#1426 := (exists (vars (?x48 T2)) #1421) +#1667 := (not #1426) +#1691 := (or #981 #728 #1667 #1476 #1617 #1634 #1652 #1656) +#1347 := (not #1346) +#1381 := (and #1238 #1347) +#1384 := (not #1381) +#1390 := (or #1384 #1387) +#1393 := (forall (vars (?x53 T2) (?x54 T2)) #1390) +#1396 := (not #1393) +#1404 := (or #167 #1396) +#1409 := (and #1393 #1404) +#1362 := (= #1364 0::int) +#1356 := (>= #1358 0::int) +#1359 := (not #1356) +#1366 := (and #1359 #1362) +#1369 := (exists (vars (?x50 T2)) #1366) +#1350 := (and #73 #1347) +#1353 := (not #1350) +#1372 := (or #1353 #1369) +#1375 := (forall (vars (?x49 T2)) #1372) +#1378 := (not #1375) +#1412 := (or #1378 #1409) +#1415 := (and #1375 #1412) +#1444 := (or #683 #665 #692 #674 #1415 #1426) +#1696 := (and #1444 #1691) +#1318 := (not #1317) +#1311 := (= #1315 0::int) +#1327 := (and #478 #1311 #1318) +#1332 := (exists (vars (?x47 T2)) #1327) +#1305 := (and #73 #1302) +#1308 := (not #1305) +#1335 := (or #1308 #1332) +#1338 := (forall (vars (?x46 T2)) #1335) +#1341 := (not #1338) +#86 := (uf_12 #10) +#1223 := (* -1::int #86) +#1250 := (+ #1223 #91) +#1251 := (+ #69 #1250) +#1273 := (= #1251 0::int) +#1224 := (+ #69 #1223) +#1222 := (>= #1224 0::int) +#1276 := (not #1222) +#1285 := (and #77 #1276 #1273) +#1290 := (exists (vars (?x38 T2)) #1285) +#1264 := (not #1263) +#1267 := (and #73 #1264) +#1270 := (not #1267) +#1293 := (or #1270 #1290) +#1296 := (forall (vars (?x37 T2)) #1293) +#1752 := (not #1296) +#1773 := (or #1749 #1752 #1341 #1696 #1719 #1735 #1744) +#1778 := (and #1296 #1773) +#1248 := (>= #1251 0::int) +#1241 := (and #77 #1238) +#1244 := (not #1241) +#1252 := (or #1244 #1248) +#1255 := (forall (vars (?x33 T2) (?x34 T2)) #1252) +#1258 := (not #1255) +#1781 := (or #1258 #1778) +#1784 := (and #1255 #1781) +#84 := (up_13 #10) +#85 := (and #78 #84) +#454 := (not #85) +#1226 := (or #454 #1222) +#1229 := (forall (vars (?x29 T2) (?x30 T2)) #1226) +#1232 := (not #1229) +#1787 := (or #1232 #1784) +#1790 := (and #1229 #1787) +#1213 := (>= #69 0::int) +#1214 := (forall (vars (?x27 T2)) #1213) +#1217 := (not #1214) +#1793 := (or #1217 #1790) +#1796 := (and #1214 #1793) +#80 := (uf_12 uf_11) +#81 := (= #80 0::int) +#1208 := (not #81) +#1799 := (or #1208 #1796) +#1802 := (and #81 #1799) +#1177 := (not #448) +#79 := (forall (vars (?x26 T2)) #78) +#1168 := (not #79) +#1825 := (or #1168 #1177 #1802 #1813) +#1830 := (not #1825) +#1 := true +#242 := (implies false true) +#229 := (+ #202 #91) +#236 := (= #224 #229) +#213 := (= #212 uf_8) +#237 := (and #213 #236) +#235 := (< #202 #224) +#238 := (and #235 #237) +#239 := (exists (vars (?x76 T2)) #238) +#233 := (< #202 uf_9) +#234 := (and #73 #233) +#240 := (implies #234 #239) +#241 := (forall (vars (?x75 T2)) #240) +#243 := (implies #241 #242) +#244 := (and #241 #243) +#230 := (<= #224 #229) +#92 := (< #91 uf_9) +#228 := (and #213 #92) +#231 := (implies #228 #230) +#232 := (forall (vars (?x71 T2) (?x72 T2)) #231) +#245 := (implies #232 #244) +#246 := (and #232 #245) +#225 := (<= #224 #202) +#222 := (= #221 uf_8) +#220 := (not #213) +#223 := (and #220 #222) +#226 := (implies #223 #225) +#227 := (forall (vars (?x67 T2) (?x68 T2)) #226) +#247 := (implies #227 #246) +#248 := (and #227 #247) +#218 := (<= 0::int #202) +#219 := (forall (vars (?x65 T2)) #218) +#249 := (implies #219 #248) +#250 := (and #219 #249) +#251 := (implies #217 #250) +#252 := (and #217 #251) +#253 := (implies true #252) +#254 := (implies true #253) +#207 := (= #202 #110) +#214 := (implies #213 #207) +#215 := (forall (vars (?x63 T2)) #214) +#255 := (implies #215 #254) +#256 := (and #215 #255) +#210 := (<= #202 #110) +#211 := (forall (vars (?x61 T2)) #210) +#257 := (implies #211 #256) +#258 := (and #211 #257) +#199 := (+ #188 #197) +#200 := (< #199 #110) +#198 := (< #197 uf_9) +#201 := (and #198 #200) +#206 := (not #201) +#208 := (implies #206 #207) +#209 := (forall (vars (?x60 T2)) #208) +#259 := (implies #209 #258) +#203 := (= #202 #199) +#204 := (implies #201 #203) +#205 := (forall (vars (?x59 T2)) #204) +#260 := (implies #205 #259) +#261 := (implies #195 #260) +#190 := (<= #188 #110) +#115 := (= #114 uf_8) +#116 := (not #115) +#191 := (implies #116 #190) +#192 := (forall (vars (?x58 T2)) #191) +#262 := (implies #192 #261) +#189 := (< #188 uf_9) +#263 := (implies #189 #262) +#186 := (= #185 uf_8) +#187 := (not #186) +#264 := (implies #187 #263) +#129 := (< #110 uf_9) +#138 := (and #116 #129) +#139 := (exists (vars (?x48 T2)) #138) +#265 := (implies #139 #264) +#266 := (implies true #265) +#267 := (implies true #266) +#168 := (implies #167 true) +#169 := (and #167 #168) +#156 := (+ #151 #91) +#163 := (<= #154 #156) +#152 := (< #151 uf_9) +#162 := (and #152 #92) +#164 := (implies #162 #163) +#165 := (forall (vars (?x53 T2) (?x54 T2)) #164) +#170 := (implies #165 #169) +#171 := (and #165 #170) +#157 := (= #154 #156) +#155 := (< #151 #154) +#158 := (and #155 #157) +#159 := (exists (vars (?x50 T2)) #158) +#153 := (and #73 #152) +#160 := (implies #153 #159) +#161 := (forall (vars (?x49 T2)) #160) +#172 := (implies #161 #171) +#173 := (and #161 #172) +#174 := (implies true #173) +#175 := (implies #150 #174) +#147 := (= uf_19 uf_14) +#176 := (implies #147 #175) +#177 := (implies #145 #176) +#142 := (= uf_16 uf_15) +#178 := (implies #142 #177) +#179 := (implies true #178) +#180 := (implies true #179) +#140 := (not #139) +#181 := (implies #140 #180) +#182 := (implies true #181) +#183 := (implies true #182) +#268 := (and #183 #267) +#269 := (implies true #268) +#125 := (+ #110 #91) +#132 := (= #120 #125) +#133 := (and #115 #132) +#131 := (< #110 #120) +#134 := (and #131 #133) +#135 := (exists (vars (?x47 T2)) #134) +#130 := (and #73 #129) +#136 := (implies #130 #135) +#137 := (forall (vars (?x46 T2)) #136) +#270 := (implies #137 #269) +#126 := (<= #120 #125) +#124 := (and #115 #92) +#127 := (implies #124 #126) +#128 := (forall (vars (?x44 T2) (?x45 T2)) #127) +#271 := (implies #128 #270) +#121 := (<= #120 #110) +#118 := (= #117 uf_8) +#119 := (and #116 #118) +#122 := (implies #119 #121) +#123 := (forall (vars (?x42 T2) (?x43 T2)) #122) +#272 := (implies #123 #271) +#111 := (<= 0::int #110) +#112 := (forall (vars (?x41 T2)) #111) +#273 := (implies #112 #272) +#274 := (implies #109 #273) +#275 := (implies true #274) +#276 := (implies true #275) +#94 := (+ #69 #91) +#101 := (= #86 #94) +#102 := (and #77 #101) +#100 := (< #69 #86) +#103 := (and #100 #102) +#104 := (exists (vars (?x38 T2)) #103) +#98 := (< #69 uf_9) +#99 := (and #73 #98) +#105 := (implies #99 #104) +#106 := (forall (vars (?x37 T2)) #105) +#277 := (implies #106 #276) +#278 := (and #106 #277) +#95 := (<= #86 #94) +#93 := (and #77 #92) +#96 := (implies #93 #95) +#97 := (forall (vars (?x33 T2) (?x34 T2)) #96) +#279 := (implies #97 #278) +#280 := (and #97 #279) +#87 := (<= #86 #69) +#88 := (implies #85 #87) +#89 := (forall (vars (?x29 T2) (?x30 T2)) #88) +#281 := (implies #89 #280) +#282 := (and #89 #281) +#82 := (<= 0::int #69) +#83 := (forall (vars (?x27 T2)) #82) +#283 := (implies #83 #282) +#284 := (and #83 #283) +#285 := (implies #81 #284) +#286 := (and #81 #285) +#287 := (implies true #286) +#288 := (implies #79 #287) +#74 := (= #69 uf_9) +#75 := (implies #73 #74) +#76 := (forall (vars (?x25 T2)) #75) +#289 := (implies #76 #288) +#71 := (implies #68 #70) +#72 := (forall (vars (?x24 T2)) #71) +#290 := (implies #72 #289) +#291 := (implies true #290) +#292 := (implies true #291) +#293 := (not #292) +#1833 := (iff #293 #1830) +#819 := (+ #91 #202) +#837 := (= #224 #819) +#840 := (and #773 #837) +#843 := (and #235 #840) +#846 := (exists (vars (?x76 T2)) #843) +#852 := (not #234) +#853 := (or #852 #846) +#858 := (forall (vars (?x75 T2)) #853) +#822 := (<= #224 #819) +#814 := (and #92 #773) +#828 := (not #814) +#829 := (or #828 #822) +#834 := (forall (vars (?x71 T2) (?x72 T2)) #829) +#880 := (not #834) +#881 := (or #880 #858) +#886 := (and #834 #881) +#803 := (or #225 #802) +#808 := (forall (vars (?x67 T2) (?x68 T2)) #803) +#892 := (not #808) +#893 := (or #892 #886) +#898 := (and #808 #893) +#904 := (not #219) +#905 := (or #904 #898) +#910 := (and #219 #905) +#788 := (= 0::int #216) +#916 := (not #788) +#917 := (or #916 #910) +#922 := (and #788 #917) +#940 := (or #939 #922) +#945 := (and #785 #940) +#951 := (not #211) +#952 := (or #951 #945) +#957 := (and #211 #952) +#765 := (or #201 #759) +#770 := (forall (vars (?x60 T2)) #765) +#963 := (not #770) +#964 := (or #963 #957) +#745 := (= #199 #202) +#751 := (or #206 #745) +#756 := (forall (vars (?x59 T2)) #751) +#972 := (not #756) +#973 := (or #972 #964) +#982 := (or #981 #973) +#737 := (or #190 #478) +#742 := (forall (vars (?x58 T2)) #737) +#990 := (not #742) +#991 := (or #990 #982) +#999 := (not #189) +#1000 := (or #999 #991) +#1008 := (or #728 #1000) +#555 := (and #129 #481) +#560 := (exists (vars (?x48 T2)) #555) +#563 := (not #560) +#1016 := (or #563 #1008) +#614 := (= 0::int #166) +#572 := (+ #91 #151) +#599 := (<= #154 #572) +#596 := (and #92 #152) +#605 := (not #596) +#606 := (or #605 #599) +#611 := (forall (vars (?x53 T2) (?x54 T2)) #606) +#634 := (not #611) +#635 := (or #634 #614) +#640 := (and #611 #635) +#575 := (= #154 #572) +#578 := (and #155 #575) +#581 := (exists (vars (?x50 T2)) #578) +#587 := (not #153) +#588 := (or #587 #581) +#593 := (forall (vars (?x49 T2)) #588) +#646 := (not #593) +#647 := (or #646 #640) +#652 := (and #593 #647) +#666 := (or #665 #652) +#675 := (or #674 #666) +#684 := (or #683 #675) +#693 := (or #692 #684) +#712 := (or #560 #693) +#1032 := (and #712 #1016) +#510 := (+ #91 #110) +#528 := (= #120 #510) +#531 := (and #478 #528) +#534 := (and #131 #531) +#537 := (exists (vars (?x47 T2)) #534) +#543 := (not #130) +#544 := (or #543 #537) +#549 := (forall (vars (?x46 T2)) #544) +#1045 := (not #549) +#1046 := (or #1045 #1032) +#513 := (<= #120 #510) +#505 := (and #92 #478) +#519 := (not #505) +#520 := (or #519 #513) +#525 := (forall (vars (?x44 T2) (?x45 T2)) #520) +#1054 := (not #525) +#1055 := (or #1054 #1046) +#494 := (or #121 #493) +#499 := (forall (vars (?x42 T2) (?x43 T2)) #494) +#1063 := (not #499) +#1064 := (or #1063 #1055) +#1072 := (not #112) +#1073 := (or #1072 #1064) +#475 := (= 0::int #108) +#1081 := (not #475) +#1082 := (or #1081 #1073) +#468 := (not #99) +#469 := (or #468 #104) +#472 := (forall (vars (?x37 T2)) #469) +#1101 := (not #472) +#1102 := (or #1101 #1082) +#1107 := (and #472 #1102) +#461 := (not #93) +#462 := (or #461 #95) +#465 := (forall (vars (?x33 T2) (?x34 T2)) #462) +#1113 := (not #465) +#1114 := (or #1113 #1107) +#1119 := (and #465 #1114) +#455 := (or #454 #87) +#458 := (forall (vars (?x29 T2) (?x30 T2)) #455) +#1125 := (not #458) +#1126 := (or #1125 #1119) +#1131 := (and #458 #1126) +#1137 := (not #83) +#1138 := (or #1137 #1131) +#1143 := (and #83 #1138) +#451 := (= 0::int #80) +#1149 := (not #451) +#1150 := (or #1149 #1143) +#1155 := (and #451 #1150) +#1169 := (or #1168 #1155) +#1178 := (or #1177 #1169) +#423 := (= 0::int #69) +#429 := (or #73 #423) +#434 := (forall (vars (?x24 T2)) #429) +#1186 := (not #434) +#1187 := (or #1186 #1178) +#1203 := (not #1187) +#1831 := (iff #1203 #1830) +#1828 := (iff #1187 #1825) +#1816 := (or #1168 #1802) +#1819 := (or #1177 #1816) +#1822 := (or #1813 #1819) +#1826 := (iff #1822 #1825) +#1827 := [rewrite]: #1826 +#1823 := (iff #1187 #1822) +#1820 := (iff #1178 #1819) +#1817 := (iff #1169 #1816) +#1803 := (iff #1155 #1802) +#1800 := (iff #1150 #1799) +#1797 := (iff #1143 #1796) +#1794 := (iff #1138 #1793) +#1791 := (iff #1131 #1790) +#1788 := (iff #1126 #1787) +#1785 := (iff #1119 #1784) +#1782 := (iff #1114 #1781) +#1779 := (iff #1107 #1778) +#1776 := (iff #1102 #1773) +#1755 := (or #1341 #1696) +#1758 := (or #1719 #1755) +#1761 := (or #1735 #1758) +#1764 := (or #1744 #1761) +#1767 := (or #1749 #1764) +#1770 := (or #1752 #1767) +#1774 := (iff #1770 #1773) +#1775 := [rewrite]: #1774 +#1771 := (iff #1102 #1770) +#1768 := (iff #1082 #1767) +#1765 := (iff #1073 #1764) +#1762 := (iff #1064 #1761) +#1759 := (iff #1055 #1758) +#1756 := (iff #1046 #1755) +#1697 := (iff #1032 #1696) +#1694 := (iff #1016 #1691) +#1670 := (or #1476 #1617) +#1673 := (or #1634 #1670) +#1676 := (or #981 #1673) +#1679 := (or #1652 #1676) +#1682 := (or #1656 #1679) +#1685 := (or #728 #1682) +#1688 := (or #1667 #1685) +#1692 := (iff #1688 #1691) +#1693 := [rewrite]: #1692 +#1689 := (iff #1016 #1688) +#1686 := (iff #1008 #1685) +#1683 := (iff #1000 #1682) +#1680 := (iff #991 #1679) +#1677 := (iff #982 #1676) +#1674 := (iff #973 #1673) +#1671 := (iff #964 #1670) +#1618 := (iff #957 #1617) +#1615 := (iff #952 #1614) +#1612 := (iff #945 #1611) +#1609 := (iff #940 #1608) +#1606 := (iff #922 #1605) +#1603 := (iff #917 #1602) +#1600 := (iff #910 #1599) +#1597 := (iff #905 #1596) +#1594 := (iff #898 #1593) +#1591 := (iff #893 #1590) +#1588 := (iff #886 #1587) +#1585 := (iff #881 #1584) +#1582 := (iff #858 #1581) +#1579 := (iff #853 #1578) +#1576 := (iff #846 #1575) +#1573 := (iff #843 #1570) +#1564 := (and #773 #1558) +#1567 := (and #1561 #1564) +#1571 := (iff #1567 #1570) +#1572 := [rewrite]: #1571 +#1568 := (iff #843 #1567) +#1565 := (iff #840 #1564) +#1559 := (iff #837 #1558) +#1560 := [rewrite]: #1559 +#1566 := [monotonicity #1560]: #1565 +#1562 := (iff #235 #1561) +#1563 := [rewrite]: #1562 +#1569 := [monotonicity #1563 #1566]: #1568 +#1574 := [trans #1569 #1572]: #1573 +#1577 := [quant-intro #1574]: #1576 +#1556 := (iff #852 #1555) +#1553 := (iff #234 #1552) +#1550 := (iff #233 #1549) +#1551 := [rewrite]: #1550 +#1554 := [monotonicity #1551]: #1553 +#1557 := [monotonicity #1554]: #1556 +#1580 := [monotonicity #1557 #1577]: #1579 +#1583 := [quant-intro #1580]: #1582 +#1545 := (iff #880 #1544) +#1542 := (iff #834 #1541) +#1539 := (iff #829 #1538) +#1535 := (iff #822 #1534) +#1537 := [rewrite]: #1535 +#1532 := (iff #828 #1531) +#1529 := (iff #814 #1526) +#1523 := (and #1238 #773) +#1527 := (iff #1523 #1526) +#1528 := [rewrite]: #1527 +#1524 := (iff #814 #1523) +#1239 := (iff #92 #1238) +#1240 := [rewrite]: #1239 +#1525 := [monotonicity #1240]: #1524 +#1530 := [trans #1525 #1528]: #1529 +#1533 := [monotonicity #1530]: #1532 +#1540 := [monotonicity #1533 #1537]: #1539 +#1543 := [quant-intro #1540]: #1542 +#1546 := [monotonicity #1543]: #1545 +#1586 := [monotonicity #1546 #1583]: #1585 +#1589 := [monotonicity #1543 #1586]: #1588 +#1521 := (iff #892 #1520) +#1518 := (iff #808 #1517) +#1515 := (iff #803 #1512) +#1509 := (or #1504 #802) +#1513 := (iff #1509 #1512) +#1514 := [rewrite]: #1513 +#1510 := (iff #803 #1509) +#1507 := (iff #225 #1504) +#1508 := [rewrite]: #1507 +#1511 := [monotonicity #1508]: #1510 +#1516 := [trans #1511 #1514]: #1515 +#1519 := [quant-intro #1516]: #1518 +#1522 := [monotonicity #1519]: #1521 +#1592 := [monotonicity #1522 #1589]: #1591 +#1595 := [monotonicity #1519 #1592]: #1594 +#1502 := (iff #904 #1501) +#1499 := (iff #219 #1498) +#1496 := (iff #218 #1495) +#1497 := [rewrite]: #1496 +#1500 := [quant-intro #1497]: #1499 +#1503 := [monotonicity #1500]: #1502 +#1598 := [monotonicity #1503 #1595]: #1597 +#1601 := [monotonicity #1500 #1598]: #1600 +#1493 := (iff #916 #1492) +#1490 := (iff #788 #217) +#1491 := [rewrite]: #1490 +#1494 := [monotonicity #1491]: #1493 +#1604 := [monotonicity #1494 #1601]: #1603 +#1607 := [monotonicity #1491 #1604]: #1606 +#1610 := [monotonicity #1607]: #1609 +#1613 := [monotonicity #1610]: #1612 +#1488 := (iff #951 #1487) +#1485 := (iff #211 #1484) +#1482 := (iff #210 #1479) +#1483 := [rewrite]: #1482 +#1486 := [quant-intro #1483]: #1485 +#1489 := [monotonicity #1486]: #1488 +#1616 := [monotonicity #1489 #1613]: #1615 +#1619 := [monotonicity #1486 #1616]: #1618 +#1477 := (iff #963 #1476) +#1474 := (iff #770 #1473) +#1471 := (iff #765 #1468) +#1465 := (or #1462 #759) +#1469 := (iff #1465 #1468) +#1470 := [rewrite]: #1469 +#1466 := (iff #765 #1465) +#1463 := (iff #201 #1462) +#1460 := (iff #200 #1459) +#1461 := [rewrite]: #1460 +#1453 := (iff #198 #1452) +#1454 := [rewrite]: #1453 +#1464 := [monotonicity #1454 #1461]: #1463 +#1467 := [monotonicity #1464]: #1466 +#1472 := [trans #1467 #1470]: #1471 +#1475 := [quant-intro #1472]: #1474 +#1478 := [monotonicity #1475]: #1477 +#1672 := [monotonicity #1478 #1619]: #1671 +#1635 := (iff #972 #1634) +#1632 := (iff #756 #1631) +#1629 := (iff #751 #1628) +#1626 := (iff #745 #1625) +#1627 := [rewrite]: #1626 +#1621 := (iff #206 #1620) +#1622 := [monotonicity #1464]: #1621 +#1630 := [monotonicity #1622 #1627]: #1629 +#1633 := [quant-intro #1630]: #1632 +#1636 := [monotonicity #1633]: #1635 +#1675 := [monotonicity #1636 #1672]: #1674 +#1678 := [monotonicity #1675]: #1677 +#1653 := (iff #990 #1652) +#1650 := (iff #742 #1649) +#1647 := (iff #737 #1644) +#1641 := (or #1637 #478) +#1645 := (iff #1641 #1644) +#1646 := [rewrite]: #1645 +#1642 := (iff #737 #1641) +#1639 := (iff #190 #1637) +#1640 := [rewrite]: #1639 +#1643 := [monotonicity #1640]: #1642 +#1648 := [trans #1643 #1646]: #1647 +#1651 := [quant-intro #1648]: #1650 +#1654 := [monotonicity #1651]: #1653 +#1681 := [monotonicity #1654 #1678]: #1680 +#1665 := (iff #999 #1656) +#1657 := (not #1656) +#1660 := (not #1657) +#1663 := (iff #1660 #1656) +#1664 := [rewrite]: #1663 +#1661 := (iff #999 #1660) +#1658 := (iff #189 #1657) +#1659 := [rewrite]: #1658 +#1662 := [monotonicity #1659]: #1661 +#1666 := [trans #1662 #1664]: #1665 +#1684 := [monotonicity #1666 #1681]: #1683 +#1687 := [monotonicity #1684]: #1686 +#1668 := (iff #563 #1667) +#1427 := (iff #560 #1426) +#1424 := (iff #555 #1421) +#1418 := (and #1302 #481) +#1422 := (iff #1418 #1421) +#1423 := [rewrite]: #1422 +#1419 := (iff #555 #1418) +#1303 := (iff #129 #1302) +#1304 := [rewrite]: #1303 +#1420 := [monotonicity #1304]: #1419 +#1425 := [trans #1420 #1423]: #1424 +#1428 := [quant-intro #1425]: #1427 +#1669 := [monotonicity #1428]: #1668 +#1690 := [monotonicity #1669 #1687]: #1689 +#1695 := [trans #1690 #1693]: #1694 +#1447 := (iff #712 #1444) +#1429 := (or #665 #1415) +#1432 := (or #674 #1429) +#1435 := (or #683 #1432) +#1438 := (or #692 #1435) +#1441 := (or #1426 #1438) +#1445 := (iff #1441 #1444) +#1446 := [rewrite]: #1445 +#1442 := (iff #712 #1441) +#1439 := (iff #693 #1438) +#1436 := (iff #684 #1435) +#1433 := (iff #675 #1432) +#1430 := (iff #666 #1429) +#1416 := (iff #652 #1415) +#1413 := (iff #647 #1412) +#1410 := (iff #640 #1409) +#1407 := (iff #635 #1404) +#1401 := (or #1396 #167) +#1405 := (iff #1401 #1404) +#1406 := [rewrite]: #1405 +#1402 := (iff #635 #1401) +#1399 := (iff #614 #167) +#1400 := [rewrite]: #1399 +#1397 := (iff #634 #1396) +#1394 := (iff #611 #1393) +#1391 := (iff #606 #1390) +#1388 := (iff #599 #1387) +#1389 := [rewrite]: #1388 +#1385 := (iff #605 #1384) +#1382 := (iff #596 #1381) +#1348 := (iff #152 #1347) +#1349 := [rewrite]: #1348 +#1383 := [monotonicity #1240 #1349]: #1382 +#1386 := [monotonicity #1383]: #1385 +#1392 := [monotonicity #1386 #1389]: #1391 +#1395 := [quant-intro #1392]: #1394 +#1398 := [monotonicity #1395]: #1397 +#1403 := [monotonicity #1398 #1400]: #1402 +#1408 := [trans #1403 #1406]: #1407 +#1411 := [monotonicity #1395 #1408]: #1410 +#1379 := (iff #646 #1378) +#1376 := (iff #593 #1375) +#1373 := (iff #588 #1372) +#1370 := (iff #581 #1369) +#1367 := (iff #578 #1366) +#1363 := (iff #575 #1362) +#1365 := [rewrite]: #1363 +#1360 := (iff #155 #1359) +#1361 := [rewrite]: #1360 +#1368 := [monotonicity #1361 #1365]: #1367 +#1371 := [quant-intro #1368]: #1370 +#1354 := (iff #587 #1353) +#1351 := (iff #153 #1350) +#1352 := [monotonicity #1349]: #1351 +#1355 := [monotonicity #1352]: #1354 +#1374 := [monotonicity #1355 #1371]: #1373 +#1377 := [quant-intro #1374]: #1376 +#1380 := [monotonicity #1377]: #1379 +#1414 := [monotonicity #1380 #1411]: #1413 +#1417 := [monotonicity #1377 #1414]: #1416 +#1431 := [monotonicity #1417]: #1430 +#1434 := [monotonicity #1431]: #1433 +#1437 := [monotonicity #1434]: #1436 +#1440 := [monotonicity #1437]: #1439 +#1443 := [monotonicity #1428 #1440]: #1442 +#1448 := [trans #1443 #1446]: #1447 +#1698 := [monotonicity #1448 #1695]: #1697 +#1342 := (iff #1045 #1341) +#1339 := (iff #549 #1338) +#1336 := (iff #544 #1335) +#1333 := (iff #537 #1332) +#1330 := (iff #534 #1327) +#1321 := (and #478 #1311) +#1324 := (and #1318 #1321) +#1328 := (iff #1324 #1327) +#1329 := [rewrite]: #1328 +#1325 := (iff #534 #1324) +#1322 := (iff #531 #1321) +#1312 := (iff #528 #1311) +#1316 := [rewrite]: #1312 +#1323 := [monotonicity #1316]: #1322 +#1319 := (iff #131 #1318) +#1320 := [rewrite]: #1319 +#1326 := [monotonicity #1320 #1323]: #1325 +#1331 := [trans #1326 #1329]: #1330 +#1334 := [quant-intro #1331]: #1333 +#1309 := (iff #543 #1308) +#1306 := (iff #130 #1305) +#1307 := [monotonicity #1304]: #1306 +#1310 := [monotonicity #1307]: #1309 +#1337 := [monotonicity #1310 #1334]: #1336 +#1340 := [quant-intro #1337]: #1339 +#1343 := [monotonicity #1340]: #1342 +#1757 := [monotonicity #1343 #1698]: #1756 +#1720 := (iff #1054 #1719) +#1717 := (iff #525 #1716) +#1714 := (iff #520 #1713) +#1711 := (iff #513 #1710) +#1712 := [rewrite]: #1711 +#1708 := (iff #519 #1707) +#1705 := (iff #505 #1702) +#1699 := (and #1238 #478) +#1703 := (iff #1699 #1702) +#1704 := [rewrite]: #1703 +#1700 := (iff #505 #1699) +#1701 := [monotonicity #1240]: #1700 +#1706 := [trans #1701 #1704]: #1705 +#1709 := [monotonicity #1706]: #1708 +#1715 := [monotonicity #1709 #1712]: #1714 +#1718 := [quant-intro #1715]: #1717 +#1721 := [monotonicity #1718]: #1720 +#1760 := [monotonicity #1721 #1757]: #1759 +#1736 := (iff #1063 #1735) +#1733 := (iff #499 #1732) +#1730 := (iff #494 #1727) +#1724 := (or #1317 #493) +#1728 := (iff #1724 #1727) +#1729 := [rewrite]: #1728 +#1725 := (iff #494 #1724) +#1722 := (iff #121 #1317) +#1723 := [rewrite]: #1722 +#1726 := [monotonicity #1723]: #1725 +#1731 := [trans #1726 #1729]: #1730 +#1734 := [quant-intro #1731]: #1733 +#1737 := [monotonicity #1734]: #1736 +#1763 := [monotonicity #1737 #1760]: #1762 +#1745 := (iff #1072 #1744) +#1742 := (iff #112 #1741) +#1739 := (iff #111 #1738) +#1740 := [rewrite]: #1739 +#1743 := [quant-intro #1740]: #1742 +#1746 := [monotonicity #1743]: #1745 +#1766 := [monotonicity #1746 #1763]: #1765 +#1750 := (iff #1081 #1749) +#1747 := (iff #475 #109) +#1748 := [rewrite]: #1747 +#1751 := [monotonicity #1748]: #1750 +#1769 := [monotonicity #1751 #1766]: #1768 +#1753 := (iff #1101 #1752) +#1297 := (iff #472 #1296) +#1294 := (iff #469 #1293) +#1291 := (iff #104 #1290) +#1288 := (iff #103 #1285) +#1279 := (and #77 #1273) +#1282 := (and #1276 #1279) +#1286 := (iff #1282 #1285) +#1287 := [rewrite]: #1286 +#1283 := (iff #103 #1282) +#1280 := (iff #102 #1279) +#1274 := (iff #101 #1273) +#1275 := [rewrite]: #1274 +#1281 := [monotonicity #1275]: #1280 +#1277 := (iff #100 #1276) +#1278 := [rewrite]: #1277 +#1284 := [monotonicity #1278 #1281]: #1283 +#1289 := [trans #1284 #1287]: #1288 +#1292 := [quant-intro #1289]: #1291 +#1271 := (iff #468 #1270) +#1268 := (iff #99 #1267) +#1265 := (iff #98 #1264) +#1266 := [rewrite]: #1265 +#1269 := [monotonicity #1266]: #1268 +#1272 := [monotonicity #1269]: #1271 +#1295 := [monotonicity #1272 #1292]: #1294 +#1298 := [quant-intro #1295]: #1297 +#1754 := [monotonicity #1298]: #1753 +#1772 := [monotonicity #1754 #1769]: #1771 +#1777 := [trans #1772 #1775]: #1776 +#1780 := [monotonicity #1298 #1777]: #1779 +#1259 := (iff #1113 #1258) +#1256 := (iff #465 #1255) +#1253 := (iff #462 #1252) +#1247 := (iff #95 #1248) +#1249 := [rewrite]: #1247 +#1245 := (iff #461 #1244) +#1242 := (iff #93 #1241) +#1243 := [monotonicity #1240]: #1242 +#1246 := [monotonicity #1243]: #1245 +#1254 := [monotonicity #1246 #1249]: #1253 +#1257 := [quant-intro #1254]: #1256 +#1260 := [monotonicity #1257]: #1259 +#1783 := [monotonicity #1260 #1780]: #1782 +#1786 := [monotonicity #1257 #1783]: #1785 +#1233 := (iff #1125 #1232) +#1230 := (iff #458 #1229) +#1227 := (iff #455 #1226) +#1221 := (iff #87 #1222) +#1225 := [rewrite]: #1221 +#1228 := [monotonicity #1225]: #1227 +#1231 := [quant-intro #1228]: #1230 +#1234 := [monotonicity #1231]: #1233 +#1789 := [monotonicity #1234 #1786]: #1788 +#1792 := [monotonicity #1231 #1789]: #1791 +#1218 := (iff #1137 #1217) +#1215 := (iff #83 #1214) +#1211 := (iff #82 #1213) +#1212 := [rewrite]: #1211 +#1216 := [quant-intro #1212]: #1215 +#1219 := [monotonicity #1216]: #1218 +#1795 := [monotonicity #1219 #1792]: #1794 +#1798 := [monotonicity #1216 #1795]: #1797 +#1209 := (iff #1149 #1208) +#1206 := (iff #451 #81) +#1207 := [rewrite]: #1206 +#1210 := [monotonicity #1207]: #1209 +#1801 := [monotonicity #1210 #1798]: #1800 +#1804 := [monotonicity #1207 #1801]: #1803 +#1818 := [monotonicity #1804]: #1817 +#1821 := [monotonicity #1818]: #1820 +#1814 := (iff #1186 #1813) +#1811 := (iff #434 #1810) +#1808 := (iff #429 #1807) +#1805 := (iff #423 #70) +#1806 := [rewrite]: #1805 +#1809 := [monotonicity #1806]: #1808 +#1812 := [quant-intro #1809]: #1811 +#1815 := [monotonicity #1812]: #1814 +#1824 := [monotonicity #1815 #1821]: #1823 +#1829 := [trans #1824 #1827]: #1828 +#1832 := [monotonicity #1829]: #1831 +#1204 := (iff #293 #1203) +#1201 := (iff #292 #1187) +#1192 := (implies true #1187) +#1195 := (iff #1192 #1187) +#1196 := [rewrite]: #1195 +#1199 := (iff #292 #1192) +#1197 := (iff #291 #1187) +#1193 := (iff #291 #1192) +#1190 := (iff #290 #1187) +#1183 := (implies #434 #1178) +#1188 := (iff #1183 #1187) +#1189 := [rewrite]: #1188 +#1184 := (iff #290 #1183) +#1181 := (iff #289 #1178) +#1174 := (implies #448 #1169) +#1179 := (iff #1174 #1178) +#1180 := [rewrite]: #1179 +#1175 := (iff #289 #1174) +#1172 := (iff #288 #1169) +#1165 := (implies #79 #1155) +#1170 := (iff #1165 #1169) +#1171 := [rewrite]: #1170 +#1166 := (iff #288 #1165) +#1163 := (iff #287 #1155) +#1158 := (implies true #1155) +#1161 := (iff #1158 #1155) +#1162 := [rewrite]: #1161 +#1159 := (iff #287 #1158) +#1156 := (iff #286 #1155) +#1153 := (iff #285 #1150) +#1146 := (implies #451 #1143) +#1151 := (iff #1146 #1150) +#1152 := [rewrite]: #1151 +#1147 := (iff #285 #1146) +#1144 := (iff #284 #1143) +#1141 := (iff #283 #1138) +#1134 := (implies #83 #1131) +#1139 := (iff #1134 #1138) +#1140 := [rewrite]: #1139 +#1135 := (iff #283 #1134) +#1132 := (iff #282 #1131) +#1129 := (iff #281 #1126) +#1122 := (implies #458 #1119) +#1127 := (iff #1122 #1126) +#1128 := [rewrite]: #1127 +#1123 := (iff #281 #1122) +#1120 := (iff #280 #1119) +#1117 := (iff #279 #1114) +#1110 := (implies #465 #1107) +#1115 := (iff #1110 #1114) +#1116 := [rewrite]: #1115 +#1111 := (iff #279 #1110) +#1108 := (iff #278 #1107) +#1105 := (iff #277 #1102) +#1098 := (implies #472 #1082) +#1103 := (iff #1098 #1102) +#1104 := [rewrite]: #1103 +#1099 := (iff #277 #1098) +#1096 := (iff #276 #1082) +#1087 := (implies true #1082) +#1090 := (iff #1087 #1082) +#1091 := [rewrite]: #1090 +#1094 := (iff #276 #1087) +#1092 := (iff #275 #1082) +#1088 := (iff #275 #1087) +#1085 := (iff #274 #1082) +#1078 := (implies #475 #1073) +#1083 := (iff #1078 #1082) +#1084 := [rewrite]: #1083 +#1079 := (iff #274 #1078) +#1076 := (iff #273 #1073) +#1069 := (implies #112 #1064) +#1074 := (iff #1069 #1073) +#1075 := [rewrite]: #1074 +#1070 := (iff #273 #1069) +#1067 := (iff #272 #1064) +#1060 := (implies #499 #1055) +#1065 := (iff #1060 #1064) +#1066 := [rewrite]: #1065 +#1061 := (iff #272 #1060) +#1058 := (iff #271 #1055) +#1051 := (implies #525 #1046) +#1056 := (iff #1051 #1055) +#1057 := [rewrite]: #1056 +#1052 := (iff #271 #1051) +#1049 := (iff #270 #1046) +#1042 := (implies #549 #1032) +#1047 := (iff #1042 #1046) +#1048 := [rewrite]: #1047 +#1043 := (iff #270 #1042) +#1040 := (iff #269 #1032) +#1035 := (implies true #1032) +#1038 := (iff #1035 #1032) +#1039 := [rewrite]: #1038 +#1036 := (iff #269 #1035) +#1033 := (iff #268 #1032) +#1030 := (iff #267 #1016) +#1021 := (implies true #1016) +#1024 := (iff #1021 #1016) +#1025 := [rewrite]: #1024 +#1028 := (iff #267 #1021) +#1026 := (iff #266 #1016) +#1022 := (iff #266 #1021) +#1019 := (iff #265 #1016) +#1013 := (implies #560 #1008) +#1017 := (iff #1013 #1016) +#1018 := [rewrite]: #1017 +#1014 := (iff #265 #1013) +#1011 := (iff #264 #1008) +#731 := (not #728) +#1005 := (implies #731 #1000) +#1009 := (iff #1005 #1008) +#1010 := [rewrite]: #1009 +#1006 := (iff #264 #1005) +#1003 := (iff #263 #1000) +#996 := (implies #189 #991) +#1001 := (iff #996 #1000) +#1002 := [rewrite]: #1001 +#997 := (iff #263 #996) +#994 := (iff #262 #991) +#987 := (implies #742 #982) +#992 := (iff #987 #991) +#993 := [rewrite]: #992 +#988 := (iff #262 #987) +#985 := (iff #261 #982) +#978 := (implies #195 #973) +#983 := (iff #978 #982) +#984 := [rewrite]: #983 +#979 := (iff #261 #978) +#976 := (iff #260 #973) +#969 := (implies #756 #964) +#974 := (iff #969 #973) +#975 := [rewrite]: #974 +#970 := (iff #260 #969) +#967 := (iff #259 #964) +#960 := (implies #770 #957) +#965 := (iff #960 #964) +#966 := [rewrite]: #965 +#961 := (iff #259 #960) +#958 := (iff #258 #957) +#955 := (iff #257 #952) +#948 := (implies #211 #945) +#953 := (iff #948 #952) +#954 := [rewrite]: #953 +#949 := (iff #257 #948) +#946 := (iff #256 #945) +#943 := (iff #255 #940) +#936 := (implies #785 #922) +#941 := (iff #936 #940) +#942 := [rewrite]: #941 +#937 := (iff #255 #936) +#934 := (iff #254 #922) +#925 := (implies true #922) +#928 := (iff #925 #922) +#929 := [rewrite]: #928 +#932 := (iff #254 #925) +#930 := (iff #253 #922) +#926 := (iff #253 #925) +#923 := (iff #252 #922) +#920 := (iff #251 #917) +#913 := (implies #788 #910) +#918 := (iff #913 #917) +#919 := [rewrite]: #918 +#914 := (iff #251 #913) +#911 := (iff #250 #910) +#908 := (iff #249 #905) +#901 := (implies #219 #898) +#906 := (iff #901 #905) +#907 := [rewrite]: #906 +#902 := (iff #249 #901) +#899 := (iff #248 #898) +#896 := (iff #247 #893) +#889 := (implies #808 #886) +#894 := (iff #889 #893) +#895 := [rewrite]: #894 +#890 := (iff #247 #889) +#887 := (iff #246 #886) +#884 := (iff #245 #881) +#877 := (implies #834 #858) +#882 := (iff #877 #881) +#883 := [rewrite]: #882 +#878 := (iff #245 #877) +#875 := (iff #244 #858) +#870 := (and #858 true) +#873 := (iff #870 #858) +#874 := [rewrite]: #873 +#871 := (iff #244 #870) +#868 := (iff #243 true) +#863 := (implies #858 true) +#866 := (iff #863 true) +#867 := [rewrite]: #866 +#864 := (iff #243 #863) +#861 := (iff #242 true) +#862 := [rewrite]: #861 +#859 := (iff #241 #858) +#856 := (iff #240 #853) +#849 := (implies #234 #846) +#854 := (iff #849 #853) +#855 := [rewrite]: #854 +#850 := (iff #240 #849) +#847 := (iff #239 #846) +#844 := (iff #238 #843) +#841 := (iff #237 #840) +#838 := (iff #236 #837) +#820 := (= #229 #819) +#821 := [rewrite]: #820 +#839 := [monotonicity #821]: #838 +#774 := (iff #213 #773) +#775 := [rewrite]: #774 +#842 := [monotonicity #775 #839]: #841 +#845 := [monotonicity #842]: #844 +#848 := [quant-intro #845]: #847 +#851 := [monotonicity #848]: #850 +#857 := [trans #851 #855]: #856 +#860 := [quant-intro #857]: #859 +#865 := [monotonicity #860 #862]: #864 +#869 := [trans #865 #867]: #868 +#872 := [monotonicity #860 #869]: #871 +#876 := [trans #872 #874]: #875 +#835 := (iff #232 #834) +#832 := (iff #231 #829) +#825 := (implies #814 #822) +#830 := (iff #825 #829) +#831 := [rewrite]: #830 +#826 := (iff #231 #825) +#823 := (iff #230 #822) +#824 := [monotonicity #821]: #823 +#817 := (iff #228 #814) +#811 := (and #773 #92) +#815 := (iff #811 #814) +#816 := [rewrite]: #815 +#812 := (iff #228 #811) +#813 := [monotonicity #775]: #812 +#818 := [trans #813 #816]: #817 +#827 := [monotonicity #818 #824]: #826 +#833 := [trans #827 #831]: #832 +#836 := [quant-intro #833]: #835 +#879 := [monotonicity #836 #876]: #878 +#885 := [trans #879 #883]: #884 +#888 := [monotonicity #836 #885]: #887 +#809 := (iff #227 #808) +#806 := (iff #226 #803) +#799 := (implies #796 #225) +#804 := (iff #799 #803) +#805 := [rewrite]: #804 +#800 := (iff #226 #799) +#797 := (iff #223 #796) +#794 := (iff #222 #793) +#795 := [rewrite]: #794 +#791 := (iff #220 #779) +#792 := [monotonicity #775]: #791 +#798 := [monotonicity #792 #795]: #797 +#801 := [monotonicity #798]: #800 +#807 := [trans #801 #805]: #806 +#810 := [quant-intro #807]: #809 +#891 := [monotonicity #810 #888]: #890 +#897 := [trans #891 #895]: #896 +#900 := [monotonicity #810 #897]: #899 +#903 := [monotonicity #900]: #902 +#909 := [trans #903 #907]: #908 +#912 := [monotonicity #909]: #911 +#789 := (iff #217 #788) +#790 := [rewrite]: #789 +#915 := [monotonicity #790 #912]: #914 +#921 := [trans #915 #919]: #920 +#924 := [monotonicity #790 #921]: #923 +#927 := [monotonicity #924]: #926 +#931 := [trans #927 #929]: #930 +#933 := [monotonicity #931]: #932 +#935 := [trans #933 #929]: #934 +#786 := (iff #215 #785) +#783 := (iff #214 #780) +#776 := (implies #773 #759) +#781 := (iff #776 #780) +#782 := [rewrite]: #781 +#777 := (iff #214 #776) +#760 := (iff #207 #759) +#761 := [rewrite]: #760 +#778 := [monotonicity #775 #761]: #777 +#784 := [trans #778 #782]: #783 +#787 := [quant-intro #784]: #786 +#938 := [monotonicity #787 #935]: #937 +#944 := [trans #938 #942]: #943 +#947 := [monotonicity #787 #944]: #946 +#950 := [monotonicity #947]: #949 +#956 := [trans #950 #954]: #955 +#959 := [monotonicity #956]: #958 +#771 := (iff #209 #770) +#768 := (iff #208 #765) +#762 := (implies #206 #759) +#766 := (iff #762 #765) +#767 := [rewrite]: #766 +#763 := (iff #208 #762) +#764 := [monotonicity #761]: #763 +#769 := [trans #764 #767]: #768 +#772 := [quant-intro #769]: #771 +#962 := [monotonicity #772 #959]: #961 +#968 := [trans #962 #966]: #967 +#757 := (iff #205 #756) +#754 := (iff #204 #751) +#748 := (implies #201 #745) +#752 := (iff #748 #751) +#753 := [rewrite]: #752 +#749 := (iff #204 #748) +#746 := (iff #203 #745) +#747 := [rewrite]: #746 +#750 := [monotonicity #747]: #749 +#755 := [trans #750 #753]: #754 +#758 := [quant-intro #755]: #757 +#971 := [monotonicity #758 #968]: #970 +#977 := [trans #971 #975]: #976 +#980 := [monotonicity #977]: #979 +#986 := [trans #980 #984]: #985 +#743 := (iff #192 #742) +#740 := (iff #191 #737) +#734 := (implies #481 #190) +#738 := (iff #734 #737) +#739 := [rewrite]: #738 +#735 := (iff #191 #734) +#482 := (iff #116 #481) +#479 := (iff #115 #478) +#480 := [rewrite]: #479 +#483 := [monotonicity #480]: #482 +#736 := [monotonicity #483]: #735 +#741 := [trans #736 #739]: #740 +#744 := [quant-intro #741]: #743 +#989 := [monotonicity #744 #986]: #988 +#995 := [trans #989 #993]: #994 +#998 := [monotonicity #995]: #997 +#1004 := [trans #998 #1002]: #1003 +#732 := (iff #187 #731) +#729 := (iff #186 #728) +#730 := [rewrite]: #729 +#733 := [monotonicity #730]: #732 +#1007 := [monotonicity #733 #1004]: #1006 +#1012 := [trans #1007 #1010]: #1011 +#561 := (iff #139 #560) +#558 := (iff #138 #555) +#552 := (and #481 #129) +#556 := (iff #552 #555) +#557 := [rewrite]: #556 +#553 := (iff #138 #552) +#554 := [monotonicity #483]: #553 +#559 := [trans #554 #557]: #558 +#562 := [quant-intro #559]: #561 +#1015 := [monotonicity #562 #1012]: #1014 +#1020 := [trans #1015 #1018]: #1019 +#1023 := [monotonicity #1020]: #1022 +#1027 := [trans #1023 #1025]: #1026 +#1029 := [monotonicity #1027]: #1028 +#1031 := [trans #1029 #1025]: #1030 +#726 := (iff #183 #712) +#717 := (implies true #712) +#720 := (iff #717 #712) +#721 := [rewrite]: #720 +#724 := (iff #183 #717) +#722 := (iff #182 #712) +#718 := (iff #182 #717) +#715 := (iff #181 #712) +#709 := (implies #563 #693) +#713 := (iff #709 #712) +#714 := [rewrite]: #713 +#710 := (iff #181 #709) +#707 := (iff #180 #693) +#698 := (implies true #693) +#701 := (iff #698 #693) +#702 := [rewrite]: #701 +#705 := (iff #180 #698) +#703 := (iff #179 #693) +#699 := (iff #179 #698) +#696 := (iff #178 #693) +#689 := (implies #566 #684) +#694 := (iff #689 #693) +#695 := [rewrite]: #694 +#690 := (iff #178 #689) +#687 := (iff #177 #684) +#680 := (implies #145 #675) +#685 := (iff #680 #684) +#686 := [rewrite]: #685 +#681 := (iff #177 #680) +#678 := (iff #176 #675) +#671 := (implies #569 #666) +#676 := (iff #671 #675) +#677 := [rewrite]: #676 +#672 := (iff #176 #671) +#669 := (iff #175 #666) +#662 := (implies #150 #652) +#667 := (iff #662 #666) +#668 := [rewrite]: #667 +#663 := (iff #175 #662) +#660 := (iff #174 #652) +#655 := (implies true #652) +#658 := (iff #655 #652) +#659 := [rewrite]: #658 +#656 := (iff #174 #655) +#653 := (iff #173 #652) +#650 := (iff #172 #647) +#643 := (implies #593 #640) +#648 := (iff #643 #647) +#649 := [rewrite]: #648 +#644 := (iff #172 #643) +#641 := (iff #171 #640) +#638 := (iff #170 #635) +#631 := (implies #611 #614) +#636 := (iff #631 #635) +#637 := [rewrite]: #636 +#632 := (iff #170 #631) +#629 := (iff #169 #614) +#624 := (and #614 true) +#627 := (iff #624 #614) +#628 := [rewrite]: #627 +#625 := (iff #169 #624) +#622 := (iff #168 true) +#617 := (implies #614 true) +#620 := (iff #617 true) +#621 := [rewrite]: #620 +#618 := (iff #168 #617) +#615 := (iff #167 #614) +#616 := [rewrite]: #615 +#619 := [monotonicity #616]: #618 +#623 := [trans #619 #621]: #622 +#626 := [monotonicity #616 #623]: #625 +#630 := [trans #626 #628]: #629 +#612 := (iff #165 #611) +#609 := (iff #164 #606) +#602 := (implies #596 #599) +#607 := (iff #602 #606) +#608 := [rewrite]: #607 +#603 := (iff #164 #602) +#600 := (iff #163 #599) +#573 := (= #156 #572) +#574 := [rewrite]: #573 +#601 := [monotonicity #574]: #600 +#597 := (iff #162 #596) +#598 := [rewrite]: #597 +#604 := [monotonicity #598 #601]: #603 +#610 := [trans #604 #608]: #609 +#613 := [quant-intro #610]: #612 +#633 := [monotonicity #613 #630]: #632 +#639 := [trans #633 #637]: #638 +#642 := [monotonicity #613 #639]: #641 +#594 := (iff #161 #593) +#591 := (iff #160 #588) +#584 := (implies #153 #581) +#589 := (iff #584 #588) +#590 := [rewrite]: #589 +#585 := (iff #160 #584) +#582 := (iff #159 #581) +#579 := (iff #158 #578) +#576 := (iff #157 #575) +#577 := [monotonicity #574]: #576 +#580 := [monotonicity #577]: #579 +#583 := [quant-intro #580]: #582 +#586 := [monotonicity #583]: #585 +#592 := [trans #586 #590]: #591 +#595 := [quant-intro #592]: #594 +#645 := [monotonicity #595 #642]: #644 +#651 := [trans #645 #649]: #650 +#654 := [monotonicity #595 #651]: #653 +#657 := [monotonicity #654]: #656 +#661 := [trans #657 #659]: #660 +#664 := [monotonicity #661]: #663 +#670 := [trans #664 #668]: #669 +#570 := (iff #147 #569) +#571 := [rewrite]: #570 +#673 := [monotonicity #571 #670]: #672 +#679 := [trans #673 #677]: #678 +#682 := [monotonicity #679]: #681 +#688 := [trans #682 #686]: #687 +#567 := (iff #142 #566) +#568 := [rewrite]: #567 +#691 := [monotonicity #568 #688]: #690 +#697 := [trans #691 #695]: #696 +#700 := [monotonicity #697]: #699 +#704 := [trans #700 #702]: #703 +#706 := [monotonicity #704]: #705 +#708 := [trans #706 #702]: #707 +#564 := (iff #140 #563) +#565 := [monotonicity #562]: #564 +#711 := [monotonicity #565 #708]: #710 +#716 := [trans #711 #714]: #715 +#719 := [monotonicity #716]: #718 +#723 := [trans #719 #721]: #722 +#725 := [monotonicity #723]: #724 +#727 := [trans #725 #721]: #726 +#1034 := [monotonicity #727 #1031]: #1033 +#1037 := [monotonicity #1034]: #1036 +#1041 := [trans #1037 #1039]: #1040 +#550 := (iff #137 #549) +#547 := (iff #136 #544) +#540 := (implies #130 #537) +#545 := (iff #540 #544) +#546 := [rewrite]: #545 +#541 := (iff #136 #540) +#538 := (iff #135 #537) +#535 := (iff #134 #534) +#532 := (iff #133 #531) +#529 := (iff #132 #528) +#511 := (= #125 #510) +#512 := [rewrite]: #511 +#530 := [monotonicity #512]: #529 +#533 := [monotonicity #480 #530]: #532 +#536 := [monotonicity #533]: #535 +#539 := [quant-intro #536]: #538 +#542 := [monotonicity #539]: #541 +#548 := [trans #542 #546]: #547 +#551 := [quant-intro #548]: #550 +#1044 := [monotonicity #551 #1041]: #1043 +#1050 := [trans #1044 #1048]: #1049 +#526 := (iff #128 #525) +#523 := (iff #127 #520) +#516 := (implies #505 #513) +#521 := (iff #516 #520) +#522 := [rewrite]: #521 +#517 := (iff #127 #516) +#514 := (iff #126 #513) +#515 := [monotonicity #512]: #514 +#508 := (iff #124 #505) +#502 := (and #478 #92) +#506 := (iff #502 #505) +#507 := [rewrite]: #506 +#503 := (iff #124 #502) +#504 := [monotonicity #480]: #503 +#509 := [trans #504 #507]: #508 +#518 := [monotonicity #509 #515]: #517 +#524 := [trans #518 #522]: #523 +#527 := [quant-intro #524]: #526 +#1053 := [monotonicity #527 #1050]: #1052 +#1059 := [trans #1053 #1057]: #1058 +#500 := (iff #123 #499) +#497 := (iff #122 #494) +#490 := (implies #487 #121) +#495 := (iff #490 #494) +#496 := [rewrite]: #495 +#491 := (iff #122 #490) +#488 := (iff #119 #487) +#485 := (iff #118 #484) +#486 := [rewrite]: #485 +#489 := [monotonicity #483 #486]: #488 +#492 := [monotonicity #489]: #491 +#498 := [trans #492 #496]: #497 +#501 := [quant-intro #498]: #500 +#1062 := [monotonicity #501 #1059]: #1061 +#1068 := [trans #1062 #1066]: #1067 +#1071 := [monotonicity #1068]: #1070 +#1077 := [trans #1071 #1075]: #1076 +#476 := (iff #109 #475) +#477 := [rewrite]: #476 +#1080 := [monotonicity #477 #1077]: #1079 +#1086 := [trans #1080 #1084]: #1085 +#1089 := [monotonicity #1086]: #1088 +#1093 := [trans #1089 #1091]: #1092 +#1095 := [monotonicity #1093]: #1094 +#1097 := [trans #1095 #1091]: #1096 +#473 := (iff #106 #472) +#470 := (iff #105 #469) +#471 := [rewrite]: #470 +#474 := [quant-intro #471]: #473 +#1100 := [monotonicity #474 #1097]: #1099 +#1106 := [trans #1100 #1104]: #1105 +#1109 := [monotonicity #474 #1106]: #1108 +#466 := (iff #97 #465) +#463 := (iff #96 #462) +#464 := [rewrite]: #463 +#467 := [quant-intro #464]: #466 +#1112 := [monotonicity #467 #1109]: #1111 +#1118 := [trans #1112 #1116]: #1117 +#1121 := [monotonicity #467 #1118]: #1120 +#459 := (iff #89 #458) +#456 := (iff #88 #455) +#457 := [rewrite]: #456 +#460 := [quant-intro #457]: #459 +#1124 := [monotonicity #460 #1121]: #1123 +#1130 := [trans #1124 #1128]: #1129 +#1133 := [monotonicity #460 #1130]: #1132 +#1136 := [monotonicity #1133]: #1135 +#1142 := [trans #1136 #1140]: #1141 +#1145 := [monotonicity #1142]: #1144 +#452 := (iff #81 #451) +#453 := [rewrite]: #452 +#1148 := [monotonicity #453 #1145]: #1147 +#1154 := [trans #1148 #1152]: #1153 +#1157 := [monotonicity #453 #1154]: #1156 +#1160 := [monotonicity #1157]: #1159 +#1164 := [trans #1160 #1162]: #1163 +#1167 := [monotonicity #1164]: #1166 +#1173 := [trans #1167 #1171]: #1172 +#449 := (iff #76 #448) +#446 := (iff #75 #443) +#440 := (implies #73 #437) +#444 := (iff #440 #443) +#445 := [rewrite]: #444 +#441 := (iff #75 #440) +#438 := (iff #74 #437) +#439 := [rewrite]: #438 +#442 := [monotonicity #439]: #441 +#447 := [trans #442 #445]: #446 +#450 := [quant-intro #447]: #449 +#1176 := [monotonicity #450 #1173]: #1175 +#1182 := [trans #1176 #1180]: #1181 +#435 := (iff #72 #434) +#432 := (iff #71 #429) +#426 := (implies #68 #423) +#430 := (iff #426 #429) +#431 := [rewrite]: #430 +#427 := (iff #71 #426) +#424 := (iff #70 #423) +#425 := [rewrite]: #424 +#428 := [monotonicity #425]: #427 +#433 := [trans #428 #431]: #432 +#436 := [quant-intro #433]: #435 +#1185 := [monotonicity #436 #1182]: #1184 +#1191 := [trans #1185 #1189]: #1190 +#1194 := [monotonicity #1191]: #1193 +#1198 := [trans #1194 #1196]: #1197 +#1200 := [monotonicity #1198]: #1199 +#1202 := [trans #1200 #1196]: #1201 +#1205 := [monotonicity #1202]: #1204 +#1834 := [trans #1205 #1832]: #1833 +#422 := [asserted]: #293 +#1835 := [mp #422 #1834]: #1830 +#1837 := [not-or-elim #1835]: #448 +#1902 := [mp~ #1837 #1863]: #448 +#4263 := [mp #1902 #4262]: #4258 +#4706 := (not #4258) +#5054 := (or #4706 #2574 #5019) +#1992 := (= ?x37!5 uf_11) +#5025 := (or #1992 #5019) +#5056 := (or #4706 #5025) +#5111 := (iff #5056 #5054) +#5048 := (or #2574 #5019) +#5057 := (or #4706 #5048) +#5087 := (iff #5057 #5054) +#5088 := [rewrite]: #5087 +#5059 := (iff #5056 #5057) +#5049 := (iff #5025 #5048) +#2575 := (iff #1992 #2574) +#2576 := [rewrite]: #2575 +#5053 := [monotonicity #2576]: #5049 +#5060 := [monotonicity #5053]: #5059 +#5120 := [trans #5060 #5088]: #5111 +#5047 := [quant-inst]: #5056 +#5121 := [mp #5047 #5120]: #5054 +#5195 := [unit-resolution #5121 #4263 #5192 #5194]: false +#5196 := [lemma #5195]: #4299 +#4585 := (or #4302 #4582) +#4588 := (not #4585) +#3123 := (or #78 #1237 #1248) +#4282 := (forall (vars (?x33 T2) (?x34 T2)) (:pat #4281) #3123) +#4287 := (not #4282) +#4591 := (or #4287 #4588) +#4594 := (not #4591) +decl ?x34!3 :: T2 +#1946 := ?x34!3 +#1953 := (uf_12 ?x34!3) +decl ?x33!4 :: T2 +#1947 := ?x33!4 +#1950 := (uf_12 ?x33!4) +#1951 := (* -1::int #1950) +#2561 := (+ #1951 #1953) +#1948 := (uf_1 ?x34!3 ?x33!4) +#1949 := (uf_10 #1948) +#2562 := (+ #1949 #2561) +#2565 := (>= #2562 0::int) +#1960 := (up_13 ?x34!3) +#3086 := (not #1960) +#1956 := (* -1::int #1949) +#1957 := (+ uf_9 #1956) +#1958 := (<= #1957 0::int) +#3101 := (or #1958 #3086 #2565) +#5147 := [hypothesis]: #1960 +#4251 := (forall (vars (?x26 T2)) (:pat #4250) #78) +#4254 := (iff #79 #4251) +#4252 := (iff #78 #78) +#4253 := [refl]: #4252 +#4255 := [quant-intro #4253]: #4254 +#1860 := (~ #79 #79) +#1897 := (~ #78 #78) +#1898 := [refl]: #1897 +#1861 := [nnf-pos #1898]: #1860 +#1836 := [not-or-elim #1835]: #79 +#1899 := [mp~ #1836 #1861]: #79 +#4256 := [mp #1899 #4255]: #4251 +#4844 := (not #4251) +#4845 := (or #4844 #3086) +#4846 := [quant-inst]: #4845 +#5148 := [unit-resolution #4846 #4256 #5147]: false +#5157 := [lemma #5148]: #3086 +#3862 := (or #3101 #1960) +#3866 := [def-axiom]: #3862 +#4921 := [unit-resolution #3866 #5157]: #3101 +#3106 := (not #3101) +#4597 := (or #3106 #4594) +#4600 := (not #4597) +#4272 := (pattern #77 #84) +#2527 := (not #84) +#3078 := (or #77 #2527 #1222) +#4273 := (forall (vars (?x29 T2) (?x30 T2)) (:pat #4272) #3078) +#4278 := (not #4273) +#4603 := (or #4278 #4600) +#4606 := (not #4603) +decl ?x30!1 :: T2 +#1921 := ?x30!1 +#1925 := (uf_12 ?x30!1) +#2542 := (* -1::int #1925) +decl ?x29!2 :: T2 +#1922 := ?x29!2 +#1923 := (uf_12 ?x29!2) +#2543 := (+ #1923 #2542) +#2544 := (<= #2543 0::int) +#1929 := (up_13 ?x30!1) +#1928 := (up_13 ?x29!2) +#1968 := (not #1928) +#2171 := (or #1968 #1929 #2544) +#2248 := (not #2171) +#4609 := (or #2248 #4606) +#4612 := (not #4609) +#4264 := (forall (vars (?x27 T2)) (:pat #4257) #1213) +#4269 := (not #4264) +#4615 := (or #4269 #4612) +#4618 := (not #4615) +decl ?x27!0 :: T2 +#1906 := ?x27!0 +#1907 := (uf_12 ?x27!0) +#1908 := (>= #1907 0::int) +#1909 := (not #1908) +#4621 := (or #1909 #4618) +#4624 := (not #4621) +#4627 := (or #1208 #4624) +#4630 := (not #4627) +#4637 := (forall (vars (?x24 T2)) (:pat #4257) #1807) +#4640 := (iff #1810 #4637) +#4638 := (iff #1807 #1807) +#4639 := [refl]: #4638 +#4641 := [quant-intro #4639]: #4640 +#2061 := (~ #1810 #1810) +#2287 := (~ #1807 #1807) +#2288 := [refl]: #2287 +#2062 := [nnf-pos #2288]: #2061 +#1840 := [not-or-elim #1835]: #1810 +#1967 := [mp~ #1840 #2062]: #1810 +#4642 := [mp #1967 #4641]: #4637 +#4660 := [hypothesis]: #1208 +#3800 := (not #4637) +#3794 := (or #3800 #81) +#3912 := (= uf_11 uf_11) +#3913 := (not #3912) +#3914 := (or #3913 #81) +#3784 := (or #3800 #3914) +#3782 := (iff #3784 #3794) +#3786 := (iff #3794 #3794) +#4657 := [rewrite]: #3786 +#3799 := (iff #3914 #81) +#3778 := (or false #81) +#3797 := (iff #3778 #81) +#3798 := [rewrite]: #3797 +#3776 := (iff #3914 #3778) +#3775 := (iff #3913 false) +#8605 := (not true) +#8608 := (iff #8605 false) +#8609 := [rewrite]: #8608 +#3783 := (iff #3913 #8605) +#3915 := (iff #3912 true) +#3787 := [rewrite]: #3915 +#3788 := [monotonicity #3787]: #3783 +#3777 := [trans #3788 #8609]: #3775 +#3779 := [monotonicity #3777]: #3776 +#3791 := [trans #3779 #3798]: #3799 +#3785 := [monotonicity #3791]: #3782 +#4658 := [trans #3785 #4657]: #3782 +#3781 := [quant-inst]: #3784 +#4659 := [mp #3781 #4658]: #3794 +#4661 := [unit-resolution #4659 #4660 #4642]: false +#4656 := [lemma #4661]: #81 +#4633 := (or #1208 #4630) +#3536 := (forall (vars (?x76 T2)) #3525) +#3543 := (not #3536) +#3521 := (forall (vars (?x71 T2) (?x72 T2)) #3516) +#3542 := (not #3521) +#3544 := (or #2367 #2934 #3542 #3543) +#3545 := (not #3544) +#3550 := (or #3499 #3545) +#3557 := (not #3550) +#3476 := (forall (vars (?x67 T2) (?x68 T2)) #3471) +#3556 := (not #3476) +#3558 := (or #3556 #3557) +#3559 := (not #3558) +#3564 := (or #3453 #3559) +#3570 := (not #3564) +#3571 := (or #1501 #3570) +#3572 := (not #3571) +#3577 := (or #2284 #3572) +#3583 := (not #3577) +#3584 := (or #1492 #3583) +#3585 := (not #3584) +#3590 := (or #1492 #3585) +#3596 := (not #3590) +#3597 := (or #939 #3596) +#3598 := (not #3597) +#3603 := (or #2877 #3598) +#3609 := (not #3603) +#3610 := (or #1487 #3609) +#3611 := (not #3610) +#3616 := (or #2863 #3611) +#3624 := (not #3616) +#3430 := (forall (vars (?x59 T2)) #3425) +#3623 := (not #3430) +#3412 := (forall (vars (?x60 T2)) #3409) +#3622 := (not #3412) +#3625 := (or #981 #728 #1652 #1656 #2218 #2221 #3622 #3623 #3624) +#3626 := (not #3625) +#3342 := (forall (vars (?x53 T2) (?x54 T2)) #3337) +#3348 := (not #3342) +#3349 := (or #167 #3348) +#3350 := (not #3349) +#3377 := (or #3350 #3374) +#3384 := (not #3377) +#3320 := (forall (vars (?x49 T2)) #3315) +#3383 := (not #3320) +#3385 := (or #3383 #3384) +#3386 := (not #3385) +#3283 := (forall (vars (?x50 T2)) #3272) +#3289 := (not #3283) +#3290 := (or #2097 #2712 #3289) +#3291 := (not #3290) +#3391 := (or #3291 #3386) +#3398 := (not #3391) +#3268 := (forall (vars (?x48 T2)) #3257) +#3397 := (not #3268) +#3399 := (or #683 #665 #692 #674 #3397 #3398) +#3400 := (not #3399) +#3631 := (or #3400 #3626) +#3641 := (not #3631) +#3254 := (forall (vars (?x46 T2)) #3249) +#3640 := (not #3254) +#3226 := (forall (vars (?x37 T2)) #3221) +#3639 := (not #3226) +#3198 := (forall (vars (?x42 T2) (?x43 T2)) #3193) +#3638 := (not #3198) +#3175 := (forall (vars (?x44 T2) (?x45 T2)) #3170) +#3637 := (not #3175) +#3642 := (or #1749 #1744 #3637 #3638 #3639 #3640 #3641) +#3643 := (not #3642) +#3143 := (forall (vars (?x38 T2)) #3132) +#3149 := (not #3143) +#3150 := (or #1990 #2574 #3149) +#3151 := (not #3150) +#3648 := (or #3151 #3643) +#3655 := (not #3648) +#3128 := (forall (vars (?x33 T2) (?x34 T2)) #3123) +#3654 := (not #3128) +#3656 := (or #3654 #3655) +#3657 := (not #3656) +#3662 := (or #3106 #3657) +#3669 := (not #3662) +#3083 := (forall (vars (?x29 T2) (?x30 T2)) #3078) +#3668 := (not #3083) +#3670 := (or #3668 #3669) +#3671 := (not #3670) +#3676 := (or #2248 #3671) +#3682 := (not #3676) +#3683 := (or #1217 #3682) +#3684 := (not #3683) +#3689 := (or #1909 #3684) +#3695 := (not #3689) +#3696 := (or #1208 #3695) +#3697 := (not #3696) +#3702 := (or #1208 #3697) +#4634 := (iff #3702 #4633) +#4631 := (iff #3697 #4630) +#4628 := (iff #3696 #4627) +#4625 := (iff #3695 #4624) +#4622 := (iff #3689 #4621) +#4619 := (iff #3684 #4618) +#4616 := (iff #3683 #4615) +#4613 := (iff #3682 #4612) +#4610 := (iff #3676 #4609) +#4607 := (iff #3671 #4606) +#4604 := (iff #3670 #4603) +#4601 := (iff #3669 #4600) +#4598 := (iff #3662 #4597) +#4595 := (iff #3657 #4594) +#4592 := (iff #3656 #4591) +#4589 := (iff #3655 #4588) +#4586 := (iff #3648 #4585) +#4583 := (iff #3643 #4582) +#4580 := (iff #3642 #4579) +#4577 := (iff #3641 #4576) +#4574 := (iff #3631 #4573) +#4571 := (iff #3626 #4570) +#4568 := (iff #3625 #4567) +#4565 := (iff #3624 #4564) +#4562 := (iff #3616 #4561) +#4559 := (iff #3611 #4558) +#4556 := (iff #3610 #4555) +#4553 := (iff #3609 #4552) +#4550 := (iff #3603 #4549) +#4547 := (iff #3598 #4546) +#4544 := (iff #3597 #4543) +#4541 := (iff #3596 #4540) +#4538 := (iff #3590 #4537) +#4535 := (iff #3585 #4534) +#4532 := (iff #3584 #4531) +#4529 := (iff #3583 #4528) +#4526 := (iff #3577 #4525) +#4523 := (iff #3572 #4522) +#4520 := (iff #3571 #4519) +#4517 := (iff #3570 #4516) +#4514 := (iff #3564 #4513) +#4511 := (iff #3559 #4510) +#4508 := (iff #3558 #4507) +#4505 := (iff #3557 #4504) +#4502 := (iff #3550 #4501) +#4499 := (iff #3545 #4498) +#4496 := (iff #3544 #4495) +#4493 := (iff #3543 #4492) +#4490 := (iff #3536 #4487) +#4488 := (iff #3525 #3525) +#4489 := [refl]: #4488 +#4491 := [quant-intro #4489]: #4490 +#4494 := [monotonicity #4491]: #4493 +#4484 := (iff #3542 #4483) +#4481 := (iff #3521 #4478) +#4479 := (iff #3516 #3516) +#4480 := [refl]: #4479 +#4482 := [quant-intro #4480]: #4481 +#4485 := [monotonicity #4482]: #4484 +#4497 := [monotonicity #4485 #4494]: #4496 +#4500 := [monotonicity #4497]: #4499 +#4503 := [monotonicity #4500]: #4502 +#4506 := [monotonicity #4503]: #4505 +#4476 := (iff #3556 #4475) +#4473 := (iff #3476 #4470) +#4471 := (iff #3471 #3471) +#4472 := [refl]: #4471 +#4474 := [quant-intro #4472]: #4473 +#4477 := [monotonicity #4474]: #4476 +#4509 := [monotonicity #4477 #4506]: #4508 +#4512 := [monotonicity #4509]: #4511 +#4515 := [monotonicity #4512]: #4514 +#4518 := [monotonicity #4515]: #4517 +#4467 := (iff #1501 #4466) +#4464 := (iff #1498 #4461) +#4462 := (iff #1495 #1495) +#4463 := [refl]: #4462 +#4465 := [quant-intro #4463]: #4464 +#4468 := [monotonicity #4465]: #4467 +#4521 := [monotonicity #4468 #4518]: #4520 +#4524 := [monotonicity #4521]: #4523 +#4527 := [monotonicity #4524]: #4526 +#4530 := [monotonicity #4527]: #4529 +#4533 := [monotonicity #4530]: #4532 +#4536 := [monotonicity #4533]: #4535 +#4539 := [monotonicity #4536]: #4538 +#4542 := [monotonicity #4539]: #4541 +#4459 := (iff #939 #4458) +#4456 := (iff #785 #4453) +#4454 := (iff #780 #780) +#4455 := [refl]: #4454 +#4457 := [quant-intro #4455]: #4456 +#4460 := [monotonicity #4457]: #4459 +#4545 := [monotonicity #4460 #4542]: #4544 +#4548 := [monotonicity #4545]: #4547 +#4551 := [monotonicity #4548]: #4550 +#4554 := [monotonicity #4551]: #4553 +#4450 := (iff #1487 #4449) +#4447 := (iff #1484 #4444) +#4445 := (iff #1479 #1479) +#4446 := [refl]: #4445 +#4448 := [quant-intro #4446]: #4447 +#4451 := [monotonicity #4448]: #4450 +#4557 := [monotonicity #4451 #4554]: #4556 +#4560 := [monotonicity #4557]: #4559 +#4563 := [monotonicity #4560]: #4562 +#4566 := [monotonicity #4563]: #4565 +#4442 := (iff #3623 #4441) +#4439 := (iff #3430 #4436) +#4437 := (iff #3425 #3425) +#4438 := [refl]: #4437 +#4440 := [quant-intro #4438]: #4439 +#4443 := [monotonicity #4440]: #4442 +#4434 := (iff #3622 #4433) +#4431 := (iff #3412 #4428) +#4429 := (iff #3409 #3409) +#4430 := [refl]: #4429 +#4432 := [quant-intro #4430]: #4431 +#4435 := [monotonicity #4432]: #4434 +#4424 := (iff #1652 #4423) +#4421 := (iff #1649 #4418) +#4419 := (iff #1644 #1644) +#4420 := [refl]: #4419 +#4422 := [quant-intro #4420]: #4421 +#4425 := [monotonicity #4422]: #4424 +#4569 := [monotonicity #4425 #4435 #4443 #4566]: #4568 +#4572 := [monotonicity #4569]: #4571 +#4416 := (iff #3400 #4415) +#4413 := (iff #3399 #4412) +#4410 := (iff #3398 #4409) +#4407 := (iff #3391 #4406) +#4404 := (iff #3386 #4403) +#4401 := (iff #3385 #4400) +#4398 := (iff #3384 #4397) +#4395 := (iff #3377 #4394) +#4392 := (iff #3350 #4391) +#4389 := (iff #3349 #4388) +#4386 := (iff #3348 #4385) +#4383 := (iff #3342 #4380) +#4381 := (iff #3337 #3337) +#4382 := [refl]: #4381 +#4384 := [quant-intro #4382]: #4383 +#4387 := [monotonicity #4384]: #4386 +#4390 := [monotonicity #4387]: #4389 +#4393 := [monotonicity #4390]: #4392 +#4396 := [monotonicity #4393]: #4395 +#4399 := [monotonicity #4396]: #4398 +#4378 := (iff #3383 #4377) +#4375 := (iff #3320 #4372) +#4373 := (iff #3315 #3315) +#4374 := [refl]: #4373 +#4376 := [quant-intro #4374]: #4375 +#4379 := [monotonicity #4376]: #4378 +#4402 := [monotonicity #4379 #4399]: #4401 +#4405 := [monotonicity #4402]: #4404 +#4370 := (iff #3291 #4369) +#4367 := (iff #3290 #4366) +#4364 := (iff #3289 #4363) +#4361 := (iff #3283 #4358) +#4359 := (iff #3272 #3272) +#4360 := [refl]: #4359 +#4362 := [quant-intro #4360]: #4361 +#4365 := [monotonicity #4362]: #4364 +#4368 := [monotonicity #4365]: #4367 +#4371 := [monotonicity #4368]: #4370 +#4408 := [monotonicity #4371 #4405]: #4407 +#4411 := [monotonicity #4408]: #4410 +#4354 := (iff #3397 #4353) +#4351 := (iff #3268 #4348) +#4349 := (iff #3257 #3257) +#4350 := [refl]: #4349 +#4352 := [quant-intro #4350]: #4351 +#4355 := [monotonicity #4352]: #4354 +#4414 := [monotonicity #4355 #4411]: #4413 +#4417 := [monotonicity #4414]: #4416 +#4575 := [monotonicity #4417 #4572]: #4574 +#4578 := [monotonicity #4575]: #4577 +#4345 := (iff #3640 #4344) +#4342 := (iff #3254 #4339) +#4340 := (iff #3249 #3249) +#4341 := [refl]: #4340 +#4343 := [quant-intro #4341]: #4342 +#4346 := [monotonicity #4343]: #4345 +#4337 := (iff #3639 #4336) +#4334 := (iff #3226 #4331) +#4332 := (iff #3221 #3221) +#4333 := [refl]: #4332 +#4335 := [quant-intro #4333]: #4334 +#4338 := [monotonicity #4335]: #4337 +#4329 := (iff #3638 #4328) +#4326 := (iff #3198 #4323) +#4324 := (iff #3193 #3193) +#4325 := [refl]: #4324 +#4327 := [quant-intro #4325]: #4326 +#4330 := [monotonicity #4327]: #4329 +#4320 := (iff #3637 #4319) +#4317 := (iff #3175 #4314) +#4315 := (iff #3170 #3170) +#4316 := [refl]: #4315 +#4318 := [quant-intro #4316]: #4317 +#4321 := [monotonicity #4318]: #4320 +#4312 := (iff #1744 #4311) +#4309 := (iff #1741 #4306) +#4307 := (iff #1738 #1738) +#4308 := [refl]: #4307 +#4310 := [quant-intro #4308]: #4309 +#4313 := [monotonicity #4310]: #4312 +#4581 := [monotonicity #4313 #4321 #4330 #4338 #4346 #4578]: #4580 +#4584 := [monotonicity #4581]: #4583 +#4303 := (iff #3151 #4302) +#4300 := (iff #3150 #4299) +#4297 := (iff #3149 #4296) +#4294 := (iff #3143 #4291) +#4292 := (iff #3132 #3132) +#4293 := [refl]: #4292 +#4295 := [quant-intro #4293]: #4294 +#4298 := [monotonicity #4295]: #4297 +#4301 := [monotonicity #4298]: #4300 +#4304 := [monotonicity #4301]: #4303 +#4587 := [monotonicity #4304 #4584]: #4586 +#4590 := [monotonicity #4587]: #4589 +#4288 := (iff #3654 #4287) +#4285 := (iff #3128 #4282) +#4283 := (iff #3123 #3123) +#4284 := [refl]: #4283 +#4286 := [quant-intro #4284]: #4285 +#4289 := [monotonicity #4286]: #4288 +#4593 := [monotonicity #4289 #4590]: #4592 +#4596 := [monotonicity #4593]: #4595 +#4599 := [monotonicity #4596]: #4598 +#4602 := [monotonicity #4599]: #4601 +#4279 := (iff #3668 #4278) +#4276 := (iff #3083 #4273) +#4274 := (iff #3078 #3078) +#4275 := [refl]: #4274 +#4277 := [quant-intro #4275]: #4276 +#4280 := [monotonicity #4277]: #4279 +#4605 := [monotonicity #4280 #4602]: #4604 +#4608 := [monotonicity #4605]: #4607 +#4611 := [monotonicity #4608]: #4610 +#4614 := [monotonicity #4611]: #4613 +#4270 := (iff #1217 #4269) +#4267 := (iff #1214 #4264) +#4265 := (iff #1213 #1213) +#4266 := [refl]: #4265 +#4268 := [quant-intro #4266]: #4267 +#4271 := [monotonicity #4268]: #4270 +#4617 := [monotonicity #4271 #4614]: #4616 +#4620 := [monotonicity #4617]: #4619 +#4623 := [monotonicity #4620]: #4622 +#4626 := [monotonicity #4623]: #4625 +#4629 := [monotonicity #4626]: #4628 +#4632 := [monotonicity #4629]: #4631 +#4635 := [monotonicity #4632]: #4634 +#2363 := (not #2362) +#2965 := (and #773 #2363 #2962) +#2968 := (not #2965) +#2971 := (forall (vars (?x76 T2)) #2968) +#2937 := (not #2934) +#2368 := (not #2367) +#2980 := (and #1541 #2368 #2937 #2971) +#2907 := (and #2336 #2338) +#2910 := (not #2907) +#2928 := (or #2910 #2923) +#2931 := (not #2928) +#2985 := (or #2931 #2980) +#2988 := (and #1517 #2985) +#2307 := (not #2306) +#2882 := (and #2304 #2307) +#2885 := (not #2882) +#2901 := (or #2885 #2896) +#2904 := (not #2901) +#2991 := (or #2904 #2988) +#2994 := (and #1498 #2991) +#2997 := (or #2284 #2994) +#3000 := (and #217 #2997) +#3003 := (or #1492 #3000) +#3006 := (and #785 #3003) +#3009 := (or #2877 #3006) +#3012 := (and #1484 #3009) +#3015 := (or #2863 #3012) +#2222 := (not #2221) +#2219 := (not #2218) +#3021 := (and #195 #731 #1473 #1631 #1649 #1657 #2219 #2222 #3015) +#2164 := (not #2163) +#2160 := (not #2159) +#2800 := (and #2160 #2164) +#2803 := (not #2800) +#2820 := (or #2803 #2815) +#2823 := (not #2820) +#2173 := (not #167) +#2183 := (and #2173 #1393) +#2829 := (or #2183 #2823) +#2773 := (not #2768) +#2791 := (and #2773 #2786) +#2794 := (or #1353 #2791) +#2797 := (forall (vars (?x49 T2)) #2794) +#2834 := (and #2797 #2829) +#2093 := (not #2092) +#2743 := (and #2093 #2740) +#2746 := (not #2743) +#2749 := (forall (vars (?x50 T2)) #2746) +#2715 := (not #2712) +#2098 := (not #2097) +#2755 := (and #2098 #2715 #2749) +#2837 := (or #2755 #2834) +#2199 := (not #1421) +#2202 := (forall (vars (?x48 T2)) #2199) +#2843 := (and #145 #150 #566 #569 #2202 #2837) +#3026 := (or #2843 #3021) +#2692 := (not #2687) +#2695 := (and #2057 #2675 #2692) +#2698 := (or #1308 #2695) +#2701 := (forall (vars (?x46 T2)) #2698) +#2637 := (not #2632) +#2655 := (and #2030 #2637 #2650) +#2658 := (or #1270 #2655) +#2661 := (forall (vars (?x37 T2)) #2658) +#3032 := (and #109 #1716 #1732 #1741 #2661 #2701 #3026) +#1986 := (not #1985) +#2605 := (and #77 #1986 #2602) +#2608 := (not #2605) +#2611 := (forall (vars (?x38 T2)) #2608) +#2617 := (and #1991 #2577 #2611) +#3037 := (or #2617 #3032) +#3040 := (and #1255 #3037) +#1959 := (not #1958) +#2555 := (and #1959 #1960) +#2558 := (not #2555) +#2568 := (or #2558 #2565) +#2571 := (not #2568) +#3043 := (or #2571 #3040) +#3046 := (and #1229 #3043) +#1930 := (not #1929) +#2530 := (and #1928 #1930) +#2533 := (not #2530) +#2549 := (or #2533 #2544) +#2552 := (not #2549) +#3049 := (or #2552 #3046) +#3052 := (and #1214 #3049) +#3055 := (or #1909 #3052) +#3058 := (and #81 #3055) +#3061 := (or #1208 #3058) +#3703 := (iff #3061 #3702) +#3700 := (iff #3058 #3697) +#3692 := (and #81 #3689) +#3698 := (iff #3692 #3697) +#3699 := [rewrite]: #3698 +#3693 := (iff #3058 #3692) +#3690 := (iff #3055 #3689) +#3687 := (iff #3052 #3684) +#3679 := (and #1214 #3676) +#3685 := (iff #3679 #3684) +#3686 := [rewrite]: #3685 +#3680 := (iff #3052 #3679) +#3677 := (iff #3049 #3676) +#3674 := (iff #3046 #3671) +#3665 := (and #3083 #3662) +#3672 := (iff #3665 #3671) +#3673 := [rewrite]: #3672 +#3666 := (iff #3046 #3665) +#3663 := (iff #3043 #3662) +#3660 := (iff #3040 #3657) +#3651 := (and #3128 #3648) +#3658 := (iff #3651 #3657) +#3659 := [rewrite]: #3658 +#3652 := (iff #3040 #3651) +#3649 := (iff #3037 #3648) +#3646 := (iff #3032 #3643) +#3634 := (and #109 #3175 #3198 #1741 #3226 #3254 #3631) +#3644 := (iff #3634 #3643) +#3645 := [rewrite]: #3644 +#3635 := (iff #3032 #3634) +#3632 := (iff #3026 #3631) +#3629 := (iff #3021 #3626) +#3619 := (and #195 #731 #3412 #3430 #1649 #1657 #2219 #2222 #3616) +#3627 := (iff #3619 #3626) +#3628 := [rewrite]: #3627 +#3620 := (iff #3021 #3619) +#3617 := (iff #3015 #3616) +#3614 := (iff #3012 #3611) +#3606 := (and #1484 #3603) +#3612 := (iff #3606 #3611) +#3613 := [rewrite]: #3612 +#3607 := (iff #3012 #3606) +#3604 := (iff #3009 #3603) +#3601 := (iff #3006 #3598) +#3593 := (and #785 #3590) +#3599 := (iff #3593 #3598) +#3600 := [rewrite]: #3599 +#3594 := (iff #3006 #3593) +#3591 := (iff #3003 #3590) +#3588 := (iff #3000 #3585) +#3580 := (and #217 #3577) +#3586 := (iff #3580 #3585) +#3587 := [rewrite]: #3586 +#3581 := (iff #3000 #3580) +#3578 := (iff #2997 #3577) +#3575 := (iff #2994 #3572) +#3567 := (and #1498 #3564) +#3573 := (iff #3567 #3572) +#3574 := [rewrite]: #3573 +#3568 := (iff #2994 #3567) +#3565 := (iff #2991 #3564) +#3562 := (iff #2988 #3559) +#3553 := (and #3476 #3550) +#3560 := (iff #3553 #3559) +#3561 := [rewrite]: #3560 +#3554 := (iff #2988 #3553) +#3551 := (iff #2985 #3550) +#3548 := (iff #2980 #3545) +#3539 := (and #3521 #2368 #2937 #3536) +#3546 := (iff #3539 #3545) +#3547 := [rewrite]: #3546 +#3540 := (iff #2980 #3539) +#3537 := (iff #2971 #3536) +#3534 := (iff #2968 #3525) +#3526 := (not #3525) +#3529 := (not #3526) +#3532 := (iff #3529 #3525) +#3533 := [rewrite]: #3532 +#3530 := (iff #2968 #3529) +#3527 := (iff #2965 #3526) +#3528 := [rewrite]: #3527 +#3531 := [monotonicity #3528]: #3530 +#3535 := [trans #3531 #3533]: #3534 +#3538 := [quant-intro #3535]: #3537 +#3522 := (iff #1541 #3521) +#3519 := (iff #1538 #3516) +#3502 := (or #779 #1237) +#3513 := (or #3502 #1534) +#3517 := (iff #3513 #3516) +#3518 := [rewrite]: #3517 +#3514 := (iff #1538 #3513) +#3511 := (iff #1531 #3502) +#3503 := (not #3502) +#3506 := (not #3503) +#3509 := (iff #3506 #3502) +#3510 := [rewrite]: #3509 +#3507 := (iff #1531 #3506) +#3504 := (iff #1526 #3503) +#3505 := [rewrite]: #3504 +#3508 := [monotonicity #3505]: #3507 +#3512 := [trans #3508 #3510]: #3511 +#3515 := [monotonicity #3512]: #3514 +#3520 := [trans #3515 #3518]: #3519 +#3523 := [quant-intro #3520]: #3522 +#3541 := [monotonicity #3523 #3538]: #3540 +#3549 := [trans #3541 #3547]: #3548 +#3500 := (iff #2931 #3499) +#3497 := (iff #2928 #3494) +#3480 := (or #2335 #3479) +#3491 := (or #3480 #2923) +#3495 := (iff #3491 #3494) +#3496 := [rewrite]: #3495 +#3492 := (iff #2928 #3491) +#3489 := (iff #2910 #3480) +#3481 := (not #3480) +#3484 := (not #3481) +#3487 := (iff #3484 #3480) +#3488 := [rewrite]: #3487 +#3485 := (iff #2910 #3484) +#3482 := (iff #2907 #3481) +#3483 := [rewrite]: #3482 +#3486 := [monotonicity #3483]: #3485 +#3490 := [trans #3486 #3488]: #3489 +#3493 := [monotonicity #3490]: #3492 +#3498 := [trans #3493 #3496]: #3497 +#3501 := [monotonicity #3498]: #3500 +#3552 := [monotonicity #3501 #3549]: #3551 +#3477 := (iff #1517 #3476) +#3474 := (iff #1512 #3471) +#3457 := (or #773 #3456) +#3468 := (or #3457 #1504) +#3472 := (iff #3468 #3471) +#3473 := [rewrite]: #3472 +#3469 := (iff #1512 #3468) +#3466 := (iff #802 #3457) +#3458 := (not #3457) +#3461 := (not #3458) +#3464 := (iff #3461 #3457) +#3465 := [rewrite]: #3464 +#3462 := (iff #802 #3461) +#3459 := (iff #796 #3458) +#3460 := [rewrite]: #3459 +#3463 := [monotonicity #3460]: #3462 +#3467 := [trans #3463 #3465]: #3466 +#3470 := [monotonicity #3467]: #3469 +#3475 := [trans #3470 #3473]: #3474 +#3478 := [quant-intro #3475]: #3477 +#3555 := [monotonicity #3478 #3552]: #3554 +#3563 := [trans #3555 #3561]: #3562 +#3454 := (iff #2904 #3453) +#3451 := (iff #2901 #3448) +#3434 := (or #3433 #2306) +#3445 := (or #3434 #2896) +#3449 := (iff #3445 #3448) +#3450 := [rewrite]: #3449 +#3446 := (iff #2901 #3445) +#3443 := (iff #2885 #3434) +#3435 := (not #3434) +#3438 := (not #3435) +#3441 := (iff #3438 #3434) +#3442 := [rewrite]: #3441 +#3439 := (iff #2885 #3438) +#3436 := (iff #2882 #3435) +#3437 := [rewrite]: #3436 +#3440 := [monotonicity #3437]: #3439 +#3444 := [trans #3440 #3442]: #3443 +#3447 := [monotonicity #3444]: #3446 +#3452 := [trans #3447 #3450]: #3451 +#3455 := [monotonicity #3452]: #3454 +#3566 := [monotonicity #3455 #3563]: #3565 +#3569 := [monotonicity #3566]: #3568 +#3576 := [trans #3569 #3574]: #3575 +#3579 := [monotonicity #3576]: #3578 +#3582 := [monotonicity #3579]: #3581 +#3589 := [trans #3582 #3587]: #3588 +#3592 := [monotonicity #3589]: #3591 +#3595 := [monotonicity #3592]: #3594 +#3602 := [trans #3595 #3600]: #3601 +#3605 := [monotonicity #3602]: #3604 +#3608 := [monotonicity #3605]: #3607 +#3615 := [trans #3608 #3613]: #3614 +#3618 := [monotonicity #3615]: #3617 +#3431 := (iff #1631 #3430) +#3428 := (iff #1628 #3425) +#3422 := (or #3405 #1625) +#3426 := (iff #3422 #3425) +#3427 := [rewrite]: #3426 +#3423 := (iff #1628 #3422) +#3420 := (iff #1620 #3405) +#3415 := (not #3406) +#3418 := (iff #3415 #3405) +#3419 := [rewrite]: #3418 +#3416 := (iff #1620 #3415) +#3407 := (iff #1462 #3406) +#3408 := [rewrite]: #3407 +#3417 := [monotonicity #3408]: #3416 +#3421 := [trans #3417 #3419]: #3420 +#3424 := [monotonicity #3421]: #3423 +#3429 := [trans #3424 #3427]: #3428 +#3432 := [quant-intro #3429]: #3431 +#3413 := (iff #1473 #3412) +#3410 := (iff #1468 #3409) +#3411 := [monotonicity #3408]: #3410 +#3414 := [quant-intro #3411]: #3413 +#3621 := [monotonicity #3414 #3432 #3618]: #3620 +#3630 := [trans #3621 #3628]: #3629 +#3403 := (iff #2843 #3400) +#3394 := (and #145 #150 #566 #569 #3268 #3391) +#3401 := (iff #3394 #3400) +#3402 := [rewrite]: #3401 +#3395 := (iff #2843 #3394) +#3392 := (iff #2837 #3391) +#3389 := (iff #2834 #3386) +#3380 := (and #3320 #3377) +#3387 := (iff #3380 #3386) +#3388 := [rewrite]: #3387 +#3381 := (iff #2834 #3380) +#3378 := (iff #2829 #3377) +#3375 := (iff #2823 #3374) +#3372 := (iff #2820 #3369) +#3355 := (or #2159 #2163) +#3366 := (or #3355 #2815) +#3370 := (iff #3366 #3369) +#3371 := [rewrite]: #3370 +#3367 := (iff #2820 #3366) +#3364 := (iff #2803 #3355) +#3356 := (not #3355) +#3359 := (not #3356) +#3362 := (iff #3359 #3355) +#3363 := [rewrite]: #3362 +#3360 := (iff #2803 #3359) +#3357 := (iff #2800 #3356) +#3358 := [rewrite]: #3357 +#3361 := [monotonicity #3358]: #3360 +#3365 := [trans #3361 #3363]: #3364 +#3368 := [monotonicity #3365]: #3367 +#3373 := [trans #3368 #3371]: #3372 +#3376 := [monotonicity #3373]: #3375 +#3353 := (iff #2183 #3350) +#3345 := (and #2173 #3342) +#3351 := (iff #3345 #3350) +#3352 := [rewrite]: #3351 +#3346 := (iff #2183 #3345) +#3343 := (iff #1393 #3342) +#3340 := (iff #1390 #3337) +#3323 := (or #1237 #1346) +#3334 := (or #3323 #1387) +#3338 := (iff #3334 #3337) +#3339 := [rewrite]: #3338 +#3335 := (iff #1390 #3334) +#3332 := (iff #1384 #3323) +#3324 := (not #3323) +#3327 := (not #3324) +#3330 := (iff #3327 #3323) +#3331 := [rewrite]: #3330 +#3328 := (iff #1384 #3327) +#3325 := (iff #1381 #3324) +#3326 := [rewrite]: #3325 +#3329 := [monotonicity #3326]: #3328 +#3333 := [trans #3329 #3331]: #3332 +#3336 := [monotonicity #3333]: #3335 +#3341 := [trans #3336 #3339]: #3340 +#3344 := [quant-intro #3341]: #3343 +#3347 := [monotonicity #3344]: #3346 +#3354 := [trans #3347 #3352]: #3353 +#3379 := [monotonicity #3354 #3376]: #3378 +#3321 := (iff #2797 #3320) +#3318 := (iff #2794 #3315) +#3296 := (or #68 #1346) +#3312 := (or #3296 #3309) +#3316 := (iff #3312 #3315) +#3317 := [rewrite]: #3316 +#3313 := (iff #2794 #3312) +#3310 := (iff #2791 #3309) +#3311 := [rewrite]: #3310 +#3305 := (iff #1353 #3296) +#3297 := (not #3296) +#3300 := (not #3297) +#3303 := (iff #3300 #3296) +#3304 := [rewrite]: #3303 +#3301 := (iff #1353 #3300) +#3298 := (iff #1350 #3297) +#3299 := [rewrite]: #3298 +#3302 := [monotonicity #3299]: #3301 +#3306 := [trans #3302 #3304]: #3305 +#3314 := [monotonicity #3306 #3311]: #3313 +#3319 := [trans #3314 #3317]: #3318 +#3322 := [quant-intro #3319]: #3321 +#3382 := [monotonicity #3322 #3379]: #3381 +#3390 := [trans #3382 #3388]: #3389 +#3294 := (iff #2755 #3291) +#3286 := (and #2098 #2715 #3283) +#3292 := (iff #3286 #3291) +#3293 := [rewrite]: #3292 +#3287 := (iff #2755 #3286) +#3284 := (iff #2749 #3283) +#3281 := (iff #2746 #3272) +#3273 := (not #3272) +#3276 := (not #3273) +#3279 := (iff #3276 #3272) +#3280 := [rewrite]: #3279 +#3277 := (iff #2746 #3276) +#3274 := (iff #2743 #3273) +#3275 := [rewrite]: #3274 +#3278 := [monotonicity #3275]: #3277 +#3282 := [trans #3278 #3280]: #3281 +#3285 := [quant-intro #3282]: #3284 +#3288 := [monotonicity #3285]: #3287 +#3295 := [trans #3288 #3293]: #3294 +#3393 := [monotonicity #3295 #3390]: #3392 +#3269 := (iff #2202 #3268) +#3266 := (iff #2199 #3257) +#3258 := (not #3257) +#3261 := (not #3258) +#3264 := (iff #3261 #3257) +#3265 := [rewrite]: #3264 +#3262 := (iff #2199 #3261) +#3259 := (iff #1421 #3258) +#3260 := [rewrite]: #3259 +#3263 := [monotonicity #3260]: #3262 +#3267 := [trans #3263 #3265]: #3266 +#3270 := [quant-intro #3267]: #3269 +#3396 := [monotonicity #3270 #3393]: #3395 +#3404 := [trans #3396 #3402]: #3403 +#3633 := [monotonicity #3404 #3630]: #3632 +#3255 := (iff #2701 #3254) +#3252 := (iff #2698 #3249) +#3229 := (or #68 #1301) +#3246 := (or #3229 #3243) +#3250 := (iff #3246 #3249) +#3251 := [rewrite]: #3250 +#3247 := (iff #2698 #3246) +#3244 := (iff #2695 #3243) +#3245 := [rewrite]: #3244 +#3238 := (iff #1308 #3229) +#3230 := (not #3229) +#3233 := (not #3230) +#3236 := (iff #3233 #3229) +#3237 := [rewrite]: #3236 +#3234 := (iff #1308 #3233) +#3231 := (iff #1305 #3230) +#3232 := [rewrite]: #3231 +#3235 := [monotonicity #3232]: #3234 +#3239 := [trans #3235 #3237]: #3238 +#3248 := [monotonicity #3239 #3245]: #3247 +#3253 := [trans #3248 #3251]: #3252 +#3256 := [quant-intro #3253]: #3255 +#3227 := (iff #2661 #3226) +#3224 := (iff #2658 #3221) +#3201 := (or #68 #1263) +#3218 := (or #3201 #3215) +#3222 := (iff #3218 #3221) +#3223 := [rewrite]: #3222 +#3219 := (iff #2658 #3218) +#3216 := (iff #2655 #3215) +#3217 := [rewrite]: #3216 +#3210 := (iff #1270 #3201) +#3202 := (not #3201) +#3205 := (not #3202) +#3208 := (iff #3205 #3201) +#3209 := [rewrite]: #3208 +#3206 := (iff #1270 #3205) +#3203 := (iff #1267 #3202) +#3204 := [rewrite]: #3203 +#3207 := [monotonicity #3204]: #3206 +#3211 := [trans #3207 #3209]: #3210 +#3220 := [monotonicity #3211 #3217]: #3219 +#3225 := [trans #3220 #3223]: #3224 +#3228 := [quant-intro #3225]: #3227 +#3199 := (iff #1732 #3198) +#3196 := (iff #1727 #3193) +#3179 := (or #478 #3178) +#3190 := (or #3179 #1317) +#3194 := (iff #3190 #3193) +#3195 := [rewrite]: #3194 +#3191 := (iff #1727 #3190) +#3188 := (iff #493 #3179) +#3180 := (not #3179) +#3183 := (not #3180) +#3186 := (iff #3183 #3179) +#3187 := [rewrite]: #3186 +#3184 := (iff #493 #3183) +#3181 := (iff #487 #3180) +#3182 := [rewrite]: #3181 +#3185 := [monotonicity #3182]: #3184 +#3189 := [trans #3185 #3187]: #3188 +#3192 := [monotonicity #3189]: #3191 +#3197 := [trans #3192 #3195]: #3196 +#3200 := [quant-intro #3197]: #3199 +#3176 := (iff #1716 #3175) +#3173 := (iff #1713 #3170) +#3156 := (or #481 #1237) +#3167 := (or #3156 #1710) +#3171 := (iff #3167 #3170) +#3172 := [rewrite]: #3171 +#3168 := (iff #1713 #3167) +#3165 := (iff #1707 #3156) +#3157 := (not #3156) +#3160 := (not #3157) +#3163 := (iff #3160 #3156) +#3164 := [rewrite]: #3163 +#3161 := (iff #1707 #3160) +#3158 := (iff #1702 #3157) +#3159 := [rewrite]: #3158 +#3162 := [monotonicity #3159]: #3161 +#3166 := [trans #3162 #3164]: #3165 +#3169 := [monotonicity #3166]: #3168 +#3174 := [trans #3169 #3172]: #3173 +#3177 := [quant-intro #3174]: #3176 +#3636 := [monotonicity #3177 #3200 #3228 #3256 #3633]: #3635 +#3647 := [trans #3636 #3645]: #3646 +#3154 := (iff #2617 #3151) +#3146 := (and #1991 #2577 #3143) +#3152 := (iff #3146 #3151) +#3153 := [rewrite]: #3152 +#3147 := (iff #2617 #3146) +#3144 := (iff #2611 #3143) +#3141 := (iff #2608 #3132) +#3133 := (not #3132) +#3136 := (not #3133) +#3139 := (iff #3136 #3132) +#3140 := [rewrite]: #3139 +#3137 := (iff #2608 #3136) +#3134 := (iff #2605 #3133) +#3135 := [rewrite]: #3134 +#3138 := [monotonicity #3135]: #3137 +#3142 := [trans #3138 #3140]: #3141 +#3145 := [quant-intro #3142]: #3144 +#3148 := [monotonicity #3145]: #3147 +#3155 := [trans #3148 #3153]: #3154 +#3650 := [monotonicity #3155 #3647]: #3649 +#3129 := (iff #1255 #3128) +#3126 := (iff #1252 #3123) +#3109 := (or #78 #1237) +#3120 := (or #3109 #1248) +#3124 := (iff #3120 #3123) +#3125 := [rewrite]: #3124 +#3121 := (iff #1252 #3120) +#3118 := (iff #1244 #3109) +#3110 := (not #3109) +#3113 := (not #3110) +#3116 := (iff #3113 #3109) +#3117 := [rewrite]: #3116 +#3114 := (iff #1244 #3113) +#3111 := (iff #1241 #3110) +#3112 := [rewrite]: #3111 +#3115 := [monotonicity #3112]: #3114 +#3119 := [trans #3115 #3117]: #3118 +#3122 := [monotonicity #3119]: #3121 +#3127 := [trans #3122 #3125]: #3126 +#3130 := [quant-intro #3127]: #3129 +#3653 := [monotonicity #3130 #3650]: #3652 +#3661 := [trans #3653 #3659]: #3660 +#3107 := (iff #2571 #3106) +#3104 := (iff #2568 #3101) +#3087 := (or #1958 #3086) +#3098 := (or #3087 #2565) +#3102 := (iff #3098 #3101) +#3103 := [rewrite]: #3102 +#3099 := (iff #2568 #3098) +#3096 := (iff #2558 #3087) +#3088 := (not #3087) +#3091 := (not #3088) +#3094 := (iff #3091 #3087) +#3095 := [rewrite]: #3094 +#3092 := (iff #2558 #3091) +#3089 := (iff #2555 #3088) +#3090 := [rewrite]: #3089 +#3093 := [monotonicity #3090]: #3092 +#3097 := [trans #3093 #3095]: #3096 +#3100 := [monotonicity #3097]: #3099 +#3105 := [trans #3100 #3103]: #3104 +#3108 := [monotonicity #3105]: #3107 +#3664 := [monotonicity #3108 #3661]: #3663 +#3084 := (iff #1229 #3083) +#3081 := (iff #1226 #3078) +#3064 := (or #77 #2527) +#3075 := (or #3064 #1222) +#3079 := (iff #3075 #3078) +#3080 := [rewrite]: #3079 +#3076 := (iff #1226 #3075) +#3073 := (iff #454 #3064) +#3065 := (not #3064) +#3068 := (not #3065) +#3071 := (iff #3068 #3064) +#3072 := [rewrite]: #3071 +#3069 := (iff #454 #3068) +#3066 := (iff #85 #3065) +#3067 := [rewrite]: #3066 +#3070 := [monotonicity #3067]: #3069 +#3074 := [trans #3070 #3072]: #3073 +#3077 := [monotonicity #3074]: #3076 +#3082 := [trans #3077 #3080]: #3081 +#3085 := [quant-intro #3082]: #3084 +#3667 := [monotonicity #3085 #3664]: #3666 +#3675 := [trans #3667 #3673]: #3674 +#2345 := (iff #2552 #2248) +#2137 := (iff #2549 #2171) +#1937 := (or #1968 #1929) +#2268 := (or #1937 #2544) +#2172 := (iff #2268 #2171) +#2136 := [rewrite]: #2172 +#2226 := (iff #2549 #2268) +#2035 := (iff #2533 #1937) +#1868 := (not #1937) +#2314 := (not #1868) +#1913 := (iff #2314 #1937) +#2034 := [rewrite]: #1913 +#2315 := (iff #2533 #2314) +#1869 := (iff #2530 #1868) +#1938 := [rewrite]: #1869 +#1912 := [monotonicity #1938]: #2315 +#2267 := [trans #1912 #2034]: #2035 +#2227 := [monotonicity #2267]: #2226 +#2247 := [trans #2227 #2136]: #2137 +#2346 := [monotonicity #2247]: #2345 +#3678 := [monotonicity #2346 #3675]: #3677 +#3681 := [monotonicity #3678]: #3680 +#3688 := [trans #3681 #3686]: #3687 +#3691 := [monotonicity #3688]: #3690 +#3694 := [monotonicity #3691]: #3693 +#3701 := [trans #3694 #3699]: #3700 +#3704 := [monotonicity #3701]: #3703 +#2360 := (+ #2359 #2357) +#2361 := (= #2360 0::int) +#2364 := (and #773 #2363 #2361) +#2381 := (not #2364) +#2384 := (forall (vars (?x76 T2)) #2381) +#2369 := (= ?x75!20 uf_11) +#2370 := (not #2369) +#2371 := (and #2370 #2368) +#2372 := (not #2371) +#2378 := (not #2372) +#2388 := (and #2378 #2384) +#2393 := (and #1541 #2388) +#2326 := (* -1::int #2325) +#2328 := (+ #2327 #2326) +#2331 := (+ #2330 #2328) +#2332 := (>= #2331 0::int) +#2339 := (and #2338 #2336) +#2340 := (not #2339) +#2341 := (or #2340 #2332) +#2342 := (not #2341) +#2397 := (or #2342 #2393) +#2401 := (and #1517 #2397) +#2299 := (* -1::int #2298) +#2301 := (+ #2300 #2299) +#2302 := (>= #2301 0::int) +#2308 := (and #2307 #2304) +#2309 := (not #2308) +#2310 := (or #2309 #2302) +#2311 := (not #2310) +#2405 := (or #2311 #2401) +#2409 := (and #1498 #2405) +#2413 := (or #2284 #2409) +#2278 := (not #1492) +#2417 := (and #2278 #2413) +#2421 := (or #1492 #2417) +#2425 := (and #785 #2421) +#2262 := (= #2261 #2260) +#2263 := (or #2262 #2259) +#2264 := (not #2263) +#2429 := (or #2264 #2425) +#2433 := (and #1484 #2429) +#2240 := (* -1::int #2239) +#2242 := (+ #2241 #2240) +#2243 := (>= #2242 0::int) +#2244 := (not #2243) +#2437 := (or #2244 #2433) +#2223 := (and #2222 #2219) +#2209 := (not #981) +#2457 := (and #2209 #731 #2223 #1473 #2437 #1631 #1649 #1657) +#2150 := (* -1::int #2149) +#2152 := (+ #2151 #2150) +#2155 := (+ #2154 #2152) +#2156 := (>= #2155 0::int) +#2165 := (and #2164 #2160) +#2166 := (not #2165) +#2167 := (or #2166 #2156) +#2168 := (not #2167) +#2187 := (or #2168 #2183) +#2126 := (+ #2125 #1344) +#2129 := (+ #2128 #2126) +#2130 := (= #2129 0::int) +#2131 := (>= #2126 0::int) +#2132 := (not #2131) +#2133 := (and #2132 #2130) +#2138 := (or #1353 #2133) +#2141 := (forall (vars (?x49 T2)) #2138) +#2191 := (and #2141 #2187) +#2090 := (+ #2089 #2087) +#2091 := (= #2090 0::int) +#2094 := (and #2093 #2091) +#2110 := (not #2094) +#2113 := (forall (vars (?x50 T2)) #2110) +#2099 := (= ?x49!8 uf_11) +#2100 := (not #2099) +#2101 := (and #2100 #2098) +#2102 := (not #2101) +#2107 := (not #2102) +#2117 := (and #2107 #2113) +#2195 := (or #2117 #2191) +#2081 := (not #674) +#2078 := (not #692) +#2075 := (not #665) +#2072 := (not #683) +#2205 := (and #2072 #2075 #2078 #2081 #2195 #2202) +#2461 := (or #2205 #2457) +#2049 := (+ #2048 #1299) +#2050 := (>= #2049 0::int) +#2051 := (not #2050) +#2054 := (+ #2053 #2049) +#2055 := (= #2054 0::int) +#2058 := (and #2057 #2055 #2051) +#2063 := (or #1308 #2058) +#2066 := (forall (vars (?x46 T2)) #2063) +#2023 := (+ #1261 #2022) +#2025 := (+ #2024 #2023) +#2026 := (= #2025 0::int) +#2027 := (+ #2024 #1261) +#2028 := (>= #2027 0::int) +#2029 := (not #2028) +#2031 := (and #2030 #2029 #2026) +#2036 := (or #1270 #2031) +#2039 := (forall (vars (?x37 T2)) #2036) +#2015 := (not #1749) +#2486 := (and #2015 #2039 #2066 #2461 #1716 #1732 #1741) +#1981 := (+ #1980 #1978) +#1982 := (+ #69 #1981) +#1983 := (= #1982 0::int) +#1987 := (and #77 #1986 #1983) +#2003 := (not #1987) +#2006 := (forall (vars (?x38 T2)) #2003) +#1993 := (not #1992) +#1994 := (and #1993 #1991) +#1995 := (not #1994) +#2000 := (not #1995) +#2010 := (and #2000 #2006) +#2490 := (or #2010 #2486) +#2494 := (and #1255 #2490) +#1952 := (+ #1951 #1949) +#1954 := (+ #1953 #1952) +#1955 := (>= #1954 0::int) +#1961 := (and #1960 #1959) +#1962 := (not #1961) +#1963 := (or #1962 #1955) +#1964 := (not #1963) +#2498 := (or #1964 #2494) +#2502 := (and #1229 #2498) +#1924 := (* -1::int #1923) +#1926 := (+ #1925 #1924) +#1927 := (>= #1926 0::int) +#1931 := (and #1930 #1928) +#1932 := (not #1931) +#1933 := (or #1932 #1927) +#1934 := (not #1933) +#2506 := (or #1934 #2502) +#2510 := (and #1214 #2506) +#2514 := (or #1909 #2510) +#1864 := (not #1208) +#2518 := (and #1864 #2514) +#2522 := (or #1208 #2518) +#3062 := (iff #2522 #3061) +#3059 := (iff #2518 #3058) +#3056 := (iff #2514 #3055) +#3053 := (iff #2510 #3052) +#3050 := (iff #2506 #3049) +#3047 := (iff #2502 #3046) +#3044 := (iff #2498 #3043) +#3041 := (iff #2494 #3040) +#3038 := (iff #2490 #3037) +#3035 := (iff #2486 #3032) +#3029 := (and #109 #2661 #2701 #3026 #1716 #1732 #1741) +#3033 := (iff #3029 #3032) +#3034 := [rewrite]: #3033 +#3030 := (iff #2486 #3029) +#3027 := (iff #2461 #3026) +#3024 := (iff #2457 #3021) +#3018 := (and #195 #731 #2223 #1473 #3015 #1631 #1649 #1657) +#3022 := (iff #3018 #3021) +#3023 := [rewrite]: #3022 +#3019 := (iff #2457 #3018) +#3016 := (iff #2437 #3015) +#3013 := (iff #2433 #3012) +#3010 := (iff #2429 #3009) +#3007 := (iff #2425 #3006) +#3004 := (iff #2421 #3003) +#3001 := (iff #2417 #3000) +#2998 := (iff #2413 #2997) +#2995 := (iff #2409 #2994) +#2992 := (iff #2405 #2991) +#2989 := (iff #2401 #2988) +#2986 := (iff #2397 #2985) +#2983 := (iff #2393 #2980) +#2943 := (and #2368 #2937) +#2974 := (and #2943 #2971) +#2977 := (and #1541 #2974) +#2981 := (iff #2977 #2980) +#2982 := [rewrite]: #2981 +#2978 := (iff #2393 #2977) +#2975 := (iff #2388 #2974) +#2972 := (iff #2384 #2971) +#2969 := (iff #2381 #2968) +#2966 := (iff #2364 #2965) +#2963 := (iff #2361 #2962) +#2960 := (= #2360 #2959) +#2961 := [rewrite]: #2960 +#2964 := [monotonicity #2961]: #2963 +#2967 := [monotonicity #2964]: #2966 +#2970 := [monotonicity #2967]: #2969 +#2973 := [quant-intro #2970]: #2972 +#2956 := (iff #2378 #2943) +#2948 := (not #2943) +#2951 := (not #2948) +#2954 := (iff #2951 #2943) +#2955 := [rewrite]: #2954 +#2952 := (iff #2378 #2951) +#2949 := (iff #2372 #2948) +#2946 := (iff #2371 #2943) +#2940 := (and #2937 #2368) +#2944 := (iff #2940 #2943) +#2945 := [rewrite]: #2944 +#2941 := (iff #2371 #2940) +#2938 := (iff #2370 #2937) +#2935 := (iff #2369 #2934) +#2936 := [rewrite]: #2935 +#2939 := [monotonicity #2936]: #2938 +#2942 := [monotonicity #2939]: #2941 +#2947 := [trans #2942 #2945]: #2946 +#2950 := [monotonicity #2947]: #2949 +#2953 := [monotonicity #2950]: #2952 +#2957 := [trans #2953 #2955]: #2956 +#2976 := [monotonicity #2957 #2973]: #2975 +#2979 := [monotonicity #2976]: #2978 +#2984 := [trans #2979 #2982]: #2983 +#2932 := (iff #2342 #2931) +#2929 := (iff #2341 #2928) +#2926 := (iff #2332 #2923) +#2913 := (+ #2327 #2330) +#2914 := (+ #2326 #2913) +#2917 := (>= #2914 0::int) +#2924 := (iff #2917 #2923) +#2925 := [rewrite]: #2924 +#2918 := (iff #2332 #2917) +#2915 := (= #2331 #2914) +#2916 := [rewrite]: #2915 +#2919 := [monotonicity #2916]: #2918 +#2927 := [trans #2919 #2925]: #2926 +#2911 := (iff #2340 #2910) +#2908 := (iff #2339 #2907) +#2909 := [rewrite]: #2908 +#2912 := [monotonicity #2909]: #2911 +#2930 := [monotonicity #2912 #2927]: #2929 +#2933 := [monotonicity #2930]: #2932 +#2987 := [monotonicity #2933 #2984]: #2986 +#2990 := [monotonicity #2987]: #2989 +#2905 := (iff #2311 #2904) +#2902 := (iff #2310 #2901) +#2899 := (iff #2302 #2896) +#2888 := (+ #2299 #2300) +#2891 := (>= #2888 0::int) +#2897 := (iff #2891 #2896) +#2898 := [rewrite]: #2897 +#2892 := (iff #2302 #2891) +#2889 := (= #2301 #2888) +#2890 := [rewrite]: #2889 +#2893 := [monotonicity #2890]: #2892 +#2900 := [trans #2893 #2898]: #2899 +#2886 := (iff #2309 #2885) +#2883 := (iff #2308 #2882) +#2884 := [rewrite]: #2883 +#2887 := [monotonicity #2884]: #2886 +#2903 := [monotonicity #2887 #2900]: #2902 +#2906 := [monotonicity #2903]: #2905 +#2993 := [monotonicity #2906 #2990]: #2992 +#2996 := [monotonicity #2993]: #2995 +#2999 := [monotonicity #2996]: #2998 +#2880 := (iff #2278 #217) +#2881 := [rewrite]: #2880 +#3002 := [monotonicity #2881 #2999]: #3001 +#3005 := [monotonicity #3002]: #3004 +#3008 := [monotonicity #3005]: #3007 +#2878 := (iff #2264 #2877) +#2875 := (iff #2263 #2872) +#2869 := (or #2866 #2259) +#2873 := (iff #2869 #2872) +#2874 := [rewrite]: #2873 +#2870 := (iff #2263 #2869) +#2867 := (iff #2262 #2866) +#2868 := [rewrite]: #2867 +#2871 := [monotonicity #2868]: #2870 +#2876 := [trans #2871 #2874]: #2875 +#2879 := [monotonicity #2876]: #2878 +#3011 := [monotonicity #2879 #3008]: #3010 +#3014 := [monotonicity #3011]: #3013 +#2864 := (iff #2244 #2863) +#2861 := (iff #2243 #2858) +#2850 := (+ #2240 #2241) +#2853 := (>= #2850 0::int) +#2859 := (iff #2853 #2858) +#2860 := [rewrite]: #2859 +#2854 := (iff #2243 #2853) +#2851 := (= #2242 #2850) +#2852 := [rewrite]: #2851 +#2855 := [monotonicity #2852]: #2854 +#2862 := [trans #2855 #2860]: #2861 +#2865 := [monotonicity #2862]: #2864 +#3017 := [monotonicity #2865 #3014]: #3016 +#2848 := (iff #2209 #195) +#2849 := [rewrite]: #2848 +#3020 := [monotonicity #2849 #3017]: #3019 +#3025 := [trans #3020 #3023]: #3024 +#2846 := (iff #2205 #2843) +#2840 := (and #145 #150 #566 #569 #2837 #2202) +#2844 := (iff #2840 #2843) +#2845 := [rewrite]: #2844 +#2841 := (iff #2205 #2840) +#2838 := (iff #2195 #2837) +#2835 := (iff #2191 #2834) +#2832 := (iff #2187 #2829) +#2826 := (or #2823 #2183) +#2830 := (iff #2826 #2829) +#2831 := [rewrite]: #2830 +#2827 := (iff #2187 #2826) +#2824 := (iff #2168 #2823) +#2821 := (iff #2167 #2820) +#2818 := (iff #2156 #2815) +#2806 := (+ #2151 #2154) +#2807 := (+ #2150 #2806) +#2810 := (>= #2807 0::int) +#2816 := (iff #2810 #2815) +#2817 := [rewrite]: #2816 +#2811 := (iff #2156 #2810) +#2808 := (= #2155 #2807) +#2809 := [rewrite]: #2808 +#2812 := [monotonicity #2809]: #2811 +#2819 := [trans #2812 #2817]: #2818 +#2804 := (iff #2166 #2803) +#2801 := (iff #2165 #2800) +#2802 := [rewrite]: #2801 +#2805 := [monotonicity #2802]: #2804 +#2822 := [monotonicity #2805 #2819]: #2821 +#2825 := [monotonicity #2822]: #2824 +#2828 := [monotonicity #2825]: #2827 +#2833 := [trans #2828 #2831]: #2832 +#2798 := (iff #2141 #2797) +#2795 := (iff #2138 #2794) +#2792 := (iff #2133 #2791) +#2789 := (iff #2130 #2786) +#2776 := (+ #2125 #2128) +#2777 := (+ #1344 #2776) +#2780 := (= #2777 0::int) +#2787 := (iff #2780 #2786) +#2788 := [rewrite]: #2787 +#2781 := (iff #2130 #2780) +#2778 := (= #2129 #2777) +#2779 := [rewrite]: #2778 +#2782 := [monotonicity #2779]: #2781 +#2790 := [trans #2782 #2788]: #2789 +#2774 := (iff #2132 #2773) +#2771 := (iff #2131 #2768) +#2760 := (+ #1344 #2125) +#2763 := (>= #2760 0::int) +#2769 := (iff #2763 #2768) +#2770 := [rewrite]: #2769 +#2764 := (iff #2131 #2763) +#2761 := (= #2126 #2760) +#2762 := [rewrite]: #2761 +#2765 := [monotonicity #2762]: #2764 +#2772 := [trans #2765 #2770]: #2771 +#2775 := [monotonicity #2772]: #2774 +#2793 := [monotonicity #2775 #2790]: #2792 +#2796 := [monotonicity #2793]: #2795 +#2799 := [quant-intro #2796]: #2798 +#2836 := [monotonicity #2799 #2833]: #2835 +#2758 := (iff #2117 #2755) +#2721 := (and #2098 #2715) +#2752 := (and #2721 #2749) +#2756 := (iff #2752 #2755) +#2757 := [rewrite]: #2756 +#2753 := (iff #2117 #2752) +#2750 := (iff #2113 #2749) +#2747 := (iff #2110 #2746) +#2744 := (iff #2094 #2743) +#2741 := (iff #2091 #2740) +#2738 := (= #2090 #2737) +#2739 := [rewrite]: #2738 +#2742 := [monotonicity #2739]: #2741 +#2745 := [monotonicity #2742]: #2744 +#2748 := [monotonicity #2745]: #2747 +#2751 := [quant-intro #2748]: #2750 +#2734 := (iff #2107 #2721) +#2726 := (not #2721) +#2729 := (not #2726) +#2732 := (iff #2729 #2721) +#2733 := [rewrite]: #2732 +#2730 := (iff #2107 #2729) +#2727 := (iff #2102 #2726) +#2724 := (iff #2101 #2721) +#2718 := (and #2715 #2098) +#2722 := (iff #2718 #2721) +#2723 := [rewrite]: #2722 +#2719 := (iff #2101 #2718) +#2716 := (iff #2100 #2715) +#2713 := (iff #2099 #2712) +#2714 := [rewrite]: #2713 +#2717 := [monotonicity #2714]: #2716 +#2720 := [monotonicity #2717]: #2719 +#2725 := [trans #2720 #2723]: #2724 +#2728 := [monotonicity #2725]: #2727 +#2731 := [monotonicity #2728]: #2730 +#2735 := [trans #2731 #2733]: #2734 +#2754 := [monotonicity #2735 #2751]: #2753 +#2759 := [trans #2754 #2757]: #2758 +#2839 := [monotonicity #2759 #2836]: #2838 +#2710 := (iff #2081 #569) +#2711 := [rewrite]: #2710 +#2708 := (iff #2078 #566) +#2709 := [rewrite]: #2708 +#2706 := (iff #2075 #150) +#2707 := [rewrite]: #2706 +#2704 := (iff #2072 #145) +#2705 := [rewrite]: #2704 +#2842 := [monotonicity #2705 #2707 #2709 #2711 #2839]: #2841 +#2847 := [trans #2842 #2845]: #2846 +#3028 := [monotonicity #2847 #3025]: #3027 +#2702 := (iff #2066 #2701) +#2699 := (iff #2063 #2698) +#2696 := (iff #2058 #2695) +#2693 := (iff #2051 #2692) +#2690 := (iff #2050 #2687) +#2680 := (+ #1299 #2048) +#2683 := (>= #2680 0::int) +#2688 := (iff #2683 #2687) +#2689 := [rewrite]: #2688 +#2684 := (iff #2050 #2683) +#2681 := (= #2049 #2680) +#2682 := [rewrite]: #2681 +#2685 := [monotonicity #2682]: #2684 +#2691 := [trans #2685 #2689]: #2690 +#2694 := [monotonicity #2691]: #2693 +#2678 := (iff #2055 #2675) +#2664 := (+ #2048 #2053) +#2665 := (+ #1299 #2664) +#2668 := (= #2665 0::int) +#2676 := (iff #2668 #2675) +#2677 := [rewrite]: #2676 +#2669 := (iff #2055 #2668) +#2666 := (= #2054 #2665) +#2667 := [rewrite]: #2666 +#2670 := [monotonicity #2667]: #2669 +#2679 := [trans #2670 #2677]: #2678 +#2697 := [monotonicity #2679 #2694]: #2696 +#2700 := [monotonicity #2697]: #2699 +#2703 := [quant-intro #2700]: #2702 +#2662 := (iff #2039 #2661) +#2659 := (iff #2036 #2658) +#2656 := (iff #2031 #2655) +#2653 := (iff #2026 #2650) +#2640 := (+ #2022 #2024) +#2641 := (+ #1261 #2640) +#2644 := (= #2641 0::int) +#2651 := (iff #2644 #2650) +#2652 := [rewrite]: #2651 +#2645 := (iff #2026 #2644) +#2642 := (= #2025 #2641) +#2643 := [rewrite]: #2642 +#2646 := [monotonicity #2643]: #2645 +#2654 := [trans #2646 #2652]: #2653 +#2638 := (iff #2029 #2637) +#2635 := (iff #2028 #2632) +#2624 := (+ #1261 #2024) +#2627 := (>= #2624 0::int) +#2633 := (iff #2627 #2632) +#2634 := [rewrite]: #2633 +#2628 := (iff #2028 #2627) +#2625 := (= #2027 #2624) +#2626 := [rewrite]: #2625 +#2629 := [monotonicity #2626]: #2628 +#2636 := [trans #2629 #2634]: #2635 +#2639 := [monotonicity #2636]: #2638 +#2657 := [monotonicity #2639 #2654]: #2656 +#2660 := [monotonicity #2657]: #2659 +#2663 := [quant-intro #2660]: #2662 +#2622 := (iff #2015 #109) +#2623 := [rewrite]: #2622 +#3031 := [monotonicity #2623 #2663 #2703 #3028]: #3030 +#3036 := [trans #3031 #3034]: #3035 +#2620 := (iff #2010 #2617) +#2583 := (and #1991 #2577) +#2614 := (and #2583 #2611) +#2618 := (iff #2614 #2617) +#2619 := [rewrite]: #2618 +#2615 := (iff #2010 #2614) +#2612 := (iff #2006 #2611) +#2609 := (iff #2003 #2608) +#2606 := (iff #1987 #2605) +#2603 := (iff #1983 #2602) +#2600 := (= #1982 #2599) +#2601 := [rewrite]: #2600 +#2604 := [monotonicity #2601]: #2603 +#2607 := [monotonicity #2604]: #2606 +#2610 := [monotonicity #2607]: #2609 +#2613 := [quant-intro #2610]: #2612 +#2596 := (iff #2000 #2583) +#2588 := (not #2583) +#2591 := (not #2588) +#2594 := (iff #2591 #2583) +#2595 := [rewrite]: #2594 +#2592 := (iff #2000 #2591) +#2589 := (iff #1995 #2588) +#2586 := (iff #1994 #2583) +#2580 := (and #2577 #1991) +#2584 := (iff #2580 #2583) +#2585 := [rewrite]: #2584 +#2581 := (iff #1994 #2580) +#2578 := (iff #1993 #2577) +#2579 := [monotonicity #2576]: #2578 +#2582 := [monotonicity #2579]: #2581 +#2587 := [trans #2582 #2585]: #2586 +#2590 := [monotonicity #2587]: #2589 +#2593 := [monotonicity #2590]: #2592 +#2597 := [trans #2593 #2595]: #2596 +#2616 := [monotonicity #2597 #2613]: #2615 +#2621 := [trans #2616 #2619]: #2620 +#3039 := [monotonicity #2621 #3036]: #3038 +#3042 := [monotonicity #3039]: #3041 +#2572 := (iff #1964 #2571) +#2569 := (iff #1963 #2568) +#2566 := (iff #1955 #2565) +#2563 := (= #1954 #2562) +#2564 := [rewrite]: #2563 +#2567 := [monotonicity #2564]: #2566 +#2559 := (iff #1962 #2558) +#2556 := (iff #1961 #2555) +#2557 := [rewrite]: #2556 +#2560 := [monotonicity #2557]: #2559 +#2570 := [monotonicity #2560 #2567]: #2569 +#2573 := [monotonicity #2570]: #2572 +#3045 := [monotonicity #2573 #3042]: #3044 +#3048 := [monotonicity #3045]: #3047 +#2553 := (iff #1934 #2552) +#2550 := (iff #1933 #2549) +#2547 := (iff #1927 #2544) +#2536 := (+ #1924 #1925) +#2539 := (>= #2536 0::int) +#2545 := (iff #2539 #2544) +#2546 := [rewrite]: #2545 +#2540 := (iff #1927 #2539) +#2537 := (= #1926 #2536) +#2538 := [rewrite]: #2537 +#2541 := [monotonicity #2538]: #2540 +#2548 := [trans #2541 #2546]: #2547 +#2534 := (iff #1932 #2533) +#2531 := (iff #1931 #2530) +#2532 := [rewrite]: #2531 +#2535 := [monotonicity #2532]: #2534 +#2551 := [monotonicity #2535 #2548]: #2550 +#2554 := [monotonicity #2551]: #2553 +#3051 := [monotonicity #2554 #3048]: #3050 +#3054 := [monotonicity #3051]: #3053 +#3057 := [monotonicity #3054]: #3056 +#2528 := (iff #1864 #81) +#2529 := [rewrite]: #2528 +#3060 := [monotonicity #2529 #3057]: #3059 +#3063 := [monotonicity #3060]: #3062 +#1838 := (not #1802) +#2523 := (~ #1838 #2522) +#2519 := (not #1799) +#2520 := (~ #2519 #2518) +#2515 := (not #1796) +#2516 := (~ #2515 #2514) +#2511 := (not #1793) +#2512 := (~ #2511 #2510) +#2507 := (not #1790) +#2508 := (~ #2507 #2506) +#2503 := (not #1787) +#2504 := (~ #2503 #2502) +#2499 := (not #1784) +#2500 := (~ #2499 #2498) +#2495 := (not #1781) +#2496 := (~ #2495 #2494) +#2491 := (not #1778) +#2492 := (~ #2491 #2490) +#2487 := (not #1773) +#2488 := (~ #2487 #2486) +#2483 := (not #1744) +#2484 := (~ #2483 #1741) +#2481 := (~ #1741 #1741) +#2479 := (~ #1738 #1738) +#2480 := [refl]: #2479 +#2482 := [nnf-pos #2480]: #2481 +#2485 := [nnf-neg #2482]: #2484 +#2476 := (not #1735) +#2477 := (~ #2476 #1732) +#2474 := (~ #1732 #1732) +#2472 := (~ #1727 #1727) +#2473 := [refl]: #2472 +#2475 := [nnf-pos #2473]: #2474 +#2478 := [nnf-neg #2475]: #2477 +#2469 := (not #1719) +#2470 := (~ #2469 #1716) +#2467 := (~ #1716 #1716) +#2465 := (~ #1713 #1713) +#2466 := [refl]: #2465 +#2468 := [nnf-pos #2466]: #2467 +#2471 := [nnf-neg #2468]: #2470 +#2462 := (not #1696) +#2463 := (~ #2462 #2461) +#2458 := (not #1691) +#2459 := (~ #2458 #2457) +#2455 := (~ #1657 #1657) +#2456 := [refl]: #2455 +#2452 := (not #1652) +#2453 := (~ #2452 #1649) +#2450 := (~ #1649 #1649) +#2448 := (~ #1644 #1644) +#2449 := [refl]: #2448 +#2451 := [nnf-pos #2449]: #2450 +#2454 := [nnf-neg #2451]: #2453 +#2445 := (not #1634) +#2446 := (~ #2445 #1631) +#2443 := (~ #1631 #1631) +#2441 := (~ #1628 #1628) +#2442 := [refl]: #2441 +#2444 := [nnf-pos #2442]: #2443 +#2447 := [nnf-neg #2444]: #2446 +#2438 := (not #1617) +#2439 := (~ #2438 #2437) +#2434 := (not #1614) +#2435 := (~ #2434 #2433) +#2430 := (not #1611) +#2431 := (~ #2430 #2429) +#2426 := (not #1608) +#2427 := (~ #2426 #2425) +#2422 := (not #1605) +#2423 := (~ #2422 #2421) +#2418 := (not #1602) +#2419 := (~ #2418 #2417) +#2414 := (not #1599) +#2415 := (~ #2414 #2413) +#2410 := (not #1596) +#2411 := (~ #2410 #2409) +#2406 := (not #1593) +#2407 := (~ #2406 #2405) +#2402 := (not #1590) +#2403 := (~ #2402 #2401) +#2398 := (not #1587) +#2399 := (~ #2398 #2397) +#2394 := (not #1584) +#2395 := (~ #2394 #2393) +#2375 := (not #1581) +#2391 := (~ #2375 #2388) +#2365 := (exists (vars (?x76 T2)) #2364) +#2373 := (or #2372 #2365) +#2374 := (not #2373) +#2389 := (~ #2374 #2388) +#2385 := (not #2365) +#2386 := (~ #2385 #2384) +#2382 := (~ #2381 #2381) +#2383 := [refl]: #2382 +#2387 := [nnf-neg #2383]: #2386 +#2379 := (~ #2378 #2378) +#2380 := [refl]: #2379 +#2390 := [nnf-neg #2380 #2387]: #2389 +#2376 := (~ #2375 #2374) +#2377 := [sk]: #2376 +#2392 := [trans #2377 #2390]: #2391 +#2351 := (not #1544) +#2352 := (~ #2351 #1541) +#2349 := (~ #1541 #1541) +#2347 := (~ #1538 #1538) +#2348 := [refl]: #2347 +#2350 := [nnf-pos #2348]: #2349 +#2353 := [nnf-neg #2350]: #2352 +#2396 := [nnf-neg #2353 #2392]: #2395 +#2343 := (~ #1544 #2342) +#2344 := [sk]: #2343 +#2400 := [nnf-neg #2344 #2396]: #2399 +#2320 := (not #1520) +#2321 := (~ #2320 #1517) +#2318 := (~ #1517 #1517) +#2316 := (~ #1512 #1512) +#2317 := [refl]: #2316 +#2319 := [nnf-pos #2317]: #2318 +#2322 := [nnf-neg #2319]: #2321 +#2404 := [nnf-neg #2322 #2400]: #2403 +#2312 := (~ #1520 #2311) +#2313 := [sk]: #2312 +#2408 := [nnf-neg #2313 #2404]: #2407 +#2293 := (not #1501) +#2294 := (~ #2293 #1498) +#2291 := (~ #1498 #1498) +#2289 := (~ #1495 #1495) +#2290 := [refl]: #2289 +#2292 := [nnf-pos #2290]: #2291 +#2295 := [nnf-neg #2292]: #2294 +#2412 := [nnf-neg #2295 #2408]: #2411 +#2285 := (~ #1501 #2284) +#2286 := [sk]: #2285 +#2416 := [nnf-neg #2286 #2412]: #2415 +#2279 := (~ #2278 #2278) +#2280 := [refl]: #2279 +#2420 := [nnf-neg #2280 #2416]: #2419 +#2276 := (~ #1492 #1492) +#2277 := [refl]: #2276 +#2424 := [nnf-neg #2277 #2420]: #2423 +#2273 := (not #939) +#2274 := (~ #2273 #785) +#2271 := (~ #785 #785) +#2269 := (~ #780 #780) +#2270 := [refl]: #2269 +#2272 := [nnf-pos #2270]: #2271 +#2275 := [nnf-neg #2272]: #2274 +#2428 := [nnf-neg #2275 #2424]: #2427 +#2265 := (~ #939 #2264) +#2266 := [sk]: #2265 +#2432 := [nnf-neg #2266 #2428]: #2431 +#2253 := (not #1487) +#2254 := (~ #2253 #1484) +#2251 := (~ #1484 #1484) +#2249 := (~ #1479 #1479) +#2250 := [refl]: #2249 +#2252 := [nnf-pos #2250]: #2251 +#2255 := [nnf-neg #2252]: #2254 +#2436 := [nnf-neg #2255 #2432]: #2435 +#2245 := (~ #1487 #2244) +#2246 := [sk]: #2245 +#2440 := [nnf-neg #2246 #2436]: #2439 +#2235 := (not #1476) +#2236 := (~ #2235 #1473) +#2233 := (~ #1473 #1473) +#2231 := (~ #1468 #1468) +#2232 := [refl]: #2231 +#2234 := [nnf-pos #2232]: #2233 +#2237 := [nnf-neg #2234]: #2236 +#2228 := (not #1667) +#2229 := (~ #2228 #2223) +#2224 := (~ #1426 #2223) +#2225 := [sk]: #2224 +#2230 := [nnf-neg #2225]: #2229 +#2212 := (~ #731 #731) +#2213 := [refl]: #2212 +#2210 := (~ #2209 #2209) +#2211 := [refl]: #2210 +#2460 := [nnf-neg #2211 #2213 #2230 #2237 #2440 #2447 #2454 #2456]: #2459 +#2206 := (not #1444) +#2207 := (~ #2206 #2205) +#2203 := (~ #1667 #2202) +#2200 := (~ #2199 #2199) +#2201 := [refl]: #2200 +#2204 := [nnf-neg #2201]: #2203 +#2196 := (not #1415) +#2197 := (~ #2196 #2195) +#2192 := (not #1412) +#2193 := (~ #2192 #2191) +#2188 := (not #1409) +#2189 := (~ #2188 #2187) +#2184 := (not #1404) +#2185 := (~ #2184 #2183) +#2180 := (not #1396) +#2181 := (~ #2180 #1393) +#2178 := (~ #1393 #1393) +#2176 := (~ #1390 #1390) +#2177 := [refl]: #2176 +#2179 := [nnf-pos #2177]: #2178 +#2182 := [nnf-neg #2179]: #2181 +#2174 := (~ #2173 #2173) +#2175 := [refl]: #2174 +#2186 := [nnf-neg #2175 #2182]: #2185 +#2169 := (~ #1396 #2168) +#2170 := [sk]: #2169 +#2190 := [nnf-neg #2170 #2186]: #2189 +#2144 := (not #1378) +#2145 := (~ #2144 #2141) +#2142 := (~ #1375 #2141) +#2139 := (~ #1372 #2138) +#2134 := (~ #1369 #2133) +#2135 := [sk]: #2134 +#2122 := (~ #1353 #1353) +#2123 := [refl]: #2122 +#2140 := [monotonicity #2123 #2135]: #2139 +#2143 := [nnf-pos #2140]: #2142 +#2146 := [nnf-neg #2143]: #2145 +#2194 := [nnf-neg #2146 #2190]: #2193 +#2120 := (~ #1378 #2117) +#2095 := (exists (vars (?x50 T2)) #2094) +#2103 := (or #2102 #2095) +#2104 := (not #2103) +#2118 := (~ #2104 #2117) +#2114 := (not #2095) +#2115 := (~ #2114 #2113) +#2111 := (~ #2110 #2110) +#2112 := [refl]: #2111 +#2116 := [nnf-neg #2112]: #2115 +#2108 := (~ #2107 #2107) +#2109 := [refl]: #2108 +#2119 := [nnf-neg #2109 #2116]: #2118 +#2105 := (~ #1378 #2104) +#2106 := [sk]: #2105 +#2121 := [trans #2106 #2119]: #2120 +#2198 := [nnf-neg #2121 #2194]: #2197 +#2082 := (~ #2081 #2081) +#2083 := [refl]: #2082 +#2079 := (~ #2078 #2078) +#2080 := [refl]: #2079 +#2076 := (~ #2075 #2075) +#2077 := [refl]: #2076 +#2073 := (~ #2072 #2072) +#2074 := [refl]: #2073 +#2208 := [nnf-neg #2074 #2077 #2080 #2083 #2198 #2204]: #2207 +#2464 := [nnf-neg #2208 #2460]: #2463 +#2069 := (not #1341) +#2070 := (~ #2069 #2066) +#2067 := (~ #1338 #2066) +#2064 := (~ #1335 #2063) +#2059 := (~ #1332 #2058) +#2060 := [sk]: #2059 +#2045 := (~ #1308 #1308) +#2046 := [refl]: #2045 +#2065 := [monotonicity #2046 #2060]: #2064 +#2068 := [nnf-pos #2065]: #2067 +#2071 := [nnf-neg #2068]: #2070 +#2042 := (not #1752) +#2043 := (~ #2042 #2039) +#2040 := (~ #1296 #2039) +#2037 := (~ #1293 #2036) +#2032 := (~ #1290 #2031) +#2033 := [sk]: #2032 +#2018 := (~ #1270 #1270) +#2019 := [refl]: #2018 +#2038 := [monotonicity #2019 #2033]: #2037 +#2041 := [nnf-pos #2038]: #2040 +#2044 := [nnf-neg #2041]: #2043 +#2016 := (~ #2015 #2015) +#2017 := [refl]: #2016 +#2489 := [nnf-neg #2017 #2044 #2071 #2464 #2471 #2478 #2485]: #2488 +#2013 := (~ #1752 #2010) +#1988 := (exists (vars (?x38 T2)) #1987) +#1996 := (or #1995 #1988) +#1997 := (not #1996) +#2011 := (~ #1997 #2010) +#2007 := (not #1988) +#2008 := (~ #2007 #2006) +#2004 := (~ #2003 #2003) +#2005 := [refl]: #2004 +#2009 := [nnf-neg #2005]: #2008 +#2001 := (~ #2000 #2000) +#2002 := [refl]: #2001 +#2012 := [nnf-neg #2002 #2009]: #2011 +#1998 := (~ #1752 #1997) +#1999 := [sk]: #1998 +#2014 := [trans #1999 #2012]: #2013 +#2493 := [nnf-neg #2014 #2489]: #2492 +#1973 := (not #1258) +#1974 := (~ #1973 #1255) +#1971 := (~ #1255 #1255) +#1969 := (~ #1252 #1252) +#1970 := [refl]: #1969 +#1972 := [nnf-pos #1970]: #1971 +#1975 := [nnf-neg #1972]: #1974 +#2497 := [nnf-neg #1975 #2493]: #2496 +#1965 := (~ #1258 #1964) +#1966 := [sk]: #1965 +#2501 := [nnf-neg #1966 #2497]: #2500 +#1943 := (not #1232) +#1944 := (~ #1943 #1229) +#1941 := (~ #1229 #1229) +#1939 := (~ #1226 #1226) +#1940 := [refl]: #1939 +#1942 := [nnf-pos #1940]: #1941 +#1945 := [nnf-neg #1942]: #1944 +#2505 := [nnf-neg #1945 #2501]: #2504 +#1935 := (~ #1232 #1934) +#1936 := [sk]: #1935 +#2509 := [nnf-neg #1936 #2505]: #2508 +#1918 := (not #1217) +#1919 := (~ #1918 #1214) +#1916 := (~ #1214 #1214) +#1914 := (~ #1213 #1213) +#1915 := [refl]: #1914 +#1917 := [nnf-pos #1915]: #1916 +#1920 := [nnf-neg #1917]: #1919 +#2513 := [nnf-neg #1920 #2509]: #2512 +#1910 := (~ #1217 #1909) +#1911 := [sk]: #1910 +#2517 := [nnf-neg #1911 #2513]: #2516 +#1865 := (~ #1864 #1864) +#1905 := [refl]: #1865 +#2521 := [nnf-neg #1905 #2517]: #2520 +#1903 := (~ #1208 #1208) +#1904 := [refl]: #1903 +#2524 := [nnf-neg #1904 #2521]: #2523 +#1839 := [not-or-elim #1835]: #1838 +#2525 := [mp~ #1839 #2524]: #2522 +#2526 := [mp #2525 #3063]: #3061 +#3705 := [mp #2526 #3704]: #3702 +#4636 := [mp #3705 #4635]: #4633 +#4922 := [unit-resolution #4636 #4656]: #4630 +#3960 := (or #4627 #4621) +#3961 := [def-axiom]: #3960 +#4948 := [unit-resolution #3961 #4922]: #4621 +#373 := (<= uf_9 0::int) +#374 := (not #373) +#57 := (< 0::int uf_9) +#375 := (iff #57 #374) +#376 := [rewrite]: #375 +#369 := [asserted]: #57 +#377 := [mp #369 #376]: #374 +#4731 := (* -1::int #1907) +#4773 := (+ uf_9 #4731) +#4774 := (<= #4773 0::int) +#4662 := (= uf_9 #1907) +#4665 := (= uf_11 ?x27!0) +#4779 := (not #4665) +#4776 := (= #1907 0::int) +#4795 := (not #4776) +#4789 := [hypothesis]: #1909 +#4796 := (or #4795 #1908) +#4797 := [th-lemma]: #4796 +#4798 := [unit-resolution #4797 #4789]: #4795 +#4767 := (or #3800 #4779 #4776) +#4663 := (= ?x27!0 uf_11) +#4777 := (not #4663) +#4778 := (or #4777 #4776) +#4762 := (or #3800 #4778) +#4791 := (iff #4762 #4767) +#4764 := (or #4779 #4776) +#4769 := (or #3800 #4764) +#4772 := (iff #4769 #4767) +#4790 := [rewrite]: #4772 +#4770 := (iff #4762 #4769) +#4765 := (iff #4778 #4764) +#4780 := (iff #4777 #4779) +#4666 := (iff #4663 #4665) +#4718 := [rewrite]: #4666 +#4763 := [monotonicity #4718]: #4780 +#4766 := [monotonicity #4763]: #4765 +#4771 := [monotonicity #4766]: #4770 +#4792 := [trans #4771 #4790]: #4791 +#4768 := [quant-inst]: #4762 +#4793 := [mp #4768 #4792]: #4767 +#4799 := [unit-resolution #4793 #4642 #4798]: #4779 +#4722 := (or #4662 #4665) +#4707 := (or #4706 #4662 #4665) +#4664 := (or #4663 #4662) +#4708 := (or #4706 #4664) +#4714 := (iff #4708 #4707) +#4710 := (or #4706 #4722) +#4712 := (iff #4710 #4707) +#4713 := [rewrite]: #4712 +#4705 := (iff #4708 #4710) +#4725 := (iff #4664 #4722) +#4719 := (or #4665 #4662) +#4723 := (iff #4719 #4722) +#4724 := [rewrite]: #4723 +#4720 := (iff #4664 #4719) +#4721 := [monotonicity #4718]: #4720 +#4726 := [trans #4721 #4724]: #4725 +#4711 := [monotonicity #4726]: #4705 +#4715 := [trans #4711 #4713]: #4714 +#4709 := [quant-inst]: #4708 +#4730 := [mp #4709 #4715]: #4707 +#4851 := [unit-resolution #4730 #4263]: #4722 +#4852 := [unit-resolution #4851 #4799]: #4662 +#4853 := (not #4662) +#4854 := (or #4853 #4774) +#4855 := [th-lemma]: #4854 +#4856 := [unit-resolution #4855 #4852]: #4774 +#4794 := (<= #1907 0::int) +#4857 := (or #4794 #1908) +#4858 := [th-lemma]: #4857 +#4859 := [unit-resolution #4858 #4789]: #4794 +#4839 := [th-lemma #4859 #4856 #377]: false +#4840 := [lemma #4839]: #1908 +#3955 := (or #4624 #1909 #4618) +#3956 := [def-axiom]: #3955 +#5086 := [unit-resolution #3956 #4840 #4948]: #4618 +#3979 := (or #4615 #4609) +#3980 := [def-axiom]: #3979 +#5292 := [unit-resolution #3980 #5086]: #4609 +#5416 := [hypothesis]: #1928 +#5028 := (or #4844 #1968) +#5029 := [quant-inst]: #5028 +#5422 := [unit-resolution #5029 #4256 #5416]: false +#5423 := [lemma #5422]: #1968 +#3772 := (or #2171 #1928) +#3859 := [def-axiom]: #3772 +#5293 := [unit-resolution #3859 #5423]: #2171 +#3973 := (or #4612 #2248 #4606) +#3975 := [def-axiom]: #3973 +#5417 := [unit-resolution #3975 #5293 #5292]: #4606 +#4000 := (or #4603 #4597) +#4001 := [def-axiom]: #4000 +#5418 := [unit-resolution #4001 #5417]: #4597 +#3996 := (or #4600 #3106 #4594) +#3997 := [def-axiom]: #3996 +#5452 := [unit-resolution #3997 #5418 #4921]: #4594 +#4010 := (or #4591 #4585) +#4031 := [def-axiom]: #4010 +#5507 := [unit-resolution #4031 #5452]: #4585 +#4018 := (or #4588 #4302 #4582) +#4019 := [def-axiom]: #4018 +#5496 := [unit-resolution #4019 #5507 #5196]: #4582 +#4049 := (or #4579 #109) +#4050 := [def-axiom]: #4049 +#5498 := [unit-resolution #4050 #5496]: #109 +#9048 := (= #166 #108) +#9042 := [hypothesis]: #4415 +#4151 := (or #4412 #569) +#4152 := [def-axiom]: #4151 +#9043 := [unit-resolution #4152 #9042]: #569 +#8965 := [symm #9043]: #147 +#9049 := [monotonicity #8965]: #9048 +#9047 := [trans #9049 #5498]: #167 +#3867 := (or #4412 #4348) +#4149 := [def-axiom]: #3867 +#9030 := [unit-resolution #4149 #9042]: #4348 +#7181 := (or #3369 #4353 #674) +#9174 := [hypothesis]: #569 +#9183 := [hypothesis]: #4348 +#9172 := [hypothesis]: #3374 +#4168 := (or #3369 #2164) +#4169 := [def-axiom]: #4168 +#7182 := [unit-resolution #4169 #9172]: #2164 +#9200 := (or #3369 #2163 #4353 #674) +#8980 := (uf_4 uf_14 ?x54!10) +#7073 := (uf_4 uf_14 ?x53!11) +#7171 := (* -1::int #7073) +#9058 := (+ #7171 #8980) +#9059 := (+ #2154 #9058) +#9062 := (>= #9059 0::int) +#8984 := (uf_6 uf_15 ?x54!10) +#8985 := (= uf_8 #8984) +#8981 := (* -1::int #8980) +#8982 := (+ uf_9 #8981) +#8983 := (<= #8982 0::int) +#9168 := (not #8983) +#4166 := (or #3369 #2160) +#4167 := [def-axiom]: #4166 +#9173 := [unit-resolution #4167 #9172]: #2160 +#9161 := (+ #2151 #8981) +#9163 := (>= #9161 0::int) +#9160 := (= #2151 #8980) +#9175 := (= #8980 #2151) +#9176 := [monotonicity #9174]: #9175 +#9177 := [symm #9176]: #9160 +#9178 := (not #9160) +#9179 := (or #9178 #9163) +#9180 := [th-lemma]: #9179 +#9181 := [unit-resolution #9180 #9177]: #9163 +#9169 := (not #9163) +#9170 := (or #9168 #9169 #2159) +#9164 := [hypothesis]: #2160 +#9165 := [hypothesis]: #8983 +#9166 := [hypothesis]: #9163 +#9167 := [th-lemma #9166 #9165 #9164]: false +#9171 := [lemma #9167]: #9170 +#9182 := [unit-resolution #9171 #9181 #9173]: #9168 +#8987 := (or #8983 #8985) +#8990 := (or #4353 #8983 #8985) +#8986 := (or #8985 #8983) +#8991 := (or #4353 #8986) +#8998 := (iff #8991 #8990) +#8993 := (or #4353 #8987) +#8996 := (iff #8993 #8990) +#8997 := [rewrite]: #8996 +#8994 := (iff #8991 #8993) +#8988 := (iff #8986 #8987) +#8989 := [rewrite]: #8988 +#8995 := [monotonicity #8989]: #8994 +#8999 := [trans #8995 #8997]: #8998 +#8992 := [quant-inst]: #8991 +#9000 := [mp #8992 #8999]: #8990 +#9184 := [unit-resolution #9000 #9183]: #8987 +#9185 := [unit-resolution #9184 #9182]: #8985 +#9056 := (not #8985) +#9188 := (or #9056 #9062) +#9186 := [hypothesis]: #2164 +#4052 := (or #4579 #4314) +#4032 := [def-axiom]: #4052 +#9187 := [unit-resolution #4032 #5496]: #4314 +#9073 := (or #4319 #2163 #9056 #9062) +#9051 := (+ #8980 #7171) +#9052 := (+ #2154 #9051) +#9055 := (>= #9052 0::int) +#9057 := (or #9056 #2163 #9055) +#9074 := (or #4319 #9057) +#9081 := (iff #9074 #9073) +#9068 := (or #2163 #9056 #9062) +#9076 := (or #4319 #9068) +#9079 := (iff #9076 #9073) +#9080 := [rewrite]: #9079 +#9077 := (iff #9074 #9076) +#9071 := (iff #9057 #9068) +#9065 := (or #9056 #2163 #9062) +#9069 := (iff #9065 #9068) +#9070 := [rewrite]: #9069 +#9066 := (iff #9057 #9065) +#9063 := (iff #9055 #9062) +#9060 := (= #9052 #9059) +#9061 := [rewrite]: #9060 +#9064 := [monotonicity #9061]: #9063 +#9067 := [monotonicity #9064]: #9066 +#9072 := [trans #9067 #9070]: #9071 +#9078 := [monotonicity #9072]: #9077 +#9082 := [trans #9078 #9080]: #9081 +#9075 := [quant-inst]: #9074 +#9083 := [mp #9075 #9082]: #9073 +#9189 := [unit-resolution #9083 #9187 #9186]: #9188 +#9190 := [unit-resolution #9189 #9185]: #9062 +#4164 := (not #2815) +#4170 := (or #3369 #4164) +#3808 := [def-axiom]: #4170 +#9191 := [unit-resolution #3808 #9172]: #4164 +#9155 := (+ #2149 #7171) +#9158 := (<= #9155 0::int) +#9154 := (= #2149 #7073) +#9192 := (= #7073 #2149) +#9193 := [monotonicity #9174]: #9192 +#9194 := [symm #9193]: #9154 +#9195 := (not #9154) +#9196 := (or #9195 #9158) +#9197 := [th-lemma]: #9196 +#9198 := [unit-resolution #9197 #9194]: #9158 +#9199 := [th-lemma #9198 #9191 #9181 #9190]: false +#9201 := [lemma #9199]: #9200 +#7183 := [unit-resolution #9201 #7182 #9172 #9183 #9174]: false +#7082 := [lemma #7183]: #7181 +#9031 := [unit-resolution #7082 #9030 #9043]: #3369 +#4153 := (or #4412 #4406) +#4150 := [def-axiom]: #4153 +#8964 := [unit-resolution #4150 #9042]: #4406 +#9037 := (or #4412 #2097) +#5396 := (uf_4 uf_14 ?x49!8) +#5456 := (* -1::int #5396) +#5457 := (+ uf_9 #5456) +#5458 := (<= #5457 0::int) +#6793 := (not #5458) +#6788 := [hypothesis]: #2098 +#6206 := (+ #2085 #5456) +#6232 := (>= #6206 0::int) +#5257 := (= #2085 #5396) +#9044 := (= #5396 #2085) +#9045 := [monotonicity #9043]: #9044 +#8975 := [symm #9045]: #5257 +#8976 := (not #5257) +#8977 := (or #8976 #6232) +#8978 := [th-lemma]: #8977 +#8979 := [unit-resolution #8978 #8975]: #6232 +#6794 := (not #6232) +#6792 := (or #6793 #6794 #2097) +#6786 := [hypothesis]: #6232 +#6790 := [hypothesis]: #5458 +#6791 := [th-lemma #6790 #6786 #6788]: false +#6816 := [lemma #6791]: #6792 +#8974 := [unit-resolution #6816 #8979 #6788]: #6793 +#6231 := (<= #6206 0::int) +#8961 := (or #8976 #6231) +#8962 := [th-lemma]: #8961 +#8963 := [unit-resolution #8962 #8975]: #6231 +#4163 := (or #4388 #2173) +#3826 := [def-axiom]: #4163 +#9029 := [unit-resolution #3826 #9047]: #4388 +#4175 := (or #4397 #4391 #3374) +#4161 := [def-axiom]: #4175 +#9032 := [unit-resolution #4161 #9031 #9029]: #4397 +#3885 := (or #4400 #4394) +#3886 := [def-axiom]: #3885 +#9046 := [unit-resolution #3886 #9032]: #4400 +#4155 := (or #4409 #4369 #4403) +#4159 := [def-axiom]: #4155 +#9033 := [unit-resolution #4159 #9046 #8964]: #4369 +#5178 := (?x47!7 ?x49!8) +#6376 := (uf_4 uf_19 #5178) +#6600 := (* -1::int #6376) +#5179 := (uf_4 uf_14 #5178) +#8481 := (+ #5179 #6600) +#6172 := (<= #8481 0::int) +#5887 := (= #5179 #6376) +#9028 := [monotonicity #9043]: #5887 +#9034 := (not #5887) +#9035 := (or #9034 #6172) +#9036 := [th-lemma]: #9035 +#9038 := [unit-resolution #9036 #9028]: #6172 +#8515 := (>= #8481 0::int) +#9053 := (or #9034 #8515) +#9054 := [th-lemma]: #9053 +#9109 := [unit-resolution #9054 #9028]: #8515 +#9290 := (not #6172) +#9289 := (not #6231) +#8950 := (not #8515) +#9263 := (or #4366 #6794 #8950 #9289 #9290 #5458) +#5641 := (+ #5179 #5456) +#5665 := (>= #5641 0::int) +#7930 := (not #5665) +#5450 := (uf_1 #5178 ?x49!8) +#5451 := (uf_10 #5450) +#5631 := (+ #5456 #5451) +#5637 := (+ #5179 #5631) +#5526 := (= #5637 0::int) +#5525 := (not #5526) +#5508 := (uf_6 uf_15 #5178) +#5517 := (= uf_8 #5508) +#5518 := (not #5517) +#5697 := (or #5518 #5525 #5665) +#5707 := (not #5697) +#9240 := [hypothesis]: #6793 +#9241 := [hypothesis]: #4369 +#3837 := (or #4366 #2715) +#3842 := [def-axiom]: #3837 +#9242 := [unit-resolution #3842 #9241]: #2715 +#4039 := (or #4579 #4339) +#4034 := [def-axiom]: #4039 +#9243 := [unit-resolution #4034 #5496]: #4339 +#7629 := (or #4344 #2712 #5458 #5707) +#5424 := (* -1::int #5179) +#5425 := (+ #5396 #5424) +#5426 := (<= #5425 0::int) +#5509 := (* -1::int #5451) +#5514 := (+ #5424 #5509) +#5515 := (+ #5396 #5514) +#5513 := (= #5515 0::int) +#5516 := (not #5513) +#5523 := (or #5518 #5516 #5426) +#5524 := (not #5523) +#5522 := (or #2099 #5458 #5524) +#7083 := (or #4344 #5522) +#7435 := (iff #7083 #7629) +#5819 := (or #2712 #5458 #5707) +#7261 := (or #4344 #5819) +#7430 := (iff #7261 #7629) +#7436 := [rewrite]: #7430 +#7210 := (iff #7083 #7261) +#5798 := (iff #5522 #5819) +#5801 := (iff #5524 #5707) +#5705 := (iff #5523 #5697) +#5638 := (iff #5426 #5665) +#5536 := (+ #5424 #5396) +#5639 := (<= #5536 0::int) +#5666 := (iff #5639 #5665) +#5703 := [rewrite]: #5666 +#5640 := (iff #5426 #5639) +#5497 := (= #5425 #5536) +#5495 := [rewrite]: #5497 +#5634 := [monotonicity #5495]: #5640 +#5704 := [trans #5634 #5703]: #5638 +#5534 := (iff #5516 #5525) +#5533 := (iff #5513 #5526) +#5499 := (+ #5396 #5509) +#5500 := (+ #5424 #5499) +#5504 := (= #5500 0::int) +#5527 := (iff #5504 #5526) +#5532 := [rewrite]: #5527 +#5635 := (iff #5513 #5504) +#5505 := (= #5515 #5500) +#5506 := [rewrite]: #5505 +#5636 := [monotonicity #5506]: #5635 +#5531 := [trans #5636 #5532]: #5533 +#5535 := [monotonicity #5531]: #5534 +#5706 := [monotonicity #5535 #5704]: #5705 +#5818 := [monotonicity #5706]: #5801 +#5817 := [monotonicity #2714 #5818]: #5798 +#6900 := [monotonicity #5817]: #7210 +#7432 := [trans #6900 #7436]: #7435 +#7217 := [quant-inst]: #7083 +#7437 := [mp #7217 #7432]: #7629 +#9244 := [unit-resolution #7437 #9243 #9242 #9240]: #5707 +#7888 := (or #5697 #7930) +#7931 := [def-axiom]: #7888 +#9239 := [unit-resolution #7931 #9244]: #7930 +#6592 := (+ #2085 #6600) +#6591 := (<= #6592 0::int) +#6800 := (+ #5509 #6600) +#6802 := (+ #2085 #6800) +#6783 := (= #6802 0::int) +#8973 := (<= #6802 0::int) +#9282 := [hypothesis]: #6172 +#9283 := [hypothesis]: #6231 +#5866 := (>= #5637 0::int) +#7549 := (or #5697 #5526) +#7928 := [def-axiom]: #7549 +#9245 := [unit-resolution #7928 #9244]: #5526 +#9246 := (or #5525 #5866) +#9247 := [th-lemma]: #9246 +#9248 := [unit-resolution #9247 #9245]: #5866 +#9288 := (not #5866) +#9291 := (or #8973 #9288 #9289 #9290) +#9284 := [hypothesis]: #5866 +#9285 := (not #8973) +#9286 := [hypothesis]: #9285 +#9287 := [th-lemma #9286 #9284 #9283 #9282]: false +#9292 := [lemma #9287]: #9291 +#9249 := [unit-resolution #9292 #9248 #9283 #9282]: #8973 +#9256 := (or #6783 #9285) +#7263 := (>= #6802 0::int) +#8944 := [hypothesis]: #8515 +#5863 := (<= #5637 0::int) +#9251 := (or #5525 #5863) +#9252 := [th-lemma]: #9251 +#9253 := [unit-resolution #9252 #9245]: #5863 +#8949 := (not #5863) +#8951 := (or #7263 #8949 #6794 #8950) +#8945 := [hypothesis]: #5863 +#8946 := (not #7263) +#8947 := [hypothesis]: #8946 +#8948 := [th-lemma #8947 #8945 #6786 #8944]: false +#8952 := [lemma #8948]: #8951 +#9254 := [unit-resolution #8952 #9253 #6786 #8944]: #7263 +#9255 := (or #6783 #9285 #8946) +#9250 := [th-lemma]: #9255 +#9257 := [unit-resolution #9250 #9254]: #9256 +#9258 := [unit-resolution #9257 #9249]: #6783 +#6818 := (not #6783) +#6815 := (or #6591 #6818) +#4178 := (or #4366 #4358) +#3838 := [def-axiom]: #4178 +#9259 := [unit-resolution #3838 #9241]: #4358 +#8898 := (or #4363 #6591 #6818) +#6374 := (+ #2086 #5451) +#6377 := (+ #6376 #6374) +#6477 := (= #6377 0::int) +#6478 := (not #6477) +#6366 := (+ #6376 #2086) +#6479 := (>= #6366 0::int) +#6480 := (or #6479 #6478) +#8899 := (or #4363 #6480) +#8905 := (iff #8899 #8898) +#8124 := (or #4363 #6815) +#8903 := (iff #8124 #8898) +#8904 := [rewrite]: #8903 +#8901 := (iff #8899 #8124) +#6821 := (iff #6480 #6815) +#6819 := (iff #6478 #6818) +#6782 := (iff #6477 #6783) +#6572 := (+ #5451 #6376) +#6011 := (+ #2086 #6572) +#6083 := (= #6011 0::int) +#6787 := (iff #6083 #6783) +#6789 := [rewrite]: #6787 +#6797 := (iff #6477 #6083) +#6795 := (= #6377 #6011) +#6796 := [rewrite]: #6795 +#6799 := [monotonicity #6796]: #6797 +#6817 := [trans #6799 #6789]: #6782 +#6820 := [monotonicity #6817]: #6819 +#6626 := (iff #6479 #6591) +#6481 := (+ #2086 #6376) +#6597 := (>= #6481 0::int) +#6627 := (iff #6597 #6591) +#6628 := [rewrite]: #6627 +#6598 := (iff #6479 #6597) +#6476 := (= #6366 #6481) +#6482 := [rewrite]: #6476 +#6599 := [monotonicity #6482]: #6598 +#6629 := [trans #6599 #6628]: #6626 +#6784 := [monotonicity #6629 #6820]: #6821 +#8902 := [monotonicity #6784]: #8901 +#7262 := [trans #8902 #8904]: #8905 +#8900 := [quant-inst]: #8899 +#7283 := [mp #8900 #7262]: #8898 +#9260 := [unit-resolution #7283 #9259]: #6815 +#9264 := [unit-resolution #9260 #9258]: #6591 +#9265 := [th-lemma #6786 #8944 #9264 #9239]: false +#9266 := [lemma #9265]: #9263 +#9110 := [unit-resolution #9266 #9109 #9038 #9033 #8963 #8979 #8974]: false +#9111 := [lemma #9110]: #9037 +#5419 := [unit-resolution #9111 #9042]: #2097 +#4187 := (or #4366 #2098) +#4177 := [def-axiom]: #4187 +#5867 := [unit-resolution #4177 #5419]: #4366 +#5862 := [unit-resolution #4159 #5867 #8964]: #4403 +#5363 := [unit-resolution #3886 #5862]: #4394 +#7926 := [unit-resolution #4161 #5363 #9031]: #4391 +#7932 := [unit-resolution #3826 #7926]: #2173 +#7884 := [unit-resolution #7932 #9047]: false +#8887 := [lemma #7884]: #4412 +#4040 := (or #4579 #4573) +#4008 := [def-axiom]: #4040 +#10724 := [unit-resolution #4008 #5496]: #4573 +#4045 := (or #4576 #4415 #4570) +#4046 := [def-axiom]: #4045 +#10725 := [unit-resolution #4046 #10724]: #4573 +#10726 := [unit-resolution #10725 #8887]: #4570 +#4068 := (or #4567 #195) +#4069 := [def-axiom]: #4068 +#13581 := [unit-resolution #4069 #10726]: #195 +#13578 := [symm #13581]: #7680 +#13576 := (= #10323 #194) #48 := (:var 0 T5) #47 := (:var 2 T4) #49 := (uf_7 #47 #10 #48) -#4329 := (pattern #49) +#4231 := (pattern #49) #360 := (= uf_8 #48) #50 := (uf_6 #49 #10) #356 := (= uf_8 #50) #363 := (iff #356 #360) -#4330 := (forall (vars (?x17 T4) (?x18 T2) (?x19 T5)) (:pat #4329) #363) +#4232 := (forall (vars (?x17 T4) (?x18 T2) (?x19 T5)) (:pat #4231) #363) #366 := (forall (vars (?x17 T4) (?x18 T2) (?x19 T5)) #363) -#4333 := (iff #366 #4330) -#4331 := (iff #363 #363) -#4332 := [refl]: #4331 -#4334 := [quant-intro #4332]: #4333 -#1957 := (~ #366 #366) -#1995 := (~ #363 #363) -#1996 := [refl]: #1995 -#1958 := [nnf-pos #1996]: #1957 +#4235 := (iff #366 #4232) +#4233 := (iff #363 #363) +#4234 := [refl]: #4233 +#4236 := [quant-intro #4234]: #4235 +#1852 := (~ #366 #366) +#1890 := (~ #363 #363) +#1891 := [refl]: #1890 +#1853 := [nnf-pos #1891]: #1852 #52 := (= #48 uf_8) #51 := (= #50 uf_8) #53 := (iff #51 #52) @@ -4731,274 +4637,150 @@ #368 := [quant-intro #365]: #367 #355 := [asserted]: #54 #371 := [mp #355 #368]: #366 -#1997 := [mp~ #371 #1958]: #366 -#4335 := [mp #1997 #4334]: #4330 -#7014 := (not #4330) -#7015 := (or #7014 #5314) -#5318 := (= uf_8 uf_8) -#5320 := (iff #5314 #5318) -#7018 := (or #7014 #5320) -#7020 := (iff #7018 #7015) -#7022 := (iff #7015 #7015) -#7023 := [rewrite]: #7022 -#5344 := (iff #5320 #5314) -#5323 := (iff #5314 true) -#5342 := (iff #5323 #5314) -#5343 := [rewrite]: #5342 -#5324 := (iff #5320 #5323) -#5321 := (iff #5318 true) -#5322 := [rewrite]: #5321 -#5340 := [monotonicity #5322]: #5324 -#5345 := [trans #5340 #5343]: #5344 -#7021 := [monotonicity #5345]: #7020 -#7024 := [trans #7021 #7023]: #7020 -#7019 := [quant-inst]: #7018 -#7025 := [mp #7019 #7024]: #7015 -#8579 := [unit-resolution #7025 #4335]: #5314 -#8580 := [symm #8579]: #6089 -#8035 := (= #7128 uf_16) -#7129 := (= uf_16 #7128) -#16 := (uf_2 #12) -#325 := (= #10 #16) -#4301 := (forall (vars (?x4 T2) (?x5 T2)) (:pat #4294) #325) -#329 := (forall (vars (?x4 T2) (?x5 T2)) #325) -#4304 := (iff #329 #4301) -#4302 := (iff #325 #325) -#4303 := [refl]: #4302 -#4305 := [quant-intro #4303]: #4304 -#1949 := (~ #329 #329) -#1983 := (~ #325 #325) -#1984 := [refl]: #1983 -#1950 := [nnf-pos #1984]: #1949 -#17 := (= #16 #10) -#18 := (forall (vars (?x4 T2) (?x5 T2)) #17) -#330 := (iff #18 #329) -#327 := (iff #17 #325) -#328 := [rewrite]: #327 -#331 := [quant-intro #328]: #330 -#324 := [asserted]: #18 -#334 := [mp #324 #331]: #329 -#1985 := [mp~ #334 #1950]: #329 -#4306 := [mp #1985 #4305]: #4301 -#7136 := (not #4301) -#7154 := (or #7136 #7129) -#7155 := [quant-inst]: #7154 -#8034 := [unit-resolution #7155 #4306]: #7129 -#8036 := [symm #8034]: #8035 -#8593 := [monotonicity #8036 #8580]: #8592 -#8591 := [trans #8593 #8595]: #8596 -#32982 := [monotonicity #8591 #32984]: #32980 -#32970 := [monotonicity #32982]: #32985 -#32965 := [symm #32970]: #32955 -#32971 := [monotonicity #32965]: #32968 -#32956 := (not #15367) -#32950 := [hypothesis]: #32956 -#15373 := (or #9846 #15367) -#8693 := (= #144 #2212) -#8633 := (= #2212 #144) -#6559 := (= ?x46!9 uf_16) -#7707 := (= ?x46!9 #7128) -#6330 := (uf_6 uf_15 ?x46!9) -#6365 := (= uf_8 #6330) -#7717 := (ite #7707 #5314 #6365) -#7711 := (uf_6 #7203 ?x46!9) -#7714 := (= uf_8 #7711) -#7720 := (iff #7714 #7717) -#8351 := (or #7026 #7720) -#7708 := (ite #7707 #6089 #6365) -#7712 := (= #7711 uf_8) -#7713 := (iff #7712 #7708) -#8361 := (or #7026 #7713) -#8363 := (iff #8361 #8351) -#8365 := (iff #8351 #8351) -#8366 := [rewrite]: #8365 -#7721 := (iff #7713 #7720) -#7718 := (iff #7708 #7717) -#7719 := [monotonicity #6102]: #7718 -#7715 := (iff #7712 #7714) -#7716 := [rewrite]: #7715 -#7722 := [monotonicity #7716 #7719]: #7721 -#8364 := [monotonicity #7722]: #8363 -#8367 := [trans #8364 #8366]: #8363 -#8362 := [quant-inst]: #8361 -#8368 := [mp #8362 #8367]: #8351 -#8576 := [unit-resolution #8368 #4320]: #7720 -#8601 := (= #2208 #7711) -#8597 := (= #7711 #2208) -#8598 := [monotonicity #8591]: #8597 -#8625 := [symm #8598]: #8601 -#8571 := [hypothesis]: #2836 -#4285 := (or #2831 #2209) -#4275 := [def-axiom]: #4285 -#8577 := [unit-resolution #4275 #8571]: #2209 -#8626 := [trans #8577 #8625]: #7714 -#8404 := (not #7714) -#8401 := (not #7720) -#8405 := (or #8401 #8404 #7717) -#8400 := [def-axiom]: #8405 -#8627 := [unit-resolution #8400 #8626 #8576]: #7717 -#6393 := (uf_1 uf_16 ?x46!9) -#6394 := (uf_10 #6393) -#6337 := (* -1::int #2212) -#6411 := (+ #6337 #6394) -#6412 := (+ #144 #6411) -#6413 := (>= #6412 0::int) -#8287 := (not #6413) -#6395 := (* -1::int #6394) -#6396 := (+ uf_9 #6395) -#6397 := (<= #6396 0::int) -#6421 := (or #6397 #6413) -#6426 := (not #6421) -#3935 := (not #2825) -#3940 := (or #2831 #3935) -#4276 := [def-axiom]: #3940 -#8572 := [unit-resolution #4276 #8571]: #3935 -#4212 := (or #4595 #4456) -#4213 := [def-axiom]: #4212 -#8574 := [unit-resolution #4213 #7389]: #4456 -#8302 := (or #4461 #2825 #6426) -#6398 := (+ #1449 #6395) -#6399 := (+ #2212 #6398) -#6400 := (<= #6399 0::int) -#6401 := (or #6400 #6397) -#6402 := (not #6401) -#6403 := (or #2213 #6402) -#8303 := (or #4461 #6403) -#8310 := (iff #8303 #8302) -#6429 := (or #2825 #6426) -#8305 := (or #4461 #6429) -#8308 := (iff #8305 #8302) -#8309 := [rewrite]: #8308 -#8306 := (iff #8303 #8305) -#6430 := (iff #6403 #6429) -#6427 := (iff #6402 #6426) -#6424 := (iff #6401 #6421) -#6418 := (or #6413 #6397) -#6422 := (iff #6418 #6421) -#6423 := [rewrite]: #6422 -#6419 := (iff #6401 #6418) -#6416 := (iff #6400 #6413) -#6404 := (+ #2212 #6395) -#6405 := (+ #1449 #6404) -#6408 := (<= #6405 0::int) -#6414 := (iff #6408 #6413) -#6415 := [rewrite]: #6414 -#6409 := (iff #6400 #6408) -#6406 := (= #6399 #6405) -#6407 := [rewrite]: #6406 -#6410 := [monotonicity #6407]: #6409 -#6417 := [trans #6410 #6415]: #6416 -#6420 := [monotonicity #6417]: #6419 -#6425 := [trans #6420 #6423]: #6424 -#6428 := [monotonicity #6425]: #6427 -#6431 := [monotonicity #2827 #6428]: #6430 -#8307 := [monotonicity #6431]: #8306 -#8311 := [trans #8307 #8309]: #8310 -#8304 := [quant-inst]: #8303 -#8312 := [mp #8304 #8311]: #8302 -#8628 := [unit-resolution #8312 #8574 #8572]: #6426 -#8288 := (or #6421 #8287) -#8289 := [def-axiom]: #8288 -#8629 := [unit-resolution #8289 #8628]: #8287 -#8369 := (not #7717) -#9187 := (or #7707 #6413 #8369) -#7754 := (uf_1 #7128 ?x46!9) -#7838 := (uf_3 #7754) -#9086 := (uf_4 uf_14 #7838) -#9087 := (* -1::int #9086) -#7168 := (uf_4 uf_14 #7128) -#9088 := (+ #7168 #9087) -#9089 := (>= #9088 0::int) -#9090 := (uf_6 uf_15 #7838) -#9091 := (= uf_8 #9090) -#9139 := (= #6330 #9090) -#9135 := (= #9090 #6330) -#9133 := (= #7838 ?x46!9) -#7839 := (= ?x46!9 #7838) -#8519 := (or #5378 #7839) -#8257 := [quant-inst]: #8519 -#9132 := [unit-resolution #8257 #4300]: #7839 -#9134 := [symm #9132]: #9133 -#9136 := [monotonicity #9134]: #9135 -#9140 := [symm #9136]: #9139 -#9129 := [hypothesis]: #7717 -#7733 := (not #7707) -#9130 := [hypothesis]: #7733 -#8347 := (or #8369 #7707 #6365) -#8354 := [def-axiom]: #8347 -#9131 := [unit-resolution #8354 #9130 #9129]: #6365 -#9141 := [trans #9131 #9140]: #9091 -#9092 := (not #9091) -#9154 := (or #9089 #9092) -#7246 := (uf_6 uf_15 #7128) -#7247 := (= uf_8 #7246) -#9149 := (not #7247) -#9150 := (iff #588 #9149) -#9147 := (iff #585 #7247) -#9145 := (iff #7247 #585) -#9143 := (= #7246 #141) -#9144 := [monotonicity #8036]: #9143 -#9146 := [monotonicity #9144]: #9145 -#9148 := [symm #9146]: #9147 -#9151 := [monotonicity #9148]: #9150 -#4127 := (or #4595 #588) -#4220 := [def-axiom]: #4127 -#9142 := [unit-resolution #4220 #7389]: #588 -#9152 := [mp #9142 #9151]: #9149 -#4076 := (or #4677 #4421) -#4131 := [def-axiom]: #4076 -#9153 := [unit-resolution #4131 #7304]: #4421 -#9097 := (or #4426 #7247 #9089 #9092) -#9093 := (or #9092 #7247 #9089) -#9098 := (or #4426 #9093) -#9105 := (iff #9098 #9097) -#9094 := (or #7247 #9089 #9092) -#9100 := (or #4426 #9094) -#9103 := (iff #9100 #9097) -#9104 := [rewrite]: #9103 -#9101 := (iff #9098 #9100) -#9095 := (iff #9093 #9094) -#9096 := [rewrite]: #9095 -#9102 := [monotonicity #9096]: #9101 -#9106 := [trans #9102 #9104]: #9105 -#9099 := [quant-inst]: #9098 -#9107 := [mp #9099 #9106]: #9097 -#9155 := [unit-resolution #9107 #9153 #9152]: #9154 -#9156 := [unit-resolution #9155 #9141]: #9089 -#9157 := [hypothesis]: #8287 -#7755 := (uf_10 #7754) -#7756 := (* -1::int #7755) -#8520 := (+ #6394 #7756) -#8524 := (>= #8520 0::int) -#8517 := (= #6394 #7755) -#9160 := (= #7755 #6394) -#9158 := (= #7754 #6393) -#9159 := [monotonicity #8036]: #9158 -#9161 := [monotonicity #9159]: #9160 -#9162 := [symm #9161]: #8517 -#9163 := (not #8517) -#9164 := (or #9163 #8524) -#9165 := [th-lemma]: #9164 -#9166 := [unit-resolution #9165 #9162]: #8524 -#8514 := (>= #7755 0::int) -#7798 := (<= #7755 0::int) -#7799 := (not #7798) -#7804 := (or #7707 #7799) +#1892 := [mp~ #371 #1853]: #366 +#4237 := [mp #1892 #4236]: #4232 +#3883 := (not #4232) +#3893 := (or #3883 #3895) +#3909 := (= uf_8 uf_8) +#3896 := (iff #3895 #3909) +#3874 := (or #3883 #3896) +#3876 := (iff #3874 #3893) +#3877 := (iff #3893 #3893) +#3878 := [rewrite]: #3877 +#3891 := (iff #3896 #3895) +#3899 := (iff #3895 true) +#3890 := (iff #3899 #3895) +#3882 := [rewrite]: #3890 +#3888 := (iff #3896 #3899) +#3898 := (iff #3909 true) +#3897 := [rewrite]: #3898 +#3889 := [monotonicity #3897]: #3888 +#3892 := [trans #3889 #3882]: #3891 +#3868 := [monotonicity #3892]: #3876 +#3869 := [trans #3868 #3878]: #3876 +#3875 := [quant-inst]: #3874 +#3879 := [mp #3875 #3869]: #3893 +#7677 := [unit-resolution #3879 #4237]: #3895 +#7678 := [symm #7677]: #4958 +#13577 := [monotonicity #10708 #7678]: #13576 +#13583 := [trans #13577 #13578]: #13582 +#13585 := [monotonicity #13583]: #13584 +#13596 := [symm #13585]: #13595 +#13580 := [hypothesis]: #2258 +#13597 := [trans #13580 #13596]: #10327 +#13450 := (not #10327) +#13447 := (not #10333) +#13451 := (or #13447 #13450 #10330) +#13454 := [def-axiom]: #13451 +#13593 := [unit-resolution #13454 #13597 #13579]: #10330 +#13335 := (not #9519) +#11136 := (uf_6 uf_15 #10448) +#11137 := (= uf_8 #11136) +#11138 := (not #11137) +#11133 := (* -1::int #11132) +#9701 := (uf_4 uf_14 #9695) +#11134 := (+ #9701 #11133) +#11135 := (>= #11134 0::int) +#13999 := (not #11135) +#9563 := (uf_1 uf_22 ?x63!14) +#9564 := (uf_10 #9563) +#9508 := (* -1::int #2261) +#9581 := (+ #9508 #9564) +#9582 := (+ #188 #9581) +#9583 := (>= #9582 0::int) +#12179 := (not #9583) +#9565 := (* -1::int #9564) +#9569 := (+ uf_9 #9565) +#9570 := (<= #9569 0::int) +#9588 := (or #9570 #9583) +#9591 := (not #9588) +#13188 := [hypothesis]: #4140 +#4061 := (or #4567 #4428) +#4056 := [def-axiom]: #4061 +#10727 := [unit-resolution #4056 #10726]: #4428 +#13444 := (or #4433 #2866 #9591) +#9566 := (+ #1455 #9565) +#9567 := (+ #2261 #9566) +#9568 := (<= #9567 0::int) +#9571 := (or #9570 #9568) +#9572 := (not #9571) +#9573 := (or #2262 #9572) +#11618 := (or #4433 #9573) +#11883 := (iff #11618 #13444) +#9594 := (or #2866 #9591) +#11662 := (or #4433 #9594) +#11855 := (iff #11662 #13444) +#11841 := [rewrite]: #11855 +#11663 := (iff #11618 #11662) +#9595 := (iff #9573 #9594) +#9592 := (iff #9572 #9591) +#9589 := (iff #9571 #9588) +#9586 := (iff #9568 #9583) +#9574 := (+ #2261 #9565) +#9575 := (+ #1455 #9574) +#9578 := (<= #9575 0::int) +#9584 := (iff #9578 #9583) +#9585 := [rewrite]: #9584 +#9579 := (iff #9568 #9578) +#9576 := (= #9567 #9575) +#9577 := [rewrite]: #9576 +#9580 := [monotonicity #9577]: #9579 +#9587 := [trans #9580 #9585]: #9586 +#9590 := [monotonicity #9587]: #9589 +#9593 := [monotonicity #9590]: #9592 +#9596 := [monotonicity #2868 #9593]: #9595 +#11801 := [monotonicity #9596]: #11663 +#11930 := [trans #11801 #11841]: #11883 +#11417 := [quant-inst]: #11618 +#12087 := [mp #11417 #11930]: #13444 +#13594 := [unit-resolution #12087 #10727 #13188]: #9591 +#12344 := (or #9588 #12179) +#11143 := [def-axiom]: #12344 +#13599 := [unit-resolution #11143 #13594]: #12179 +#14000 := (or #13999 #9583) +#13979 := [hypothesis]: #11135 +#13980 := [hypothesis]: #12179 +#10368 := (uf_10 #10367) +#10372 := (* -1::int #10368) +#11940 := (+ #9564 #10372) +#11942 := (>= #11940 0::int) +#9867 := (= #9564 #10368) +#13987 := (= #9563 #10367) +#13988 := [monotonicity #10706]: #13987 +#13989 := [monotonicity #13988]: #9867 +#13990 := (not #9867) +#13991 := (or #13990 #11942) +#13992 := [th-lemma]: #13991 +#13993 := [unit-resolution #13992 #13989]: #11942 +#11219 := (>= #10368 0::int) +#10318 := (= #10368 0::int) +#11983 := (not #10318) +#11981 := (not #11219) +#11982 := [hypothesis]: #11981 +#11984 := (or #11983 #11219) +#11985 := [th-lemma]: #11984 +#11986 := [unit-resolution #11985 #11982]: #11983 +#10404 := (<= #10368 0::int) +#11987 := (or #11219 #10404) +#11988 := [th-lemma]: #11987 +#11989 := [unit-resolution #11988 #11982]: #10404 +#10405 := (not #10404) +#10410 := (or #10319 #10405) #59 := (uf_10 #12) #409 := (<= #59 0::int) #410 := (not #409) #58 := (= #10 #11) #413 := (or #58 #410) -#4342 := (forall (vars (?x22 T2) (?x23 T2)) (:pat #4294) #413) +#4244 := (forall (vars (?x22 T2) (?x23 T2)) (:pat #4196) #413) #416 := (forall (vars (?x22 T2) (?x23 T2)) #413) -#4345 := (iff #416 #4342) -#4343 := (iff #413 #413) -#4344 := [refl]: #4343 -#4346 := [quant-intro #4344]: #4345 -#1963 := (~ #416 #416) -#1962 := (~ #413 #413) -#2000 := [refl]: #1962 -#1964 := [nnf-pos #2000]: #1963 +#4247 := (iff #416 #4244) +#4245 := (iff #413 #413) +#4246 := [refl]: #4245 +#4248 := [quant-intro #4246]: #4247 +#1858 := (~ #416 #416) +#1857 := (~ #413 #413) +#1895 := [refl]: #1857 +#1859 := [nnf-pos #1895]: #1858 #64 := (< 0::int #59) #63 := (not #58) #65 := (implies #63 #64) @@ -5019,98 +4801,42 @@ #420 := [trans #408 #418]: #419 #402 := [asserted]: #66 #421 := [mp #402 #420]: #416 -#2001 := [mp~ #421 #1964]: #416 -#4347 := [mp #2001 #4346]: #4342 -#7093 := (not #4342) -#8435 := (or #7093 #7707 #7799) -#7800 := (= #7128 ?x46!9) -#7801 := (or #7800 #7799) -#8436 := (or #7093 #7801) -#8447 := (iff #8436 #8435) -#8437 := (or #7093 #7804) -#8443 := (iff #8437 #8435) -#8444 := [rewrite]: #8443 -#8438 := (iff #8436 #8437) -#7805 := (iff #7801 #7804) -#7802 := (iff #7800 #7707) -#7803 := [rewrite]: #7802 -#7806 := [monotonicity #7803]: #7805 -#8439 := [monotonicity #7806]: #8438 -#8448 := [trans #8439 #8444]: #8447 -#8434 := [quant-inst]: #8436 -#8482 := [mp #8434 #8448]: #8435 -#9167 := [unit-resolution #8482 #4347]: #7804 -#9168 := [unit-resolution #9167 #9130]: #7799 -#9169 := (or #8514 #7798) -#9170 := [th-lemma]: #9169 -#9171 := [unit-resolution #9170 #9168]: #8514 -#9126 := (+ #2212 #9087) -#9127 := (<= #9126 0::int) -#9125 := (= #2212 #9086) -#9172 := (= #9086 #2212) -#9173 := [monotonicity #9134]: #9172 -#9174 := [symm #9173]: #9125 -#9175 := (not #9125) -#9176 := (or #9175 #9127) -#9177 := [th-lemma]: #9176 -#9178 := [unit-resolution #9177 #9174]: #9127 -#7162 := (* -1::int #7168) -#7568 := (+ #144 #7162) -#7572 := (>= #7568 0::int) -#7233 := (= #144 #7168) -#9179 := (= #7168 #144) -#9180 := [monotonicity #8036]: #9179 -#9181 := [symm #9180]: #7233 -#9182 := (not #7233) -#9183 := (or #9182 #7572) -#9184 := [th-lemma]: #9183 -#9185 := [unit-resolution #9184 #9181]: #7572 -#9186 := [th-lemma #9185 #9178 #9171 #9166 #9157 #9156]: false -#9188 := [lemma #9186]: #9187 -#8624 := [unit-resolution #9188 #8629 #8627]: #7707 -#8630 := [trans #8624 #8036]: #6559 -#8634 := [monotonicity #8630]: #8633 -#8694 := [symm #8634]: #8693 -#8695 := (= #2211 #144) -#5856 := (uf_18 uf_16) -#8641 := (= #5856 #144) -#5857 := (= #144 #5856) -#5844 := (uf_1 uf_16 uf_16) -#5845 := (uf_10 #5844) -#5864 := (>= #5845 0::int) -#5848 := (* -1::int #5845) -#5849 := (+ uf_9 #5848) -#5850 := (<= #5849 0::int) -#5872 := (or #5850 #5864) -#7965 := (uf_1 #7128 #7128) -#7966 := (uf_10 #7965) -#7967 := (* -1::int #7966) -#8029 := (+ #5845 #7967) -#8033 := (>= #8029 0::int) -#8028 := (= #5845 #7966) -#8039 := (= #5844 #7965) -#8037 := (= #7965 #5844) -#8038 := [monotonicity #8036 #8036]: #8037 -#8040 := [symm #8038]: #8039 -#8041 := [monotonicity #8040]: #8028 -#8042 := (not #8028) -#8043 := (or #8042 #8033) -#8044 := [th-lemma]: #8043 -#8045 := [unit-resolution #8044 #8041]: #8033 -#7976 := (>= #7966 0::int) -#7998 := (= #7966 0::int) +#1896 := [mp~ #421 #1859]: #416 +#4249 := [mp #1896 #4248]: #4244 +#10302 := (not #4244) +#11162 := (or #10302 #10319 #10405) +#10406 := (= #9695 ?x63!14) +#10407 := (or #10406 #10405) +#11163 := (or #10302 #10407) +#11205 := (iff #11163 #11162) +#11188 := (or #10302 #10410) +#11203 := (iff #11188 #11162) +#11204 := [rewrite]: #11203 +#11189 := (iff #11163 #11188) +#10411 := (iff #10407 #10410) +#10408 := (iff #10406 #10319) +#10409 := [rewrite]: #10408 +#10412 := [monotonicity #10409]: #10411 +#11202 := [monotonicity #10412]: #11189 +#11206 := [trans #11202 #11204]: #11205 +#11164 := [quant-inst]: #11163 +#11201 := [mp #11164 #11206]: #11162 +#11990 := [unit-resolution #11201 #4249]: #10410 +#11991 := [unit-resolution #11990 #11989]: #10319 +#10346 := (not #10319) +#10431 := (or #10318 #10346) #60 := (= #59 0::int) #393 := (or #63 #60) -#4336 := (forall (vars (?x20 T2) (?x21 T2)) (:pat #4294) #393) +#4238 := (forall (vars (?x20 T2) (?x21 T2)) (:pat #4196) #393) #396 := (forall (vars (?x20 T2) (?x21 T2)) #393) -#4339 := (iff #396 #4336) -#4337 := (iff #393 #393) -#4338 := [refl]: #4337 -#4340 := [quant-intro #4338]: #4339 -#1959 := (~ #396 #396) -#1998 := (~ #393 #393) -#1999 := [refl]: #1998 -#1960 := [nnf-pos #1999]: #1959 +#4241 := (iff #396 #4238) +#4239 := (iff #393 #393) +#4240 := [refl]: #4239 +#4242 := [quant-intro #4240]: #4241 +#1854 := (~ #396 #396) +#1893 := (~ #393 #393) +#1894 := [refl]: #1893 +#1855 := [nnf-pos #1894]: #1854 #61 := (implies #58 #60) #62 := (forall (vars (?x20 T2) (?x21 T2)) #61) #399 := (iff #62 #396) @@ -5137,1945 +4863,2267 @@ #400 := [trans #390 #398]: #399 #370 := [asserted]: #62 #401 := [mp #370 #400]: #396 -#1961 := [mp~ #401 #1960]: #396 -#4341 := [mp #1961 #4340]: #4336 -#6863 := (not #4336) -#8012 := (or #6863 #7998) -#7248 := (= #7128 #7128) -#7999 := (not #7248) -#8000 := (or #7999 #7998) -#8013 := (or #6863 #8000) -#8015 := (iff #8013 #8012) -#8017 := (iff #8012 #8012) -#8018 := [rewrite]: #8017 -#8010 := (iff #8000 #7998) -#8005 := (or false #7998) -#8008 := (iff #8005 #7998) -#8009 := [rewrite]: #8008 -#8006 := (iff #8000 #8005) -#8003 := (iff #7999 false) -#8001 := (iff #7999 #6849) -#7256 := (iff #7248 true) -#7257 := [rewrite]: #7256 -#8002 := [monotonicity #7257]: #8001 -#8004 := [trans #8002 #6853]: #8003 -#8007 := [monotonicity #8004]: #8006 -#8011 := [trans #8007 #8009]: #8010 -#8016 := [monotonicity #8011]: #8015 -#8019 := [trans #8016 #8018]: #8015 -#8014 := [quant-inst]: #8013 -#8020 := [mp #8014 #8019]: #8012 -#8046 := [unit-resolution #8020 #4341]: #7998 -#8047 := (not #7998) -#8048 := (or #8047 #7976) -#8049 := [th-lemma]: #8048 -#8050 := [unit-resolution #8049 #8046]: #7976 -#6974 := (not #5864) -#8051 := [hypothesis]: #6974 -#8052 := [th-lemma #8051 #8050 #8045]: false -#8053 := [lemma #8052]: #5864 -#6975 := (or #5872 #6974) -#6976 := [def-axiom]: #6975 -#8573 := [unit-resolution #6976 #8053]: #5872 -#5877 := (not #5872) -#5880 := (or #5857 #5877) -#6955 := (or #4461 #5857 #5877) -#5851 := (+ #1449 #5848) -#5852 := (+ #144 #5851) -#5853 := (<= #5852 0::int) -#5854 := (or #5853 #5850) -#5855 := (not #5854) -#5858 := (or #5857 #5855) -#6956 := (or #4461 #5858) -#6967 := (iff #6956 #6955) -#6962 := (or #4461 #5880) -#6965 := (iff #6962 #6955) -#6966 := [rewrite]: #6965 -#6963 := (iff #6956 #6962) -#5881 := (iff #5858 #5880) -#5878 := (iff #5855 #5877) -#5875 := (iff #5854 #5872) -#5869 := (or #5864 #5850) -#5873 := (iff #5869 #5872) -#5874 := [rewrite]: #5873 -#5870 := (iff #5854 #5869) -#5867 := (iff #5853 #5864) -#5861 := (<= #5848 0::int) -#5865 := (iff #5861 #5864) -#5866 := [rewrite]: #5865 -#5862 := (iff #5853 #5861) -#5859 := (= #5852 #5848) -#5860 := [rewrite]: #5859 -#5863 := [monotonicity #5860]: #5862 -#5868 := [trans #5863 #5866]: #5867 -#5871 := [monotonicity #5868]: #5870 -#5876 := [trans #5871 #5874]: #5875 -#5879 := [monotonicity #5876]: #5878 -#5882 := [monotonicity #5879]: #5881 -#6964 := [monotonicity #5882]: #6963 -#6968 := [trans #6964 #6966]: #6967 -#6961 := [quant-inst]: #6956 -#6969 := [mp #6961 #6968]: #6955 -#8575 := [unit-resolution #6969 #8574]: #5880 -#8570 := [unit-resolution #8575 #8573]: #5857 -#8642 := [symm #8570]: #8641 -#8631 := (= #2211 #5856) -#8632 := [monotonicity #8630]: #8631 -#8696 := [trans #8632 #8642]: #8695 -#8697 := [trans #8696 #8694]: #2825 -#8698 := [unit-resolution #8572 #8697]: false -#8699 := [lemma #8698]: #2831 -#4203 := (or #4595 #4589) -#4204 := [def-axiom]: #4203 -#9435 := [unit-resolution #4204 #7389]: #4589 -#4209 := (or #4595 #4464) -#4214 := [def-axiom]: #4209 -#7390 := [unit-resolution #4214 #7389]: #4464 -#6363 := (or #2817 #4469 #4461) -#6139 := (uf_1 uf_16 ?x45!8) -#6140 := (uf_10 #6139) -#6165 := (+ #2191 #6140) -#6166 := (+ #144 #6165) -#6192 := (>= #6166 0::int) -#6169 := (= #6166 0::int) -#6144 := (* -1::int #6140) -#6145 := (+ uf_9 #6144) -#6146 := (<= #6145 0::int) -#6226 := (not #6146) -#6158 := (+ #2815 #6140) -#6159 := (+ #144 #6158) -#6160 := (>= #6159 0::int) -#6203 := (or #6146 #6160) -#6208 := (not #6203) -#6197 := (= #2190 #2192) -#6342 := (not #6197) -#6341 := [hypothesis]: #2822 -#6345 := (or #6342 #2817) -#6346 := [th-lemma]: #6345 -#6347 := [unit-resolution #6346 #6341]: #6342 -#6348 := [hypothesis]: #4456 -#6214 := (or #4461 #6197 #6208) -#6147 := (+ #1449 #6144) -#6148 := (+ #2192 #6147) -#6149 := (<= #6148 0::int) -#6193 := (or #6149 #6146) -#6194 := (not #6193) -#6195 := (= #2192 #2190) -#6196 := (or #6195 #6194) -#6215 := (or #4461 #6196) -#6222 := (iff #6215 #6214) -#6211 := (or #6197 #6208) -#6217 := (or #4461 #6211) -#6220 := (iff #6217 #6214) -#6221 := [rewrite]: #6220 -#6218 := (iff #6215 #6217) -#6212 := (iff #6196 #6211) -#6209 := (iff #6194 #6208) -#6206 := (iff #6193 #6203) -#6200 := (or #6160 #6146) -#6204 := (iff #6200 #6203) -#6205 := [rewrite]: #6204 -#6201 := (iff #6193 #6200) -#6163 := (iff #6149 #6160) -#6151 := (+ #2192 #6144) -#6152 := (+ #1449 #6151) -#6155 := (<= #6152 0::int) -#6161 := (iff #6155 #6160) -#6162 := [rewrite]: #6161 -#6156 := (iff #6149 #6155) -#6153 := (= #6148 #6152) -#6154 := [rewrite]: #6153 -#6157 := [monotonicity #6154]: #6156 -#6164 := [trans #6157 #6162]: #6163 -#6202 := [monotonicity #6164]: #6201 -#6207 := [trans #6202 #6205]: #6206 -#6210 := [monotonicity #6207]: #6209 -#6198 := (iff #6195 #6197) -#6199 := [rewrite]: #6198 -#6213 := [monotonicity #6199 #6210]: #6212 -#6219 := [monotonicity #6213]: #6218 -#6223 := [trans #6219 #6221]: #6222 -#6216 := [quant-inst]: #6215 -#6224 := [mp #6216 #6223]: #6214 -#6349 := [unit-resolution #6224 #6348 #6347]: #6208 -#6227 := (or #6203 #6226) -#6228 := [def-axiom]: #6227 -#6350 := [unit-resolution #6228 #6349]: #6226 -#6229 := (not #6160) -#6230 := (or #6203 #6229) -#6231 := [def-axiom]: #6230 -#6351 := [unit-resolution #6231 #6349]: #6229 -#6175 := (or #6146 #6160 #6169) -#6352 := [hypothesis]: #4464 -#6180 := (or #4469 #6146 #6160 #6169) -#6141 := (+ #6140 #2191) -#6142 := (+ #144 #6141) -#6143 := (= #6142 0::int) -#6150 := (or #6149 #6146 #6143) -#6181 := (or #4469 #6150) -#6188 := (iff #6181 #6180) -#6183 := (or #4469 #6175) -#6186 := (iff #6183 #6180) -#6187 := [rewrite]: #6186 -#6184 := (iff #6181 #6183) -#6178 := (iff #6150 #6175) -#6172 := (or #6160 #6146 #6169) -#6176 := (iff #6172 #6175) -#6177 := [rewrite]: #6176 -#6173 := (iff #6150 #6172) -#6170 := (iff #6143 #6169) -#6167 := (= #6142 #6166) -#6168 := [rewrite]: #6167 -#6171 := [monotonicity #6168]: #6170 -#6174 := [monotonicity #6164 #6171]: #6173 -#6179 := [trans #6174 #6177]: #6178 -#6185 := [monotonicity #6179]: #6184 -#6189 := [trans #6185 #6187]: #6188 -#6182 := [quant-inst]: #6181 -#6190 := [mp #6182 #6189]: #6180 -#6353 := [unit-resolution #6190 #6352]: #6175 -#6354 := [unit-resolution #6353 #6351 #6350]: #6169 -#6355 := (not #6169) -#6356 := (or #6355 #6192) -#6357 := [th-lemma]: #6356 -#6358 := [unit-resolution #6357 #6354]: #6192 -#6225 := (>= #2816 0::int) -#6359 := (or #6225 #2817) -#6360 := [th-lemma]: #6359 -#6361 := [unit-resolution #6360 #6341]: #6225 -#6362 := [th-lemma #6361 #6351 #6358]: false -#6364 := [lemma #6362]: #6363 -#9436 := [unit-resolution #6364 #7390 #8574]: #2817 -#4123 := (or #4592 #2822 #4586) -#4124 := [def-axiom]: #4123 -#9437 := [unit-resolution #4124 #9436 #9435]: #4586 -#4215 := (or #4583 #4577) -#4216 := [def-axiom]: #4215 -#9438 := [unit-resolution #4216 #9437]: #4577 -#4111 := (or #4580 #2836 #4574) -#4070 := [def-axiom]: #4111 -#9439 := [unit-resolution #4070 #9438]: #4577 -#9422 := [unit-resolution #9439 #8699]: #4574 -#4068 := (or #4571 #4481) -#4069 := [def-axiom]: #4068 -#9423 := [unit-resolution #4069 #9422]: #4481 -#10877 := (or #4486 #9846 #15367) -#15363 := (= #15362 #2306) -#15366 := (or #15363 #9846) -#14529 := (or #4486 #15366) -#14658 := (iff #14529 #10877) -#14681 := (or #4486 #15373) -#14554 := (iff #14681 #10877) -#14690 := [rewrite]: #14554 -#12540 := (iff #14529 #14681) -#15376 := (iff #15366 #15373) -#15370 := (or #15367 #9846) -#15374 := (iff #15370 #15373) -#15375 := [rewrite]: #15374 -#15371 := (iff #15366 #15370) -#15368 := (iff #15363 #15367) -#15369 := [rewrite]: #15368 -#15372 := [monotonicity #15369]: #15371 -#15377 := [trans #15372 #15375]: #15376 -#14677 := [monotonicity #15377]: #12540 -#14520 := [trans #14677 #14690]: #14658 -#14692 := [quant-inst]: #14529 -#10887 := [mp #14692 #14520]: #10877 -#32978 := [unit-resolution #10887 #9423]: #15373 -#32979 := [unit-resolution #32978 #32950]: #9846 -#32981 := [mp #32979 #32971]: #30313 -#18779 := (= ?x52!15 #7128) -#32989 := (iff #18779 #32602) -#32770 := (iff #32602 #18779) -#25219 := (= #7128 ?x52!15) -#25223 := (iff #25219 #18779) -#29787 := [commutativity]: #25223 -#32974 := (iff #32602 #25219) -#32975 := [monotonicity #32984]: #32974 -#32986 := [trans #32975 #29787]: #32770 -#32991 := [symm #32986]: #32989 -#15413 := (uf_1 uf_16 ?x52!15) -#15414 := (uf_10 #15413) -#15439 := (+ #2307 #15414) -#15440 := (+ #144 #15439) -#15443 := (= #15440 0::int) -#15432 := (+ #15397 #15414) -#15433 := (+ #144 #15432) -#15434 := (>= #15433 0::int) -#15418 := (* -1::int #15414) -#15419 := (+ uf_9 #15418) -#15420 := (<= #15419 0::int) -#15473 := (or #15420 #15434) -#15478 := (not #15473) -#15481 := (or #15367 #15478) -#14730 := (or #4461 #15367 #15478) -#15421 := (+ #1449 #15418) -#15422 := (+ #15362 #15421) -#15423 := (<= #15422 0::int) -#15467 := (or #15423 #15420) -#15468 := (not #15467) -#15469 := (or #15363 #15468) -#14738 := (or #4461 #15469) -#14986 := (iff #14738 #14730) -#14556 := (or #4461 #15481) -#14896 := (iff #14556 #14730) -#15034 := [rewrite]: #14896 -#14832 := (iff #14738 #14556) -#15482 := (iff #15469 #15481) -#15479 := (iff #15468 #15478) -#15476 := (iff #15467 #15473) -#15470 := (or #15434 #15420) -#15474 := (iff #15470 #15473) -#15475 := [rewrite]: #15474 -#15471 := (iff #15467 #15470) -#15437 := (iff #15423 #15434) -#15425 := (+ #15362 #15418) -#15426 := (+ #1449 #15425) -#15429 := (<= #15426 0::int) -#15435 := (iff #15429 #15434) -#15436 := [rewrite]: #15435 -#15430 := (iff #15423 #15429) -#15427 := (= #15422 #15426) -#15428 := [rewrite]: #15427 -#15431 := [monotonicity #15428]: #15430 -#15438 := [trans #15431 #15436]: #15437 -#15472 := [monotonicity #15438]: #15471 -#15477 := [trans #15472 #15475]: #15476 -#15480 := [monotonicity #15477]: #15479 -#15483 := [monotonicity #15369 #15480]: #15482 -#15031 := [monotonicity #15483]: #14832 -#14985 := [trans #15031 #15034]: #14986 -#12501 := [quant-inst]: #14738 -#15678 := [mp #12501 #14985]: #14730 -#32969 := [unit-resolution #15678 #8574]: #15481 -#32952 := [unit-resolution #32969 #32950]: #15478 -#29629 := (or #15473 #15443) -#25301 := (not #15443) -#29623 := [hypothesis]: #25301 -#15187 := (not #15420) -#29624 := [hypothesis]: #15478 -#14833 := (or #15473 #15187) -#15233 := [def-axiom]: #14833 -#29625 := [unit-resolution #15233 #29624]: #15187 -#8899 := (not #15434) -#15050 := (or #15473 #8899) -#15028 := [def-axiom]: #15050 -#29626 := [unit-resolution #15028 #29624]: #8899 -#15449 := (or #15420 #15434 #15443) -#12503 := (or #4469 #15420 #15434 #15443) -#15415 := (+ #15414 #2307) -#15416 := (+ #144 #15415) -#15417 := (= #15416 0::int) -#15424 := (or #15423 #15420 #15417) -#12502 := (or #4469 #15424) -#14824 := (iff #12502 #12503) -#14693 := (or #4469 #15449) -#14698 := (iff #14693 #12503) -#14734 := [rewrite]: #14698 -#14675 := (iff #12502 #14693) -#15452 := (iff #15424 #15449) -#15446 := (or #15434 #15420 #15443) -#15450 := (iff #15446 #15449) -#15451 := [rewrite]: #15450 -#15447 := (iff #15424 #15446) -#15444 := (iff #15417 #15443) -#15441 := (= #15416 #15440) -#15442 := [rewrite]: #15441 -#15445 := [monotonicity #15442]: #15444 -#15448 := [monotonicity #15438 #15445]: #15447 -#15453 := [trans #15448 #15451]: #15452 -#14648 := [monotonicity #15453]: #14675 -#14687 := [trans #14648 #14734]: #14824 -#14736 := [quant-inst]: #12502 -#13488 := [mp #14736 #14687]: #12503 -#29627 := [unit-resolution #13488 #7390]: #15449 -#29628 := [unit-resolution #29627 #29626 #29625 #29623]: false -#29630 := [lemma #29628]: #29629 -#32972 := [unit-resolution #29630 #32952]: #15443 -#29799 := (or #25301 #18779) -#7126 := (uf_3 #6008) -#15598 := (uf_1 #7126 ?x52!15) -#27533 := (uf_3 #15598) -#28710 := (uf_1 #7128 #27533) -#28711 := (uf_10 #28710) -#28714 := (* -1::int #28711) -#28814 := (+ #15414 #28714) -#28506 := (>= #28814 0::int) -#28505 := (= #15414 #28711) -#29767 := (= #28711 #15414) -#29765 := (= #28710 #15413) -#29763 := (= #27533 ?x52!15) -#27534 := (= ?x52!15 #27533) -#27563 := (or #5378 #27534) -#27564 := [quant-inst]: #27563 -#29762 := [unit-resolution #27564 #4300]: #27534 -#29764 := [symm #29762]: #29763 -#29766 := [monotonicity #8036 #29764]: #29765 -#29768 := [monotonicity #29766]: #29767 -#29769 := [symm #29768]: #28505 -#29770 := (not #28505) -#29771 := (or #29770 #28506) -#29772 := [th-lemma]: #29771 -#29773 := [unit-resolution #29772 #29769]: #28506 -#5902 := (* -1::int #5856) -#6232 := (+ #144 #5902) -#6233 := (>= #6232 0::int) -#4218 := (or #4583 #4472) -#4120 := [def-axiom]: #4218 -#14110 := [unit-resolution #4120 #9437]: #4472 -#6959 := (or #4477 #6233) -#6960 := [quant-inst]: #6959 -#12934 := [unit-resolution #6960 #14110]: #6233 -#7167 := (uf_18 #7128) -#8141 := (* -1::int #7167) -#10187 := (+ #5856 #8141) -#7462 := (>= #10187 0::int) -#10181 := (= #5856 #7167) -#14102 := (= #7167 #5856) -#14103 := [monotonicity #8036]: #14102 -#14104 := [symm #14103]: #10181 -#14105 := (not #10181) -#25236 := (or #14105 #7462) -#25243 := [th-lemma]: #25236 -#25235 := [unit-resolution #25243 #14104]: #7462 -#14406 := (<= #15440 0::int) -#25300 := [hypothesis]: #15443 -#25302 := (or #25301 #14406) -#25303 := [th-lemma]: #25302 -#25304 := [unit-resolution #25303 #25300]: #14406 -#15344 := (+ #2306 #8141) -#15518 := (<= #15344 0::int) -#7164 := (uf_6 uf_17 #7128) -#7165 := (= uf_8 #7164) -#25365 := (= #5319 #7164) -#25361 := (= #7164 #5319) -#25364 := [monotonicity #8578 #8036]: #25361 -#25366 := [symm #25364]: #25365 -#25367 := [trans #8579 #25366]: #7165 -#15503 := (uf_1 #7128 ?x52!15) -#15504 := (uf_10 #15503) -#15530 := (* -1::int #15504) -#15531 := (+ #8141 #15530) -#15532 := (+ #2306 #15531) -#15533 := (= #15532 0::int) -#25324 := (or #25301 #15533) -#15538 := (not #15533) -#25256 := [hypothesis]: #15538 -#14445 := (>= #15532 0::int) -#14444 := (+ #15414 #15530) -#14494 := (>= #14444 0::int) -#14488 := (= #15414 #15504) -#25275 := (= #15504 #15414) -#25257 := (= #15503 #15413) -#25274 := [monotonicity #8036]: #25257 -#25270 := [monotonicity #25274]: #25275 -#25276 := [symm #25270]: #14488 -#25277 := (not #14488) -#25278 := (or #25277 #14494) -#25279 := [th-lemma]: #25278 -#25280 := [unit-resolution #25279 #25276]: #14494 -#25306 := (not #14406) -#25305 := (not #14494) -#13409 := (not #6233) -#25299 := (not #7462) -#25307 := (or #14445 #25299 #13409 #25305 #25306) -#25308 := [th-lemma]: #25307 -#25309 := [unit-resolution #25308 #25304 #25235 #12934 #25280]: #14445 -#14391 := (<= #15532 0::int) -#14441 := (<= #14444 0::int) -#25316 := (or #25277 #14441) -#25317 := [th-lemma]: #25316 -#25315 := [unit-resolution #25317 #25276]: #14441 -#6970 := (<= #6232 0::int) -#14098 := (not #5857) -#14099 := (or #14098 #6970) -#14100 := [th-lemma]: #14099 -#14101 := [unit-resolution #14100 #8570]: #6970 -#10188 := (<= #10187 0::int) -#14106 := (or #14105 #10188) -#14107 := [th-lemma]: #14106 -#14108 := [unit-resolution #14107 #14104]: #10188 -#14435 := (>= #15440 0::int) -#25318 := (or #25301 #14435) -#25319 := [th-lemma]: #25318 -#25320 := [unit-resolution #25319 #25300]: #14435 -#25333 := (not #14435) -#25332 := (not #14441) -#12642 := (not #6970) -#25331 := (not #10188) -#25334 := (or #14391 #25331 #12642 #25332 #25333) -#25335 := [th-lemma]: #25334 -#25336 := [unit-resolution #25335 #25320 #14108 #14101 #25315]: #14391 -#25314 := (not #14445) -#25337 := (not #14391) -#25321 := (or #15533 #25337 #25314) -#25322 := [th-lemma]: #25321 -#25323 := [unit-resolution #25322 #25336 #25309 #25256]: false -#25313 := [lemma #25323]: #25324 -#29774 := [unit-resolution #25313 #25300]: #15533 -#7166 := (not #7165) -#15541 := (or #7166 #15518 #15538) -#6002 := (+ #108 #1449) -#7864 := (<= #6002 0::int) -#23377 := (= #108 #144) -#12197 := (= #144 #108) -#5945 := (= uf_16 uf_11) -#5947 := (= uf_11 uf_16) -#5928 := (?x40!7 uf_16) -#5932 := (uf_4 uf_14 #5928) -#5933 := (* -1::int #5932) -#5929 := (uf_1 #5928 uf_16) -#5930 := (uf_10 #5929) -#5931 := (* -1::int #5930) -#5950 := (+ #5931 #5933) -#5951 := (+ #144 #5950) -#5954 := (= #5951 0::int) -#5957 := (not #5954) -#5940 := (uf_6 uf_15 #5928) -#5941 := (= uf_8 #5940) -#5942 := (not #5941) -#5938 := (+ #144 #5933) -#5939 := (<= #5938 0::int) -#5963 := (or #5939 #5942 #5957) -#6003 := (>= #6002 0::int) -#9539 := (not #7864) -#23470 := [hypothesis]: #9539 -#23505 := (or #7864 #6003) -#23465 := [th-lemma]: #23505 -#23464 := [unit-resolution #23465 #23470]: #6003 -#9672 := (not #6003) -#18009 := (or #9672 #5939) -#7466 := (>= #5932 0::int) -#8252 := (not #7466) -#8253 := [hypothesis]: #8252 -#8212 := (or #4409 #7466) -#8206 := [quant-inst]: #8212 -#8263 := [unit-resolution #8206 #7305 #8253]: false -#8264 := [lemma #8263]: #7466 -#17999 := (or #9672 #8252 #5939) -#4050 := (<= #108 0::int) -#7308 := (or #1854 #4050) -#7309 := [th-lemma]: #7308 -#7310 := [unit-resolution #7309 #7307]: #4050 -#5512 := (not #4050) -#17980 := (or #9672 #5512 #8252 #5939) -#17982 := [th-lemma]: #17980 -#17995 := [unit-resolution #17982 #7310]: #17999 -#18007 := [unit-resolution #17995 #8264]: #18009 -#23469 := [unit-resolution #18007 #23464]: #5939 -#7005 := (not #5939) -#7006 := (or #5963 #7005) -#7007 := [def-axiom]: #7006 -#23506 := [unit-resolution #7007 #23469]: #5963 -#5968 := (not #5963) -#18000 := (or #5947 #5968) -#4208 := (or #4595 #1657) -#4210 := [def-axiom]: #4208 -#16348 := [unit-resolution #4210 #7389]: #1657 -#6992 := (or #4442 #1656 #5947 #5968) -#5934 := (+ #5933 #5931) -#5935 := (+ #144 #5934) -#5936 := (= #5935 0::int) -#5937 := (not #5936) -#5943 := (or #5942 #5939 #5937) -#5944 := (not #5943) -#5946 := (or #5945 #1656 #5944) -#6993 := (or #4442 #5946) -#7000 := (iff #6993 #6992) -#5974 := (or #1656 #5947 #5968) -#6995 := (or #4442 #5974) -#6998 := (iff #6995 #6992) -#6999 := [rewrite]: #6998 -#6996 := (iff #6993 #6995) -#5977 := (iff #5946 #5974) -#5971 := (or #5947 #1656 #5968) -#5975 := (iff #5971 #5974) -#5976 := [rewrite]: #5975 -#5972 := (iff #5946 #5971) -#5969 := (iff #5944 #5968) -#5966 := (iff #5943 #5963) -#5960 := (or #5942 #5939 #5957) -#5964 := (iff #5960 #5963) -#5965 := [rewrite]: #5964 -#5961 := (iff #5943 #5960) -#5958 := (iff #5937 #5957) -#5955 := (iff #5936 #5954) -#5952 := (= #5935 #5951) -#5953 := [rewrite]: #5952 -#5956 := [monotonicity #5953]: #5955 -#5959 := [monotonicity #5956]: #5958 -#5962 := [monotonicity #5959]: #5961 -#5967 := [trans #5962 #5965]: #5966 -#5970 := [monotonicity #5967]: #5969 -#5948 := (iff #5945 #5947) -#5949 := [rewrite]: #5948 -#5973 := [monotonicity #5949 #5970]: #5972 -#5978 := [trans #5973 #5976]: #5977 -#6997 := [monotonicity #5978]: #6996 -#7001 := [trans #6997 #6999]: #7000 -#6994 := [quant-inst]: #6993 -#7002 := [mp #6994 #7001]: #6992 -#18030 := [unit-resolution #7002 #8030 #16348]: #18000 -#23510 := [unit-resolution #18030 #23506]: #5947 -#12009 := [symm #23510]: #5945 -#12163 := [monotonicity #12009]: #12197 -#12023 := [symm #12163]: #23377 -#12215 := (not #23377) -#12216 := (or #12215 #7864) -#11464 := [th-lemma]: #12216 -#12217 := [unit-resolution #11464 #23470 #12023]: false -#12311 := [lemma #12217]: #7864 -#9540 := (or #2234 #9539) -#6554 := (uf_1 uf_16 ?x47!10) -#6555 := (uf_10 #6554) -#6435 := (* -1::int #2233) -#6597 := (+ #6435 #6555) -#6598 := (+ #144 #6597) -#8322 := (<= #6598 0::int) -#6601 := (= #6598 0::int) -#6538 := (* -1::int #6555) -#6539 := (+ uf_9 #6538) -#6540 := (<= #6539 0::int) -#8336 := (not #6540) -#6327 := (uf_4 uf_14 ?x47!10) -#6471 := (* -1::int #6327) -#6590 := (+ #6471 #6555) -#6591 := (+ #144 #6590) -#6592 := (>= #6591 0::int) -#6650 := (or #6540 #6592) -#6660 := (not #6650) -#6344 := (= #2233 #6327) -#9327 := (not #6344) -#6472 := (+ #2233 #6471) -#8301 := (>= #6472 0::int) -#9317 := (not #8301) -#9321 := [hypothesis]: #2235 -#6814 := (>= #6327 0::int) -#7656 := (or #4409 #6814) -#7657 := [quant-inst]: #7656 -#9322 := [unit-resolution #7657 #7305]: #6814 -#9323 := (not #6814) -#9324 := (or #9317 #2234 #9323) -#9325 := [th-lemma]: #9324 -#9326 := [unit-resolution #9325 #9322 #9321]: #9317 -#9347 := (or #9327 #8301) -#9348 := [th-lemma]: #9347 -#9349 := [unit-resolution #9348 #9326]: #9327 -#6662 := (or #6344 #6660) -#8339 := (or #4461 #6344 #6660) -#6541 := (+ #1449 #6538) -#6536 := (+ #6327 #6541) -#6542 := (<= #6536 0::int) -#6641 := (or #6542 #6540) -#6642 := (not #6641) -#6328 := (= #6327 #2233) -#6643 := (or #6328 #6642) -#8314 := (or #4461 #6643) -#8333 := (iff #8314 #8339) -#8281 := (or #4461 #6662) -#8343 := (iff #8281 #8339) -#8332 := [rewrite]: #8343 -#8341 := (iff #8314 #8281) -#6663 := (iff #6643 #6662) -#6661 := (iff #6642 #6660) -#6655 := (iff #6641 #6650) -#6649 := (or #6592 #6540) -#6653 := (iff #6649 #6650) -#6654 := [rewrite]: #6653 -#6651 := (iff #6641 #6649) -#6595 := (iff #6542 #6592) -#6544 := (+ #6327 #6538) -#6545 := (+ #1449 #6544) -#6562 := (<= #6545 0::int) -#6593 := (iff #6562 #6592) -#6594 := [rewrite]: #6593 -#6587 := (iff #6542 #6562) -#6546 := (= #6536 #6545) -#6561 := [rewrite]: #6546 -#6589 := [monotonicity #6561]: #6587 -#6596 := [trans #6589 #6594]: #6595 -#6652 := [monotonicity #6596]: #6651 -#6658 := [trans #6652 #6654]: #6655 -#6659 := [monotonicity #6658]: #6661 -#6383 := (iff #6328 #6344) -#6384 := [rewrite]: #6383 -#6664 := [monotonicity #6384 #6659]: #6663 -#8342 := [monotonicity #6664]: #8341 -#8334 := [trans #8342 #8332]: #8333 -#8340 := [quant-inst]: #8314 -#8335 := [mp #8340 #8334]: #8339 -#9350 := [unit-resolution #8335 #8574]: #6662 -#9351 := [unit-resolution #9350 #9349]: #6660 -#8327 := (or #6650 #8336) -#8352 := [def-axiom]: #8327 -#9346 := [unit-resolution #8352 #9351]: #8336 -#8360 := (not #6592) -#8348 := (or #6650 #8360) -#8353 := [def-axiom]: #8348 -#9352 := [unit-resolution #8353 #9351]: #8360 -#6607 := (or #6540 #6592 #6601) -#8318 := (or #4469 #6540 #6592 #6601) -#6556 := (+ #6555 #6435) -#6557 := (+ #144 #6556) -#6537 := (= #6557 0::int) -#6543 := (or #6542 #6540 #6537) -#8344 := (or #4469 #6543) -#8331 := (iff #8344 #8318) -#8349 := (or #4469 #6607) -#8330 := (iff #8349 #8318) -#8325 := [rewrite]: #8330 -#8328 := (iff #8344 #8349) -#6578 := (iff #6543 #6607) -#6604 := (or #6592 #6540 #6601) -#6579 := (iff #6604 #6607) -#6580 := [rewrite]: #6579 -#6605 := (iff #6543 #6604) -#6602 := (iff #6537 #6601) -#6599 := (= #6557 #6598) -#6600 := [rewrite]: #6599 -#6603 := [monotonicity #6600]: #6602 -#6606 := [monotonicity #6596 #6603]: #6605 -#6581 := [trans #6606 #6580]: #6578 -#8329 := [monotonicity #6581]: #8328 -#8320 := [trans #8329 #8325]: #8331 -#8345 := [quant-inst]: #8344 -#8321 := [mp #8345 #8320]: #8318 -#9353 := [unit-resolution #8321 #7390]: #6607 -#9354 := [unit-resolution #9353 #9352 #9346]: #6601 -#9355 := (not #6601) -#9356 := (or #9355 #8322) -#9367 := [th-lemma]: #9356 -#9368 := [unit-resolution #9367 #9354]: #8322 -#9369 := [hypothesis]: #7864 -#4041 := (>= #108 0::int) -#9370 := (or #1854 #4041) -#9371 := [th-lemma]: #9370 -#9366 := [unit-resolution #9371 #7307]: #4041 -#8900 := (uf_1 #7128 ?x47!10) -#8901 := (uf_10 #8900) -#8908 := (* -1::int #8901) -#9307 := (+ #6555 #8908) -#9320 := (>= #9307 0::int) -#9304 := (= #6555 #8901) -#9374 := (= #8901 #6555) -#9372 := (= #8900 #6554) -#9373 := [monotonicity #8036]: #9372 -#9375 := [monotonicity #9373]: #9374 -#9376 := [symm #9375]: #9304 -#9382 := (not #9304) -#9383 := (or #9382 #9320) -#9432 := [th-lemma]: #9383 -#9433 := [unit-resolution #9432 #9376]: #9320 -#9287 := (>= #8901 0::int) -#9080 := (<= #8901 0::int) -#9207 := (not #9080) -#8373 := (= ?x47!10 #7128) -#8730 := (not #8373) -#6710 := (uf_6 uf_15 ?x47!10) -#6711 := (= uf_8 #6710) -#8599 := (ite #8373 #5314 #6711) -#8729 := (not #8599) -#7658 := (uf_6 #7203 ?x47!10) -#8338 := (= uf_8 #7658) -#8700 := (iff #8338 #8599) -#8684 := (or #7026 #8700) -#7640 := (ite #8373 #6089 #6711) -#7653 := (= #7658 uf_8) -#8337 := (iff #7653 #7640) -#8723 := (or #7026 #8337) -#8725 := (iff #8723 #8684) -#8726 := (iff #8684 #8684) -#8722 := [rewrite]: #8726 -#8682 := (iff #8337 #8700) -#8326 := (iff #7640 #8599) -#8600 := [monotonicity #6102]: #8326 -#8410 := (iff #7653 #8338) -#8521 := [rewrite]: #8410 -#8683 := [monotonicity #8521 #8600]: #8682 -#8720 := [monotonicity #8683]: #8725 -#8727 := [trans #8720 #8722]: #8725 -#8724 := [quant-inst]: #8723 -#8728 := [mp #8724 #8727]: #8684 -#9434 := [unit-resolution #8728 #4320]: #8700 -#8895 := (not #8338) -#6323 := (uf_6 uf_17 ?x47!10) -#6325 := (= uf_8 #6323) -#6326 := (not #6325) -#9431 := (iff #6326 #8895) -#9429 := (iff #6325 #8338) -#9427 := (iff #8338 #6325) -#9426 := (= #7658 #6323) -#9421 := [monotonicity #8591]: #9426 -#9428 := [monotonicity #9421]: #9427 -#9430 := [symm #9428]: #9429 -#9460 := [monotonicity #9430]: #9431 -#6382 := (or #6326 #6344) -#8268 := (or #4486 #6326 #6344) -#6343 := (or #6328 #6326) -#8269 := (or #4486 #6343) -#8297 := (iff #8269 #8268) -#8294 := (or #4486 #6382) -#8296 := (iff #8294 #8268) -#8283 := [rewrite]: #8296 -#8284 := (iff #8269 #8294) -#6390 := (iff #6343 #6382) -#6385 := (or #6344 #6326) -#6388 := (iff #6385 #6382) -#6389 := [rewrite]: #6388 -#6386 := (iff #6343 #6385) -#6387 := [monotonicity #6384]: #6386 -#6391 := [trans #6387 #6389]: #6390 -#8295 := [monotonicity #6391]: #8284 -#8298 := [trans #8295 #8283]: #8297 -#8293 := [quant-inst]: #8269 -#8299 := [mp #8293 #8298]: #8268 -#9424 := [unit-resolution #8299 #9423]: #6382 -#9425 := [unit-resolution #9424 #9349]: #6326 -#9461 := [mp #9425 #9460]: #8895 -#8917 := (not #8700) -#8920 := (or #8917 #8338 #8729) -#8894 := [def-axiom]: #8920 -#9462 := [unit-resolution #8894 #9461 #9434]: #8729 -#9463 := (or #8599 #8730) -#7040 := (not #5314) -#8907 := (or #8599 #8730 #7040) -#8910 := [def-axiom]: #8907 -#9464 := [unit-resolution #8910 #8579]: #9463 -#9459 := [unit-resolution #9464 #9462]: #8730 -#9206 := (or #8373 #9207) -#9214 := (or #7093 #8373 #9207) -#9208 := (= #7128 ?x47!10) -#9209 := (or #9208 #9207) -#9215 := (or #7093 #9209) -#9241 := (iff #9215 #9214) -#9242 := (or #7093 #9206) -#9245 := (iff #9242 #9214) -#9246 := [rewrite]: #9245 -#9243 := (iff #9215 #9242) -#9212 := (iff #9209 #9206) -#9210 := (iff #9208 #8373) -#9211 := [rewrite]: #9210 -#9213 := [monotonicity #9211]: #9212 -#9244 := [monotonicity #9213]: #9243 -#9247 := [trans #9244 #9246]: #9241 -#9216 := [quant-inst]: #9215 -#9248 := [mp #9216 #9247]: #9214 -#9465 := [unit-resolution #9248 #4347]: #9206 -#9466 := [unit-resolution #9465 #9459]: #9207 -#9467 := (or #9287 #9080) -#9468 := [th-lemma]: #9467 -#9469 := [unit-resolution #9468 #9466]: #9287 -#9538 := [th-lemma #9321 #9469 #9433 #9366 #9369 #9368]: false -#9541 := [lemma #9538]: #9540 -#29775 := [unit-resolution #9541 #12311]: #2234 -#4222 := (or #4571 #4565) -#4223 := [def-axiom]: #4222 -#25326 := [unit-resolution #4223 #9422]: #4565 -#25357 := (or #4568 #4562) -#6056 := (= #108 #172) -#25354 := (iff #6056 #173) -#25353 := [commutativity]: #1490 -#25351 := (iff #6056 #645) -#25352 := [monotonicity #7307]: #25351 -#25355 := [trans #25352 #25353]: #25354 -#6015 := (uf_10 #6008) -#6019 := (* -1::int #6015) -#6022 := (+ #1449 #6019) -#6023 := (+ #108 #6022) -#6024 := (<= #6023 0::int) -#6020 := (+ uf_9 #6019) -#6021 := (<= #6020 0::int) -#6058 := (or #6021 #6024) -#7125 := (>= #6015 0::int) -#7105 := (= #6015 0::int) -#7087 := (<= #6015 0::int) -#4062 := (not #6024) -#7293 := [hypothesis]: #4062 -#7312 := (or #7087 #6024) -#7088 := (not #7087) -#7292 := [hypothesis]: #7088 -#6001 := (>= #144 0::int) -#7016 := (or #4409 #6001) -#7017 := [quant-inst]: #7016 -#7306 := [unit-resolution #7017 #7305]: #6001 -#7311 := [th-lemma #7310 #7306 #7293 #7292]: false -#7313 := [lemma #7311]: #7312 -#7186 := [unit-resolution #7313 #7293]: #7087 -#7090 := (or #5947 #7088) -#7094 := (or #7093 #5947 #7088) -#7089 := (or #5945 #7088) -#7095 := (or #7093 #7089) -#7102 := (iff #7095 #7094) -#7097 := (or #7093 #7090) -#7100 := (iff #7097 #7094) -#7101 := [rewrite]: #7100 -#7098 := (iff #7095 #7097) -#7091 := (iff #7089 #7090) -#7092 := [monotonicity #5949]: #7091 -#7099 := [monotonicity #7092]: #7098 -#7103 := [trans #7099 #7101]: #7102 -#7096 := [quant-inst]: #7095 -#7104 := [mp #7096 #7103]: #7094 -#7187 := [unit-resolution #7104 #4347]: #7090 -#7182 := [unit-resolution #7187 #7186]: #5947 -#7108 := (not #5947) -#7111 := (or #7108 #7105) -#7114 := (or #6863 #7108 #7105) -#7106 := (not #5945) -#7107 := (or #7106 #7105) -#7115 := (or #6863 #7107) -#7122 := (iff #7115 #7114) -#7117 := (or #6863 #7111) -#7120 := (iff #7117 #7114) -#7121 := [rewrite]: #7120 -#7118 := (iff #7115 #7117) -#7112 := (iff #7107 #7111) -#7109 := (iff #7106 #7108) -#7110 := [monotonicity #5949]: #7109 -#7113 := [monotonicity #7110]: #7112 -#7119 := [monotonicity #7113]: #7118 -#7123 := [trans #7119 #7121]: #7122 -#7116 := [quant-inst]: #7115 -#7124 := [mp #7116 #7123]: #7114 -#7188 := [unit-resolution #7124 #4341]: #7111 -#7189 := [unit-resolution #7188 #7182]: #7105 -#7190 := (not #7105) -#7191 := (or #7190 #7125) -#7192 := [th-lemma]: #7191 -#7196 := [unit-resolution #7192 #7189]: #7125 -#7197 := [th-lemma #7310 #7306 #7293 #7196]: false -#7195 := [lemma #7197]: #6024 -#5757 := (or #6058 #4062) -#5573 := [def-axiom]: #5757 -#25327 := [unit-resolution #5573 #7195]: #6058 -#6061 := (not #6058) -#6064 := (or #6056 #6061) -#6681 := (or #4461 #6056 #6061) -#6054 := (or #6024 #6021) -#6055 := (not #6054) -#6057 := (or #6056 #6055) -#6682 := (or #4461 #6057) -#6844 := (iff #6682 #6681) -#6138 := (or #4461 #6064) -#6392 := (iff #6138 #6681) -#6683 := [rewrite]: #6392 -#6118 := (iff #6682 #6138) -#6065 := (iff #6057 #6064) -#6062 := (iff #6055 #6061) -#6059 := (iff #6054 #6058) -#6060 := [rewrite]: #6059 -#6063 := [monotonicity #6060]: #6062 -#6066 := [monotonicity #6063]: #6065 -#6735 := [monotonicity #6066]: #6118 -#6845 := [trans #6735 #6683]: #6844 -#6137 := [quant-inst]: #6682 -#6878 := [mp #6137 #6845]: #6681 -#25328 := [unit-resolution #6878 #8574]: #6064 -#25329 := [unit-resolution #25328 #25327]: #6056 -#25356 := [mp #25329 #25355]: #173 -#4237 := (or #4568 #1492 #4562) -#4066 := [def-axiom]: #4237 -#25358 := [unit-resolution #4066 #25356]: #25357 -#25359 := [unit-resolution #25358 #25326]: #4562 -#4232 := (or #4559 #4553) -#4233 := [def-axiom]: #4232 -#25339 := [unit-resolution #4233 #25359]: #4553 -#4087 := (or #4556 #2235 #4550) -#4088 := [def-axiom]: #4087 -#25340 := [unit-resolution #4088 #25339]: #4553 -#29776 := [unit-resolution #25340 #29775]: #4550 -#4242 := (or #4547 #4541) -#4243 := [def-axiom]: #4242 -#29777 := [unit-resolution #4243 #29776]: #4541 -#25343 := (or #4544 #4538) -#12812 := (= #2249 #5856) -#12998 := (= ?x48!12 uf_16) -#10849 := (= ?x48!12 #7128) -#10847 := (uf_6 uf_15 ?x48!12) -#10848 := (= uf_8 #10847) -#10857 := (ite #10849 #5314 #10848) -#10851 := (uf_6 #7203 ?x48!12) -#10854 := (= uf_8 #10851) -#10860 := (iff #10854 #10857) -#12152 := (or #7026 #10860) -#10850 := (ite #10849 #6089 #10848) -#10852 := (= #10851 uf_8) -#10853 := (iff #10852 #10850) -#12155 := (or #7026 #10853) -#10823 := (iff #12155 #12152) -#10879 := (iff #12152 #12152) -#10880 := [rewrite]: #10879 -#10861 := (iff #10853 #10860) -#10858 := (iff #10850 #10857) -#10859 := [monotonicity #6102]: #10858 -#10855 := (iff #10852 #10854) -#10856 := [rewrite]: #10855 -#10862 := [monotonicity #10856 #10859]: #10861 -#10824 := [monotonicity #10862]: #10823 -#11111 := [trans #10824 #10880]: #10823 -#12156 := [quant-inst]: #12155 -#11091 := [mp #12156 #11111]: #12152 -#13286 := [unit-resolution #11091 #4320]: #10860 -#12615 := (= #2254 #10851) -#12608 := (= #10851 #2254) -#12613 := [monotonicity #8591]: #12608 -#12730 := [symm #12613]: #12615 -#12965 := [hypothesis]: #3403 -#3920 := (or #3398 #2255) -#4261 := [def-axiom]: #3920 -#12611 := [unit-resolution #4261 #12965]: #2255 -#13209 := [trans #12611 #12730]: #10854 -#11918 := (not #10854) -#10920 := (not #10860) -#11919 := (or #10920 #11918 #10857) -#12057 := [def-axiom]: #11919 -#13217 := [unit-resolution #12057 #13209 #13286]: #10857 -#10236 := (not #10848) -#11183 := (uf_4 uf_14 ?x48!12) -#11200 := (* -1::int #11183) -#13667 := (+ #7168 #11200) -#13668 := (>= #13667 0::int) -#13764 := (not #13668) -#12813 := (+ #2249 #5902) -#12814 := (<= #12813 0::int) -#13407 := (not #12814) -#11572 := (uf_4 uf_14 ?x49!11) -#11589 := (* -1::int #11572) -#11709 := (+ #144 #11589) -#11710 := (<= #11709 0::int) -#11467 := (uf_6 uf_15 ?x49!11) -#11468 := (= uf_8 #11467) -#12026 := (not #11468) -#11469 := (= ?x49!11 #7128) -#11477 := (ite #11469 #5314 #11468) -#12038 := (not #11477) -#11471 := (uf_6 #7203 ?x49!11) -#11474 := (= uf_8 #11471) -#11480 := (iff #11474 #11477) -#12030 := (or #7026 #11480) -#11470 := (ite #11469 #6089 #11468) -#11472 := (= #11471 uf_8) -#11473 := (iff #11472 #11470) -#12028 := (or #7026 #11473) -#12024 := (iff #12028 #12030) -#12033 := (iff #12030 #12030) -#12035 := [rewrite]: #12033 -#11481 := (iff #11473 #11480) -#11478 := (iff #11470 #11477) -#11479 := [monotonicity #6102]: #11478 -#11475 := (iff #11472 #11474) -#11476 := [rewrite]: #11475 -#11482 := [monotonicity #11476 #11479]: #11481 -#12032 := [monotonicity #11482]: #12024 -#12036 := [trans #12032 #12035]: #12024 -#12031 := [quant-inst]: #12028 -#12034 := [mp #12031 #12036]: #12030 -#13262 := [unit-resolution #12034 #4320]: #11480 -#12051 := (not #11474) -#13387 := (iff #2258 #12051) -#13353 := (iff #2257 #11474) -#12939 := (iff #11474 #2257) -#13219 := (= #11471 #2256) -#13243 := [monotonicity #8591]: #13219 -#13039 := [monotonicity #13243]: #12939 -#13377 := [symm #13039]: #13353 -#13388 := [monotonicity #13377]: #13387 -#3924 := (or #3398 #2258) -#3925 := [def-axiom]: #3924 -#12937 := [unit-resolution #3925 #12965]: #2258 -#12543 := [mp #12937 #13388]: #12051 -#12047 := (not #11480) -#12048 := (or #12047 #11474 #12038) -#12050 := [def-axiom]: #12048 -#12544 := [unit-resolution #12050 #12543 #13262]: #12038 -#12039 := (not #11469) -#12539 := (or #11477 #12039) -#12042 := (or #11477 #12039 #7040) -#12043 := [def-axiom]: #12042 -#12545 := [unit-resolution #12043 #8579]: #12539 -#12546 := [unit-resolution #12545 #12544]: #12039 -#12044 := (or #11477 #11469 #12026) -#12045 := [def-axiom]: #12044 -#12548 := [unit-resolution #12045 #12546 #12544]: #12026 -#11715 := (or #11468 #11710) -#4217 := (or #4595 #4446) -#4221 := [def-axiom]: #4217 -#12574 := [unit-resolution #4221 #7389]: #4446 -#12438 := (or #4451 #11468 #11710) -#11700 := (+ #11572 #1449) -#11701 := (>= #11700 0::int) -#11702 := (or #11468 #11701) -#12444 := (or #4451 #11702) -#12454 := (iff #12444 #12438) -#12448 := (or #4451 #11715) -#12452 := (iff #12448 #12438) -#12453 := [rewrite]: #12452 -#12450 := (iff #12444 #12448) -#11716 := (iff #11702 #11715) -#11713 := (iff #11701 #11710) -#11703 := (+ #1449 #11572) -#11706 := (>= #11703 0::int) -#11711 := (iff #11706 #11710) -#11712 := [rewrite]: #11711 -#11707 := (iff #11701 #11706) -#11704 := (= #11700 #11703) -#11705 := [rewrite]: #11704 -#11708 := [monotonicity #11705]: #11707 -#11714 := [trans #11708 #11712]: #11713 -#11717 := [monotonicity #11714]: #11716 -#12451 := [monotonicity #11717]: #12450 -#12449 := [trans #12451 #12453]: #12454 -#12447 := [quant-inst]: #12444 -#12455 := [mp #12447 #12449]: #12438 -#12575 := [unit-resolution #12455 #12574]: #11715 -#12576 := [unit-resolution #12575 #12548]: #11710 -#3926 := (not #2855) -#3927 := (or #3398 #3926) -#4263 := [def-axiom]: #3927 -#12577 := [unit-resolution #4263 #12965]: #3926 -#13397 := (not #11710) -#12612 := (or #13407 #2855 #11469 #13397) -#11605 := (uf_1 uf_16 ?x49!11) -#11606 := (uf_10 #11605) -#11631 := (+ #2853 #11606) -#11632 := (+ #144 #11631) -#12233 := (<= #11632 0::int) -#11635 := (= #11632 0::int) -#11610 := (* -1::int #11606) -#11611 := (+ uf_9 #11610) -#11612 := (<= #11611 0::int) -#12253 := (not #11612) -#11624 := (+ #11589 #11606) -#11625 := (+ #144 #11624) -#11626 := (>= #11625 0::int) -#11669 := (or #11612 #11626) -#11674 := (not #11669) -#11663 := (= #2251 #11572) -#13437 := (not #11663) -#11590 := (+ #2251 #11589) -#12252 := (>= #11590 0::int) -#13367 := (not #12252) -#13284 := [hypothesis]: #11710 -#13478 := [hypothesis]: #3926 -#13215 := [hypothesis]: #12814 -#13389 := (or #13367 #13397 #2855 #13407 #13409) -#13410 := [th-lemma]: #13389 -#13436 := [unit-resolution #13410 #13215 #13478 #13284 #12934]: #13367 -#13434 := (or #13437 #12252) -#12573 := [th-lemma]: #13434 -#13419 := [unit-resolution #12573 #13436]: #13437 -#11677 := (or #11663 #11674) -#12241 := (or #4461 #11663 #11674) -#11613 := (+ #1449 #11610) -#11614 := (+ #11572 #11613) -#11615 := (<= #11614 0::int) -#11659 := (or #11615 #11612) -#11660 := (not #11659) -#11661 := (= #11572 #2251) -#11662 := (or #11661 #11660) -#12242 := (or #4461 #11662) -#12249 := (iff #12242 #12241) -#12245 := (or #4461 #11677) -#12247 := (iff #12245 #12241) -#12248 := [rewrite]: #12247 -#12239 := (iff #12242 #12245) -#11678 := (iff #11662 #11677) -#11675 := (iff #11660 #11674) -#11672 := (iff #11659 #11669) -#11666 := (or #11626 #11612) -#11670 := (iff #11666 #11669) -#11671 := [rewrite]: #11670 -#11667 := (iff #11659 #11666) -#11629 := (iff #11615 #11626) -#11617 := (+ #11572 #11610) -#11618 := (+ #1449 #11617) -#11621 := (<= #11618 0::int) -#11627 := (iff #11621 #11626) -#11628 := [rewrite]: #11627 -#11622 := (iff #11615 #11621) -#11619 := (= #11614 #11618) -#11620 := [rewrite]: #11619 -#11623 := [monotonicity #11620]: #11622 -#11630 := [trans #11623 #11628]: #11629 -#11668 := [monotonicity #11630]: #11667 -#11673 := [trans #11668 #11671]: #11672 -#11676 := [monotonicity #11673]: #11675 -#11664 := (iff #11661 #11663) -#11665 := [rewrite]: #11664 -#11679 := [monotonicity #11665 #11676]: #11678 -#12246 := [monotonicity #11679]: #12239 -#12244 := [trans #12246 #12248]: #12249 -#12243 := [quant-inst]: #12242 -#12251 := [mp #12243 #12244]: #12241 -#13417 := [unit-resolution #12251 #8574]: #11677 -#13423 := [unit-resolution #13417 #13419]: #11674 -#12254 := (or #11669 #12253) -#12258 := [def-axiom]: #12254 -#13426 := [unit-resolution #12258 #13423]: #12253 -#12250 := (not #11626) -#12259 := (or #11669 #12250) -#12257 := [def-axiom]: #12259 -#13412 := [unit-resolution #12257 #13423]: #12250 -#11641 := (or #11612 #11626 #11635) -#12229 := (or #4469 #11612 #11626 #11635) -#11607 := (+ #11606 #2853) -#11608 := (+ #144 #11607) -#11609 := (= #11608 0::int) -#11616 := (or #11615 #11612 #11609) -#12222 := (or #4469 #11616) -#12236 := (iff #12222 #12229) -#12231 := (or #4469 #11641) -#12227 := (iff #12231 #12229) -#12235 := [rewrite]: #12227 -#12232 := (iff #12222 #12231) -#11644 := (iff #11616 #11641) -#11638 := (or #11626 #11612 #11635) -#11642 := (iff #11638 #11641) -#11643 := [rewrite]: #11642 -#11639 := (iff #11616 #11638) -#11636 := (iff #11609 #11635) -#11633 := (= #11608 #11632) -#11634 := [rewrite]: #11633 -#11637 := [monotonicity #11634]: #11636 -#11640 := [monotonicity #11630 #11637]: #11639 -#11645 := [trans #11640 #11643]: #11644 -#12234 := [monotonicity #11645]: #12232 -#12237 := [trans #12234 #12235]: #12236 -#12230 := [quant-inst]: #12222 -#12238 := [mp #12230 #12237]: #12229 -#13413 := [unit-resolution #12238 #7390]: #11641 -#13433 := [unit-resolution #13413 #13412 #13426]: #11635 -#13370 := (not #11635) -#13390 := (or #13370 #12233) -#13391 := [th-lemma]: #13390 -#13385 := [unit-resolution #13391 #13433]: #12233 -#12787 := (uf_1 #7128 ?x49!11) -#12788 := (uf_10 #12787) -#12790 := (* -1::int #12788) -#13285 := (+ #11606 #12790) -#13280 := (>= #13285 0::int) -#13216 := (= #11606 #12788) -#12644 := (= #12788 #11606) -#12645 := (= #12787 #11605) -#13418 := [monotonicity #8036]: #12645 -#12646 := [monotonicity #13418]: #12644 -#12578 := [symm #12646]: #13216 -#12641 := (not #13216) -#12647 := (or #12641 #13280) -#12643 := [th-lemma]: #12647 -#12582 := [unit-resolution #12643 #12578]: #13280 -#13068 := (<= #12788 0::int) -#13063 := (not #13068) -#12581 := [hypothesis]: #12039 -#13211 := (or #7093 #11469 #13063) -#12936 := (= #7128 ?x49!11) -#13162 := (or #12936 #13063) -#13263 := (or #7093 #13162) -#13354 := (iff #13263 #13211) -#13255 := (or #11469 #13063) -#13161 := (or #7093 #13255) -#13351 := (iff #13161 #13211) -#13352 := [rewrite]: #13351 -#13071 := (iff #13263 #13161) -#13069 := (iff #13162 #13255) -#13204 := (iff #12936 #11469) -#13265 := [rewrite]: #13204 -#13210 := [monotonicity #13265]: #13069 -#13163 := [monotonicity #13210]: #13071 -#13067 := [trans #13163 #13352]: #13354 -#13160 := [quant-inst]: #13263 -#13355 := [mp #13160 #13067]: #13211 -#12609 := [unit-resolution #13355 #4347 #12581]: #13063 -#12610 := [th-lemma #13478 #13215 #12934 #12609 #12582 #13385]: false -#12583 := [lemma #12610]: #12612 -#12780 := [unit-resolution #12583 #12577 #12546 #12576]: #13407 -#11201 := (+ #2249 #11200) -#11202 := (<= #11201 0::int) -#12087 := (or #4477 #11202) -#11190 := (+ #11183 #2250) -#11193 := (>= #11190 0::int) -#12088 := (or #4477 #11193) -#12090 := (iff #12088 #12087) -#12092 := (iff #12087 #12087) -#12093 := [rewrite]: #12092 -#11205 := (iff #11193 #11202) -#11194 := (+ #2250 #11183) -#11197 := (>= #11194 0::int) -#11203 := (iff #11197 #11202) -#11204 := [rewrite]: #11203 -#11198 := (iff #11193 #11197) -#11195 := (= #11190 #11194) -#11196 := [rewrite]: #11195 -#11199 := [monotonicity #11196]: #11198 -#11206 := [trans #11199 #11204]: #11205 -#12091 := [monotonicity #11206]: #12090 -#12095 := [trans #12091 #12093]: #12090 -#12085 := [quant-inst]: #12088 -#12097 := [mp #12085 #12095]: #12087 -#13218 := [unit-resolution #12097 #14110]: #11202 -#12617 := (not #11202) -#12729 := (not #7572) -#12728 := (or #13764 #12729 #12814 #12617 #12642) -#12733 := [th-lemma]: #12728 -#12616 := [unit-resolution #12733 #13218 #9185 #14101 #12780]: #13764 -#13844 := (or #10236 #13668) -#13842 := [hypothesis]: #13764 -#13843 := [hypothesis]: #10848 -#13801 := (or #4426 #7247 #10236 #13668) -#13669 := (or #10236 #7247 #13668) -#13802 := (or #4426 #13669) -#13788 := (iff #13802 #13801) -#13670 := (or #7247 #10236 #13668) -#13804 := (or #4426 #13670) -#13786 := (iff #13804 #13801) -#13787 := [rewrite]: #13786 -#13805 := (iff #13802 #13804) -#13665 := (iff #13669 #13670) -#13671 := [rewrite]: #13665 -#13806 := [monotonicity #13671]: #13805 -#13789 := [trans #13806 #13787]: #13788 -#13803 := [quant-inst]: #13802 -#13790 := [mp #13803 #13789]: #13801 -#13838 := [unit-resolution #13790 #9153 #9152 #13843 #13842]: false -#13845 := [lemma #13838]: #13844 -#12734 := [unit-resolution #13845 #12616]: #10236 -#11030 := (not #10857) -#11307 := (or #11030 #10849 #10848) -#11294 := [def-axiom]: #11307 -#12742 := [unit-resolution #11294 #12734 #13217]: #10849 -#12732 := [trans #12742 #8036]: #12998 -#12743 := [monotonicity #12732]: #12812 -#12969 := (not #12812) -#12967 := (or #12969 #12814) -#12973 := [th-lemma]: #12967 -#12786 := [unit-resolution #12973 #12780]: #12969 -#12771 := [unit-resolution #12786 #12743]: false -#12735 := [lemma #12771]: #3398 -#4251 := (or #4544 #3403 #4538) -#4248 := [def-axiom]: #4251 -#25338 := [unit-resolution #4248 #12735]: #25343 -#29778 := [unit-resolution #25338 #29777]: #4538 -#3968 := (or #4535 #4529) -#3969 := [def-axiom]: #3968 -#29779 := [unit-resolution #3969 #29778]: #4529 -#25346 := (or #4532 #4526) -#17148 := [hypothesis]: #3449 -#4266 := (or #3444 #2287) -#4267 := [def-axiom]: #4266 -#17149 := [unit-resolution #4267 #17148]: #2287 -#9931 := (uf_1 uf_16 ?x50!14) -#9932 := (uf_10 #9931) -#9936 := (* -1::int #9932) -#17081 := (+ #2281 #9936) -#17083 := (>= #17081 0::int) -#17080 := (= #2281 #9932) -#17180 := (= #2280 #9931) -#17179 := (= ?x51!13 uf_16) -#10336 := (= ?x51!13 #7128) -#10334 := (uf_6 uf_15 ?x51!13) -#10335 := (= uf_8 #10334) -#10367 := (not #10335) -#10193 := (uf_4 uf_14 ?x51!13) -#9874 := (uf_4 uf_14 ?x50!14) -#9915 := (* -1::int #9874) -#10384 := (+ #9915 #10193) -#10385 := (+ #2281 #10384) -#10388 := (>= #10385 0::int) -#17156 := (not #10388) -#9916 := (+ #2276 #9915) -#9917 := (<= #9916 0::int) -#16708 := (or #4477 #9917) -#9907 := (+ #9874 #2277) -#9908 := (>= #9907 0::int) -#16709 := (or #4477 #9908) -#16711 := (iff #16709 #16708) -#16713 := (iff #16708 #16708) -#16714 := [rewrite]: #16713 -#9920 := (iff #9908 #9917) -#9909 := (+ #2277 #9874) -#9912 := (>= #9909 0::int) -#9918 := (iff #9912 #9917) -#9919 := [rewrite]: #9918 -#9913 := (iff #9908 #9912) -#9910 := (= #9907 #9909) -#9911 := [rewrite]: #9910 -#9914 := [monotonicity #9911]: #9913 -#9921 := [trans #9914 #9919]: #9920 -#16712 := [monotonicity #9921]: #16711 -#16715 := [trans #16712 #16714]: #16711 -#16710 := [quant-inst]: #16709 -#16716 := [mp #16710 #16715]: #16708 -#17308 := [unit-resolution #16716 #14110]: #9917 -#3906 := (not #2882) -#4269 := (or #3444 #3906) -#4271 := [def-axiom]: #4269 -#17151 := [unit-resolution #4271 #17148]: #3906 -#10228 := (* -1::int #10193) -#10229 := (+ #2278 #10228) -#15878 := (>= #10229 0::int) -#10198 := (= #2278 #10193) -#4262 := (or #3444 #2289) -#4268 := [def-axiom]: #4262 -#17152 := [unit-resolution #4268 #17148]: #2289 -#16493 := (or #4486 #3429 #10198) -#10194 := (= #10193 #2278) -#10197 := (or #10194 #3429) -#16494 := (or #4486 #10197) -#16503 := (iff #16494 #16493) -#10204 := (or #3429 #10198) -#16498 := (or #4486 #10204) -#16501 := (iff #16498 #16493) -#16502 := [rewrite]: #16501 -#16499 := (iff #16494 #16498) -#10207 := (iff #10197 #10204) -#10201 := (or #10198 #3429) -#10205 := (iff #10201 #10204) -#10206 := [rewrite]: #10205 -#10202 := (iff #10197 #10201) -#10199 := (iff #10194 #10198) -#10200 := [rewrite]: #10199 -#10203 := [monotonicity #10200]: #10202 -#10208 := [trans #10203 #10206]: #10207 -#16500 := [monotonicity #10208]: #16499 -#16504 := [trans #16500 #16502]: #16503 -#16497 := [quant-inst]: #16494 -#16505 := [mp #16497 #16504]: #16493 -#17150 := [unit-resolution #16505 #9423 #17152]: #10198 -#17153 := (not #10198) -#17154 := (or #17153 #15878) -#17155 := [th-lemma]: #17154 -#17147 := [unit-resolution #17155 #17150]: #15878 -#17313 := (not #9917) -#17315 := (not #15878) -#17157 := (or #17156 #17315 #17313 #2882) -#17158 := [th-lemma]: #17157 -#17159 := [unit-resolution #17158 #17147 #17151 #17308]: #17156 -#17146 := (or #10367 #10388) -#16749 := (or #4417 #2286 #10367 #10388) -#10380 := (+ #10193 #9915) -#10381 := (+ #2281 #10380) -#10382 := (>= #10381 0::int) -#10383 := (or #10367 #2286 #10382) -#16750 := (or #4417 #10383) -#16757 := (iff #16750 #16749) -#10394 := (or #2286 #10367 #10388) -#16752 := (or #4417 #10394) -#16755 := (iff #16752 #16749) -#16756 := [rewrite]: #16755 -#16753 := (iff #16750 #16752) -#10397 := (iff #10383 #10394) -#10391 := (or #10367 #2286 #10388) -#10395 := (iff #10391 #10394) -#10396 := [rewrite]: #10395 -#10392 := (iff #10383 #10391) -#10389 := (iff #10382 #10388) -#10386 := (= #10381 #10385) -#10387 := [rewrite]: #10386 -#10390 := [monotonicity #10387]: #10389 -#10393 := [monotonicity #10390]: #10392 -#10398 := [trans #10393 #10396]: #10397 -#16754 := [monotonicity #10398]: #16753 -#16758 := [trans #16754 #16756]: #16757 -#16751 := [quant-inst]: #16750 -#16759 := [mp #16751 #16758]: #16749 -#17160 := [unit-resolution #16759 #8027 #17149]: #17146 -#17161 := [unit-resolution #17160 #17159]: #10367 -#10344 := (ite #10336 #5314 #10335) -#10338 := (uf_6 #7203 ?x51!13) -#10341 := (= uf_8 #10338) -#10347 := (iff #10341 #10344) -#16506 := (or #7026 #10347) -#10337 := (ite #10336 #6089 #10335) -#10339 := (= #10338 uf_8) -#10340 := (iff #10339 #10337) -#16507 := (or #7026 #10340) -#16509 := (iff #16507 #16506) -#16511 := (iff #16506 #16506) -#16512 := [rewrite]: #16511 -#10348 := (iff #10340 #10347) -#10345 := (iff #10337 #10344) -#10346 := [monotonicity #6102]: #10345 -#10342 := (iff #10339 #10341) -#10343 := [rewrite]: #10342 -#10349 := [monotonicity #10343 #10346]: #10348 -#16510 := [monotonicity #10349]: #16509 -#16513 := [trans #16510 #16512]: #16509 -#16508 := [quant-inst]: #16507 -#16514 := [mp #16508 #16513]: #16506 -#17162 := [unit-resolution #16514 #4320]: #10347 -#17171 := (= #2288 #10338) -#17163 := (= #10338 #2288) -#17164 := [monotonicity #8591]: #17163 -#17174 := [symm #17164]: #17171 -#17175 := [trans #17152 #17174]: #10341 -#16528 := (not #10341) -#16525 := (not #10347) -#16529 := (or #16525 #16528 #10344) -#16530 := [def-axiom]: #16529 -#17176 := [unit-resolution #16530 #17175 #17162]: #10344 -#16515 := (not #10344) -#16519 := (or #16515 #10336 #10335) -#16520 := [def-axiom]: #16519 -#17178 := [unit-resolution #16520 #17176 #17161]: #10336 -#17177 := [trans #17178 #8036]: #17179 -#17181 := [monotonicity #17177]: #17180 -#17182 := [monotonicity #17181]: #17080 -#17187 := (not #17080) -#17188 := (or #17187 #17083) -#17186 := [th-lemma]: #17188 -#17189 := [unit-resolution #17186 #17182]: #17083 -#9937 := (+ uf_9 #9936) -#9938 := (<= #9937 0::int) -#9950 := (+ #9915 #9932) -#9951 := (+ #144 #9950) -#9952 := (>= #9951 0::int) -#16744 := (not #9952) -#10475 := (uf_2 #2280) -#11002 := (uf_4 uf_14 #10475) -#11016 := (* -1::int #11002) -#16798 := (+ #10193 #11016) -#16800 := (>= #16798 0::int) -#16797 := (= #10193 #11002) -#10476 := (= ?x51!13 #10475) -#16793 := (or #7136 #10476) -#16794 := [quant-inst]: #16793 -#17292 := [unit-resolution #16794 #4306]: #10476 -#17295 := [monotonicity #17292]: #16797 -#17296 := (not #16797) -#17297 := (or #17296 #16800) -#17298 := [th-lemma]: #17297 -#17299 := [unit-resolution #17298 #17295]: #16800 -#11017 := (+ #144 #11016) -#11018 := (<= #11017 0::int) -#17065 := (= #144 #11002) -#17202 := (= #11002 #144) -#17194 := (= #10475 uf_16) -#17192 := (= #10475 #7128) -#17190 := (= #10475 ?x51!13) -#17191 := [symm #17292]: #17190 -#17193 := [trans #17191 #17178]: #17192 -#17195 := [trans #17193 #8036]: #17194 -#17203 := [monotonicity #17195]: #17202 -#17204 := [symm #17203]: #17065 -#17205 := (not #17065) -#17206 := (or #17205 #11018) -#17201 := [th-lemma]: #17206 -#17207 := [unit-resolution #17201 #17204]: #11018 -#17316 := (not #11018) -#17314 := (not #16800) -#17208 := (not #17083) -#17209 := (or #16744 #17208 #17313 #2882 #17314 #17315 #17316) -#17210 := [th-lemma]: #17209 -#17211 := [unit-resolution #17210 #17207 #17308 #17147 #17151 #17299 #17189]: #16744 -#9957 := (+ #2277 #9932) -#9958 := (+ #144 #9957) -#9961 := (= #9958 0::int) -#17223 := (not #9961) -#16729 := (>= #9958 0::int) -#17219 := (not #16729) -#17220 := (or #17219 #17208 #2882 #17314 #17315 #17316) -#17221 := [th-lemma]: #17220 -#17222 := [unit-resolution #17221 #17207 #17147 #17151 #17299 #17189]: #17219 -#17218 := (or #17223 #16729) -#17224 := [th-lemma]: #17218 -#17225 := [unit-resolution #17224 #17222]: #17223 -#9967 := (or #9938 #9952 #9961) -#16717 := (or #4469 #9938 #9952 #9961) -#9933 := (+ #9932 #2277) -#9934 := (+ #144 #9933) -#9935 := (= #9934 0::int) -#9939 := (+ #1449 #9936) -#9940 := (+ #9874 #9939) -#9941 := (<= #9940 0::int) -#9942 := (or #9941 #9938 #9935) -#16718 := (or #4469 #9942) -#16725 := (iff #16718 #16717) -#16720 := (or #4469 #9967) -#16723 := (iff #16720 #16717) -#16724 := [rewrite]: #16723 -#16721 := (iff #16718 #16720) -#9970 := (iff #9942 #9967) -#9964 := (or #9952 #9938 #9961) -#9968 := (iff #9964 #9967) -#9969 := [rewrite]: #9968 -#9965 := (iff #9942 #9964) -#9962 := (iff #9935 #9961) -#9959 := (= #9934 #9958) -#9960 := [rewrite]: #9959 -#9963 := [monotonicity #9960]: #9962 -#9955 := (iff #9941 #9952) -#9943 := (+ #9874 #9936) -#9944 := (+ #1449 #9943) -#9947 := (<= #9944 0::int) -#9953 := (iff #9947 #9952) -#9954 := [rewrite]: #9953 -#9948 := (iff #9941 #9947) -#9945 := (= #9940 #9944) -#9946 := [rewrite]: #9945 -#9949 := [monotonicity #9946]: #9948 -#9956 := [trans #9949 #9954]: #9955 -#9966 := [monotonicity #9956 #9963]: #9965 -#9971 := [trans #9966 #9969]: #9970 -#16722 := [monotonicity #9971]: #16721 -#16726 := [trans #16722 #16724]: #16725 -#16719 := [quant-inst]: #16718 -#16727 := [mp #16719 #16726]: #16717 -#17226 := [unit-resolution #16727 #7390]: #9967 -#17227 := [unit-resolution #17226 #17225 #17211]: #9938 -#17228 := [th-lemma #17227 #17189 #17149]: false -#17253 := [lemma #17228]: #3444 -#4253 := (or #4532 #3449 #4526) -#4257 := [def-axiom]: #4253 -#25347 := [unit-resolution #4257 #17253]: #25346 -#29780 := [unit-resolution #25347 #29779]: #4526 -#3983 := (or #4523 #4515) -#3984 := [def-axiom]: #3983 -#29781 := [unit-resolution #3984 #29780]: #4515 -#20547 := (or #4520 #7166 #15518 #15538) -#15505 := (+ #2307 #15504) -#15506 := (+ #7167 #15505) -#15507 := (= #15506 0::int) -#15508 := (not #15507) -#15509 := (+ #7167 #2307) -#15510 := (>= #15509 0::int) -#15511 := (or #7166 #15510 #15508) -#17626 := (or #4520 #15511) -#20578 := (iff #17626 #20547) -#20549 := (or #4520 #15541) -#20516 := (iff #20549 #20547) -#20517 := [rewrite]: #20516 -#20386 := (iff #17626 #20549) -#15542 := (iff #15511 #15541) -#15539 := (iff #15508 #15538) -#15536 := (iff #15507 #15533) -#15523 := (+ #7167 #15504) -#15524 := (+ #2307 #15523) -#15527 := (= #15524 0::int) -#15534 := (iff #15527 #15533) -#15535 := [rewrite]: #15534 -#15528 := (iff #15507 #15527) -#15525 := (= #15506 #15524) -#15526 := [rewrite]: #15525 -#15529 := [monotonicity #15526]: #15528 -#15537 := [trans #15529 #15535]: #15536 -#15540 := [monotonicity #15537]: #15539 -#15521 := (iff #15510 #15518) -#15512 := (+ #2307 #7167) -#15515 := (>= #15512 0::int) -#15519 := (iff #15515 #15518) -#15520 := [rewrite]: #15519 -#15516 := (iff #15510 #15515) -#15513 := (= #15509 #15512) -#15514 := [rewrite]: #15513 -#15517 := [monotonicity #15514]: #15516 -#15522 := [trans #15517 #15520]: #15521 -#15543 := [monotonicity #15522 #15540]: #15542 -#20388 := [monotonicity #15543]: #20386 -#20485 := [trans #20388 #20517]: #20578 -#20612 := [quant-inst]: #17626 -#20795 := [mp #20612 #20485]: #20547 -#29782 := [unit-resolution #20795 #29781]: #15541 -#29783 := [unit-resolution #29782 #29774 #25367]: #15518 -#28771 := (<= #28711 0::int) -#28772 := (not #28771) -#28773 := (= #7128 #27533) -#29792 := (not #28773) -#28056 := (not #18779) -#29793 := (iff #28056 #29792) -#29790 := (iff #18779 #28773) -#29788 := (iff #28773 #18779) -#29785 := (iff #28773 #25219) -#29786 := [monotonicity #29764]: #29785 -#29789 := [trans #29786 #29787]: #29788 -#29791 := [symm #29789]: #29790 -#29794 := [monotonicity #29791]: #29793 -#29784 := [hypothesis]: #28056 -#29795 := [mp #29784 #29794]: #29792 -#28775 := (or #28772 #28773) -#29743 := (or #7093 #28772 #28773) -#28774 := (or #28773 #28772) -#29744 := (or #7093 #28774) -#29751 := (iff #29744 #29743) -#29746 := (or #7093 #28775) -#29749 := (iff #29746 #29743) -#29750 := [rewrite]: #29749 -#29747 := (iff #29744 #29746) -#28776 := (iff #28774 #28775) -#28777 := [rewrite]: #28776 -#29748 := [monotonicity #28777]: #29747 -#29752 := [trans #29748 #29750]: #29751 -#29745 := [quant-inst]: #29744 -#29753 := [mp #29745 #29752]: #29743 -#29796 := [unit-resolution #29753 #4347]: #28775 -#29797 := [unit-resolution #29796 #29795]: #28772 -#29798 := [th-lemma #29797 #29783 #25304 #25235 #12934 #29773]: false -#29800 := [lemma #29798]: #29799 -#32973 := [unit-resolution #29800 #32972]: #18779 -#32992 := [mp #32973 #32991]: #32602 -#32632 := (not #32602) -#32993 := (or #32591 #32632) -#29702 := (or #32591 #32632 #7040) -#30315 := [def-axiom]: #29702 -#32995 := [unit-resolution #30315 #8579]: #32993 -#32996 := [unit-resolution #32995 #32992]: #32591 -#32599 := (not #32591) -#28920 := (not #32564) -#30916 := (or #28920 #32538 #32599) -#31002 := [def-axiom]: #30916 -#32987 := [unit-resolution #31002 #32996 #32981 #32954]: false -#32997 := [lemma #32987]: #15367 -#39375 := (or #32956 #14218) -#39376 := [th-lemma]: #39375 -#39377 := [unit-resolution #39376 #32997]: #14218 -#34351 := (not #14218) -#34357 := (or #34324 #34351) -#4272 := (or #4523 #2319) -#4270 := [def-axiom]: #4272 -#34292 := [unit-resolution #4270 #29780]: #2319 -#34293 := [hypothesis]: #14218 -#34291 := [hypothesis]: #16010 -#34294 := [th-lemma #34291 #34293 #34292]: false -#34641 := [lemma #34294]: #34357 -#39378 := [unit-resolution #34641 #39377]: #34324 -#39380 := (or #16010 #16030) -#4273 := (or #4523 #2896) -#4259 := [def-axiom]: #4273 -#39379 := [unit-resolution #4259 #29780]: #2896 -#17785 := (or #4442 #2893 #16010 #16030) -#15998 := (+ #15997 #15995) -#15999 := (+ #15362 #15998) -#16000 := (= #15999 0::int) -#16001 := (not #16000) -#16007 := (or #16006 #16003 #16001) -#16008 := (not #16007) -#16011 := (or #2320 #16010 #16008) -#20897 := (or #4442 #16011) -#21453 := (iff #20897 #17785) -#16033 := (or #2893 #16010 #16030) -#21128 := (or #4442 #16033) -#18084 := (iff #21128 #17785) -#22003 := [rewrite]: #18084 -#20788 := (iff #20897 #21128) -#16034 := (iff #16011 #16033) -#16031 := (iff #16008 #16030) -#16028 := (iff #16007 #16025) -#16022 := (or #16006 #16003 #16019) -#16026 := (iff #16022 #16025) -#16027 := [rewrite]: #16026 -#16023 := (iff #16007 #16022) -#16020 := (iff #16001 #16019) -#16017 := (iff #16000 #16016) -#16014 := (= #15999 #16013) -#16015 := [rewrite]: #16014 -#16018 := [monotonicity #16015]: #16017 -#16021 := [monotonicity #16018]: #16020 -#16024 := [monotonicity #16021]: #16023 -#16029 := [trans #16024 #16027]: #16028 -#16032 := [monotonicity #16029]: #16031 -#16035 := [monotonicity #2895 #16032]: #16034 -#21237 := [monotonicity #16035]: #20788 -#21330 := [trans #21237 #22003]: #21453 -#21234 := [quant-inst]: #20897 -#21506 := [mp #21234 #21330]: #17785 -#39381 := [unit-resolution #21506 #8030 #39379]: #39380 -#39382 := [unit-resolution #39381 #39378]: #16030 -#22030 := (or #16025 #16016) -#18058 := [def-axiom]: #22030 -#32365 := [unit-resolution #18058 #39382]: #16016 -#32368 := (or #16019 #22006) -#29568 := [th-lemma]: #32368 -#29612 := [unit-resolution #29568 #32365]: #22006 -#20411 := (+ uf_9 #15995) -#20412 := (<= #20411 0::int) -#20215 := (uf_6 uf_17 #15992) -#20216 := (= uf_8 #20215) -#20606 := (uf_2 #15993) -#39327 := (uf_6 #7203 #20606) -#38901 := (= #39327 #20215) -#38905 := (= #20215 #39327) -#20607 := (= #15992 #20606) -#25402 := (or #7136 #20607) -#25393 := [quant-inst]: #25402 -#39384 := [unit-resolution #25393 #4306]: #20607 -#8247 := (= uf_17 #7203) -#8291 := (= #150 #7203) -#8151 := [symm #8593]: #8291 -#8423 := [trans #8578 #8151]: #8247 -#38906 := [monotonicity #8423 #39384]: #38905 -#38907 := [symm #38906]: #38901 -#39330 := (= uf_8 #39327) -#21345 := (uf_6 uf_15 #20606) -#21346 := (= uf_8 #21345) -#39333 := (= #7128 #20606) -#39336 := (ite #39333 #5314 #21346) -#39339 := (iff #39330 #39336) -#38867 := (or #7026 #39339) -#39325 := (= #20606 #7128) -#39326 := (ite #39325 #6089 #21346) -#39328 := (= #39327 uf_8) -#39329 := (iff #39328 #39326) -#38877 := (or #7026 #39329) -#38879 := (iff #38877 #38867) -#38890 := (iff #38867 #38867) -#38888 := [rewrite]: #38890 -#39340 := (iff #39329 #39339) -#39337 := (iff #39326 #39336) -#39334 := (iff #39325 #39333) -#39335 := [rewrite]: #39334 -#39338 := [monotonicity #39335 #6102]: #39337 -#39331 := (iff #39328 #39330) -#39332 := [rewrite]: #39331 -#39341 := [monotonicity #39332 #39338]: #39340 -#38889 := [monotonicity #39341]: #38879 -#38894 := [trans #38889 #38888]: #38879 -#38878 := [quant-inst]: #38877 -#38893 := [mp #38878 #38894]: #38867 -#38919 := [unit-resolution #38893 #4320]: #39339 -#38895 := (not #39339) -#38922 := (or #38895 #39330) -#39351 := (not #39336) -#39371 := [hypothesis]: #39351 -#39352 := (not #39333) -#39372 := (or #39336 #39352) -#39357 := (or #39336 #39352 #7040) -#39358 := [def-axiom]: #39357 -#39373 := [unit-resolution #39358 #8579]: #39372 -#39374 := [unit-resolution #39373 #39371]: #39352 -#39394 := (or #39336 #39333) -#39391 := (= #16004 #21345) -#39387 := (= #21345 #16004) -#39385 := (= #20606 #15992) -#39386 := [symm #39384]: #39385 -#39388 := [monotonicity #39386]: #39387 -#39392 := [symm #39388]: #39391 -#21698 := (or #16025 #16005) -#21997 := [def-axiom]: #21698 -#39383 := [unit-resolution #21997 #39382]: #16005 -#39393 := [trans #39383 #39392]: #21346 -#21347 := (not #21346) -#39359 := (or #39336 #39333 #21347) -#39360 := [def-axiom]: #39359 -#39395 := [unit-resolution #39360 #39393]: #39394 -#39396 := [unit-resolution #39395 #39374 #39371]: false -#39397 := [lemma #39396]: #39336 -#38896 := (or #38895 #39330 #39351) -#38897 := [def-axiom]: #38896 -#38902 := [unit-resolution #38897 #39397]: #38922 -#38903 := [unit-resolution #38902 #38919]: #39330 -#38908 := [trans #38903 #38907]: #20216 -#20217 := (not #20216) -#38921 := [hypothesis]: #20217 -#38924 := [unit-resolution #38921 #38908]: false -#38927 := [lemma #38924]: #20216 -#20218 := (uf_18 #15992) -#20235 := (* -1::int #20218) -#20421 := (+ #15995 #20235) -#20422 := (+ #2306 #20421) -#20423 := (<= #20422 0::int) -#31764 := (not #20423) -#26473 := (>= #20422 0::int) -#20236 := (+ #15996 #20235) -#20237 := (>= #20236 0::int) -#25927 := (or #4477 #20237) -#26030 := [quant-inst]: #25927 -#27294 := [unit-resolution #26030 #14110]: #20237 -#32338 := (not #20237) -#29920 := (not #22006) -#29919 := (or #26473 #29920 #34351 #32338) -#32347 := [th-lemma]: #29919 -#32336 := [unit-resolution #32347 #29612 #27294 #39377]: #26473 -#20465 := (= #20422 0::int) -#20470 := (not #20465) -#20454 := (+ #2306 #20235) -#20455 := (<= #20454 0::int) -#32331 := (not #20455) -#22027 := (not #16003) -#21587 := (or #16025 #22027) -#21112 := [def-axiom]: #21587 -#32344 := [unit-resolution #21112 #39382]: #22027 -#32343 := (or #32331 #16003 #34351 #32338) -#32335 := [th-lemma]: #32343 -#32375 := [unit-resolution #32335 #32344 #27294 #39377]: #32331 -#20473 := (or #20217 #20455 #20470) -#25436 := (or #4520 #20217 #20455 #20470) -#20442 := (+ #2307 #15994) -#20443 := (+ #20218 #20442) -#20444 := (= #20443 0::int) -#20445 := (not #20444) -#20406 := (+ #20218 #2307) -#20446 := (>= #20406 0::int) -#20447 := (or #20217 #20446 #20445) -#26505 := (or #4520 #20447) -#26705 := (iff #26505 #25436) -#11899 := (or #4520 #20473) -#26697 := (iff #11899 #25436) -#26704 := [rewrite]: #26697 -#25435 := (iff #26505 #11899) -#20474 := (iff #20447 #20473) -#20471 := (iff #20445 #20470) -#20468 := (iff #20444 #20465) -#20414 := (+ #15994 #20218) -#20415 := (+ #2307 #20414) -#20462 := (= #20415 0::int) -#20466 := (iff #20462 #20465) -#20467 := [rewrite]: #20466 -#20463 := (iff #20444 #20462) -#20460 := (= #20443 #20415) -#20461 := [rewrite]: #20460 -#20464 := [monotonicity #20461]: #20463 -#20469 := [trans #20464 #20467]: #20468 -#20472 := [monotonicity #20469]: #20471 -#20458 := (iff #20446 #20455) -#20448 := (+ #2307 #20218) -#20451 := (>= #20448 0::int) -#20456 := (iff #20451 #20455) -#20457 := [rewrite]: #20456 -#20452 := (iff #20446 #20451) -#20449 := (= #20406 #20448) -#20450 := [rewrite]: #20449 -#20453 := [monotonicity #20450]: #20452 -#20459 := [trans #20453 #20457]: #20458 -#20475 := [monotonicity #20459 #20472]: #20474 -#26696 := [monotonicity #20475]: #25435 -#26698 := [trans #26696 #26704]: #26705 -#26504 := [quant-inst]: #26505 -#26615 := [mp #26504 #26698]: #25436 -#29644 := [unit-resolution #26615 #29781]: #20473 -#29611 := [unit-resolution #29644 #38927 #32375]: #20470 -#32349 := (not #26473) -#31762 := (or #20465 #31764 #32349) -#29679 := [th-lemma]: #31762 -#32366 := [unit-resolution #29679 #29611 #32336]: #31764 -#20428 := (or #20217 #20412 #20423) -#4260 := (or #4523 #4506) -#3982 := [def-axiom]: #4260 -#29894 := [unit-resolution #3982 #29780]: #4506 -#26338 := (or #4511 #20217 #20412 #20423) -#20407 := (+ #15994 #20406) -#20410 := (>= #20407 0::int) -#20413 := (or #20217 #20412 #20410) -#26340 := (or #4511 #20413) -#25469 := (iff #26340 #26338) -#10685 := (or #4511 #20428) -#25478 := (iff #10685 #26338) -#25474 := [rewrite]: #25478 -#25437 := (iff #26340 #10685) -#20429 := (iff #20413 #20428) -#20426 := (iff #20410 #20423) -#20418 := (>= #20415 0::int) -#20424 := (iff #20418 #20423) -#20425 := [rewrite]: #20424 -#20419 := (iff #20410 #20418) -#20416 := (= #20407 #20415) -#20417 := [rewrite]: #20416 -#20420 := [monotonicity #20417]: #20419 -#20427 := [trans #20420 #20425]: #20426 -#20430 := [monotonicity #20427]: #20429 -#25388 := [monotonicity #20430]: #25437 -#25466 := [trans #25388 #25474]: #25469 -#25477 := [quant-inst]: #26340 -#25432 := [mp #25477 #25466]: #26338 -#31826 := [unit-resolution #25432 #29894]: #20428 -#29616 := [unit-resolution #31826 #32366 #38927]: #20412 -[th-lemma #34292 #39377 #29616 #29612 #29563]: false +#1856 := [mp~ #401 #1855]: #396 +#4243 := [mp #1856 #4242]: #4238 +#6331 := (not #4238) +#11098 := (or #6331 #10318 #10346) +#10424 := (not #10406) +#10425 := (or #10424 #10318) +#11207 := (or #6331 #10425) +#11222 := (iff #11207 #11098) +#11209 := (or #6331 #10431) +#11220 := (iff #11209 #11098) +#11221 := [rewrite]: #11220 +#11210 := (iff #11207 #11209) +#10434 := (iff #10425 #10431) +#10428 := (or #10346 #10318) +#10432 := (iff #10428 #10431) +#10433 := [rewrite]: #10432 +#10429 := (iff #10425 #10428) +#10426 := (iff #10424 #10346) +#10427 := [monotonicity #10409]: #10426 +#10430 := [monotonicity #10427]: #10429 +#10435 := [trans #10430 #10433]: #10434 +#11211 := [monotonicity #10435]: #11210 +#11223 := [trans #11211 #11221]: #11222 +#11208 := [quant-inst]: #11207 +#11224 := [mp #11208 #11223]: #11098 +#11992 := [unit-resolution #11224 #4243]: #10431 +#11993 := [unit-resolution #11992 #11991 #11986]: false +#11994 := [lemma #11993]: #11219 +#10057 := (+ #2261 #11133) +#11955 := (<= #10057 0::int) +#13994 := (not #12385) +#13995 := (or #13994 #11955) +#13996 := [th-lemma]: #13995 +#13997 := [unit-resolution #13996 #13203]: #11955 +#9707 := (* -1::int #9701) +#10565 := (+ #188 #9707) +#10581 := (>= #10565 0::int) +#9923 := (= #188 #9701) +#12919 := (= #9701 #188) +#12920 := [monotonicity #10708]: #12919 +#12921 := [symm #12920]: #9923 +#12922 := (not #9923) +#12923 := (or #12922 #10581) +#12924 := [th-lemma]: #12923 +#12925 := [unit-resolution #12924 #12921]: #10581 +#13998 := [th-lemma #12925 #13997 #11994 #13993 #13980 #13979]: false +#14001 := [lemma #13998]: #14000 +#13600 := [unit-resolution #14001 #13599]: #13999 +#13971 := (or #11138 #11135) +#13935 := [hypothesis]: #13999 +#13936 := [hypothesis]: #11137 +#9685 := (uf_6 uf_15 #9695) +#9686 := (= uf_8 #9685) +#13964 := (not #9686) +#13965 := (iff #731 #13964) +#13957 := (iff #728 #9686) +#13955 := (iff #9686 #728) +#13948 := (= #9685 #185) +#13954 := [monotonicity #10708]: #13948 +#13956 := [monotonicity #13954]: #13955 +#13958 := [symm #13956]: #13957 +#13968 := [monotonicity #13958]: #13965 +#4070 := (or #4567 #731) +#4065 := [def-axiom]: #4070 +#13953 := [unit-resolution #4065 #10726]: #731 +#13969 := [mp #13953 #13968]: #13964 +#3978 := (or #4579 #4323) +#4033 := [def-axiom]: #3978 +#13967 := [unit-resolution #4033 #5496]: #4323 +#13906 := (or #4328 #9686 #11135 #11138) +#11139 := (or #9686 #11138 #11135) +#13907 := (or #4328 #11139) +#13933 := (iff #13907 #13906) +#11140 := (or #9686 #11135 #11138) +#13912 := (or #4328 #11140) +#13931 := (iff #13912 #13906) +#13932 := [rewrite]: #13931 +#13929 := (iff #13907 #13912) +#11141 := (iff #11139 #11140) +#11142 := [rewrite]: #11141 +#13930 := [monotonicity #11142]: #13929 +#13928 := [trans #13930 #13932]: #13933 +#13911 := [quant-inst]: #13907 +#13934 := [mp #13911 #13928]: #13906 +#13970 := [unit-resolution #13934 #13967 #13969 #13936 #13935]: false +#13972 := [lemma #13970]: #13971 +#13598 := [unit-resolution #13972 #13600]: #11138 +#13978 := (or #13335 #11137) +#13974 := (iff #9519 #11137) +#13973 := (iff #11137 #9519) +#13736 := (= #11136 #9518) +#13693 := (= #10448 ?x63!14) +#13735 := [symm #13195]: #13693 +#13905 := [monotonicity #13735]: #13736 +#13966 := [monotonicity #13905]: #13973 +#13975 := [symm #13966]: #13974 +#13619 := [hypothesis]: #9519 +#13976 := [mp #13619 #13975]: #11137 +#13625 := [hypothesis]: #11138 +#13977 := [unit-resolution #13625 #13976]: false +#13982 := [lemma #13977]: #13978 +#13601 := [unit-resolution #13982 #13598]: #13335 +#13052 := (not #10330) +#13244 := (or #13052 #10319 #9519) +#13279 := [def-axiom]: #13244 +#13602 := [unit-resolution #13279 #13601 #13593]: #10319 +#13607 := [trans #13735 #13602]: #13609 +#13611 := [trans #13607 #10708]: #13610 +#13622 := [monotonicity #13611]: #13621 +#13629 := [symm #13622]: #13623 +#13634 := (= #2260 #188) +#4740 := (uf_24 uf_22) +#10619 := (= #4740 #188) +#4741 := (= #188 #4740) +#4729 := (uf_10 #4728) +#4748 := (>= #4729 0::int) +#4732 := (* -1::int #4729) +#4736 := (+ uf_9 #4732) +#4737 := (<= #4736 0::int) +#4753 := (or #4737 #4748) +#9615 := (= #4729 0::int) +#9682 := (or #6331 #9615) +#4959 := (= uf_22 uf_22) +#9598 := (not #4959) +#9599 := (or #9598 #9615) +#9683 := (or #6331 #9599) +#9749 := (iff #9683 #9682) +#9751 := (iff #9682 #9682) +#9752 := [rewrite]: #9751 +#9631 := (iff #9599 #9615) +#9603 := (or false #9615) +#9606 := (iff #9603 #9615) +#9607 := [rewrite]: #9606 +#9604 := (iff #9599 #9603) +#9602 := (iff #9598 false) +#9600 := (iff #9598 #8605) +#4968 := (iff #4959 true) +#4969 := [rewrite]: #4968 +#9601 := [monotonicity #4969]: #9600 +#9597 := [trans #9601 #8609]: #9602 +#9605 := [monotonicity #9597]: #9604 +#9632 := [trans #9605 #9607]: #9631 +#9750 := [monotonicity #9632]: #9749 +#9747 := [trans #9750 #9752]: #9749 +#9748 := [quant-inst]: #9683 +#9724 := [mp #9748 #9747]: #9682 +#10718 := [unit-resolution #9724 #4243]: #9615 +#10719 := (not #9615) +#10720 := (or #10719 #4748) +#10721 := [th-lemma]: #10720 +#10722 := [unit-resolution #10721 #10718]: #4748 +#9223 := (not #4748) +#9224 := (or #4753 #9223) +#9225 := [def-axiom]: #9224 +#10723 := [unit-resolution #9225 #10722]: #4753 +#4756 := (not #4753) +#4759 := (or #4741 #4756) +#7499 := (or #4433 #4741 #4756) +#4733 := (+ #1455 #4732) +#4734 := (+ #188 #4733) +#4735 := (<= #4734 0::int) +#4738 := (or #4737 #4735) +#4739 := (not #4738) +#4742 := (or #4741 #4739) +#7438 := (or #4433 #4742) +#9202 := (iff #7438 #7499) +#9040 := (or #4433 #4759) +#9159 := (iff #9040 #7499) +#9162 := [rewrite]: #9159 +#9149 := (iff #7438 #9040) +#4760 := (iff #4742 #4759) +#4757 := (iff #4739 #4756) +#4754 := (iff #4738 #4753) +#4751 := (iff #4735 #4748) +#4745 := (<= #4732 0::int) +#4749 := (iff #4745 #4748) +#4750 := [rewrite]: #4749 +#4746 := (iff #4735 #4745) +#4743 := (= #4734 #4732) +#4744 := [rewrite]: #4743 +#4747 := [monotonicity #4744]: #4746 +#4752 := [trans #4747 #4750]: #4751 +#4755 := [monotonicity #4752]: #4754 +#4758 := [monotonicity #4755]: #4757 +#4761 := [monotonicity #4758]: #4760 +#8886 := [monotonicity #4761]: #9149 +#9203 := [trans #8886 #9162]: #9202 +#9039 := [quant-inst]: #7438 +#9204 := [mp #9039 #9203]: #7499 +#10728 := [unit-resolution #9204 #10727]: #4759 +#10729 := [unit-resolution #10728 #10723]: #4741 +#13620 := [symm #10729]: #10619 +#13608 := (= #2260 #4740) +#9704 := (= ?x63!14 uf_22) +#13603 := [trans #13602 #10708]: #9704 +#13612 := [monotonicity #13603]: #13608 +#13641 := [trans #13612 #13620]: #13634 +#13633 := [trans #13641 #13629]: #13222 +#13644 := [trans #13633 #13213]: #2866 +#13646 := [unit-resolution #13188 #13644]: false +#13639 := [lemma #13646]: #2872 +#10564 := [unit-resolution #13639 #10294 #10559]: false +#10587 := [lemma #10564]: #2872 +#4036 := (or #4567 #4561) +#4037 := [def-axiom]: #4036 +#10784 := [unit-resolution #4037 #10726]: #4561 +#4062 := (or #4567 #4436) +#4035 := [def-axiom]: #4062 +#10785 := [unit-resolution #4035 #10726]: #4436 +#9537 := (or #2858 #4441 #4433) +#9339 := (uf_1 uf_22 ?x61!13) +#9340 := (uf_10 #9339) +#9365 := (+ #2240 #9340) +#9366 := (+ #188 #9365) +#9387 := (>= #9366 0::int) +#9369 := (= #9366 0::int) +#9344 := (* -1::int #9340) +#9348 := (+ uf_9 #9344) +#9349 := (<= #9348 0::int) +#9416 := (not #9349) +#9358 := (+ #2856 #9340) +#9359 := (+ #188 #9358) +#9360 := (>= #9359 0::int) +#9395 := (or #9349 #9360) +#9398 := (not #9395) +#9392 := (= #2239 #2241) +#9517 := (not #9392) +#9516 := [hypothesis]: #2863 +#9520 := (or #9517 #2858) +#9521 := [th-lemma]: #9520 +#9522 := [unit-resolution #9521 #9516]: #9517 +#9523 := [hypothesis]: #4428 +#9404 := (or #4433 #9392 #9398) +#9345 := (+ #1455 #9344) +#9346 := (+ #2241 #9345) +#9347 := (<= #9346 0::int) +#9388 := (or #9349 #9347) +#9389 := (not #9388) +#9390 := (= #2241 #2239) +#9391 := (or #9390 #9389) +#9405 := (or #4433 #9391) +#9412 := (iff #9405 #9404) +#9401 := (or #9392 #9398) +#9407 := (or #4433 #9401) +#9410 := (iff #9407 #9404) +#9411 := [rewrite]: #9410 +#9408 := (iff #9405 #9407) +#9402 := (iff #9391 #9401) +#9399 := (iff #9389 #9398) +#9396 := (iff #9388 #9395) +#9363 := (iff #9347 #9360) +#9351 := (+ #2241 #9344) +#9352 := (+ #1455 #9351) +#9355 := (<= #9352 0::int) +#9361 := (iff #9355 #9360) +#9362 := [rewrite]: #9361 +#9356 := (iff #9347 #9355) +#9353 := (= #9346 #9352) +#9354 := [rewrite]: #9353 +#9357 := [monotonicity #9354]: #9356 +#9364 := [trans #9357 #9362]: #9363 +#9397 := [monotonicity #9364]: #9396 +#9400 := [monotonicity #9397]: #9399 +#9393 := (iff #9390 #9392) +#9394 := [rewrite]: #9393 +#9403 := [monotonicity #9394 #9400]: #9402 +#9409 := [monotonicity #9403]: #9408 +#9413 := [trans #9409 #9411]: #9412 +#9406 := [quant-inst]: #9405 +#9414 := [mp #9406 #9413]: #9404 +#9524 := [unit-resolution #9414 #9523 #9522]: #9398 +#9417 := (or #9395 #9416) +#9418 := [def-axiom]: #9417 +#9525 := [unit-resolution #9418 #9524]: #9416 +#9419 := (not #9360) +#9420 := (or #9395 #9419) +#9421 := [def-axiom]: #9420 +#9526 := [unit-resolution #9421 #9524]: #9419 +#9372 := (or #9349 #9360 #9369) +#7593 := [hypothesis]: #4436 +#9375 := (or #4441 #9349 #9360 #9369) +#9341 := (+ #9340 #2240) +#9342 := (+ #188 #9341) +#9343 := (= #9342 0::int) +#9350 := (or #9349 #9347 #9343) +#9376 := (or #4441 #9350) +#9383 := (iff #9376 #9375) +#9378 := (or #4441 #9372) +#9381 := (iff #9378 #9375) +#9382 := [rewrite]: #9381 +#9379 := (iff #9376 #9378) +#9373 := (iff #9350 #9372) +#9370 := (iff #9343 #9369) +#9367 := (= #9342 #9366) +#9368 := [rewrite]: #9367 +#9371 := [monotonicity #9368]: #9370 +#9374 := [monotonicity #9364 #9371]: #9373 +#9380 := [monotonicity #9374]: #9379 +#9384 := [trans #9380 #9382]: #9383 +#9377 := [quant-inst]: #9376 +#9385 := [mp #9377 #9384]: #9375 +#9527 := [unit-resolution #9385 #7593]: #9372 +#9528 := [unit-resolution #9527 #9526 #9525]: #9369 +#9529 := (not #9369) +#9530 := (or #9529 #9387) +#9531 := [th-lemma]: #9530 +#9532 := [unit-resolution #9531 #9528]: #9387 +#9415 := (>= #2857 0::int) +#9533 := (or #9415 #2858) +#9534 := [th-lemma]: #9533 +#9535 := [unit-resolution #9534 #9516]: #9415 +#9536 := [th-lemma #9535 #9526 #9532]: false +#9538 := [lemma #9536]: #9537 +#10786 := [unit-resolution #9538 #10785 #10727]: #2858 +#4066 := (or #4564 #2863 #4558) +#4067 := [def-axiom]: #4066 +#10787 := [unit-resolution #4067 #10786 #10784]: #4558 +#4081 := (or #4555 #4549) +#4082 := [def-axiom]: #4081 +#25696 := [unit-resolution #4082 #10787]: #4549 +#4077 := (or #4552 #2877 #4546) +#4078 := [def-axiom]: #4077 +#25697 := [unit-resolution #4078 #25696]: #4549 +#25698 := [unit-resolution #25697 #10587]: #4546 +#4087 := (or #4543 #4453) +#4089 := [def-axiom]: #4087 +#25699 := [unit-resolution #4089 #25698]: #4453 +#17073 := (or #3494 #2335 #4458) +#4079 := (or #4555 #4444) +#4080 := [def-axiom]: #4079 +#10788 := [unit-resolution #4080 #10787]: #4444 +#4071 := (or #4567 #4418) +#4057 := [def-axiom]: #4071 +#17666 := [unit-resolution #4057 #10726]: #4418 +#17049 := (or #3494 #2335 #4441 #4423 #981 #4449 #4458) +#7717 := (or #3494 #4319 #2335 #4441 #4423 #981 #4449 #4458) +#6095 := (uf_4 uf_14 ?x72!18) +#6194 := (* -1::int #6095) +#6195 := (+ #2327 #6194) +#7308 := (>= #6195 0::int) +#6100 := (= #2327 #6095) +#4138 := (or #3494 #2338) +#4132 := [def-axiom]: #4138 +#7656 := [unit-resolution #4132 #7658]: #2338 +#7708 := [hypothesis]: #4453 +#6849 := (or #4458 #3479 #6100) +#6096 := (= #6095 #2327) +#6099 := (or #6096 #3479) +#6850 := (or #4458 #6099) +#6859 := (iff #6850 #6849) +#6106 := (or #3479 #6100) +#6854 := (or #4458 #6106) +#6857 := (iff #6854 #6849) +#6858 := [rewrite]: #6857 +#6855 := (iff #6850 #6854) +#6109 := (iff #6099 #6106) +#6103 := (or #6100 #3479) +#6107 := (iff #6103 #6106) +#6108 := [rewrite]: #6107 +#6104 := (iff #6099 #6103) +#6101 := (iff #6096 #6100) +#6102 := [rewrite]: #6101 +#6105 := [monotonicity #6102]: #6104 +#6110 := [trans #6105 #6108]: #6109 +#6856 := [monotonicity #6110]: #6855 +#6860 := [trans #6856 #6858]: #6859 +#6853 := [quant-inst]: #6850 +#6861 := [mp #6853 #6860]: #6849 +#7667 := [unit-resolution #6861 #7708 #7656]: #6100 +#7668 := (not #6100) +#7666 := (or #7668 #7308) +#7669 := [th-lemma]: #7666 +#7670 := [unit-resolution #7669 #7667]: #7308 +#4139 := (not #2923) +#3968 := (or #3494 #4139) +#3970 := [def-axiom]: #3968 +#7671 := [unit-resolution #3970 #7658]: #4139 +#7057 := (uf_4 uf_14 ?x71!19) +#7092 := (* -1::int #7057) +#7093 := (+ #2325 #7092) +#7094 := (<= #7093 0::int) +#7672 := [hypothesis]: #4444 +#7099 := (or #4449 #7094) +#7084 := (+ #7057 #2326) +#7085 := (>= #7084 0::int) +#7100 := (or #4449 #7085) +#7102 := (iff #7100 #7099) +#7104 := (iff #7099 #7099) +#7105 := [rewrite]: #7104 +#7097 := (iff #7085 #7094) +#7086 := (+ #2326 #7057) +#7089 := (>= #7086 0::int) +#7095 := (iff #7089 #7094) +#7096 := [rewrite]: #7095 +#7090 := (iff #7085 #7089) +#7087 := (= #7084 #7086) +#7088 := [rewrite]: #7087 +#7091 := [monotonicity #7088]: #7090 +#7098 := [trans #7091 #7096]: #7097 +#7103 := [monotonicity #7098]: #7102 +#7106 := [trans #7103 #7105]: #7102 +#7101 := [quant-inst]: #7100 +#7107 := [mp #7101 #7106]: #7099 +#7673 := [unit-resolution #7107 #7672]: #7094 +#7218 := (+ #6095 #7092) +#7219 := (+ #2330 #7218) +#7220 := (>= #7219 0::int) +#6129 := (uf_6 uf_15 ?x72!18) +#6130 := (= uf_8 #6129) +decl uf_2 :: (-> T1 T2) +#7303 := (uf_2 #2329) +#7315 := (uf_6 uf_15 #7303) +#7316 := (= uf_8 #7315) +#7618 := (iff #7316 #6130) +#7616 := (= #7315 #6129) +#7707 := (= #6129 #7315) +#7304 := (= ?x72!18 #7303) +#16 := (uf_2 #12) +#325 := (= #10 #16) +#4203 := (forall (vars (?x4 T2) (?x5 T2)) (:pat #4196) #325) +#329 := (forall (vars (?x4 T2) (?x5 T2)) #325) +#4206 := (iff #329 #4203) +#4204 := (iff #325 #325) +#4205 := [refl]: #4204 +#4207 := [quant-intro #4205]: #4206 +#1844 := (~ #329 #329) +#1878 := (~ #325 #325) +#1879 := [refl]: #1878 +#1845 := [nnf-pos #1879]: #1844 +#17 := (= #16 #10) +#18 := (forall (vars (?x4 T2) (?x5 T2)) #17) +#330 := (iff #18 #329) +#327 := (iff #17 #325) +#328 := [rewrite]: #327 +#331 := [quant-intro #328]: #330 +#324 := [asserted]: #18 +#334 := [mp #324 #331]: #329 +#1880 := [mp~ #334 #1845]: #329 +#4208 := [mp #1880 #4207]: #4203 +#7310 := (not #4203) +#7311 := (or #7310 #7304) +#7312 := [quant-inst]: #7311 +#7862 := [unit-resolution #7312 #4208]: #7304 +#7751 := [monotonicity #7862]: #7707 +#7710 := [symm #7751]: #7616 +#7711 := [monotonicity #7710]: #7618 +#7575 := [hypothesis]: #4418 +#6147 := (= uf_22 ?x72!18) +#6150 := (ite #6147 #3895 #6130) +#4961 := (uf_7 uf_15 uf_22 #3894) +#6141 := (uf_6 #4961 ?x72!18) +#6144 := (= uf_8 #6141) +#6153 := (iff #6144 #6150) +#7188 := (or #4987 #6153) +#6139 := (= ?x72!18 uf_22) +#6140 := (ite #6139 #4958 #6130) +#6142 := (= #6141 uf_8) +#6143 := (iff #6142 #6140) +#7189 := (or #4987 #6143) +#7191 := (iff #7189 #7188) +#7193 := (iff #7188 #7188) +#7194 := [rewrite]: #7193 +#6154 := (iff #6143 #6153) +#6151 := (iff #6140 #6150) +#6148 := (iff #6139 #6147) +#6149 := [rewrite]: #6148 +#6152 := [monotonicity #6149 #4971]: #6151 +#6145 := (iff #6142 #6144) +#6146 := [rewrite]: #6145 +#6155 := [monotonicity #6146 #6152]: #6154 +#7192 := [monotonicity #6155]: #7191 +#7195 := [trans #7192 #7194]: #7191 +#7190 := [quant-inst]: #7189 +#7196 := [mp #7190 #7195]: #7188 +#7674 := [unit-resolution #7196 #4222]: #6153 +#7702 := (= #2337 #6141) +#7684 := (= #6141 #2337) +#7682 := (= #4961 uf_23) +#7718 := [hypothesis]: #195 +#7681 := [symm #7718]: #7680 +#7676 := (= #4961 #194) +#7679 := [monotonicity #7678]: #7676 +#7683 := [trans #7679 #7681]: #7682 +#7700 := [monotonicity #7683]: #7684 +#7703 := [symm #7700]: #7702 +#7704 := [trans #7656 #7703]: #6144 +#7211 := (not #6144) +#7208 := (not #6153) +#7212 := (or #7208 #7211 #6150) +#7213 := [def-axiom]: #7212 +#7705 := [unit-resolution #7213 #7704 #7674]: #6150 +#7587 := [hypothesis]: #2336 +#7197 := (not #6150) +#7876 := (not #7308) +#7626 := (not #7094) +#7627 := (or #7316 #7626 #2923 #7876 #4441 #2335 #7197 #4423) +#7857 := [hypothesis]: #7308 +#7858 := [hypothesis]: #4139 +#7108 := (uf_1 uf_22 ?x71!19) +#7109 := (uf_10 #7108) +#7113 := (* -1::int #7109) +#7842 := (+ #2330 #7113) +#7844 := (>= #7842 0::int) +#7841 := (= #2330 #7109) +#7843 := (= #2329 #7108) +#7559 := [hypothesis]: #6150 +#7205 := (not #6130) +#7614 := (not #7316) +#7615 := [hypothesis]: #7614 +#7625 := (or #7205 #7316) +#7620 := (iff #6130 #7316) +#7863 := (= #7303 ?x72!18) +#7864 := [symm #7862]: #7863 +#7617 := [monotonicity #7864]: #7616 +#7619 := [monotonicity #7617]: #7618 +#7621 := [symm #7619]: #7620 +#7613 := [hypothesis]: #6130 +#7665 := [mp #7613 #7621]: #7316 +#7624 := [unit-resolution #7615 #7665]: false +#7623 := [lemma #7624]: #7625 +#7565 := [unit-resolution #7623 #7615]: #7205 +#7201 := (or #7197 #6147 #6130) +#7202 := [def-axiom]: #7201 +#7566 := [unit-resolution #7202 #7565 #7559]: #6147 +#7567 := [symm #7566]: #6139 +#7568 := [monotonicity #7567]: #7843 +#7569 := [monotonicity #7568]: #7841 +#7847 := (not #7841) +#7848 := (or #7847 #7844) +#7849 := [th-lemma]: #7848 +#7576 := [unit-resolution #7849 #7569]: #7844 +#7577 := [hypothesis]: #7094 +#7127 := (+ #7092 #7109) +#7128 := (+ #188 #7127) +#7129 := (>= #7128 0::int) +#7134 := (+ #2326 #7109) +#7135 := (+ #188 #7134) +#7138 := (= #7135 0::int) +#7584 := (not #7138) +#7156 := (>= #7135 0::int) +#7875 := (not #7156) +#7309 := (uf_4 uf_14 #7303) +#7324 := (* -1::int #7309) +#7325 := (+ #188 #7324) +#7326 := (<= #7325 0::int) +#7331 := (or #7316 #7326) +#7334 := (or #4423 #7316 #7326) +#7313 := (+ #7309 #1455) +#7314 := (>= #7313 0::int) +#7317 := (or #7316 #7314) +#7335 := (or #4423 #7317) +#7342 := (iff #7335 #7334) +#7337 := (or #4423 #7331) +#7340 := (iff #7337 #7334) +#7341 := [rewrite]: #7340 +#7338 := (iff #7335 #7337) +#7332 := (iff #7317 #7331) +#7329 := (iff #7314 #7326) +#7318 := (+ #1455 #7309) +#7321 := (>= #7318 0::int) +#7327 := (iff #7321 #7326) +#7328 := [rewrite]: #7327 +#7322 := (iff #7314 #7321) +#7319 := (= #7313 #7318) +#7320 := [rewrite]: #7319 +#7323 := [monotonicity #7320]: #7322 +#7330 := [trans #7323 #7328]: #7329 +#7333 := [monotonicity #7330]: #7332 +#7339 := [monotonicity #7333]: #7338 +#7343 := [trans #7339 #7341]: #7342 +#7336 := [quant-inst]: #7335 +#7344 := [mp #7336 #7343]: #7334 +#7578 := [unit-resolution #7344 #7575]: #7331 +#7579 := [unit-resolution #7578 #7615]: #7326 +#7874 := (not #7326) +#7873 := (not #7844) +#7877 := (or #7873 #7874 #7875 #2923 #7876) +#7859 := [hypothesis]: #7156 +#7860 := [hypothesis]: #7844 +#7861 := [hypothesis]: #7326 +#7503 := (+ #6095 #7324) +#7507 := (>= #7503 0::int) +#7502 := (= #6095 #7309) +#7865 := (= #7309 #6095) +#7866 := [monotonicity #7864]: #7865 +#7867 := [symm #7866]: #7502 +#7868 := (not #7502) +#7869 := (or #7868 #7507) +#7870 := [th-lemma]: #7869 +#7871 := [unit-resolution #7870 #7867]: #7507 +#7872 := [th-lemma #7871 #7861 #7860 #7859 #7858 #7857]: false +#7878 := [lemma #7872]: #7877 +#7580 := [unit-resolution #7878 #7576 #7579 #7858 #7857]: #7875 +#7585 := (or #7584 #7156) +#7581 := [th-lemma]: #7585 +#7586 := [unit-resolution #7581 #7580]: #7584 +#7117 := (+ uf_9 #7113) +#7118 := (<= #7117 0::int) +#7180 := (not #7118) +#7590 := (or #7180 #2335 #7873) +#7591 := [th-lemma]: #7590 +#7592 := [unit-resolution #7591 #7576 #7587]: #7180 +#7141 := (or #7118 #7129 #7138) +#7144 := (or #4441 #7118 #7129 #7138) +#7110 := (+ #7109 #2326) +#7111 := (+ #188 #7110) +#7112 := (= #7111 0::int) +#7114 := (+ #1455 #7113) +#7115 := (+ #7057 #7114) +#7116 := (<= #7115 0::int) +#7119 := (or #7118 #7116 #7112) +#7145 := (or #4441 #7119) +#7152 := (iff #7145 #7144) +#7147 := (or #4441 #7141) +#7150 := (iff #7147 #7144) +#7151 := [rewrite]: #7150 +#7148 := (iff #7145 #7147) +#7142 := (iff #7119 #7141) +#7139 := (iff #7112 #7138) +#7136 := (= #7111 #7135) +#7137 := [rewrite]: #7136 +#7140 := [monotonicity #7137]: #7139 +#7132 := (iff #7116 #7129) +#7120 := (+ #7057 #7113) +#7121 := (+ #1455 #7120) +#7124 := (<= #7121 0::int) +#7130 := (iff #7124 #7129) +#7131 := [rewrite]: #7130 +#7125 := (iff #7116 #7124) +#7122 := (= #7115 #7121) +#7123 := [rewrite]: #7122 +#7126 := [monotonicity #7123]: #7125 +#7133 := [trans #7126 #7131]: #7132 +#7143 := [monotonicity #7133 #7140]: #7142 +#7149 := [monotonicity #7143]: #7148 +#7153 := [trans #7149 #7151]: #7152 +#7146 := [quant-inst]: #7145 +#7154 := [mp #7146 #7153]: #7144 +#7594 := [unit-resolution #7154 #7593]: #7141 +#7595 := [unit-resolution #7594 #7592 #7586]: #7129 +#7632 := [th-lemma #7871 #7579 #7595 #7577 #7576 #7858 #7857]: false +#7628 := [lemma #7632]: #7627 +#7706 := [unit-resolution #7628 #7673 #7671 #7670 #7593 #7587 #7705 #7575]: #7316 +#7709 := [mp #7706 #7711]: #6130 +#7713 := (or #7205 #7220) +#7712 := [hypothesis]: #4314 +#7225 := (or #4319 #2335 #7205 #7220) +#7221 := (or #7205 #2335 #7220) +#7226 := (or #4319 #7221) +#7233 := (iff #7226 #7225) +#7222 := (or #2335 #7205 #7220) +#7228 := (or #4319 #7222) +#7231 := (iff #7228 #7225) +#7232 := [rewrite]: #7231 +#7229 := (iff #7226 #7228) +#7223 := (iff #7221 #7222) +#7224 := [rewrite]: #7223 +#7230 := [monotonicity #7224]: #7229 +#7234 := [trans #7230 #7232]: #7233 +#7227 := [quant-inst]: #7226 +#7235 := [mp #7227 #7234]: #7225 +#7714 := [unit-resolution #7235 #7712 #7587]: #7713 +#7715 := [unit-resolution #7714 #7709]: #7220 +#7716 := [th-lemma #7715 #7673 #7671 #7670]: false +#7761 := [lemma #7716]: #7717 +#16979 := [unit-resolution #7761 #9187]: #17049 +#17090 := [unit-resolution #16979 #10785 #17666 #13581 #10788]: #17073 +#16348 := [unit-resolution #17090 #25699]: #17077 +#17046 := [unit-resolution #16348 #17052 #7658]: false +#16349 := [lemma #17046]: #3494 +#5169 := (uf_6 uf_23 ?x75!20) +#5170 := (= uf_8 #5169) +#21788 := (uf_6 #10323 ?x75!20) +#20832 := (= #21788 #5169) +#20857 := (= #5169 #21788) +#14425 := (= uf_23 #10323) +#14423 := (= #194 #10323) +#14424 := [symm #13577]: #14423 +#14426 := [trans #13581 #14424]: #14425 +#20702 := [monotonicity #14426]: #20857 +#20877 := [symm #20702]: #20832 +#21791 := (= uf_8 #21788) +#6014 := (uf_6 uf_15 ?x75!20) +#6015 := (= uf_8 #6014) +#21786 := (= ?x75!20 #9695) +#21794 := (ite #21786 #3895 #6015) +#21797 := (iff #21791 #21794) +#19678 := (or #4987 #21797) +#21787 := (ite #21786 #4958 #6015) +#21789 := (= #21788 uf_8) +#21790 := (iff #21789 #21787) +#19676 := (or #4987 #21790) +#19311 := (iff #19676 #19678) +#19682 := (iff #19678 #19678) +#19685 := [rewrite]: #19682 +#21798 := (iff #21790 #21797) +#21795 := (iff #21787 #21794) +#21796 := [monotonicity #4971]: #21795 +#21792 := (iff #21789 #21791) +#21793 := [rewrite]: #21792 +#21799 := [monotonicity #21793 #21796]: #21798 +#19681 := [monotonicity #21799]: #19311 +#19733 := [trans #19681 #19685]: #19311 +#19684 := [quant-inst]: #19676 +#19731 := [mp #19684 #19733]: #19678 +#20449 := [unit-resolution #19731 #4222]: #21797 +#15125 := (uf_1 #9695 ?x75!20) +#15126 := (uf_10 #15125) +#19741 := (<= #15126 0::int) +#4781 := (* -1::int #4740) +#5008 := (+ #188 #4781) +#5009 := (>= #5008 0::int) +#9232 := (or #4449 #5009) +#7564 := [quant-inst]: #9232 +#9311 := [unit-resolution #7564 #10788]: #5009 +#11895 := (uf_24 #9695) +#10002 := (* -1::int #11895) +#14620 := (+ #2355 #10002) +#15146 := (<= #14620 0::int) +#14621 := (uf_6 uf_23 #9695) +#14622 := (= uf_8 #14621) +#21303 := (= #3894 #14621) +#21293 := (= #14621 #3894) +#21294 := [monotonicity #13581 #10708]: #21293 +#21304 := [symm #21294]: #21303 +#21305 := [trans #7677 #21304]: #14622 +#15158 := (* -1::int #15126) +#15159 := (+ #10002 #15158) +#15160 := (+ #2355 #15159) +#15161 := (= #15160 0::int) +#23339 := (<= #15160 0::int) +#5042 := (<= #5008 0::int) +#10190 := (not #4741) +#10191 := (or #10190 #5042) +#10192 := [th-lemma]: #10191 +#17097 := [unit-resolution #10192 #10729]: #5042 +#10025 := (+ #4740 #10002) +#10040 := (<= #10025 0::int) +#9964 := (= #4740 #11895) +#17098 := (= #11895 #4740) +#17107 := [monotonicity #10708]: #17098 +#17108 := [symm #17107]: #9964 +#17109 := (not #9964) +#17110 := (or #17109 #10040) +#17111 := [th-lemma]: #17110 +#17112 := [unit-resolution #17111 #17108]: #10040 +#5826 := (uf_1 uf_22 ?x75!20) +#5827 := (uf_10 #5826) +#23569 := (+ #5827 #15158) +#23570 := (<= #23569 0::int) +#23568 := (= #5827 #15126) +#20561 := (= #5826 #15125) +#25948 := (= #15125 #5826) +#25949 := [monotonicity #10708]: #25948 +#20537 := [symm #25949]: #20561 +#20580 := [monotonicity #20537]: #23568 +#25953 := (not #23568) +#25961 := (or #25953 #23570) +#25962 := [th-lemma]: #25961 +#20545 := [unit-resolution #25962 #20580]: #23570 +#5852 := (+ #2356 #5827) +#5853 := (+ #188 #5852) +#23301 := (>= #5853 0::int) +#20066 := [hypothesis]: #4498 +#4125 := (or #4495 #2368) +#4127 := [def-axiom]: #4125 +#20064 := [unit-resolution #4127 #20066]: #2368 +#29068 := (or #23301 #2367) +#4860 := (>= #188 0::int) +#4051 := (or #4579 #4306) +#4047 := [def-axiom]: #4051 +#10596 := [unit-resolution #4047 #5496]: #4306 +#9276 := (or #4311 #4860) +#9277 := [quant-inst]: #9276 +#12530 := [unit-resolution #9277 #10596]: #4860 +#25950 := (= #15126 #5827) +#25951 := [monotonicity #25949]: #25950 +#25952 := [symm #25951]: #23568 +#25963 := [unit-resolution #25962 #25952]: #23570 +#23340 := (>= #15160 0::int) +#10021 := (>= #10025 0::int) +#25933 := (or #17109 #10021) +#25934 := [th-lemma]: #25933 +#25935 := [unit-resolution #25934 #17108]: #10021 +#23571 := (>= #23569 0::int) +#25954 := (or #25953 #23571) +#25955 := [th-lemma]: #25954 +#25956 := [unit-resolution #25955 #25952]: #23571 +#23300 := (<= #5853 0::int) +#25968 := (not #23301) +#25609 := [hypothesis]: #25968 +#29059 := (or #23300 #23301) +#29062 := [th-lemma]: #29059 +#29061 := [unit-resolution #29062 #25609]: #23300 +#25957 := (not #23571) +#25944 := (not #23300) +#8759 := (not #5009) +#25942 := (not #10021) +#25958 := (or #23340 #25942 #8759 #25944 #25957) +#25959 := [th-lemma]: #25958 +#29063 := [unit-resolution #25959 #29061 #25956 #25935 #9311]: #23340 +#24500 := [hypothesis]: #2368 +#5831 := (* -1::int #5827) +#5835 := (+ uf_9 #5831) +#5836 := (<= #5835 0::int) +#25729 := (or #23301 #5836) +#23313 := (not #5836) +#25608 := [hypothesis]: #23313 +#5856 := (= #5853 0::int) +#25937 := (not #5856) +#25964 := (or #25937 #23301) +#25965 := [th-lemma]: #25964 +#25722 := [unit-resolution #25965 #25609]: #25937 +#5775 := (uf_4 uf_14 ?x75!20) +#5810 := (* -1::int #5775) +#5845 := (+ #5810 #5827) +#5846 := (+ #188 #5845) +#5847 := (>= #5846 0::int) +#23316 := (not #5847) +#5811 := (+ #2355 #5810) +#5812 := (<= #5811 0::int) +#23280 := (or #4449 #5812) +#5802 := (+ #5775 #2356) +#5803 := (>= #5802 0::int) +#23281 := (or #4449 #5803) +#23283 := (iff #23281 #23280) +#23285 := (iff #23280 #23280) +#23286 := [rewrite]: #23285 +#5815 := (iff #5803 #5812) +#5804 := (+ #2356 #5775) +#5807 := (>= #5804 0::int) +#5813 := (iff #5807 #5812) +#5814 := [rewrite]: #5813 +#5808 := (iff #5803 #5807) +#5805 := (= #5802 #5804) +#5806 := [rewrite]: #5805 +#5809 := [monotonicity #5806]: #5808 +#5816 := [trans #5809 #5814]: #5815 +#23284 := [monotonicity #5816]: #23283 +#23287 := [trans #23284 #23286]: #23283 +#23282 := [quant-inst]: #23281 +#23288 := [mp #23282 #23287]: #23280 +#25723 := [unit-resolution #23288 #10788]: #5812 +#25724 := (not #5812) +#25725 := (or #23301 #23316 #25724) +#25726 := [th-lemma]: #25725 +#25721 := [unit-resolution #25726 #25609 #25723]: #23316 +#5859 := (or #5836 #5847 #5856) +#23289 := (or #4441 #5836 #5847 #5856) +#5828 := (+ #5827 #2356) +#5829 := (+ #188 #5828) +#5830 := (= #5829 0::int) +#5832 := (+ #1455 #5831) +#5833 := (+ #5775 #5832) +#5834 := (<= #5833 0::int) +#5837 := (or #5836 #5834 #5830) +#23290 := (or #4441 #5837) +#23297 := (iff #23290 #23289) +#23292 := (or #4441 #5859) +#23295 := (iff #23292 #23289) +#23296 := [rewrite]: #23295 +#23293 := (iff #23290 #23292) +#5860 := (iff #5837 #5859) +#5857 := (iff #5830 #5856) +#5854 := (= #5829 #5853) +#5855 := [rewrite]: #5854 +#5858 := [monotonicity #5855]: #5857 +#5850 := (iff #5834 #5847) +#5838 := (+ #5775 #5831) +#5839 := (+ #1455 #5838) +#5842 := (<= #5839 0::int) +#5848 := (iff #5842 #5847) +#5849 := [rewrite]: #5848 +#5843 := (iff #5834 #5842) +#5840 := (= #5833 #5839) +#5841 := [rewrite]: #5840 +#5844 := [monotonicity #5841]: #5843 +#5851 := [trans #5844 #5849]: #5850 +#5861 := [monotonicity #5851 #5858]: #5860 +#23294 := [monotonicity #5861]: #23293 +#23298 := [trans #23294 #23296]: #23297 +#23291 := [quant-inst]: #23290 +#23299 := [mp #23291 #23298]: #23289 +#25727 := [unit-resolution #23299 #10785]: #5859 +#25728 := [unit-resolution #25727 #25721 #25722 #25608]: false +#25730 := [lemma #25728]: #25729 +#29064 := [unit-resolution #25730 #25609]: #5836 +#29069 := [th-lemma #29064 #24500 #29063 #17112 #17097 #25963 #12530]: false +#29071 := [lemma #29069]: #29068 +#20448 := [unit-resolution #29071 #20064]: #23301 +#25969 := (not #23570) +#8760 := (not #5042) +#25967 := (not #10040) +#25970 := (or #23339 #25967 #8760 #25968 #25969) +#25971 := [th-lemma]: #25970 +#20530 := [unit-resolution #25971 #20448 #20545 #17112 #17097]: #23339 +#20756 := [unit-resolution #25955 #20580]: #23571 +#5878 := (or #5836 #5847) +#5881 := (not #5878) +#5780 := (= #2355 #5775) +#20062 := (not #5780) +#24414 := (>= #5811 0::int) +#24509 := (not #24414) +#4128 := (or #4495 #2937) +#4126 := [def-axiom]: #4128 +#20016 := [unit-resolution #4126 #20066]: #2937 +#4012 := (or #4495 #4487) +#4013 := [def-axiom]: #4012 +#20068 := [unit-resolution #4013 #20066]: #4487 +#23167 := (or #24509 #4492 #2934 #2367) +#6043 := (?x47!7 ?x75!20) +#6048 := (uf_1 #6043 ?x75!20) +#24986 := (uf_2 #6048) +#25837 := (uf_6 uf_15 #24986) +#25838 := (= uf_8 #25837) +#21703 := (= #9695 #24986) +#22526 := (ite #21703 #3895 #25838) +#23044 := (not #22526) +#21277 := (uf_6 #10323 #24986) +#21622 := (= uf_8 #21277) +#22242 := (iff #21622 #22526) +#22524 := (or #4987 #22242) +#21246 := (= #24986 #9695) +#21220 := (ite #21246 #4958 #25838) +#21600 := (= #21277 uf_8) +#21601 := (iff #21600 #21220) +#23024 := (or #4987 #21601) +#23028 := (iff #23024 #22524) +#23042 := (iff #22524 #22524) +#23043 := [rewrite]: #23042 +#22533 := (iff #21601 #22242) +#22529 := (iff #21220 #22526) +#22515 := (iff #21246 #21703) +#22520 := [rewrite]: #22515 +#22530 := [monotonicity #22520 #4971]: #22529 +#21227 := (iff #21600 #21622) +#21608 := [rewrite]: #21227 +#22592 := [monotonicity #21608 #22530]: #22533 +#23029 := [monotonicity #22592]: #23028 +#23041 := [trans #23029 #23043]: #23028 +#23025 := [quant-inst]: #23024 +#22593 := [mp #23025 #23041]: #22524 +#23122 := [unit-resolution #22593 #4222]: #22242 +#23057 := (not #21622) +#24996 := (uf_6 uf_23 #24986) +#24997 := (= uf_8 #24996) +#24998 := (not #24997) +#23123 := (iff #24998 #23057) +#23153 := (iff #24997 #21622) +#23146 := (iff #21622 #24997) +#23151 := (= #21277 #24996) +#23149 := [monotonicity #13583]: #23151 +#23152 := [monotonicity #23149]: #23146 +#23127 := [symm #23152]: #23153 +#23124 := [monotonicity #23127]: #23123 +#24506 := [hypothesis]: #24414 +#6049 := (uf_10 #6048) +#6050 := (* -1::int #6049) +#6044 := (uf_4 uf_14 #6043) +#6045 := (* -1::int #6044) +#6051 := (+ #6045 #6050) +#6052 := (+ #5775 #6051) +#18721 := (>= #6052 0::int) +#6053 := (= #6052 0::int) +#6055 := (uf_6 uf_15 #6043) +#6056 := (= uf_8 #6055) +#6057 := (not #6056) +#6054 := (not #6053) +#6046 := (+ #5775 #6045) +#6047 := (<= #6046 0::int) +#6063 := (or #6047 #6054 #6057) +#6066 := (not #6063) +#6060 := (+ uf_9 #5810) +#6061 := (<= #6060 0::int) +#24508 := (not #6061) +#24510 := (or #24508 #24509 #2367) +#24505 := [hypothesis]: #6061 +#24507 := [th-lemma #24506 #24505 #24500]: false +#24511 := [lemma #24507]: #24510 +#23125 := [unit-resolution #24511 #24506 #24500]: #24508 +#23139 := (or #6061 #6066) +#23138 := [hypothesis]: #2937 +#19063 := (or #4344 #2934 #6061 #6066) +#6058 := (or #6057 #6054 #6047) +#6059 := (not #6058) +#6062 := (or #2369 #6061 #6059) +#19066 := (or #4344 #6062) +#18390 := (iff #19066 #19063) +#6069 := (or #2934 #6061 #6066) +#19050 := (or #4344 #6069) +#19090 := (iff #19050 #19063) +#19045 := [rewrite]: #19090 +#19051 := (iff #19066 #19050) +#6070 := (iff #6062 #6069) +#6067 := (iff #6059 #6066) +#6064 := (iff #6058 #6063) +#6065 := [rewrite]: #6064 +#6068 := [monotonicity #6065]: #6067 +#6071 := [monotonicity #2936 #6068]: #6070 +#19061 := [monotonicity #6071]: #19051 +#18720 := [trans #19061 #19045]: #18390 +#18371 := [quant-inst]: #19066 +#19065 := [mp #18371 #18720]: #19063 +#23137 := [unit-resolution #19065 #9243 #23138]: #23139 +#23140 := [unit-resolution #23137 #23125]: #6066 +#19081 := (or #6063 #6053) +#12850 := [def-axiom]: #19081 +#23135 := [unit-resolution #12850 #23140]: #6053 +#23141 := (or #6054 #18721) +#23143 := [th-lemma]: #23141 +#23144 := [unit-resolution #23143 #23135]: #18721 +#25978 := [hypothesis]: #4487 +#23633 := (<= #6052 0::int) +#23142 := (or #6054 #23633) +#23145 := [th-lemma]: #23142 +#23147 := [unit-resolution #23145 #23135]: #23633 +#19055 := (not #6047) +#18317 := (or #6063 #19055) +#19075 := [def-axiom]: #18317 +#23148 := [unit-resolution #19075 #23140]: #19055 +#30651 := (not #18721) +#30650 := (not #23633) +#30652 := (or #24998 #6047 #30650 #4492 #30651 #24509) +#28637 := (uf_1 #24986 ?x75!20) +#28638 := (uf_10 #28637) +#28651 := (* -1::int #28638) +#24990 := (uf_24 #24986) +#24991 := (* -1::int #24990) +#28652 := (+ #24991 #28651) +#28653 := (+ #2355 #28652) +#28682 := (>= #28653 0::int) +#28654 := (= #28653 0::int) +#19910 := (uf_3 #15125) +#22479 := (uf_1 #24986 #19910) +#22480 := (uf_10 #22479) +#22498 := (* -1::int #22480) +#22603 := (+ #22498 #24991) +#21494 := (uf_24 #19910) +#22604 := (+ #21494 #22603) +#30636 := (= #22604 0::int) +#30564 := [hypothesis]: #23633 +#29534 := [hypothesis]: #18721 +#29010 := (+ #6049 #28651) +#21152 := (<= #29010 0::int) +#21012 := (= #6049 #28638) +#30043 := (= #6048 #28637) +#24987 := (= #6043 #24986) +#20827 := (or #7310 #24987) +#20849 := [quant-inst]: #20827 +#30089 := [unit-resolution #20849 #4208]: #24987 +#30044 := [monotonicity #30089]: #30043 +#30097 := [monotonicity #30044]: #21012 +#30102 := (not #21012) +#30581 := (or #30102 #21152) +#30588 := [th-lemma]: #30581 +#30589 := [unit-resolution #30588 #30097]: #21152 +#29012 := (>= #29010 0::int) +#30590 := (or #30102 #29012) +#30591 := [th-lemma]: #30590 +#30592 := [unit-resolution #30591 #30097]: #29012 +#19879 := (+ #22480 #28651) +#21287 := (<= #19879 0::int) +#21209 := (= #22480 #28638) +#30596 := (= #22479 #28637) +#30594 := (= #19910 ?x75!20) +#19950 := (= ?x75!20 #19910) +#19896 := (or #8139 #19950) +#19953 := [quant-inst]: #19896 +#30593 := [unit-resolution #19953 #4202]: #19950 +#30595 := [symm #30593]: #30594 +#30597 := [monotonicity #30595]: #30596 +#30598 := [monotonicity #30597]: #21209 +#30599 := (not #21209) +#30600 := (or #30599 #21287) +#30601 := [th-lemma]: #30600 +#30602 := [unit-resolution #30601 #30598]: #21287 +#20127 := (>= #19879 0::int) +#30603 := (or #30599 #20127) +#30604 := [th-lemma]: #30603 +#30605 := [unit-resolution #30604 #30598]: #20127 +#22047 := (* -1::int #21494) +#20392 := (+ #2355 #22047) +#20387 := (<= #20392 0::int) +#19773 := (= #2355 #21494) +#30606 := (= #21494 #2355) +#30607 := [monotonicity #30595]: #30606 +#30608 := [symm #30607]: #19773 +#30609 := (not #19773) +#30610 := (or #30609 #20387) +#30611 := [th-lemma]: #30610 +#30612 := [unit-resolution #30611 #30608]: #20387 +#20393 := (>= #20392 0::int) +#30613 := (or #30609 #20393) +#30614 := [th-lemma]: #30613 +#30615 := [unit-resolution #30614 #30608]: #20393 +#25833 := (uf_4 uf_14 #24986) +#25834 := (* -1::int #25833) +#21068 := (+ #6044 #25834) +#21067 := (<= #21068 0::int) +#21078 := (= #6044 #25833) +#30618 := (= #25833 #6044) +#30616 := (= #24986 #6043) +#30617 := [symm #30089]: #30616 +#30619 := [monotonicity #30617]: #30618 +#30620 := [symm #30619]: #21078 +#30621 := (not #21078) +#30622 := (or #30621 #21067) +#30623 := [th-lemma]: #30622 +#30624 := [unit-resolution #30623 #30620]: #21067 +#21183 := (>= #21068 0::int) +#30625 := (or #30621 #21183) +#30626 := [th-lemma]: #30625 +#30627 := [unit-resolution #30626 #30620]: #21183 +#28762 := (+ #24990 #25834) +#28763 := (<= #28762 0::int) +#22244 := (or #4449 #28763) +#28754 := (+ #25833 #24991) +#28755 := (>= #28754 0::int) +#22245 := (or #4449 #28755) +#22246 := (iff #22245 #22244) +#22240 := (iff #22244 #22244) +#22286 := [rewrite]: #22240 +#28766 := (iff #28755 #28763) +#28756 := (+ #24991 #25833) +#28759 := (>= #28756 0::int) +#28764 := (iff #28759 #28763) +#28765 := [rewrite]: #28764 +#28760 := (iff #28755 #28759) +#28757 := (= #28754 #28756) +#28758 := [rewrite]: #28757 +#28761 := [monotonicity #28758]: #28760 +#28767 := [trans #28761 #28765]: #28766 +#22285 := [monotonicity #28767]: #22246 +#22287 := [trans #22285 #22286]: #22246 +#22187 := [quant-inst]: #22245 +#22288 := [mp #22187 #22287]: #22244 +#30628 := [unit-resolution #22288 #10788]: #28763 +#30563 := (>= #28762 0::int) +#28732 := (= #24990 #25833) +#30629 := [hypothesis]: #24997 +#28738 := (or #24998 #28732) +#22151 := (or #4458 #24998 #28732) +#28728 := (= #25833 #24990) +#28729 := (or #28728 #24998) +#22180 := (or #4458 #28729) +#22189 := (iff #22180 #22151) +#22183 := (or #4458 #28738) +#22186 := (iff #22183 #22151) +#22188 := [rewrite]: #22186 +#22179 := (iff #22180 #22183) +#28741 := (iff #28729 #28738) +#28735 := (or #28732 #24998) +#28739 := (iff #28735 #28738) +#28740 := [rewrite]: #28739 +#28736 := (iff #28729 #28735) +#28733 := (iff #28728 #28732) +#28734 := [rewrite]: #28733 +#28737 := [monotonicity #28734]: #28736 +#28742 := [trans #28737 #28740]: #28741 +#22185 := [monotonicity #28742]: #22179 +#22184 := [trans #22185 #22188]: #22189 +#22154 := [quant-inst]: #22180 +#22243 := [mp #22154 #22184]: #22151 +#30630 := [unit-resolution #22243 #25699]: #28738 +#30631 := [unit-resolution #30630 #30629]: #28732 +#30632 := (not #28732) +#30633 := (or #30632 #30563) +#30634 := [th-lemma]: #30633 +#30635 := [unit-resolution #30634 #30631]: #30563 +#30637 := [th-lemma #30635 #30628 #30627 #30624 #30615 #30612 #30605 #30602 #30592 #30589 #29534 #30564 #24506 #25723]: #30636 +#30640 := (= #28653 #22604) +#30638 := (= #22604 #28653) +#30639 := [th-lemma #30615 #30612 #30605 #30602]: #30638 +#30641 := [symm #30639]: #30640 +#30642 := [trans #30641 #30637]: #28654 +#28659 := (not #28654) +#30643 := (or #28659 #28682) +#30644 := [th-lemma]: #30643 +#30645 := [unit-resolution #30644 #30642]: #28682 +#16155 := (+ #2355 #24991) +#15436 := (<= #16155 0::int) +#28665 := (or #15436 #24998 #28659) +#21333 := (or #4492 #15436 #24998 #28659) +#28639 := (+ #2356 #28638) +#28640 := (+ #24990 #28639) +#28641 := (= #28640 0::int) +#28642 := (not #28641) +#15907 := (+ #24990 #2356) +#16051 := (>= #15907 0::int) +#28643 := (or #24998 #16051 #28642) +#21024 := (or #4492 #28643) +#20975 := (iff #21024 #21333) +#21318 := (or #4492 #28665) +#21314 := (iff #21318 #21333) +#21310 := [rewrite]: #21314 +#21153 := (iff #21024 #21318) +#28668 := (iff #28643 #28665) +#28662 := (or #24998 #15436 #28659) +#28666 := (iff #28662 #28665) +#28667 := [rewrite]: #28666 +#28663 := (iff #28643 #28662) +#28660 := (iff #28642 #28659) +#28657 := (iff #28641 #28654) +#28644 := (+ #24990 #28638) +#28645 := (+ #2356 #28644) +#28648 := (= #28645 0::int) +#28655 := (iff #28648 #28654) +#28656 := [rewrite]: #28655 +#28649 := (iff #28641 #28648) +#28646 := (= #28640 #28645) +#28647 := [rewrite]: #28646 +#28650 := [monotonicity #28647]: #28649 +#28658 := [trans #28650 #28656]: #28657 +#28661 := [monotonicity #28658]: #28660 +#15965 := (iff #16051 #15436) +#13744 := (+ #2356 #24990) +#15070 := (>= #13744 0::int) +#14692 := (iff #15070 #15436) +#16264 := [rewrite]: #14692 +#16123 := (iff #16051 #15070) +#16156 := (= #15907 #13744) +#15494 := [rewrite]: #16156 +#15349 := [monotonicity #15494]: #16123 +#16187 := [trans #15349 #16264]: #15965 +#28664 := [monotonicity #16187 #28661]: #28663 +#28669 := [trans #28664 #28667]: #28668 +#21320 := [monotonicity #28669]: #21153 +#21319 := [trans #21320 #21310]: #20975 +#21323 := [quant-inst]: #21024 +#20612 := [mp #21323 #21319]: #21333 +#30646 := [unit-resolution #20612 #25978]: #28665 +#30647 := [unit-resolution #30646 #30642 #30629]: #15436 +#30648 := [hypothesis]: #19055 +#30649 := [th-lemma #30648 #30589 #30564 #30647 #30645]: false +#30653 := [lemma #30649]: #30652 +#23150 := [unit-resolution #30653 #23148 #23147 #25978 #23144 #24506]: #24998 +#23131 := [mp #23150 #23124]: #23057 +#23056 := (not #22242) +#23047 := (or #23056 #21622 #23044) +#23051 := [def-axiom]: #23047 +#23132 := [unit-resolution #23051 #23131 #23122]: #23044 +#23045 := (not #21703) +#23130 := (or #22526 #23045) +#3924 := (not #3895) +#23054 := (or #22526 #23045 #3924) +#23050 := [def-axiom]: #23054 +#23157 := [unit-resolution #23050 #7677]: #23130 +#23158 := [unit-resolution #23157 #23132]: #23045 +#23164 := (or #22526 #21703) +#23162 := (= #6055 #25837) +#23129 := (= #25837 #6055) +#23156 := [monotonicity #30617]: #23129 +#23128 := [symm #23156]: #23162 +#19077 := (or #6063 #6056) +#19078 := [def-axiom]: #19077 +#23133 := [unit-resolution #19078 #23140]: #6056 +#23163 := [trans #23133 #23128]: #25838 +#25839 := (not #25838) +#23055 := (or #22526 #21703 #25839) +#23053 := [def-axiom]: #23055 +#23165 := [unit-resolution #23053 #23163]: #23164 +#23169 := [unit-resolution #23165 #23158 #23132]: false +#23168 := [lemma #23169]: #23167 +#19986 := [unit-resolution #23168 #20068 #20016 #20064]: #24509 +#20023 := (or #20062 #24414) +#20453 := [th-lemma]: #20023 +#20441 := [unit-resolution #20453 #19986]: #20062 +#5884 := (or #5780 #5881) +#18349 := (or #4433 #5780 #5881) +#5875 := (or #5836 #5834) +#5876 := (not #5875) +#5776 := (= #5775 #2355) +#5877 := (or #5776 #5876) +#18388 := (or #4433 #5877) +#18375 := (iff #18388 #18349) +#18374 := (or #4433 #5884) +#18380 := (iff #18374 #18349) +#18414 := [rewrite]: #18380 +#18370 := (iff #18388 #18374) +#5885 := (iff #5877 #5884) +#5882 := (iff #5876 #5881) +#5879 := (iff #5875 #5878) +#5880 := [monotonicity #5851]: #5879 +#5883 := [monotonicity #5880]: #5882 +#5781 := (iff #5776 #5780) +#5782 := [rewrite]: #5781 +#5886 := [monotonicity #5782 #5883]: #5885 +#18411 := [monotonicity #5886]: #18370 +#18381 := [trans #18411 #18414]: #18375 +#18415 := [quant-inst]: #18388 +#18379 := [mp #18415 #18381]: #18349 +#20748 := [unit-resolution #18379 #10727]: #5884 +#20750 := [unit-resolution #20748 #20441]: #5881 +#18373 := (or #5878 #23313) +#18372 := [def-axiom]: #18373 +#20751 := [unit-resolution #18372 #20750]: #23313 +#18384 := (or #5878 #23316) +#18417 := [def-axiom]: #18384 +#20753 := [unit-resolution #18417 #20750]: #23316 +#20775 := [unit-resolution #25727 #20753 #20751]: #5856 +#25938 := (or #25937 #23300) +#25939 := [th-lemma]: #25938 +#20664 := [unit-resolution #25939 #20775]: #23300 +#20698 := [unit-resolution #25959 #20664 #20756 #25935 #9311]: #23340 +#25974 := (not #23340) +#25973 := (not #23339) +#25975 := (or #15161 #25973 #25974) +#25976 := [th-lemma]: #25975 +#20700 := [unit-resolution #25976 #20698 #20530]: #15161 +#15166 := (not #15161) +#14623 := (not #14622) +#15169 := (or #14623 #15146 #15166) +#23321 := (or #4492 #14623 #15146 #15166) +#15133 := (+ #2356 #15126) +#15134 := (+ #11895 #15133) +#15135 := (= #15134 0::int) +#15136 := (not #15135) +#15137 := (+ #11895 #2356) +#15138 := (>= #15137 0::int) +#15139 := (or #14623 #15138 #15136) +#23322 := (or #4492 #15139) +#23336 := (iff #23322 #23321) +#23331 := (or #4492 #15169) +#23334 := (iff #23331 #23321) +#23335 := [rewrite]: #23334 +#23332 := (iff #23322 #23331) +#15170 := (iff #15139 #15169) +#15167 := (iff #15136 #15166) +#15164 := (iff #15135 #15161) +#15151 := (+ #11895 #15126) +#15152 := (+ #2356 #15151) +#15155 := (= #15152 0::int) +#15162 := (iff #15155 #15161) +#15163 := [rewrite]: #15162 +#15156 := (iff #15135 #15155) +#15153 := (= #15134 #15152) +#15154 := [rewrite]: #15153 +#15157 := [monotonicity #15154]: #15156 +#15165 := [trans #15157 #15163]: #15164 +#15168 := [monotonicity #15165]: #15167 +#15149 := (iff #15138 #15146) +#15140 := (+ #2356 #11895) +#15143 := (>= #15140 0::int) +#15147 := (iff #15143 #15146) +#15148 := [rewrite]: #15147 +#15144 := (iff #15138 #15143) +#15141 := (= #15137 #15140) +#15142 := [rewrite]: #15141 +#15145 := [monotonicity #15142]: #15144 +#15150 := [trans #15145 #15148]: #15149 +#15171 := [monotonicity #15150 #15168]: #15170 +#23333 := [monotonicity #15171]: #23332 +#23337 := [trans #23333 #23335]: #23336 +#23330 := [quant-inst]: #23322 +#23338 := [mp #23330 #23337]: #23321 +#20754 := [unit-resolution #23338 #20068]: #15169 +#20663 := [unit-resolution #20754 #20700 #21305]: #15146 +#25941 := (not #15146) +#20620 := (or #19741 #25957 #25944 #25941 #25942 #8759) +#20724 := [th-lemma]: #20620 +#20713 := [unit-resolution #20724 #20664 #25935 #20756 #20663 #9311]: #19741 +#19738 := (not #19741) +#20001 := (or #10302 #19738 #21786) +#19729 := (= #9695 ?x75!20) +#19725 := (or #19729 #19738) +#19974 := (or #10302 #19725) +#19977 := (iff #19974 #20001) +#19723 := (or #19738 #21786) +#19976 := (or #10302 #19723) +#19980 := (iff #19976 #20001) +#19985 := [rewrite]: #19980 +#19981 := (iff #19974 #19976) +#19759 := (iff #19725 #19723) +#19727 := (or #21786 #19738) +#19755 := (iff #19727 #19723) +#19757 := [rewrite]: #19755 +#19728 := (iff #19725 #19727) +#19742 := (iff #19729 #21786) +#19724 := [rewrite]: #19742 +#19726 := [monotonicity #19724]: #19728 +#19753 := [trans #19726 #19757]: #19759 +#19982 := [monotonicity #19753]: #19981 +#19947 := [trans #19982 #19985]: #19977 +#20005 := [quant-inst]: #19974 +#19997 := [mp #20005 #19947]: #20001 +#20716 := [unit-resolution #19997 #4249 #20713]: #21786 +#19787 := (not #21786) +#20752 := (or #21794 #19787) +#19745 := (or #21794 #19787 #3924) +#19739 := [def-axiom]: #19745 +#20774 := [unit-resolution #19739 #7677]: #20752 +#20858 := [unit-resolution #20774 #20716]: #21794 +#19746 := (not #21794) +#19695 := (not #21797) +#19694 := (or #19695 #21791 #19746) +#19761 := [def-axiom]: #19694 +#20783 := [unit-resolution #19761 #20858 #20449]: #21791 +#20859 := [trans #20783 #20877]: #5170 +#5171 := (not #5170) +#5786 := (or #5171 #5780) +#18330 := (or #4458 #5171 #5780) +#5779 := (or #5776 #5171) +#18283 := (or #4458 #5779) +#18352 := (iff #18283 #18330) +#18337 := (or #4458 #5786) +#18335 := (iff #18337 #18330) +#18351 := [rewrite]: #18335 +#18336 := (iff #18283 #18337) +#5789 := (iff #5779 #5786) +#5783 := (or #5780 #5171) +#5787 := (iff #5783 #5786) +#5788 := [rewrite]: #5787 +#5784 := (iff #5779 #5783) +#5785 := [monotonicity #5782]: #5784 +#5790 := [trans #5785 #5788]: #5789 +#18331 := [monotonicity #5790]: #18336 +#18358 := [trans #18331 #18351]: #18352 +#18350 := [quant-inst]: #18283 +#18412 := [mp #18350 #18358]: #18330 +#20455 := [unit-resolution #18412 #25699]: #5786 +#20452 := [unit-resolution #20455 #20441]: #5171 +#20896 := [unit-resolution #20452 #20859]: false +#20973 := [lemma #20896]: #4495 +#4006 := (or #4504 #3499 #4498) +#4021 := [def-axiom]: #4006 +#19204 := [unit-resolution #4021 #20973 #16349]: #4504 +#4017 := (or #4507 #4501) +#4025 := [def-axiom]: #4017 +#19217 := [unit-resolution #4025 #19204]: #4507 +#11869 := (or #2283 #4458) +#8633 := (uf_1 uf_22 ?x65!15) +#8634 := (uf_10 #8633) +#5075 := (* -1::int #2282) +#8695 := (+ #5075 #8634) +#8696 := (+ #188 #8695) +#13182 := (<= #8696 0::int) +#8699 := (= #8696 0::int) +#8635 := (* -1::int #8634) +#8639 := (+ uf_9 #8635) +#8640 := (<= #8639 0::int) +#13547 := (not #8640) +#8579 := (uf_4 uf_14 ?x65!15) +#8589 := (* -1::int #8579) +#8655 := (+ #8589 #8634) +#8656 := (+ #188 #8655) +#8657 := (>= #8656 0::int) +#8662 := (or #8640 #8657) +#8665 := (not #8662) +#8645 := (= #2282 #8579) +#13653 := (not #8645) +#8618 := (+ #2282 #8589) +#13938 := (>= #8618 0::int) +#14032 := (not #13938) +#8826 := [hypothesis]: #2284 +#14033 := (or #14032 #2283) +#14029 := [hypothesis]: #13938 +#8778 := (>= #8579 0::int) +#13651 := (or #4311 #8778) +#13643 := [quant-inst]: #13651 +#14030 := [unit-resolution #13643 #10596]: #8778 +#14031 := [th-lemma #8826 #14030 #14029]: false +#14034 := [lemma #14031]: #14033 +#14041 := [unit-resolution #14034 #8826]: #14032 +#13631 := (or #13653 #13938) +#13654 := [th-lemma]: #13631 +#11935 := [unit-resolution #13654 #14041]: #13653 +#13403 := (or #4433 #8645 #8665) +#8636 := (+ #1455 #8635) +#8637 := (+ #8579 #8636) +#8638 := (<= #8637 0::int) +#8641 := (or #8640 #8638) +#8642 := (not #8641) +#8643 := (= #8579 #2282) +#8644 := (or #8643 #8642) +#13406 := (or #4433 #8644) +#13514 := (iff #13406 #13403) +#8668 := (or #8645 #8665) +#12773 := (or #4433 #8668) +#13027 := (iff #12773 #13403) +#13487 := [rewrite]: #13027 +#12976 := (iff #13406 #12773) +#8669 := (iff #8644 #8668) +#8666 := (iff #8642 #8665) +#8663 := (iff #8641 #8662) +#8660 := (iff #8638 #8657) +#8648 := (+ #8579 #8635) +#8649 := (+ #1455 #8648) +#8652 := (<= #8649 0::int) +#8658 := (iff #8652 #8657) +#8659 := [rewrite]: #8658 +#8653 := (iff #8638 #8652) +#8650 := (= #8637 #8649) +#8651 := [rewrite]: #8650 +#8654 := [monotonicity #8651]: #8653 +#8661 := [trans #8654 #8659]: #8660 +#8664 := [monotonicity #8661]: #8663 +#8667 := [monotonicity #8664]: #8666 +#8646 := (iff #8643 #8645) +#8647 := [rewrite]: #8646 +#8670 := [monotonicity #8647 #8667]: #8669 +#13475 := [monotonicity #8670]: #12976 +#13534 := [trans #13475 #13487]: #13514 +#12762 := [quant-inst]: #13406 +#13535 := [mp #12762 #13534]: #13403 +#11121 := [unit-resolution #13535 #10727 #11935]: #8665 +#13548 := (or #8662 #13547) +#13558 := [def-axiom]: #13548 +#11762 := [unit-resolution #13558 #11121]: #13547 +#13559 := (not #8657) +#13562 := (or #8662 #13559) +#13563 := [def-axiom]: #13562 +#11759 := [unit-resolution #13563 #11121]: #13559 +#8702 := (or #8640 #8657 #8699) +#13329 := (or #4441 #8640 #8657 #8699) +#8691 := (+ #8634 #5075) +#8692 := (+ #188 #8691) +#8693 := (= #8692 0::int) +#8694 := (or #8640 #8638 #8693) +#13330 := (or #4441 #8694) +#13187 := (iff #13330 #13329) +#13332 := (or #4441 #8702) +#13183 := (iff #13332 #13329) +#13184 := [rewrite]: #13183 +#13355 := (iff #13330 #13332) +#8703 := (iff #8694 #8702) +#8700 := (iff #8693 #8699) +#8697 := (= #8692 #8696) +#8698 := [rewrite]: #8697 +#8701 := [monotonicity #8698]: #8700 +#8704 := [monotonicity #8661 #8701]: #8703 +#13107 := [monotonicity #8704]: #13355 +#13265 := [trans #13107 #13184]: #13187 +#13331 := [quant-inst]: #13330 +#13285 := [mp #13331 #13265]: #13329 +#14047 := [unit-resolution #13285 #10785]: #8702 +#11757 := [unit-resolution #14047 #11759 #11762]: #8699 +#14049 := (not #8699) +#14050 := (or #14049 #13182) +#14051 := [th-lemma]: #14050 +#11817 := [unit-resolution #14051 #11757]: #13182 +#13904 := (uf_1 #9695 ?x65!15) +#13630 := (uf_10 #13904) +#13902 := (* -1::int #13630) +#13836 := (+ #8634 #13902) +#14015 := (>= #13836 0::int) +#13832 := (= #8634 #13630) +#14026 := (= #13630 #8634) +#14022 := (= #13904 #8633) +#14023 := [monotonicity #10708]: #14022 +#14027 := [monotonicity #14023]: #14026 +#14028 := [symm #14027]: #13832 +#14035 := (not #13832) +#14036 := (or #14035 #14015) +#14037 := [th-lemma]: #14036 +#14038 := [unit-resolution #14037 #14028]: #14015 +#13835 := (>= #13630 0::int) +#13660 := (<= #13630 0::int) +#13692 := (not #13660) +#13628 := (= ?x65!15 #9695) +#13668 := (not #13628) +#8269 := (uf_6 uf_15 ?x65!15) +#8270 := (= uf_8 #8269) +#13741 := (ite #13628 #3895 #8270) +#13822 := (not #13741) +#13691 := (uf_6 #10323 ?x65!15) +#13734 := (= uf_8 #13691) +#13768 := (iff #13734 #13741) +#13775 := (or #4987 #13768) +#13690 := (ite #13628 #4958 #8270) +#13689 := (= #13691 uf_8) +#13739 := (iff #13689 #13690) +#13769 := (or #4987 #13739) +#13699 := (iff #13769 #13775) +#13700 := (iff #13775 #13775) +#13746 := [rewrite]: #13700 +#13770 := (iff #13739 #13768) +#13772 := (iff #13690 #13741) +#13773 := [monotonicity #4971]: #13772 +#13738 := (iff #13689 #13734) +#13740 := [rewrite]: #13738 +#13774 := [monotonicity #13740 #13773]: #13770 +#13776 := [monotonicity #13774]: #13699 +#13747 := [trans #13776 #13746]: #13699 +#13777 := [quant-inst]: #13769 +#13821 := [mp #13777 #13747]: #13775 +#14053 := [unit-resolution #13821 #4222]: #13768 +#13869 := (not #13734) +#5078 := (uf_6 uf_23 ?x65!15) +#5079 := (= uf_8 #5078) +#5080 := (not #5079) +#14062 := (iff #5080 #13869) +#14060 := (iff #5079 #13734) +#14058 := (iff #13734 #5079) +#14056 := (= #13691 #5078) +#14057 := [monotonicity #13583]: #14056 +#14059 := [monotonicity #14057]: #14058 +#14061 := [symm #14059]: #14060 +#14063 := [monotonicity #14061]: #14062 +#12807 := (or #4458 #5080 #8645) +#12793 := (or #8643 #5080) +#12815 := (or #4458 #12793) +#12826 := (iff #12815 #12807) +#12790 := (or #5080 #8645) +#12845 := (or #4458 #12790) +#12847 := (iff #12845 #12807) +#12848 := [rewrite]: #12847 +#12846 := (iff #12815 #12845) +#12791 := (iff #12793 #12790) +#12794 := (or #8645 #5080) +#12785 := (iff #12794 #12790) +#12787 := [rewrite]: #12785 +#12788 := (iff #12793 #12794) +#12789 := [monotonicity #8647]: #12788 +#12814 := [trans #12789 #12787]: #12791 +#12844 := [monotonicity #12814]: #12846 +#12827 := [trans #12844 #12848]: #12826 +#12816 := [quant-inst]: #12815 +#12828 := [mp #12816 #12827]: #12807 +#10280 := [unit-resolution #12828 #7708 #11935]: #5080 +#10395 := [mp #10280 #14063]: #13869 +#13829 := (not #13768) +#13844 := (or #13829 #13734 #13822) +#13874 := [def-axiom]: #13844 +#11835 := [unit-resolution #13874 #10395 #14053]: #13822 +#14066 := (or #13741 #13668) +#13827 := (or #13741 #13668 #3924) +#13840 := [def-axiom]: #13827 +#14067 := [unit-resolution #13840 #7677]: #14066 +#12190 := [unit-resolution #14067 #11835]: #13668 +#13657 := (or #13628 #13692) +#14009 := (or #10302 #13628 #13692) +#13626 := (= #9695 ?x65!15) +#13627 := (or #13626 #13692) +#14004 := (or #10302 #13627) +#14021 := (iff #14004 #14009) +#14011 := (or #10302 #13657) +#14014 := (iff #14011 #14009) +#14020 := [rewrite]: #14014 +#14012 := (iff #14004 #14011) +#13742 := (iff #13627 #13657) +#13642 := (iff #13626 #13628) +#13640 := [rewrite]: #13642 +#13743 := [monotonicity #13640]: #13742 +#14013 := [monotonicity #13743]: #14012 +#14024 := [trans #14013 #14020]: #14021 +#14010 := [quant-inst]: #14004 +#14025 := [mp #14010 #14024]: #14009 +#14069 := [unit-resolution #14025 #4249]: #13657 +#12191 := [unit-resolution #14069 #12190]: #13692 +#14071 := (or #13835 #13660) +#14072 := [th-lemma]: #14071 +#12192 := [unit-resolution #14072 #12191]: #13835 +#11148 := [th-lemma #12192 #14038 #8826 #11817 #12530]: false +#11885 := [lemma #11148]: #11869 +#19815 := [unit-resolution #11885 #25699]: #2283 +#4090 := (or #4543 #4537) +#4091 := [def-axiom]: #4090 +#19829 := [unit-resolution #4091 #25698]: #4537 +#19826 := (or #4540 #4534) +#10106 := (uf_1 #9695 uf_11) +#10107 := (uf_10 #10106) +#10111 := (* -1::int #10107) +#4883 := (uf_1 uf_22 uf_11) +#4884 := (uf_10 #4883) +#10686 := (+ #4884 #10111) +#10690 := (>= #10686 0::int) +#10685 := (= #4884 #10107) +#10711 := (= #10107 #4884) +#10709 := (= #10106 #4883) +#10710 := [monotonicity #10708]: #10709 +#10712 := [monotonicity #10710]: #10711 +#10713 := [symm #10712]: #10685 +#10714 := (not #10685) +#10715 := (or #10714 #10690) +#10716 := [th-lemma]: #10715 +#10717 := [unit-resolution #10716 #10713]: #10690 +#3952 := (<= #108 0::int) +#5799 := (or #1749 #3952) +#5800 := [th-lemma]: #5799 +#6367 := [unit-resolution #5800 #5498]: #3952 +#4802 := (?x47!7 uf_22) +#4803 := (uf_4 uf_14 #4802) +#4804 := (* -1::int #4803) +#4805 := (+ #188 #4804) +#4806 := (<= #4805 0::int) +#9262 := (not #4806) +#4814 := (uf_6 uf_15 #4802) +#4815 := (= uf_8 #4814) +#4816 := (not #4815) +#4807 := (uf_1 #4802 uf_22) +#4808 := (uf_10 #4807) +#4809 := (* -1::int #4808) +#4810 := (+ #4804 #4809) +#4811 := (+ #188 #4810) +#4812 := (= #4811 0::int) +#4813 := (not #4812) +#4824 := (or #4806 #4813 #4816) +#4827 := (not #4824) +#4821 := (= uf_11 uf_22) +#8243 := (not #4821) +#10613 := [hypothesis]: #1492 +#10629 := (or #8243 #217 #10190) +#10625 := (= #216 #108) +#10621 := (= #188 #108) +#4819 := (= uf_22 uf_11) +#10614 := [hypothesis]: #4821 +#10615 := [symm #10614]: #4819 +#10622 := [monotonicity #10615]: #10621 +#10623 := (= #216 #188) +#10616 := [hypothesis]: #4741 +#10620 := [symm #10616]: #10619 +#10617 := (= #216 #4740) +#10618 := [monotonicity #10614]: #10617 +#10624 := [trans #10618 #10620]: #10623 +#10626 := [trans #10624 #10622]: #10625 +#10627 := [trans #10626 #5498]: #217 +#10628 := [unit-resolution #10613 #10627]: false +#10630 := [lemma #10628]: #10629 +#10730 := [unit-resolution #10630 #10613 #10729]: #8243 +#10732 := (or #4821 #4827) +#4053 := (or #4567 #1657) +#4054 := [def-axiom]: #4053 +#10731 := [unit-resolution #4054 #10726]: #1657 +#8960 := (or #4344 #1656 #4821 #4827) +#4817 := (or #4816 #4813 #4806) +#4818 := (not #4817) +#4820 := (or #4819 #1656 #4818) +#8966 := (or #4344 #4820) +#9267 := (iff #8966 #8960) +#4833 := (or #1656 #4821 #4827) +#9153 := (or #4344 #4833) +#8906 := (iff #9153 #8960) +#9205 := [rewrite]: #8906 +#9156 := (iff #8966 #9153) +#4836 := (iff #4820 #4833) +#4830 := (or #4821 #1656 #4827) +#4834 := (iff #4830 #4833) +#4835 := [rewrite]: #4834 +#4831 := (iff #4820 #4830) +#4828 := (iff #4818 #4827) +#4825 := (iff #4817 #4824) +#4826 := [rewrite]: #4825 +#4829 := [monotonicity #4826]: #4828 +#4822 := (iff #4819 #4821) +#4823 := [rewrite]: #4822 +#4832 := [monotonicity #4823 #4829]: #4831 +#4837 := [trans #4832 #4835]: #4836 +#9157 := [monotonicity #4837]: #9156 +#9268 := [trans #9157 #9205]: #9267 +#9217 := [quant-inst]: #8966 +#9238 := [mp #9217 #9268]: #8960 +#10733 := [unit-resolution #9238 #9243 #10731]: #10732 +#10734 := [unit-resolution #10733 #10730]: #4827 +#9269 := (or #4824 #9262) +#9261 := [def-axiom]: #9269 +#10735 := [unit-resolution #9261 #10734]: #9262 +#6905 := (>= #4803 0::int) +#10502 := (not #6905) +#10503 := [hypothesis]: #10502 +#10442 := (or #4311 #6905) +#10443 := [quant-inst]: #10442 +#10607 := [unit-resolution #10443 #10596 #10503]: false +#10608 := [lemma #10607]: #6905 +#4888 := (* -1::int #4884) +#4889 := (+ #1455 #4888) +#4890 := (+ #108 #4889) +#4891 := (<= #4890 0::int) +#9338 := (not #4891) +#4892 := (+ uf_9 #4888) +#4893 := (<= #4892 0::int) +#4927 := (or #4891 #4893) +#4930 := (not #4927) +#4925 := (= #108 #216) +#10743 := (not #4925) +#10744 := (iff #1492 #10743) +#10741 := (iff #217 #4925) +#10739 := (iff #4925 #217) +#10738 := [commutativity]: #1490 +#10736 := (iff #4925 #788) +#10737 := [monotonicity #5498]: #10736 +#10740 := [trans #10737 #10738]: #10739 +#10742 := [symm #10740]: #10741 +#10745 := [monotonicity #10742]: #10744 +#10746 := [mp #10613 #10745]: #10743 +#4933 := (or #4925 #4930) +#9308 := (or #4433 #4925 #4930) +#4923 := (or #4893 #4891) +#4924 := (not #4923) +#4926 := (or #4925 #4924) +#9309 := (or #4433 #4926) +#9334 := (iff #9309 #9308) +#9329 := (or #4433 #4933) +#9332 := (iff #9329 #9308) +#9333 := [rewrite]: #9332 +#9330 := (iff #9309 #9329) +#4934 := (iff #4926 #4933) +#4931 := (iff #4924 #4930) +#4928 := (iff #4923 #4927) +#4929 := [rewrite]: #4928 +#4932 := [monotonicity #4929]: #4931 +#4935 := [monotonicity #4932]: #4934 +#9331 := [monotonicity #4935]: #9330 +#9335 := [trans #9331 #9333]: #9334 +#9310 := [quant-inst]: #9309 +#9336 := [mp #9310 #9335]: #9308 +#10747 := [unit-resolution #9336 #10727]: #4933 +#10748 := [unit-resolution #10747 #10746]: #4930 +#9321 := (or #4927 #9338) +#9322 := [def-axiom]: #9321 +#10749 := [unit-resolution #9322 #10748]: #9338 +#10647 := (>= #10107 0::int) +#9978 := (<= #10107 0::int) +#9979 := (not #9978) +#10042 := (= uf_11 #9695) +#10207 := (not #10042) +#10754 := (iff #8243 #10207) +#10752 := (iff #4821 #10042) +#10750 := (iff #10042 #4821) +#10751 := [monotonicity #10708]: #10750 +#10753 := [symm #10751]: #10752 +#10755 := [monotonicity #10753]: #10754 +#10756 := [mp #10730 #10755]: #10207 +#10049 := (or #9979 #10042) +#10397 := (or #10302 #9979 #10042) +#10035 := (= #9695 uf_11) +#10036 := (or #10035 #9979) +#10418 := (or #10302 #10036) +#10633 := (iff #10418 #10397) +#10308 := (or #10302 #10049) +#10631 := (iff #10308 #10397) +#10632 := [rewrite]: #10631 +#10609 := (iff #10418 #10308) +#10065 := (iff #10036 #10049) +#10046 := (or #10042 #9979) +#10050 := (iff #10046 #10049) +#10051 := [rewrite]: #10050 +#10047 := (iff #10036 #10046) +#10044 := (iff #10035 #10042) +#10045 := [rewrite]: #10044 +#10048 := [monotonicity #10045]: #10047 +#10165 := [trans #10048 #10051]: #10065 +#10612 := [monotonicity #10165]: #10609 +#10634 := [trans #10612 #10632]: #10633 +#10413 := [quant-inst]: #10418 +#10635 := [mp #10413 #10634]: #10397 +#10757 := [unit-resolution #10635 #4249]: #10049 +#10758 := [unit-resolution #10757 #10756]: #9979 +#10759 := (or #10647 #9978) +#10760 := [th-lemma]: #10759 +#10761 := [unit-resolution #10760 #10758]: #10647 +#10762 := [th-lemma #10761 #10749 #10608 #10735 #6367 #10717]: false +#10763 := [lemma #10762]: #217 +#4100 := (or #4540 #1492 #4534) +#4086 := [def-axiom]: #4100 +#19860 := [unit-resolution #4086 #10763]: #19826 +#19861 := [unit-resolution #19860 #19829]: #4534 +#4109 := (or #4531 #4525) +#4093 := [def-axiom]: #4109 +#19854 := [unit-resolution #4093 #19861]: #4525 +#4106 := (or #4528 #2284 #4522) +#4107 := [def-axiom]: #4106 +#19851 := [unit-resolution #4107 #19854 #19815]: #4522 +#4101 := (or #4519 #4513) +#4103 := [def-axiom]: #4101 +#19863 := [unit-resolution #4103 #19851]: #4513 +#4123 := (or #4516 #3453 #4510) +#4110 := [def-axiom]: #4123 +#19864 := [unit-resolution #4110 #19863]: #4513 +#19859 := [unit-resolution #19864 #19217]: #3453 +#4134 := (or #3448 #4133) +#4135 := [def-axiom]: #4134 +#19869 := [unit-resolution #4135 #19859]: #4133 +#4148 := (or #3448 #2304) +#3989 := [def-axiom]: #4148 +#19866 := [unit-resolution #3989 #19859]: #2304 +#3990 := (or #3448 #2307) +#3991 := [def-axiom]: #3990 +#19868 := [unit-resolution #3991 #19859]: #2307 +#17736 := (or #3433 #2896 #2306) +#12004 := [hypothesis]: #4133 +#6675 := (uf_1 uf_22 ?x68!16) +#6676 := (uf_10 #6675) +#6701 := (+ #2894 #6676) +#6702 := (+ #188 #6701) +#16997 := (<= #6702 0::int) +#6705 := (= #6702 0::int) +#6642 := (uf_4 uf_14 ?x68!16) +#6659 := (* -1::int #6642) +#6694 := (+ #6659 #6676) +#6695 := (+ #188 #6694) +#6696 := (>= #6695 0::int) +#6680 := (* -1::int #6676) +#6684 := (+ uf_9 #6680) +#6685 := (<= #6684 0::int) +#6731 := (or #6685 #6696) +#6734 := (not #6731) +#6728 := (= #2300 #6642) +#14098 := (not #6728) +#6660 := (+ #2300 #6659) +#17022 := (>= #6660 0::int) +#14117 := (not #17022) +#6472 := (+ #188 #6659) +#6473 := (<= #6472 0::int) +#6496 := (uf_6 uf_15 ?x68!16) +#6497 := (= uf_8 #6496) +#16679 := (not #6497) +#13671 := (= ?x68!16 #9695) +#13592 := (ite #13671 #3895 #6497) +#15564 := (not #13592) +#13670 := (uf_6 #10323 ?x68!16) +#13649 := (= uf_8 #13670) +#13820 := (iff #13592 #13649) +#14101 := (or #4987 #13820) +#13645 := (ite #13671 #4958 #6497) +#13647 := (= #13670 uf_8) +#13648 := (iff #13647 #13645) +#15296 := (or #4987 #13648) +#15178 := (iff #15296 #14101) +#15561 := (iff #14101 #14101) +#15556 := [rewrite]: #15561 +#13826 := (iff #13648 #13820) +#13663 := (iff #13649 #13592) +#13819 := (iff #13663 #13820) +#13825 := [rewrite]: #13819 +#13664 := (iff #13648 #13663) +#13638 := (iff #13645 #13592) +#13662 := [monotonicity #4971]: #13638 +#13655 := (iff #13647 #13649) +#13661 := [rewrite]: #13655 +#13665 := [monotonicity #13661 #13662]: #13664 +#13828 := [trans #13665 #13825]: #13826 +#15560 := [monotonicity #13828]: #15178 +#15558 := [trans #15560 #15556]: #15178 +#15177 := [quant-inst]: #15296 +#15563 := [mp #15177 #15558]: #14101 +#17427 := [unit-resolution #15563 #4222]: #13820 +#16690 := (not #13649) +#17349 := (iff #2307 #16690) +#17657 := (iff #2306 #13649) +#17654 := (iff #13649 #2306) +#17353 := (= #13670 #2305) +#17653 := [monotonicity #13583]: #17353 +#17655 := [monotonicity #17653]: #17654 +#17300 := [symm #17655]: #17657 +#17434 := [monotonicity #17300]: #17349 +#17656 := [hypothesis]: #2307 +#17463 := [mp #17656 #17434]: #16690 +#16689 := (not #13820) +#16795 := (or #16689 #15564 #13649) +#16793 := [def-axiom]: #16795 +#17402 := [unit-resolution #16793 #17463 #17427]: #15564 +#15565 := (not #13671) +#17466 := (or #13592 #15565) +#15618 := (or #13592 #15565 #3924) +#15814 := [def-axiom]: #15618 +#17660 := [unit-resolution #15814 #7677]: #17466 +#17661 := [unit-resolution #17660 #17402]: #15565 +#16681 := (or #13592 #13671 #16679) +#16685 := [def-axiom]: #16681 +#17658 := [unit-resolution #16685 #17661 #17402]: #16679 +#6530 := (or #6473 #6497) +#17091 := (or #4423 #6473 #6497) +#6493 := (+ #6642 #1455) +#6494 := (>= #6493 0::int) +#6495 := (or #6497 #6494) +#17092 := (or #4423 #6495) +#16354 := (iff #17092 #17091) +#17099 := (or #4423 #6530) +#16352 := (iff #17099 #17091) +#16353 := [rewrite]: #16352 +#16351 := (iff #17092 #17099) +#6533 := (iff #6495 #6530) +#6527 := (or #6497 #6473) +#6531 := (iff #6527 #6530) +#6532 := [rewrite]: #6531 +#6528 := (iff #6495 #6527) +#6525 := (iff #6494 #6473) +#6467 := (+ #1455 #6642) +#6469 := (>= #6467 0::int) +#6474 := (iff #6469 #6473) +#6524 := [rewrite]: #6474 +#6470 := (iff #6494 #6469) +#6468 := (= #6493 #6467) +#6466 := [rewrite]: #6468 +#6471 := [monotonicity #6466]: #6470 +#6526 := [trans #6471 #6524]: #6525 +#6529 := [monotonicity #6526]: #6528 +#6534 := [trans #6529 #6532]: #6533 +#17076 := [monotonicity #6534]: #16351 +#16356 := [trans #17076 #16353]: #16354 +#17101 := [quant-inst]: #17092 +#16358 := [mp #17101 #16356]: #17091 +#17669 := [unit-resolution #16358 #17666]: #6530 +#17659 := [unit-resolution #17669 #17658]: #6473 +#6393 := (+ #2298 #4781) +#18006 := (<= #6393 0::int) +#17296 := (= #2298 #4740) +#6449 := (= ?x67!17 uf_22) +#14077 := (= ?x67!17 #9695) +#6439 := (uf_6 uf_15 ?x67!17) +#6440 := (= uf_8 #6439) +#14085 := (ite #14077 #3895 #6440) +#14079 := (uf_6 #10323 ?x67!17) +#14082 := (= uf_8 #14079) +#14088 := (iff #14082 #14085) +#16842 := (or #4987 #14088) +#14078 := (ite #14077 #4958 #6440) +#14080 := (= #14079 uf_8) +#14081 := (iff #14080 #14078) +#16843 := (or #4987 #14081) +#16827 := (iff #16843 #16842) +#16829 := (iff #16842 #16842) +#16830 := [rewrite]: #16829 +#14089 := (iff #14081 #14088) +#14086 := (iff #14078 #14085) +#14087 := [monotonicity #4971]: #14086 +#14083 := (iff #14080 #14082) +#14084 := [rewrite]: #14083 +#14090 := [monotonicity #14084 #14087]: #14089 +#16828 := [monotonicity #14090]: #16827 +#16825 := [trans #16828 #16830]: #16827 +#16822 := [quant-inst]: #16843 +#16780 := [mp #16822 #16825]: #16842 +#17670 := [unit-resolution #16780 #4222]: #14088 +#17684 := (= #2303 #14079) +#17672 := (= #14079 #2303) +#17688 := [monotonicity #13583]: #17672 +#17680 := [symm #17688]: #17684 +#17671 := [hypothesis]: #2304 +#17705 := [trans #17671 #17680]: #14082 +#16844 := (not #14082) +#16848 := (not #14088) +#16851 := (or #16848 #16844 #14085) +#16852 := [def-axiom]: #16851 +#17703 := [unit-resolution #16852 #17705 #17670]: #14085 +#16816 := (not #6440) +#6405 := (uf_4 uf_14 ?x67!17) +#17321 := (+ #6405 #9707) +#17316 := (<= #17321 0::int) +#14110 := (not #17316) +#14118 := (not #6473) +#17620 := (or #14110 #2896 #13671 #14118) +#17021 := (not #6685) +#5574 := (* -1::int #6405) +#5674 := (+ #2298 #5574) +#5698 := (<= #5674 0::int) +#16850 := (or #4449 #5698) +#5667 := (+ #6405 #2299) +#5668 := (>= #5667 0::int) +#16871 := (or #4449 #5668) +#16875 := (iff #16871 #16850) +#16878 := (iff #16850 #16850) +#16872 := [rewrite]: #16878 +#5701 := (iff #5668 #5698) +#5669 := (+ #2299 #6405) +#5671 := (>= #5669 0::int) +#5699 := (iff #5671 #5698) +#5700 := [rewrite]: #5699 +#5672 := (iff #5668 #5671) +#5664 := (= #5667 #5669) +#5670 := [rewrite]: #5664 +#5673 := [monotonicity #5670]: #5672 +#5702 := [trans #5673 #5700]: #5701 +#16877 := [monotonicity #5702]: #16875 +#16845 := [trans #16877 #16872]: #16875 +#16831 := [quant-inst]: #16871 +#16858 := [mp #16831 #16845]: #16850 +#12150 := [unit-resolution #16858 #10788]: #5698 +#14116 := [hypothesis]: #6473 +#14115 := [hypothesis]: #17316 +#14119 := (not #10581) +#14113 := (not #5698) +#14120 := (or #14117 #14118 #14113 #2896 #14110 #14119) +#14127 := [th-lemma]: #14120 +#14128 := [unit-resolution #14127 #14115 #14116 #12150 #12004 #12925]: #14117 +#14167 := (or #14098 #17022) +#14099 := [th-lemma]: #14167 +#14170 := [unit-resolution #14099 #14128]: #14098 +#6737 := (or #6728 #6734) +#17003 := (or #4433 #6728 #6734) +#6681 := (+ #1455 #6680) +#6682 := (+ #6642 #6681) +#6683 := (<= #6682 0::int) +#6724 := (or #6685 #6683) +#6725 := (not #6724) +#6726 := (= #6642 #2300) +#6727 := (or #6726 #6725) +#17006 := (or #4433 #6727) +#17015 := (iff #17006 #17003) +#16994 := (or #4433 #6737) +#17013 := (iff #16994 #17003) +#17012 := [rewrite]: #17013 +#17007 := (iff #17006 #16994) +#6738 := (iff #6727 #6737) +#6735 := (iff #6725 #6734) +#6732 := (iff #6724 #6731) +#6699 := (iff #6683 #6696) +#6687 := (+ #6642 #6680) +#6688 := (+ #1455 #6687) +#6691 := (<= #6688 0::int) +#6697 := (iff #6691 #6696) +#6698 := [rewrite]: #6697 +#6692 := (iff #6683 #6691) +#6689 := (= #6682 #6688) +#6690 := [rewrite]: #6689 +#6693 := [monotonicity #6690]: #6692 +#6700 := [trans #6693 #6698]: #6699 +#6733 := [monotonicity #6700]: #6732 +#6736 := [monotonicity #6733]: #6735 +#6729 := (iff #6726 #6728) +#6730 := [rewrite]: #6729 +#6739 := [monotonicity #6730 #6736]: #6738 +#17008 := [monotonicity #6739]: #17007 +#17016 := [trans #17008 #17012]: #17015 +#17005 := [quant-inst]: #17006 +#17017 := [mp #17005 #17016]: #17003 +#14323 := [unit-resolution #17017 #10727]: #6737 +#14462 := [unit-resolution #14323 #14170]: #6734 +#17024 := (or #6731 #17021) +#17014 := [def-axiom]: #17024 +#14131 := [unit-resolution #17014 #14462]: #17021 +#17023 := (not #6696) +#17025 := (or #6731 #17023) +#17026 := [def-axiom]: #17025 +#14463 := [unit-resolution #17026 #14462]: #17023 +#6708 := (or #6685 #6696 #6705) +#16986 := (or #4441 #6685 #6696 #6705) +#6677 := (+ #6676 #2894) +#6678 := (+ #188 #6677) +#6679 := (= #6678 0::int) +#6686 := (or #6685 #6683 #6679) +#16942 := (or #4441 #6686) +#16967 := (iff #16942 #16986) +#16939 := (or #4441 #6708) +#16965 := (iff #16939 #16986) +#16966 := [rewrite]: #16965 +#16958 := (iff #16942 #16939) +#6709 := (iff #6686 #6708) +#6706 := (iff #6679 #6705) +#6703 := (= #6678 #6702) +#6704 := [rewrite]: #6703 +#6707 := [monotonicity #6704]: #6706 +#6710 := [monotonicity #6700 #6707]: #6709 +#16981 := [monotonicity #6710]: #16958 +#16968 := [trans #16981 #16966]: #16967 +#16943 := [quant-inst]: #16942 +#16957 := [mp #16943 #16968]: #16986 +#14091 := [unit-resolution #16957 #10785]: #6708 +#15021 := [unit-resolution #14091 #14463 #14131]: #6705 +#15016 := (not #6705) +#15301 := (or #15016 #16997) +#16191 := [th-lemma]: #15301 +#16210 := [unit-resolution #16191 #15021]: #16997 +#17555 := (uf_1 #9695 ?x68!16) +#17556 := (uf_10 #17555) +#11882 := (* -1::int #17556) +#11971 := (+ #6676 #11882) +#12082 := (>= #11971 0::int) +#12012 := (= #6676 #17556) +#16213 := (= #17556 #6676) +#15939 := (= #17555 #6675) +#16211 := [monotonicity #10708]: #15939 +#16214 := [monotonicity #16211]: #16213 +#16541 := [symm #16214]: #12012 +#16559 := (not #12012) +#16520 := (or #16559 #12082) +#16540 := [th-lemma]: #16520 +#16560 := [unit-resolution #16540 #16541]: #12082 +#17561 := (<= #17556 0::int) +#17562 := (not #17561) +#16607 := [hypothesis]: #15565 +#17589 := (or #10302 #13671 #17562) +#17574 := (= #9695 ?x68!16) +#17571 := (or #17574 #17562) +#17590 := (or #10302 #17571) +#11861 := (iff #17590 #17589) +#17587 := (or #13671 #17562) +#17606 := (or #10302 #17587) +#17600 := (iff #17606 #17589) +#17601 := [rewrite]: #17600 +#17593 := (iff #17590 #17606) +#17569 := (iff #17571 #17587) +#17585 := (iff #17574 #13671) +#17586 := [rewrite]: #17585 +#17588 := [monotonicity #17586]: #17569 +#17599 := [monotonicity #17588]: #17593 +#17598 := [trans #17599 #17601]: #11861 +#17583 := [quant-inst]: #17590 +#11866 := [mp #17583 #17598]: #17589 +#17621 := [unit-resolution #11866 #4249 #16607]: #17562 +#17166 := [th-lemma #12150 #12004 #14115 #12925 #17621 #16560 #16210]: false +#17622 := [lemma #17166]: #17620 +#17722 := [unit-resolution #17622 #17661 #12004 #17659]: #14110 +#18133 := (or #16816 #17316) +#18126 := [hypothesis]: #14110 +#18130 := [hypothesis]: #6440 +#18002 := (or #4328 #16816 #9686 #17316) +#17278 := (+ #9701 #5574) +#17277 := (>= #17278 0::int) +#17276 := (or #9686 #16816 #17277) +#18007 := (or #4328 #17276) +#17984 := (iff #18007 #18002) +#17337 := (or #16816 #9686 #17316) +#17982 := (or #4328 #17337) +#17989 := (iff #17982 #18002) +#17990 := [rewrite]: #17989 +#17983 := (iff #18007 #17982) +#17345 := (iff #17276 #17337) +#17333 := (or #9686 #16816 #17316) +#17346 := (iff #17333 #17337) +#17341 := [rewrite]: #17346 +#17336 := (iff #17276 #17333) +#17332 := (iff #17277 #17316) +#17303 := (+ #5574 #9701) +#17297 := (>= #17303 0::int) +#17320 := (iff #17297 #17316) +#17326 := [rewrite]: #17320 +#17306 := (iff #17277 #17297) +#17304 := (= #17278 #17303) +#17305 := [rewrite]: #17304 +#17317 := [monotonicity #17305]: #17306 +#17325 := [trans #17317 #17326]: #17332 +#17324 := [monotonicity #17325]: #17336 +#17347 := [trans #17324 #17341]: #17345 +#17971 := [monotonicity #17347]: #17983 +#17988 := [trans #17971 #17990]: #17984 +#17980 := [quant-inst]: #18007 +#17991 := [mp #17980 #17988]: #18002 +#18132 := [unit-resolution #17991 #13967 #13969 #18130 #18126]: false +#18134 := [lemma #18132]: #18133 +#17723 := [unit-resolution #18134 #17722]: #16816 +#16801 := (not #14085) +#16815 := (or #16801 #14077 #6440) +#16832 := [def-axiom]: #16815 +#17724 := [unit-resolution #16832 #17723 #17703]: #14077 +#17720 := [trans #17724 #10708]: #6449 +#17697 := [monotonicity #17720]: #17296 +#17700 := (not #17296) +#17701 := (or #17700 #18006) +#17693 := [th-lemma]: #17701 +#17673 := [unit-resolution #17693 #17697]: #18006 +#18044 := (not #18006) +#18045 := (or #18044 #14117 #2896 #14118) +#18043 := [hypothesis]: #18006 +#18041 := [hypothesis]: #17022 +#18042 := [th-lemma #18041 #12004 #14116 #9311 #18043]: false +#18046 := [lemma #18042]: #18045 +#17699 := [unit-resolution #18046 #17673 #12004 #17659]: #14117 +#17676 := [unit-resolution #14099 #17699]: #14098 +#17677 := [unit-resolution #14323 #17676]: #6734 +#17164 := (or #6731 #6705) +#17180 := [hypothesis]: #15016 +#17209 := [hypothesis]: #6734 +#17210 := [unit-resolution #17014 #17209]: #17021 +#17207 := [unit-resolution #17026 #17209]: #17023 +#17212 := [unit-resolution #14091 #17207 #17210 #17180]: false +#17211 := [lemma #17212]: #17164 +#17678 := [unit-resolution #17211 #17677]: #6705 +#17681 := [unit-resolution #16191 #17678]: #16997 +#17675 := [unit-resolution #11866 #4249 #17661]: #17562 +#17735 := [th-lemma #17673 #9311 #17675 #16560 #17681 #12004]: false +#17721 := [lemma #17735]: #17736 +[unit-resolution #17721 #19868 #19866 #19869]: false unsat diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Boogie/Examples/cert/Boogie_max --- a/src/HOL/Boogie/Examples/cert/Boogie_max Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Boogie/Examples/cert/Boogie_max Sat Nov 14 09:40:27 2009 +0100 @@ -1,19 +1,19 @@ (benchmark Isabelle :extrafuns ( - (uf_5 Int) - (uf_7 Int) + (uf_4 Int) (uf_11 Int) - (uf_4 Int) + (uf_7 Int) + (uf_5 Int) + (uf_13 Int) (uf_9 Int) - (uf_13 Int) (uf_2 Int) (uf_6 Int) (uf_10 Int) + (uf_12 Int) (uf_8 Int) - (uf_12 Int) (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 (<= 1 1) (and (<= 1 1) (and (<= 0 0) (<= 0 0)))) (and (implies (forall (?x1 Int) (implies (and (< ?x1 1) (<= 0 ?x1)) (<= (uf_3 ?x1) uf_2))) (and (implies (= (uf_3 0) uf_2) (implies true (implies (and (<= 1 uf_4) (<= 0 uf_5)) (implies (forall (?x2 Int) (implies (and (< ?x2 uf_4) (<= 0 ?x2)) (<= (uf_3 ?x2) uf_6))) (implies (= (uf_3 uf_5) uf_6) (implies (and (<= 1 uf_4) (<= 0 uf_5)) (and (implies true (implies (and (<= 1 uf_4) (<= 0 uf_5)) (implies (< uf_4 uf_1) (implies (and (<= 1 uf_4) (<= 0 uf_5)) (and (implies true (implies (and (<= 1 uf_4) (<= 0 uf_5)) (implies (<= (uf_3 uf_4) uf_6) (implies (and (<= 1 uf_4) (<= 0 uf_5)) (implies true (implies (= uf_7 uf_5) (implies (= uf_8 uf_6) (implies true (implies (and (<= 1 uf_4) (<= 0 uf_7)) (implies (= uf_9 (+ uf_4 1)) (implies (and (<= 2 uf_9) (<= 0 uf_7)) (implies true (and (implies (forall (?x3 Int) (implies (and (< ?x3 uf_9) (<= 0 ?x3)) (<= (uf_3 ?x3) uf_8))) (and (implies (= (uf_3 uf_7) uf_8) (implies false true)) (= (uf_3 uf_7) uf_8))) (forall (?x4 Int) (implies (and (< ?x4 uf_9) (<= 0 ?x4)) (<= (uf_3 ?x4) uf_8)))))))))))))))) (implies true (implies (and (<= 1 uf_4) (<= 0 uf_5)) (implies (< uf_6 (uf_3 uf_4)) (implies (= uf_10 (uf_3 uf_4)) (implies (and (<= 1 uf_4) (<= 1 uf_4)) (implies true (implies (= uf_7 uf_4) (implies (= uf_8 uf_10) (implies true (implies (and (<= 1 uf_4) (<= 0 uf_7)) (implies (= uf_9 (+ uf_4 1)) (implies (and (<= 2 uf_9) (<= 0 uf_7)) (implies true (and (implies (forall (?x5 Int) (implies (and (< ?x5 uf_9) (<= 0 ?x5)) (<= (uf_3 ?x5) uf_8))) (and (implies (= (uf_3 uf_7) uf_8) (implies false true)) (= (uf_3 uf_7) uf_8))) (forall (?x6 Int) (implies (and (< ?x6 uf_9) (<= 0 ?x6)) (<= (uf_3 ?x6) uf_8)))))))))))))))))))))) (implies true (implies (and (<= 1 uf_4) (<= 0 uf_5)) (implies (<= uf_1 uf_4) (implies (and (<= 1 uf_4) (<= 0 uf_5)) (implies true (implies (= uf_11 uf_5) (implies (= uf_12 uf_6) (implies (= uf_13 uf_4) (implies true (and (implies (exists (?x7 Int) (implies (and (< ?x7 uf_1) (<= 0 ?x7)) (= (uf_3 ?x7) uf_12))) (and (implies (forall (?x8 Int) (implies (and (< ?x8 uf_1) (<= 0 ?x8)) (<= (uf_3 ?x8) uf_12))) true) (forall (?x9 Int) (implies (and (< ?x9 uf_1) (<= 0 ?x9)) (<= (uf_3 ?x9) uf_12))))) (exists (?x10 Int) (implies (and (< ?x10 uf_1) (<= 0 ?x10)) (= (uf_3 ?x10) uf_12)))))))))))))))))))) (= (uf_3 0) uf_2))) (forall (?x11 Int) (implies (and (< ?x11 1) (<= 0 ?x11)) (<= (uf_3 ?x11) uf_2)))))))))) +: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)))))))))))))))))))))))))))))))))))))) :formula true ) diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Boogie/Examples/cert/Boogie_max.proof --- a/src/HOL/Boogie/Examples/cert/Boogie_max.proof Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Boogie/Examples/cert/Boogie_max.proof Sat Nov 14 09:40:27 2009 +0100 @@ -1,361 +1,258 @@ #2 := false #4 := 0::int decl uf_3 :: (-> int int) -decl ?x3!1 :: int -#1188 := ?x3!1 -#1195 := (uf_3 ?x3!1) -#760 := -1::int -#1381 := (* -1::int #1195) -decl uf_4 :: int -#25 := uf_4 -#39 := (uf_3 uf_4) -#2763 := (+ #39 #1381) -#2765 := (>= #2763 0::int) -#2762 := (= #39 #1195) -#2631 := (= uf_4 ?x3!1) -#1394 := (* -1::int ?x3!1) -#2564 := (+ uf_4 #1394) -#2628 := (>= #2564 0::int) -decl uf_9 :: int -#47 := uf_9 -#806 := (* -1::int uf_9) -#838 := (+ uf_4 #806) -#1977 := (>= #838 -1::int) -#837 := (= #838 -1::int) -#1395 := (+ uf_9 #1394) -#1396 := (<= #1395 0::int) -decl uf_8 :: int -#43 := uf_8 -#1382 := (+ uf_8 #1381) -#1383 := (>= #1382 0::int) -#1192 := (>= ?x3!1 0::int) -#1625 := (not #1192) -#1640 := (or #1625 #1383 #1396) -#1645 := (not #1640) -#16 := (:var 0 int) -#20 := (uf_3 #16) -#2303 := (pattern #20) -#807 := (+ #16 #806) -#805 := (>= #807 0::int) -#799 := (* -1::int uf_8) -#800 := (+ #20 #799) -#801 := (<= #800 0::int) -#753 := (>= #16 0::int) -#1548 := (not #753) -#1607 := (or #1548 #801 #805) -#2320 := (forall (vars (?x3 int)) (:pat #2303) #1607) -#2325 := (not #2320) -decl uf_7 :: int -#41 := uf_7 -#58 := (uf_3 uf_7) -#221 := (= uf_8 #58) -#2328 := (or #221 #2325) -#2331 := (not #2328) -#2334 := (or #2331 #1645) -#2337 := (not #2334) -#851 := (* -1::int #39) -decl uf_6 :: int -#32 := uf_6 -#852 := (+ uf_6 #851) -#850 := (>= #852 0::int) -#840 := (not #837) -#50 := 2::int -#791 := (>= uf_9 2::int) -#1656 := (not #791) -#788 := (>= uf_7 0::int) -#1655 := (not #788) -decl uf_5 :: int -#27 := uf_5 -#779 := (>= uf_5 0::int) -#1654 := (not #779) -#10 := 1::int -#776 := (>= uf_4 1::int) -#886 := (not #776) -#361 := (= uf_4 uf_7) -#376 := (not #361) -decl uf_10 :: int -#78 := uf_10 -#356 := (= #39 uf_10) -#401 := (not #356) -#82 := (= uf_8 uf_10) -#367 := (not #82) -#2346 := (or #367 #401 #376 #886 #1654 #1655 #1656 #840 #850 #2337) -#2349 := (not #2346) -#854 := (not #850) -#194 := (= uf_6 uf_8) -#301 := (not #194) -#191 := (= uf_5 uf_7) -#310 := (not #191) -#2340 := (or #310 #301 #886 #1654 #1655 #1656 #840 #854 #2337) -#2343 := (not #2340) -#2352 := (or #2343 #2349) -#2355 := (not #2352) -#924 := (* -1::int uf_4) -decl uf_1 :: int -#5 := uf_1 -#925 := (+ uf_1 #924) -#926 := (<= #925 0::int) -#2358 := (or #886 #1654 #926 #2355) -#2361 := (not #2358) -decl ?x7!2 :: int -#1279 := ?x7!2 -#1287 := (uf_3 ?x7!2) -decl uf_12 :: int -#99 := uf_12 -#1469 := (= uf_12 #1287) -#1284 := (>= ?x7!2 0::int) -#1711 := (not #1284) -#1280 := (* -1::int ?x7!2) -#1281 := (+ uf_1 #1280) -#1282 := (<= #1281 0::int) -#1726 := (or #1282 #1711 #1469) -#1757 := (not #1726) -decl ?x8!3 :: int -#1297 := ?x8!3 -#1298 := (uf_3 ?x8!3) -#1493 := (* -1::int #1298) -#1494 := (+ uf_12 #1493) -#1495 := (>= #1494 0::int) -#1305 := (>= ?x8!3 0::int) -#1731 := (not #1305) -#1301 := (* -1::int ?x8!3) -#1302 := (+ uf_1 #1301) -#1303 := (<= #1302 0::int) -#1795 := (or #1303 #1731 #1495 #1757) -#1798 := (not #1795) -#951 := (* -1::int #16) -#952 := (+ uf_1 #951) -#953 := (<= #952 0::int) -#105 := (= #20 uf_12) -#1700 := (or #105 #1548 #953) -#1705 := (not #1700) -#2364 := (forall (vars (?x7 int)) (:pat #2303) #1705) -#2369 := (or #2364 #1798) -#2372 := (not #2369) -#927 := (not #926) -decl uf_13 :: int -#101 := uf_13 -#472 := (= uf_4 uf_13) -#542 := (not #472) -#469 := (= uf_6 uf_12) -#551 := (not #469) -decl uf_11 :: int -#97 := uf_11 -#466 := (= uf_5 uf_11) -#560 := (not #466) -#2375 := (or #560 #551 #542 #886 #1654 #927 #2372) -#2378 := (not #2375) -#2381 := (or #2361 #2378) -#2384 := (not #2381) -#1030 := (+ #16 #924) -#1029 := (>= #1030 0::int) -#1024 := (* -1::int uf_6) -#1025 := (+ #20 #1024) -#1026 := (<= #1025 0::int) -#1585 := (or #1548 #1026 #1029) -#2312 := (forall (vars (?x2 int)) (:pat #2303) #1585) -#2317 := (not #2312) -#763 := (* -1::int #20) +#8 := (uf_3 0::int) +#714 := -1::int +#2117 := (* -1::int #8) decl uf_2 :: int #7 := uf_2 -#764 := (+ uf_2 #763) -#762 := (>= #764 0::int) -#749 := (>= #16 1::int) -#1563 := (or #749 #1548 #762) -#2304 := (forall (vars (?x1 int)) (:pat #2303) #1563) -#2309 := (not #2304) -#36 := (uf_3 uf_5) -#188 := (= uf_6 #36) -#619 := (not #188) -#2387 := (or #619 #886 #1654 #2309 #2317 #2384) -#2390 := (not #2387) -decl ?x1!0 :: int -#1152 := ?x1!0 -#1156 := (>= ?x1!0 1::int) -#1155 := (>= ?x1!0 0::int) -#1164 := (not #1155) -#1153 := (uf_3 ?x1!0) -#1150 := (* -1::int #1153) -#1151 := (+ uf_2 #1150) -#1154 := (>= #1151 0::int) -#1540 := (or #1154 #1164 #1156) -#2202 := (= uf_2 #1153) -#8 := (uf_3 0::int) -#2191 := (= #8 #1153) -#2188 := (= #1153 #8) -#2207 := (= ?x1!0 0::int) -#1157 := (not #1156) -#1545 := (not #1540) -#2205 := [hypothesis]: #1545 -#1976 := (or #1540 #1157) -#1967 := [def-axiom]: #1976 -#2206 := [unit-resolution #1967 #2205]: #1157 -#1975 := (or #1540 #1155) -#1890 := [def-axiom]: #1975 -#2203 := [unit-resolution #1890 #2205]: #1155 -#2187 := [th-lemma #2203 #2206]: #2207 -#2190 := [monotonicity #2187]: #2188 -#2192 := [symm #2190]: #2191 +#2122 := (+ uf_2 #2117) +#2118 := (>= #2122 0::int) #9 := (= uf_2 #8) -#1078 := (<= uf_1 0::int) -#1031 := (not #1029) -#1034 := (and #753 #1031) -#1037 := (not #1034) -#1040 := (or #1026 #1037) -#1043 := (forall (vars (?x2 int)) #1040) -#1046 := (not #1043) -#972 := (* -1::int uf_12) -#973 := (+ #20 #972) -#974 := (<= #973 0::int) -#954 := (not #953) -#957 := (and #753 #954) -#960 := (not #957) -#980 := (or #960 #974) -#985 := (forall (vars (?x8 int)) #980) -#963 := (or #105 #960) -#966 := (exists (vars (?x7 int)) #963) -#969 := (not #966) -#988 := (or #969 #985) -#991 := (and #966 #988) -#781 := (and #776 #779) -#784 := (not #781) -#1016 := (or #560 #551 #542 #784 #927 #991) -#843 := (and #776 #788) -#846 := (not #843) -#804 := (not #805) -#810 := (and #753 #804) -#813 := (not #810) -#816 := (or #801 #813) -#819 := (forall (vars (?x3 int)) #816) -#822 := (not #819) -#828 := (or #221 #822) -#833 := (and #819 #828) -#793 := (and #788 #791) -#796 := (not #793) -#916 := (or #367 #401 #376 #886 #784 #796 #833 #840 #846 #850) -#881 := (or #310 #301 #784 #796 #833 #840 #846 #854) -#921 := (and #881 #916) -#946 := (or #784 #921 #926) -#1021 := (and #946 #1016) -#652 := (not #9) -#1064 := (or #652 #619 #784 #1021 #1046) -#1069 := (and #9 #1064) -#747 := (not #749) -#754 := (and #747 #753) -#757 := (not #754) -#766 := (or #757 #762) -#769 := (forall (vars (?x1 int)) #766) -#772 := (not #769) -#1072 := (or #772 #1069) -#1075 := (and #769 #1072) -#1098 := (or #652 #1075 #1078) -#1103 := (not #1098) -#21 := (<= #20 uf_2) -#18 := (<= 0::int #16) -#17 := (< #16 1::int) -#19 := (and #17 #18) -#22 := (implies #19 #21) -#23 := (forall (vars (?x1 int)) #22) -#24 := (= #8 uf_2) -#103 := (< #16 uf_1) -#104 := (and #103 #18) -#106 := (implies #104 #105) -#107 := (exists (vars (?x7 int)) #106) -#108 := (<= #20 uf_12) -#109 := (implies #104 #108) -#110 := (forall (vars (?x8 int)) #109) +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) +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) +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) +decl uf_9 :: int +#43 := uf_9 +#189 := (= uf_5 uf_9) +#241 := (not #189) +#186 := (= uf_6 uf_8) +#250 := (not #186) +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) #1 := true -#111 := (implies #110 true) -#112 := (and #111 #110) -#113 := (implies #107 #112) -#114 := (and #113 #107) -#115 := (implies true #114) -#102 := (= uf_13 uf_4) -#116 := (implies #102 #115) -#100 := (= uf_12 uf_6) -#117 := (implies #100 #116) -#98 := (= uf_11 uf_5) -#118 := (implies #98 #117) -#119 := (implies true #118) -#28 := (<= 0::int uf_5) -#26 := (<= 1::int uf_4) +#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) -#96 := (<= uf_1 uf_4) -#121 := (implies #96 #120) +#66 := (< uf_5 uf_1) +#121 := (implies #66 #120) #122 := (implies #29 #121) #123 := (implies true #122) -#55 := (<= #20 uf_8) -#53 := (< #16 uf_9) -#54 := (and #53 #18) -#56 := (implies #54 #55) -#57 := (forall (vars (?x3 int)) #56) -#59 := (= #58 uf_8) -#60 := (implies false true) -#61 := (implies #59 #60) -#62 := (and #61 #59) -#63 := (implies #57 #62) -#64 := (and #63 #57) +#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) -#45 := (<= 0::int uf_7) -#51 := (<= 2::int uf_9) -#52 := (and #51 #45) -#66 := (implies #52 #65) -#48 := (+ uf_4 1::int) -#49 := (= uf_9 #48) -#67 := (implies #49 #66) -#46 := (and #26 #45) -#68 := (implies #46 #67) -#69 := (implies true #68) -#83 := (implies #82 #69) -#81 := (= uf_7 uf_4) -#84 := (implies #81 #83) -#85 := (implies true #84) -#80 := (and #26 #26) -#86 := (implies #80 #85) -#79 := (= uf_10 #39) -#87 := (implies #79 #86) -#77 := (< uf_6 #39) -#88 := (implies #77 #87) -#89 := (implies #29 #88) -#90 := (implies true #89) -#44 := (= uf_8 uf_6) -#70 := (implies #44 #69) -#42 := (= uf_7 uf_5) -#71 := (implies #42 #70) -#72 := (implies true #71) -#73 := (implies #29 #72) -#40 := (<= #39 uf_6) -#74 := (implies #40 #73) -#75 := (implies #29 #74) -#76 := (implies true #75) -#91 := (and #76 #90) -#92 := (implies #29 #91) -#38 := (< uf_4 uf_1) -#93 := (implies #38 #92) -#94 := (implies #29 #93) -#95 := (implies true #94) -#124 := (and #95 #123) +#124 := (and #65 #123) #125 := (implies #29 #124) #37 := (= #36 uf_6) #126 := (implies #37 #125) #33 := (<= #20 uf_6) -#30 := (< #16 uf_4) -#31 := (and #30 #18) +#30 := (< #16 uf_5) +#31 := (and #17 #30) #34 := (implies #31 #33) -#35 := (forall (vars (?x2 int)) #34) +#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 #130 #24) +#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 #132 #23) -#12 := (<= 0::int 0::int) +#133 := (and #23 #132) +#12 := (<= 1::int 1::int) #13 := (and #12 #12) -#11 := (<= 1::int 1::int) -#14 := (and #11 #13) -#15 := (and #11 #14) +#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) @@ -363,1967 +260,2032 @@ #137 := (implies #6 #136) #138 := (implies true #137) #139 := (not #138) -#1106 := (iff #139 #1103) -#475 := (and #18 #103) -#481 := (not #475) -#493 := (or #108 #481) -#498 := (forall (vars (?x8 int)) #493) -#482 := (or #105 #481) -#487 := (exists (vars (?x7 int)) #482) -#518 := (not #487) -#519 := (or #518 #498) -#527 := (and #487 #519) -#543 := (or #542 #527) -#552 := (or #551 #543) -#561 := (or #560 #552) -#326 := (not #29) -#576 := (or #326 #561) -#584 := (not #96) -#585 := (or #584 #576) -#593 := (or #326 #585) -#206 := (and #18 #53) -#212 := (not #206) -#213 := (or #55 #212) -#218 := (forall (vars (?x3 int)) #213) -#243 := (not #218) -#244 := (or #243 #221) -#252 := (and #218 #244) -#203 := (and #45 #51) -#267 := (not #203) -#268 := (or #267 #252) -#197 := (+ 1::int uf_4) -#200 := (= uf_9 #197) -#276 := (not #200) -#277 := (or #276 #268) -#285 := (not #46) -#286 := (or #285 #277) -#368 := (or #367 #286) -#377 := (or #376 #368) -#392 := (not #26) -#393 := (or #392 #377) -#402 := (or #401 #393) -#410 := (not #77) -#411 := (or #410 #402) -#419 := (or #326 #411) -#302 := (or #301 #286) -#311 := (or #310 #302) -#327 := (or #326 #311) -#335 := (not #40) -#336 := (or #335 #327) -#344 := (or #326 #336) -#431 := (and #344 #419) -#437 := (or #326 #431) -#445 := (not #38) -#446 := (or #445 #437) -#454 := (or #326 #446) -#605 := (and #454 #593) -#611 := (or #326 #605) -#620 := (or #619 #611) -#173 := (and #18 #30) -#179 := (not #173) -#180 := (or #33 #179) -#185 := (forall (vars (?x2 int)) #180) -#628 := (not #185) -#629 := (or #628 #620) -#637 := (or #326 #629) -#653 := (or #652 #637) -#661 := (and #9 #653) +#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) -#669 := (not #168) -#670 := (or #669 #661) -#678 := (and #168 #670) -#158 := (and #11 #12) -#161 := (and #11 #158) -#686 := (not #161) -#687 := (or #686 #678) -#695 := (or #652 #687) -#710 := (not #6) -#711 := (or #710 #695) -#723 := (not #711) -#1104 := (iff #723 #1103) -#1101 := (iff #711 #1098) -#1089 := (or false #1075) -#1092 := (or #652 #1089) -#1095 := (or #1078 #1092) -#1099 := (iff #1095 #1098) -#1100 := [rewrite]: #1099 -#1096 := (iff #711 #1095) -#1093 := (iff #695 #1092) -#1090 := (iff #687 #1089) -#1076 := (iff #678 #1075) -#1073 := (iff #670 #1072) -#1070 := (iff #661 #1069) -#1067 := (iff #653 #1064) -#1049 := (or #784 #1021) -#1052 := (or #619 #1049) -#1055 := (or #1046 #1052) -#1058 := (or #784 #1055) -#1061 := (or #652 #1058) -#1065 := (iff #1061 #1064) -#1066 := [rewrite]: #1065 -#1062 := (iff #653 #1061) -#1059 := (iff #637 #1058) -#1056 := (iff #629 #1055) -#1053 := (iff #620 #1052) -#1050 := (iff #611 #1049) -#1022 := (iff #605 #1021) -#1019 := (iff #593 #1016) -#998 := (or #542 #991) -#1001 := (or #551 #998) -#1004 := (or #560 #1001) -#1007 := (or #784 #1004) -#1010 := (or #927 #1007) -#1013 := (or #784 #1010) -#1017 := (iff #1013 #1016) -#1018 := [rewrite]: #1017 -#1014 := (iff #593 #1013) -#1011 := (iff #585 #1010) -#1008 := (iff #576 #1007) -#1005 := (iff #561 #1004) -#1002 := (iff #552 #1001) -#999 := (iff #543 #998) -#992 := (iff #527 #991) -#989 := (iff #519 #988) -#986 := (iff #498 #985) -#983 := (iff #493 #980) -#977 := (or #974 #960) -#981 := (iff #977 #980) +#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 -#978 := (iff #493 #977) -#961 := (iff #481 #960) -#958 := (iff #475 #957) -#955 := (iff #103 #954) -#956 := [rewrite]: #955 -#751 := (iff #18 #753) -#752 := [rewrite]: #751 -#959 := [monotonicity #752 #956]: #958 -#962 := [monotonicity #959]: #961 -#975 := (iff #108 #974) -#976 := [rewrite]: #975 -#979 := [monotonicity #976 #962]: #978 -#984 := [trans #979 #982]: #983 -#987 := [quant-intro #984]: #986 -#970 := (iff #518 #969) -#967 := (iff #487 #966) -#964 := (iff #482 #963) -#965 := [monotonicity #962]: #964 -#968 := [quant-intro #965]: #967 -#971 := [monotonicity #968]: #970 -#990 := [monotonicity #971 #987]: #989 -#993 := [monotonicity #968 #990]: #992 -#1000 := [monotonicity #993]: #999 -#1003 := [monotonicity #1000]: #1002 -#1006 := [monotonicity #1003]: #1005 -#785 := (iff #326 #784) -#782 := (iff #29 #781) -#778 := (iff #28 #779) -#780 := [rewrite]: #778 -#775 := (iff #26 #776) -#777 := [rewrite]: #775 -#783 := [monotonicity #777 #780]: #782 -#786 := [monotonicity #783]: #785 -#1009 := [monotonicity #786 #1006]: #1008 -#996 := (iff #584 #927) -#994 := (iff #96 #926) -#995 := [rewrite]: #994 -#997 := [monotonicity #995]: #996 -#1012 := [monotonicity #997 #1009]: #1011 -#1015 := [monotonicity #786 #1012]: #1014 -#1020 := [trans #1015 #1018]: #1019 -#949 := (iff #454 #946) -#937 := (or #784 #921) -#940 := (or #926 #937) -#943 := (or #784 #940) -#947 := (iff #943 #946) -#948 := [rewrite]: #947 -#944 := (iff #454 #943) -#941 := (iff #446 #940) -#938 := (iff #437 #937) -#922 := (iff #431 #921) -#919 := (iff #419 #916) -#857 := (or #796 #833) -#860 := (or #840 #857) -#863 := (or #846 #860) -#898 := (or #367 #863) -#901 := (or #376 #898) -#904 := (or #886 #901) -#907 := (or #401 #904) -#910 := (or #850 #907) -#913 := (or #784 #910) -#917 := (iff #913 #916) -#918 := [rewrite]: #917 -#914 := (iff #419 #913) -#911 := (iff #411 #910) -#908 := (iff #402 #907) -#905 := (iff #393 #904) -#902 := (iff #377 #901) -#899 := (iff #368 #898) -#864 := (iff #286 #863) -#861 := (iff #277 #860) -#858 := (iff #268 #857) -#834 := (iff #252 #833) -#831 := (iff #244 #828) -#825 := (or #822 #221) -#829 := (iff #825 #828) -#830 := [rewrite]: #829 -#826 := (iff #244 #825) -#823 := (iff #243 #822) -#820 := (iff #218 #819) -#817 := (iff #213 #816) -#814 := (iff #212 #813) -#811 := (iff #206 #810) -#808 := (iff #53 #804) -#809 := [rewrite]: #808 -#812 := [monotonicity #752 #809]: #811 -#815 := [monotonicity #812]: #814 -#802 := (iff #55 #801) -#803 := [rewrite]: #802 -#818 := [monotonicity #803 #815]: #817 -#821 := [quant-intro #818]: #820 -#824 := [monotonicity #821]: #823 -#827 := [monotonicity #824]: #826 -#832 := [trans #827 #830]: #831 -#835 := [monotonicity #821 #832]: #834 -#797 := (iff #267 #796) -#794 := (iff #203 #793) -#790 := (iff #51 #791) -#792 := [rewrite]: #790 -#787 := (iff #45 #788) -#789 := [rewrite]: #787 -#795 := [monotonicity #789 #792]: #794 -#798 := [monotonicity #795]: #797 -#859 := [monotonicity #798 #835]: #858 -#841 := (iff #276 #840) -#836 := (iff #200 #837) -#839 := [rewrite]: #836 -#842 := [monotonicity #839]: #841 -#862 := [monotonicity #842 #859]: #861 -#847 := (iff #285 #846) -#844 := (iff #46 #843) -#845 := [monotonicity #777 #789]: #844 -#848 := [monotonicity #845]: #847 -#865 := [monotonicity #848 #862]: #864 -#900 := [monotonicity #865]: #899 -#903 := [monotonicity #900]: #902 -#887 := (iff #392 #886) -#888 := [monotonicity #777]: #887 -#906 := [monotonicity #888 #903]: #905 -#909 := [monotonicity #906]: #908 -#896 := (iff #410 #850) -#891 := (not #854) -#894 := (iff #891 #850) -#895 := [rewrite]: #894 -#892 := (iff #410 #891) -#889 := (iff #77 #854) -#890 := [rewrite]: #889 -#893 := [monotonicity #890]: #892 -#897 := [trans #893 #895]: #896 -#912 := [monotonicity #897 #909]: #911 -#915 := [monotonicity #786 #912]: #914 -#920 := [trans #915 #918]: #919 -#884 := (iff #344 #881) -#866 := (or #301 #863) -#869 := (or #310 #866) -#872 := (or #784 #869) -#875 := (or #854 #872) -#878 := (or #784 #875) -#882 := (iff #878 #881) -#883 := [rewrite]: #882 -#879 := (iff #344 #878) -#876 := (iff #336 #875) -#873 := (iff #327 #872) -#870 := (iff #311 #869) -#867 := (iff #302 #866) -#868 := [monotonicity #865]: #867 -#871 := [monotonicity #868]: #870 -#874 := [monotonicity #786 #871]: #873 -#855 := (iff #335 #854) -#849 := (iff #40 #850) -#853 := [rewrite]: #849 -#856 := [monotonicity #853]: #855 -#877 := [monotonicity #856 #874]: #876 -#880 := [monotonicity #786 #877]: #879 -#885 := [trans #880 #883]: #884 -#923 := [monotonicity #885 #920]: #922 -#939 := [monotonicity #786 #923]: #938 -#935 := (iff #445 #926) -#930 := (not #927) -#933 := (iff #930 #926) -#934 := [rewrite]: #933 -#931 := (iff #445 #930) -#928 := (iff #38 #927) -#929 := [rewrite]: #928 -#932 := [monotonicity #929]: #931 -#936 := [trans #932 #934]: #935 -#942 := [monotonicity #936 #939]: #941 -#945 := [monotonicity #786 #942]: #944 -#950 := [trans #945 #948]: #949 -#1023 := [monotonicity #950 #1020]: #1022 -#1051 := [monotonicity #786 #1023]: #1050 -#1054 := [monotonicity #1051]: #1053 -#1047 := (iff #628 #1046) -#1044 := (iff #185 #1043) -#1041 := (iff #180 #1040) -#1038 := (iff #179 #1037) -#1035 := (iff #173 #1034) -#1032 := (iff #30 #1031) -#1033 := [rewrite]: #1032 -#1036 := [monotonicity #752 #1033]: #1035 -#1039 := [monotonicity #1036]: #1038 -#1027 := (iff #33 #1026) -#1028 := [rewrite]: #1027 -#1042 := [monotonicity #1028 #1039]: #1041 -#1045 := [quant-intro #1042]: #1044 +#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) +#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 -#1057 := [monotonicity #1048 #1054]: #1056 -#1060 := [monotonicity #786 #1057]: #1059 -#1063 := [monotonicity #1060]: #1062 -#1068 := [trans #1063 #1066]: #1067 -#1071 := [monotonicity #1068]: #1070 -#773 := (iff #669 #772) -#770 := (iff #168 #769) -#767 := (iff #165 #766) -#761 := (iff #21 #762) -#765 := [rewrite]: #761 -#758 := (iff #164 #757) -#755 := (iff #19 #754) -#748 := (iff #17 #747) -#750 := [rewrite]: #748 -#756 := [monotonicity #750 #752]: #755 -#759 := [monotonicity #756]: #758 -#768 := [monotonicity #759 #765]: #767 -#771 := [quant-intro #768]: #770 -#774 := [monotonicity #771]: #773 -#1074 := [monotonicity #774 #1071]: #1073 -#1077 := [monotonicity #771 #1074]: #1076 -#745 := (iff #686 false) -#740 := (not true) -#743 := (iff #740 false) -#744 := [rewrite]: #743 -#741 := (iff #686 #740) -#738 := (iff #161 true) -#730 := (and true true) -#733 := (and true #730) -#736 := (iff #733 true) -#737 := [rewrite]: #736 -#734 := (iff #161 #733) -#731 := (iff #158 #730) -#728 := (iff #12 true) -#729 := [rewrite]: #728 -#726 := (iff #11 true) -#727 := [rewrite]: #726 -#732 := [monotonicity #727 #729]: #731 -#735 := [monotonicity #727 #732]: #734 -#739 := [trans #735 #737]: #738 -#742 := [monotonicity #739]: #741 -#746 := [trans #742 #744]: #745 -#1091 := [monotonicity #746 #1077]: #1090 -#1094 := [monotonicity #1091]: #1093 -#1087 := (iff #710 #1078) -#1079 := (not #1078) -#1082 := (not #1079) -#1085 := (iff #1082 #1078) -#1086 := [rewrite]: #1085 -#1083 := (iff #710 #1082) -#1080 := (iff #6 #1079) -#1081 := [rewrite]: #1080 -#1084 := [monotonicity #1081]: #1083 -#1088 := [trans #1084 #1086]: #1087 -#1097 := [monotonicity #1088 #1094]: #1096 -#1102 := [trans #1097 #1100]: #1101 -#1105 := [monotonicity #1102]: #1104 -#724 := (iff #139 #723) -#721 := (iff #138 #711) -#716 := (implies true #711) -#719 := (iff #716 #711) -#720 := [rewrite]: #719 -#717 := (iff #138 #716) -#714 := (iff #137 #711) -#707 := (implies #6 #695) -#712 := (iff #707 #711) -#713 := [rewrite]: #712 -#708 := (iff #137 #707) -#705 := (iff #136 #695) -#700 := (implies true #695) -#703 := (iff #700 #695) -#704 := [rewrite]: #703 -#701 := (iff #136 #700) -#698 := (iff #135 #695) -#692 := (implies #9 #687) -#696 := (iff #692 #695) -#697 := [rewrite]: #696 -#693 := (iff #135 #692) -#690 := (iff #134 #687) -#683 := (implies #161 #678) -#688 := (iff #683 #687) -#689 := [rewrite]: #688 -#684 := (iff #134 #683) -#681 := (iff #133 #678) -#675 := (and #670 #168) -#679 := (iff #675 #678) -#680 := [rewrite]: #679 -#676 := (iff #133 #675) +#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) +#586 := [rewrite]: #585 +#581 := (iff #126 #580) +#578 := (iff #125 #575) +#572 := (implies #29 #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 +#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 -#673 := (iff #132 #670) -#666 := (implies #168 #661) -#671 := (iff #666 #670) -#672 := [rewrite]: #671 -#667 := (iff #132 #666) -#664 := (iff #131 #661) -#658 := (and #653 #9) -#662 := (iff #658 #661) -#663 := [rewrite]: #662 -#659 := (iff #131 #658) -#171 := (iff #24 #9) -#172 := [rewrite]: #171 -#656 := (iff #130 #653) -#649 := (implies #9 #637) -#654 := (iff #649 #653) -#655 := [rewrite]: #654 -#650 := (iff #130 #649) -#647 := (iff #129 #637) -#642 := (implies true #637) -#645 := (iff #642 #637) -#646 := [rewrite]: #645 -#643 := (iff #129 #642) -#640 := (iff #128 #637) -#634 := (implies #29 #629) -#638 := (iff #634 #637) -#639 := [rewrite]: #638 -#635 := (iff #128 #634) -#632 := (iff #127 #629) -#625 := (implies #185 #620) -#630 := (iff #625 #629) -#631 := [rewrite]: #630 -#626 := (iff #127 #625) -#623 := (iff #126 #620) -#616 := (implies #188 #611) -#621 := (iff #616 #620) -#622 := [rewrite]: #621 -#617 := (iff #126 #616) -#614 := (iff #125 #611) -#608 := (implies #29 #605) -#612 := (iff #608 #611) -#613 := [rewrite]: #612 -#609 := (iff #125 #608) -#606 := (iff #124 #605) -#603 := (iff #123 #593) -#598 := (implies true #593) -#601 := (iff #598 #593) -#602 := [rewrite]: #601 -#599 := (iff #123 #598) -#596 := (iff #122 #593) -#590 := (implies #29 #585) -#594 := (iff #590 #593) -#595 := [rewrite]: #594 -#591 := (iff #122 #590) -#588 := (iff #121 #585) -#581 := (implies #96 #576) -#586 := (iff #581 #585) -#587 := [rewrite]: #586 -#582 := (iff #121 #581) -#579 := (iff #120 #576) -#573 := (implies #29 #561) -#577 := (iff #573 #576) -#578 := [rewrite]: #577 -#574 := (iff #120 #573) -#571 := (iff #119 #561) -#566 := (implies true #561) -#569 := (iff #566 #561) -#570 := [rewrite]: #569 -#567 := (iff #119 #566) -#564 := (iff #118 #561) -#557 := (implies #466 #552) -#562 := (iff #557 #561) -#563 := [rewrite]: #562 -#558 := (iff #118 #557) -#555 := (iff #117 #552) -#548 := (implies #469 #543) -#553 := (iff #548 #552) -#554 := [rewrite]: #553 -#549 := (iff #117 #548) -#546 := (iff #116 #543) -#539 := (implies #472 #527) -#544 := (iff #539 #543) -#545 := [rewrite]: #544 -#540 := (iff #116 #539) -#537 := (iff #115 #527) -#532 := (implies true #527) -#535 := (iff #532 #527) -#536 := [rewrite]: #535 -#533 := (iff #115 #532) -#530 := (iff #114 #527) -#524 := (and #519 #487) -#528 := (iff #524 #527) -#529 := [rewrite]: #528 -#525 := (iff #114 #524) -#488 := (iff #107 #487) -#485 := (iff #106 #482) -#478 := (implies #475 #105) -#483 := (iff #478 #482) -#484 := [rewrite]: #483 -#479 := (iff #106 #478) -#476 := (iff #104 #475) -#477 := [rewrite]: #476 -#480 := [monotonicity #477]: #479 -#486 := [trans #480 #484]: #485 -#489 := [quant-intro #486]: #488 -#522 := (iff #113 #519) -#515 := (implies #487 #498) -#520 := (iff #515 #519) -#521 := [rewrite]: #520 -#516 := (iff #113 #515) -#513 := (iff #112 #498) -#508 := (and true #498) -#511 := (iff #508 #498) -#512 := [rewrite]: #511 -#509 := (iff #112 #508) -#499 := (iff #110 #498) -#496 := (iff #109 #493) -#490 := (implies #475 #108) -#494 := (iff #490 #493) -#495 := [rewrite]: #494 -#491 := (iff #109 #490) -#492 := [monotonicity #477]: #491 -#497 := [trans #492 #495]: #496 -#500 := [quant-intro #497]: #499 -#506 := (iff #111 true) -#501 := (implies #498 true) -#504 := (iff #501 true) -#505 := [rewrite]: #504 -#502 := (iff #111 #501) -#503 := [monotonicity #500]: #502 -#507 := [trans #503 #505]: #506 -#510 := [monotonicity #507 #500]: #509 -#514 := [trans #510 #512]: #513 -#517 := [monotonicity #489 #514]: #516 -#523 := [trans #517 #521]: #522 -#526 := [monotonicity #523 #489]: #525 -#531 := [trans #526 #529]: #530 -#534 := [monotonicity #531]: #533 -#538 := [trans #534 #536]: #537 -#473 := (iff #102 #472) -#474 := [rewrite]: #473 -#541 := [monotonicity #474 #538]: #540 -#547 := [trans #541 #545]: #546 -#470 := (iff #100 #469) -#471 := [rewrite]: #470 -#550 := [monotonicity #471 #547]: #549 -#556 := [trans #550 #554]: #555 -#467 := (iff #98 #466) -#468 := [rewrite]: #467 -#559 := [monotonicity #468 #556]: #558 -#565 := [trans #559 #563]: #564 -#568 := [monotonicity #565]: #567 -#572 := [trans #568 #570]: #571 -#575 := [monotonicity #572]: #574 -#580 := [trans #575 #578]: #579 -#583 := [monotonicity #580]: #582 -#589 := [trans #583 #587]: #588 -#592 := [monotonicity #589]: #591 -#597 := [trans #592 #595]: #596 -#600 := [monotonicity #597]: #599 -#604 := [trans #600 #602]: #603 -#464 := (iff #95 #454) -#459 := (implies true #454) -#462 := (iff #459 #454) -#463 := [rewrite]: #462 -#460 := (iff #95 #459) -#457 := (iff #94 #454) -#451 := (implies #29 #446) -#455 := (iff #451 #454) -#456 := [rewrite]: #455 -#452 := (iff #94 #451) -#449 := (iff #93 #446) -#442 := (implies #38 #437) -#447 := (iff #442 #446) -#448 := [rewrite]: #447 -#443 := (iff #93 #442) -#440 := (iff #92 #437) -#434 := (implies #29 #431) -#438 := (iff #434 #437) -#439 := [rewrite]: #438 -#435 := (iff #92 #434) -#432 := (iff #91 #431) -#429 := (iff #90 #419) -#424 := (implies true #419) -#427 := (iff #424 #419) -#428 := [rewrite]: #427 -#425 := (iff #90 #424) -#422 := (iff #89 #419) -#416 := (implies #29 #411) -#420 := (iff #416 #419) -#421 := [rewrite]: #420 -#417 := (iff #89 #416) -#414 := (iff #88 #411) -#407 := (implies #77 #402) -#412 := (iff #407 #411) -#413 := [rewrite]: #412 -#408 := (iff #88 #407) -#405 := (iff #87 #402) -#398 := (implies #356 #393) -#403 := (iff #398 #402) -#404 := [rewrite]: #403 -#399 := (iff #87 #398) -#396 := (iff #86 #393) -#389 := (implies #26 #377) -#394 := (iff #389 #393) -#395 := [rewrite]: #394 -#390 := (iff #86 #389) -#387 := (iff #85 #377) -#382 := (implies true #377) -#385 := (iff #382 #377) -#386 := [rewrite]: #385 -#383 := (iff #85 #382) -#380 := (iff #84 #377) -#373 := (implies #361 #368) -#378 := (iff #373 #377) -#379 := [rewrite]: #378 -#374 := (iff #84 #373) -#371 := (iff #83 #368) -#364 := (implies #82 #286) -#369 := (iff #364 #368) -#370 := [rewrite]: #369 -#365 := (iff #83 #364) -#296 := (iff #69 #286) -#291 := (implies true #286) -#294 := (iff #291 #286) -#295 := [rewrite]: #294 -#292 := (iff #69 #291) -#289 := (iff #68 #286) -#282 := (implies #46 #277) -#287 := (iff #282 #286) -#288 := [rewrite]: #287 -#283 := (iff #68 #282) -#280 := (iff #67 #277) -#273 := (implies #200 #268) -#278 := (iff #273 #277) -#279 := [rewrite]: #278 -#274 := (iff #67 #273) -#271 := (iff #66 #268) -#264 := (implies #203 #252) -#269 := (iff #264 #268) -#270 := [rewrite]: #269 -#265 := (iff #66 #264) -#262 := (iff #65 #252) -#257 := (implies true #252) -#260 := (iff #257 #252) -#261 := [rewrite]: #260 -#258 := (iff #65 #257) -#255 := (iff #64 #252) -#249 := (and #244 #218) -#253 := (iff #249 #252) -#254 := [rewrite]: #253 -#250 := (iff #64 #249) -#219 := (iff #57 #218) -#216 := (iff #56 #213) -#209 := (implies #206 #55) -#214 := (iff #209 #213) -#215 := [rewrite]: #214 -#210 := (iff #56 #209) -#207 := (iff #54 #206) -#208 := [rewrite]: #207 -#211 := [monotonicity #208]: #210 -#217 := [trans #211 #215]: #216 -#220 := [quant-intro #217]: #219 -#247 := (iff #63 #244) -#240 := (implies #218 #221) -#245 := (iff #240 #244) -#246 := [rewrite]: #245 -#241 := (iff #63 #240) -#238 := (iff #62 #221) -#233 := (and true #221) -#236 := (iff #233 #221) -#237 := [rewrite]: #236 -#234 := (iff #62 #233) -#222 := (iff #59 #221) -#223 := [rewrite]: #222 -#231 := (iff #61 true) -#226 := (implies #221 true) -#229 := (iff #226 true) -#230 := [rewrite]: #229 -#227 := (iff #61 #226) -#224 := (iff #60 true) -#225 := [rewrite]: #224 -#228 := [monotonicity #223 #225]: #227 -#232 := [trans #228 #230]: #231 -#235 := [monotonicity #232 #223]: #234 -#239 := [trans #235 #237]: #238 -#242 := [monotonicity #220 #239]: #241 -#248 := [trans #242 #246]: #247 -#251 := [monotonicity #248 #220]: #250 -#256 := [trans #251 #254]: #255 -#259 := [monotonicity #256]: #258 -#263 := [trans #259 #261]: #262 -#204 := (iff #52 #203) -#205 := [rewrite]: #204 -#266 := [monotonicity #205 #263]: #265 -#272 := [trans #266 #270]: #271 -#201 := (iff #49 #200) -#198 := (= #48 #197) -#199 := [rewrite]: #198 -#202 := [monotonicity #199]: #201 -#275 := [monotonicity #202 #272]: #274 -#281 := [trans #275 #279]: #280 -#284 := [monotonicity #281]: #283 -#290 := [trans #284 #288]: #289 -#293 := [monotonicity #290]: #292 -#297 := [trans #293 #295]: #296 -#366 := [monotonicity #297]: #365 -#372 := [trans #366 #370]: #371 -#362 := (iff #81 #361) -#363 := [rewrite]: #362 -#375 := [monotonicity #363 #372]: #374 -#381 := [trans #375 #379]: #380 -#384 := [monotonicity #381]: #383 -#388 := [trans #384 #386]: #387 -#359 := (iff #80 #26) -#360 := [rewrite]: #359 -#391 := [monotonicity #360 #388]: #390 -#397 := [trans #391 #395]: #396 -#357 := (iff #79 #356) -#358 := [rewrite]: #357 -#400 := [monotonicity #358 #397]: #399 -#406 := [trans #400 #404]: #405 -#409 := [monotonicity #406]: #408 -#415 := [trans #409 #413]: #414 -#418 := [monotonicity #415]: #417 -#423 := [trans #418 #421]: #422 -#426 := [monotonicity #423]: #425 -#430 := [trans #426 #428]: #429 -#354 := (iff #76 #344) -#349 := (implies true #344) -#352 := (iff #349 #344) -#353 := [rewrite]: #352 -#350 := (iff #76 #349) -#347 := (iff #75 #344) -#341 := (implies #29 #336) -#345 := (iff #341 #344) -#346 := [rewrite]: #345 -#342 := (iff #75 #341) -#339 := (iff #74 #336) -#332 := (implies #40 #327) -#337 := (iff #332 #336) -#338 := [rewrite]: #337 -#333 := (iff #74 #332) -#330 := (iff #73 #327) -#323 := (implies #29 #311) -#328 := (iff #323 #327) -#329 := [rewrite]: #328 -#324 := (iff #73 #323) -#321 := (iff #72 #311) -#316 := (implies true #311) -#319 := (iff #316 #311) -#320 := [rewrite]: #319 -#317 := (iff #72 #316) -#314 := (iff #71 #311) -#307 := (implies #191 #302) -#312 := (iff #307 #311) -#313 := [rewrite]: #312 -#308 := (iff #71 #307) -#305 := (iff #70 #302) -#298 := (implies #194 #286) -#303 := (iff #298 #302) -#304 := [rewrite]: #303 -#299 := (iff #70 #298) -#195 := (iff #44 #194) -#196 := [rewrite]: #195 -#300 := [monotonicity #196 #297]: #299 -#306 := [trans #300 #304]: #305 -#192 := (iff #42 #191) -#193 := [rewrite]: #192 -#309 := [monotonicity #193 #306]: #308 -#315 := [trans #309 #313]: #314 -#318 := [monotonicity #315]: #317 -#322 := [trans #318 #320]: #321 -#325 := [monotonicity #322]: #324 -#331 := [trans #325 #329]: #330 -#334 := [monotonicity #331]: #333 -#340 := [trans #334 #338]: #339 -#343 := [monotonicity #340]: #342 -#348 := [trans #343 #346]: #347 -#351 := [monotonicity #348]: #350 -#355 := [trans #351 #353]: #354 -#433 := [monotonicity #355 #430]: #432 -#436 := [monotonicity #433]: #435 -#441 := [trans #436 #439]: #440 -#444 := [monotonicity #441]: #443 -#450 := [trans #444 #448]: #449 -#453 := [monotonicity #450]: #452 -#458 := [trans #453 #456]: #457 -#461 := [monotonicity #458]: #460 -#465 := [trans #461 #463]: #464 -#607 := [monotonicity #465 #604]: #606 -#610 := [monotonicity #607]: #609 -#615 := [trans #610 #613]: #614 -#189 := (iff #37 #188) -#190 := [rewrite]: #189 -#618 := [monotonicity #190 #615]: #617 -#624 := [trans #618 #622]: #623 -#186 := (iff #35 #185) -#183 := (iff #34 #180) -#176 := (implies #173 #33) -#181 := (iff #176 #180) -#182 := [rewrite]: #181 -#177 := (iff #34 #176) -#174 := (iff #31 #173) -#175 := [rewrite]: #174 -#178 := [monotonicity #175]: #177 -#184 := [trans #178 #182]: #183 -#187 := [quant-intro #184]: #186 -#627 := [monotonicity #187 #624]: #626 +#627 := [monotonicity #170 #624]: #626 #633 := [trans #627 #631]: #632 -#636 := [monotonicity #633]: #635 -#641 := [trans #636 #639]: #640 -#644 := [monotonicity #641]: #643 -#648 := [trans #644 #646]: #647 -#651 := [monotonicity #172 #648]: #650 -#657 := [trans #651 #655]: #656 -#660 := [monotonicity #657 #172]: #659 -#665 := [trans #660 #663]: #664 -#668 := [monotonicity #170 #665]: #667 -#674 := [trans #668 #672]: #673 -#677 := [monotonicity #674 #170]: #676 -#682 := [trans #677 #680]: #681 +#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 -#685 := [monotonicity #163 #682]: #684 -#691 := [trans #685 #689]: #690 -#694 := [monotonicity #691]: #693 -#699 := [trans #694 #697]: #698 -#702 := [monotonicity #699]: #701 -#706 := [trans #702 #704]: #705 -#709 := [monotonicity #706]: #708 -#715 := [trans #709 #713]: #714 -#718 := [monotonicity #715]: #717 -#722 := [trans #718 #720]: #721 -#725 := [monotonicity #722]: #724 -#1107 := [trans #725 #1105]: #1106 +#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 -#1108 := [mp #155 #1107]: #1103 -#1109 := [not-or-elim #1108]: #9 -#2193 := [trans #1109 #2192]: #2202 -#1888 := (not #1154) -#1974 := (or #1540 #1888) -#1889 := [def-axiom]: #1974 -#2194 := [unit-resolution #1889 #2205]: #1888 -#2195 := (not #2202) -#2196 := (or #2195 #1154) -#2197 := [th-lemma]: #2196 -#2198 := [unit-resolution #2197 #2194 #2193]: false -#2199 := [lemma #2198]: #1540 -#2393 := (or #1545 #2390) -#1708 := (forall (vars (?x7 int)) #1705) -#1801 := (or #1708 #1798) -#1804 := (not #1801) -#1807 := (or #560 #551 #542 #886 #1654 #927 #1804) -#1810 := (not #1807) -#1612 := (forall (vars (?x3 int)) #1607) -#1618 := (not #1612) -#1619 := (or #221 #1618) -#1620 := (not #1619) -#1648 := (or #1620 #1645) -#1657 := (not #1648) -#1667 := (or #367 #401 #376 #886 #1654 #1655 #1656 #840 #850 #1657) -#1668 := (not #1667) -#1658 := (or #310 #301 #886 #1654 #1655 #1656 #840 #854 #1657) -#1659 := (not #1658) -#1673 := (or #1659 #1668) -#1679 := (not #1673) -#1680 := (or #886 #1654 #926 #1679) -#1681 := (not #1680) -#1813 := (or #1681 #1810) -#1816 := (not #1813) -#1590 := (forall (vars (?x2 int)) #1585) -#1784 := (not #1590) -#1568 := (forall (vars (?x1 int)) #1563) -#1783 := (not #1568) -#1819 := (or #619 #886 #1654 #1783 #1784 #1816) -#1822 := (not #1819) -#1825 := (or #1545 #1822) -#2394 := (iff #1825 #2393) -#2391 := (iff #1822 #2390) -#2388 := (iff #1819 #2387) -#2385 := (iff #1816 #2384) -#2382 := (iff #1813 #2381) -#2379 := (iff #1810 #2378) -#2376 := (iff #1807 #2375) -#2373 := (iff #1804 #2372) -#2370 := (iff #1801 #2369) -#2367 := (iff #1708 #2364) -#2365 := (iff #1705 #1705) -#2366 := [refl]: #2365 -#2368 := [quant-intro #2366]: #2367 +#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 -#2374 := [monotonicity #2371]: #2373 -#2377 := [monotonicity #2374]: #2376 -#2380 := [monotonicity #2377]: #2379 -#2362 := (iff #1681 #2361) -#2359 := (iff #1680 #2358) -#2356 := (iff #1679 #2355) -#2353 := (iff #1673 #2352) -#2350 := (iff #1668 #2349) -#2347 := (iff #1667 #2346) -#2338 := (iff #1657 #2337) -#2335 := (iff #1648 #2334) -#2332 := (iff #1620 #2331) -#2329 := (iff #1619 #2328) -#2326 := (iff #1618 #2325) -#2323 := (iff #1612 #2320) -#2321 := (iff #1607 #1607) -#2322 := [refl]: #2321 -#2324 := [quant-intro #2322]: #2323 -#2327 := [monotonicity #2324]: #2326 +#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 -#2348 := [monotonicity #2339]: #2347 -#2351 := [monotonicity #2348]: #2350 -#2344 := (iff #1659 #2343) -#2341 := (iff #1658 #2340) -#2342 := [monotonicity #2339]: #2341 -#2345 := [monotonicity #2342]: #2344 -#2354 := [monotonicity #2345 #2351]: #2353 -#2357 := [monotonicity #2354]: #2356 -#2360 := [monotonicity #2357]: #2359 -#2363 := [monotonicity #2360]: #2362 -#2383 := [monotonicity #2363 #2380]: #2382 -#2386 := [monotonicity #2383]: #2385 -#2318 := (iff #1784 #2317) -#2315 := (iff #1590 #2312) -#2313 := (iff #1585 #1585) -#2314 := [refl]: #2313 -#2316 := [quant-intro #2314]: #2315 +#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 -#2310 := (iff #1783 #2309) -#2307 := (iff #1568 #2304) -#2305 := (iff #1563 #1563) -#2306 := [refl]: #2305 -#2308 := [quant-intro #2306]: #2307 -#2311 := [monotonicity #2308]: #2310 -#2389 := [monotonicity #2311 #2319 #2386]: #2388 -#2392 := [monotonicity #2389]: #2391 -#2395 := [monotonicity #2392]: #2394 -#1304 := (not #1303) -#1481 := (and #1304 #1305) -#1484 := (not #1481) -#1500 := (or #1484 #1495) -#1503 := (not #1500) -#1283 := (not #1282) -#1472 := (and #1283 #1284) -#1475 := (not #1472) -#1478 := (or #1469 #1475) -#1506 := (and #1478 #1503) -#1273 := (not #963) -#1276 := (forall (vars (?x7 int)) #1273) -#1509 := (or #1276 #1506) -#1515 := (and #466 #469 #472 #776 #779 #926 #1509) -#1401 := (not #1396) -#1404 := (and #1192 #1401) -#1407 := (not #1404) -#1410 := (or #1383 #1407) -#1413 := (not #1410) -#1204 := (not #221) -#1214 := (and #1204 #819) -#1419 := (or #1214 #1413) -#1447 := (and #82 #356 #361 #776 #779 #788 #791 #837 #854 #1419) -#1431 := (and #191 #194 #776 #779 #788 #791 #837 #850 #1419) -#1452 := (or #1431 #1447) -#1458 := (and #776 #779 #927 #1452) -#1520 := (or #1458 #1515) -#1526 := (and #188 #769 #776 #779 #1043 #1520) -#1348 := (and #1155 #1157) -#1351 := (not #1348) -#1357 := (or #1154 #1351) -#1362 := (not #1357) -#1531 := (or #1362 #1526) -#1828 := (iff #1531 #1825) -#1746 := (or #1303 #1731 #1495) -#1758 := (or #1757 #1746) -#1759 := (not #1758) -#1764 := (or #1708 #1759) -#1770 := (not #1764) -#1771 := (or #560 #551 #542 #886 #1654 #927 #1770) -#1772 := (not #1771) -#1777 := (or #1681 #1772) -#1785 := (not #1777) -#1786 := (or #619 #886 #1654 #1783 #1784 #1785) -#1787 := (not #1786) -#1792 := (or #1545 #1787) -#1826 := (iff #1792 #1825) -#1823 := (iff #1787 #1822) -#1820 := (iff #1786 #1819) -#1817 := (iff #1785 #1816) -#1814 := (iff #1777 #1813) -#1811 := (iff #1772 #1810) -#1808 := (iff #1771 #1807) -#1805 := (iff #1770 #1804) -#1802 := (iff #1764 #1801) -#1799 := (iff #1759 #1798) -#1796 := (iff #1758 #1795) -#1797 := [rewrite]: #1796 -#1800 := [monotonicity #1797]: #1799 -#1803 := [monotonicity #1800]: #1802 -#1806 := [monotonicity #1803]: #1805 -#1809 := [monotonicity #1806]: #1808 -#1812 := [monotonicity #1809]: #1811 -#1815 := [monotonicity #1812]: #1814 -#1818 := [monotonicity #1815]: #1817 -#1821 := [monotonicity #1818]: #1820 -#1824 := [monotonicity #1821]: #1823 -#1827 := [monotonicity #1824]: #1826 -#1793 := (iff #1531 #1792) -#1790 := (iff #1526 #1787) -#1780 := (and #188 #1568 #776 #779 #1590 #1777) -#1788 := (iff #1780 #1787) -#1789 := [rewrite]: #1788 -#1781 := (iff #1526 #1780) -#1778 := (iff #1520 #1777) -#1775 := (iff #1515 #1772) -#1767 := (and #466 #469 #472 #776 #779 #926 #1764) -#1773 := (iff #1767 #1772) -#1774 := [rewrite]: #1773 -#1768 := (iff #1515 #1767) -#1765 := (iff #1509 #1764) -#1762 := (iff #1506 #1759) -#1751 := (not #1746) -#1754 := (and #1726 #1751) -#1760 := (iff #1754 #1759) -#1761 := [rewrite]: #1760 -#1755 := (iff #1506 #1754) -#1752 := (iff #1503 #1751) -#1749 := (iff #1500 #1746) -#1732 := (or #1303 #1731) -#1743 := (or #1732 #1495) -#1747 := (iff #1743 #1746) -#1748 := [rewrite]: #1747 -#1744 := (iff #1500 #1743) -#1741 := (iff #1484 #1732) -#1733 := (not #1732) -#1736 := (not #1733) -#1739 := (iff #1736 #1732) -#1740 := [rewrite]: #1739 -#1737 := (iff #1484 #1736) -#1734 := (iff #1481 #1733) -#1735 := [rewrite]: #1734 -#1738 := [monotonicity #1735]: #1737 -#1742 := [trans #1738 #1740]: #1741 -#1745 := [monotonicity #1742]: #1744 -#1750 := [trans #1745 #1748]: #1749 -#1753 := [monotonicity #1750]: #1752 -#1729 := (iff #1478 #1726) -#1712 := (or #1282 #1711) -#1723 := (or #1469 #1712) -#1727 := (iff #1723 #1726) +#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 -#1724 := (iff #1478 #1723) -#1721 := (iff #1475 #1712) -#1713 := (not #1712) -#1716 := (not #1713) -#1719 := (iff #1716 #1712) -#1720 := [rewrite]: #1719 -#1717 := (iff #1475 #1716) -#1714 := (iff #1472 #1713) +#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 -#1718 := [monotonicity #1715]: #1717 -#1722 := [trans #1718 #1720]: #1721 -#1725 := [monotonicity #1722]: #1724 -#1730 := [trans #1725 #1728]: #1729 -#1756 := [monotonicity #1730 #1753]: #1755 -#1763 := [trans #1756 #1761]: #1762 -#1709 := (iff #1276 #1708) -#1706 := (iff #1273 #1705) -#1703 := (iff #963 #1700) -#1686 := (or #1548 #953) -#1697 := (or #105 #1686) -#1701 := (iff #1697 #1700) -#1702 := [rewrite]: #1701 -#1698 := (iff #963 #1697) -#1695 := (iff #960 #1686) -#1687 := (not #1686) -#1690 := (not #1687) -#1693 := (iff #1690 #1686) -#1694 := [rewrite]: #1693 -#1691 := (iff #960 #1690) -#1688 := (iff #957 #1687) -#1689 := [rewrite]: #1688 -#1692 := [monotonicity #1689]: #1691 -#1696 := [trans #1692 #1694]: #1695 +#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 -#1704 := [trans #1699 #1702]: #1703 -#1707 := [monotonicity #1704]: #1706 -#1710 := [quant-intro #1707]: #1709 -#1766 := [monotonicity #1710 #1763]: #1765 -#1769 := [monotonicity #1766]: #1768 -#1776 := [trans #1769 #1774]: #1775 -#1684 := (iff #1458 #1681) -#1676 := (and #776 #779 #927 #1673) -#1682 := (iff #1676 #1681) -#1683 := [rewrite]: #1682 -#1677 := (iff #1458 #1676) -#1674 := (iff #1452 #1673) -#1671 := (iff #1447 #1668) -#1664 := (and #82 #356 #361 #776 #779 #788 #791 #837 #854 #1648) -#1669 := (iff #1664 #1668) -#1670 := [rewrite]: #1669 -#1665 := (iff #1447 #1664) -#1649 := (iff #1419 #1648) -#1646 := (iff #1413 #1645) -#1643 := (iff #1410 #1640) -#1626 := (or #1625 #1396) -#1637 := (or #1383 #1626) -#1641 := (iff #1637 #1640) -#1642 := [rewrite]: #1641 -#1638 := (iff #1410 #1637) -#1635 := (iff #1407 #1626) -#1627 := (not #1626) -#1630 := (not #1627) -#1633 := (iff #1630 #1626) -#1634 := [rewrite]: #1633 -#1631 := (iff #1407 #1630) -#1628 := (iff #1404 #1627) -#1629 := [rewrite]: #1628 -#1632 := [monotonicity #1629]: #1631 -#1636 := [trans #1632 #1634]: #1635 -#1639 := [monotonicity #1636]: #1638 -#1644 := [trans #1639 #1642]: #1643 -#1647 := [monotonicity #1644]: #1646 -#1623 := (iff #1214 #1620) -#1615 := (and #1204 #1612) +#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 #1214 #1615) -#1613 := (iff #819 #1612) -#1610 := (iff #816 #1607) -#1593 := (or #1548 #805) -#1604 := (or #801 #1593) +#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 #816 #1604) -#1602 := (iff #813 #1593) +#1605 := (iff #1360 #1604) +#1602 := (iff #1344 #1593) #1594 := (not #1593) #1597 := (not #1594) #1600 := (iff #1597 #1593) #1601 := [rewrite]: #1600 -#1598 := (iff #813 #1597) -#1595 := (iff #810 #1594) +#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 := [quant-intro #1611]: #1613 -#1617 := [monotonicity #1614]: #1616 +#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 -#1650 := [monotonicity #1624 #1647]: #1649 -#1666 := [monotonicity #1650]: #1665 -#1672 := [trans #1666 #1670]: #1671 -#1662 := (iff #1431 #1659) -#1651 := (and #191 #194 #776 #779 #788 #791 #837 #850 #1648) -#1660 := (iff #1651 #1659) -#1661 := [rewrite]: #1660 -#1652 := (iff #1431 #1651) -#1653 := [monotonicity #1650]: #1652 -#1663 := [trans #1653 #1661]: #1662 -#1675 := [monotonicity #1663 #1672]: #1674 -#1678 := [monotonicity #1675]: #1677 -#1685 := [trans #1678 #1683]: #1684 -#1779 := [monotonicity #1685 #1776]: #1778 -#1591 := (iff #1043 #1590) -#1588 := (iff #1040 #1585) -#1571 := (or #1548 #1029) -#1582 := (or #1026 #1571) -#1586 := (iff #1582 #1585) -#1587 := [rewrite]: #1586 -#1583 := (iff #1040 #1582) -#1580 := (iff #1037 #1571) -#1572 := (not #1571) -#1575 := (not #1572) -#1578 := (iff #1575 #1571) -#1579 := [rewrite]: #1578 -#1576 := (iff #1037 #1575) -#1573 := (iff #1034 #1572) -#1574 := [rewrite]: #1573 -#1577 := [monotonicity #1574]: #1576 -#1581 := [trans #1577 #1579]: #1580 -#1584 := [monotonicity #1581]: #1583 -#1589 := [trans #1584 #1587]: #1588 -#1592 := [quant-intro #1589]: #1591 -#1569 := (iff #769 #1568) -#1566 := (iff #766 #1563) -#1549 := (or #749 #1548) -#1560 := (or #1549 #762) -#1564 := (iff #1560 #1563) -#1565 := [rewrite]: #1564 -#1561 := (iff #766 #1560) -#1558 := (iff #757 #1549) -#1550 := (not #1549) -#1553 := (not #1550) -#1556 := (iff #1553 #1549) -#1557 := [rewrite]: #1556 -#1554 := (iff #757 #1553) -#1551 := (iff #754 #1550) -#1552 := [rewrite]: #1551 -#1555 := [monotonicity #1552]: #1554 -#1559 := [trans #1555 #1557]: #1558 -#1562 := [monotonicity #1559]: #1561 -#1567 := [trans #1562 #1565]: #1566 -#1570 := [quant-intro #1567]: #1569 -#1782 := [monotonicity #1570 #1592 #1779]: #1781 -#1791 := [trans #1782 #1789]: #1790 -#1546 := (iff #1362 #1545) -#1543 := (iff #1357 #1540) -#1165 := (or #1164 #1156) -#1537 := (or #1154 #1165) -#1541 := (iff #1537 #1540) -#1542 := [rewrite]: #1541 -#1538 := (iff #1357 #1537) -#1535 := (iff #1351 #1165) -#1292 := (not #1165) -#1314 := (not #1292) -#1347 := (iff #1314 #1165) -#1534 := [rewrite]: #1347 -#1202 := (iff #1351 #1314) -#1293 := (iff #1348 #1292) -#1313 := [rewrite]: #1293 -#1203 := [monotonicity #1313]: #1202 -#1536 := [trans #1203 #1534]: #1535 -#1539 := [monotonicity #1536]: #1538 -#1544 := [trans #1539 #1542]: #1543 -#1547 := [monotonicity #1544]: #1546 -#1794 := [monotonicity #1547 #1791]: #1793 -#1829 := [trans #1794 #1827]: #1828 -#1299 := (+ #1298 #972) -#1300 := (<= #1299 0::int) -#1306 := (and #1305 #1304) -#1307 := (not #1306) -#1308 := (or #1307 #1300) -#1309 := (not #1308) -#1285 := (and #1284 #1283) -#1286 := (not #1285) -#1288 := (= #1287 uf_12) -#1289 := (or #1288 #1286) -#1315 := (and #1289 #1309) -#1319 := (or #1276 #1315) -#1176 := (not #784) -#1268 := (not #542) -#1265 := (not #551) -#1262 := (not #560) -#1323 := (and #1262 #1265 #1268 #1176 #930 #1319) -#1225 := (not #846) -#1222 := (not #840) -#1189 := (+ ?x3!1 #806) -#1190 := (>= #1189 0::int) -#1191 := (not #1190) -#1193 := (and #1192 #1191) -#1194 := (not #1193) -#1196 := (+ #1195 #799) -#1197 := (<= #1196 0::int) -#1198 := (or #1197 #1194) -#1199 := (not #1198) -#1218 := (or #1199 #1214) -#1185 := (not #796) -#1243 := (not #886) -#1240 := (not #376) -#1237 := (not #401) -#1234 := (not #367) -#1248 := (and #1234 #1237 #1240 #1243 #1176 #1185 #1218 #1222 #1225 #854) -#1182 := (not #301) -#1179 := (not #310) -#1230 := (and #1179 #1182 #1176 #1185 #1218 #1222 #1225 #891) -#1252 := (or #1230 #1248) -#1258 := (and #1176 #1252 #927) -#1327 := (or #1258 #1323) -#1166 := (not #619) -#1338 := (and #1166 #769 #1176 #1327 #1043) -#1158 := (and #1157 #1155) -#1159 := (not #1158) -#1160 := (or #1159 #1154) -#1161 := (not #1160) -#1342 := (or #1161 #1338) -#1532 := (iff #1342 #1531) -#1529 := (iff #1338 #1526) -#1523 := (and #188 #769 #781 #1520 #1043) -#1527 := (iff #1523 #1526) +#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 -#1524 := (iff #1338 #1523) -#1521 := (iff #1327 #1520) -#1518 := (iff #1323 #1515) -#1512 := (and #466 #469 #472 #781 #926 #1509) -#1516 := (iff #1512 #1515) -#1517 := [rewrite]: #1516 -#1513 := (iff #1323 #1512) -#1510 := (iff #1319 #1509) -#1507 := (iff #1315 #1506) -#1504 := (iff #1309 #1503) -#1501 := (iff #1308 #1500) -#1498 := (iff #1300 #1495) -#1487 := (+ #972 #1298) -#1490 := (<= #1487 0::int) -#1496 := (iff #1490 #1495) -#1497 := [rewrite]: #1496 -#1491 := (iff #1300 #1490) -#1488 := (= #1299 #1487) -#1489 := [rewrite]: #1488 -#1492 := [monotonicity #1489]: #1491 -#1499 := [trans #1492 #1497]: #1498 -#1485 := (iff #1307 #1484) -#1482 := (iff #1306 #1481) -#1483 := [rewrite]: #1482 -#1486 := [monotonicity #1483]: #1485 -#1502 := [monotonicity #1486 #1499]: #1501 -#1505 := [monotonicity #1502]: #1504 -#1479 := (iff #1289 #1478) -#1476 := (iff #1286 #1475) -#1473 := (iff #1285 #1472) -#1474 := [rewrite]: #1473 -#1477 := [monotonicity #1474]: #1476 -#1470 := (iff #1288 #1469) +#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 -#1480 := [monotonicity #1471 #1477]: #1479 -#1508 := [monotonicity #1480 #1505]: #1507 -#1511 := [monotonicity #1508]: #1510 -#1367 := (iff #1176 #781) -#1368 := [rewrite]: #1367 -#1467 := (iff #1268 #472) -#1468 := [rewrite]: #1467 -#1465 := (iff #1265 #469) -#1466 := [rewrite]: #1465 -#1463 := (iff #1262 #466) -#1464 := [rewrite]: #1463 -#1514 := [monotonicity #1464 #1466 #1468 #1368 #934 #1511]: #1513 -#1519 := [trans #1514 #1517]: #1518 -#1461 := (iff #1258 #1458) -#1455 := (and #781 #1452 #927) +#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 #1258 #1455) -#1453 := (iff #1252 #1452) -#1450 := (iff #1248 #1447) -#1444 := (and #82 #356 #361 #776 #781 #793 #1419 #837 #843 #854) -#1448 := (iff #1444 #1447) -#1449 := [rewrite]: #1448 -#1445 := (iff #1248 #1444) -#1426 := (iff #1225 #843) -#1427 := [rewrite]: #1426 -#1424 := (iff #1222 #837) -#1425 := [rewrite]: #1424 -#1422 := (iff #1218 #1419) -#1416 := (or #1413 #1214) -#1420 := (iff #1416 #1419) -#1421 := [rewrite]: #1420 -#1417 := (iff #1218 #1416) -#1414 := (iff #1199 #1413) -#1411 := (iff #1198 #1410) -#1408 := (iff #1194 #1407) -#1405 := (iff #1193 #1404) -#1402 := (iff #1191 #1401) -#1399 := (iff #1190 #1396) -#1388 := (+ #806 ?x3!1) -#1391 := (>= #1388 0::int) -#1397 := (iff #1391 #1396) -#1398 := [rewrite]: #1397 -#1392 := (iff #1190 #1391) -#1389 := (= #1189 #1388) -#1390 := [rewrite]: #1389 -#1393 := [monotonicity #1390]: #1392 -#1400 := [trans #1393 #1398]: #1399 -#1403 := [monotonicity #1400]: #1402 -#1406 := [monotonicity #1403]: #1405 -#1409 := [monotonicity #1406]: #1408 -#1386 := (iff #1197 #1383) -#1375 := (+ #799 #1195) -#1378 := (<= #1375 0::int) -#1384 := (iff #1378 #1383) -#1385 := [rewrite]: #1384 -#1379 := (iff #1197 #1378) -#1376 := (= #1196 #1375) -#1377 := [rewrite]: #1376 -#1380 := [monotonicity #1377]: #1379 -#1387 := [trans #1380 #1385]: #1386 -#1412 := [monotonicity #1387 #1409]: #1411 -#1415 := [monotonicity #1412]: #1414 -#1418 := [monotonicity #1415]: #1417 -#1423 := [trans #1418 #1421]: #1422 -#1373 := (iff #1185 #793) -#1374 := [rewrite]: #1373 -#1442 := (iff #1243 #776) -#1443 := [rewrite]: #1442 -#1440 := (iff #1240 #361) -#1441 := [rewrite]: #1440 -#1438 := (iff #1237 #356) -#1439 := [rewrite]: #1438 -#1436 := (iff #1234 #82) -#1437 := [rewrite]: #1436 -#1446 := [monotonicity #1437 #1439 #1441 #1443 #1368 #1374 #1423 #1425 #1427]: #1445 -#1451 := [trans #1446 #1449]: #1450 -#1434 := (iff #1230 #1431) -#1428 := (and #191 #194 #781 #793 #1419 #837 #843 #850) -#1432 := (iff #1428 #1431) -#1433 := [rewrite]: #1432 -#1429 := (iff #1230 #1428) -#1371 := (iff #1182 #194) -#1372 := [rewrite]: #1371 -#1369 := (iff #1179 #191) -#1370 := [rewrite]: #1369 -#1430 := [monotonicity #1370 #1372 #1368 #1374 #1423 #1425 #1427 #895]: #1429 -#1435 := [trans #1430 #1433]: #1434 -#1454 := [monotonicity #1435 #1451]: #1453 -#1457 := [monotonicity #1368 #1454]: #1456 +#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 -#1522 := [monotonicity #1462 #1519]: #1521 -#1365 := (iff #1166 #188) -#1366 := [rewrite]: #1365 -#1525 := [monotonicity #1366 #1368 #1522]: #1524 -#1530 := [trans #1525 #1528]: #1529 -#1363 := (iff #1161 #1362) -#1360 := (iff #1160 #1357) -#1354 := (or #1351 #1154) -#1358 := (iff #1354 #1357) -#1359 := [rewrite]: #1358 -#1355 := (iff #1160 #1354) -#1352 := (iff #1159 #1351) -#1349 := (iff #1158 #1348) -#1350 := [rewrite]: #1349 -#1353 := [monotonicity #1350]: #1352 -#1356 := [monotonicity #1353]: #1355 -#1361 := [trans #1356 #1359]: #1360 -#1364 := [monotonicity #1361]: #1363 -#1533 := [monotonicity #1364 #1530]: #1532 -#1138 := (or #619 #772 #784 #1021 #1046) -#1143 := (and #769 #1138) -#1146 := (not #1143) -#1343 := (~ #1146 #1342) -#1339 := (not #1138) -#1340 := (~ #1339 #1338) -#1335 := (not #1046) -#1336 := (~ #1335 #1043) -#1333 := (~ #1043 #1043) -#1331 := (~ #1040 #1040) -#1332 := [refl]: #1331 -#1334 := [nnf-pos #1332]: #1333 -#1337 := [nnf-neg #1334]: #1336 -#1328 := (not #1021) -#1329 := (~ #1328 #1327) -#1324 := (not #1016) -#1325 := (~ #1324 #1323) -#1320 := (not #991) -#1321 := (~ #1320 #1319) -#1316 := (not #988) -#1317 := (~ #1316 #1315) -#1310 := (not #985) -#1311 := (~ #1310 #1309) -#1312 := [sk]: #1311 -#1294 := (not #969) -#1295 := (~ #1294 #1289) -#1290 := (~ #966 #1289) -#1291 := [sk]: #1290 -#1296 := [nnf-neg #1291]: #1295 -#1318 := [nnf-neg #1296 #1312]: #1317 -#1277 := (~ #969 #1276) -#1274 := (~ #1273 #1273) -#1275 := [refl]: #1274 -#1278 := [nnf-neg #1275]: #1277 -#1322 := [nnf-neg #1278 #1318]: #1321 -#1271 := (~ #930 #930) -#1272 := [refl]: #1271 -#1177 := (~ #1176 #1176) -#1178 := [refl]: #1177 -#1269 := (~ #1268 #1268) -#1270 := [refl]: #1269 -#1266 := (~ #1265 #1265) -#1267 := [refl]: #1266 -#1263 := (~ #1262 #1262) -#1264 := [refl]: #1263 -#1326 := [nnf-neg #1264 #1267 #1270 #1178 #1272 #1322]: #1325 -#1259 := (not #946) -#1260 := (~ #1259 #1258) -#1256 := (~ #927 #927) -#1257 := [refl]: #1256 -#1253 := (not #921) -#1254 := (~ #1253 #1252) -#1249 := (not #916) -#1250 := (~ #1249 #1248) -#1246 := (~ #854 #854) -#1247 := [refl]: #1246 -#1226 := (~ #1225 #1225) -#1227 := [refl]: #1226 -#1223 := (~ #1222 #1222) -#1224 := [refl]: #1223 -#1219 := (not #833) -#1220 := (~ #1219 #1218) -#1215 := (not #828) -#1216 := (~ #1215 #1214) -#1211 := (not #822) -#1212 := (~ #1211 #819) -#1209 := (~ #819 #819) -#1207 := (~ #816 #816) +#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 -#1210 := [nnf-pos #1208]: #1209 -#1213 := [nnf-neg #1210]: #1212 -#1205 := (~ #1204 #1204) -#1206 := [refl]: #1205 -#1217 := [nnf-neg #1206 #1213]: #1216 -#1200 := (~ #822 #1199) -#1201 := [sk]: #1200 -#1221 := [nnf-neg #1201 #1217]: #1220 -#1186 := (~ #1185 #1185) -#1187 := [refl]: #1186 -#1244 := (~ #1243 #1243) -#1245 := [refl]: #1244 -#1241 := (~ #1240 #1240) -#1242 := [refl]: #1241 -#1238 := (~ #1237 #1237) -#1239 := [refl]: #1238 -#1235 := (~ #1234 #1234) -#1236 := [refl]: #1235 -#1251 := [nnf-neg #1236 #1239 #1242 #1245 #1178 #1187 #1221 #1224 #1227 #1247]: #1250 -#1231 := (not #881) -#1232 := (~ #1231 #1230) -#1228 := (~ #891 #891) -#1229 := [refl]: #1228 -#1183 := (~ #1182 #1182) -#1184 := [refl]: #1183 -#1180 := (~ #1179 #1179) -#1181 := [refl]: #1180 -#1233 := [nnf-neg #1181 #1184 #1178 #1187 #1221 #1224 #1227 #1229]: #1232 -#1255 := [nnf-neg #1233 #1251]: #1254 -#1261 := [nnf-neg #1178 #1255 #1257]: #1260 -#1330 := [nnf-neg #1261 #1326]: #1329 -#1173 := (not #772) -#1174 := (~ #1173 #769) -#1171 := (~ #769 #769) -#1169 := (~ #766 #766) -#1170 := [refl]: #1169 -#1172 := [nnf-pos #1170]: #1171 -#1175 := [nnf-neg #1172]: #1174 -#1167 := (~ #1166 #1166) -#1168 := [refl]: #1167 -#1341 := [nnf-neg #1168 #1175 #1178 #1330 #1337]: #1340 -#1162 := (~ #772 #1161) -#1163 := [sk]: #1162 -#1344 := [nnf-neg #1163 #1341]: #1343 -#1110 := (not #1075) -#1147 := (iff #1110 #1146) -#1144 := (iff #1075 #1143) -#1141 := (iff #1072 #1138) -#1123 := (or #619 #784 #1021 #1046) -#1135 := (or #772 #1123) -#1139 := (iff #1135 #1138) -#1140 := [rewrite]: #1139 -#1136 := (iff #1072 #1135) -#1133 := (iff #1069 #1123) -#1128 := (and true #1123) -#1131 := (iff #1128 #1123) -#1132 := [rewrite]: #1131 -#1129 := (iff #1069 #1128) -#1126 := (iff #1064 #1123) -#1120 := (or false #619 #784 #1021 #1046) -#1124 := (iff #1120 #1123) -#1125 := [rewrite]: #1124 -#1121 := (iff #1064 #1120) -#1118 := (iff #652 false) -#1116 := (iff #652 #740) -#1115 := (iff #9 true) -#1113 := [iff-true #1109]: #1115 -#1117 := [monotonicity #1113]: #1116 -#1119 := [trans #1117 #744]: #1118 -#1122 := [monotonicity #1119]: #1121 -#1127 := [trans #1122 #1125]: #1126 -#1130 := [monotonicity #1113 #1127]: #1129 -#1134 := [trans #1130 #1132]: #1133 -#1137 := [monotonicity #1134]: #1136 -#1142 := [trans #1137 #1140]: #1141 -#1145 := [monotonicity #1142]: #1144 -#1148 := [monotonicity #1145]: #1147 -#1111 := [not-or-elim #1108]: #1110 -#1149 := [mp #1111 #1148]: #1146 -#1345 := [mp~ #1149 #1344]: #1342 -#1346 := [mp #1345 #1533]: #1531 -#1830 := [mp #1346 #1829]: #1825 -#2396 := [mp #1830 #2395]: #2393 -#1909 := [unit-resolution #2396 #2199]: #2390 -#2214 := (or #2387 #2381) -#2210 := [def-axiom]: #2214 -#2414 := [unit-resolution #2210 #1909]: #2381 -#2426 := (uf_3 uf_11) -#2430 := (= uf_12 #2426) -#2480 := (= #36 #2426) -#2478 := (= #2426 #36) -#2463 := [hypothesis]: #2378 -#2138 := (or #2375 #466) -#2139 := [def-axiom]: #2138 -#2474 := [unit-resolution #2139 #2463]: #466 -#2475 := [symm #2474]: #98 -#2479 := [monotonicity #2475]: #2478 -#2481 := [symm #2479]: #2480 -#2482 := (= uf_12 #36) -#2221 := (or #2387 #188) -#2222 := [def-axiom]: #2221 -#2476 := [unit-resolution #2222 #1909]: #188 -#2132 := (or #2375 #469) -#2140 := [def-axiom]: #2132 -#2466 := [unit-resolution #2140 #2463]: #469 -#2477 := [symm #2466]: #100 -#2483 := [trans #2477 #2476]: #2482 -#2484 := [trans #2483 #2481]: #2430 -#2458 := (not #2430) -#2424 := (>= uf_11 0::int) -#2425 := (not #2424) -#2421 := (* -1::int uf_11) -#2422 := (+ uf_1 #2421) -#2423 := (<= #2422 0::int) -#2436 := (or #2423 #2425 #2430) -#2441 := (not #2436) -#2227 := (or #2375 #2369) -#2219 := [def-axiom]: #2227 -#2464 := [unit-resolution #2219 #2463]: #2369 -#2238 := (or #2375 #926) -#2225 := [def-axiom]: #2238 -#2465 := [unit-resolution #2225 #2463]: #926 -#2031 := (+ uf_6 #972) -#2032 := (<= #2031 0::int) -#2467 := (or #551 #2032) -#2468 := [th-lemma]: #2467 -#2469 := [unit-resolution #2468 #2466]: #2032 -#1925 := (not #2032) -#1915 := (or #1795 #1925 #927) -#2004 := (+ uf_4 #1301) -#2005 := (<= #2004 0::int) -#1918 := (not #2005) -#1910 := [hypothesis]: #926 -#1911 := [hypothesis]: #1798 -#2086 := (or #1795 #1304) -#2239 := [def-axiom]: #2086 -#1916 := [unit-resolution #2239 #1911]: #1304 -#1919 := (or #1918 #927 #1303) -#1921 := [th-lemma]: #1919 -#1917 := [unit-resolution #1921 #1916 #1910]: #1918 -#2021 := (+ uf_6 #1493) -#2022 := (>= #2021 0::int) -#1930 := (not #2022) -#1936 := [hypothesis]: #2032 -#2243 := (not #1495) -#2241 := (or #1795 #2243) -#2244 := [def-axiom]: #2241 -#1922 := [unit-resolution #2244 #1911]: #2243 -#1931 := (or #1930 #1495 #1925) -#1924 := [hypothesis]: #2243 -#1926 := [hypothesis]: #2022 -#1927 := [th-lemma #1926 #1924 #1936]: false -#1906 := [lemma #1927]: #1931 -#1905 := [unit-resolution #1906 #1922 #1936]: #1930 -#1913 := (or #2005 #2022) -#2240 := (or #1795 #1305) -#2242 := [def-axiom]: #2240 -#1908 := [unit-resolution #2242 #1911]: #1305 -#2212 := (or #2387 #2312) -#2213 := [def-axiom]: #2212 -#1912 := [unit-resolution #2213 #1909]: #2312 -#1994 := (or #2317 #1731 #2005 #2022) -#2034 := (+ ?x8!3 #924) -#2035 := (>= #2034 0::int) -#2036 := (+ #1298 #1024) -#2037 := (<= #2036 0::int) -#2026 := (or #1731 #2037 #2035) -#1961 := (or #2317 #2026) -#1971 := (iff #1961 #1994) -#1991 := (or #1731 #2005 #2022) -#1964 := (or #2317 #1991) -#1969 := (iff #1964 #1994) -#1970 := [rewrite]: #1969 -#1955 := (iff #1961 #1964) -#1993 := (iff #2026 #1991) -#2008 := (or #1731 #2022 #2005) -#1983 := (iff #2008 #1991) -#1992 := [rewrite]: #1983 -#1989 := (iff #2026 #2008) -#2007 := (iff #2035 #2005) -#2010 := (+ #924 ?x8!3) -#2012 := (>= #2010 0::int) -#1997 := (iff #2012 #2005) -#2006 := [rewrite]: #1997 -#2014 := (iff #2035 #2012) -#2011 := (= #2034 #2010) -#2013 := [rewrite]: #2011 -#2003 := [monotonicity #2013]: #2014 -#1998 := [trans #2003 #2006]: #2007 -#2024 := (iff #2037 #2022) -#2038 := (+ #1024 #1298) -#2018 := (<= #2038 0::int) -#2023 := (iff #2018 #2022) -#2016 := [rewrite]: #2023 -#2019 := (iff #2037 #2018) -#2015 := (= #2036 #2038) -#2017 := [rewrite]: #2015 -#2020 := [monotonicity #2017]: #2019 -#2009 := [trans #2020 #2016]: #2024 -#1990 := [monotonicity #2009 #1998]: #1989 -#1984 := [trans #1990 #1992]: #1993 -#1968 := [monotonicity #1984]: #1955 -#1972 := [trans #1968 #1970]: #1971 -#1963 := [quant-inst]: #1961 -#1962 := [mp #1963 #1972]: #1994 -#1914 := [unit-resolution #1962 #1912 #1908]: #1913 -#1907 := [unit-resolution #1914 #1905 #1917]: false -#1900 := [lemma #1907]: #1915 -#2470 := [unit-resolution #1900 #2469 #2465]: #1795 -#2121 := (or #2372 #2364 #1798) -#2136 := [def-axiom]: #2121 -#2471 := [unit-resolution #2136 #2470 #2464]: #2364 -#2235 := (not #2364) -#2444 := (or #2235 #2441) -#2427 := (= #2426 uf_12) -#2428 := (or #2427 #2425 #2423) -#2429 := (not #2428) -#2445 := (or #2235 #2429) -#2447 := (iff #2445 #2444) -#2449 := (iff #2444 #2444) -#2450 := [rewrite]: #2449 -#2442 := (iff #2429 #2441) -#2439 := (iff #2428 #2436) -#2433 := (or #2430 #2425 #2423) -#2437 := (iff #2433 #2436) -#2438 := [rewrite]: #2437 -#2434 := (iff #2428 #2433) -#2431 := (iff #2427 #2430) -#2432 := [rewrite]: #2431 -#2435 := [monotonicity #2432]: #2434 -#2440 := [trans #2435 #2438]: #2439 -#2443 := [monotonicity #2440]: #2442 -#2448 := [monotonicity #2443]: #2447 -#2451 := [trans #2448 #2450]: #2447 -#2446 := [quant-inst]: #2445 -#2452 := [mp #2446 #2451]: #2444 -#2472 := [unit-resolution #2452 #2471]: #2441 -#2459 := (or #2436 #2458) -#2460 := [def-axiom]: #2459 -#2473 := [unit-resolution #2460 #2472]: #2458 -#2485 := [unit-resolution #2473 #2484]: false -#2486 := [lemma #2485]: #2375 -#2231 := (or #2384 #2361 #2378) -#2220 := [def-axiom]: #2231 -#2415 := [unit-resolution #2220 #2486 #2414]: #2361 -#2106 := (or #2358 #2352) -#2248 := [def-axiom]: #2106 -#2416 := [unit-resolution #2248 #2415]: #2352 -#2417 := [hypothesis]: #840 -#2285 := (or #2340 #837) -#1923 := [def-axiom]: #2285 -#2418 := [unit-resolution #1923 #2417]: #2340 -#1987 := (or #2346 #837) -#1988 := [def-axiom]: #1987 -#2419 := [unit-resolution #1988 #2417]: #2346 -#2255 := (or #2355 #2343 #2349) -#2256 := [def-axiom]: #2255 -#2420 := [unit-resolution #2256 #2419 #2418 #2416]: false -#2403 := [lemma #2420]: #837 -#2690 := (or #840 #1977) -#2691 := [th-lemma]: #2690 -#2692 := [unit-resolution #2691 #2403]: #1977 -#2661 := [hypothesis]: #2349 -#2272 := (or #2346 #361) -#2273 := [def-axiom]: #2272 -#2662 := [unit-resolution #2273 #2661]: #361 -#2629 := (= #58 #1195) -#2642 := (not #2629) -#2630 := (+ #58 #1381) -#2632 := (>= #2630 0::int) -#2636 := (not #2632) -#2402 := (+ #39 #799) -#2405 := (<= #2402 0::int) -#2404 := (= #39 uf_8) -#2665 := (= uf_10 uf_8) -#2000 := (or #2346 #82) -#2001 := [def-axiom]: #2000 -#2663 := [unit-resolution #2001 #2661]: #82 -#2666 := [symm #2663]: #2665 -#2002 := (or #2346 #356) -#1896 := [def-axiom]: #2002 -#2664 := [unit-resolution #1896 #2661]: #356 -#2667 := [trans #2664 #2666]: #2404 -#2668 := (not #2404) -#2669 := (or #2668 #2405) -#2670 := [th-lemma]: #2669 -#2671 := [unit-resolution #2670 #2667]: #2405 -#1966 := (not #1383) -#1982 := (or #2346 #2334) -#2264 := [def-axiom]: #1982 -#2672 := [unit-resolution #2264 #2661]: #2334 -#2537 := (= #39 #58) -#2674 := (= #58 #39) -#2673 := [symm #2662]: #81 -#2675 := [monotonicity #2673]: #2674 -#2677 := [symm #2675]: #2537 -#2678 := (= uf_8 #39) -#2676 := [symm #2664]: #79 -#2679 := [trans #2663 #2676]: #2678 -#2680 := [trans #2679 #2677]: #221 -#1981 := (or #2328 #1204) -#1960 := [def-axiom]: #1981 -#2681 := [unit-resolution #1960 #2680]: #2328 -#1953 := (or #2337 #2331 #1645) -#2294 := [def-axiom]: #1953 -#2682 := [unit-resolution #2294 #2681 #2672]: #1645 -#2298 := (or #1640 #1966) -#2299 := [def-axiom]: #2298 -#2683 := [unit-resolution #2299 #2682]: #1966 -#2033 := (* -1::int #58) -#2538 := (+ #39 #2033) -#2540 := (>= #2538 0::int) -#2684 := (not #2537) -#2685 := (or #2684 #2540) -#2686 := [th-lemma]: #2685 -#2687 := [unit-resolution #2686 #2677]: #2540 -#2617 := (not #2405) -#2637 := (not #2540) -#2638 := (or #2636 #2637 #1383 #2617) -#2633 := [hypothesis]: #2632 -#2609 := [hypothesis]: #2405 -#2610 := [hypothesis]: #1966 -#2634 := [hypothesis]: #2540 -#2635 := [th-lemma #2634 #2610 #2609 #2633]: false -#2639 := [lemma #2635]: #2638 -#2688 := [unit-resolution #2639 #2687 #2683 #2671]: #2636 -#2643 := (or #2642 #2632) -#2644 := [th-lemma]: #2643 -#2689 := [unit-resolution #2644 #2688]: #2642 -#2300 := (or #1640 #1401) -#2301 := [def-axiom]: #2300 -#2693 := [unit-resolution #2301 #2682]: #1401 -#2694 := (not #1977) -#2695 := (or #2628 #1396 #2694) -#2696 := [th-lemma]: #2695 -#2697 := [unit-resolution #2696 #2693 #2692]: #2628 -#2565 := (<= #2564 0::int) -#2552 := (+ uf_6 #1381) -#2553 := (>= #2552 0::int) -#2699 := (not #2553) -#2266 := (or #2346 #854) -#2267 := [def-axiom]: #2266 -#2698 := [unit-resolution #2267 #2661]: #854 -#2700 := (or #2699 #1383 #2617 #850) -#2701 := [th-lemma]: #2700 -#2702 := [unit-resolution #2701 #2683 #2671 #2698]: #2699 -#2704 := (or #2553 #2565) -#2291 := (or #1640 #1192) -#1965 := [def-axiom]: #2291 -#2703 := [unit-resolution #1965 #2682]: #1192 -#2573 := (or #2317 #1625 #2553 #2565) -#2541 := (+ ?x3!1 #924) -#2542 := (>= #2541 0::int) -#2543 := (+ #1195 #1024) -#2544 := (<= #2543 0::int) -#2545 := (or #1625 #2544 #2542) -#2574 := (or #2317 #2545) -#2581 := (iff #2574 #2573) -#2570 := (or #1625 #2553 #2565) -#2576 := (or #2317 #2570) -#2579 := (iff #2576 #2573) -#2580 := [rewrite]: #2579 -#2577 := (iff #2574 #2576) -#2571 := (iff #2545 #2570) -#2568 := (iff #2542 #2565) -#2558 := (+ #924 ?x3!1) -#2561 := (>= #2558 0::int) -#2566 := (iff #2561 #2565) -#2567 := [rewrite]: #2566 -#2562 := (iff #2542 #2561) -#2559 := (= #2541 #2558) -#2560 := [rewrite]: #2559 -#2563 := [monotonicity #2560]: #2562 -#2569 := [trans #2563 #2567]: #2568 -#2556 := (iff #2544 #2553) -#2546 := (+ #1024 #1195) -#2549 := (<= #2546 0::int) -#2554 := (iff #2549 #2553) -#2555 := [rewrite]: #2554 -#2550 := (iff #2544 #2549) -#2547 := (= #2543 #2546) -#2548 := [rewrite]: #2547 -#2551 := [monotonicity #2548]: #2550 -#2557 := [trans #2551 #2555]: #2556 -#2572 := [monotonicity #2557 #2569]: #2571 -#2578 := [monotonicity #2572]: #2577 -#2582 := [trans #2578 #2580]: #2581 -#2575 := [quant-inst]: #2574 -#2583 := [mp #2575 #2582]: #2573 -#2705 := [unit-resolution #2583 #1912 #2703]: #2704 -#2706 := [unit-resolution #2705 #2702]: #2565 -#2708 := (not #2628) -#2707 := (not #2565) -#2709 := (or #2631 #2707 #2708) -#2710 := [th-lemma]: #2709 -#2711 := [unit-resolution #2710 #2706 #2697]: #2631 -#2658 := (not #2631) -#2659 := (or #2658 #2629 #376) -#2654 := (= #1195 #58) -#2652 := (= ?x3!1 uf_7) -#2648 := [hypothesis]: #361 -#2650 := (= ?x3!1 uf_4) -#2649 := [hypothesis]: #2631 -#2651 := [symm #2649]: #2650 -#2653 := [trans #2651 #2648]: #2652 -#2655 := [monotonicity #2653]: #2654 -#2656 := [symm #2655]: #2629 -#2647 := [hypothesis]: #2642 -#2657 := [unit-resolution #2647 #2656]: false -#2660 := [lemma #2657]: #2659 -#2712 := [unit-resolution #2660 #2711 #2689 #2662]: false -#2713 := [lemma #2712]: #2346 -#2766 := [unit-resolution #2256 #2713 #2416]: #2343 -#1928 := (or #2340 #2334) -#1929 := [def-axiom]: #1928 -#2767 := [unit-resolution #1929 #2766]: #2334 -#2515 := (= #36 #58) -#2772 := (= #58 #36) -#1937 := (or #2340 #191) -#2278 := [def-axiom]: #1937 -#2768 := [unit-resolution #2278 #2766]: #191 -#2769 := [symm #2768]: #42 -#2773 := [monotonicity #2769]: #2772 -#2774 := [symm #2773]: #2515 -#2775 := (= uf_8 #36) -#1941 := (or #2340 #194) -#1942 := [def-axiom]: #1941 -#2770 := [unit-resolution #1942 #2766]: #194 -#2771 := [symm #2770]: #44 -#2776 := [trans #2771 #2476]: #2775 -#2777 := [trans #2776 #2774]: #221 -#2778 := [unit-resolution #1960 #2777]: #2328 -#2779 := [unit-resolution #2294 #2778 #2767]: #1645 -#2780 := [unit-resolution #2301 #2779]: #1401 -#2781 := [unit-resolution #2696 #2780 #2692]: #2628 -#2510 := (+ uf_6 #799) -#2511 := (<= #2510 0::int) -#2782 := (or #301 #2511) -#2783 := [th-lemma]: #2782 -#2784 := [unit-resolution #2783 #2770]: #2511 -#2785 := [unit-resolution #2299 #2779]: #1966 -#2786 := (not #2511) -#2787 := (or #2699 #1383 #2786) -#2788 := [th-lemma]: #2787 -#2789 := [unit-resolution #2788 #2785 #2784]: #2699 -#2790 := [unit-resolution #1965 #2779]: #1192 -#2791 := [unit-resolution #2583 #1912 #2790 #2789]: #2565 -#2792 := [unit-resolution #2710 #2791 #2781]: #2631 -#2793 := [monotonicity #2792]: #2762 -#2794 := (not #2762) -#2795 := (or #2794 #2765) -#2796 := [th-lemma]: #2795 -#2797 := [unit-resolution #2796 #2793]: #2765 -#2286 := (or #2340 #850) -#2288 := [def-axiom]: #2286 -#2798 := [unit-resolution #2288 #2766]: #850 -[th-lemma #2798 #2785 #2784 #2797]: false +#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 unsat diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Boogie/Examples/cert/VCC_maximum --- a/src/HOL/Boogie/Examples/cert/VCC_maximum Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Boogie/Examples/cert/VCC_maximum Sat Nov 14 09:40:27 2009 +0100 @@ -8,7 +8,7 @@ (uf_66 T5 Int T3 T5) (uf_43 T3 Int T5) (uf_116 T5 Int) - (uf_15 T5 T3) + (uf_13 T5 T3) (uf_81 Int Int Int) (uf_80 Int Int Int) (uf_46 T4 T4 T5 T3 T2) @@ -18,7 +18,7 @@ (uf_72 T3 Int Int Int) (uf_124 T3 Int T3) (uf_25 T4 T5 T5) - (uf_27 T4 T5 T2) + (uf_24 T4 T5 T2) (uf_255 T3) (uf_254 T3) (uf_94 T3) @@ -35,11 +35,11 @@ (uf_101 T3 Int Int Int) (uf_100 Int Int Int) (uf_71 T3 Int Int Int) - (uf_24 T4 T5 T2) - (uf_10 T4 T5 T6) + (uf_27 T4 T5 T2) + (uf_16 T4 T5 T6) (uf_128 T4 T5 T6) (uf_20 T4 T9) - (uf_138 T3 Int) + (uf_139 T3 Int) (uf_5 T3) (uf_291 T1) (uf_79 Int Int) @@ -51,31 +51,31 @@ (uf_258 T3) (uf_240 T3) (uf_284 T16) - (uf_95 Int) - (uf_92 Int) - (uf_88 Int) - (uf_85 Int) + (uf_96 Int) + (uf_93 Int) + (uf_89 Int) + (uf_86 Int) (uf_78 Int) (uf_77 Int) (uf_76 Int) (uf_75 Int) (uf_253 Int) - (uf_96 Int) - (uf_93 Int) - (uf_89 Int) - (uf_86 Int) + (uf_95 Int) + (uf_92 Int) + (uf_88 Int) + (uf_85 Int) (uf_6 T3 T3) (uf_224 T17 T17 T2) - (uf_173 T4 T5 T5 T11) + (uf_171 T4 T5 T5 T11) (uf_153 T6 T6 T2) - (uf_13 T5 T6 T2) - (uf_136 T14 T5) + (uf_15 T5 T6 T2) + (uf_135 T14 T5) (uf_37 T3) (uf_279 T1) (uf_281 T1) (uf_287 T1) (uf_122 T2 T2) - (uf_14 T3 T8) + (uf_12 T3 T8) (uf_114 T4 T5 Int) (uf_113 T4 T5 Int) (uf_112 T4 T5 Int) @@ -88,7 +88,7 @@ (uf_145 T5 T6 T2) (uf_147 T5 T6 T2) (uf_41 T4 T12) - (uf_170 T4 T5 Int) + (uf_172 T4 T5 Int) (uf_82 T3 Int Int) (uf_232 T4 T5 T18) (uf_188 T4 T5 T5 T5 T5) @@ -97,15 +97,15 @@ (uf_230 T17) (uf_179 T4 T4 T5 T3 T2) (uf_215 T11 T5) - (uf_172 T12 T5 T11 T12) + (uf_170 T12 T5 T11 T12) (uf_251 T13 T5 T14 T13) (uf_266 T3 T3) - (uf_233 T18 T4) + (uf_234 T18 T4) (uf_257 T3) (uf_99 Int Int Int Int Int Int) (uf_55 T4 T2) (uf_60 Int T3 T5) - (uf_246 Int T5) + (uf_245 Int T5) (uf_220 T5 T15 Int) (uf_196 T4 T5 T5 T2) (uf_264 T3 T3) @@ -115,7 +115,7 @@ (uf_58 T13 T5 T14) (uf_152 T6) (uf_157 T6 T6 T6) - (uf_178 T9 T5 Int T9) + (uf_177 T9 T5 Int T9) (uf_174 T4 T5 T5 T4) (uf_106 T3 Int Int Int) (uf_103 T3 Int Int Int) @@ -124,7 +124,7 @@ (uf_105 T3 Int Int Int) (uf_241 T15 Int T15) (uf_50 T5 T5 T2) - (uf_245 Int T15) + (uf_246 Int T15) (uf_51 T4 T2) (uf_195 T4 T5 T5 T2) (uf_262 T8) @@ -141,12 +141,12 @@ (uf_268 T3) (uf_289 T1) (uf_132 T5 T3 Int T6) - (uf_139 T5 T5 T2) + (uf_138 T5 T5 T2) (uf_276 T19 Int) (uf_130 T5 T6) (uf_44 T4 T2) (uf_261 T8) - (uf_248 T3 T3 Int) + (uf_250 T3 T3 Int) (uf_249 T3 T3 Int) (uf_181 T4 T4 T2) (uf_117 T5 Int) @@ -166,16 +166,16 @@ (uf_28 Int T5) (uf_141 T3 T2) (uf_260 T3 T2) - (uf_23 T3 T2) + (uf_22 T3 T2) (uf_159 T5 T5 T5) (uf_29 T5 Int) (uf_201 T4 T5 T3 T5) - (uf_12 T4 T5 T7) + (uf_11 T4 T5 T7) (uf_131 T6 T6 T2) (uf_149 T6) (uf_39 T11 Int) (uf_217 T11 Int) - (uf_67 T4 T5 T2) + (uf_68 T4 T5 T2) (uf_275 T1) (uf_134 T5 T3 Int T6) (uf_189 T5 T7) @@ -184,26 +184,26 @@ (uf_221 Int Int T2) (uf_151 T5 T6) (uf_162 T4 T5 T6) - (uf_234 T18 Int) + (uf_233 T18 Int) (uf_256 T3) (uf_286 T1) (uf_288 T1) (uf_295 T1) (uf_290 T1) - (uf_305 T1) + (uf_301 T1) (uf_243 T15 T15) - (uf_242 T15 Int) + (uf_244 T15 Int) (uf_45 T4 T5 T2) (uf_203 T4 T2) (uf_148 T5 T2) (uf_283 Int T5 T2) (uf_57 T3 T2) (uf_263 T8) - (uf_16 T8) + (uf_14 T8) (uf_156 T6 T6 T6) - (uf_303 T1) (uf_306 T1) - (uf_177 T4 T4 T2) + (uf_302 T1) + (uf_178 T4 T4 T2) (uf_183 T10 T5 Int) (uf_62 Int Int) (uf_63 Int Int) @@ -225,20 +225,20 @@ (uf_214 T3 T15) (uf_155 T6 T6 T6) (uf_206 T4 T4 T5 T3 T2) - (uf_135 T14 T2) + (uf_136 T14 T2) (uf_33 T7 Int) (uf_236 T5 T15 T5) - (uf_171 T4 Int) + (uf_173 T4 Int) (uf_133 T5 T6 T6 Int) (uf_186 T5 T5 T2) (uf_247 T3 T3 Int Int T2) (uf_227 T3 T15 T3 T2) (uf_127 T4 T5 T6) - (uf_22 T3 T2) + (uf_23 T3 T2) (uf_184 T4 T5 T10) (uf_97 Int Int Int Int Int) (uf_8 T4 T4 T5 T6 T2) - (uf_11 T7 T5 Int) + (uf_10 T7 T5 Int) (uf_238 T15 T3) (uf_210 T4 T5 T2) (uf_180 T3 T15 T2) @@ -270,21 +270,21 @@ (uf_273 T4) (uf_270 Int) (uf_294 Int) - (uf_302 Int) + (uf_305 Int) (uf_297 Int) (uf_269 Int) (uf_274 Int) (uf_272 Int) (uf_285 Int) (uf_292 Int) - (uf_304 Int) (uf_300 Int) + (uf_303 Int) (uf_296 Int) (uf_299 Int) (uf_271 Int) (uf_282 Int) (uf_293 Int) - (uf_301 Int) + (uf_304 Int) (uf_298 Int) ) :extrapreds ( @@ -292,20 +292,20 @@ (up_146 T5 T6) (up_213 T14) (up_209 T4 T5 T3) - (up_250 T3 T3) + (up_248 T3 T3) (up_218 T11) (up_36 T3) (up_1 Int T1) (up_212 T11) (up_3 Int T3) (up_182 Int) - (up_244 T15) + (up_242 T15) (up_216) (up_193 T2) - (up_280 T4 T1 T1 Int T3) + (up_278 T4 T1 T1 Int T3) (up_52 T6) - (up_68 T14) - (up_278 T4 T1 T1 T5 T3) + (up_67 T14) + (up_280 T4 T1 T1 T5 T3) (up_197 T3) (up_165 T4) (up_205 T4 T4 T5 T3) @@ -313,9 +313,9 @@ :assumption (up_1 1 uf_2) :assumption (up_3 1 uf_4) :assumption (= uf_5 (uf_6 uf_7)) -:assumption (forall (?x1 T4) (?x2 T4) (?x3 T5) (?x4 T6) (iff (= (uf_8 ?x1 ?x2 ?x3 ?x4) uf_9) (and (= (uf_10 ?x1 ?x3) (uf_10 ?x2 ?x3)) (forall (?x5 T5) (implies (and (not (= (uf_13 ?x5 ?x4) uf_9)) (= (uf_14 (uf_15 ?x5)) uf_16)) (= (uf_11 (uf_12 ?x1 ?x3) ?x5) (uf_11 (uf_12 ?x2 ?x3) ?x5))) :pat { (uf_11 (uf_12 ?x2 ?x3) ?x5) }))) :pat { (uf_8 ?x1 ?x2 ?x3 ?x4) }) -:assumption (forall (?x6 T4) (?x7 T4) (?x8 T6) (implies (forall (?x9 T5) (implies (and (not (= (uf_14 (uf_15 ?x9)) uf_16)) (= (uf_13 ?x9 ?x8) uf_9)) (or (= (uf_8 ?x6 ?x7 ?x9 ?x8) uf_9) (= (uf_19 (uf_20 ?x6) ?x9) (uf_19 (uf_20 ?x7) ?x9)))) :pat { (uf_18 ?x9) }) (= (uf_17 ?x6 ?x7 ?x8) uf_9)) :pat { (uf_17 ?x6 ?x7 ?x8) }) -:assumption (forall (?x10 T4) (?x11 T4) (?x12 T6) (implies (forall (?x13 T5) (implies (or (= (uf_22 (uf_15 ?x13)) uf_9) (= (uf_23 (uf_15 ?x13)) uf_9)) (implies (and (not (or (and (= (uf_24 ?x10 ?x13) uf_9) (= (uf_14 (uf_15 ?x13)) uf_16)) (not (= (uf_25 ?x10 ?x13) uf_26)))) (= (uf_27 ?x10 ?x13) uf_9)) (or (= (uf_13 ?x13 ?x12) uf_9) (= (uf_19 (uf_20 ?x10) ?x13) (uf_19 (uf_20 ?x11) ?x13))))) :pat { (uf_18 ?x13) }) (= (uf_21 ?x10 ?x11 ?x12) uf_9)) :pat { (uf_21 ?x10 ?x11 ?x12) }) +:assumption (forall (?x1 T4) (?x2 T4) (?x3 T5) (?x4 T6) (iff (= (uf_8 ?x1 ?x2 ?x3 ?x4) uf_9) (and (forall (?x5 T5) (implies (and (= (uf_12 (uf_13 ?x5)) uf_14) (not (= (uf_15 ?x5 ?x4) uf_9))) (= (uf_10 (uf_11 ?x1 ?x3) ?x5) (uf_10 (uf_11 ?x2 ?x3) ?x5))) :pat { (uf_10 (uf_11 ?x2 ?x3) ?x5) }) (= (uf_16 ?x1 ?x3) (uf_16 ?x2 ?x3)))) :pat { (uf_8 ?x1 ?x2 ?x3 ?x4) }) +:assumption (forall (?x6 T4) (?x7 T4) (?x8 T6) (implies (forall (?x9 T5) (implies (and (= (uf_15 ?x9 ?x8) uf_9) (not (= (uf_12 (uf_13 ?x9)) uf_14))) (or (= (uf_19 (uf_20 ?x6) ?x9) (uf_19 (uf_20 ?x7) ?x9)) (= (uf_8 ?x6 ?x7 ?x9 ?x8) uf_9))) :pat { (uf_18 ?x9) }) (= (uf_17 ?x6 ?x7 ?x8) uf_9)) :pat { (uf_17 ?x6 ?x7 ?x8) }) +:assumption (forall (?x10 T4) (?x11 T4) (?x12 T6) (implies (forall (?x13 T5) (implies (or (= (uf_22 (uf_13 ?x13)) uf_9) (= (uf_23 (uf_13 ?x13)) uf_9)) (implies (and (= (uf_24 ?x10 ?x13) uf_9) (not (or (not (= (uf_25 ?x10 ?x13) uf_26)) (and (= (uf_12 (uf_13 ?x13)) uf_14) (= (uf_27 ?x10 ?x13) uf_9))))) (or (= (uf_19 (uf_20 ?x10) ?x13) (uf_19 (uf_20 ?x11) ?x13)) (= (uf_15 ?x13 ?x12) uf_9)))) :pat { (uf_18 ?x13) }) (= (uf_21 ?x10 ?x11 ?x12) uf_9)) :pat { (uf_21 ?x10 ?x11 ?x12) }) :assumption (forall (?x14 T5) (= (uf_28 (uf_29 ?x14)) ?x14)) :assumption (forall (?x15 T10) (= (uf_30 (uf_31 ?x15)) ?x15)) :assumption (forall (?x16 T7) (= (uf_32 (uf_33 ?x16)) ?x16)) @@ -323,24 +323,24 @@ :assumption (up_36 uf_37) :assumption (forall (?x18 T4) (?x19 T5) (= (uf_38 ?x18 ?x19) (uf_39 (uf_40 (uf_41 ?x18) ?x19))) :pat { (uf_38 ?x18 ?x19) }) :assumption (= uf_42 (uf_43 uf_37 0)) -:assumption (forall (?x20 T4) (?x21 T5) (implies (and (= (uf_45 ?x20 ?x21) uf_9) (= (uf_44 ?x20) uf_9)) (= (uf_46 ?x20 ?x20 ?x21 (uf_15 ?x21)) uf_9)) :pat { (uf_44 ?x20) (uf_45 ?x20 ?x21) }) -:assumption (forall (?x22 T4) (?x23 T5) (iff (= (uf_45 ?x22 ?x23) uf_9) (= (uf_24 ?x22 ?x23) uf_9)) :pat { (uf_45 ?x22 ?x23) }) -:assumption (forall (?x24 T4) (?x25 T5) (iff (= (uf_47 ?x24 ?x25) uf_9) (and (or (= (uf_38 ?x24 ?x25) 0) (not (up_36 (uf_15 ?x25)))) (and (= (uf_22 (uf_15 ?x25)) uf_9) (and (not (= (uf_14 (uf_15 ?x25)) uf_16)) (and (= (uf_27 ?x24 ?x25) uf_9) (and (= (uf_48 ?x25 (uf_15 ?x25)) uf_9) (and (= (uf_25 ?x24 ?x25) uf_26) (= (uf_24 ?x24 ?x25) uf_9)))))))) :pat { (uf_47 ?x24 ?x25) }) -:assumption (forall (?x26 T4) (?x27 T5) (?x28 Int) (implies (and (= (uf_50 ?x27 (uf_43 uf_37 ?x28)) uf_9) (= (uf_49 ?x26 ?x27) uf_9)) (= (uf_49 ?x26 (uf_43 uf_37 ?x28)) uf_9)) :pat { (uf_49 ?x26 ?x27) (uf_50 ?x27 (uf_43 uf_37 ?x28)) }) -:assumption (forall (?x29 T4) (?x30 T5) (?x31 T5) (implies (and (= (uf_50 ?x30 ?x31) uf_9) (= (uf_49 ?x29 ?x30) uf_9)) (= (uf_46 ?x29 ?x29 ?x31 (uf_15 ?x31)) uf_9)) :pat { (uf_49 ?x29 ?x30) (uf_50 ?x30 ?x31) }) -:assumption (forall (?x32 T4) (?x33 T5) (?x34 T5) (implies (= (uf_51 ?x32) uf_9) (implies (and (= (uf_24 ?x32 ?x33) uf_9) (= (uf_50 ?x33 ?x34) uf_9)) (and (< 0 (uf_38 ?x32 ?x34)) (and (= (uf_24 ?x32 ?x34) uf_9) (up_52 (uf_53 ?x32 ?x34)))))) :pat { (uf_24 ?x32 ?x33) (uf_50 ?x33 ?x34) }) -:assumption (forall (?x35 T4) (?x36 T5) (?x37 T5) (implies (and (= (uf_54 ?x36 ?x37) uf_9) (= (uf_49 ?x35 ?x36) uf_9)) (= (uf_49 ?x35 ?x37) uf_9)) :pat { (uf_49 ?x35 ?x36) (uf_54 ?x36 ?x37) }) -:assumption (forall (?x38 T5) (?x39 T5) (implies (and (forall (?x40 T4) (implies (= (uf_49 ?x40 ?x38) uf_9) (= (uf_24 ?x40 ?x39) uf_9))) (and (= (uf_48 ?x39 uf_37) uf_9) (= (uf_48 ?x38 uf_37) uf_9))) (= (uf_54 ?x38 ?x39) uf_9)) :pat { (uf_54 ?x38 ?x39) }) -:assumption (forall (?x41 T4) (?x42 T5) (implies (= (uf_49 ?x41 ?x42) uf_9) (and (= (uf_44 ?x41) uf_9) (= (uf_24 ?x41 ?x42) uf_9))) :pat { (uf_49 ?x41 ?x42) }) -:assumption (forall (?x43 T4) (?x44 T5) (implies (and (= (uf_24 ?x43 ?x44) uf_9) (= (uf_55 ?x43) uf_9)) (= (uf_49 ?x43 ?x44) uf_9)) :pat { (uf_55 ?x43) (uf_49 ?x43 ?x44) }) -:assumption (forall (?x45 T3) (implies (= (uf_56 ?x45) uf_9) (= (uf_23 ?x45) uf_9)) :pat { (uf_56 ?x45) }) -:assumption (forall (?x46 T3) (implies (= (uf_57 ?x46) uf_9) (= (uf_23 ?x46) uf_9)) :pat { (uf_57 ?x46) }) -:assumption (forall (?x47 T4) (?x48 Int) (?x49 T3) (implies (and (= (uf_51 ?x47) uf_9) (= (uf_56 ?x49) uf_9)) (= (uf_61 ?x47 (uf_60 ?x48 ?x49)) uf_9)) :pat { (uf_58 (uf_59 ?x47) (uf_60 ?x48 ?x49)) } :pat { (uf_40 (uf_41 ?x47) (uf_60 ?x48 ?x49)) }) +:assumption (forall (?x20 T4) (?x21 T5) (implies (and (= (uf_44 ?x20) uf_9) (= (uf_45 ?x20 ?x21) uf_9)) (= (uf_46 ?x20 ?x20 ?x21 (uf_13 ?x21)) uf_9)) :pat { (uf_44 ?x20) (uf_45 ?x20 ?x21) }) +:assumption (forall (?x22 T4) (?x23 T5) (iff (= (uf_45 ?x22 ?x23) uf_9) (= (uf_27 ?x22 ?x23) uf_9)) :pat { (uf_45 ?x22 ?x23) }) +:assumption (forall (?x24 T4) (?x25 T5) (iff (= (uf_47 ?x24 ?x25) uf_9) (and (= (uf_27 ?x24 ?x25) uf_9) (and (= (uf_25 ?x24 ?x25) uf_26) (and (= (uf_48 ?x25 (uf_13 ?x25)) uf_9) (and (= (uf_24 ?x24 ?x25) uf_9) (and (not (= (uf_12 (uf_13 ?x25)) uf_14)) (and (= (uf_23 (uf_13 ?x25)) uf_9) (or (not (up_36 (uf_13 ?x25))) (= (uf_38 ?x24 ?x25) 0))))))))) :pat { (uf_47 ?x24 ?x25) }) +:assumption (forall (?x26 T4) (?x27 T5) (?x28 Int) (implies (and (= (uf_49 ?x26 ?x27) uf_9) (= (uf_50 ?x27 (uf_43 uf_37 ?x28)) uf_9)) (= (uf_49 ?x26 (uf_43 uf_37 ?x28)) uf_9)) :pat { (uf_49 ?x26 ?x27) (uf_50 ?x27 (uf_43 uf_37 ?x28)) }) +:assumption (forall (?x29 T4) (?x30 T5) (?x31 T5) (implies (and (= (uf_49 ?x29 ?x30) uf_9) (= (uf_50 ?x30 ?x31) uf_9)) (= (uf_46 ?x29 ?x29 ?x31 (uf_13 ?x31)) uf_9)) :pat { (uf_49 ?x29 ?x30) (uf_50 ?x30 ?x31) }) +:assumption (forall (?x32 T4) (?x33 T5) (?x34 T5) (implies (= (uf_51 ?x32) uf_9) (implies (and (= (uf_50 ?x33 ?x34) uf_9) (= (uf_27 ?x32 ?x33) uf_9)) (and (up_52 (uf_53 ?x32 ?x34)) (and (= (uf_27 ?x32 ?x34) uf_9) (< 0 (uf_38 ?x32 ?x34)))))) :pat { (uf_27 ?x32 ?x33) (uf_50 ?x33 ?x34) }) +:assumption (forall (?x35 T4) (?x36 T5) (?x37 T5) (implies (and (= (uf_49 ?x35 ?x36) uf_9) (= (uf_54 ?x36 ?x37) uf_9)) (= (uf_49 ?x35 ?x37) uf_9)) :pat { (uf_49 ?x35 ?x36) (uf_54 ?x36 ?x37) }) +:assumption (forall (?x38 T5) (?x39 T5) (implies (and (= (uf_48 ?x38 uf_37) uf_9) (and (= (uf_48 ?x39 uf_37) uf_9) (forall (?x40 T4) (implies (= (uf_49 ?x40 ?x38) uf_9) (= (uf_27 ?x40 ?x39) uf_9))))) (= (uf_54 ?x38 ?x39) uf_9)) :pat { (uf_54 ?x38 ?x39) }) +:assumption (forall (?x41 T4) (?x42 T5) (implies (= (uf_49 ?x41 ?x42) uf_9) (and (= (uf_27 ?x41 ?x42) uf_9) (= (uf_44 ?x41) uf_9))) :pat { (uf_49 ?x41 ?x42) }) +:assumption (forall (?x43 T4) (?x44 T5) (implies (and (= (uf_55 ?x43) uf_9) (= (uf_27 ?x43 ?x44) uf_9)) (= (uf_49 ?x43 ?x44) uf_9)) :pat { (uf_55 ?x43) (uf_49 ?x43 ?x44) }) +:assumption (forall (?x45 T3) (implies (= (uf_56 ?x45) uf_9) (= (uf_22 ?x45) uf_9)) :pat { (uf_56 ?x45) }) +:assumption (forall (?x46 T3) (implies (= (uf_57 ?x46) uf_9) (= (uf_22 ?x46) uf_9)) :pat { (uf_57 ?x46) }) +:assumption (forall (?x47 T4) (?x48 Int) (?x49 T3) (implies (and (= (uf_56 ?x49) uf_9) (= (uf_51 ?x47) uf_9)) (= (uf_61 ?x47 (uf_60 ?x48 ?x49)) uf_9)) :pat { (uf_58 (uf_59 ?x47) (uf_60 ?x48 ?x49)) } :pat { (uf_40 (uf_41 ?x47) (uf_60 ?x48 ?x49)) }) :assumption (forall (?x50 Int) (= (uf_62 (uf_63 ?x50)) ?x50)) :assumption (forall (?x51 Int) (?x52 T3) (= (uf_60 ?x51 ?x52) (uf_43 ?x52 (uf_63 ?x51))) :pat { (uf_60 ?x51 ?x52) }) -:assumption (forall (?x53 Int) (?x54 Int) (?x55 T4) (implies (= (uf_51 ?x55) uf_9) (and (forall (?x56 Int) (implies (and (< ?x56 ?x54) (<= 0 ?x56)) (and (= (uf_67 ?x55 (uf_66 (uf_64 ?x53 ?x54) ?x56 uf_7)) uf_9) (and (= (uf_48 (uf_66 (uf_64 ?x53 ?x54) ?x56 uf_7) uf_7) uf_9) (up_68 (uf_58 (uf_59 ?x55) (uf_66 (uf_64 ?x53 ?x54) ?x56 uf_7)))))) :pat { (uf_40 (uf_41 ?x55) (uf_66 (uf_64 ?x53 ?x54) ?x56 uf_7)) } :pat { (uf_58 (uf_59 ?x55) (uf_66 (uf_64 ?x53 ?x54) ?x56 uf_7)) }) (= (uf_27 ?x55 (uf_64 ?x53 ?x54)) uf_9))) :pat { (uf_27 ?x55 (uf_64 ?x53 ?x54)) } :pat { (uf_65 ?x55 (uf_64 ?x53 ?x54) uf_7 ?x54) }) +:assumption (forall (?x53 Int) (?x54 Int) (?x55 T4) (implies (= (uf_51 ?x55) uf_9) (and (= (uf_24 ?x55 (uf_64 ?x53 ?x54)) uf_9) (forall (?x56 Int) (implies (and (<= 0 ?x56) (< ?x56 ?x54)) (and (up_67 (uf_58 (uf_59 ?x55) (uf_66 (uf_64 ?x53 ?x54) ?x56 uf_7))) (and (= (uf_48 (uf_66 (uf_64 ?x53 ?x54) ?x56 uf_7) uf_7) uf_9) (= (uf_68 ?x55 (uf_66 (uf_64 ?x53 ?x54) ?x56 uf_7)) uf_9)))) :pat { (uf_40 (uf_41 ?x55) (uf_66 (uf_64 ?x53 ?x54) ?x56 uf_7)) } :pat { (uf_58 (uf_59 ?x55) (uf_66 (uf_64 ?x53 ?x54) ?x56 uf_7)) }))) :pat { (uf_24 ?x55 (uf_64 ?x53 ?x54)) } :pat { (uf_65 ?x55 (uf_64 ?x53 ?x54) uf_7 ?x54) }) :assumption (forall (?x57 Int) (?x58 Int) (= (uf_48 (uf_64 ?x57 ?x58) uf_7) uf_9) :pat { (uf_64 ?x57 ?x58) }) -:assumption (forall (?x59 Int) (?x60 Int) (= (uf_69 ?x59 ?x60) (+ ?x59 ?x60)) :pat { (uf_69 ?x59 ?x60) }) +:assumption (forall (?x59 Int) (?x60 Int) (= (uf_69 ?x59 ?x60) (* ?x59 ?x60)) :pat { (uf_69 ?x59 ?x60) }) :assumption (forall (?x61 T3) (?x62 Int) (?x63 Int) (= (uf_70 ?x61 ?x62 ?x63) (uf_70 ?x61 ?x63 ?x62)) :pat { (uf_70 ?x61 ?x62 ?x63) }) :assumption (forall (?x64 T3) (?x65 Int) (?x66 Int) (= (uf_71 ?x64 ?x65 ?x66) (uf_71 ?x64 ?x66 ?x65)) :pat { (uf_71 ?x64 ?x65 ?x66) }) :assumption (forall (?x67 T3) (?x68 Int) (?x69 Int) (= (uf_72 ?x67 ?x68 ?x69) (uf_72 ?x67 ?x69 ?x68)) :pat { (uf_72 ?x67 ?x68 ?x69) }) @@ -359,84 +359,84 @@ :assumption (forall (?x96 T3) (?x97 Int) (= (uf_70 ?x96 ?x97 (uf_73 ?x96 ?x97)) 0) :pat { (uf_70 ?x96 ?x97 (uf_73 ?x96 ?x97)) }) :assumption (forall (?x98 T3) (?x99 Int) (= (uf_72 ?x98 ?x99 (uf_73 ?x98 ?x99)) (uf_73 ?x98 0)) :pat { (uf_72 ?x98 ?x99 (uf_73 ?x98 ?x99)) }) :assumption (forall (?x100 T3) (?x101 Int) (= (uf_74 ?x100 (uf_73 ?x100 ?x101)) uf_9) :pat { (uf_73 ?x100 ?x101) }) -:assumption (forall (?x102 T3) (?x103 Int) (?x104 Int) (implies (and (<= ?x104 uf_75) (and (<= 0 ?x104) (and (<= ?x103 uf_75) (<= 0 ?x103)))) (and (<= (uf_71 ?x102 ?x103 ?x104) uf_75) (<= 0 (uf_71 ?x102 ?x103 ?x104)))) :pat { (uf_71 ?x102 ?x103 ?x104) }) -:assumption (forall (?x105 T3) (?x106 Int) (?x107 Int) (implies (and (<= ?x107 uf_76) (and (<= 0 ?x107) (and (<= ?x106 uf_76) (<= 0 ?x106)))) (and (<= (uf_71 ?x105 ?x106 ?x107) uf_76) (<= 0 (uf_71 ?x105 ?x106 ?x107)))) :pat { (uf_71 ?x105 ?x106 ?x107) }) -:assumption (forall (?x108 T3) (?x109 Int) (?x110 Int) (implies (and (<= ?x110 uf_77) (and (<= 0 ?x110) (and (<= ?x109 uf_77) (<= 0 ?x109)))) (and (<= (uf_71 ?x108 ?x109 ?x110) uf_77) (<= 0 (uf_71 ?x108 ?x109 ?x110)))) :pat { (uf_71 ?x108 ?x109 ?x110) }) -:assumption (forall (?x111 T3) (?x112 Int) (?x113 Int) (implies (and (<= ?x113 uf_78) (and (<= 0 ?x113) (and (<= ?x112 uf_78) (<= 0 ?x112)))) (and (<= (uf_71 ?x111 ?x112 ?x113) uf_78) (<= 0 (uf_71 ?x111 ?x112 ?x113)))) :pat { (uf_71 ?x111 ?x112 ?x113) }) -:assumption (forall (?x114 T3) (?x115 Int) (?x116 Int) (implies (and (<= ?x116 uf_75) (and (<= 0 ?x116) (and (<= ?x115 uf_75) (<= 0 ?x115)))) (and (<= (uf_70 ?x114 ?x115 ?x116) uf_75) (<= 0 (uf_70 ?x114 ?x115 ?x116)))) :pat { (uf_70 ?x114 ?x115 ?x116) }) -:assumption (forall (?x117 T3) (?x118 Int) (?x119 Int) (implies (and (<= ?x119 uf_76) (and (<= 0 ?x119) (and (<= ?x118 uf_76) (<= 0 ?x118)))) (and (<= (uf_70 ?x117 ?x118 ?x119) uf_76) (<= 0 (uf_70 ?x117 ?x118 ?x119)))) :pat { (uf_70 ?x117 ?x118 ?x119) }) -:assumption (forall (?x120 T3) (?x121 Int) (?x122 Int) (implies (and (<= ?x122 uf_77) (and (<= 0 ?x122) (and (<= ?x121 uf_77) (<= 0 ?x121)))) (and (<= (uf_70 ?x120 ?x121 ?x122) uf_77) (<= 0 (uf_70 ?x120 ?x121 ?x122)))) :pat { (uf_70 ?x120 ?x121 ?x122) }) -:assumption (forall (?x123 T3) (?x124 Int) (?x125 Int) (implies (and (<= ?x125 uf_78) (and (<= 0 ?x125) (and (<= ?x124 uf_78) (<= 0 ?x124)))) (and (<= (uf_70 ?x123 ?x124 ?x125) uf_78) (<= 0 (uf_70 ?x123 ?x124 ?x125)))) :pat { (uf_70 ?x123 ?x124 ?x125) }) -:assumption (forall (?x126 T3) (?x127 Int) (?x128 Int) (implies (and (<= ?x128 uf_75) (and (<= 0 ?x128) (and (<= ?x127 uf_75) (<= 0 ?x127)))) (and (<= (uf_72 ?x126 ?x127 ?x128) uf_75) (<= 0 (uf_72 ?x126 ?x127 ?x128)))) :pat { (uf_72 ?x126 ?x127 ?x128) }) -:assumption (forall (?x129 T3) (?x130 Int) (?x131 Int) (implies (and (<= ?x131 uf_76) (and (<= 0 ?x131) (and (<= ?x130 uf_76) (<= 0 ?x130)))) (and (<= (uf_72 ?x129 ?x130 ?x131) uf_76) (<= 0 (uf_72 ?x129 ?x130 ?x131)))) :pat { (uf_72 ?x129 ?x130 ?x131) }) -:assumption (forall (?x132 T3) (?x133 Int) (?x134 Int) (implies (and (<= ?x134 uf_77) (and (<= 0 ?x134) (and (<= ?x133 uf_77) (<= 0 ?x133)))) (and (<= (uf_72 ?x132 ?x133 ?x134) uf_77) (<= 0 (uf_72 ?x132 ?x133 ?x134)))) :pat { (uf_72 ?x132 ?x133 ?x134) }) -:assumption (forall (?x135 T3) (?x136 Int) (?x137 Int) (implies (and (<= ?x137 uf_78) (and (<= 0 ?x137) (and (<= ?x136 uf_78) (<= 0 ?x136)))) (and (<= (uf_72 ?x135 ?x136 ?x137) uf_78) (<= 0 (uf_72 ?x135 ?x136 ?x137)))) :pat { (uf_72 ?x135 ?x136 ?x137) }) -:assumption (forall (?x138 T3) (?x139 Int) (?x140 Int) (?x141 Int) (implies (and (= (uf_74 ?x138 ?x140) uf_9) (and (= (uf_74 ?x138 ?x139) uf_9) (and (< ?x140 (uf_79 ?x141)) (and (< ?x139 (uf_79 ?x141)) (and (< ?x141 64) (and (<= 0 ?x141) (and (<= 0 ?x140) (<= 0 ?x139)))))))) (< (uf_72 ?x138 ?x139 ?x140) (uf_79 ?x141))) :pat { (uf_72 ?x138 ?x139 ?x140) (uf_79 ?x141) }) -:assumption (forall (?x142 T3) (?x143 Int) (?x144 Int) (implies (and (= (uf_74 ?x142 ?x144) uf_9) (and (= (uf_74 ?x142 ?x143) uf_9) (and (<= 0 ?x144) (<= 0 ?x143)))) (and (<= ?x144 (uf_72 ?x142 ?x143 ?x144)) (<= ?x143 (uf_72 ?x142 ?x143 ?x144)))) :pat { (uf_72 ?x142 ?x143 ?x144) }) -:assumption (forall (?x145 T3) (?x146 Int) (?x147 Int) (implies (and (= (uf_74 ?x145 ?x147) uf_9) (and (= (uf_74 ?x145 ?x146) uf_9) (and (<= 0 ?x147) (<= 0 ?x146)))) (and (<= (uf_72 ?x145 ?x146 ?x147) (+ ?x146 ?x147)) (<= 0 (uf_72 ?x145 ?x146 ?x147)))) :pat { (uf_72 ?x145 ?x146 ?x147) }) -:assumption (forall (?x148 T3) (?x149 Int) (?x150 Int) (implies (and (= (uf_74 ?x148 ?x150) uf_9) (and (= (uf_74 ?x148 ?x149) uf_9) (and (<= 0 ?x150) (<= 0 ?x149)))) (and (<= (uf_70 ?x148 ?x149 ?x150) ?x150) (<= (uf_70 ?x148 ?x149 ?x150) ?x149))) :pat { (uf_70 ?x148 ?x149 ?x150) }) -:assumption (forall (?x151 T3) (?x152 Int) (?x153 Int) (implies (and (= (uf_74 ?x151 ?x152) uf_9) (<= 0 ?x152)) (and (<= (uf_70 ?x151 ?x152 ?x153) ?x152) (<= 0 (uf_70 ?x151 ?x152 ?x153)))) :pat { (uf_70 ?x151 ?x152 ?x153) }) -:assumption (forall (?x154 Int) (?x155 Int) (implies (and (< ?x155 0) (<= ?x154 0)) (and (<= (uf_80 ?x154 ?x155) 0) (< ?x155 (uf_80 ?x154 ?x155)))) :pat { (uf_80 ?x154 ?x155) }) -:assumption (forall (?x156 Int) (?x157 Int) (implies (and (< 0 ?x157) (<= ?x156 0)) (and (<= (uf_80 ?x156 ?x157) 0) (< (+ 0 ?x157) (uf_80 ?x156 ?x157)))) :pat { (uf_80 ?x156 ?x157) }) -:assumption (forall (?x158 Int) (?x159 Int) (implies (and (< ?x159 0) (<= 0 ?x158)) (and (< (uf_80 ?x158 ?x159) (+ 0 ?x159)) (<= 0 (uf_80 ?x158 ?x159)))) :pat { (uf_80 ?x158 ?x159) }) -:assumption (forall (?x160 Int) (?x161 Int) (implies (and (< 0 ?x161) (<= 0 ?x160)) (and (< (uf_80 ?x160 ?x161) ?x161) (<= 0 (uf_80 ?x160 ?x161)))) :pat { (uf_80 ?x160 ?x161) }) -:assumption (forall (?x162 Int) (?x163 Int) (= (uf_80 ?x162 ?x163) (+ ?x162 (+ (uf_81 ?x162 ?x163) ?x163))) :pat { (uf_80 ?x162 ?x163) } :pat { (uf_81 ?x162 ?x163) }) +:assumption (forall (?x102 T3) (?x103 Int) (?x104 Int) (implies (and (<= 0 ?x103) (and (<= ?x103 uf_75) (and (<= 0 ?x104) (<= ?x104 uf_75)))) (and (<= 0 (uf_71 ?x102 ?x103 ?x104)) (<= (uf_71 ?x102 ?x103 ?x104) uf_75))) :pat { (uf_71 ?x102 ?x103 ?x104) }) +:assumption (forall (?x105 T3) (?x106 Int) (?x107 Int) (implies (and (<= 0 ?x106) (and (<= ?x106 uf_76) (and (<= 0 ?x107) (<= ?x107 uf_76)))) (and (<= 0 (uf_71 ?x105 ?x106 ?x107)) (<= (uf_71 ?x105 ?x106 ?x107) uf_76))) :pat { (uf_71 ?x105 ?x106 ?x107) }) +:assumption (forall (?x108 T3) (?x109 Int) (?x110 Int) (implies (and (<= 0 ?x109) (and (<= ?x109 uf_77) (and (<= 0 ?x110) (<= ?x110 uf_77)))) (and (<= 0 (uf_71 ?x108 ?x109 ?x110)) (<= (uf_71 ?x108 ?x109 ?x110) uf_77))) :pat { (uf_71 ?x108 ?x109 ?x110) }) +:assumption (forall (?x111 T3) (?x112 Int) (?x113 Int) (implies (and (<= 0 ?x112) (and (<= ?x112 uf_78) (and (<= 0 ?x113) (<= ?x113 uf_78)))) (and (<= 0 (uf_71 ?x111 ?x112 ?x113)) (<= (uf_71 ?x111 ?x112 ?x113) uf_78))) :pat { (uf_71 ?x111 ?x112 ?x113) }) +:assumption (forall (?x114 T3) (?x115 Int) (?x116 Int) (implies (and (<= 0 ?x115) (and (<= ?x115 uf_75) (and (<= 0 ?x116) (<= ?x116 uf_75)))) (and (<= 0 (uf_70 ?x114 ?x115 ?x116)) (<= (uf_70 ?x114 ?x115 ?x116) uf_75))) :pat { (uf_70 ?x114 ?x115 ?x116) }) +:assumption (forall (?x117 T3) (?x118 Int) (?x119 Int) (implies (and (<= 0 ?x118) (and (<= ?x118 uf_76) (and (<= 0 ?x119) (<= ?x119 uf_76)))) (and (<= 0 (uf_70 ?x117 ?x118 ?x119)) (<= (uf_70 ?x117 ?x118 ?x119) uf_76))) :pat { (uf_70 ?x117 ?x118 ?x119) }) +:assumption (forall (?x120 T3) (?x121 Int) (?x122 Int) (implies (and (<= 0 ?x121) (and (<= ?x121 uf_77) (and (<= 0 ?x122) (<= ?x122 uf_77)))) (and (<= 0 (uf_70 ?x120 ?x121 ?x122)) (<= (uf_70 ?x120 ?x121 ?x122) uf_77))) :pat { (uf_70 ?x120 ?x121 ?x122) }) +:assumption (forall (?x123 T3) (?x124 Int) (?x125 Int) (implies (and (<= 0 ?x124) (and (<= ?x124 uf_78) (and (<= 0 ?x125) (<= ?x125 uf_78)))) (and (<= 0 (uf_70 ?x123 ?x124 ?x125)) (<= (uf_70 ?x123 ?x124 ?x125) uf_78))) :pat { (uf_70 ?x123 ?x124 ?x125) }) +:assumption (forall (?x126 T3) (?x127 Int) (?x128 Int) (implies (and (<= 0 ?x127) (and (<= ?x127 uf_75) (and (<= 0 ?x128) (<= ?x128 uf_75)))) (and (<= 0 (uf_72 ?x126 ?x127 ?x128)) (<= (uf_72 ?x126 ?x127 ?x128) uf_75))) :pat { (uf_72 ?x126 ?x127 ?x128) }) +:assumption (forall (?x129 T3) (?x130 Int) (?x131 Int) (implies (and (<= 0 ?x130) (and (<= ?x130 uf_76) (and (<= 0 ?x131) (<= ?x131 uf_76)))) (and (<= 0 (uf_72 ?x129 ?x130 ?x131)) (<= (uf_72 ?x129 ?x130 ?x131) uf_76))) :pat { (uf_72 ?x129 ?x130 ?x131) }) +:assumption (forall (?x132 T3) (?x133 Int) (?x134 Int) (implies (and (<= 0 ?x133) (and (<= ?x133 uf_77) (and (<= 0 ?x134) (<= ?x134 uf_77)))) (and (<= 0 (uf_72 ?x132 ?x133 ?x134)) (<= (uf_72 ?x132 ?x133 ?x134) uf_77))) :pat { (uf_72 ?x132 ?x133 ?x134) }) +:assumption (forall (?x135 T3) (?x136 Int) (?x137 Int) (implies (and (<= 0 ?x136) (and (<= ?x136 uf_78) (and (<= 0 ?x137) (<= ?x137 uf_78)))) (and (<= 0 (uf_72 ?x135 ?x136 ?x137)) (<= (uf_72 ?x135 ?x136 ?x137) uf_78))) :pat { (uf_72 ?x135 ?x136 ?x137) }) +:assumption (forall (?x138 T3) (?x139 Int) (?x140 Int) (?x141 Int) (implies (and (<= 0 ?x139) (and (<= 0 ?x140) (and (<= 0 ?x141) (and (< ?x141 64) (and (< ?x139 (uf_79 ?x141)) (and (< ?x140 (uf_79 ?x141)) (and (= (uf_74 ?x138 ?x139) uf_9) (= (uf_74 ?x138 ?x140) uf_9)))))))) (< (uf_72 ?x138 ?x139 ?x140) (uf_79 ?x141))) :pat { (uf_72 ?x138 ?x139 ?x140) (uf_79 ?x141) }) +:assumption (forall (?x142 T3) (?x143 Int) (?x144 Int) (implies (and (<= 0 ?x143) (and (<= 0 ?x144) (and (= (uf_74 ?x142 ?x143) uf_9) (= (uf_74 ?x142 ?x144) uf_9)))) (and (<= ?x143 (uf_72 ?x142 ?x143 ?x144)) (<= ?x144 (uf_72 ?x142 ?x143 ?x144)))) :pat { (uf_72 ?x142 ?x143 ?x144) }) +:assumption (forall (?x145 T3) (?x146 Int) (?x147 Int) (implies (and (<= 0 ?x146) (and (<= 0 ?x147) (and (= (uf_74 ?x145 ?x146) uf_9) (= (uf_74 ?x145 ?x147) uf_9)))) (and (<= 0 (uf_72 ?x145 ?x146 ?x147)) (<= (uf_72 ?x145 ?x146 ?x147) (+ ?x146 ?x147)))) :pat { (uf_72 ?x145 ?x146 ?x147) }) +:assumption (forall (?x148 T3) (?x149 Int) (?x150 Int) (implies (and (<= 0 ?x149) (and (<= 0 ?x150) (and (= (uf_74 ?x148 ?x149) uf_9) (= (uf_74 ?x148 ?x150) uf_9)))) (and (<= (uf_70 ?x148 ?x149 ?x150) ?x149) (<= (uf_70 ?x148 ?x149 ?x150) ?x150))) :pat { (uf_70 ?x148 ?x149 ?x150) }) +:assumption (forall (?x151 T3) (?x152 Int) (?x153 Int) (implies (and (<= 0 ?x152) (= (uf_74 ?x151 ?x152) uf_9)) (and (<= 0 (uf_70 ?x151 ?x152 ?x153)) (<= (uf_70 ?x151 ?x152 ?x153) ?x152))) :pat { (uf_70 ?x151 ?x152 ?x153) }) +:assumption (forall (?x154 Int) (?x155 Int) (implies (and (<= ?x154 0) (< ?x155 0)) (and (< ?x155 (uf_80 ?x154 ?x155)) (<= (uf_80 ?x154 ?x155) 0))) :pat { (uf_80 ?x154 ?x155) }) +:assumption (forall (?x156 Int) (?x157 Int) (implies (and (<= ?x156 0) (< 0 ?x157)) (and (< (- 0 ?x157) (uf_80 ?x156 ?x157)) (<= (uf_80 ?x156 ?x157) 0))) :pat { (uf_80 ?x156 ?x157) }) +:assumption (forall (?x158 Int) (?x159 Int) (implies (and (<= 0 ?x158) (< ?x159 0)) (and (<= 0 (uf_80 ?x158 ?x159)) (< (uf_80 ?x158 ?x159) (- 0 ?x159)))) :pat { (uf_80 ?x158 ?x159) }) +:assumption (forall (?x160 Int) (?x161 Int) (implies (and (<= 0 ?x160) (< 0 ?x161)) (and (<= 0 (uf_80 ?x160 ?x161)) (< (uf_80 ?x160 ?x161) ?x161))) :pat { (uf_80 ?x160 ?x161) }) +:assumption (forall (?x162 Int) (?x163 Int) (= (uf_80 ?x162 ?x163) (- ?x162 (* (uf_81 ?x162 ?x163) ?x163))) :pat { (uf_80 ?x162 ?x163) } :pat { (uf_81 ?x162 ?x163) }) :assumption (forall (?x164 Int) (implies (not (= ?x164 0)) (= (uf_81 ?x164 ?x164) 1)) :pat { (uf_81 ?x164 ?x164) }) -:assumption (forall (?x165 Int) (?x166 Int) (implies (and (< 0 ?x166) (< 0 ?x165)) (and (<= (+ (uf_81 ?x165 ?x166) ?x166) ?x165) (< (+ ?x165 ?x166) (+ (uf_81 ?x165 ?x166) ?x166)))) :pat { (uf_81 ?x165 ?x166) }) -:assumption (forall (?x167 Int) (?x168 Int) (implies (and (< 0 ?x168) (<= 0 ?x167)) (<= (uf_81 ?x167 ?x168) ?x167)) :pat { (uf_81 ?x167 ?x168) }) -:assumption (forall (?x169 T3) (?x170 Int) (?x171 Int) (?x172 Int) (implies (and (<= 0 ?x170) (and (= (uf_74 ?x169 (+ (uf_79 ?x171) 1)) uf_9) (= (uf_74 ?x169 ?x170) uf_9))) (= (uf_80 ?x170 (uf_79 ?x171)) (uf_70 ?x169 ?x170 (+ (uf_79 ?x171) 1)))) :pat { (uf_80 ?x170 (uf_79 ?x171)) (uf_70 ?x169 ?x170 ?x172) }) -:assumption (forall (?x173 Int) (implies (and (<= ?x173 uf_85) (<= uf_86 ?x173)) (= (uf_82 uf_83 (uf_82 uf_84 ?x173)) ?x173)) :pat { (uf_82 uf_83 (uf_82 uf_84 ?x173)) }) -:assumption (forall (?x174 Int) (implies (and (<= ?x174 uf_88) (<= uf_89 ?x174)) (= (uf_82 uf_87 (uf_82 uf_4 ?x174)) ?x174)) :pat { (uf_82 uf_87 (uf_82 uf_4 ?x174)) }) -:assumption (forall (?x175 Int) (implies (and (<= ?x175 uf_92) (<= uf_93 ?x175)) (= (uf_82 uf_90 (uf_82 uf_91 ?x175)) ?x175)) :pat { (uf_82 uf_90 (uf_82 uf_91 ?x175)) }) -:assumption (forall (?x176 Int) (implies (and (<= ?x176 uf_95) (<= uf_96 ?x176)) (= (uf_82 uf_94 (uf_82 uf_7 ?x176)) ?x176)) :pat { (uf_82 uf_94 (uf_82 uf_7 ?x176)) }) -:assumption (forall (?x177 Int) (implies (and (<= ?x177 uf_75) (<= 0 ?x177)) (= (uf_82 uf_84 (uf_82 uf_83 ?x177)) ?x177)) :pat { (uf_82 uf_84 (uf_82 uf_83 ?x177)) }) -:assumption (forall (?x178 Int) (implies (and (<= ?x178 uf_76) (<= 0 ?x178)) (= (uf_82 uf_4 (uf_82 uf_87 ?x178)) ?x178)) :pat { (uf_82 uf_4 (uf_82 uf_87 ?x178)) }) -:assumption (forall (?x179 Int) (implies (and (<= ?x179 uf_77) (<= 0 ?x179)) (= (uf_82 uf_91 (uf_82 uf_90 ?x179)) ?x179)) :pat { (uf_82 uf_91 (uf_82 uf_90 ?x179)) }) -:assumption (forall (?x180 Int) (implies (and (<= ?x180 uf_78) (<= 0 ?x180)) (= (uf_82 uf_7 (uf_82 uf_94 ?x180)) ?x180)) :pat { (uf_82 uf_7 (uf_82 uf_94 ?x180)) }) +:assumption (forall (?x165 Int) (?x166 Int) (implies (and (< 0 ?x165) (< 0 ?x166)) (and (< (- ?x165 ?x166) (* (uf_81 ?x165 ?x166) ?x166)) (<= (* (uf_81 ?x165 ?x166) ?x166) ?x165))) :pat { (uf_81 ?x165 ?x166) }) +:assumption (forall (?x167 Int) (?x168 Int) (implies (and (<= 0 ?x167) (< 0 ?x168)) (<= (uf_81 ?x167 ?x168) ?x167)) :pat { (uf_81 ?x167 ?x168) }) +:assumption (forall (?x169 T3) (?x170 Int) (?x171 Int) (?x172 Int) (implies (and (= (uf_74 ?x169 ?x170) uf_9) (and (= (uf_74 ?x169 (- (uf_79 ?x171) 1)) uf_9) (<= 0 ?x170))) (= (uf_80 ?x170 (uf_79 ?x171)) (uf_70 ?x169 ?x170 (- (uf_79 ?x171) 1)))) :pat { (uf_80 ?x170 (uf_79 ?x171)) (uf_70 ?x169 ?x170 ?x172) }) +:assumption (forall (?x173 Int) (implies (and (<= uf_85 ?x173) (<= ?x173 uf_86)) (= (uf_82 uf_83 (uf_82 uf_84 ?x173)) ?x173)) :pat { (uf_82 uf_83 (uf_82 uf_84 ?x173)) }) +:assumption (forall (?x174 Int) (implies (and (<= uf_88 ?x174) (<= ?x174 uf_89)) (= (uf_82 uf_87 (uf_82 uf_4 ?x174)) ?x174)) :pat { (uf_82 uf_87 (uf_82 uf_4 ?x174)) }) +:assumption (forall (?x175 Int) (implies (and (<= uf_92 ?x175) (<= ?x175 uf_93)) (= (uf_82 uf_90 (uf_82 uf_91 ?x175)) ?x175)) :pat { (uf_82 uf_90 (uf_82 uf_91 ?x175)) }) +:assumption (forall (?x176 Int) (implies (and (<= uf_95 ?x176) (<= ?x176 uf_96)) (= (uf_82 uf_94 (uf_82 uf_7 ?x176)) ?x176)) :pat { (uf_82 uf_94 (uf_82 uf_7 ?x176)) }) +:assumption (forall (?x177 Int) (implies (and (<= 0 ?x177) (<= ?x177 uf_75)) (= (uf_82 uf_84 (uf_82 uf_83 ?x177)) ?x177)) :pat { (uf_82 uf_84 (uf_82 uf_83 ?x177)) }) +:assumption (forall (?x178 Int) (implies (and (<= 0 ?x178) (<= ?x178 uf_76)) (= (uf_82 uf_4 (uf_82 uf_87 ?x178)) ?x178)) :pat { (uf_82 uf_4 (uf_82 uf_87 ?x178)) }) +:assumption (forall (?x179 Int) (implies (and (<= 0 ?x179) (<= ?x179 uf_77)) (= (uf_82 uf_91 (uf_82 uf_90 ?x179)) ?x179)) :pat { (uf_82 uf_91 (uf_82 uf_90 ?x179)) }) +:assumption (forall (?x180 Int) (implies (and (<= 0 ?x180) (<= ?x180 uf_78)) (= (uf_82 uf_7 (uf_82 uf_94 ?x180)) ?x180)) :pat { (uf_82 uf_7 (uf_82 uf_94 ?x180)) }) :assumption (forall (?x181 T3) (?x182 Int) (= (uf_74 ?x181 (uf_82 ?x181 ?x182)) uf_9) :pat { (uf_82 ?x181 ?x182) }) :assumption (forall (?x183 T3) (?x184 Int) (implies (= (uf_74 ?x183 ?x184) uf_9) (= (uf_82 ?x183 ?x184) ?x184)) :pat { (uf_82 ?x183 ?x184) }) -:assumption (forall (?x185 Int) (iff (= (uf_74 uf_84 ?x185) uf_9) (and (<= ?x185 uf_75) (<= 0 ?x185))) :pat { (uf_74 uf_84 ?x185) }) -:assumption (forall (?x186 Int) (iff (= (uf_74 uf_4 ?x186) uf_9) (and (<= ?x186 uf_76) (<= 0 ?x186))) :pat { (uf_74 uf_4 ?x186) }) -:assumption (forall (?x187 Int) (iff (= (uf_74 uf_91 ?x187) uf_9) (and (<= ?x187 uf_77) (<= 0 ?x187))) :pat { (uf_74 uf_91 ?x187) }) -:assumption (forall (?x188 Int) (iff (= (uf_74 uf_7 ?x188) uf_9) (and (<= ?x188 uf_78) (<= 0 ?x188))) :pat { (uf_74 uf_7 ?x188) }) -:assumption (forall (?x189 Int) (iff (= (uf_74 uf_83 ?x189) uf_9) (and (<= ?x189 uf_85) (<= uf_86 ?x189))) :pat { (uf_74 uf_83 ?x189) }) -:assumption (forall (?x190 Int) (iff (= (uf_74 uf_87 ?x190) uf_9) (and (<= ?x190 uf_88) (<= uf_89 ?x190))) :pat { (uf_74 uf_87 ?x190) }) -:assumption (forall (?x191 Int) (iff (= (uf_74 uf_90 ?x191) uf_9) (and (<= ?x191 uf_92) (<= uf_93 ?x191))) :pat { (uf_74 uf_90 ?x191) }) -:assumption (forall (?x192 Int) (iff (= (uf_74 uf_94 ?x192) uf_9) (and (<= ?x192 uf_95) (<= uf_96 ?x192))) :pat { (uf_74 uf_94 ?x192) }) -:assumption (forall (?x193 Int) (?x194 Int) (?x195 Int) (?x196 Int) (implies (and (<= (uf_79 (+ (+ ?x194 ?x193) 1)) (uf_80 (uf_81 ?x195 (uf_79 ?x193)) (uf_79 (+ ?x194 ?x193)))) (and (<= 0 ?x195) (and (<= ?x194 ?x196) (and (< ?x193 ?x194) (<= 0 ?x193))))) (= (uf_97 ?x195 ?x196 ?x193 ?x194) (+ (uf_79 (+ (+ ?x194 ?x193) 1)) (uf_80 (uf_81 ?x195 (uf_79 ?x193)) (uf_79 (+ ?x194 ?x193)))))) :pat { (uf_97 ?x195 ?x196 ?x193 ?x194) }) -:assumption (forall (?x197 Int) (?x198 Int) (?x199 Int) (?x200 Int) (implies (and (< (uf_80 (uf_81 ?x199 (uf_79 ?x197)) (uf_79 (+ ?x198 ?x197))) (uf_79 (+ (+ ?x198 ?x197) 1))) (and (<= 0 ?x199) (and (<= ?x198 ?x200) (and (< ?x197 ?x198) (<= 0 ?x197))))) (= (uf_97 ?x199 ?x200 ?x197 ?x198) (uf_80 (uf_81 ?x199 (uf_79 ?x197)) (uf_79 (+ ?x198 ?x197))))) :pat { (uf_97 ?x199 ?x200 ?x197 ?x198) }) -:assumption (forall (?x201 Int) (?x202 Int) (?x203 Int) (?x204 Int) (implies (and (<= 0 ?x203) (and (<= ?x202 ?x204) (and (< ?x201 ?x202) (<= 0 ?x201)))) (= (uf_98 ?x203 ?x204 ?x201 ?x202) (uf_80 (uf_81 ?x203 (uf_79 ?x201)) (uf_79 (+ ?x202 ?x201))))) :pat { (uf_98 ?x203 ?x204 ?x201 ?x202) }) -:assumption (forall (?x205 Int) (?x206 Int) (?x207 Int) (implies (and (<= ?x206 ?x207) (and (< ?x205 ?x206) (<= 0 ?x205))) (= (uf_98 0 ?x207 ?x205 ?x206) 0)) :pat { (uf_98 0 ?x207 ?x205 ?x206) }) -:assumption (forall (?x208 Int) (?x209 Int) (?x210 Int) (implies (and (<= ?x209 ?x210) (and (< ?x208 ?x209) (<= 0 ?x208))) (= (uf_97 0 ?x210 ?x208 ?x209) 0)) :pat { (uf_97 0 ?x210 ?x208 ?x209) }) -:assumption (forall (?x211 Int) (?x212 Int) (?x213 Int) (?x214 Int) (?x215 Int) (?x216 Int) (?x217 Int) (implies (and (<= ?x212 ?x215) (and (< ?x211 ?x212) (<= 0 ?x211))) (implies (and (<= ?x217 ?x215) (and (< ?x216 ?x217) (<= 0 ?x216))) (implies (or (<= ?x212 ?x216) (<= ?x217 ?x211)) (= (uf_98 (uf_99 ?x214 ?x215 ?x211 ?x212 ?x213) ?x215 ?x216 ?x217) (uf_98 ?x214 ?x215 ?x216 ?x217))))) :pat { (uf_98 (uf_99 ?x214 ?x215 ?x211 ?x212 ?x213) ?x215 ?x216 ?x217) }) -:assumption (forall (?x218 Int) (?x219 Int) (?x220 Int) (?x221 Int) (?x222 Int) (?x223 Int) (?x224 Int) (implies (and (<= ?x219 ?x222) (and (< ?x218 ?x219) (<= 0 ?x218))) (implies (and (<= ?x224 ?x222) (and (< ?x223 ?x224) (<= 0 ?x223))) (implies (or (<= ?x219 ?x223) (<= ?x224 ?x218)) (= (uf_97 (uf_99 ?x221 ?x222 ?x218 ?x219 ?x220) ?x222 ?x223 ?x224) (uf_97 ?x221 ?x222 ?x223 ?x224))))) :pat { (uf_97 (uf_99 ?x221 ?x222 ?x218 ?x219 ?x220) ?x222 ?x223 ?x224) }) -:assumption (forall (?x225 Int) (?x226 Int) (?x227 Int) (?x228 Int) (implies (and (<= ?x226 ?x228) (and (< ?x225 ?x226) (<= 0 ?x225))) (and (<= (uf_98 ?x227 ?x228 ?x225 ?x226) (+ (uf_79 (+ ?x226 ?x225)) 1)) (<= 0 (uf_98 ?x227 ?x228 ?x225 ?x226)))) :pat { (uf_98 ?x227 ?x228 ?x225 ?x226) }) -:assumption (forall (?x229 Int) (?x230 Int) (?x231 Int) (?x232 Int) (implies (and (<= ?x230 ?x232) (and (< ?x229 ?x230) (<= 0 ?x229))) (and (<= (uf_97 ?x231 ?x232 ?x229 ?x230) (+ (uf_79 (+ (+ ?x230 ?x229) 1)) 1)) (<= (+ 0 (uf_79 (+ (+ ?x230 ?x229) 1))) (uf_97 ?x231 ?x232 ?x229 ?x230)))) :pat { (uf_97 ?x231 ?x232 ?x229 ?x230) }) -:assumption (forall (?x233 Int) (?x234 Int) (?x235 Int) (?x236 Int) (?x237 Int) (implies (and (<= ?x234 ?x237) (and (< ?x233 ?x234) (<= 0 ?x233))) (implies (and (< ?x235 (uf_79 (+ ?x234 ?x233))) (<= 0 ?x235)) (= (uf_98 (uf_99 ?x236 ?x237 ?x233 ?x234 ?x235) ?x237 ?x233 ?x234) ?x235))) :pat { (uf_98 (uf_99 ?x236 ?x237 ?x233 ?x234 ?x235) ?x237 ?x233 ?x234) }) -:assumption (forall (?x238 Int) (?x239 Int) (?x240 Int) (?x241 Int) (?x242 Int) (implies (and (<= ?x239 ?x242) (and (< ?x238 ?x239) (<= 0 ?x238))) (implies (and (< ?x240 (uf_79 (+ (+ ?x239 ?x238) 1))) (<= (+ 0 (uf_79 (+ (+ ?x239 ?x238) 1))) ?x240)) (= (uf_97 (uf_99 ?x241 ?x242 ?x238 ?x239 ?x240) ?x242 ?x238 ?x239) ?x240))) :pat { (uf_97 (uf_99 ?x241 ?x242 ?x238 ?x239 ?x240) ?x242 ?x238 ?x239) }) -:assumption (forall (?x243 Int) (?x244 Int) (?x245 Int) (implies (and (<= ?x244 ?x245) (and (< ?x243 ?x244) (<= 0 ?x243))) (= (uf_99 0 ?x245 ?x243 ?x244 0) 0)) :pat { (uf_99 0 ?x245 ?x243 ?x244 0) }) -:assumption (forall (?x246 Int) (?x247 Int) (?x248 Int) (?x249 Int) (?x250 Int) (implies (and (<= ?x248 ?x249) (and (< ?x247 ?x248) (<= 0 ?x247))) (implies (and (< ?x250 (uf_79 (+ ?x248 ?x247))) (<= 0 ?x250)) (and (< (uf_99 ?x246 ?x249 ?x247 ?x248 ?x250) (uf_79 ?x249)) (<= 0 (uf_99 ?x246 ?x249 ?x247 ?x248 ?x250))))) :pat { (uf_99 ?x246 ?x249 ?x247 ?x248 ?x250) }) +:assumption (forall (?x185 Int) (iff (= (uf_74 uf_84 ?x185) uf_9) (and (<= 0 ?x185) (<= ?x185 uf_75))) :pat { (uf_74 uf_84 ?x185) }) +:assumption (forall (?x186 Int) (iff (= (uf_74 uf_4 ?x186) uf_9) (and (<= 0 ?x186) (<= ?x186 uf_76))) :pat { (uf_74 uf_4 ?x186) }) +:assumption (forall (?x187 Int) (iff (= (uf_74 uf_91 ?x187) uf_9) (and (<= 0 ?x187) (<= ?x187 uf_77))) :pat { (uf_74 uf_91 ?x187) }) +:assumption (forall (?x188 Int) (iff (= (uf_74 uf_7 ?x188) uf_9) (and (<= 0 ?x188) (<= ?x188 uf_78))) :pat { (uf_74 uf_7 ?x188) }) +:assumption (forall (?x189 Int) (iff (= (uf_74 uf_83 ?x189) uf_9) (and (<= uf_85 ?x189) (<= ?x189 uf_86))) :pat { (uf_74 uf_83 ?x189) }) +:assumption (forall (?x190 Int) (iff (= (uf_74 uf_87 ?x190) uf_9) (and (<= uf_88 ?x190) (<= ?x190 uf_89))) :pat { (uf_74 uf_87 ?x190) }) +:assumption (forall (?x191 Int) (iff (= (uf_74 uf_90 ?x191) uf_9) (and (<= uf_92 ?x191) (<= ?x191 uf_93))) :pat { (uf_74 uf_90 ?x191) }) +:assumption (forall (?x192 Int) (iff (= (uf_74 uf_94 ?x192) uf_9) (and (<= uf_95 ?x192) (<= ?x192 uf_96))) :pat { (uf_74 uf_94 ?x192) }) +:assumption (forall (?x193 Int) (?x194 Int) (?x195 Int) (?x196 Int) (implies (and (<= 0 ?x193) (and (< ?x193 ?x194) (and (<= ?x194 ?x196) (and (<= 0 ?x195) (<= (uf_79 (- (- ?x194 ?x193) 1)) (uf_80 (uf_81 ?x195 (uf_79 ?x193)) (uf_79 (- ?x194 ?x193)))))))) (= (uf_97 ?x195 ?x196 ?x193 ?x194) (- (uf_79 (- (- ?x194 ?x193) 1)) (uf_80 (uf_81 ?x195 (uf_79 ?x193)) (uf_79 (- ?x194 ?x193)))))) :pat { (uf_97 ?x195 ?x196 ?x193 ?x194) }) +:assumption (forall (?x197 Int) (?x198 Int) (?x199 Int) (?x200 Int) (implies (and (<= 0 ?x197) (and (< ?x197 ?x198) (and (<= ?x198 ?x200) (and (<= 0 ?x199) (< (uf_80 (uf_81 ?x199 (uf_79 ?x197)) (uf_79 (- ?x198 ?x197))) (uf_79 (- (- ?x198 ?x197) 1))))))) (= (uf_97 ?x199 ?x200 ?x197 ?x198) (uf_80 (uf_81 ?x199 (uf_79 ?x197)) (uf_79 (- ?x198 ?x197))))) :pat { (uf_97 ?x199 ?x200 ?x197 ?x198) }) +:assumption (forall (?x201 Int) (?x202 Int) (?x203 Int) (?x204 Int) (implies (and (<= 0 ?x201) (and (< ?x201 ?x202) (and (<= ?x202 ?x204) (<= 0 ?x203)))) (= (uf_98 ?x203 ?x204 ?x201 ?x202) (uf_80 (uf_81 ?x203 (uf_79 ?x201)) (uf_79 (- ?x202 ?x201))))) :pat { (uf_98 ?x203 ?x204 ?x201 ?x202) }) +:assumption (forall (?x205 Int) (?x206 Int) (?x207 Int) (implies (and (<= 0 ?x205) (and (< ?x205 ?x206) (<= ?x206 ?x207))) (= (uf_98 0 ?x207 ?x205 ?x206) 0)) :pat { (uf_98 0 ?x207 ?x205 ?x206) }) +:assumption (forall (?x208 Int) (?x209 Int) (?x210 Int) (implies (and (<= 0 ?x208) (and (< ?x208 ?x209) (<= ?x209 ?x210))) (= (uf_97 0 ?x210 ?x208 ?x209) 0)) :pat { (uf_97 0 ?x210 ?x208 ?x209) }) +:assumption (forall (?x211 Int) (?x212 Int) (?x213 Int) (?x214 Int) (?x215 Int) (?x216 Int) (?x217 Int) (implies (and (<= 0 ?x211) (and (< ?x211 ?x212) (<= ?x212 ?x215))) (implies (and (<= 0 ?x216) (and (< ?x216 ?x217) (<= ?x217 ?x215))) (implies (or (<= ?x217 ?x211) (<= ?x212 ?x216)) (= (uf_98 (uf_99 ?x214 ?x215 ?x211 ?x212 ?x213) ?x215 ?x216 ?x217) (uf_98 ?x214 ?x215 ?x216 ?x217))))) :pat { (uf_98 (uf_99 ?x214 ?x215 ?x211 ?x212 ?x213) ?x215 ?x216 ?x217) }) +:assumption (forall (?x218 Int) (?x219 Int) (?x220 Int) (?x221 Int) (?x222 Int) (?x223 Int) (?x224 Int) (implies (and (<= 0 ?x218) (and (< ?x218 ?x219) (<= ?x219 ?x222))) (implies (and (<= 0 ?x223) (and (< ?x223 ?x224) (<= ?x224 ?x222))) (implies (or (<= ?x224 ?x218) (<= ?x219 ?x223)) (= (uf_97 (uf_99 ?x221 ?x222 ?x218 ?x219 ?x220) ?x222 ?x223 ?x224) (uf_97 ?x221 ?x222 ?x223 ?x224))))) :pat { (uf_97 (uf_99 ?x221 ?x222 ?x218 ?x219 ?x220) ?x222 ?x223 ?x224) }) +:assumption (forall (?x225 Int) (?x226 Int) (?x227 Int) (?x228 Int) (implies (and (<= 0 ?x225) (and (< ?x225 ?x226) (<= ?x226 ?x228))) (and (<= 0 (uf_98 ?x227 ?x228 ?x225 ?x226)) (<= (uf_98 ?x227 ?x228 ?x225 ?x226) (- (uf_79 (- ?x226 ?x225)) 1)))) :pat { (uf_98 ?x227 ?x228 ?x225 ?x226) }) +:assumption (forall (?x229 Int) (?x230 Int) (?x231 Int) (?x232 Int) (implies (and (<= 0 ?x229) (and (< ?x229 ?x230) (<= ?x230 ?x232))) (and (<= (- 0 (uf_79 (- (- ?x230 ?x229) 1))) (uf_97 ?x231 ?x232 ?x229 ?x230)) (<= (uf_97 ?x231 ?x232 ?x229 ?x230) (- (uf_79 (- (- ?x230 ?x229) 1)) 1)))) :pat { (uf_97 ?x231 ?x232 ?x229 ?x230) }) +:assumption (forall (?x233 Int) (?x234 Int) (?x235 Int) (?x236 Int) (?x237 Int) (implies (and (<= 0 ?x233) (and (< ?x233 ?x234) (<= ?x234 ?x237))) (implies (and (<= 0 ?x235) (< ?x235 (uf_79 (- ?x234 ?x233)))) (= (uf_98 (uf_99 ?x236 ?x237 ?x233 ?x234 ?x235) ?x237 ?x233 ?x234) ?x235))) :pat { (uf_98 (uf_99 ?x236 ?x237 ?x233 ?x234 ?x235) ?x237 ?x233 ?x234) }) +:assumption (forall (?x238 Int) (?x239 Int) (?x240 Int) (?x241 Int) (?x242 Int) (implies (and (<= 0 ?x238) (and (< ?x238 ?x239) (<= ?x239 ?x242))) (implies (and (<= (- 0 (uf_79 (- (- ?x239 ?x238) 1))) ?x240) (< ?x240 (uf_79 (- (- ?x239 ?x238) 1)))) (= (uf_97 (uf_99 ?x241 ?x242 ?x238 ?x239 ?x240) ?x242 ?x238 ?x239) ?x240))) :pat { (uf_97 (uf_99 ?x241 ?x242 ?x238 ?x239 ?x240) ?x242 ?x238 ?x239) }) +:assumption (forall (?x243 Int) (?x244 Int) (?x245 Int) (implies (and (<= 0 ?x243) (and (< ?x243 ?x244) (<= ?x244 ?x245))) (= (uf_99 0 ?x245 ?x243 ?x244 0) 0)) :pat { (uf_99 0 ?x245 ?x243 ?x244 0) }) +:assumption (forall (?x246 Int) (?x247 Int) (?x248 Int) (?x249 Int) (?x250 Int) (implies (and (<= 0 ?x247) (and (< ?x247 ?x248) (<= ?x248 ?x249))) (implies (and (<= 0 ?x250) (< ?x250 (uf_79 (- ?x248 ?x247)))) (and (<= 0 (uf_99 ?x246 ?x249 ?x247 ?x248 ?x250)) (< (uf_99 ?x246 ?x249 ?x247 ?x248 ?x250) (uf_79 ?x249))))) :pat { (uf_99 ?x246 ?x249 ?x247 ?x248 ?x250) }) :assumption (forall (?x251 Int) (?x252 Int) (= (uf_100 ?x251 ?x252) (uf_81 ?x251 (uf_79 ?x252))) :pat { (uf_100 ?x251 ?x252) }) -:assumption (forall (?x253 T3) (?x254 Int) (?x255 Int) (= (uf_101 ?x253 ?x254 ?x255) (uf_82 ?x253 (+ ?x254 (uf_79 ?x255)))) :pat { (uf_101 ?x253 ?x254 ?x255) }) +:assumption (forall (?x253 T3) (?x254 Int) (?x255 Int) (= (uf_101 ?x253 ?x254 ?x255) (uf_82 ?x253 (* ?x254 (uf_79 ?x255)))) :pat { (uf_101 ?x253 ?x254 ?x255) }) :assumption (forall (?x256 T3) (?x257 Int) (?x258 Int) (= (uf_102 ?x256 ?x257 ?x258) (uf_82 ?x256 (uf_80 ?x257 ?x258))) :pat { (uf_102 ?x256 ?x257 ?x258) }) :assumption (forall (?x259 T3) (?x260 Int) (?x261 Int) (= (uf_103 ?x259 ?x260 ?x261) (uf_82 ?x259 (uf_81 ?x260 ?x261))) :pat { (uf_103 ?x259 ?x260 ?x261) }) -:assumption (forall (?x262 T3) (?x263 Int) (?x264 Int) (= (uf_104 ?x262 ?x263 ?x264) (uf_82 ?x262 (+ ?x263 ?x264))) :pat { (uf_104 ?x262 ?x263 ?x264) }) -:assumption (forall (?x265 T3) (?x266 Int) (?x267 Int) (= (uf_105 ?x265 ?x266 ?x267) (uf_82 ?x265 (+ ?x266 ?x267))) :pat { (uf_105 ?x265 ?x266 ?x267) }) +:assumption (forall (?x262 T3) (?x263 Int) (?x264 Int) (= (uf_104 ?x262 ?x263 ?x264) (uf_82 ?x262 (* ?x263 ?x264))) :pat { (uf_104 ?x262 ?x263 ?x264) }) +:assumption (forall (?x265 T3) (?x266 Int) (?x267 Int) (= (uf_105 ?x265 ?x266 ?x267) (uf_82 ?x265 (- ?x266 ?x267))) :pat { (uf_105 ?x265 ?x266 ?x267) }) :assumption (forall (?x268 T3) (?x269 Int) (?x270 Int) (= (uf_106 ?x268 ?x269 ?x270) (uf_82 ?x268 (+ ?x269 ?x270))) :pat { (uf_106 ?x268 ?x269 ?x270) }) -:assumption (and (= (uf_79 63) 9223372036854775808) (and (= (uf_79 62) 4611686018427387904) (and (= (uf_79 61) 2305843009213693952) (and (= (uf_79 60) 1152921504606846976) (and (= (uf_79 59) 576460752303423488) (and (= (uf_79 58) 288230376151711744) (and (= (uf_79 57) 144115188075855872) (and (= (uf_79 56) 72057594037927936) (and (= (uf_79 55) 36028797018963968) (and (= (uf_79 54) 18014398509481984) (and (= (uf_79 53) 9007199254740992) (and (= (uf_79 52) 4503599627370496) (and (= (uf_79 51) 2251799813685248) (and (= (uf_79 50) 1125899906842624) (and (= (uf_79 49) 562949953421312) (and (= (uf_79 48) 281474976710656) (and (= (uf_79 47) 140737488355328) (and (= (uf_79 46) 70368744177664) (and (= (uf_79 45) 35184372088832) (and (= (uf_79 44) 17592186044416) (and (= (uf_79 43) 8796093022208) (and (= (uf_79 42) 4398046511104) (and (= (uf_79 41) 2199023255552) (and (= (uf_79 40) 1099511627776) (and (= (uf_79 39) 549755813888) (and (= (uf_79 38) 274877906944) (and (= (uf_79 37) 137438953472) (and (= (uf_79 36) 68719476736) (and (= (uf_79 35) 34359738368) (and (= (uf_79 34) 17179869184) (and (= (uf_79 33) 8589934592) (and (= (uf_79 32) 4294967296) (and (= (uf_79 31) 2147483648) (and (= (uf_79 30) 1073741824) (and (= (uf_79 29) 536870912) (and (= (uf_79 28) 268435456) (and (= (uf_79 27) 134217728) (and (= (uf_79 26) 67108864) (and (= (uf_79 25) 33554432) (and (= (uf_79 24) 16777216) (and (= (uf_79 23) 8388608) (and (= (uf_79 22) 4194304) (and (= (uf_79 21) 2097152) (and (= (uf_79 20) 1048576) (and (= (uf_79 19) 524288) (and (= (uf_79 18) 262144) (and (= (uf_79 17) 131072) (and (= (uf_79 16) 65536) (and (= (uf_79 15) 32768) (and (= (uf_79 14) 16384) (and (= (uf_79 13) 8192) (and (= (uf_79 12) 4096) (and (= (uf_79 11) 2048) (and (= (uf_79 10) 1024) (and (= (uf_79 9) 512) (and (= (uf_79 8) 256) (and (= (uf_79 7) 128) (and (= (uf_79 6) 64) (and (= (uf_79 5) 32) (and (= (uf_79 4) 16) (and (= (uf_79 3) 8) (and (= (uf_79 2) 4) (and (= (uf_79 1) 2) (= (uf_79 0) 1)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) -:assumption (forall (?x271 T4) (?x272 T5) (implies (= (uf_51 ?x271) uf_9) (and (<= (uf_107 ?x271 ?x272) uf_75) (<= 0 (uf_107 ?x271 ?x272)))) :pat { (uf_107 ?x271 ?x272) }) -:assumption (forall (?x273 T4) (?x274 T5) (implies (= (uf_51 ?x273) uf_9) (and (<= (uf_108 ?x273 ?x274) uf_76) (<= 0 (uf_108 ?x273 ?x274)))) :pat { (uf_108 ?x273 ?x274) }) -:assumption (forall (?x275 T4) (?x276 T5) (implies (= (uf_51 ?x275) uf_9) (and (<= (uf_109 ?x275 ?x276) uf_77) (<= 0 (uf_109 ?x275 ?x276)))) :pat { (uf_109 ?x275 ?x276) }) -:assumption (forall (?x277 T4) (?x278 T5) (implies (= (uf_51 ?x277) uf_9) (and (<= (uf_110 ?x277 ?x278) uf_78) (<= 0 (uf_110 ?x277 ?x278)))) :pat { (uf_110 ?x277 ?x278) }) -:assumption (forall (?x279 T4) (?x280 T5) (implies (= (uf_51 ?x279) uf_9) (and (<= (uf_111 ?x279 ?x280) uf_85) (<= uf_86 (uf_111 ?x279 ?x280)))) :pat { (uf_111 ?x279 ?x280) }) -:assumption (forall (?x281 T4) (?x282 T5) (implies (= (uf_51 ?x281) uf_9) (and (<= (uf_112 ?x281 ?x282) uf_88) (<= uf_89 (uf_112 ?x281 ?x282)))) :pat { (uf_112 ?x281 ?x282) }) -:assumption (forall (?x283 T4) (?x284 T5) (implies (= (uf_51 ?x283) uf_9) (and (<= (uf_113 ?x283 ?x284) uf_92) (<= uf_93 (uf_113 ?x283 ?x284)))) :pat { (uf_113 ?x283 ?x284) }) -:assumption (forall (?x285 T4) (?x286 T5) (implies (= (uf_51 ?x285) uf_9) (and (<= (uf_114 ?x285 ?x286) uf_95) (<= uf_96 (uf_114 ?x285 ?x286)))) :pat { (uf_114 ?x285 ?x286) }) -:assumption (forall (?x287 T5) (?x288 T5) (= (uf_115 ?x287 ?x288) (+ (uf_116 ?x287) (uf_116 ?x288))) :pat { (uf_115 ?x287 ?x288) }) -:assumption (forall (?x289 T5) (implies (and (<= (uf_116 ?x289) uf_88) (<= uf_89 (uf_116 ?x289))) (= (uf_117 ?x289) (uf_116 ?x289))) :pat { (uf_117 ?x289) }) -:assumption (forall (?x290 T5) (implies (and (<= (uf_116 ?x290) uf_76) (<= 0 (uf_116 ?x290))) (= (uf_118 ?x290) (uf_116 ?x290))) :pat { (uf_118 ?x290) }) -:assumption (forall (?x291 T5) (implies (and (<= (uf_116 ?x291) uf_85) (<= uf_86 (uf_116 ?x291))) (= (uf_119 ?x291) (uf_116 ?x291))) :pat { (uf_119 ?x291) }) -:assumption (forall (?x292 T5) (implies (and (<= (uf_116 ?x292) uf_75) (<= 0 (uf_116 ?x292))) (= (uf_120 ?x292) (uf_116 ?x292))) :pat { (uf_120 ?x292) }) +:assumption (and (= (uf_79 0) 1) (and (= (uf_79 1) 2) (and (= (uf_79 2) 4) (and (= (uf_79 3) 8) (and (= (uf_79 4) 16) (and (= (uf_79 5) 32) (and (= (uf_79 6) 64) (and (= (uf_79 7) 128) (and (= (uf_79 8) 256) (and (= (uf_79 9) 512) (and (= (uf_79 10) 1024) (and (= (uf_79 11) 2048) (and (= (uf_79 12) 4096) (and (= (uf_79 13) 8192) (and (= (uf_79 14) 16384) (and (= (uf_79 15) 32768) (and (= (uf_79 16) 65536) (and (= (uf_79 17) 131072) (and (= (uf_79 18) 262144) (and (= (uf_79 19) 524288) (and (= (uf_79 20) 1048576) (and (= (uf_79 21) 2097152) (and (= (uf_79 22) 4194304) (and (= (uf_79 23) 8388608) (and (= (uf_79 24) 16777216) (and (= (uf_79 25) 33554432) (and (= (uf_79 26) 67108864) (and (= (uf_79 27) 134217728) (and (= (uf_79 28) 268435456) (and (= (uf_79 29) 536870912) (and (= (uf_79 30) 1073741824) (and (= (uf_79 31) 2147483648) (and (= (uf_79 32) 4294967296) (and (= (uf_79 33) 8589934592) (and (= (uf_79 34) 17179869184) (and (= (uf_79 35) 34359738368) (and (= (uf_79 36) 68719476736) (and (= (uf_79 37) 137438953472) (and (= (uf_79 38) 274877906944) (and (= (uf_79 39) 549755813888) (and (= (uf_79 40) 1099511627776) (and (= (uf_79 41) 2199023255552) (and (= (uf_79 42) 4398046511104) (and (= (uf_79 43) 8796093022208) (and (= (uf_79 44) 17592186044416) (and (= (uf_79 45) 35184372088832) (and (= (uf_79 46) 70368744177664) (and (= (uf_79 47) 140737488355328) (and (= (uf_79 48) 281474976710656) (and (= (uf_79 49) 562949953421312) (and (= (uf_79 50) 1125899906842624) (and (= (uf_79 51) 2251799813685248) (and (= (uf_79 52) 4503599627370496) (and (= (uf_79 53) 9007199254740992) (and (= (uf_79 54) 18014398509481984) (and (= (uf_79 55) 36028797018963968) (and (= (uf_79 56) 72057594037927936) (and (= (uf_79 57) 144115188075855872) (and (= (uf_79 58) 288230376151711744) (and (= (uf_79 59) 576460752303423488) (and (= (uf_79 60) 1152921504606846976) (and (= (uf_79 61) 2305843009213693952) (and (= (uf_79 62) 4611686018427387904) (= (uf_79 63) 9223372036854775808)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) +:assumption (forall (?x271 T4) (?x272 T5) (implies (= (uf_51 ?x271) uf_9) (and (<= 0 (uf_107 ?x271 ?x272)) (<= (uf_107 ?x271 ?x272) uf_75))) :pat { (uf_107 ?x271 ?x272) }) +:assumption (forall (?x273 T4) (?x274 T5) (implies (= (uf_51 ?x273) uf_9) (and (<= 0 (uf_108 ?x273 ?x274)) (<= (uf_108 ?x273 ?x274) uf_76))) :pat { (uf_108 ?x273 ?x274) }) +:assumption (forall (?x275 T4) (?x276 T5) (implies (= (uf_51 ?x275) uf_9) (and (<= 0 (uf_109 ?x275 ?x276)) (<= (uf_109 ?x275 ?x276) uf_77))) :pat { (uf_109 ?x275 ?x276) }) +:assumption (forall (?x277 T4) (?x278 T5) (implies (= (uf_51 ?x277) uf_9) (and (<= 0 (uf_110 ?x277 ?x278)) (<= (uf_110 ?x277 ?x278) uf_78))) :pat { (uf_110 ?x277 ?x278) }) +:assumption (forall (?x279 T4) (?x280 T5) (implies (= (uf_51 ?x279) uf_9) (and (<= uf_85 (uf_111 ?x279 ?x280)) (<= (uf_111 ?x279 ?x280) uf_86))) :pat { (uf_111 ?x279 ?x280) }) +:assumption (forall (?x281 T4) (?x282 T5) (implies (= (uf_51 ?x281) uf_9) (and (<= uf_88 (uf_112 ?x281 ?x282)) (<= (uf_112 ?x281 ?x282) uf_89))) :pat { (uf_112 ?x281 ?x282) }) +:assumption (forall (?x283 T4) (?x284 T5) (implies (= (uf_51 ?x283) uf_9) (and (<= uf_92 (uf_113 ?x283 ?x284)) (<= (uf_113 ?x283 ?x284) uf_93))) :pat { (uf_113 ?x283 ?x284) }) +:assumption (forall (?x285 T4) (?x286 T5) (implies (= (uf_51 ?x285) uf_9) (and (<= uf_95 (uf_114 ?x285 ?x286)) (<= (uf_114 ?x285 ?x286) uf_96))) :pat { (uf_114 ?x285 ?x286) }) +:assumption (forall (?x287 T5) (?x288 T5) (= (uf_115 ?x287 ?x288) (- (uf_116 ?x287) (uf_116 ?x288))) :pat { (uf_115 ?x287 ?x288) }) +:assumption (forall (?x289 T5) (implies (and (<= uf_88 (uf_116 ?x289)) (<= (uf_116 ?x289) uf_89)) (= (uf_117 ?x289) (uf_116 ?x289))) :pat { (uf_117 ?x289) }) +:assumption (forall (?x290 T5) (implies (and (<= 0 (uf_116 ?x290)) (<= (uf_116 ?x290) uf_76)) (= (uf_118 ?x290) (uf_116 ?x290))) :pat { (uf_118 ?x290) }) +:assumption (forall (?x291 T5) (implies (and (<= uf_85 (uf_116 ?x291)) (<= (uf_116 ?x291) uf_86)) (= (uf_119 ?x291) (uf_116 ?x291))) :pat { (uf_119 ?x291) }) +:assumption (forall (?x292 T5) (implies (and (<= 0 (uf_116 ?x292)) (<= (uf_116 ?x292) uf_75)) (= (uf_120 ?x292) (uf_116 ?x292))) :pat { (uf_120 ?x292) }) :assumption (= (uf_117 uf_121) 0) :assumption (= (uf_118 uf_121) 0) :assumption (= (uf_119 uf_121) 0) @@ -449,134 +449,134 @@ :assumption (forall (?x303 T4) (?x304 T5) (= (uf_112 ?x303 ?x304) (uf_19 (uf_20 ?x303) ?x304)) :pat { (uf_112 ?x303 ?x304) }) :assumption (forall (?x305 T4) (?x306 T5) (= (uf_113 ?x305 ?x306) (uf_19 (uf_20 ?x305) ?x306)) :pat { (uf_113 ?x305 ?x306) }) :assumption (forall (?x307 T4) (?x308 T5) (= (uf_114 ?x307 ?x308) (uf_19 (uf_20 ?x307) ?x308)) :pat { (uf_114 ?x307 ?x308) }) -:assumption (= uf_75 (+ (+ (+ (+ 65536 65536) 65536) 65536) 1)) -:assumption (= uf_76 (+ (+ 65536 65536) 1)) +:assumption (= uf_75 (- (* (* (* 65536 65536) 65536) 65536) 1)) +:assumption (= uf_76 (- (* 65536 65536) 1)) :assumption (= uf_77 65535) :assumption (= uf_78 255) -:assumption (= uf_85 (+ (+ (+ (+ 65536 65536) 65536) 32768) 1)) -:assumption (= uf_86 (+ 0 (+ (+ (+ 65536 65536) 65536) 32768))) -:assumption (= uf_88 (+ (+ 65536 32768) 1)) -:assumption (= uf_89 (+ 0 (+ 65536 32768))) -:assumption (= uf_92 32767) -:assumption (= uf_93 (+ 0 32768)) -:assumption (= uf_95 127) -:assumption (= uf_96 (+ 0 128)) +:assumption (= uf_86 (- (* (* (* 65536 65536) 65536) 32768) 1)) +:assumption (= uf_85 (- 0 (* (* (* 65536 65536) 65536) 32768))) +:assumption (= uf_89 (- (* 65536 32768) 1)) +:assumption (= uf_88 (- 0 (* 65536 32768))) +:assumption (= uf_93 32767) +:assumption (= uf_92 (- 0 32768)) +:assumption (= uf_96 127) +:assumption (= uf_95 (- 0 128)) :assumption (forall (?x309 T2) (iff (= (uf_122 ?x309) uf_9) (= ?x309 uf_9)) :pat { (uf_122 ?x309) }) -:assumption (forall (?x310 T4) (?x311 T4) (?x312 T5) (?x313 T3) (?x314 Int) (implies (= (uf_23 ?x313) uf_9) (implies (= (uf_123 ?x310 ?x311 ?x312 (uf_124 ?x313 ?x314)) uf_9) (forall (?x315 Int) (implies (and (< ?x315 ?x314) (<= 0 ?x315)) (= (uf_19 (uf_20 ?x310) (uf_66 (uf_43 ?x313 (uf_116 ?x312)) ?x315 ?x313)) (uf_19 (uf_20 ?x311) (uf_66 (uf_43 ?x313 (uf_116 ?x312)) ?x315 ?x313)))) :pat { (uf_19 (uf_20 ?x311) (uf_66 (uf_43 ?x313 (uf_116 ?x312)) ?x315 ?x313)) }))) :pat { (uf_123 ?x310 ?x311 ?x312 (uf_124 ?x313 ?x314)) (uf_23 ?x313) }) -:assumption (forall (?x316 T5) (?x317 Int) (?x318 T15) (= (uf_125 (uf_126 (uf_66 ?x316 ?x317 (uf_15 ?x316)) ?x318) ?x316) ?x317) :pat { (uf_125 (uf_126 (uf_66 ?x316 ?x317 (uf_15 ?x316)) ?x318) ?x316) }) -:assumption (forall (?x319 T5) (?x320 Int) (= (uf_125 (uf_66 ?x319 ?x320 (uf_15 ?x319)) ?x319) ?x320) :pat { (uf_66 ?x319 ?x320 (uf_15 ?x319)) }) -:assumption (forall (?x321 T5) (?x322 T4) (?x323 T5) (iff (= (uf_13 ?x321 (uf_127 ?x322 ?x323)) uf_9) (and (= (uf_13 ?x321 (uf_128 ?x322 ?x323)) uf_9) (not (= (uf_116 ?x323) (uf_116 uf_121))))) :pat { (uf_13 ?x321 (uf_127 ?x322 ?x323)) }) -:assumption (forall (?x324 T5) (?x325 Int) (?x326 T3) (?x327 Int) (iff (= (uf_13 ?x324 (uf_129 (uf_43 ?x326 ?x325) ?x326 ?x327)) uf_9) (and (= (uf_13 ?x324 (uf_130 (uf_66 (uf_43 ?x326 ?x325) (uf_125 ?x324 (uf_43 ?x326 ?x325)) ?x326))) uf_9) (and (<= (uf_125 ?x324 (uf_43 ?x326 ?x325)) (+ ?x327 1)) (and (<= 0 (uf_125 ?x324 (uf_43 ?x326 ?x325))) (not (= ?x325 0)))))) :pat { (uf_13 ?x324 (uf_129 (uf_43 ?x326 ?x325) ?x326 ?x327)) }) -:assumption (forall (?x328 T5) (?x329 T3) (?x330 Int) (?x331 Int) (?x332 T6) (implies (and (< ?x331 ?x330) (<= 0 ?x331)) (= (uf_133 (uf_66 ?x328 ?x331 ?x329) ?x332 (uf_132 ?x328 ?x329 ?x330)) 2)) :pat { (uf_66 ?x328 ?x331 ?x329) (uf_131 ?x332 (uf_132 ?x328 ?x329 ?x330)) }) -:assumption (forall (?x333 T5) (?x334 T3) (?x335 Int) (?x336 Int) (?x337 T6) (implies (and (< ?x336 ?x335) (<= 0 ?x336)) (= (uf_133 (uf_66 ?x333 ?x336 ?x334) (uf_132 ?x333 ?x334 ?x335) ?x337) 1)) :pat { (uf_66 ?x333 ?x336 ?x334) (uf_131 (uf_132 ?x333 ?x334 ?x335) ?x337) }) -:assumption (forall (?x338 T5) (?x339 Int) (?x340 T3) (?x341 Int) (iff (= (uf_13 ?x338 (uf_132 (uf_43 ?x340 ?x339) ?x340 ?x341)) uf_9) (and (= (uf_13 ?x338 (uf_130 (uf_66 (uf_43 ?x340 ?x339) (uf_125 ?x338 (uf_43 ?x340 ?x339)) ?x340))) uf_9) (and (<= (uf_125 ?x338 (uf_43 ?x340 ?x339)) (+ ?x341 1)) (<= 0 (uf_125 ?x338 (uf_43 ?x340 ?x339)))))) :pat { (uf_13 ?x338 (uf_132 (uf_43 ?x340 ?x339) ?x340 ?x341)) }) -:assumption (forall (?x342 T5) (?x343 T3) (?x344 Int) (?x345 T5) (iff (= (uf_13 ?x345 (uf_134 ?x342 ?x343 ?x344)) uf_9) (and (= ?x345 (uf_66 ?x342 (uf_125 ?x345 ?x342) ?x343)) (and (<= (uf_125 ?x345 ?x342) (+ ?x344 1)) (<= 0 (uf_125 ?x345 ?x342))))) :pat { (uf_13 ?x345 (uf_134 ?x342 ?x343 ?x344)) }) -:assumption (forall (?x346 T4) (?x347 Int) (?x348 T3) (?x349 Int) (?x350 Int) (implies (= (uf_27 ?x346 (uf_43 (uf_124 ?x348 ?x349) ?x347)) uf_9) (implies (and (< ?x350 ?x349) (<= 0 ?x350)) (and (= (uf_27 ?x346 (uf_66 (uf_43 (uf_124 ?x348 ?x349) ?x347) ?x350 ?x348)) uf_9) (and (up_68 (uf_58 (uf_59 ?x346) (uf_66 (uf_43 (uf_124 ?x348 ?x349) ?x347) ?x350 ?x348))) (and (not (= (uf_135 (uf_58 (uf_59 ?x346) (uf_66 (uf_43 (uf_124 ?x348 ?x349) ?x347) ?x350 ?x348))) uf_9)) (= (uf_136 (uf_58 (uf_59 ?x346) (uf_66 (uf_43 (uf_124 ?x348 ?x349) ?x347) ?x350 ?x348))) (uf_43 (uf_124 ?x348 ?x349) ?x347))))))) :pat { (uf_40 (uf_41 ?x346) (uf_66 (uf_43 ?x348 ?x347) ?x350 ?x348)) (uf_43 (uf_124 ?x348 ?x349) ?x347) } :pat { (uf_58 (uf_59 ?x346) (uf_66 (uf_43 ?x348 ?x347) ?x350 ?x348)) (uf_43 (uf_124 ?x348 ?x349) ?x347) }) -:assumption (forall (?x351 T4) (?x352 T5) (?x353 Int) (?x354 T3) (?x355 Int) (iff (= (uf_13 ?x352 (uf_128 ?x351 (uf_43 (uf_124 ?x354 ?x355) ?x353))) uf_9) (or (and (= (uf_13 ?x352 (uf_128 ?x351 (uf_66 (uf_43 ?x354 ?x353) (uf_125 ?x352 (uf_43 ?x354 ?x353)) ?x354))) uf_9) (and (<= (uf_125 ?x352 (uf_43 ?x354 ?x353)) (+ ?x355 1)) (<= 0 (uf_125 ?x352 (uf_43 ?x354 ?x353))))) (= ?x352 (uf_43 (uf_124 ?x354 ?x355) ?x353)))) :pat { (uf_13 ?x352 (uf_128 ?x351 (uf_43 (uf_124 ?x354 ?x355) ?x353))) }) -:assumption (forall (?x356 T5) (?x357 Int) (?x358 T3) (?x359 Int) (iff (= (uf_13 ?x356 (uf_130 (uf_43 (uf_124 ?x358 ?x359) ?x357))) uf_9) (or (and (= (uf_13 ?x356 (uf_130 (uf_66 (uf_43 ?x358 ?x357) (uf_125 ?x356 (uf_43 ?x358 ?x357)) ?x358))) uf_9) (and (<= (uf_125 ?x356 (uf_43 ?x358 ?x357)) (+ ?x359 1)) (<= 0 (uf_125 ?x356 (uf_43 ?x358 ?x357))))) (= ?x356 (uf_43 (uf_124 ?x358 ?x359) ?x357)))) :pat { (uf_13 ?x356 (uf_130 (uf_43 (uf_124 ?x358 ?x359) ?x357))) }) -:assumption (forall (?x360 T4) (?x361 T5) (?x362 T3) (?x363 Int) (iff (= (uf_65 ?x360 ?x361 ?x362 ?x363) uf_9) (and (forall (?x364 Int) (implies (and (< ?x364 ?x363) (<= 0 ?x364)) (and (= (uf_27 ?x360 (uf_66 ?x361 ?x364 ?x362)) uf_9) (up_68 (uf_58 (uf_59 ?x360) (uf_66 ?x361 ?x364 ?x362))))) :pat { (uf_40 (uf_41 ?x360) (uf_66 ?x361 ?x364 ?x362)) } :pat { (uf_58 (uf_59 ?x360) (uf_66 ?x361 ?x364 ?x362)) } :pat { (uf_19 (uf_20 ?x360) (uf_66 ?x361 ?x364 ?x362)) }) (= (uf_48 ?x361 ?x362) uf_9))) :pat { (uf_65 ?x360 ?x361 ?x362 ?x363) }) -:assumption (forall (?x365 T4) (?x366 T5) (?x367 T3) (?x368 Int) (?x369 T2) (iff (= (uf_137 ?x365 ?x366 ?x367 ?x368 ?x369) uf_9) (and (forall (?x370 Int) (implies (and (< ?x370 ?x368) (<= 0 ?x370)) (and (= (uf_27 ?x365 (uf_66 ?x366 ?x370 ?x367)) uf_9) (and (up_68 (uf_58 (uf_59 ?x365) (uf_66 ?x366 ?x370 ?x367))) (iff (= (uf_135 (uf_58 (uf_59 ?x365) (uf_66 ?x366 ?x370 ?x367))) uf_9) (= ?x369 uf_9))))) :pat { (uf_40 (uf_41 ?x365) (uf_66 ?x366 ?x370 ?x367)) } :pat { (uf_58 (uf_59 ?x365) (uf_66 ?x366 ?x370 ?x367)) } :pat { (uf_19 (uf_20 ?x365) (uf_66 ?x366 ?x370 ?x367)) }) (= (uf_48 ?x366 ?x367) uf_9))) :pat { (uf_137 ?x365 ?x366 ?x367 ?x368 ?x369) }) -:assumption (forall (?x371 T5) (?x372 Int) (?x373 Int) (?x374 T3) (implies (and (not (= ?x373 0)) (not (= ?x372 0))) (= (uf_66 (uf_66 ?x371 ?x372 ?x374) ?x373 ?x374) (uf_66 ?x371 (+ ?x372 ?x373) ?x374))) :pat { (uf_66 (uf_66 ?x371 ?x372 ?x374) ?x373 ?x374) }) -:assumption (forall (?x375 T5) (?x376 Int) (?x377 T3) (and (= (uf_66 ?x375 ?x376 ?x377) (uf_43 ?x377 (+ (uf_116 ?x375) (+ ?x376 (uf_138 ?x377))))) (= (uf_139 (uf_66 ?x375 ?x376 ?x377) ?x375) uf_9)) :pat { (uf_66 ?x375 ?x376 ?x377) }) +:assumption (forall (?x310 T4) (?x311 T4) (?x312 T5) (?x313 T3) (?x314 Int) (implies (= (uf_22 ?x313) uf_9) (implies (= (uf_123 ?x310 ?x311 ?x312 (uf_124 ?x313 ?x314)) uf_9) (forall (?x315 Int) (implies (and (<= 0 ?x315) (< ?x315 ?x314)) (= (uf_19 (uf_20 ?x310) (uf_66 (uf_43 ?x313 (uf_116 ?x312)) ?x315 ?x313)) (uf_19 (uf_20 ?x311) (uf_66 (uf_43 ?x313 (uf_116 ?x312)) ?x315 ?x313)))) :pat { (uf_19 (uf_20 ?x311) (uf_66 (uf_43 ?x313 (uf_116 ?x312)) ?x315 ?x313)) }))) :pat { (uf_123 ?x310 ?x311 ?x312 (uf_124 ?x313 ?x314)) (uf_22 ?x313) }) +:assumption (forall (?x316 T5) (?x317 Int) (?x318 T15) (= (uf_125 (uf_126 (uf_66 ?x316 ?x317 (uf_13 ?x316)) ?x318) ?x316) ?x317) :pat { (uf_125 (uf_126 (uf_66 ?x316 ?x317 (uf_13 ?x316)) ?x318) ?x316) }) +:assumption (forall (?x319 T5) (?x320 Int) (= (uf_125 (uf_66 ?x319 ?x320 (uf_13 ?x319)) ?x319) ?x320) :pat { (uf_66 ?x319 ?x320 (uf_13 ?x319)) }) +:assumption (forall (?x321 T5) (?x322 T4) (?x323 T5) (iff (= (uf_15 ?x321 (uf_127 ?x322 ?x323)) uf_9) (and (not (= (uf_116 ?x323) (uf_116 uf_121))) (= (uf_15 ?x321 (uf_128 ?x322 ?x323)) uf_9))) :pat { (uf_15 ?x321 (uf_127 ?x322 ?x323)) }) +:assumption (forall (?x324 T5) (?x325 Int) (?x326 T3) (?x327 Int) (iff (= (uf_15 ?x324 (uf_129 (uf_43 ?x326 ?x325) ?x326 ?x327)) uf_9) (and (not (= ?x325 0)) (and (<= 0 (uf_125 ?x324 (uf_43 ?x326 ?x325))) (and (<= (uf_125 ?x324 (uf_43 ?x326 ?x325)) (- ?x327 1)) (= (uf_15 ?x324 (uf_130 (uf_66 (uf_43 ?x326 ?x325) (uf_125 ?x324 (uf_43 ?x326 ?x325)) ?x326))) uf_9))))) :pat { (uf_15 ?x324 (uf_129 (uf_43 ?x326 ?x325) ?x326 ?x327)) }) +:assumption (forall (?x328 T5) (?x329 T3) (?x330 Int) (?x331 Int) (?x332 T6) (implies (and (<= 0 ?x331) (< ?x331 ?x330)) (= (uf_133 (uf_66 ?x328 ?x331 ?x329) ?x332 (uf_132 ?x328 ?x329 ?x330)) 2)) :pat { (uf_66 ?x328 ?x331 ?x329) (uf_131 ?x332 (uf_132 ?x328 ?x329 ?x330)) }) +:assumption (forall (?x333 T5) (?x334 T3) (?x335 Int) (?x336 Int) (?x337 T6) (implies (and (<= 0 ?x336) (< ?x336 ?x335)) (= (uf_133 (uf_66 ?x333 ?x336 ?x334) (uf_132 ?x333 ?x334 ?x335) ?x337) 1)) :pat { (uf_66 ?x333 ?x336 ?x334) (uf_131 (uf_132 ?x333 ?x334 ?x335) ?x337) }) +:assumption (forall (?x338 T5) (?x339 Int) (?x340 T3) (?x341 Int) (iff (= (uf_15 ?x338 (uf_132 (uf_43 ?x340 ?x339) ?x340 ?x341)) uf_9) (and (<= 0 (uf_125 ?x338 (uf_43 ?x340 ?x339))) (and (<= (uf_125 ?x338 (uf_43 ?x340 ?x339)) (- ?x341 1)) (= (uf_15 ?x338 (uf_130 (uf_66 (uf_43 ?x340 ?x339) (uf_125 ?x338 (uf_43 ?x340 ?x339)) ?x340))) uf_9)))) :pat { (uf_15 ?x338 (uf_132 (uf_43 ?x340 ?x339) ?x340 ?x341)) }) +:assumption (forall (?x342 T5) (?x343 T3) (?x344 Int) (?x345 T5) (iff (= (uf_15 ?x345 (uf_134 ?x342 ?x343 ?x344)) uf_9) (and (<= 0 (uf_125 ?x345 ?x342)) (and (<= (uf_125 ?x345 ?x342) (- ?x344 1)) (= ?x345 (uf_66 ?x342 (uf_125 ?x345 ?x342) ?x343))))) :pat { (uf_15 ?x345 (uf_134 ?x342 ?x343 ?x344)) }) +:assumption (forall (?x346 T4) (?x347 Int) (?x348 T3) (?x349 Int) (?x350 Int) (implies (= (uf_24 ?x346 (uf_43 (uf_124 ?x348 ?x349) ?x347)) uf_9) (implies (and (<= 0 ?x350) (< ?x350 ?x349)) (and (= (uf_135 (uf_58 (uf_59 ?x346) (uf_66 (uf_43 (uf_124 ?x348 ?x349) ?x347) ?x350 ?x348))) (uf_43 (uf_124 ?x348 ?x349) ?x347)) (and (not (= (uf_136 (uf_58 (uf_59 ?x346) (uf_66 (uf_43 (uf_124 ?x348 ?x349) ?x347) ?x350 ?x348))) uf_9)) (and (up_67 (uf_58 (uf_59 ?x346) (uf_66 (uf_43 (uf_124 ?x348 ?x349) ?x347) ?x350 ?x348))) (= (uf_24 ?x346 (uf_66 (uf_43 (uf_124 ?x348 ?x349) ?x347) ?x350 ?x348)) uf_9)))))) :pat { (uf_40 (uf_41 ?x346) (uf_66 (uf_43 ?x348 ?x347) ?x350 ?x348)) (uf_43 (uf_124 ?x348 ?x349) ?x347) } :pat { (uf_58 (uf_59 ?x346) (uf_66 (uf_43 ?x348 ?x347) ?x350 ?x348)) (uf_43 (uf_124 ?x348 ?x349) ?x347) }) +:assumption (forall (?x351 T4) (?x352 T5) (?x353 Int) (?x354 T3) (?x355 Int) (iff (= (uf_15 ?x352 (uf_128 ?x351 (uf_43 (uf_124 ?x354 ?x355) ?x353))) uf_9) (or (= ?x352 (uf_43 (uf_124 ?x354 ?x355) ?x353)) (and (<= 0 (uf_125 ?x352 (uf_43 ?x354 ?x353))) (and (<= (uf_125 ?x352 (uf_43 ?x354 ?x353)) (- ?x355 1)) (= (uf_15 ?x352 (uf_128 ?x351 (uf_66 (uf_43 ?x354 ?x353) (uf_125 ?x352 (uf_43 ?x354 ?x353)) ?x354))) uf_9))))) :pat { (uf_15 ?x352 (uf_128 ?x351 (uf_43 (uf_124 ?x354 ?x355) ?x353))) }) +:assumption (forall (?x356 T5) (?x357 Int) (?x358 T3) (?x359 Int) (iff (= (uf_15 ?x356 (uf_130 (uf_43 (uf_124 ?x358 ?x359) ?x357))) uf_9) (or (= ?x356 (uf_43 (uf_124 ?x358 ?x359) ?x357)) (and (<= 0 (uf_125 ?x356 (uf_43 ?x358 ?x357))) (and (<= (uf_125 ?x356 (uf_43 ?x358 ?x357)) (- ?x359 1)) (= (uf_15 ?x356 (uf_130 (uf_66 (uf_43 ?x358 ?x357) (uf_125 ?x356 (uf_43 ?x358 ?x357)) ?x358))) uf_9))))) :pat { (uf_15 ?x356 (uf_130 (uf_43 (uf_124 ?x358 ?x359) ?x357))) }) +:assumption (forall (?x360 T4) (?x361 T5) (?x362 T3) (?x363 Int) (iff (= (uf_65 ?x360 ?x361 ?x362 ?x363) uf_9) (and (= (uf_48 ?x361 ?x362) uf_9) (forall (?x364 Int) (implies (and (<= 0 ?x364) (< ?x364 ?x363)) (and (up_67 (uf_58 (uf_59 ?x360) (uf_66 ?x361 ?x364 ?x362))) (= (uf_24 ?x360 (uf_66 ?x361 ?x364 ?x362)) uf_9))) :pat { (uf_40 (uf_41 ?x360) (uf_66 ?x361 ?x364 ?x362)) } :pat { (uf_58 (uf_59 ?x360) (uf_66 ?x361 ?x364 ?x362)) } :pat { (uf_19 (uf_20 ?x360) (uf_66 ?x361 ?x364 ?x362)) }))) :pat { (uf_65 ?x360 ?x361 ?x362 ?x363) }) +:assumption (forall (?x365 T4) (?x366 T5) (?x367 T3) (?x368 Int) (?x369 T2) (iff (= (uf_137 ?x365 ?x366 ?x367 ?x368 ?x369) uf_9) (and (= (uf_48 ?x366 ?x367) uf_9) (forall (?x370 Int) (implies (and (<= 0 ?x370) (< ?x370 ?x368)) (and (iff (= (uf_136 (uf_58 (uf_59 ?x365) (uf_66 ?x366 ?x370 ?x367))) uf_9) (= ?x369 uf_9)) (and (up_67 (uf_58 (uf_59 ?x365) (uf_66 ?x366 ?x370 ?x367))) (= (uf_24 ?x365 (uf_66 ?x366 ?x370 ?x367)) uf_9)))) :pat { (uf_40 (uf_41 ?x365) (uf_66 ?x366 ?x370 ?x367)) } :pat { (uf_58 (uf_59 ?x365) (uf_66 ?x366 ?x370 ?x367)) } :pat { (uf_19 (uf_20 ?x365) (uf_66 ?x366 ?x370 ?x367)) }))) :pat { (uf_137 ?x365 ?x366 ?x367 ?x368 ?x369) }) +:assumption (forall (?x371 T5) (?x372 Int) (?x373 Int) (?x374 T3) (implies (and (not (= ?x372 0)) (not (= ?x373 0))) (= (uf_66 (uf_66 ?x371 ?x372 ?x374) ?x373 ?x374) (uf_66 ?x371 (+ ?x372 ?x373) ?x374))) :pat { (uf_66 (uf_66 ?x371 ?x372 ?x374) ?x373 ?x374) }) +:assumption (forall (?x375 T5) (?x376 Int) (?x377 T3) (and (= (uf_138 (uf_66 ?x375 ?x376 ?x377) ?x375) uf_9) (= (uf_66 ?x375 ?x376 ?x377) (uf_43 ?x377 (+ (uf_116 ?x375) (* ?x376 (uf_139 ?x377)))))) :pat { (uf_66 ?x375 ?x376 ?x377) }) :assumption (forall (?x378 T5) (?x379 T3) (= (uf_140 ?x378 ?x379) ?x378) :pat { (uf_140 ?x378 ?x379) }) :assumption (forall (?x380 T3) (?x381 Int) (not (up_36 (uf_124 ?x380 ?x381))) :pat { (uf_124 ?x380 ?x381) }) :assumption (forall (?x382 T3) (?x383 Int) (= (uf_141 (uf_124 ?x382 ?x383)) uf_9) :pat { (uf_124 ?x382 ?x383) }) :assumption (forall (?x384 T3) (?x385 Int) (= (uf_142 (uf_124 ?x384 ?x385)) 0) :pat { (uf_124 ?x384 ?x385) }) :assumption (forall (?x386 T3) (?x387 Int) (= (uf_143 (uf_124 ?x386 ?x387)) ?x387) :pat { (uf_124 ?x386 ?x387) }) :assumption (forall (?x388 T3) (?x389 Int) (= (uf_144 (uf_124 ?x388 ?x389)) ?x388) :pat { (uf_124 ?x388 ?x389) }) -:assumption (forall (?x390 T5) (?x391 T6) (iff (= (uf_13 ?x390 ?x391) uf_9) (= (uf_145 ?x390 ?x391) uf_9)) :pat { (uf_145 ?x390 ?x391) }) -:assumption (forall (?x392 T5) (?x393 T6) (iff (= (uf_13 ?x392 ?x393) uf_9) (up_146 ?x392 ?x393)) :pat { (uf_13 ?x392 ?x393) }) -:assumption (forall (?x394 T5) (?x395 T6) (iff (= (uf_13 ?x394 ?x395) uf_9) (= (uf_147 ?x394 ?x395) uf_9)) :pat { (uf_13 ?x394 ?x395) }) -:assumption (forall (?x396 T5) (?x397 T4) (?x398 T5) (iff (= (uf_13 ?x396 (uf_53 ?x397 ?x398)) uf_9) (= (uf_147 ?x396 (uf_53 ?x397 ?x398)) uf_9)) :pat { (uf_147 ?x396 (uf_53 ?x397 ?x398)) (uf_148 ?x396) }) -:assumption (forall (?x399 T5) (?x400 T4) (?x401 T5) (implies (= (uf_13 ?x399 (uf_53 ?x400 ?x401)) uf_9) (= (uf_148 ?x399) uf_9)) :pat { (uf_13 ?x399 (uf_53 ?x400 ?x401)) }) -:assumption (forall (?x402 T6) (?x403 T6) (implies (forall (?x404 T5) (and (implies (= (uf_13 ?x404 ?x403) uf_9) (not (= (uf_13 ?x404 ?x402) uf_9))) (implies (= (uf_13 ?x404 ?x402) uf_9) (not (= (uf_13 ?x404 ?x403) uf_9)))) :pat { (uf_18 ?x404) }) (= (uf_131 ?x402 ?x403) uf_9)) :pat { (uf_131 ?x402 ?x403) }) -:assumption (forall (?x405 T5) (?x406 T6) (?x407 T6) (implies (and (= (uf_13 ?x405 ?x407) uf_9) (= (uf_131 ?x406 ?x407) uf_9)) (= (uf_133 ?x405 ?x406 ?x407) 2)) :pat { (uf_131 ?x406 ?x407) (uf_13 ?x405 ?x407) }) -:assumption (forall (?x408 T5) (?x409 T6) (?x410 T6) (implies (and (= (uf_13 ?x408 ?x409) uf_9) (= (uf_131 ?x409 ?x410) uf_9)) (= (uf_133 ?x408 ?x409 ?x410) 1)) :pat { (uf_131 ?x409 ?x410) (uf_13 ?x408 ?x409) }) -:assumption (forall (?x411 T5) (= (uf_13 ?x411 uf_149) uf_9) :pat { (uf_13 ?x411 uf_149) }) +:assumption (forall (?x390 T5) (?x391 T6) (iff (= (uf_15 ?x390 ?x391) uf_9) (= (uf_145 ?x390 ?x391) uf_9)) :pat { (uf_145 ?x390 ?x391) }) +:assumption (forall (?x392 T5) (?x393 T6) (iff (= (uf_15 ?x392 ?x393) uf_9) (up_146 ?x392 ?x393)) :pat { (uf_15 ?x392 ?x393) }) +:assumption (forall (?x394 T5) (?x395 T6) (iff (= (uf_15 ?x394 ?x395) uf_9) (= (uf_147 ?x394 ?x395) uf_9)) :pat { (uf_15 ?x394 ?x395) }) +:assumption (forall (?x396 T5) (?x397 T4) (?x398 T5) (iff (= (uf_15 ?x396 (uf_53 ?x397 ?x398)) uf_9) (= (uf_147 ?x396 (uf_53 ?x397 ?x398)) uf_9)) :pat { (uf_147 ?x396 (uf_53 ?x397 ?x398)) (uf_148 ?x396) }) +:assumption (forall (?x399 T5) (?x400 T4) (?x401 T5) (implies (= (uf_15 ?x399 (uf_53 ?x400 ?x401)) uf_9) (= (uf_148 ?x399) uf_9)) :pat { (uf_15 ?x399 (uf_53 ?x400 ?x401)) }) +:assumption (forall (?x402 T6) (?x403 T6) (implies (forall (?x404 T5) (and (implies (= (uf_15 ?x404 ?x402) uf_9) (not (= (uf_15 ?x404 ?x403) uf_9))) (implies (= (uf_15 ?x404 ?x403) uf_9) (not (= (uf_15 ?x404 ?x402) uf_9)))) :pat { (uf_18 ?x404) }) (= (uf_131 ?x402 ?x403) uf_9)) :pat { (uf_131 ?x402 ?x403) }) +:assumption (forall (?x405 T5) (?x406 T6) (?x407 T6) (implies (and (= (uf_131 ?x406 ?x407) uf_9) (= (uf_15 ?x405 ?x407) uf_9)) (= (uf_133 ?x405 ?x406 ?x407) 2)) :pat { (uf_131 ?x406 ?x407) (uf_15 ?x405 ?x407) }) +:assumption (forall (?x408 T5) (?x409 T6) (?x410 T6) (implies (and (= (uf_131 ?x409 ?x410) uf_9) (= (uf_15 ?x408 ?x409) uf_9)) (= (uf_133 ?x408 ?x409 ?x410) 1)) :pat { (uf_131 ?x409 ?x410) (uf_15 ?x408 ?x409) }) +:assumption (forall (?x411 T5) (= (uf_15 ?x411 uf_149) uf_9) :pat { (uf_15 ?x411 uf_149) }) :assumption (forall (?x412 T5) (= (uf_150 (uf_151 ?x412)) 1)) :assumption (= (uf_150 uf_152) 0) :assumption (forall (?x413 T6) (?x414 T6) (implies (= (uf_153 ?x413 ?x414) uf_9) (= ?x413 ?x414)) :pat { (uf_153 ?x413 ?x414) }) -:assumption (forall (?x415 T6) (?x416 T6) (implies (forall (?x417 T5) (iff (= (uf_13 ?x417 ?x415) uf_9) (= (uf_13 ?x417 ?x416) uf_9)) :pat { (uf_18 ?x417) }) (= (uf_153 ?x415 ?x416) uf_9)) :pat { (uf_153 ?x415 ?x416) }) -:assumption (forall (?x418 T6) (?x419 T6) (iff (= (uf_154 ?x418 ?x419) uf_9) (forall (?x420 T5) (implies (= (uf_13 ?x420 ?x418) uf_9) (= (uf_13 ?x420 ?x419) uf_9)) :pat { (uf_13 ?x420 ?x418) } :pat { (uf_13 ?x420 ?x419) })) :pat { (uf_154 ?x418 ?x419) }) -:assumption (forall (?x421 T6) (?x422 T6) (?x423 T5) (iff (= (uf_13 ?x423 (uf_155 ?x421 ?x422)) uf_9) (and (= (uf_13 ?x423 ?x422) uf_9) (= (uf_13 ?x423 ?x421) uf_9))) :pat { (uf_13 ?x423 (uf_155 ?x421 ?x422)) }) -:assumption (forall (?x424 T6) (?x425 T6) (?x426 T5) (iff (= (uf_13 ?x426 (uf_156 ?x424 ?x425)) uf_9) (and (not (= (uf_13 ?x426 ?x425) uf_9)) (= (uf_13 ?x426 ?x424) uf_9))) :pat { (uf_13 ?x426 (uf_156 ?x424 ?x425)) }) -:assumption (forall (?x427 T6) (?x428 T6) (?x429 T5) (iff (= (uf_13 ?x429 (uf_157 ?x427 ?x428)) uf_9) (or (= (uf_13 ?x429 ?x428) uf_9) (= (uf_13 ?x429 ?x427) uf_9))) :pat { (uf_13 ?x429 (uf_157 ?x427 ?x428)) }) -:assumption (forall (?x430 T5) (?x431 T5) (iff (= (uf_13 ?x431 (uf_158 ?x430)) uf_9) (and (not (= (uf_116 ?x430) (uf_116 uf_121))) (= ?x430 ?x431))) :pat { (uf_13 ?x431 (uf_158 ?x430)) }) -:assumption (forall (?x432 T5) (?x433 T5) (iff (= (uf_13 ?x433 (uf_151 ?x432)) uf_9) (= ?x432 ?x433)) :pat { (uf_13 ?x433 (uf_151 ?x432)) }) -:assumption (forall (?x434 T5) (not (= (uf_13 ?x434 uf_152) uf_9)) :pat { (uf_13 ?x434 uf_152) }) -:assumption (forall (?x435 T5) (?x436 T5) (= (uf_159 ?x435 ?x436) (uf_43 (uf_124 (uf_144 (uf_15 ?x435)) (+ (uf_143 (uf_15 ?x435)) (uf_143 (uf_15 ?x436)))) (uf_116 ?x435))) :pat { (uf_159 ?x435 ?x436) }) -:assumption (forall (?x437 T5) (?x438 Int) (= (uf_160 ?x437 ?x438) (uf_43 (uf_124 (uf_144 (uf_15 ?x437)) (+ (uf_143 (uf_15 ?x437)) ?x438)) (uf_116 (uf_66 (uf_43 (uf_144 (uf_15 ?x437)) (uf_116 ?x437)) ?x438 (uf_144 (uf_15 ?x437)))))) :pat { (uf_160 ?x437 ?x438) }) -:assumption (forall (?x439 T5) (?x440 Int) (= (uf_161 ?x439 ?x440) (uf_43 (uf_124 (uf_144 (uf_15 ?x439)) ?x440) (uf_116 ?x439))) :pat { (uf_161 ?x439 ?x440) }) -:assumption (forall (?x441 T4) (?x442 T5) (?x443 T5) (iff (= (uf_13 ?x442 (uf_162 ?x441 ?x443)) uf_9) (or (and (= (uf_13 ?x442 (uf_163 ?x443)) uf_9) (= (uf_135 (uf_58 (uf_59 ?x441) ?x442)) uf_9)) (= ?x442 ?x443))) :pat { (uf_13 ?x442 (uf_162 ?x441 ?x443)) }) +:assumption (forall (?x415 T6) (?x416 T6) (implies (forall (?x417 T5) (iff (= (uf_15 ?x417 ?x415) uf_9) (= (uf_15 ?x417 ?x416) uf_9)) :pat { (uf_18 ?x417) }) (= (uf_153 ?x415 ?x416) uf_9)) :pat { (uf_153 ?x415 ?x416) }) +:assumption (forall (?x418 T6) (?x419 T6) (iff (= (uf_154 ?x418 ?x419) uf_9) (forall (?x420 T5) (implies (= (uf_15 ?x420 ?x418) uf_9) (= (uf_15 ?x420 ?x419) uf_9)) :pat { (uf_15 ?x420 ?x418) } :pat { (uf_15 ?x420 ?x419) })) :pat { (uf_154 ?x418 ?x419) }) +:assumption (forall (?x421 T6) (?x422 T6) (?x423 T5) (iff (= (uf_15 ?x423 (uf_155 ?x421 ?x422)) uf_9) (and (= (uf_15 ?x423 ?x421) uf_9) (= (uf_15 ?x423 ?x422) uf_9))) :pat { (uf_15 ?x423 (uf_155 ?x421 ?x422)) }) +:assumption (forall (?x424 T6) (?x425 T6) (?x426 T5) (iff (= (uf_15 ?x426 (uf_156 ?x424 ?x425)) uf_9) (and (= (uf_15 ?x426 ?x424) uf_9) (not (= (uf_15 ?x426 ?x425) uf_9)))) :pat { (uf_15 ?x426 (uf_156 ?x424 ?x425)) }) +:assumption (forall (?x427 T6) (?x428 T6) (?x429 T5) (iff (= (uf_15 ?x429 (uf_157 ?x427 ?x428)) uf_9) (or (= (uf_15 ?x429 ?x427) uf_9) (= (uf_15 ?x429 ?x428) uf_9))) :pat { (uf_15 ?x429 (uf_157 ?x427 ?x428)) }) +:assumption (forall (?x430 T5) (?x431 T5) (iff (= (uf_15 ?x431 (uf_158 ?x430)) uf_9) (and (= ?x430 ?x431) (not (= (uf_116 ?x430) (uf_116 uf_121))))) :pat { (uf_15 ?x431 (uf_158 ?x430)) }) +:assumption (forall (?x432 T5) (?x433 T5) (iff (= (uf_15 ?x433 (uf_151 ?x432)) uf_9) (= ?x432 ?x433)) :pat { (uf_15 ?x433 (uf_151 ?x432)) }) +:assumption (forall (?x434 T5) (not (= (uf_15 ?x434 uf_152) uf_9)) :pat { (uf_15 ?x434 uf_152) }) +:assumption (forall (?x435 T5) (?x436 T5) (= (uf_159 ?x435 ?x436) (uf_43 (uf_124 (uf_144 (uf_13 ?x435)) (+ (uf_143 (uf_13 ?x435)) (uf_143 (uf_13 ?x436)))) (uf_116 ?x435))) :pat { (uf_159 ?x435 ?x436) }) +:assumption (forall (?x437 T5) (?x438 Int) (= (uf_160 ?x437 ?x438) (uf_43 (uf_124 (uf_144 (uf_13 ?x437)) (- (uf_143 (uf_13 ?x437)) ?x438)) (uf_116 (uf_66 (uf_43 (uf_144 (uf_13 ?x437)) (uf_116 ?x437)) ?x438 (uf_144 (uf_13 ?x437)))))) :pat { (uf_160 ?x437 ?x438) }) +:assumption (forall (?x439 T5) (?x440 Int) (= (uf_161 ?x439 ?x440) (uf_43 (uf_124 (uf_144 (uf_13 ?x439)) ?x440) (uf_116 ?x439))) :pat { (uf_161 ?x439 ?x440) }) +:assumption (forall (?x441 T4) (?x442 T5) (?x443 T5) (iff (= (uf_15 ?x442 (uf_162 ?x441 ?x443)) uf_9) (or (= ?x442 ?x443) (and (= (uf_136 (uf_58 (uf_59 ?x441) ?x442)) uf_9) (= (uf_15 ?x442 (uf_163 ?x443)) uf_9)))) :pat { (uf_15 ?x442 (uf_162 ?x441 ?x443)) }) :assumption (forall (?x444 T4) (implies (= (uf_164 ?x444) uf_9) (up_165 ?x444)) :pat { (uf_164 ?x444) }) :assumption (= (uf_142 uf_166) 0) :assumption (= uf_167 (uf_43 uf_166 uf_168)) -:assumption (forall (?x445 T4) (?x446 T4) (?x447 T5) (?x448 T5) (and true (and (= (uf_170 (uf_169 ?x445 ?x446 ?x447 ?x448) ?x448) (uf_171 ?x445)) (and (= (uf_38 (uf_169 ?x445 ?x446 ?x447 ?x448) ?x448) (uf_38 ?x446 ?x448)) (and (= (uf_25 (uf_169 ?x445 ?x446 ?x447 ?x448) ?x448) uf_26) (and (= (uf_24 (uf_169 ?x445 ?x446 ?x447 ?x448) ?x448) uf_9) (= (uf_41 (uf_169 ?x445 ?x446 ?x447 ?x448)) (uf_172 (uf_41 ?x446) ?x448 (uf_173 ?x446 ?x447 ?x448)))))))) :pat { (uf_169 ?x445 ?x446 ?x447 ?x448) }) -:assumption (forall (?x449 T4) (?x450 T5) (?x451 T5) (implies (not (= (uf_14 (uf_15 ?x450)) uf_16)) (and true (and (= (uf_38 (uf_174 ?x449 ?x450 ?x451) ?x451) (uf_38 ?x449 ?x451)) (and (= (uf_25 (uf_174 ?x449 ?x450 ?x451) ?x451) ?x450) (and (= (uf_24 (uf_174 ?x449 ?x450 ?x451) ?x451) uf_9) (= (uf_41 (uf_174 ?x449 ?x450 ?x451)) (uf_172 (uf_41 ?x449) ?x451 (uf_175 ?x449 ?x450 ?x451)))))))) :pat { (uf_174 ?x449 ?x450 ?x451) }) -:assumption (forall (?x452 T4) (?x453 T5) (?x454 Int) (and (= (uf_177 ?x452 (uf_176 ?x452 ?x453 ?x454)) uf_9) (and (forall (?x455 T5) (<= (uf_170 ?x452 ?x455) (uf_170 (uf_176 ?x452 ?x455 ?x454) ?x455)) :pat { (uf_170 (uf_176 ?x452 ?x455 ?x454) ?x455) }) (and (< (uf_171 ?x452) (uf_171 (uf_176 ?x452 ?x453 ?x454))) (and (= (uf_20 (uf_176 ?x452 ?x453 ?x454)) (uf_178 (uf_20 ?x452) ?x453 ?x454)) (and (= (uf_41 (uf_176 ?x452 ?x453 ?x454)) (uf_41 ?x452)) (= (uf_59 (uf_176 ?x452 ?x453 ?x454)) (uf_59 ?x452))))))) :pat { (uf_176 ?x452 ?x453 ?x454) }) -:assumption (forall (?x456 T4) (implies (= (uf_51 ?x456) uf_9) (forall (?x457 T5) (?x458 T5) (implies (and (= (uf_24 ?x456 ?x458) uf_9) (and (= (uf_13 ?x457 (uf_53 ?x456 ?x458)) uf_9) (= (uf_51 ?x456) uf_9))) (and (not (= (uf_116 ?x457) 0)) (= (uf_24 ?x456 ?x457) uf_9))) :pat { (uf_13 ?x457 (uf_53 ?x456 ?x458)) })) :pat { (uf_51 ?x456) }) -:assumption (forall (?x459 T4) (?x460 T5) (?x461 T3) (implies (and (= (uf_24 ?x459 ?x460) uf_9) (= (uf_44 ?x459) uf_9)) (= (uf_46 ?x459 ?x459 ?x460 ?x461) uf_9)) :pat { (uf_46 ?x459 ?x459 ?x460 ?x461) }) +:assumption (forall (?x445 T4) (?x446 T4) (?x447 T5) (?x448 T5) (and (= (uf_41 (uf_169 ?x445 ?x446 ?x447 ?x448)) (uf_170 (uf_41 ?x446) ?x448 (uf_171 ?x446 ?x447 ?x448))) (and (= (uf_27 (uf_169 ?x445 ?x446 ?x447 ?x448) ?x448) uf_9) (and (= (uf_25 (uf_169 ?x445 ?x446 ?x447 ?x448) ?x448) uf_26) (and (= (uf_38 (uf_169 ?x445 ?x446 ?x447 ?x448) ?x448) (uf_38 ?x446 ?x448)) (and (= (uf_172 (uf_169 ?x445 ?x446 ?x447 ?x448) ?x448) (uf_173 ?x445)) true))))) :pat { (uf_169 ?x445 ?x446 ?x447 ?x448) }) +:assumption (forall (?x449 T4) (?x450 T5) (?x451 T5) (implies (not (= (uf_12 (uf_13 ?x450)) uf_14)) (and (= (uf_41 (uf_174 ?x449 ?x450 ?x451)) (uf_170 (uf_41 ?x449) ?x451 (uf_175 ?x449 ?x450 ?x451))) (and (= (uf_27 (uf_174 ?x449 ?x450 ?x451) ?x451) uf_9) (and (= (uf_25 (uf_174 ?x449 ?x450 ?x451) ?x451) ?x450) (and (= (uf_38 (uf_174 ?x449 ?x450 ?x451) ?x451) (uf_38 ?x449 ?x451)) true))))) :pat { (uf_174 ?x449 ?x450 ?x451) }) +:assumption (forall (?x452 T4) (?x453 T5) (?x454 Int) (and (= (uf_59 (uf_176 ?x452 ?x453 ?x454)) (uf_59 ?x452)) (and (= (uf_41 (uf_176 ?x452 ?x453 ?x454)) (uf_41 ?x452)) (and (= (uf_20 (uf_176 ?x452 ?x453 ?x454)) (uf_177 (uf_20 ?x452) ?x453 ?x454)) (and (< (uf_173 ?x452) (uf_173 (uf_176 ?x452 ?x453 ?x454))) (and (forall (?x455 T5) (<= (uf_172 ?x452 ?x455) (uf_172 (uf_176 ?x452 ?x455 ?x454) ?x455)) :pat { (uf_172 (uf_176 ?x452 ?x455 ?x454) ?x455) }) (= (uf_178 ?x452 (uf_176 ?x452 ?x453 ?x454)) uf_9)))))) :pat { (uf_176 ?x452 ?x453 ?x454) }) +:assumption (forall (?x456 T4) (implies (= (uf_51 ?x456) uf_9) (forall (?x457 T5) (?x458 T5) (implies (and (= (uf_51 ?x456) uf_9) (and (= (uf_15 ?x457 (uf_53 ?x456 ?x458)) uf_9) (= (uf_27 ?x456 ?x458) uf_9))) (and (= (uf_27 ?x456 ?x457) uf_9) (not (= (uf_116 ?x457) 0)))) :pat { (uf_15 ?x457 (uf_53 ?x456 ?x458)) })) :pat { (uf_51 ?x456) }) +:assumption (forall (?x459 T4) (?x460 T5) (?x461 T3) (implies (and (= (uf_44 ?x459) uf_9) (= (uf_27 ?x459 ?x460) uf_9)) (= (uf_46 ?x459 ?x459 ?x460 ?x461) uf_9)) :pat { (uf_46 ?x459 ?x459 ?x460 ?x461) }) :assumption (forall (?x462 T4) (?x463 Int) (?x464 T3) (implies (= (uf_51 ?x462) uf_9) (implies (= (uf_141 ?x464) uf_9) (= (uf_53 ?x462 (uf_43 ?x464 ?x463)) uf_152))) :pat { (uf_53 ?x462 (uf_43 ?x464 ?x463)) (uf_141 ?x464) }) -:assumption (forall (?x465 T4) (?x466 T4) (?x467 T5) (?x468 T3) (implies (and (= (uf_15 ?x467) ?x468) (= (uf_141 ?x468) uf_9)) (and (= (uf_179 ?x465 ?x466 ?x467 ?x468) uf_9) (iff (= (uf_46 ?x465 ?x466 ?x467 ?x468) uf_9) (= (uf_27 ?x466 ?x467) uf_9)))) :pat { (uf_141 ?x468) (uf_46 ?x465 ?x466 ?x467 ?x468) }) -:assumption (forall (?x469 T4) (?x470 T5) (?x471 T5) (implies (and (= (uf_22 (uf_15 ?x470)) uf_9) (and (= (uf_24 ?x469 ?x471) uf_9) (= (uf_51 ?x469) uf_9))) (iff (= (uf_13 ?x470 (uf_53 ?x469 ?x471)) uf_9) (= (uf_25 ?x469 ?x470) ?x471))) :pat { (uf_13 ?x470 (uf_53 ?x469 ?x471)) (uf_22 (uf_15 ?x470)) }) +:assumption (forall (?x465 T4) (?x466 T4) (?x467 T5) (?x468 T3) (implies (and (= (uf_141 ?x468) uf_9) (= (uf_13 ?x467) ?x468)) (and (iff (= (uf_46 ?x465 ?x466 ?x467 ?x468) uf_9) (= (uf_24 ?x466 ?x467) uf_9)) (= (uf_179 ?x465 ?x466 ?x467 ?x468) uf_9))) :pat { (uf_141 ?x468) (uf_46 ?x465 ?x466 ?x467 ?x468) }) +:assumption (forall (?x469 T4) (?x470 T5) (?x471 T5) (implies (and (= (uf_51 ?x469) uf_9) (and (= (uf_27 ?x469 ?x471) uf_9) (= (uf_23 (uf_13 ?x470)) uf_9))) (iff (= (uf_15 ?x470 (uf_53 ?x469 ?x471)) uf_9) (= (uf_25 ?x469 ?x470) ?x471))) :pat { (uf_15 ?x470 (uf_53 ?x469 ?x471)) (uf_23 (uf_13 ?x470)) }) :assumption (forall (?x472 T4) (?x473 T4) (?x474 Int) (?x475 T3) (?x476 T15) (up_182 (uf_19 (uf_20 ?x473) (uf_126 (uf_43 ?x475 ?x474) ?x476))) :pat { (uf_180 ?x475 ?x476) (uf_181 ?x472 ?x473) (uf_19 (uf_20 ?x472) (uf_126 (uf_43 ?x475 ?x474) ?x476)) }) -:assumption (forall (?x477 T4) (?x478 Int) (?x479 T3) (?x480 T15) (implies (and (= (uf_25 ?x477 (uf_43 ?x479 ?x478)) uf_26) (and (= (uf_180 ?x479 ?x480) uf_9) (and (= (uf_24 ?x477 (uf_43 ?x479 ?x478)) uf_9) (= (uf_55 ?x477) uf_9)))) (= (uf_19 (uf_20 ?x477) (uf_126 (uf_43 ?x479 ?x478) ?x480)) (uf_183 (uf_184 ?x477 (uf_43 ?x479 ?x478)) (uf_126 (uf_43 ?x479 ?x478) ?x480)))) :pat { (uf_180 ?x479 ?x480) (uf_19 (uf_20 ?x477) (uf_126 (uf_43 ?x479 ?x478) ?x480)) }) -:assumption (forall (?x481 T4) (?x482 Int) (?x483 T3) (?x484 T15) (?x485 T15) (implies (and (or (= (uf_28 (uf_183 (uf_184 ?x481 (uf_43 ?x483 ?x482)) (uf_126 (uf_43 ?x483 ?x482) ?x484))) uf_26) (= (uf_28 (uf_19 (uf_20 ?x481) (uf_126 (uf_43 ?x483 ?x482) ?x484))) uf_26)) (and (= (uf_24 ?x481 (uf_43 ?x483 ?x482)) uf_9) (and (= (uf_185 ?x483 ?x484 ?x485) uf_9) (= (uf_55 ?x481) uf_9)))) (= (uf_19 (uf_20 ?x481) (uf_126 (uf_43 ?x483 ?x482) ?x485)) (uf_183 (uf_184 ?x481 (uf_43 ?x483 ?x482)) (uf_126 (uf_43 ?x483 ?x482) ?x485)))) :pat { (uf_185 ?x483 ?x484 ?x485) (uf_19 (uf_20 ?x481) (uf_126 (uf_43 ?x483 ?x482) ?x485)) }) +:assumption (forall (?x477 T4) (?x478 Int) (?x479 T3) (?x480 T15) (implies (and (= (uf_55 ?x477) uf_9) (and (= (uf_27 ?x477 (uf_43 ?x479 ?x478)) uf_9) (and (= (uf_180 ?x479 ?x480) uf_9) (= (uf_25 ?x477 (uf_43 ?x479 ?x478)) uf_26)))) (= (uf_19 (uf_20 ?x477) (uf_126 (uf_43 ?x479 ?x478) ?x480)) (uf_183 (uf_184 ?x477 (uf_43 ?x479 ?x478)) (uf_126 (uf_43 ?x479 ?x478) ?x480)))) :pat { (uf_180 ?x479 ?x480) (uf_19 (uf_20 ?x477) (uf_126 (uf_43 ?x479 ?x478) ?x480)) }) +:assumption (forall (?x481 T4) (?x482 Int) (?x483 T3) (?x484 T15) (?x485 T15) (implies (and (= (uf_55 ?x481) uf_9) (and (= (uf_185 ?x483 ?x484 ?x485) uf_9) (and (= (uf_27 ?x481 (uf_43 ?x483 ?x482)) uf_9) (or (= (uf_28 (uf_19 (uf_20 ?x481) (uf_126 (uf_43 ?x483 ?x482) ?x484))) uf_26) (= (uf_28 (uf_183 (uf_184 ?x481 (uf_43 ?x483 ?x482)) (uf_126 (uf_43 ?x483 ?x482) ?x484))) uf_26))))) (= (uf_19 (uf_20 ?x481) (uf_126 (uf_43 ?x483 ?x482) ?x485)) (uf_183 (uf_184 ?x481 (uf_43 ?x483 ?x482)) (uf_126 (uf_43 ?x483 ?x482) ?x485)))) :pat { (uf_185 ?x483 ?x484 ?x485) (uf_19 (uf_20 ?x481) (uf_126 (uf_43 ?x483 ?x482) ?x485)) }) :assumption (forall (?x486 T4) (?x487 T5) (= (uf_184 ?x486 ?x487) (uf_30 (uf_19 (uf_20 ?x486) ?x487))) :pat { (uf_184 ?x486 ?x487) }) -:assumption (forall (?x488 T4) (?x489 T5) (?x490 T5) (?x491 T15) (?x492 Int) (?x493 Int) (?x494 T3) (implies (and (< ?x492 ?x493) (and (<= 0 ?x492) (and (= (uf_187 ?x491 ?x493) uf_9) (and (= (uf_186 ?x489 ?x490) uf_9) (and (= (uf_24 ?x488 ?x490) uf_9) (= (uf_51 ?x488) uf_9)))))) (= (uf_19 (uf_20 ?x488) (uf_66 (uf_126 ?x489 ?x491) ?x492 ?x494)) (uf_11 (uf_189 ?x490) (uf_66 (uf_126 ?x489 ?x491) ?x492 ?x494)))) :pat { (uf_49 ?x488 ?x490) (uf_186 ?x489 ?x490) (uf_19 (uf_20 ?x488) (uf_66 (uf_126 ?x489 ?x491) ?x492 ?x494)) (uf_187 ?x491 ?x493) } :pat { (uf_188 ?x488 ?x490 ?x489 (uf_66 (uf_126 ?x489 ?x491) ?x492 ?x494)) (uf_187 ?x491 ?x493) }) -:assumption (forall (?x495 T4) (?x496 T5) (?x497 T5) (?x498 T15) (implies (and (= (uf_190 ?x498) uf_9) (and (= (uf_186 ?x496 ?x497) uf_9) (and (= (uf_24 ?x495 ?x497) uf_9) (= (uf_51 ?x495) uf_9)))) (and (= (uf_19 (uf_20 ?x495) (uf_126 ?x496 ?x498)) (uf_11 (uf_189 ?x497) (uf_126 ?x496 ?x498))) (= (uf_186 ?x496 ?x497) uf_9))) :pat { (uf_186 ?x496 ?x497) (uf_19 (uf_20 ?x495) (uf_126 ?x496 ?x498)) } :pat { (uf_188 ?x495 ?x497 ?x496 (uf_126 ?x496 ?x498)) }) +:assumption (forall (?x488 T4) (?x489 T5) (?x490 T5) (?x491 T15) (?x492 Int) (?x493 Int) (?x494 T3) (implies (and (= (uf_51 ?x488) uf_9) (and (= (uf_27 ?x488 ?x490) uf_9) (and (= (uf_186 ?x489 ?x490) uf_9) (and (= (uf_187 ?x491 ?x493) uf_9) (and (<= 0 ?x492) (< ?x492 ?x493)))))) (= (uf_19 (uf_20 ?x488) (uf_66 (uf_126 ?x489 ?x491) ?x492 ?x494)) (uf_10 (uf_189 ?x490) (uf_66 (uf_126 ?x489 ?x491) ?x492 ?x494)))) :pat { (uf_49 ?x488 ?x490) (uf_186 ?x489 ?x490) (uf_19 (uf_20 ?x488) (uf_66 (uf_126 ?x489 ?x491) ?x492 ?x494)) (uf_187 ?x491 ?x493) } :pat { (uf_188 ?x488 ?x490 ?x489 (uf_66 (uf_126 ?x489 ?x491) ?x492 ?x494)) (uf_187 ?x491 ?x493) }) +:assumption (forall (?x495 T4) (?x496 T5) (?x497 T5) (?x498 T15) (implies (and (= (uf_51 ?x495) uf_9) (and (= (uf_27 ?x495 ?x497) uf_9) (and (= (uf_186 ?x496 ?x497) uf_9) (= (uf_190 ?x498) uf_9)))) (and (= (uf_186 ?x496 ?x497) uf_9) (= (uf_19 (uf_20 ?x495) (uf_126 ?x496 ?x498)) (uf_10 (uf_189 ?x497) (uf_126 ?x496 ?x498))))) :pat { (uf_186 ?x496 ?x497) (uf_19 (uf_20 ?x495) (uf_126 ?x496 ?x498)) } :pat { (uf_188 ?x495 ?x497 ?x496 (uf_126 ?x496 ?x498)) }) :assumption (forall (?x499 T4) (?x500 T5) (?x501 T5) (?x502 T5) (= (uf_188 ?x499 ?x500 ?x501 ?x502) ?x502) :pat { (uf_188 ?x499 ?x500 ?x501 ?x502) }) -:assumption (forall (?x503 T5) (?x504 T5) (implies (forall (?x505 T4) (implies (= (uf_49 ?x505 ?x504) uf_9) (= (uf_24 ?x505 ?x503) uf_9)) :pat { (uf_191 ?x505) }) (= (uf_186 ?x503 ?x504) uf_9)) :pat { (uf_186 ?x503 ?x504) }) -:assumption (forall (?x506 T5) (?x507 T4) (?x508 T4) (?x509 T5) (up_193 (uf_13 ?x509 (uf_192 (uf_12 ?x508 ?x506)))) :pat { (uf_13 ?x509 (uf_192 (uf_12 ?x507 ?x506))) (uf_177 ?x507 ?x508) }) -:assumption (forall (?x510 T5) (?x511 T4) (?x512 T4) (?x513 T5) (up_193 (uf_13 ?x513 (uf_10 ?x512 ?x510))) :pat { (uf_13 ?x513 (uf_10 ?x511 ?x510)) (uf_177 ?x511 ?x512) }) -:assumption (forall (?x514 T4) (?x515 T5) (?x516 T15) (?x517 Int) (?x518 Int) (?x519 T3) (implies (and (< ?x518 ?x517) (and (<= 0 ?x518) (and (= (uf_194 ?x516 ?x517 ?x519) uf_9) (= (uf_51 ?x514) uf_9)))) (= (uf_135 (uf_58 (uf_59 ?x514) (uf_66 (uf_126 ?x515 ?x516) ?x518 ?x519))) uf_9)) :pat { (uf_194 ?x516 ?x517 ?x519) (uf_135 (uf_58 (uf_59 ?x514) (uf_66 (uf_126 ?x515 ?x516) ?x518 ?x519))) }) -:assumption (forall (?x520 T4) (?x521 Int) (?x522 T5) (?x523 Int) (?x524 Int) (?x525 T3) (implies (and (< ?x524 ?x523) (and (<= 0 ?x524) (and (= (uf_13 (uf_43 (uf_124 ?x525 ?x523) ?x521) (uf_10 ?x520 ?x522)) uf_9) (and (= (uf_23 ?x525) uf_9) (= (uf_55 ?x520) uf_9))))) (= (uf_19 (uf_20 ?x520) (uf_66 (uf_43 ?x525 ?x521) ?x524 ?x525)) (uf_11 (uf_12 ?x520 ?x522) (uf_66 (uf_43 ?x525 ?x521) ?x524 ?x525)))) :pat { (uf_13 (uf_43 (uf_124 ?x525 ?x523) ?x521) (uf_10 ?x520 ?x522)) (uf_19 (uf_20 ?x520) (uf_66 (uf_43 ?x525 ?x521) ?x524 ?x525)) (uf_23 ?x525) }) -:assumption (forall (?x526 T4) (?x527 Int) (?x528 T5) (?x529 Int) (?x530 Int) (?x531 T3) (implies (and (< ?x530 ?x529) (and (<= 0 ?x530) (and (= (uf_13 (uf_43 (uf_124 ?x531 ?x529) ?x527) (uf_10 ?x526 ?x528)) uf_9) (and (= (uf_23 ?x531) uf_9) (= (uf_55 ?x526) uf_9))))) (and (not (= (uf_135 (uf_58 (uf_59 ?x526) (uf_66 (uf_43 ?x531 ?x527) ?x530 ?x531))) uf_9)) (= (uf_27 ?x526 (uf_66 (uf_43 ?x531 ?x527) ?x530 ?x531)) uf_9))) :pat { (uf_13 (uf_43 (uf_124 ?x531 ?x529) ?x527) (uf_10 ?x526 ?x528)) (uf_58 (uf_59 ?x526) (uf_66 (uf_43 ?x531 ?x527) ?x530 ?x531)) (uf_23 ?x531) } :pat { (uf_13 (uf_43 (uf_124 ?x531 ?x529) ?x527) (uf_10 ?x526 ?x528)) (uf_25 ?x526 (uf_66 (uf_43 ?x531 ?x527) ?x530 ?x531)) (uf_23 ?x531) }) -:assumption (forall (?x532 T4) (?x533 T5) (?x534 T5) (?x535 T15) (?x536 Int) (?x537 Int) (?x538 T3) (implies (and (< ?x537 ?x536) (and (<= 0 ?x537) (and (= (uf_187 ?x535 ?x536) uf_9) (and (= (uf_13 ?x533 (uf_10 ?x532 ?x534)) uf_9) (= (uf_55 ?x532) uf_9))))) (and (not (= (uf_135 (uf_58 (uf_59 ?x532) (uf_66 (uf_126 ?x533 ?x535) ?x537 ?x538))) uf_9)) (= (uf_27 ?x532 (uf_66 (uf_126 ?x533 ?x535) ?x537 ?x538)) uf_9))) :pat { (uf_13 ?x533 (uf_10 ?x532 ?x534)) (uf_187 ?x535 ?x536) (uf_58 (uf_59 ?x532) (uf_66 (uf_126 ?x533 ?x535) ?x537 ?x538)) } :pat { (uf_13 ?x533 (uf_10 ?x532 ?x534)) (uf_187 ?x535 ?x536) (uf_25 ?x532 (uf_66 (uf_126 ?x533 ?x535) ?x537 ?x538)) }) -:assumption (forall (?x539 T4) (?x540 T5) (?x541 T5) (?x542 T15) (?x543 Int) (?x544 Int) (?x545 T3) (implies (and (< ?x544 ?x543) (and (<= 0 ?x544) (and (= (uf_187 ?x542 ?x543) uf_9) (and (= (uf_13 ?x540 (uf_10 ?x539 ?x541)) uf_9) (= (uf_55 ?x539) uf_9))))) (= (uf_19 (uf_20 ?x539) (uf_66 (uf_126 ?x540 ?x542) ?x544 ?x545)) (uf_11 (uf_12 ?x539 ?x541) (uf_66 (uf_126 ?x540 ?x542) ?x544 ?x545)))) :pat { (uf_13 ?x540 (uf_10 ?x539 ?x541)) (uf_187 ?x542 ?x543) (uf_19 (uf_20 ?x539) (uf_66 (uf_126 ?x540 ?x542) ?x544 ?x545)) }) -:assumption (forall (?x546 T4) (?x547 T5) (?x548 T5) (?x549 T15) (implies (and (= (uf_190 ?x549) uf_9) (and (= (uf_13 ?x547 (uf_10 ?x546 ?x548)) uf_9) (= (uf_55 ?x546) uf_9))) (and (not (= (uf_135 (uf_58 (uf_59 ?x546) (uf_126 ?x547 ?x549))) uf_9)) (= (uf_27 ?x546 (uf_126 ?x547 ?x549)) uf_9))) :pat { (uf_13 ?x547 (uf_10 ?x546 ?x548)) (uf_190 ?x549) (uf_25 ?x546 (uf_126 ?x547 ?x549)) } :pat { (uf_13 ?x547 (uf_10 ?x546 ?x548)) (uf_190 ?x549) (uf_58 (uf_59 ?x546) (uf_126 ?x547 ?x549)) }) -:assumption (forall (?x550 T4) (?x551 T5) (?x552 T5) (implies (and (= (uf_13 ?x551 (uf_10 ?x550 ?x552)) uf_9) (= (uf_55 ?x550) uf_9)) (and (not (= (uf_135 (uf_58 (uf_59 ?x550) ?x551)) uf_9)) (= (uf_27 ?x550 ?x551) uf_9))) :pat { (uf_55 ?x550) (uf_13 ?x551 (uf_10 ?x550 ?x552)) (uf_40 (uf_41 ?x550) ?x551) } :pat { (uf_55 ?x550) (uf_13 ?x551 (uf_10 ?x550 ?x552)) (uf_58 (uf_59 ?x550) ?x551) }) -:assumption (forall (?x553 T4) (?x554 T5) (?x555 T5) (?x556 T15) (implies (and (= (uf_190 ?x556) uf_9) (= (uf_13 ?x554 (uf_10 ?x553 ?x555)) uf_9)) (= (uf_19 (uf_20 ?x553) (uf_126 ?x554 ?x556)) (uf_11 (uf_12 ?x553 ?x555) (uf_126 ?x554 ?x556)))) :pat { (uf_13 ?x554 (uf_10 ?x553 ?x555)) (uf_190 ?x556) (uf_19 (uf_20 ?x553) (uf_126 ?x554 ?x556)) }) +:assumption (forall (?x503 T5) (?x504 T5) (implies (forall (?x505 T4) (implies (= (uf_49 ?x505 ?x504) uf_9) (= (uf_27 ?x505 ?x503) uf_9)) :pat { (uf_191 ?x505) }) (= (uf_186 ?x503 ?x504) uf_9)) :pat { (uf_186 ?x503 ?x504) }) +:assumption (forall (?x506 T5) (?x507 T4) (?x508 T4) (?x509 T5) (up_193 (uf_15 ?x509 (uf_192 (uf_11 ?x508 ?x506)))) :pat { (uf_15 ?x509 (uf_192 (uf_11 ?x507 ?x506))) (uf_178 ?x507 ?x508) }) +:assumption (forall (?x510 T5) (?x511 T4) (?x512 T4) (?x513 T5) (up_193 (uf_15 ?x513 (uf_16 ?x512 ?x510))) :pat { (uf_15 ?x513 (uf_16 ?x511 ?x510)) (uf_178 ?x511 ?x512) }) +:assumption (forall (?x514 T4) (?x515 T5) (?x516 T15) (?x517 Int) (?x518 Int) (?x519 T3) (implies (and (= (uf_51 ?x514) uf_9) (and (= (uf_194 ?x516 ?x517 ?x519) uf_9) (and (<= 0 ?x518) (< ?x518 ?x517)))) (= (uf_136 (uf_58 (uf_59 ?x514) (uf_66 (uf_126 ?x515 ?x516) ?x518 ?x519))) uf_9)) :pat { (uf_194 ?x516 ?x517 ?x519) (uf_136 (uf_58 (uf_59 ?x514) (uf_66 (uf_126 ?x515 ?x516) ?x518 ?x519))) }) +:assumption (forall (?x520 T4) (?x521 Int) (?x522 T5) (?x523 Int) (?x524 Int) (?x525 T3) (implies (and (= (uf_55 ?x520) uf_9) (and (= (uf_22 ?x525) uf_9) (and (= (uf_15 (uf_43 (uf_124 ?x525 ?x523) ?x521) (uf_16 ?x520 ?x522)) uf_9) (and (<= 0 ?x524) (< ?x524 ?x523))))) (= (uf_19 (uf_20 ?x520) (uf_66 (uf_43 ?x525 ?x521) ?x524 ?x525)) (uf_10 (uf_11 ?x520 ?x522) (uf_66 (uf_43 ?x525 ?x521) ?x524 ?x525)))) :pat { (uf_15 (uf_43 (uf_124 ?x525 ?x523) ?x521) (uf_16 ?x520 ?x522)) (uf_19 (uf_20 ?x520) (uf_66 (uf_43 ?x525 ?x521) ?x524 ?x525)) (uf_22 ?x525) }) +:assumption (forall (?x526 T4) (?x527 Int) (?x528 T5) (?x529 Int) (?x530 Int) (?x531 T3) (implies (and (= (uf_55 ?x526) uf_9) (and (= (uf_22 ?x531) uf_9) (and (= (uf_15 (uf_43 (uf_124 ?x531 ?x529) ?x527) (uf_16 ?x526 ?x528)) uf_9) (and (<= 0 ?x530) (< ?x530 ?x529))))) (and (= (uf_24 ?x526 (uf_66 (uf_43 ?x531 ?x527) ?x530 ?x531)) uf_9) (not (= (uf_136 (uf_58 (uf_59 ?x526) (uf_66 (uf_43 ?x531 ?x527) ?x530 ?x531))) uf_9)))) :pat { (uf_15 (uf_43 (uf_124 ?x531 ?x529) ?x527) (uf_16 ?x526 ?x528)) (uf_58 (uf_59 ?x526) (uf_66 (uf_43 ?x531 ?x527) ?x530 ?x531)) (uf_22 ?x531) } :pat { (uf_15 (uf_43 (uf_124 ?x531 ?x529) ?x527) (uf_16 ?x526 ?x528)) (uf_25 ?x526 (uf_66 (uf_43 ?x531 ?x527) ?x530 ?x531)) (uf_22 ?x531) }) +:assumption (forall (?x532 T4) (?x533 T5) (?x534 T5) (?x535 T15) (?x536 Int) (?x537 Int) (?x538 T3) (implies (and (= (uf_55 ?x532) uf_9) (and (= (uf_15 ?x533 (uf_16 ?x532 ?x534)) uf_9) (and (= (uf_187 ?x535 ?x536) uf_9) (and (<= 0 ?x537) (< ?x537 ?x536))))) (and (= (uf_24 ?x532 (uf_66 (uf_126 ?x533 ?x535) ?x537 ?x538)) uf_9) (not (= (uf_136 (uf_58 (uf_59 ?x532) (uf_66 (uf_126 ?x533 ?x535) ?x537 ?x538))) uf_9)))) :pat { (uf_15 ?x533 (uf_16 ?x532 ?x534)) (uf_187 ?x535 ?x536) (uf_58 (uf_59 ?x532) (uf_66 (uf_126 ?x533 ?x535) ?x537 ?x538)) } :pat { (uf_15 ?x533 (uf_16 ?x532 ?x534)) (uf_187 ?x535 ?x536) (uf_25 ?x532 (uf_66 (uf_126 ?x533 ?x535) ?x537 ?x538)) }) +:assumption (forall (?x539 T4) (?x540 T5) (?x541 T5) (?x542 T15) (?x543 Int) (?x544 Int) (?x545 T3) (implies (and (= (uf_55 ?x539) uf_9) (and (= (uf_15 ?x540 (uf_16 ?x539 ?x541)) uf_9) (and (= (uf_187 ?x542 ?x543) uf_9) (and (<= 0 ?x544) (< ?x544 ?x543))))) (= (uf_19 (uf_20 ?x539) (uf_66 (uf_126 ?x540 ?x542) ?x544 ?x545)) (uf_10 (uf_11 ?x539 ?x541) (uf_66 (uf_126 ?x540 ?x542) ?x544 ?x545)))) :pat { (uf_15 ?x540 (uf_16 ?x539 ?x541)) (uf_187 ?x542 ?x543) (uf_19 (uf_20 ?x539) (uf_66 (uf_126 ?x540 ?x542) ?x544 ?x545)) }) +:assumption (forall (?x546 T4) (?x547 T5) (?x548 T5) (?x549 T15) (implies (and (= (uf_55 ?x546) uf_9) (and (= (uf_15 ?x547 (uf_16 ?x546 ?x548)) uf_9) (= (uf_190 ?x549) uf_9))) (and (= (uf_24 ?x546 (uf_126 ?x547 ?x549)) uf_9) (not (= (uf_136 (uf_58 (uf_59 ?x546) (uf_126 ?x547 ?x549))) uf_9)))) :pat { (uf_15 ?x547 (uf_16 ?x546 ?x548)) (uf_190 ?x549) (uf_25 ?x546 (uf_126 ?x547 ?x549)) } :pat { (uf_15 ?x547 (uf_16 ?x546 ?x548)) (uf_190 ?x549) (uf_58 (uf_59 ?x546) (uf_126 ?x547 ?x549)) }) +:assumption (forall (?x550 T4) (?x551 T5) (?x552 T5) (implies (and (= (uf_55 ?x550) uf_9) (= (uf_15 ?x551 (uf_16 ?x550 ?x552)) uf_9)) (and (= (uf_24 ?x550 ?x551) uf_9) (not (= (uf_136 (uf_58 (uf_59 ?x550) ?x551)) uf_9)))) :pat { (uf_55 ?x550) (uf_15 ?x551 (uf_16 ?x550 ?x552)) (uf_40 (uf_41 ?x550) ?x551) } :pat { (uf_55 ?x550) (uf_15 ?x551 (uf_16 ?x550 ?x552)) (uf_58 (uf_59 ?x550) ?x551) }) +:assumption (forall (?x553 T4) (?x554 T5) (?x555 T5) (?x556 T15) (implies (and (= (uf_15 ?x554 (uf_16 ?x553 ?x555)) uf_9) (= (uf_190 ?x556) uf_9)) (= (uf_19 (uf_20 ?x553) (uf_126 ?x554 ?x556)) (uf_10 (uf_11 ?x553 ?x555) (uf_126 ?x554 ?x556)))) :pat { (uf_15 ?x554 (uf_16 ?x553 ?x555)) (uf_190 ?x556) (uf_19 (uf_20 ?x553) (uf_126 ?x554 ?x556)) }) :assumption (forall (?x557 T4) (?x558 T5) (?x559 T5) (implies (= (uf_195 ?x557 ?x558 ?x559) uf_9) (= (uf_196 ?x557 ?x558 ?x559) uf_9)) :pat { (uf_195 ?x557 ?x558 ?x559) }) -:assumption (forall (?x560 T4) (?x561 T5) (?x562 T5) (?x563 T5) (implies (and (forall (?x564 T4) (implies (and (= (uf_10 ?x564 ?x561) (uf_10 ?x560 ?x561)) (and (= (uf_12 ?x564 ?x561) (uf_12 ?x560 ?x561)) (= (uf_46 ?x564 ?x564 ?x562 (uf_15 ?x562)) uf_9))) (= (uf_145 ?x563 (uf_53 ?x564 ?x562)) uf_9))) (and (= (uf_13 ?x562 (uf_10 ?x560 ?x561)) uf_9) (up_197 (uf_15 ?x562)))) (and (= (uf_145 ?x563 (uf_53 ?x560 ?x562)) uf_9) (= (uf_195 ?x560 ?x563 ?x561) uf_9))) :pat { (uf_13 ?x562 (uf_10 ?x560 ?x561)) (uf_195 ?x560 ?x563 ?x561) }) -:assumption (forall (?x565 T4) (?x566 T5) (?x567 T5) (?x568 T5) (implies (and (= (uf_145 ?x568 (uf_53 ?x565 ?x567)) uf_9) (and (= (uf_13 ?x567 (uf_10 ?x565 ?x566)) uf_9) (not (up_197 (uf_15 ?x567))))) (and (= (uf_145 ?x568 (uf_53 ?x565 ?x567)) uf_9) (= (uf_196 ?x565 ?x568 ?x566) uf_9))) :pat { (uf_13 ?x567 (uf_10 ?x565 ?x566)) (uf_196 ?x565 ?x568 ?x566) }) -:assumption (forall (?x569 T4) (?x570 T5) (?x571 T5) (implies (and (= (uf_13 ?x571 (uf_10 ?x569 ?x570)) uf_9) (= (uf_55 ?x569) uf_9)) (= (uf_196 ?x569 ?x571 ?x570) uf_9)) :pat { (uf_196 ?x569 ?x571 ?x570) }) -:assumption (forall (?x572 T4) (?x573 T5) (implies (and (= (uf_22 (uf_15 ?x573)) uf_9) (and (not (= (uf_14 (uf_15 ?x573)) uf_16)) (and (= (uf_27 ?x572 ?x573) uf_9) (and (= (uf_48 ?x573 (uf_15 ?x573)) uf_9) (and (= (uf_25 ?x572 ?x573) uf_26) (and (= (uf_24 ?x572 ?x573) uf_9) (= (uf_55 ?x572) uf_9))))))) (= (uf_196 ?x572 ?x573 ?x573) uf_9)) :pat { (uf_196 ?x572 ?x573 ?x573) }) -:assumption (forall (?x574 T4) (?x575 T5) (?x576 T5) (implies (= (uf_196 ?x574 ?x575 ?x576) uf_9) (and (forall (?x577 T5) (implies (and (= (uf_13 ?x577 (uf_53 ?x574 ?x575)) uf_9) (not (up_197 (uf_15 ?x575)))) (= (uf_147 ?x577 (uf_192 (uf_12 ?x574 ?x576))) uf_9)) :pat { (uf_13 ?x577 (uf_53 ?x574 ?x575)) }) (and (= (uf_24 ?x574 ?x575) uf_9) (= (uf_13 ?x575 (uf_10 ?x574 ?x576)) uf_9)))) :pat { (uf_196 ?x574 ?x575 ?x576) }) +:assumption (forall (?x560 T4) (?x561 T5) (?x562 T5) (?x563 T5) (implies (and (up_197 (uf_13 ?x562)) (and (= (uf_15 ?x562 (uf_16 ?x560 ?x561)) uf_9) (forall (?x564 T4) (implies (and (= (uf_46 ?x564 ?x564 ?x562 (uf_13 ?x562)) uf_9) (and (= (uf_11 ?x564 ?x561) (uf_11 ?x560 ?x561)) (= (uf_16 ?x564 ?x561) (uf_16 ?x560 ?x561)))) (= (uf_145 ?x563 (uf_53 ?x564 ?x562)) uf_9))))) (and (= (uf_195 ?x560 ?x563 ?x561) uf_9) (= (uf_145 ?x563 (uf_53 ?x560 ?x562)) uf_9))) :pat { (uf_15 ?x562 (uf_16 ?x560 ?x561)) (uf_195 ?x560 ?x563 ?x561) }) +:assumption (forall (?x565 T4) (?x566 T5) (?x567 T5) (?x568 T5) (implies (and (not (up_197 (uf_13 ?x567))) (and (= (uf_15 ?x567 (uf_16 ?x565 ?x566)) uf_9) (= (uf_145 ?x568 (uf_53 ?x565 ?x567)) uf_9))) (and (= (uf_196 ?x565 ?x568 ?x566) uf_9) (= (uf_145 ?x568 (uf_53 ?x565 ?x567)) uf_9))) :pat { (uf_15 ?x567 (uf_16 ?x565 ?x566)) (uf_196 ?x565 ?x568 ?x566) }) +:assumption (forall (?x569 T4) (?x570 T5) (?x571 T5) (implies (and (= (uf_55 ?x569) uf_9) (= (uf_15 ?x571 (uf_16 ?x569 ?x570)) uf_9)) (= (uf_196 ?x569 ?x571 ?x570) uf_9)) :pat { (uf_196 ?x569 ?x571 ?x570) }) +:assumption (forall (?x572 T4) (?x573 T5) (implies (and (= (uf_55 ?x572) uf_9) (and (= (uf_27 ?x572 ?x573) uf_9) (and (= (uf_25 ?x572 ?x573) uf_26) (and (= (uf_48 ?x573 (uf_13 ?x573)) uf_9) (and (= (uf_24 ?x572 ?x573) uf_9) (and (not (= (uf_12 (uf_13 ?x573)) uf_14)) (= (uf_23 (uf_13 ?x573)) uf_9))))))) (= (uf_196 ?x572 ?x573 ?x573) uf_9)) :pat { (uf_196 ?x572 ?x573 ?x573) }) +:assumption (forall (?x574 T4) (?x575 T5) (?x576 T5) (implies (= (uf_196 ?x574 ?x575 ?x576) uf_9) (and (= (uf_15 ?x575 (uf_16 ?x574 ?x576)) uf_9) (and (= (uf_27 ?x574 ?x575) uf_9) (forall (?x577 T5) (implies (and (not (up_197 (uf_13 ?x575))) (= (uf_15 ?x577 (uf_53 ?x574 ?x575)) uf_9)) (= (uf_147 ?x577 (uf_192 (uf_11 ?x574 ?x576))) uf_9)) :pat { (uf_15 ?x577 (uf_53 ?x574 ?x575)) })))) :pat { (uf_196 ?x574 ?x575 ?x576) }) :assumption (forall (?x578 T4) (?x579 T5) (?x580 T5) (?x581 T16) (iff (= (uf_198 ?x578 ?x579 ?x580 ?x581) uf_9) (= (uf_195 ?x578 ?x579 ?x580) uf_9)) :pat { (uf_198 ?x578 ?x579 ?x580 ?x581) }) :assumption (forall (?x582 T4) (?x583 T5) (?x584 T5) (?x585 T16) (implies (= (uf_198 ?x582 ?x583 ?x584 ?x585) uf_9) (up_199 ?x582 ?x583 ?x585)) :pat { (uf_198 ?x582 ?x583 ?x584 ?x585) }) :assumption (forall (?x586 T4) (?x587 T5) (?x588 T5) (?x589 T16) (iff (= (uf_200 ?x586 ?x587 ?x588 ?x589) uf_9) (= (uf_196 ?x586 ?x587 ?x588) uf_9)) :pat { (uf_200 ?x586 ?x587 ?x588 ?x589) }) :assumption (forall (?x590 T4) (?x591 T5) (?x592 T5) (?x593 T16) (implies (= (uf_200 ?x590 ?x591 ?x592 ?x593) uf_9) (up_199 ?x590 ?x591 ?x593)) :pat { (uf_200 ?x590 ?x591 ?x592 ?x593) }) -:assumption (forall (?x594 T4) (?x595 T5) (= (uf_10 ?x594 ?x595) (uf_192 (uf_12 ?x594 ?x595))) :pat { (uf_10 ?x594 ?x595) }) -:assumption (forall (?x596 T4) (?x597 T5) (= (uf_12 ?x596 ?x597) (uf_32 (uf_19 (uf_20 ?x596) ?x597))) :pat { (uf_12 ?x596 ?x597) }) +:assumption (forall (?x594 T4) (?x595 T5) (= (uf_16 ?x594 ?x595) (uf_192 (uf_11 ?x594 ?x595))) :pat { (uf_16 ?x594 ?x595) }) +:assumption (forall (?x596 T4) (?x597 T5) (= (uf_11 ?x596 ?x597) (uf_32 (uf_19 (uf_20 ?x596) ?x597))) :pat { (uf_11 ?x596 ?x597) }) :assumption (forall (?x598 T4) (?x599 Int) (?x600 T3) (= (uf_43 ?x600 (uf_19 (uf_20 ?x598) (uf_43 (uf_6 ?x600) ?x599))) (uf_201 ?x598 (uf_43 (uf_6 ?x600) ?x599) ?x600)) :pat { (uf_43 ?x600 (uf_19 (uf_20 ?x598) (uf_43 (uf_6 ?x600) ?x599))) }) :assumption (forall (?x601 T1) (?x602 T4) (implies (= (uf_202 ?x601 ?x602) uf_9) (= (uf_51 ?x602) uf_9)) :pat { (uf_202 ?x601 ?x602) }) :assumption (forall (?x603 T4) (implies (= (uf_44 ?x603) uf_9) (= (uf_51 ?x603) uf_9)) :pat { (uf_44 ?x603) }) -:assumption (forall (?x604 T4) (implies (= (uf_55 ?x604) uf_9) (and (= (uf_44 ?x604) uf_9) (= (uf_51 ?x604) uf_9))) :pat { (uf_55 ?x604) }) -:assumption (forall (?x605 T4) (implies (= (uf_203 ?x605) uf_9) (and (<= 0 (uf_171 ?x605)) (= (uf_55 ?x605) uf_9))) :pat { (uf_203 ?x605) }) -:assumption (forall (?x606 T3) (implies (= (uf_23 ?x606) uf_9) (forall (?x607 T4) (?x608 Int) (?x609 T5) (iff (= (uf_13 ?x609 (uf_128 ?x607 (uf_43 ?x606 ?x608))) uf_9) (= ?x609 (uf_43 ?x606 ?x608))) :pat { (uf_13 ?x609 (uf_128 ?x607 (uf_43 ?x606 ?x608))) })) :pat { (uf_23 ?x606) }) -:assumption (forall (?x610 T3) (implies (= (uf_23 ?x610) uf_9) (forall (?x611 Int) (?x612 T5) (iff (= (uf_13 ?x612 (uf_130 (uf_43 ?x610 ?x611))) uf_9) (= ?x612 (uf_43 ?x610 ?x611))) :pat { (uf_13 ?x612 (uf_130 (uf_43 ?x610 ?x611))) })) :pat { (uf_23 ?x610) }) -:assumption (forall (?x613 T4) (?x614 T4) (?x615 T5) (?x616 T3) (iff (= (uf_204 ?x613 ?x614 ?x615 ?x616) uf_9) (and (up_205 ?x613 ?x614 ?x615 ?x616) (and (= (uf_58 (uf_59 ?x613) ?x615) (uf_58 (uf_59 ?x614) ?x615)) (= (uf_12 ?x613 ?x615) (uf_12 ?x614 ?x615))))) :pat { (uf_204 ?x613 ?x614 ?x615 ?x616) }) -:assumption (forall (?x617 T4) (?x618 T4) (?x619 T5) (?x620 T3) (iff (= (uf_206 ?x617 ?x618 ?x619 ?x620) uf_9) (and (= (uf_123 ?x617 ?x618 ?x619 ?x620) uf_9) (and (= (uf_58 (uf_59 ?x617) ?x619) (uf_58 (uf_59 ?x618) ?x619)) (and (= (uf_53 ?x617 ?x619) (uf_53 ?x618 ?x619)) (= (uf_12 ?x617 ?x619) (uf_12 ?x618 ?x619)))))) :pat { (uf_206 ?x617 ?x618 ?x619 ?x620) }) -:assumption (forall (?x621 T4) (?x622 T4) (?x623 T5) (?x624 T5) (iff (= (uf_207 ?x621 ?x622 ?x623 ?x624) uf_9) (or (= (uf_208 (uf_15 ?x623)) uf_9) (or (and (= (uf_204 ?x621 ?x622 ?x623 (uf_15 ?x623)) uf_9) (= (uf_46 ?x621 ?x622 ?x623 (uf_15 ?x623)) uf_9)) (or (and (not (= (uf_24 ?x622 ?x623) uf_9)) (not (= (uf_24 ?x621 ?x623) uf_9))) (= (uf_206 ?x621 ?x622 ?x624 (uf_15 ?x624)) uf_9))))) :pat { (uf_207 ?x621 ?x622 ?x623 ?x624) }) -:assumption (forall (?x625 T4) (?x626 T4) (?x627 T5) (?x628 T3) (iff (= (uf_179 ?x625 ?x626 ?x627 ?x628) uf_9) (implies (and (= (uf_24 ?x626 ?x627) uf_9) (= (uf_24 ?x625 ?x627) uf_9)) (= (uf_206 ?x625 ?x626 ?x627 ?x628) uf_9))) :pat { (uf_179 ?x625 ?x626 ?x627 ?x628) }) +:assumption (forall (?x604 T4) (implies (= (uf_55 ?x604) uf_9) (and (= (uf_51 ?x604) uf_9) (= (uf_44 ?x604) uf_9))) :pat { (uf_55 ?x604) }) +:assumption (forall (?x605 T4) (implies (= (uf_203 ?x605) uf_9) (and (= (uf_55 ?x605) uf_9) (<= 0 (uf_173 ?x605)))) :pat { (uf_203 ?x605) }) +:assumption (forall (?x606 T3) (implies (= (uf_22 ?x606) uf_9) (forall (?x607 T4) (?x608 Int) (?x609 T5) (iff (= (uf_15 ?x609 (uf_128 ?x607 (uf_43 ?x606 ?x608))) uf_9) (= ?x609 (uf_43 ?x606 ?x608))) :pat { (uf_15 ?x609 (uf_128 ?x607 (uf_43 ?x606 ?x608))) })) :pat { (uf_22 ?x606) }) +:assumption (forall (?x610 T3) (implies (= (uf_22 ?x610) uf_9) (forall (?x611 Int) (?x612 T5) (iff (= (uf_15 ?x612 (uf_130 (uf_43 ?x610 ?x611))) uf_9) (= ?x612 (uf_43 ?x610 ?x611))) :pat { (uf_15 ?x612 (uf_130 (uf_43 ?x610 ?x611))) })) :pat { (uf_22 ?x610) }) +:assumption (forall (?x613 T4) (?x614 T4) (?x615 T5) (?x616 T3) (iff (= (uf_204 ?x613 ?x614 ?x615 ?x616) uf_9) (and (= (uf_11 ?x613 ?x615) (uf_11 ?x614 ?x615)) (and (= (uf_58 (uf_59 ?x613) ?x615) (uf_58 (uf_59 ?x614) ?x615)) (up_205 ?x613 ?x614 ?x615 ?x616)))) :pat { (uf_204 ?x613 ?x614 ?x615 ?x616) }) +:assumption (forall (?x617 T4) (?x618 T4) (?x619 T5) (?x620 T3) (iff (= (uf_206 ?x617 ?x618 ?x619 ?x620) uf_9) (and (= (uf_11 ?x617 ?x619) (uf_11 ?x618 ?x619)) (and (= (uf_53 ?x617 ?x619) (uf_53 ?x618 ?x619)) (and (= (uf_58 (uf_59 ?x617) ?x619) (uf_58 (uf_59 ?x618) ?x619)) (= (uf_123 ?x617 ?x618 ?x619 ?x620) uf_9))))) :pat { (uf_206 ?x617 ?x618 ?x619 ?x620) }) +:assumption (forall (?x621 T4) (?x622 T4) (?x623 T5) (?x624 T5) (iff (= (uf_207 ?x621 ?x622 ?x623 ?x624) uf_9) (or (= (uf_206 ?x621 ?x622 ?x624 (uf_13 ?x624)) uf_9) (or (and (not (= (uf_27 ?x621 ?x623) uf_9)) (not (= (uf_27 ?x622 ?x623) uf_9))) (or (and (= (uf_46 ?x621 ?x622 ?x623 (uf_13 ?x623)) uf_9) (= (uf_204 ?x621 ?x622 ?x623 (uf_13 ?x623)) uf_9)) (= (uf_208 (uf_13 ?x623)) uf_9))))) :pat { (uf_207 ?x621 ?x622 ?x623 ?x624) }) +:assumption (forall (?x625 T4) (?x626 T4) (?x627 T5) (?x628 T3) (iff (= (uf_179 ?x625 ?x626 ?x627 ?x628) uf_9) (implies (and (= (uf_27 ?x625 ?x627) uf_9) (= (uf_27 ?x626 ?x627) uf_9)) (= (uf_206 ?x625 ?x626 ?x627 ?x628) uf_9))) :pat { (uf_179 ?x625 ?x626 ?x627 ?x628) }) :assumption (forall (?x629 T4) (?x630 T5) (?x631 T3) (implies (up_209 ?x629 ?x630 ?x631) (= (uf_46 ?x629 ?x629 ?x630 ?x631) uf_9)) :pat { (uf_46 ?x629 ?x629 ?x630 ?x631) }) -:assumption (forall (?x632 T4) (?x633 T5) (iff (= (uf_67 ?x632 ?x633) uf_9) (and (or (and (or (= (uf_210 ?x632 ?x633) uf_9) (= (uf_25 ?x632 ?x633) uf_26)) (not (= (uf_14 (uf_15 ?x633)) uf_16))) (and (or (= (uf_210 ?x632 (uf_136 (uf_58 (uf_59 ?x632) ?x633))) uf_9) (= (uf_25 ?x632 (uf_136 (uf_58 (uf_59 ?x632) ?x633))) uf_26)) (and (not (= (uf_14 (uf_15 (uf_136 (uf_58 (uf_59 ?x632) ?x633)))) uf_16)) (and (or (not (= (uf_24 ?x632 (uf_136 (uf_58 (uf_59 ?x632) ?x633))) uf_9)) (not (= (uf_135 (uf_58 (uf_59 ?x632) ?x633)) uf_9))) (= (uf_14 (uf_15 ?x633)) uf_16))))) (= (uf_27 ?x632 ?x633) uf_9))) :pat { (uf_67 ?x632 ?x633) }) -:assumption (forall (?x634 T4) (?x635 T5) (iff (= (uf_210 ?x634 ?x635) uf_9) (exists (?x636 T5) (and (= (uf_211 ?x634 ?x636) uf_9) (and (= (uf_22 (uf_15 ?x636)) uf_9) (and (not (= (uf_14 (uf_15 ?x636)) uf_16)) (and (= (uf_27 ?x634 ?x636) uf_9) (and (= (uf_48 ?x636 (uf_15 ?x636)) uf_9) (and (= (uf_25 ?x634 ?x636) uf_26) (and (= (uf_24 ?x634 ?x636) uf_9) (= (uf_13 ?x635 (uf_192 (uf_12 ?x634 ?x636))) uf_9)))))))) :pat { (uf_147 ?x635 (uf_192 (uf_12 ?x634 ?x636))) })) :pat { (uf_210 ?x634 ?x635) }) +:assumption (forall (?x632 T4) (?x633 T5) (iff (= (uf_68 ?x632 ?x633) uf_9) (and (= (uf_24 ?x632 ?x633) uf_9) (or (and (= (uf_12 (uf_13 ?x633)) uf_14) (and (or (not (= (uf_136 (uf_58 (uf_59 ?x632) ?x633)) uf_9)) (not (= (uf_27 ?x632 (uf_135 (uf_58 (uf_59 ?x632) ?x633))) uf_9))) (and (not (= (uf_12 (uf_13 (uf_135 (uf_58 (uf_59 ?x632) ?x633)))) uf_14)) (or (= (uf_25 ?x632 (uf_135 (uf_58 (uf_59 ?x632) ?x633))) uf_26) (= (uf_210 ?x632 (uf_135 (uf_58 (uf_59 ?x632) ?x633))) uf_9))))) (and (not (= (uf_12 (uf_13 ?x633)) uf_14)) (or (= (uf_25 ?x632 ?x633) uf_26) (= (uf_210 ?x632 ?x633) uf_9)))))) :pat { (uf_68 ?x632 ?x633) }) +:assumption (forall (?x634 T4) (?x635 T5) (iff (= (uf_210 ?x634 ?x635) uf_9) (exists (?x636 T5) (and (= (uf_15 ?x635 (uf_192 (uf_11 ?x634 ?x636))) uf_9) (and (= (uf_27 ?x634 ?x636) uf_9) (and (= (uf_25 ?x634 ?x636) uf_26) (and (= (uf_48 ?x636 (uf_13 ?x636)) uf_9) (and (= (uf_24 ?x634 ?x636) uf_9) (and (not (= (uf_12 (uf_13 ?x636)) uf_14)) (and (= (uf_23 (uf_13 ?x636)) uf_9) (= (uf_211 ?x634 ?x636) uf_9)))))))) :pat { (uf_147 ?x635 (uf_192 (uf_11 ?x634 ?x636))) })) :pat { (uf_210 ?x634 ?x635) }) :assumption (forall (?x637 T4) (?x638 T5) (iff (= (uf_211 ?x637 ?x638) uf_9) true) :pat { (uf_211 ?x637 ?x638) }) -:assumption (forall (?x639 T4) (?x640 T4) (?x641 T5) (implies (= (uf_177 ?x639 ?x640) uf_9) (up_212 (uf_40 (uf_41 ?x639) ?x641))) :pat { (uf_40 (uf_41 ?x640) ?x641) (uf_177 ?x639 ?x640) }) -:assumption (forall (?x642 T4) (?x643 T5) (implies (and (= (uf_27 ?x642 ?x643) uf_9) (= (uf_51 ?x642) uf_9)) (< 0 (uf_116 ?x643))) :pat { (uf_27 ?x642 ?x643) }) -:assumption (forall (?x644 T4) (?x645 T5) (implies (= (uf_51 ?x644) uf_9) (iff (= (uf_27 ?x644 ?x645) uf_9) (up_213 (uf_58 (uf_59 ?x644) ?x645)))) :pat { (uf_27 ?x644 ?x645) }) -:assumption (forall (?x646 T4) (?x647 T5) (iff (= (uf_61 ?x646 ?x647) uf_9) (and (not (= (uf_24 ?x646 ?x647) uf_9)) (and (= (uf_25 ?x646 ?x647) uf_26) (= (uf_27 ?x646 ?x647) uf_9)))) :pat { (uf_61 ?x646 ?x647) }) -:assumption (forall (?x648 T4) (?x649 T5) (= (uf_53 ?x648 ?x649) (uf_34 (uf_19 (uf_20 ?x648) (uf_126 ?x649 (uf_214 (uf_15 ?x649)))))) :pat { (uf_53 ?x648 ?x649) }) -:assumption (forall (?x650 T11) (and (= (uf_22 (uf_15 (uf_215 ?x650))) uf_9) (not (= (uf_14 (uf_15 (uf_215 ?x650))) uf_16))) :pat { (uf_215 ?x650) }) +:assumption (forall (?x639 T4) (?x640 T4) (?x641 T5) (implies (= (uf_178 ?x639 ?x640) uf_9) (up_212 (uf_40 (uf_41 ?x639) ?x641))) :pat { (uf_40 (uf_41 ?x640) ?x641) (uf_178 ?x639 ?x640) }) +:assumption (forall (?x642 T4) (?x643 T5) (implies (and (= (uf_51 ?x642) uf_9) (= (uf_24 ?x642 ?x643) uf_9)) (< 0 (uf_116 ?x643))) :pat { (uf_24 ?x642 ?x643) }) +:assumption (forall (?x644 T4) (?x645 T5) (implies (= (uf_51 ?x644) uf_9) (iff (= (uf_24 ?x644 ?x645) uf_9) (up_213 (uf_58 (uf_59 ?x644) ?x645)))) :pat { (uf_24 ?x644 ?x645) }) +:assumption (forall (?x646 T4) (?x647 T5) (iff (= (uf_61 ?x646 ?x647) uf_9) (and (= (uf_24 ?x646 ?x647) uf_9) (and (= (uf_25 ?x646 ?x647) uf_26) (not (= (uf_27 ?x646 ?x647) uf_9))))) :pat { (uf_61 ?x646 ?x647) }) +:assumption (forall (?x648 T4) (?x649 T5) (= (uf_53 ?x648 ?x649) (uf_34 (uf_19 (uf_20 ?x648) (uf_126 ?x649 (uf_214 (uf_13 ?x649)))))) :pat { (uf_53 ?x648 ?x649) }) +:assumption (forall (?x650 T11) (and (not (= (uf_12 (uf_13 (uf_215 ?x650))) uf_14)) (= (uf_23 (uf_13 (uf_215 ?x650))) uf_9)) :pat { (uf_215 ?x650) }) :assumption up_216 -:assumption (forall (?x651 T4) (?x652 T5) (implies (= (uf_22 (uf_15 ?x652)) uf_9) (= (uf_170 ?x651 ?x652) (uf_217 (uf_40 (uf_41 ?x651) ?x652)))) :pat { (uf_22 (uf_15 ?x652)) (uf_170 ?x651 ?x652) }) -:assumption (forall (?x653 T4) (?x654 T5) (implies (= (uf_23 (uf_15 ?x654)) uf_9) (= (uf_170 ?x653 ?x654) (uf_217 (uf_40 (uf_41 ?x653) (uf_136 (uf_58 (uf_59 ?x653) ?x654)))))) :pat { (uf_23 (uf_15 ?x654)) (uf_170 ?x653 ?x654) }) -:assumption (forall (?x655 T4) (?x656 T5) (implies (= (uf_22 (uf_15 ?x656)) uf_9) (iff (= (uf_24 ?x655 ?x656) uf_9) (up_218 (uf_40 (uf_41 ?x655) ?x656)))) :pat { (uf_22 (uf_15 ?x656)) (uf_24 ?x655 ?x656) }) -:assumption (forall (?x657 T4) (?x658 T5) (implies (= (uf_23 (uf_15 ?x658)) uf_9) (iff (= (uf_24 ?x657 ?x658) uf_9) (up_218 (uf_40 (uf_41 ?x657) (uf_136 (uf_58 (uf_59 ?x657) ?x658)))))) :pat { (uf_23 (uf_15 ?x658)) (uf_24 ?x657 ?x658) }) -:assumption (forall (?x659 T4) (?x660 T5) (implies (= (uf_22 (uf_15 ?x660)) uf_9) (= (uf_25 ?x659 ?x660) (uf_215 (uf_40 (uf_41 ?x659) ?x660)))) :pat { (uf_22 (uf_15 ?x660)) (uf_25 ?x659 ?x660) }) -:assumption (forall (?x661 T4) (?x662 T5) (implies (= (uf_23 (uf_15 ?x662)) uf_9) (= (uf_25 ?x661 ?x662) (uf_25 ?x661 (uf_136 (uf_58 (uf_59 ?x661) ?x662))))) :pat { (uf_23 (uf_15 ?x662)) (uf_25 ?x661 ?x662) }) +:assumption (forall (?x651 T4) (?x652 T5) (implies (= (uf_23 (uf_13 ?x652)) uf_9) (= (uf_172 ?x651 ?x652) (uf_217 (uf_40 (uf_41 ?x651) ?x652)))) :pat { (uf_23 (uf_13 ?x652)) (uf_172 ?x651 ?x652) }) +:assumption (forall (?x653 T4) (?x654 T5) (implies (= (uf_22 (uf_13 ?x654)) uf_9) (= (uf_172 ?x653 ?x654) (uf_217 (uf_40 (uf_41 ?x653) (uf_135 (uf_58 (uf_59 ?x653) ?x654)))))) :pat { (uf_22 (uf_13 ?x654)) (uf_172 ?x653 ?x654) }) +:assumption (forall (?x655 T4) (?x656 T5) (implies (= (uf_23 (uf_13 ?x656)) uf_9) (iff (= (uf_27 ?x655 ?x656) uf_9) (up_218 (uf_40 (uf_41 ?x655) ?x656)))) :pat { (uf_23 (uf_13 ?x656)) (uf_27 ?x655 ?x656) }) +:assumption (forall (?x657 T4) (?x658 T5) (implies (= (uf_22 (uf_13 ?x658)) uf_9) (iff (= (uf_27 ?x657 ?x658) uf_9) (up_218 (uf_40 (uf_41 ?x657) (uf_135 (uf_58 (uf_59 ?x657) ?x658)))))) :pat { (uf_22 (uf_13 ?x658)) (uf_27 ?x657 ?x658) }) +:assumption (forall (?x659 T4) (?x660 T5) (implies (= (uf_23 (uf_13 ?x660)) uf_9) (= (uf_25 ?x659 ?x660) (uf_215 (uf_40 (uf_41 ?x659) ?x660)))) :pat { (uf_23 (uf_13 ?x660)) (uf_25 ?x659 ?x660) }) +:assumption (forall (?x661 T4) (?x662 T5) (implies (= (uf_22 (uf_13 ?x662)) uf_9) (= (uf_25 ?x661 ?x662) (uf_25 ?x661 (uf_135 (uf_58 (uf_59 ?x661) ?x662))))) :pat { (uf_22 (uf_13 ?x662)) (uf_25 ?x661 ?x662) }) :assumption (forall (?x663 T5) (?x664 T3) (= (uf_126 ?x663 (uf_214 ?x664)) (uf_43 uf_219 (uf_220 ?x663 (uf_214 ?x664)))) :pat { (uf_126 ?x663 (uf_214 ?x664)) }) :assumption (up_197 uf_37) :assumption (forall (?x665 T17) (?x666 T17) (?x667 T15) (implies (= (uf_224 (uf_225 (uf_222 ?x665 ?x667)) (uf_225 (uf_222 ?x666 ?x667))) uf_9) (= (uf_221 (uf_222 ?x665 ?x667) (uf_222 ?x666 ?x667)) uf_9)) :pat { (uf_221 (uf_222 ?x665 ?x667) (uf_222 ?x666 (uf_223 ?x667))) }) @@ -584,70 +584,70 @@ :assumption (forall (?x671 T17) (= (uf_225 (uf_226 ?x671)) ?x671)) :assumption (forall (?x672 Int) (?x673 Int) (iff (= (uf_221 ?x672 ?x673) uf_9) (= ?x672 ?x673)) :pat { (uf_221 ?x672 ?x673) }) :assumption (forall (?x674 T17) (?x675 T17) (iff (= (uf_224 ?x674 ?x675) uf_9) (= ?x674 ?x675)) :pat { (uf_224 ?x674 ?x675) }) -:assumption (forall (?x676 T3) (?x677 T15) (?x678 T3) (implies (and (= (uf_228 ?x678) uf_9) (= (uf_227 ?x676 ?x677 ?x678) uf_9)) (= (uf_223 ?x677) ?x677)) :pat { (uf_227 ?x676 ?x677 ?x678) (uf_228 ?x678) }) -:assumption (forall (?x679 T3) (implies (= (uf_228 ?x679) uf_9) (= (uf_23 ?x679) uf_9)) :pat { (uf_228 ?x679) }) -:assumption (forall (?x680 T17) (?x681 T15) (?x682 T15) (?x683 Int) (or (= ?x681 ?x682) (= (uf_222 (uf_229 ?x680 ?x681 ?x683) ?x682) (uf_222 ?x680 ?x682))) :pat { (uf_222 (uf_229 ?x680 ?x681 ?x683) ?x682) }) +:assumption (forall (?x676 T3) (?x677 T15) (?x678 T3) (implies (and (= (uf_227 ?x676 ?x677 ?x678) uf_9) (= (uf_228 ?x678) uf_9)) (= (uf_223 ?x677) ?x677)) :pat { (uf_227 ?x676 ?x677 ?x678) (uf_228 ?x678) }) +:assumption (forall (?x679 T3) (implies (= (uf_228 ?x679) uf_9) (= (uf_22 ?x679) uf_9)) :pat { (uf_228 ?x679) }) +:assumption (forall (?x680 T17) (?x681 T15) (?x682 T15) (?x683 Int) (or (= (uf_222 (uf_229 ?x680 ?x681 ?x683) ?x682) (uf_222 ?x680 ?x682)) (= ?x681 ?x682)) :pat { (uf_222 (uf_229 ?x680 ?x681 ?x683) ?x682) }) :assumption (forall (?x684 T17) (?x685 T15) (?x686 Int) (= (uf_222 (uf_229 ?x684 ?x685 ?x686) ?x685) ?x686) :pat { (uf_222 (uf_229 ?x684 ?x685 ?x686) ?x685) }) :assumption (forall (?x687 T15) (= (uf_222 uf_230 ?x687) 0)) :assumption (forall (?x688 T17) (?x689 T15) (?x690 Int) (?x691 Int) (?x692 Int) (?x693 Int) (= (uf_231 ?x688 ?x689 ?x690 ?x691 ?x692 ?x693) (uf_229 ?x688 ?x689 (uf_99 (uf_222 ?x688 ?x689) ?x690 ?x691 ?x692 ?x693))) :pat { (uf_231 ?x688 ?x689 ?x690 ?x691 ?x692 ?x693) }) -:assumption (forall (?x694 T4) (?x695 T5) (implies (= (uf_51 ?x694) uf_9) (and (= (uf_233 (uf_232 ?x694 ?x695)) ?x694) (= (uf_234 (uf_232 ?x694 ?x695)) (uf_116 ?x695)))) :pat { (uf_232 ?x694 ?x695) }) -:assumption (forall (?x696 T18) (= (uf_51 (uf_233 ?x696)) uf_9)) -:assumption (= (uf_51 (uf_233 uf_235)) uf_9) -:assumption (forall (?x697 T4) (?x698 T5) (or (not (up_213 (uf_58 (uf_59 ?x697) ?x698))) (<= (uf_170 ?x697 ?x698) (uf_171 ?x697))) :pat { (uf_40 (uf_41 ?x697) ?x698) }) -:assumption (forall (?x699 T4) (?x700 T5) (implies (and (= (uf_135 (uf_58 (uf_59 ?x699) ?x700)) uf_9) (= (uf_51 ?x699) uf_9)) (= (uf_14 (uf_15 ?x700)) uf_16)) :pat { (uf_135 (uf_58 (uf_59 ?x699) ?x700)) }) -:assumption (forall (?x701 T4) (?x702 T5) (implies (= (uf_27 ?x701 ?x702) uf_9) (= (uf_27 ?x701 (uf_136 (uf_58 (uf_59 ?x701) ?x702))) uf_9)) :pat { (uf_27 ?x701 ?x702) (uf_58 (uf_59 ?x701) (uf_136 (uf_58 (uf_59 ?x701) ?x702))) }) -:assumption (forall (?x703 T14) (and (= (uf_22 (uf_15 (uf_136 ?x703))) uf_9) (not (= (uf_14 (uf_15 (uf_136 ?x703))) uf_16))) :pat { (uf_136 ?x703) }) +:assumption (forall (?x694 T4) (?x695 T5) (implies (= (uf_51 ?x694) uf_9) (and (= (uf_233 (uf_232 ?x694 ?x695)) (uf_116 ?x695)) (= (uf_234 (uf_232 ?x694 ?x695)) ?x694))) :pat { (uf_232 ?x694 ?x695) }) +:assumption (forall (?x696 T18) (= (uf_51 (uf_234 ?x696)) uf_9)) +:assumption (= (uf_51 (uf_234 uf_235)) uf_9) +:assumption (forall (?x697 T4) (?x698 T5) (or (<= (uf_172 ?x697 ?x698) (uf_173 ?x697)) (not (up_213 (uf_58 (uf_59 ?x697) ?x698)))) :pat { (uf_40 (uf_41 ?x697) ?x698) }) +:assumption (forall (?x699 T4) (?x700 T5) (implies (and (= (uf_51 ?x699) uf_9) (= (uf_136 (uf_58 (uf_59 ?x699) ?x700)) uf_9)) (= (uf_12 (uf_13 ?x700)) uf_14)) :pat { (uf_136 (uf_58 (uf_59 ?x699) ?x700)) }) +:assumption (forall (?x701 T4) (?x702 T5) (implies (= (uf_24 ?x701 ?x702) uf_9) (= (uf_24 ?x701 (uf_135 (uf_58 (uf_59 ?x701) ?x702))) uf_9)) :pat { (uf_24 ?x701 ?x702) (uf_58 (uf_59 ?x701) (uf_135 (uf_58 (uf_59 ?x701) ?x702))) }) +:assumption (forall (?x703 T14) (and (not (= (uf_12 (uf_13 (uf_135 ?x703))) uf_14)) (= (uf_23 (uf_13 (uf_135 ?x703))) uf_9)) :pat { (uf_135 ?x703) }) :assumption (forall (?x704 T5) (?x705 T15) (implies (<= 0 (uf_237 ?x705)) (= (uf_116 (uf_126 (uf_236 ?x704 ?x705) ?x705)) (uf_116 ?x704))) :pat { (uf_126 (uf_236 ?x704 ?x705) ?x705) }) :assumption (forall (?x706 T5) (?x707 T15) (= (uf_236 ?x706 ?x707) (uf_43 (uf_238 ?x707) (uf_239 ?x706 ?x707))) :pat { (uf_236 ?x706 ?x707) }) :assumption (forall (?x708 Int) (?x709 T15) (= (uf_236 (uf_126 (uf_43 (uf_238 ?x709) ?x708) ?x709) ?x709) (uf_43 (uf_238 ?x709) ?x708)) :pat { (uf_236 (uf_126 (uf_43 (uf_238 ?x709) ?x708) ?x709) ?x709) }) :assumption (forall (?x710 T5) (?x711 T3) (implies (= (uf_48 ?x710 ?x711) uf_9) (= ?x710 (uf_43 ?x711 (uf_116 ?x710)))) :pat { (uf_48 ?x710 ?x711) }) -:assumption (forall (?x712 T5) (?x713 T3) (iff (= (uf_48 ?x712 ?x713) uf_9) (= (uf_15 ?x712) ?x713))) +:assumption (forall (?x712 T5) (?x713 T3) (iff (= (uf_48 ?x712 ?x713) uf_9) (= (uf_13 ?x712) ?x713))) :assumption (= uf_121 (uf_43 uf_240 0)) -:assumption (forall (?x714 T15) (?x715 Int) (and (= (uf_242 (uf_241 ?x714 ?x715)) ?x715) (and (= (uf_243 (uf_241 ?x714 ?x715)) ?x714) (not (up_244 (uf_241 ?x714 ?x715))))) :pat { (uf_241 ?x714 ?x715) }) -:assumption (forall (?x716 T5) (?x717 T15) (and (= (uf_245 (uf_220 ?x716 ?x717)) ?x717) (= (uf_246 (uf_220 ?x716 ?x717)) ?x716)) :pat { (uf_220 ?x716 ?x717) }) +:assumption (forall (?x714 T15) (?x715 Int) (and (not (up_242 (uf_241 ?x714 ?x715))) (and (= (uf_243 (uf_241 ?x714 ?x715)) ?x714) (= (uf_244 (uf_241 ?x714 ?x715)) ?x715))) :pat { (uf_241 ?x714 ?x715) }) +:assumption (forall (?x716 T5) (?x717 T15) (and (= (uf_245 (uf_220 ?x716 ?x717)) ?x716) (= (uf_246 (uf_220 ?x716 ?x717)) ?x717)) :pat { (uf_220 ?x716 ?x717) }) :assumption (forall (?x718 T3) (?x719 Int) (= (uf_116 (uf_43 ?x718 ?x719)) ?x719)) -:assumption (forall (?x720 T3) (?x721 Int) (= (uf_15 (uf_43 ?x720 ?x721)) ?x720)) -:assumption (forall (?x722 T3) (?x723 T3) (?x724 Int) (?x725 Int) (iff (= (uf_247 ?x722 ?x723 ?x724 ?x725) uf_9) (and (= (uf_248 ?x722 ?x723) ?x725) (and (= (uf_249 ?x722 ?x723) ?x724) (up_250 ?x722 ?x723)))) :pat { (uf_247 ?x722 ?x723 ?x724 ?x725) }) -:assumption (forall (?x726 T5) (= (uf_139 ?x726 ?x726) uf_9) :pat { (uf_15 ?x726) }) -:assumption (forall (?x727 T5) (?x728 T5) (?x729 T5) (implies (and (= (uf_139 ?x728 ?x729) uf_9) (= (uf_139 ?x727 ?x728) uf_9)) (= (uf_139 ?x727 ?x729) uf_9)) :pat { (uf_139 ?x727 ?x728) (uf_139 ?x728 ?x729) }) -:assumption (forall (?x730 T12) (?x731 T5) (?x732 T5) (?x733 T11) (or (= (uf_40 (uf_172 ?x730 ?x731 ?x733) ?x732) (uf_40 ?x730 ?x732)) (= ?x731 ?x732))) -:assumption (forall (?x734 T12) (?x735 T5) (?x736 T11) (= (uf_40 (uf_172 ?x734 ?x735 ?x736) ?x735) ?x736)) -:assumption (forall (?x737 T13) (?x738 T5) (?x739 T5) (?x740 T14) (or (= (uf_58 (uf_251 ?x737 ?x738 ?x740) ?x739) (uf_58 ?x737 ?x739)) (= ?x738 ?x739))) +:assumption (forall (?x720 T3) (?x721 Int) (= (uf_13 (uf_43 ?x720 ?x721)) ?x720)) +:assumption (forall (?x722 T3) (?x723 T3) (?x724 Int) (?x725 Int) (iff (= (uf_247 ?x722 ?x723 ?x724 ?x725) uf_9) (and (up_248 ?x722 ?x723) (and (= (uf_249 ?x722 ?x723) ?x724) (= (uf_250 ?x722 ?x723) ?x725)))) :pat { (uf_247 ?x722 ?x723 ?x724 ?x725) }) +:assumption (forall (?x726 T5) (= (uf_138 ?x726 ?x726) uf_9) :pat { (uf_13 ?x726) }) +:assumption (forall (?x727 T5) (?x728 T5) (?x729 T5) (implies (and (= (uf_138 ?x727 ?x728) uf_9) (= (uf_138 ?x728 ?x729) uf_9)) (= (uf_138 ?x727 ?x729) uf_9)) :pat { (uf_138 ?x727 ?x728) (uf_138 ?x728 ?x729) }) +:assumption (forall (?x730 T12) (?x731 T5) (?x732 T5) (?x733 T11) (or (= ?x731 ?x732) (= (uf_40 (uf_170 ?x730 ?x731 ?x733) ?x732) (uf_40 ?x730 ?x732)))) +:assumption (forall (?x734 T12) (?x735 T5) (?x736 T11) (= (uf_40 (uf_170 ?x734 ?x735 ?x736) ?x735) ?x736)) +:assumption (forall (?x737 T13) (?x738 T5) (?x739 T5) (?x740 T14) (or (= ?x738 ?x739) (= (uf_58 (uf_251 ?x737 ?x738 ?x740) ?x739) (uf_58 ?x737 ?x739)))) :assumption (forall (?x741 T13) (?x742 T5) (?x743 T14) (= (uf_58 (uf_251 ?x741 ?x742 ?x743) ?x742) ?x743)) -:assumption (forall (?x744 T9) (?x745 T5) (?x746 T5) (?x747 Int) (or (= (uf_19 (uf_178 ?x744 ?x745 ?x747) ?x746) (uf_19 ?x744 ?x746)) (= ?x745 ?x746))) -:assumption (forall (?x748 T9) (?x749 T5) (?x750 Int) (= (uf_19 (uf_178 ?x748 ?x749 ?x750) ?x749) ?x750)) +:assumption (forall (?x744 T9) (?x745 T5) (?x746 T5) (?x747 Int) (or (= ?x745 ?x746) (= (uf_19 (uf_177 ?x744 ?x745 ?x747) ?x746) (uf_19 ?x744 ?x746)))) +:assumption (forall (?x748 T9) (?x749 T5) (?x750 Int) (= (uf_19 (uf_177 ?x748 ?x749 ?x750) ?x749) ?x750)) :assumption (= uf_26 (uf_43 uf_252 uf_253)) -:assumption (= (uf_23 uf_254) uf_9) -:assumption (= (uf_23 uf_255) uf_9) -:assumption (= (uf_23 uf_84) uf_9) -:assumption (= (uf_23 uf_4) uf_9) -:assumption (= (uf_23 uf_91) uf_9) -:assumption (= (uf_23 uf_7) uf_9) -:assumption (= (uf_23 uf_83) uf_9) -:assumption (= (uf_23 uf_87) uf_9) -:assumption (= (uf_23 uf_90) uf_9) -:assumption (= (uf_23 uf_94) uf_9) +:assumption (= (uf_22 uf_254) uf_9) +:assumption (= (uf_22 uf_255) uf_9) +:assumption (= (uf_22 uf_84) uf_9) +:assumption (= (uf_22 uf_4) uf_9) +:assumption (= (uf_22 uf_91) uf_9) +:assumption (= (uf_22 uf_7) uf_9) +:assumption (= (uf_22 uf_83) uf_9) +:assumption (= (uf_22 uf_87) uf_9) +:assumption (= (uf_22 uf_90) uf_9) +:assumption (= (uf_22 uf_94) uf_9) :assumption (= (uf_208 uf_252) uf_9) -:assumption (= (uf_23 uf_256) uf_9) -:assumption (= (uf_23 uf_219) uf_9) -:assumption (= (uf_23 uf_257) uf_9) -:assumption (= (uf_23 uf_258) uf_9) -:assumption (= (uf_23 uf_240) uf_9) -:assumption (forall (?x751 T3) (implies (= (uf_23 ?x751) uf_9) (not (up_36 ?x751))) :pat { (uf_23 ?x751) }) -:assumption (forall (?x752 T3) (= (uf_23 (uf_6 ?x752)) uf_9) :pat { (uf_6 ?x752) }) -:assumption (forall (?x753 T3) (?x754 T3) (= (uf_23 (uf_259 ?x753 ?x754)) uf_9) :pat { (uf_259 ?x753 ?x754) }) -:assumption (forall (?x755 T3) (implies (= (uf_208 ?x755) uf_9) (= (uf_22 ?x755) uf_9)) :pat { (uf_208 ?x755) }) -:assumption (forall (?x756 T3) (implies (= (uf_141 ?x756) uf_9) (= (uf_22 ?x756) uf_9)) :pat { (uf_141 ?x756) }) -:assumption (forall (?x757 T3) (implies (= (uf_260 ?x757) uf_9) (= (uf_22 ?x757) uf_9)) :pat { (uf_260 ?x757) }) -:assumption (forall (?x758 T3) (iff (= (uf_208 ?x758) uf_9) (= (uf_14 ?x758) uf_261)) :pat { (uf_208 ?x758) }) -:assumption (forall (?x759 T3) (iff (= (uf_141 ?x759) uf_9) (= (uf_14 ?x759) uf_262)) :pat { (uf_141 ?x759) }) -:assumption (forall (?x760 T3) (iff (= (uf_260 ?x760) uf_9) (= (uf_14 ?x760) uf_263)) :pat { (uf_260 ?x760) }) -:assumption (forall (?x761 T3) (iff (= (uf_23 ?x761) uf_9) (= (uf_14 ?x761) uf_16)) :pat { (uf_23 ?x761) }) +:assumption (= (uf_22 uf_256) uf_9) +:assumption (= (uf_22 uf_219) uf_9) +:assumption (= (uf_22 uf_257) uf_9) +:assumption (= (uf_22 uf_258) uf_9) +:assumption (= (uf_22 uf_240) uf_9) +:assumption (forall (?x751 T3) (implies (= (uf_22 ?x751) uf_9) (not (up_36 ?x751))) :pat { (uf_22 ?x751) }) +:assumption (forall (?x752 T3) (= (uf_22 (uf_6 ?x752)) uf_9) :pat { (uf_6 ?x752) }) +:assumption (forall (?x753 T3) (?x754 T3) (= (uf_22 (uf_259 ?x753 ?x754)) uf_9) :pat { (uf_259 ?x753 ?x754) }) +:assumption (forall (?x755 T3) (implies (= (uf_208 ?x755) uf_9) (= (uf_23 ?x755) uf_9)) :pat { (uf_208 ?x755) }) +:assumption (forall (?x756 T3) (implies (= (uf_141 ?x756) uf_9) (= (uf_23 ?x756) uf_9)) :pat { (uf_141 ?x756) }) +:assumption (forall (?x757 T3) (implies (= (uf_260 ?x757) uf_9) (= (uf_23 ?x757) uf_9)) :pat { (uf_260 ?x757) }) +:assumption (forall (?x758 T3) (iff (= (uf_208 ?x758) uf_9) (= (uf_12 ?x758) uf_261)) :pat { (uf_208 ?x758) }) +:assumption (forall (?x759 T3) (iff (= (uf_141 ?x759) uf_9) (= (uf_12 ?x759) uf_262)) :pat { (uf_141 ?x759) }) +:assumption (forall (?x760 T3) (iff (= (uf_260 ?x760) uf_9) (= (uf_12 ?x760) uf_263)) :pat { (uf_260 ?x760) }) +:assumption (forall (?x761 T3) (iff (= (uf_22 ?x761) uf_9) (= (uf_12 ?x761) uf_14)) :pat { (uf_22 ?x761) }) :assumption (forall (?x762 T3) (?x763 T3) (= (uf_142 (uf_259 ?x762 ?x763)) (+ (uf_142 ?x762) 23)) :pat { (uf_259 ?x762 ?x763) }) :assumption (forall (?x764 T3) (= (uf_142 (uf_6 ?x764)) (+ (uf_142 ?x764) 17)) :pat { (uf_6 ?x764) }) :assumption (forall (?x765 T3) (?x766 T3) (= (uf_264 (uf_259 ?x765 ?x766)) ?x765) :pat { (uf_259 ?x765 ?x766) }) :assumption (forall (?x767 T3) (?x768 T3) (= (uf_265 (uf_259 ?x767 ?x768)) ?x768) :pat { (uf_259 ?x767 ?x768) }) -:assumption (forall (?x769 T3) (= (uf_138 (uf_6 ?x769)) 8) :pat { (uf_6 ?x769) }) +:assumption (forall (?x769 T3) (= (uf_139 (uf_6 ?x769)) 8) :pat { (uf_6 ?x769) }) :assumption (forall (?x770 T3) (= (uf_266 (uf_6 ?x770)) ?x770) :pat { (uf_6 ?x770) }) :assumption (= (uf_260 uf_267) uf_9) :assumption (= (uf_260 uf_37) uf_9) @@ -670,18 +670,18 @@ :assumption (= (uf_142 uf_87) 0) :assumption (= (uf_142 uf_90) 0) :assumption (= (uf_142 uf_94) 0) -:assumption (= (uf_138 uf_219) 1) -:assumption (= (uf_138 uf_252) 1) -:assumption (= (uf_138 uf_254) 8) -:assumption (= (uf_138 uf_255) 4) -:assumption (= (uf_138 uf_84) 8) -:assumption (= (uf_138 uf_4) 4) -:assumption (= (uf_138 uf_91) 2) -:assumption (= (uf_138 uf_7) 1) -:assumption (= (uf_138 uf_83) 8) -:assumption (= (uf_138 uf_87) 4) -:assumption (= (uf_138 uf_90) 2) -:assumption (= (uf_138 uf_94) 1) -:assumption (not (implies true (implies (and (<= uf_269 uf_78) (<= 0 uf_269)) (implies (and (<= uf_270 uf_76) (<= 0 uf_270)) (implies (and (<= uf_271 uf_76) (<= 0 uf_271)) (implies (< uf_272 1099511627776) (implies (< 0 uf_272) (implies (and (= (uf_22 (uf_124 uf_7 uf_272)) uf_9) (and (not (= (uf_14 (uf_124 uf_7 uf_272)) uf_16)) (and (= (uf_27 uf_273 (uf_43 (uf_124 uf_7 uf_272) (uf_116 (uf_43 uf_7 uf_274)))) uf_9) (and (= (uf_48 (uf_43 (uf_124 uf_7 uf_272) (uf_116 (uf_43 uf_7 uf_274))) (uf_124 uf_7 uf_272)) uf_9) (and (= (uf_25 uf_273 (uf_43 (uf_124 uf_7 uf_272) (uf_116 (uf_43 uf_7 uf_274)))) uf_26) (= (uf_24 uf_273 (uf_43 (uf_124 uf_7 uf_272) (uf_116 (uf_43 uf_7 uf_274)))) uf_9)))))) (implies true (implies (= (uf_203 uf_273) uf_9) (implies (and (= (uf_55 uf_273) uf_9) (= (uf_202 uf_275 uf_273) uf_9)) (implies (forall (?x771 T19) (< (uf_276 ?x771) uf_277) :pat { (uf_276 ?x771) }) (implies (and (up_278 uf_273 uf_275 uf_279 (uf_43 uf_7 uf_274) (uf_6 uf_7)) (up_280 uf_273 uf_275 uf_279 (uf_29 (uf_43 uf_7 uf_274)) (uf_6 uf_7))) (implies (up_280 uf_273 uf_275 uf_281 uf_272 uf_4) (implies (= uf_282 (uf_171 uf_273)) (implies (forall (?x772 T5) (iff (= (uf_283 uf_282 ?x772) uf_9) false) :pat { (uf_283 uf_282 ?x772) }) (implies (and (<= uf_272 uf_76) (<= 0 uf_272)) (and (implies (= (uf_200 uf_273 (uf_43 (uf_124 uf_7 uf_272) (uf_116 (uf_43 uf_7 uf_274))) (uf_43 (uf_124 uf_7 uf_272) (uf_116 (uf_43 uf_7 uf_274))) uf_284) uf_9) (and (implies (and (= (uf_27 uf_273 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7)) uf_9) (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7) uf_7) uf_9)) (and (implies (and (= (uf_67 uf_273 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7)) uf_9) (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7) uf_7) uf_9)) (implies (= uf_285 (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7))) (implies (up_280 uf_273 uf_286 uf_287 uf_285 uf_7) (implies (up_280 uf_273 uf_288 uf_289 0 uf_4) (implies (up_280 uf_273 uf_290 uf_291 1 uf_4) (implies (and (<= 0 0) (and (<= 0 0) (and (<= 1 1) (<= 1 1)))) (and (implies (<= 1 uf_272) (and (implies (forall (?x773 Int) (implies (and (<= ?x773 uf_76) (<= 0 ?x773)) (implies (< ?x773 1) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x773 uf_7)) uf_285)))) (and (implies (and (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7)) uf_285) (< 0 uf_272)) (implies true (implies (and (<= uf_292 uf_78) (<= 0 uf_292)) (implies (and (<= uf_293 uf_76) (<= 0 uf_293)) (implies (and (<= uf_294 uf_76) (<= 0 uf_294)) (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies (<= uf_294 uf_272) (implies (forall (?x774 Int) (implies (and (<= ?x774 uf_76) (<= 0 ?x774)) (implies (< ?x774 uf_294) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x774 uf_7)) uf_292)))) (implies (and (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_293 uf_7)) uf_292) (< uf_293 uf_272)) (implies (and (<= 0 uf_293) (<= 1 uf_294)) (and (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies (and (= (uf_177 uf_273 uf_273) uf_9) (and (forall (?x775 T5) (<= (uf_170 uf_273 ?x775) (uf_170 uf_273 ?x775)) :pat { (uf_170 uf_273 ?x775) }) (and (<= (uf_171 uf_273) (uf_171 uf_273)) (and (forall (?x776 T5) (implies (= (uf_67 uf_273 ?x776) uf_9) (and (= (uf_67 uf_273 ?x776) uf_9) (= (uf_58 (uf_59 uf_273) ?x776) (uf_58 (uf_59 uf_273) ?x776)))) :pat { (uf_58 (uf_59 uf_273) ?x776) }) (and (forall (?x777 T5) (implies (= (uf_67 uf_273 ?x777) uf_9) (and (= (uf_67 uf_273 ?x777) uf_9) (= (uf_40 (uf_41 uf_273) ?x777) (uf_40 (uf_41 uf_273) ?x777)))) :pat { (uf_40 (uf_41 uf_273) ?x777) }) (and (forall (?x778 T5) (implies (= (uf_67 uf_273 ?x778) uf_9) (and (= (uf_67 uf_273 ?x778) uf_9) (= (uf_19 (uf_20 uf_273) ?x778) (uf_19 (uf_20 uf_273) ?x778)))) :pat { (uf_19 (uf_20 uf_273) ?x778) }) (forall (?x779 T5) (implies (not (= (uf_14 (uf_15 (uf_25 uf_273 ?x779))) uf_261)) (not (= (uf_14 (uf_15 (uf_25 uf_273 ?x779))) uf_261))) :pat { (uf_40 (uf_41 uf_273) ?x779) }))))))) (implies (and (= (uf_177 uf_273 uf_273) uf_9) (and (forall (?x780 T5) (<= (uf_170 uf_273 ?x780) (uf_170 uf_273 ?x780)) :pat { (uf_170 uf_273 ?x780) }) (<= (uf_171 uf_273) (uf_171 uf_273)))) (implies (and (= (uf_55 uf_273) uf_9) (= (uf_202 uf_295 uf_273) uf_9)) (implies (up_280 uf_273 uf_295 uf_291 uf_294 uf_4) (implies (up_280 uf_273 uf_295 uf_289 uf_293 uf_4) (implies (up_280 uf_273 uf_295 uf_287 uf_292 uf_7) (implies (up_280 uf_273 uf_295 uf_281 uf_272 uf_4) (implies (and (up_278 uf_273 uf_295 uf_279 (uf_43 uf_7 uf_274) (uf_6 uf_7)) (up_280 uf_273 uf_295 uf_279 (uf_29 (uf_43 uf_7 uf_274)) (uf_6 uf_7))) (implies (and (= (uf_41 uf_273) (uf_41 uf_273)) (= (uf_59 uf_273) (uf_59 uf_273))) (implies (and (<= 0 uf_293) (<= 1 uf_294)) (and (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies (<= uf_272 uf_294) (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (and (implies up_216 (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies true (implies (= uf_296 uf_292) (implies (= uf_297 uf_294) (implies (= uf_298 uf_293) (implies (= uf_299 uf_292) (implies true (and (implies (forall (?x781 Int) (implies (and (<= ?x781 uf_76) (<= 0 ?x781)) (implies (< ?x781 uf_272) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x781 uf_7)) uf_299)))) (and (implies (exists (?x782 Int) (and (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x782 uf_7)) uf_299) (and (< ?x782 uf_272) (and (<= ?x782 uf_76) (<= 0 ?x782))))) true) (exists (?x783 Int) (and (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x783 uf_7)) uf_299) (and (< ?x783 uf_272) (and (<= ?x783 uf_76) (<= 0 ?x783))))))) (forall (?x784 Int) (implies (and (<= ?x784 uf_76) (<= 0 ?x784)) (implies (< ?x784 uf_272) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x784 uf_7)) uf_299)))))))))))))))) up_216)))))))))) (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies (< uf_294 uf_272) (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (and (implies (and (= (uf_27 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7)) uf_9) (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7) uf_7) uf_9)) (and (implies (and (= (uf_67 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7)) uf_9) (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7) uf_7) uf_9)) (implies (and (<= 0 uf_293) (<= 1 uf_294)) (and (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7)) uf_292) (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies true (implies (= uf_300 uf_292) (implies (= uf_301 uf_293) (implies true (implies (and (<= 0 uf_301) (<= 1 uf_294)) (and (implies (and (<= (+ uf_294 1) uf_76) (<= 0 (+ uf_294 1))) (implies (= uf_302 (+ uf_294 1)) (implies (up_280 uf_273 uf_303 uf_291 uf_302 uf_4) (implies (and (<= 0 uf_301) (<= 2 uf_302)) (implies true (and (implies (<= uf_302 uf_272) (and (implies (forall (?x785 Int) (implies (and (<= ?x785 uf_76) (<= 0 ?x785)) (implies (< ?x785 uf_302) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x785 uf_7)) uf_300)))) (and (implies (and (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_301 uf_7)) uf_300) (< uf_301 uf_272)) (implies false true)) (and (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_301 uf_7)) uf_300) (< uf_301 uf_272)))) (forall (?x786 Int) (implies (and (<= ?x786 uf_76) (<= 0 ?x786)) (implies (< ?x786 uf_302) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x786 uf_7)) uf_300)))))) (<= uf_302 uf_272))))))) (and (<= (+ uf_294 1) uf_76) (<= 0 (+ uf_294 1)))))))))))))))) (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies (< uf_292 (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7))) (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (and (implies (and (= (uf_27 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7)) uf_9) (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7) uf_7) uf_9)) (and (implies (and (= (uf_67 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7)) uf_9) (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7) uf_7) uf_9)) (implies (= uf_304 (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7))) (implies (up_280 uf_273 uf_305 uf_287 uf_304 uf_7) (implies (up_280 uf_273 uf_306 uf_289 uf_294 uf_4) (implies (and (<= 1 uf_294) (<= 1 uf_294)) (implies true (implies (= uf_300 uf_304) (implies (= uf_301 uf_294) (implies true (implies (and (<= 0 uf_301) (<= 1 uf_294)) (and (implies (and (<= (+ uf_294 1) uf_76) (<= 0 (+ uf_294 1))) (implies (= uf_302 (+ uf_294 1)) (implies (up_280 uf_273 uf_303 uf_291 uf_302 uf_4) (implies (and (<= 0 uf_301) (<= 2 uf_302)) (implies true (and (implies (<= uf_302 uf_272) (and (implies (forall (?x787 Int) (implies (and (<= ?x787 uf_76) (<= 0 ?x787)) (implies (< ?x787 uf_302) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x787 uf_7)) uf_300)))) (and (implies (and (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_301 uf_7)) uf_300) (< uf_301 uf_272)) (implies false true)) (and (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_301 uf_7)) uf_300) (< uf_301 uf_272)))) (forall (?x788 Int) (implies (and (<= ?x788 uf_76) (<= 0 ?x788)) (implies (< ?x788 uf_302) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x788 uf_7)) uf_300)))))) (<= uf_302 uf_272))))))) (and (<= (+ uf_294 1) uf_76) (<= 0 (+ uf_294 1)))))))))))))) (and (= (uf_67 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7)) uf_9) (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7) uf_7) uf_9)))) (and (= (uf_27 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7)) uf_9) (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7) uf_7) uf_9)))))))))))) (and (= (uf_67 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7)) uf_9) (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7) uf_7) uf_9)))) (and (= (uf_27 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7)) uf_9) (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7) uf_7) uf_9)))))))))))))))))))))))))) (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies (not true) (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies (and (= (uf_55 uf_273) uf_9) (= (uf_202 uf_295 uf_273) uf_9)) (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (and (implies up_216 (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies true (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies (and (<= 0 uf_293) (<= 1 uf_294)) (implies true (implies (= uf_296 uf_292) (implies (= uf_297 uf_294) (implies (= uf_298 uf_293) (implies (= uf_299 uf_292) (implies true (and (implies (forall (?x789 Int) (implies (and (<= ?x789 uf_76) (<= 0 ?x789)) (implies (< ?x789 uf_272) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x789 uf_7)) uf_299)))) (and (implies (exists (?x790 Int) (and (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x790 uf_7)) uf_299) (and (< ?x790 uf_272) (and (<= ?x790 uf_76) (<= 0 ?x790))))) true) (exists (?x791 Int) (and (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x791 uf_7)) uf_299) (and (< ?x791 uf_272) (and (<= ?x791 uf_76) (<= 0 ?x791))))))) (forall (?x792 Int) (implies (and (<= ?x792 uf_76) (<= 0 ?x792)) (implies (< ?x792 uf_272) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x792 uf_7)) uf_299)))))))))))))))) up_216)))))))))))))))))))))) (and (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7)) uf_285) (< 0 uf_272)))) (forall (?x793 Int) (implies (and (<= ?x793 uf_76) (<= 0 ?x793)) (implies (< ?x793 1) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x793 uf_7)) uf_285)))))) (<= 1 uf_272)))))))) (and (= (uf_67 uf_273 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7)) uf_9) (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7) uf_7) uf_9)))) (and (= (uf_27 uf_273 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7)) uf_9) (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7) uf_7) uf_9)))) (= (uf_200 uf_273 (uf_43 (uf_124 uf_7 uf_272) (uf_116 (uf_43 uf_7 uf_274))) (uf_43 (uf_124 uf_7 uf_272) (uf_116 (uf_43 uf_7 uf_274))) uf_284) uf_9))))))))))))))))))) +:assumption (= (uf_139 uf_219) 1) +:assumption (= (uf_139 uf_252) 1) +:assumption (= (uf_139 uf_254) 8) +:assumption (= (uf_139 uf_255) 4) +:assumption (= (uf_139 uf_84) 8) +:assumption (= (uf_139 uf_4) 4) +:assumption (= (uf_139 uf_91) 2) +:assumption (= (uf_139 uf_7) 1) +:assumption (= (uf_139 uf_83) 8) +:assumption (= (uf_139 uf_87) 4) +:assumption (= (uf_139 uf_90) 2) +:assumption (= (uf_139 uf_94) 1) +:assumption (not (implies true (implies (and (<= 0 uf_269) (<= uf_269 uf_78)) (implies (and (<= 0 uf_270) (<= uf_270 uf_76)) (implies (and (<= 0 uf_271) (<= uf_271 uf_76)) (implies (< uf_272 1099511627776) (implies (< 0 uf_272) (implies (and (= (uf_27 uf_273 (uf_43 (uf_124 uf_7 uf_272) (uf_116 (uf_43 uf_7 uf_274)))) uf_9) (and (= (uf_25 uf_273 (uf_43 (uf_124 uf_7 uf_272) (uf_116 (uf_43 uf_7 uf_274)))) uf_26) (and (= (uf_48 (uf_43 (uf_124 uf_7 uf_272) (uf_116 (uf_43 uf_7 uf_274))) (uf_124 uf_7 uf_272)) uf_9) (and (= (uf_24 uf_273 (uf_43 (uf_124 uf_7 uf_272) (uf_116 (uf_43 uf_7 uf_274)))) uf_9) (and (not (= (uf_12 (uf_124 uf_7 uf_272)) uf_14)) (= (uf_23 (uf_124 uf_7 uf_272)) uf_9)))))) (implies true (implies (= (uf_203 uf_273) uf_9) (implies (and (= (uf_202 uf_275 uf_273) uf_9) (= (uf_55 uf_273) uf_9)) (implies (forall (?x771 T19) (< (uf_276 ?x771) uf_277) :pat { (uf_276 ?x771) }) (implies (and (up_278 uf_273 uf_275 uf_279 (uf_29 (uf_43 uf_7 uf_274)) (uf_6 uf_7)) (up_280 uf_273 uf_275 uf_279 (uf_43 uf_7 uf_274) (uf_6 uf_7))) (implies (up_278 uf_273 uf_275 uf_281 uf_272 uf_4) (implies (= uf_282 (uf_173 uf_273)) (implies (forall (?x772 T5) (iff (= (uf_283 uf_282 ?x772) uf_9) false) :pat { (uf_283 uf_282 ?x772) }) (implies (and (<= 0 uf_272) (<= uf_272 uf_76)) (and (= (uf_200 uf_273 (uf_43 (uf_124 uf_7 uf_272) (uf_116 (uf_43 uf_7 uf_274))) (uf_43 (uf_124 uf_7 uf_272) (uf_116 (uf_43 uf_7 uf_274))) uf_284) uf_9) (implies (= (uf_200 uf_273 (uf_43 (uf_124 uf_7 uf_272) (uf_116 (uf_43 uf_7 uf_274))) (uf_43 (uf_124 uf_7 uf_272) (uf_116 (uf_43 uf_7 uf_274))) uf_284) uf_9) (and (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7) uf_7) uf_9) (= (uf_24 uf_273 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7)) uf_9)) (implies (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7) uf_7) uf_9) (= (uf_24 uf_273 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7)) uf_9)) (and (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7) uf_7) uf_9) (= (uf_68 uf_273 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7)) uf_9)) (implies (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7) uf_7) uf_9) (= (uf_68 uf_273 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7)) uf_9)) (implies (= uf_285 (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7))) (implies (up_278 uf_273 uf_286 uf_287 uf_285 uf_7) (implies (up_278 uf_273 uf_288 uf_289 0 uf_4) (implies (up_278 uf_273 uf_290 uf_291 1 uf_4) (implies (and (<= 1 1) (and (<= 1 1) (and (<= 0 0) (<= 0 0)))) (and (<= 1 uf_272) (implies (<= 1 uf_272) (and (forall (?x773 Int) (implies (and (<= 0 ?x773) (<= ?x773 uf_76)) (implies (< ?x773 1) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x773 uf_7)) uf_285)))) (implies (forall (?x774 Int) (implies (and (<= 0 ?x774) (<= ?x774 uf_76)) (implies (< ?x774 1) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x774 uf_7)) uf_285)))) (and (and (< 0 uf_272) (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7)) uf_285)) (implies (and (< 0 uf_272) (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) 0 uf_7)) uf_285)) (implies true (implies (and (<= 0 uf_292) (<= uf_292 uf_78)) (implies (and (<= 0 uf_293) (<= uf_293 uf_76)) (implies (and (<= 0 uf_294) (<= uf_294 uf_76)) (implies (and (<= 1 uf_294) (<= 0 uf_293)) (implies (<= uf_294 uf_272) (implies (forall (?x775 Int) (implies (and (<= 0 ?x775) (<= ?x775 uf_76)) (implies (< ?x775 uf_294) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x775 uf_7)) uf_292)))) (implies (and (< uf_293 uf_272) (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_293 uf_7)) uf_292)) (implies (and (<= 1 uf_294) (<= 0 uf_293)) (and (implies true (implies (and (<= 1 uf_294) (<= 0 uf_293)) (implies (not true) (implies (and (<= 1 uf_294) (<= 0 uf_293)) (implies true (implies (and (<= 1 uf_294) (<= 0 uf_293)) (implies (and (= (uf_202 uf_295 uf_273) uf_9) (= (uf_55 uf_273) uf_9)) (implies (and (<= 1 uf_294) (<= 0 uf_293)) (implies true (implies (and (<= 1 uf_294) (<= 0 uf_293)) (and up_216 (implies up_216 (implies (and (<= 1 uf_294) (<= 0 uf_293)) (implies true (implies (and (<= 1 uf_294) (<= 0 uf_293)) (implies (and (<= 1 uf_294) (<= 0 uf_293)) (implies true (implies (= uf_296 uf_292) (implies (= uf_297 uf_294) (implies (= uf_298 uf_293) (implies (= uf_299 uf_292) (implies true (and (forall (?x776 Int) (implies (and (<= 0 ?x776) (<= ?x776 uf_76)) (implies (< ?x776 uf_272) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x776 uf_7)) uf_299)))) (implies (forall (?x777 Int) (implies (and (<= 0 ?x777) (<= ?x777 uf_76)) (implies (< ?x777 uf_272) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x777 uf_7)) uf_299)))) (and (exists (?x778 Int) (and (<= 0 ?x778) (and (<= ?x778 uf_76) (and (< ?x778 uf_272) (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x778 uf_7)) uf_299))))) (implies (exists (?x779 Int) (and (<= 0 ?x779) (and (<= ?x779 uf_76) (and (< ?x779 uf_272) (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x779 uf_7)) uf_299))))) true)))))))))))))))))))))))))) (implies true (implies (and (<= 1 uf_294) (<= 0 uf_293)) (implies true (implies (and (<= 1 uf_294) (<= 0 uf_293)) (implies true (implies (and (<= 1 uf_294) (<= 0 uf_293)) (implies (and (forall (?x780 T5) (implies (not (= (uf_12 (uf_13 (uf_25 uf_273 ?x780))) uf_261)) (not (= (uf_12 (uf_13 (uf_25 uf_273 ?x780))) uf_261))) :pat { (uf_40 (uf_41 uf_273) ?x780) }) (and (forall (?x781 T5) (implies (= (uf_68 uf_273 ?x781) uf_9) (and (= (uf_19 (uf_20 uf_273) ?x781) (uf_19 (uf_20 uf_273) ?x781)) (= (uf_68 uf_273 ?x781) uf_9))) :pat { (uf_19 (uf_20 uf_273) ?x781) }) (and (forall (?x782 T5) (implies (= (uf_68 uf_273 ?x782) uf_9) (and (= (uf_40 (uf_41 uf_273) ?x782) (uf_40 (uf_41 uf_273) ?x782)) (= (uf_68 uf_273 ?x782) uf_9))) :pat { (uf_40 (uf_41 uf_273) ?x782) }) (and (forall (?x783 T5) (implies (= (uf_68 uf_273 ?x783) uf_9) (and (= (uf_58 (uf_59 uf_273) ?x783) (uf_58 (uf_59 uf_273) ?x783)) (= (uf_68 uf_273 ?x783) uf_9))) :pat { (uf_58 (uf_59 uf_273) ?x783) }) (and (<= (uf_173 uf_273) (uf_173 uf_273)) (and (forall (?x784 T5) (<= (uf_172 uf_273 ?x784) (uf_172 uf_273 ?x784)) :pat { (uf_172 uf_273 ?x784) }) (= (uf_178 uf_273 uf_273) uf_9))))))) (implies (and (<= (uf_173 uf_273) (uf_173 uf_273)) (and (forall (?x785 T5) (<= (uf_172 uf_273 ?x785) (uf_172 uf_273 ?x785)) :pat { (uf_172 uf_273 ?x785) }) (= (uf_178 uf_273 uf_273) uf_9))) (implies (and (= (uf_202 uf_295 uf_273) uf_9) (= (uf_55 uf_273) uf_9)) (implies (up_278 uf_273 uf_295 uf_291 uf_294 uf_4) (implies (up_278 uf_273 uf_295 uf_289 uf_293 uf_4) (implies (up_278 uf_273 uf_295 uf_287 uf_292 uf_7) (implies (up_278 uf_273 uf_295 uf_281 uf_272 uf_4) (implies (and (up_278 uf_273 uf_295 uf_279 (uf_29 (uf_43 uf_7 uf_274)) (uf_6 uf_7)) (up_280 uf_273 uf_295 uf_279 (uf_43 uf_7 uf_274) (uf_6 uf_7))) (implies (and (= (uf_59 uf_273) (uf_59 uf_273)) (= (uf_41 uf_273) (uf_41 uf_273))) (implies (and (<= 1 uf_294) (<= 0 uf_293)) (and (implies true (implies (and (<= 1 uf_294) (<= 0 uf_293)) (implies (< uf_294 uf_272) (implies (and (<= 1 uf_294) (<= 0 uf_293)) (implies true (implies (and (<= 1 uf_294) (<= 0 uf_293)) (and (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7) uf_7) uf_9) (= (uf_24 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7)) uf_9)) (implies (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7) uf_7) uf_9) (= (uf_24 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7)) uf_9)) (and (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7) uf_7) uf_9) (= (uf_68 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7)) uf_9)) (implies (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7) uf_7) uf_9) (= (uf_68 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7)) uf_9)) (implies (and (<= 1 uf_294) (<= 0 uf_293)) (and (implies true (implies (and (<= 1 uf_294) (<= 0 uf_293)) (implies (< uf_292 (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7))) (implies (and (<= 1 uf_294) (<= 0 uf_293)) (implies true (implies (and (<= 1 uf_294) (<= 0 uf_293)) (and (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7) uf_7) uf_9) (= (uf_24 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7)) uf_9)) (implies (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7) uf_7) uf_9) (= (uf_24 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7)) uf_9)) (and (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7) uf_7) uf_9) (= (uf_68 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7)) uf_9)) (implies (and (= (uf_48 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7) uf_7) uf_9) (= (uf_68 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7)) uf_9)) (implies (= uf_300 (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7))) (implies (up_278 uf_273 uf_301 uf_287 uf_300 uf_7) (implies (up_278 uf_273 uf_302 uf_289 uf_294 uf_4) (implies (and (<= 1 uf_294) (<= 1 uf_294)) (implies true (implies (= uf_303 uf_300) (implies (= uf_304 uf_294) (implies true (implies (and (<= 1 uf_294) (<= 0 uf_304)) (and (and (<= 0 (+ uf_294 1)) (<= (+ uf_294 1) uf_76)) (implies (and (<= 0 (+ uf_294 1)) (<= (+ uf_294 1) uf_76)) (implies (= uf_305 (+ uf_294 1)) (implies (up_278 uf_273 uf_306 uf_291 uf_305 uf_4) (implies (and (<= 2 uf_305) (<= 0 uf_304)) (implies true (and (<= uf_305 uf_272) (implies (<= uf_305 uf_272) (and (forall (?x786 Int) (implies (and (<= 0 ?x786) (<= ?x786 uf_76)) (implies (< ?x786 uf_305) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x786 uf_7)) uf_303)))) (implies (forall (?x787 Int) (implies (and (<= 0 ?x787) (<= ?x787 uf_76)) (implies (< ?x787 uf_305) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x787 uf_7)) uf_303)))) (and (and (< uf_304 uf_272) (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_304 uf_7)) uf_303)) (implies (and (< uf_304 uf_272) (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_304 uf_7)) uf_303)) (implies false true)))))))))))))))))))))))))))))))) (implies true (implies (and (<= 1 uf_294) (<= 0 uf_293)) (implies (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_294 uf_7)) uf_292) (implies (and (<= 1 uf_294) (<= 0 uf_293)) (implies true (implies (and (<= 1 uf_294) (<= 0 uf_293)) (implies (and (<= 1 uf_294) (<= 0 uf_293)) (implies true (implies (= uf_303 uf_292) (implies (= uf_304 uf_293) (implies true (implies (and (<= 1 uf_294) (<= 0 uf_304)) (and (and (<= 0 (+ uf_294 1)) (<= (+ uf_294 1) uf_76)) (implies (and (<= 0 (+ uf_294 1)) (<= (+ uf_294 1) uf_76)) (implies (= uf_305 (+ uf_294 1)) (implies (up_278 uf_273 uf_306 uf_291 uf_305 uf_4) (implies (and (<= 2 uf_305) (<= 0 uf_304)) (implies true (and (<= uf_305 uf_272) (implies (<= uf_305 uf_272) (and (forall (?x788 Int) (implies (and (<= 0 ?x788) (<= ?x788 uf_76)) (implies (< ?x788 uf_305) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x788 uf_7)) uf_303)))) (implies (forall (?x789 Int) (implies (and (<= 0 ?x789) (<= ?x789 uf_76)) (implies (< ?x789 uf_305) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x789 uf_7)) uf_303)))) (and (and (< uf_304 uf_272) (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_304 uf_7)) uf_303)) (implies (and (< uf_304 uf_272) (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) uf_304 uf_7)) uf_303)) (implies false true))))))))))))))))))))))))))))))))))))) (implies true (implies (and (<= 1 uf_294) (<= 0 uf_293)) (implies (<= uf_272 uf_294) (implies (and (<= 1 uf_294) (<= 0 uf_293)) (implies true (implies (and (<= 1 uf_294) (<= 0 uf_293)) (implies (and (<= 1 uf_294) (<= 0 uf_293)) (implies true (implies (and (<= 1 uf_294) (<= 0 uf_293)) (and up_216 (implies up_216 (implies (and (<= 1 uf_294) (<= 0 uf_293)) (implies true (implies (and (<= 1 uf_294) (<= 0 uf_293)) (implies (and (<= 1 uf_294) (<= 0 uf_293)) (implies true (implies (= uf_296 uf_292) (implies (= uf_297 uf_294) (implies (= uf_298 uf_293) (implies (= uf_299 uf_292) (implies true (and (forall (?x790 Int) (implies (and (<= 0 ?x790) (<= ?x790 uf_76)) (implies (< ?x790 uf_272) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x790 uf_7)) uf_299)))) (implies (forall (?x791 Int) (implies (and (<= 0 ?x791) (<= ?x791 uf_76)) (implies (< ?x791 uf_272) (<= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x791 uf_7)) uf_299)))) (and (exists (?x792 Int) (and (<= 0 ?x792) (and (<= ?x792 uf_76) (and (< ?x792 uf_272) (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x792 uf_7)) uf_299))))) (implies (exists (?x793 Int) (and (<= 0 ?x793) (and (<= ?x793 uf_76) (and (< ?x793 uf_272) (= (uf_110 uf_273 (uf_66 (uf_43 uf_7 uf_274) ?x793 uf_7)) uf_299))))) true)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))) :formula true ) diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Boogie/Examples/cert/VCC_maximum.proof --- a/src/HOL/Boogie/Examples/cert/VCC_maximum.proof Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Boogie/Examples/cert/VCC_maximum.proof Sat Nov 14 09:40:27 2009 +0100 @@ -1,8070 +1,8005 @@ #2 := false -#121 := 0::int decl uf_110 :: (-> T4 T5 int) decl uf_66 :: (-> T5 int T3 T5) decl uf_7 :: T3 #10 := uf_7 -decl ?x785!14 :: int -#19054 := ?x785!14 +decl ?x786!14 :: int +#18483 := ?x786!14 decl uf_43 :: (-> T3 int T5) decl uf_274 :: int #2959 := uf_274 #2960 := (uf_43 uf_7 uf_274) -#19059 := (uf_66 #2960 ?x785!14 uf_7) +#18488 := (uf_66 #2960 ?x786!14 uf_7) decl uf_273 :: T4 -#2958 := uf_273 -#19060 := (uf_110 uf_273 #19059) -#4076 := -1::int -#19385 := (* -1::int #19060) -decl uf_300 :: int -#3186 := uf_300 -#19386 := (+ uf_300 #19385) -#19387 := (>= #19386 0::int) -#23584 := (not #19387) -#19372 := (* -1::int ?x785!14) -decl uf_302 :: int -#3196 := uf_302 -#19373 := (+ uf_302 #19372) -#19374 := (<= #19373 0::int) -#19056 := (>= ?x785!14 0::int) -#22816 := (not #19056) -#7878 := 131073::int -#19055 := (<= ?x785!14 131073::int) -#22815 := (not #19055) -#22831 := (or #22815 #22816 #19374 #19387) -#22836 := (not #22831) -#161 := (:var 0 int) -#3039 := (uf_66 #2960 #161 uf_7) -#23745 := (pattern #3039) -#15606 := (<= #161 131073::int) -#20064 := (not #15606) -#14120 := (* -1::int uf_300) -#3040 := (uf_110 uf_273 #3039) -#14121 := (+ #3040 #14120) -#14122 := (<= #14121 0::int) -#14101 := (* -1::int uf_302) -#14110 := (+ #161 #14101) -#14109 := (>= #14110 0::int) -#4084 := (>= #161 0::int) -#5113 := (not #4084) -#22797 := (or #5113 #14109 #14122 #20064) -#23762 := (forall (vars (?x785 int)) (:pat #23745) #22797) -#23767 := (not #23762) -decl uf_301 :: int -#3188 := uf_301 -#14142 := (* -1::int uf_301) -decl uf_272 :: int -#2949 := uf_272 -#14143 := (+ uf_272 #14142) -#14144 := (<= #14143 0::int) -#3208 := (uf_66 #2960 uf_301 uf_7) -#3209 := (uf_110 uf_273 #3208) -#12862 := (= uf_300 #3209) -#22782 := (not #12862) -#22783 := (or #22782 #14144) -#22784 := (not #22783) -#23770 := (or #22784 #23767) -#14145 := (not #14144) +#2957 := uf_273 +#18489 := (uf_110 uf_273 #18488) decl uf_294 :: int -#3055 := uf_294 -#14044 := (* -1::int uf_294) -#14045 := (+ uf_272 #14044) -#14046 := (<= #14045 0::int) -#14049 := (not #14046) -decl uf_125 :: (-> T5 T5 int) -decl uf_28 :: (-> int T5) -decl uf_29 :: (-> T5 int) -#2992 := (uf_29 #2960) -#23223 := (uf_28 #2992) -decl uf_15 :: (-> T5 T3) -#26404 := (uf_15 #23223) -decl uf_293 :: int -#3051 := uf_293 -#26963 := (uf_66 #23223 uf_293 #26404) -#26964 := (uf_125 #26963 #23223) -#27037 := (>= #26964 0::int) -#13947 := (>= uf_293 0::int) -decl ?x781!15 :: int -#19190 := ?x781!15 -#19195 := (uf_66 #2960 ?x781!15 uf_7) -#19196 := (uf_110 uf_273 #19195) -#19541 := (* -1::int #19196) -decl uf_299 :: int -#3138 := uf_299 -#19542 := (+ uf_299 #19541) -#19543 := (>= #19542 0::int) -#19528 := (* -1::int ?x781!15) -#19529 := (+ uf_272 #19528) -#19530 := (<= #19529 0::int) -#19192 := (>= ?x781!15 0::int) -#22993 := (not #19192) -#19191 := (<= ?x781!15 131073::int) -#22992 := (not #19191) -#23008 := (or #22992 #22993 #19530 #19543) -#23013 := (not #23008) -#13873 := (* -1::int uf_272) -#13960 := (+ #161 #13873) -#13959 := (>= #13960 0::int) -#3145 := (= #3040 uf_299) -#22966 := (not #3145) -#22967 := (or #22966 #5113 #13959 #20064) -#23886 := (forall (vars (?x782 int)) (:pat #23745) #22967) -#23891 := (not #23886) -#13970 := (* -1::int uf_299) -#13971 := (+ #3040 #13970) -#13972 := (<= #13971 0::int) -#22958 := (or #5113 #13959 #13972 #20064) -#23878 := (forall (vars (?x781 int)) (:pat #23745) #22958) -#23883 := (not #23878) -#23894 := (or #23883 #23891) -#23897 := (not #23894) -#23900 := (or #23897 #23013) -#23903 := (not #23900) -#4 := 1::int -#13950 := (>= uf_294 1::int) -#14243 := (not #13950) -#22873 := (not #13947) -decl uf_292 :: int -#3047 := uf_292 -#12576 := (= uf_292 uf_299) -#12644 := (not #12576) -decl uf_298 :: int -#3136 := uf_298 -#12573 := (= uf_293 uf_298) -#12653 := (not #12573) -decl uf_297 :: int -#3134 := uf_297 -#12570 := (= uf_294 uf_297) -#12662 := (not #12570) -decl uf_296 :: int -#3132 := uf_296 -#12567 := (= uf_292 uf_296) -#12671 := (not #12567) -#23906 := (or #12671 #12662 #12653 #12644 #22873 #14243 #14049 #23903) -#23909 := (not #23906) -#23773 := (not #23770) -#23776 := (or #23773 #22836) -#23779 := (not #23776) -#14102 := (+ uf_272 #14101) -#14100 := (>= #14102 0::int) -#14105 := (not #14100) -#23782 := (or #14105 #23779) -#23785 := (not #23782) -#23788 := (or #14105 #23785) -#23791 := (not #23788) -#1066 := 131072::int -#16368 := (<= uf_294 131072::int) -#19037 := (not #16368) -#14169 := (+ uf_294 #14101) -#14168 := (= #14169 -1::int) -#14172 := (not #14168) -#1120 := 2::int -#14092 := (>= uf_302 2::int) -#22859 := (not #14092) -#14088 := (>= uf_294 -1::int) -#19034 := (not #14088) -#14076 := (>= uf_301 0::int) -#22858 := (not #14076) -decl up_280 :: (-> T4 T1 T1 int T3 bool) +#3060 := uf_294 +#3180 := (uf_66 #2960 uf_294 uf_7) +#3189 := (uf_110 uf_273 #3180) +#29114 := (= #3189 #18489) +#29119 := (not #29114) +#121 := 0::int +#4066 := -1::int +#18810 := (* -1::int #18489) +#29118 := (+ #3189 #18810) +#29120 := (>= #29118 0::int) +#29130 := (not #29120) +decl uf_303 :: int +#3198 := uf_303 +#13466 := (* -1::int uf_303) +#28824 := (+ #3189 #13466) +#28825 := (<= #28824 0::int) +#28823 := (= #3189 uf_303) +decl uf_300 :: int +#3191 := uf_300 +#12365 := (= uf_300 uf_303) +#18811 := (+ uf_303 #18810) +#18812 := (>= #18811 0::int) +#18797 := (* -1::int ?x786!14) +decl uf_305 :: int +#3208 := uf_305 +#18798 := (+ uf_305 #18797) +#18799 := (<= #18798 0::int) +#18485 := (>= ?x786!14 0::int) +#22265 := (not #18485) +#7495 := 4294967295::int +#18484 := (<= ?x786!14 4294967295::int) +#22264 := (not #18484) +#22280 := (or #22264 #22265 #18799 #18812) +#22285 := (not #22280) +#161 := (:var 0 int) +#3044 := (uf_66 #2960 #161 uf_7) +#23194 := (pattern #3044) +#15097 := (<= #161 4294967295::int) +#19482 := (not #15097) +#3045 := (uf_110 uf_273 #3044) +#13467 := (+ #3045 #13466) +#13468 := (<= #13467 0::int) +#13447 := (* -1::int uf_305) +#13455 := (+ #161 #13447) +#13454 := (>= #13455 0::int) +#4065 := (>= #161 0::int) +#4987 := (not #4065) +#22246 := (or #4987 #13454 #13468 #19482) +#23211 := (forall (vars (?x786 int)) (:pat #23194) #22246) +#23216 := (not #23211) +decl uf_304 :: int +#3200 := uf_304 +#13488 := (* -1::int uf_304) +decl uf_272 :: int +#2954 := uf_272 +#13489 := (+ uf_272 #13488) +#13490 := (<= #13489 0::int) +#3221 := (uf_66 #2960 uf_304 uf_7) +#3222 := (uf_110 uf_273 #3221) +#12404 := (= uf_303 #3222) +#22231 := (not #12404) +#22232 := (or #22231 #13490) +#22233 := (not #22232) +#23219 := (or #22233 #23216) +#23222 := (not #23219) +#23225 := (or #23222 #22285) +#23228 := (not #23225) +#13448 := (+ uf_272 #13447) +#13446 := (>= #13448 0::int) +#13451 := (not #13446) +#23231 := (or #13451 #23228) +#23234 := (not #23231) +#23237 := (or #13451 #23234) +#23240 := (not #23237) +#15795 := 4294967294::int +#15796 := (<= uf_294 4294967294::int) +#18466 := (not #15796) +#13515 := (+ uf_294 #13447) +#13514 := (= #13515 -1::int) +#13518 := (not #13514) +#892 := 2::int +#13438 := (>= uf_305 2::int) +#22308 := (not #13438) +#13430 := (>= uf_294 -1::int) +#18463 := (not #13430) +#13421 := (>= uf_304 0::int) +#22307 := (not #13421) +decl up_278 :: (-> T4 T1 T1 int T3 bool) decl uf_4 :: T3 #7 := uf_4 decl uf_291 :: T1 -#3030 := uf_291 -decl uf_303 :: T1 -#3198 := uf_303 -#3199 := (up_280 uf_273 uf_303 uf_291 uf_302 uf_4) -#12942 := (not #3199) -#23794 := (or #12942 #22858 #19034 #22859 #14172 #19037 #23791) -#23797 := (not #23794) -#23800 := (or #19034 #19037 #23797) -#23803 := (not #23800) -#13075 := (= uf_294 uf_301) -#13081 := (not #13075) -decl uf_304 :: int -#3239 := uf_304 -#3175 := (uf_66 #2960 uf_294 uf_7) -#3184 := (uf_110 uf_273 #3175) -#13070 := (= #3184 uf_304) -#13133 := (not #13070) -decl uf_67 :: (-> T4 T5 T2) -#3181 := (uf_67 uf_273 #3175) +#3035 := uf_291 +decl uf_306 :: T1 +#3210 := uf_306 +#3211 := (up_278 uf_273 uf_306 uf_291 uf_305 uf_4) +#12469 := (not #3211) +#23243 := (or #12469 #22307 #18463 #22308 #13518 #18466 #23240) +#23246 := (not #23243) +#23249 := (or #18463 #18466 #23246) +#23252 := (not #23249) +#4 := 1::int +#13412 := (>= uf_294 1::int) +#13552 := (not #13412) +#12368 := (= uf_294 uf_304) +#12515 := (not #12368) +#12524 := (not #12365) +#12360 := (= #3189 uf_300) +#12567 := (not #12360) +decl uf_68 :: (-> T4 T5 T2) +#3186 := (uf_68 uf_273 #3180) decl uf_9 :: T2 #19 := uf_9 -#12812 := (= uf_9 #3181) -#19017 := (not #12812) +#12354 := (= uf_9 #3186) +#18434 := (not #12354) decl uf_48 :: (-> T5 T3 T2) -#3178 := (uf_48 #3175 uf_7) -#12806 := (= uf_9 #3178) -#19011 := (not #12806) -#3246 := (= uf_300 uf_304) -#13090 := (not #3246) +#3181 := (uf_48 #3180 uf_7) +#12345 := (= uf_9 #3181) +#18425 := (not #12345) decl uf_289 :: T1 -#3027 := uf_289 -decl uf_306 :: T1 -#3243 := uf_306 -#3244 := (up_280 uf_273 uf_306 uf_289 uf_294 uf_4) -#13115 := (not #3244) +#3032 := uf_289 +decl uf_302 :: T1 +#3195 := uf_302 +#3196 := (up_278 uf_273 uf_302 uf_289 uf_294 uf_4) +#12549 := (not #3196) decl uf_287 :: T1 -#3024 := uf_287 -decl uf_305 :: T1 -#3241 := uf_305 -#3242 := (up_280 uf_273 uf_305 uf_287 uf_304 uf_7) -#13124 := (not #3242) -#23812 := (or #13124 #13115 #13090 #19011 #19017 #13133 #13081 #14243 #22858 #23803) -#23815 := (not #23812) -#23818 := (or #19011 #19017 #23815) -#23821 := (not #23818) +#3029 := uf_287 +decl uf_301 :: T1 +#3193 := uf_301 +#3194 := (up_278 uf_273 uf_301 uf_287 uf_300 uf_7) +#12558 := (not #3194) +#23255 := (or #12558 #12549 #18425 #18434 #12567 #12524 #12515 #13552 #22307 #23252) +#23258 := (not #23255) +decl uf_25 :: (-> T4 T5 T5) +decl uf_135 :: (-> T14 T5) +decl uf_58 :: (-> T13 T5 T14) +decl uf_59 :: (-> T4 T13) +#3149 := (uf_59 uf_273) +#26583 := (uf_58 #3149 #3180) +#27027 := (uf_135 #26583) +#27032 := (uf_25 uf_273 #27027) +decl uf_26 :: T5 +#77 := uf_26 +#27033 := (= uf_26 #27032) +decl uf_210 :: (-> T4 T5 T2) +#27028 := (uf_210 uf_273 #27027) +#27031 := (= uf_9 #27028) +#27089 := (or #27031 #27033) +#27092 := (not #27089) +decl uf_136 :: (-> T14 T2) +#27042 := (uf_136 #26583) +#27043 := (= uf_9 #27042) +#27044 := (not #27043) decl uf_27 :: (-> T4 T5 T2) -#3176 := (uf_27 uf_273 #3175) -#12803 := (= uf_9 #3176) -#19008 := (not #12803) -#23824 := (or #19008 #19011 #23821) -#23827 := (not #23824) -#23830 := (or #19008 #19011 #23827) -#23833 := (not #23830) -#14208 := (* -1::int #3184) -#14209 := (+ uf_292 #14208) -#14207 := (>= #14209 0::int) -#23836 := (or #22873 #14243 #14207 #23833) -#23839 := (not #23836) -#14211 := (not #14207) -#12826 := (= uf_293 uf_301) -#12993 := (not #12826) -#12823 := (= uf_292 uf_300) -#13002 := (not #12823) -#23806 := (or #13002 #12993 #22873 #14243 #22858 #14211 #23803) -#23809 := (not #23806) -#23842 := (or #23809 #23839) -#23845 := (not #23842) -#23848 := (or #19011 #19017 #22873 #14243 #23845) -#23851 := (not #23848) -#23854 := (or #19011 #19017 #23851) -#23857 := (not #23854) -#23860 := (or #19008 #19011 #23857) -#23863 := (not #23860) -#23866 := (or #19008 #19011 #23863) -#23869 := (not #23866) -#23872 := (or #22873 #14243 #14046 #23869) -#23875 := (not #23872) -#23912 := (or #23875 #23909) -#23915 := (not #23912) -#14431 := (* -1::int uf_292) -#14432 := (+ #3040 #14431) -#14433 := (<= #14432 0::int) -#14421 := (+ #161 #14044) -#14420 := (>= #14421 0::int) -#22774 := (or #5113 #14420 #14433 #20064) -#23754 := (forall (vars (?x774 int)) (:pat #23745) #22774) -#23759 := (not #23754) -#1322 := 255::int -#16349 := (<= uf_292 255::int) -#23043 := (not #16349) -#16332 := (<= uf_293 131073::int) -#23042 := (not #16332) -#16310 := (<= uf_294 131073::int) -#23041 := (not #16310) -#14490 := (>= uf_292 0::int) -#23039 := (not #14490) -#14462 := (>= uf_294 0::int) -#23038 := (not #14462) -#14453 := (>= #14045 0::int) -#14456 := (not #14453) -#14402 := (* -1::int uf_293) -#14403 := (+ uf_272 #14402) -#14404 := (<= #14403 0::int) -#13942 := (<= uf_272 0::int) -decl uf_202 :: (-> T1 T4 T2) -decl uf_295 :: T1 -#3117 := uf_295 -#3118 := (uf_202 uf_295 uf_273) -#12553 := (= uf_9 #3118) -#15709 := (not #12553) -decl uf_177 :: (-> T4 T4 T2) -#3072 := (uf_177 uf_273 uf_273) -#12437 := (= uf_9 #3072) -#14399 := (not #12437) -#3067 := (uf_66 #2960 uf_293 uf_7) -#3068 := (uf_110 uf_273 #3067) -#12426 := (= uf_292 #3068) -#23037 := (not #12426) -decl uf_6 :: (-> T3 T3) -#11 := (uf_6 uf_7) -decl uf_279 :: T1 -#2990 := uf_279 -#3126 := (up_280 uf_273 uf_295 uf_279 #2992 #11) -#23036 := (not #3126) -decl up_278 :: (-> T4 T1 T1 T5 T3 bool) -#3125 := (up_278 uf_273 uf_295 uf_279 #2960 #11) -#23035 := (not #3125) -decl uf_281 :: T1 -#2995 := uf_281 -#3124 := (up_280 uf_273 uf_295 uf_281 uf_272 uf_4) -#13340 := (not #3124) -#3123 := (up_280 uf_273 uf_295 uf_287 uf_292 uf_7) -#13349 := (not #3123) -#3122 := (up_280 uf_273 uf_295 uf_289 uf_293 uf_4) -#13358 := (not #3122) -#3121 := (up_280 uf_273 uf_295 uf_291 uf_294 uf_4) -#13367 := (not #3121) -#3011 := (uf_66 #2960 0::int uf_7) -#3021 := (uf_110 uf_273 #3011) -decl uf_285 :: int -#3020 := uf_285 -#3022 := (= uf_285 #3021) -#13672 := (not #3022) -#23918 := (or #13672 #13367 #13358 #13349 #13340 #23035 #23036 #23037 #14399 #15709 #13942 #22873 #14243 #14404 #14456 #23038 #23039 #23041 #23042 #23043 #23759 #23915) -#23921 := (not #23918) -#23924 := (or #13672 #13942 #23921) -#23927 := (not #23924) -#13922 := (* -1::int #3040) -#13923 := (+ uf_285 #13922) -#13921 := (>= #13923 0::int) -#13910 := (>= #161 1::int) -#22763 := (or #5113 #13910 #13921 #20064) -#23746 := (forall (vars (?x773 int)) (:pat #23745) #22763) -#23751 := (not #23746) -#23930 := (or #23751 #23927) -#23933 := (not #23930) -decl ?x773!13 :: int -#18929 := ?x773!13 -#18939 := (>= ?x773!13 1::int) -#18934 := (uf_66 #2960 ?x773!13 uf_7) -#18935 := (uf_110 uf_273 #18934) -#18936 := (* -1::int #18935) -#18937 := (+ uf_285 #18936) -#18938 := (>= #18937 0::int) -#18931 := (>= ?x773!13 0::int) -#22737 := (not #18931) -#18930 := (<= ?x773!13 131073::int) -#22736 := (not #18930) -#22752 := (or #22736 #22737 #18938 #18939) -#22757 := (not #22752) -#23936 := (or #22757 #23933) -#23939 := (not #23936) -#13903 := (>= uf_272 1::int) -#13906 := (not #13903) -#23942 := (or #13906 #23939) -#23945 := (not #23942) -#23948 := (or #13906 #23945) -#23951 := (not #23948) -#3017 := (uf_67 uf_273 #3011) -#12367 := (= uf_9 #3017) -#18906 := (not #12367) -#3014 := (uf_48 #3011 uf_7) -#12361 := (= uf_9 #3014) -#18900 := (not #12361) -decl uf_290 :: T1 -#3029 := uf_290 -#3031 := (up_280 uf_273 uf_290 uf_291 1::int uf_4) -#13645 := (not #3031) -decl uf_288 :: T1 -#3026 := uf_288 -#3028 := (up_280 uf_273 uf_288 uf_289 0::int uf_4) -#13654 := (not #3028) -decl uf_286 :: T1 -#3023 := uf_286 -#3025 := (up_280 uf_273 uf_286 uf_287 uf_285 uf_7) -#13663 := (not #3025) -#23954 := (or #13672 #13663 #13654 #13645 #18900 #18906 #23951) -#23957 := (not #23954) -#23960 := (or #18900 #18906 #23957) -#23963 := (not #23960) -#3012 := (uf_27 uf_273 #3011) -#12358 := (= uf_9 #3012) -#18897 := (not #12358) -#23966 := (or #18897 #18900 #23963) -#23969 := (not #23966) -#23972 := (or #18897 #18900 #23969) -#23975 := (not #23972) -decl uf_200 :: (-> T4 T5 T5 T16 T2) -decl uf_284 :: T16 -#3008 := uf_284 +#27039 := (uf_27 uf_273 #27027) +#27040 := (= uf_9 #27039) +#27041 := (not #27040) +#27083 := (or #27041 #27044) +#27086 := (not #27083) +decl uf_12 :: (-> T3 T8) +decl uf_13 :: (-> T5 T3) +#26922 := (uf_13 #3180) +#27047 := (uf_12 #26922) +decl uf_14 :: T8 +#28 := uf_14 +#27065 := (= uf_14 #27047) +#27080 := (not #27065) +#27036 := (uf_13 #27027) +#27037 := (uf_12 #27036) +#27038 := (= uf_14 #27037) +#27098 := (or #27038 #27080 #27086 #27092) +#27103 := (not #27098) +#27054 := (uf_25 uf_273 #3180) +#27055 := (= uf_26 #27054) +#27052 := (uf_210 uf_273 #3180) +#27053 := (= uf_9 #27052) +#27068 := (or #27053 #27055) +#27071 := (not #27068) +#27074 := (or #27065 #27071) +#27077 := (not #27074) +#27106 := (or #27077 #27103) +#27109 := (not #27106) +decl uf_24 :: (-> T4 T5 T2) +#3183 := (uf_24 uf_273 #3180) +#12348 := (= uf_9 #3183) +#18428 := (not #12348) +#27112 := (or #18428 #27109) +#27115 := (not #27112) +#27118 := (iff #12354 #27115) +#28635 := (not #27118) +#28695 := [hypothesis]: #28635 +#23 := (:var 0 T5) +#47 := (:var 1 T4) +#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) +#10502 := (= uf_9 #2399) +#2397 := (uf_25 #47 #2388) +#10499 := (= uf_26 #2397) +#10505 := (or #10499 #10502) +#21939 := (not #10505) +#2393 := (uf_13 #2388) +#2394 := (uf_12 #2393) +#10493 := (= uf_14 #2394) +#2389 := (uf_27 #47 #2388) +#10484 := (= uf_9 #2389) +#10487 := (not #10484) +#2385 := (uf_136 #2384) +#10478 := (= uf_9 #2385) +#10481 := (not #10478) +#10490 := (or #10481 #10487) +#21938 := (not #10490) +#26 := (uf_13 #23) +#27 := (uf_12 #26) +#29 := (= #27 uf_14) +#52 := (not #29) +#21940 := (or #52 #21938 #10493 #21939) +#21941 := (not #21940) +#2405 := (uf_210 #47 #23) +#10517 := (= uf_9 #2405) +#142 := (uf_25 #47 #23) +#3639 := (= uf_26 #142) +#10520 := (or #3639 #10517) +#21933 := (not #10520) +#21934 := (or #29 #21933) +#21935 := (not #21934) +#21944 := (or #21935 #21941) +#21950 := (not #21944) +#146 := (uf_24 #47 #23) +#3645 := (= uf_9 #146) +#11090 := (not #3645) +#21951 := (or #11090 #21950) +#21952 := (not #21951) +#10474 := (= uf_9 #2381) +#21957 := (iff #10474 #21952) +#21960 := (forall (vars (?x632 T4) (?x633 T5)) (:pat #2382) #21957) +#10496 := (not #10493) +#10538 := (and #29 #10490 #10496 #10505) +#10523 := (and #52 #10520) +#10544 := (or #10523 #10538) +#10549 := (and #3645 #10544) +#10552 := (iff #10474 #10549) +#10555 := (forall (vars (?x632 T4) (?x633 T5)) (:pat #2382) #10552) +#21961 := (iff #10555 #21960) +#21958 := (iff #10552 #21957) +#21955 := (iff #10549 #21952) +#21947 := (and #3645 #21944) +#21953 := (iff #21947 #21952) +#21954 := [rewrite]: #21953 +#21948 := (iff #10549 #21947) +#21945 := (iff #10544 #21944) +#21942 := (iff #10538 #21941) +#21943 := [rewrite]: #21942 +#21936 := (iff #10523 #21935) +#21937 := [rewrite]: #21936 +#21946 := [monotonicity #21937 #21943]: #21945 +#21949 := [monotonicity #21946]: #21948 +#21956 := [trans #21949 #21954]: #21955 +#21959 := [monotonicity #21956]: #21958 +#21962 := [quant-intro #21959]: #21961 +#17883 := (~ #10555 #10555) +#17881 := (~ #10552 #10552) +#17882 := [refl]: #17881 +#17884 := [nnf-pos #17882]: #17883 +#2406 := (= #2405 uf_9) +#143 := (= #142 uf_26) +#2407 := (or #143 #2406) +#2408 := (and #52 #2407) +#2400 := (= #2399 uf_9) +#2398 := (= #2397 uf_26) +#2401 := (or #2398 #2400) +#2395 := (= #2394 uf_14) +#2396 := (not #2395) +#2402 := (and #2396 #2401) +#2390 := (= #2389 uf_9) +#2391 := (not #2390) +#2386 := (= #2385 uf_9) +#2387 := (not #2386) +#2392 := (or #2387 #2391) +#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) +#10558 := (iff #2412 #10555) +#10508 := (and #10496 #10505) +#10511 := (and #10490 #10508) +#10514 := (and #29 #10511) +#10526 := (or #10514 #10523) +#10529 := (and #3645 #10526) +#10532 := (iff #10474 #10529) +#10535 := (forall (vars (?x632 T4) (?x633 T5)) (:pat #2382) #10532) +#10556 := (iff #10535 #10555) +#10553 := (iff #10532 #10552) +#10550 := (iff #10529 #10549) +#10547 := (iff #10526 #10544) +#10541 := (or #10538 #10523) +#10545 := (iff #10541 #10544) +#10546 := [rewrite]: #10545 +#10542 := (iff #10526 #10541) +#10539 := (iff #10514 #10538) +#10540 := [rewrite]: #10539 +#10543 := [monotonicity #10540]: #10542 +#10548 := [trans #10543 #10546]: #10547 +#10551 := [monotonicity #10548]: #10550 +#10554 := [monotonicity #10551]: #10553 +#10557 := [quant-intro #10554]: #10556 +#10536 := (iff #2412 #10535) +#10533 := (iff #2411 #10532) +#10530 := (iff #2410 #10529) +#10527 := (iff #2409 #10526) +#10524 := (iff #2408 #10523) +#10521 := (iff #2407 #10520) +#10518 := (iff #2406 #10517) +#10519 := [rewrite]: #10518 +#3640 := (iff #143 #3639) +#3641 := [rewrite]: #3640 +#10522 := [monotonicity #3641 #10519]: #10521 +#10525 := [monotonicity #10522]: #10524 +#10515 := (iff #2404 #10514) +#10512 := (iff #2403 #10511) +#10509 := (iff #2402 #10508) +#10506 := (iff #2401 #10505) +#10503 := (iff #2400 #10502) +#10504 := [rewrite]: #10503 +#10500 := (iff #2398 #10499) +#10501 := [rewrite]: #10500 +#10507 := [monotonicity #10501 #10504]: #10506 +#10497 := (iff #2396 #10496) +#10494 := (iff #2395 #10493) +#10495 := [rewrite]: #10494 +#10498 := [monotonicity #10495]: #10497 +#10510 := [monotonicity #10498 #10507]: #10509 +#10491 := (iff #2392 #10490) +#10488 := (iff #2391 #10487) +#10485 := (iff #2390 #10484) +#10486 := [rewrite]: #10485 +#10489 := [monotonicity #10486]: #10488 +#10482 := (iff #2387 #10481) +#10479 := (iff #2386 #10478) +#10480 := [rewrite]: #10479 +#10483 := [monotonicity #10480]: #10482 +#10492 := [monotonicity #10483 #10489]: #10491 +#10513 := [monotonicity #10492 #10510]: #10512 +#10516 := [monotonicity #10513]: #10515 +#10528 := [monotonicity #10516 #10525]: #10527 +#3646 := (iff #147 #3645) +#3647 := [rewrite]: #3646 +#10531 := [monotonicity #3647 #10528]: #10530 +#10476 := (iff #2383 #10474) +#10477 := [rewrite]: #10476 +#10534 := [monotonicity #10477 #10531]: #10533 +#10537 := [quant-intro #10534]: #10536 +#10559 := [trans #10537 #10557]: #10558 +#10473 := [asserted]: #2412 +#10560 := [mp #10473 #10559]: #10555 +#17885 := [mp~ #10560 #17884]: #10555 +#21963 := [mp #17885 #21962]: #21960 +#27179 := (not #21960) +#28611 := (or #27179 #27118) +#27034 := (or #27033 #27031) +#27035 := (not #27034) +#27045 := (or #27044 #27041) +#27046 := (not #27045) +#27048 := (= #27047 uf_14) +#27049 := (not #27048) +#27050 := (or #27049 #27046 #27038 #27035) +#27051 := (not #27050) +#27056 := (or #27055 #27053) +#27057 := (not #27056) +#27058 := (or #27048 #27057) +#27059 := (not #27058) +#27060 := (or #27059 #27051) +#27061 := (not #27060) +#27062 := (or #18428 #27061) +#27063 := (not #27062) +#27064 := (iff #12354 #27063) +#28614 := (or #27179 #27064) +#28616 := (iff #28614 #28611) +#28601 := (iff #28611 #28611) +#28602 := [rewrite]: #28601 +#27119 := (iff #27064 #27118) +#27116 := (iff #27063 #27115) +#27113 := (iff #27062 #27112) +#27110 := (iff #27061 #27109) +#27107 := (iff #27060 #27106) +#27104 := (iff #27051 #27103) +#27101 := (iff #27050 #27098) +#27095 := (or #27080 #27086 #27038 #27092) +#27099 := (iff #27095 #27098) +#27100 := [rewrite]: #27099 +#27096 := (iff #27050 #27095) +#27093 := (iff #27035 #27092) +#27090 := (iff #27034 #27089) +#27091 := [rewrite]: #27090 +#27094 := [monotonicity #27091]: #27093 +#27087 := (iff #27046 #27086) +#27084 := (iff #27045 #27083) +#27085 := [rewrite]: #27084 +#27088 := [monotonicity #27085]: #27087 +#27081 := (iff #27049 #27080) +#27066 := (iff #27048 #27065) +#27067 := [rewrite]: #27066 +#27082 := [monotonicity #27067]: #27081 +#27097 := [monotonicity #27082 #27088 #27094]: #27096 +#27102 := [trans #27097 #27100]: #27101 +#27105 := [monotonicity #27102]: #27104 +#27078 := (iff #27059 #27077) +#27075 := (iff #27058 #27074) +#27072 := (iff #27057 #27071) +#27069 := (iff #27056 #27068) +#27070 := [rewrite]: #27069 +#27073 := [monotonicity #27070]: #27072 +#27076 := [monotonicity #27067 #27073]: #27075 +#27079 := [monotonicity #27076]: #27078 +#27108 := [monotonicity #27079 #27105]: #27107 +#27111 := [monotonicity #27108]: #27110 +#27114 := [monotonicity #27111]: #27113 +#27117 := [monotonicity #27114]: #27116 +#27120 := [monotonicity #27117]: #27119 +#28600 := [monotonicity #27120]: #28616 +#28603 := [trans #28600 #28602]: #28616 +#28615 := [quant-inst]: #28614 +#28604 := [mp #28615 #28603]: #28611 +#28729 := [unit-resolution #28604 #21963 #28695]: false +#28730 := [lemma #28729]: #27118 +#29058 := (or #28635 #12354) +#28658 := [hypothesis]: #27098 decl uf_116 :: (-> T5 int) #2961 := (uf_116 #2960) decl uf_124 :: (-> T3 int T3) -#2952 := (uf_124 uf_7 uf_272) -#2962 := (uf_43 #2952 #2961) -#3009 := (uf_200 uf_273 #2962 #2962 uf_284) -#12355 := (= uf_9 #3009) -#13715 := (not #12355) -#23978 := (or #13715 #23975) -#23981 := (not #23978) -decl uf_14 :: (-> T3 T8) -#24016 := (uf_116 #2962) -#25404 := (uf_43 #2952 #24016) -#25815 := (uf_15 #25404) -#26092 := (uf_14 #25815) -decl uf_16 :: T8 -#35 := uf_16 -#26095 := (= uf_16 #26092) -#26297 := (not #26095) -#2955 := (uf_14 #2952) -#12296 := (= uf_16 #2955) -#12299 := (not #12296) -#26298 := (iff #12299 #26297) -#26293 := (iff #12296 #26095) -#26342 := (iff #26095 #12296) -#26340 := (= #26092 #2955) -#26338 := (= #25815 #2952) -#24234 := (uf_15 #2962) -#28358 := (= #24234 #2952) -#24237 := (= #2952 #24234) +#2958 := (uf_124 uf_7 uf_272) +#2962 := (uf_43 #2958 #2961) +#2965 := (uf_25 uf_273 #2962) +#28691 := (= #2965 #27032) +#28541 := (= #27032 #2965) +#29243 := (= #27027 #2962) +decl uf_143 :: (-> T3 int) +#23568 := (uf_143 #2958) +decl uf_144 :: (-> T3 T3) +#23566 := (uf_144 #2958) +#25879 := (uf_124 #23566 #23568) +#25880 := (uf_43 #25879 #2961) +#25867 := (= #25880 #2962) +#25850 := (= #25879 #2958) +#25848 := (= #23568 uf_272) +#23569 := (= uf_272 #23568) #326 := (:var 1 T3) -#2692 := (uf_43 #326 #161) -#23682 := (pattern #2692) -#2696 := (uf_15 #2692) -#11677 := (= #326 #2696) -#23689 := (forall (vars (?x720 T3) (?x721 int)) (:pat #23682) #11677) -#11681 := (forall (vars (?x720 T3) (?x721 int)) #11677) -#23692 := (iff #11681 #23689) -#23690 := (iff #11677 #11677) -#23691 := [refl]: #23690 -#23693 := [quant-intro #23691]: #23692 -#18759 := (~ #11681 #11681) -#18757 := (~ #11677 #11677) -#18758 := [refl]: #18757 -#18760 := [nnf-pos #18758]: #18759 -#2697 := (= #2696 #326) -#2698 := (forall (vars (?x720 T3) (?x721 int)) #2697) -#11682 := (iff #2698 #11681) -#11679 := (iff #2697 #11677) -#11680 := [rewrite]: #11679 -#11683 := [quant-intro #11680]: #11682 -#11676 := [asserted]: #2698 -#11686 := [mp #11676 #11683]: #11681 -#18761 := [mp~ #11686 #18760]: #11681 -#23694 := [mp #18761 #23693]: #23689 -#24181 := (not #23689) -#24242 := (or #24181 #24237) -#24243 := [quant-inst]: #24242 -#28006 := [unit-resolution #24243 #23694]: #24237 -#28359 := [symm #28006]: #28358 -#26336 := (= #25815 #24234) -#27940 := (= #25404 #2962) -#25411 := (= #2962 #25404) -#2965 := (uf_48 #2962 #2952) -#12305 := (= uf_9 #2965) -decl uf_24 :: (-> T4 T5 T2) +#1358 := (uf_124 #326 #161) +#1592 := (pattern #1358) +#1602 := (uf_143 #1358) +#8288 := (= #161 #1602) +#8291 := (forall (vars (?x386 T3) (?x387 int)) (:pat #1592) #8288) +#17259 := (~ #8291 #8291) +#17257 := (~ #8288 #8288) +#17258 := [refl]: #17257 +#17260 := [nnf-pos #17258]: #17259 +#1603 := (= #1602 #161) +#1604 := (forall (vars (?x386 T3) (?x387 int)) (:pat #1592) #1603) +#8292 := (iff #1604 #8291) +#8289 := (iff #1603 #8288) +#8290 := [rewrite]: #8289 +#8293 := [quant-intro #8290]: #8292 +#8287 := [asserted]: #1604 +#8296 := [mp #8287 #8293]: #8291 +#17261 := [mp~ #8296 #17260]: #8291 +#23575 := (not #8291) +#23576 := (or #23575 #23569) +#23577 := [quant-inst]: #23576 +#26235 := [unit-resolution #23577 #17261]: #23569 +#25849 := [symm #26235]: #25848 +#25689 := (= #23566 uf_7) +#23567 := (= uf_7 #23566) +#1605 := (uf_144 #1358) +#8295 := (= #326 #1605) +#8299 := (forall (vars (?x388 T3) (?x389 int)) (:pat #1592) #8295) +#17264 := (~ #8299 #8299) +#17262 := (~ #8295 #8295) +#17263 := [refl]: #17262 +#17265 := [nnf-pos #17263]: #17264 +#1606 := (= #1605 #326) +#1607 := (forall (vars (?x388 T3) (?x389 int)) (:pat #1592) #1606) +#8300 := (iff #1607 #8299) +#8297 := (iff #1606 #8295) +#8298 := [rewrite]: #8297 +#8301 := [quant-intro #8298]: #8300 +#8294 := [asserted]: #1607 +#8304 := [mp #8294 #8301]: #8299 +#17266 := [mp~ #8304 #17265]: #8299 +#23570 := (not #8299) +#23571 := (or #23570 #23567) +#23572 := [quant-inst]: #23571 +#25688 := [unit-resolution #23572 #17266]: #23567 +#25690 := [symm #25688]: #25689 +#25866 := [monotonicity #25690 #25849]: #25850 +#25865 := [monotonicity #25866]: #25867 +#29241 := (= #27027 #25880) +decl uf_125 :: (-> T5 T5 int) +decl uf_28 :: (-> int T5) +decl uf_29 :: (-> T5 int) +#2996 := (uf_29 #2960) +#22665 := (uf_28 #2996) +#25805 := (uf_116 #22665) +#25821 := (uf_43 #23566 #25805) +#26356 := (uf_13 #25821) +#27024 := (uf_66 #25821 uf_294 #26356) +#27025 := (uf_125 #27024 #25821) +#27192 := (uf_66 #25880 #27025 #23566) +#27196 := (uf_58 #3149 #27192) +#27199 := (uf_135 #27196) +#29239 := (= #27199 #25880) +#27200 := (= #25880 #27199) +decl up_67 :: (-> T14 bool) +#27202 := (up_67 #27196) +#27203 := (not #27202) +#27201 := (not #27200) +#27197 := (uf_136 #27196) +#27198 := (= uf_9 #27197) +#27193 := (uf_24 uf_273 #27192) +#27194 := (= uf_9 #27193) +#27195 := (not #27194) +#27224 := (or #27195 #27198 #27201 #27203) +#27227 := (not #27224) +#25895 := (uf_24 uf_273 #25880) +#25896 := (= uf_9 #25895) #2969 := (uf_24 uf_273 #2962) -#12311 := (= uf_9 #2969) -decl uf_25 :: (-> T4 T5 T5) -#2967 := (uf_25 uf_273 #2962) -decl uf_26 :: T5 -#78 := uf_26 -#12308 := (= uf_26 #2967) +#27444 := (= #2969 #25895) +#27389 := (= #25895 #2969) +#27448 := [monotonicity #25865]: #27389 +#27445 := [symm #27448]: #27444 +#11875 := (= uf_9 #2969) +decl uf_23 :: (-> T3 T2) +#2974 := (uf_23 #2958) +#11884 := (= uf_9 #2974) +#2971 := (uf_12 #2958) +#11878 := (= uf_14 #2971) +#11881 := (not #11878) +#2967 := (uf_48 #2962 #2958) +#11872 := (= uf_9 #2967) +#11869 := (= uf_26 #2965) #2963 := (uf_27 uf_273 #2962) -#12302 := (= uf_9 #2963) -decl uf_22 :: (-> T3 T2) -#2953 := (uf_22 #2952) -#12293 := (= uf_9 #2953) -#14658 := (and #12293 #12299 #12302 #12305 #12308 #12311) +#11866 := (= uf_9 #2963) +#14124 := (and #11866 #11869 #11872 #11875 #11881 #11884) decl uf_269 :: int -#2937 := uf_269 -#14715 := (>= uf_269 0::int) -#14711 := (* -1::int uf_269) +#2942 := uf_269 +#14180 := (* -1::int uf_269) decl uf_78 :: int #429 := uf_78 -#14712 := (+ uf_78 #14711) -#14710 := (>= #14712 0::int) -#14718 := (and #14710 #14715) -#14721 := (not #14718) +#14181 := (+ uf_78 #14180) +#14179 := (>= #14181 0::int) +#14177 := (>= uf_269 0::int) +#14184 := (and #14177 #14179) +#14187 := (not #14184) decl uf_270 :: int -#2941 := uf_270 -#14701 := (>= uf_270 0::int) -#14697 := (* -1::int uf_270) +#2946 := uf_270 +#14166 := (* -1::int uf_270) decl uf_76 :: int #409 := uf_76 -#14698 := (+ uf_76 #14697) -#14696 := (>= #14698 0::int) -#14704 := (and #14696 #14701) -#14707 := (not #14704) +#14167 := (+ uf_76 #14166) +#14165 := (>= #14167 0::int) +#14163 := (>= uf_270 0::int) +#14170 := (and #14163 #14165) +#14173 := (not #14170) decl uf_271 :: int -#2945 := uf_271 -#14687 := (>= uf_271 0::int) -#14683 := (* -1::int uf_271) -#14684 := (+ uf_76 #14683) -#14682 := (>= #14684 0::int) -#14690 := (and #14682 #14687) -#14693 := (not #14690) -#974 := 1099511627776::int -#14671 := (>= uf_272 1099511627776::int) -#14661 := (not #14658) +#2950 := uf_271 +#14152 := (* -1::int uf_271) +#14153 := (+ uf_76 #14152) +#14151 := (>= #14153 0::int) +#14149 := (>= uf_271 0::int) +#14156 := (and #14149 #14151) +#14159 := (not #14156) +#1042 := 1099511627776::int +#14137 := (>= uf_272 1099511627776::int) +#14127 := (not #14124) decl uf_276 :: (-> T19 int) -#2984 := (:var 0 T19) -#2985 := (uf_276 #2984) -#2986 := (pattern #2985) +#2989 := (:var 0 T19) +#2990 := (uf_276 #2989) +#2991 := (pattern #2990) decl uf_277 :: int -#2987 := uf_277 -#14648 := (* -1::int uf_277) -#14649 := (+ #2985 #14648) -#14647 := (>= #14649 0::int) -#14646 := (not #14647) -#14652 := (forall (vars (?x771 T19)) (:pat #2986) #14646) -#14655 := (not #14652) -#13943 := (not #13942) -#14502 := (and #3022 #13943) -#14507 := (not #14502) -#14487 := (+ uf_78 #14431) -#14486 := (>= #14487 0::int) -#14493 := (and #14486 #14490) -#14496 := (not #14493) -#14472 := (+ uf_76 #14402) -#14471 := (>= #14472 0::int) -#14478 := (and #13947 #14471) -#14483 := (not #14478) -#14085 := (+ uf_76 #14044) -#14459 := (>= #14085 0::int) -#14465 := (and #14459 #14462) -#14468 := (not #14465) -#4413 := (* -1::int uf_76) -#4418 := (+ #161 #4413) -#4419 := (<= #4418 0::int) -#5736 := (and #4084 #4419) -#5739 := (not #5736) -#14442 := (or #5739 #14420 #14433) -#14447 := (forall (vars (?x774 int)) #14442) -#14450 := (not #14447) -#14405 := (not #14404) -#14411 := (and #12426 #14405) -#14416 := (not #14411) -#14084 := (>= #14085 1::int) -#14175 := (and #14084 #14088) -#14178 := (not #14175) -#14151 := (and #12862 #14145) -#14131 := (or #5739 #14109 #14122) -#14136 := (forall (vars (?x785 int)) #14131) -#14139 := (not #14136) -#14156 := (or #14139 #14151) -#14159 := (and #14136 #14156) -#14162 := (or #14105 #14159) -#14165 := (and #14100 #14162) -#14094 := (and #14076 #14092) -#14097 := (not #14094) -#14193 := (or #12942 #14097 #14165 #14172 #14178) -#14201 := (and #14084 #14088 #14193) -#14078 := (and #13950 #14076) -#14081 := (not #14078) -#12818 := (and #12806 #12812) -#13142 := (not #12818) -#14267 := (or #13124 #13115 #13090 #13142 #13133 #13081 #14243 #14081 #14201) -#14275 := (and #12806 #12812 #14267) -#12809 := (and #12803 #12806) -#13159 := (not #12809) -#14280 := (or #13159 #14275) -#14286 := (and #12803 #12806 #14280) -#13952 := (and #13947 #13950) -#13955 := (not #13952) -#14312 := (or #13955 #14207 #14286) -#14238 := (or #13002 #12993 #13955 #14081 #14201 #14211) -#14317 := (and #14238 #14312) -#14326 := (or #13142 #13955 #14317) -#14334 := (and #12806 #12812 #14326) -#14339 := (or #13159 #14334) -#14345 := (and #12803 #12806 #14339) -#14371 := (or #13955 #14046 #14345) -#13958 := (not #13959) -#13998 := (and #3145 #4084 #4419 #13958) -#14003 := (exists (vars (?x782 int)) #13998) -#13981 := (or #5739 #13959 #13972) -#13986 := (forall (vars (?x781 int)) #13981) -#13989 := (not #13986) -#14006 := (or #13989 #14003) -#14009 := (and #13986 #14006) +#2992 := uf_277 +#14114 := (* -1::int uf_277) +#14115 := (+ #2990 #14114) +#14113 := (>= #14115 0::int) +#14112 := (not #14113) +#14118 := (forall (vars (?x771 T19)) (:pat #2991) #14112) +#14121 := (not #14118) +#13404 := (<= uf_272 0::int) +#13405 := (not #13404) +#3016 := (uf_66 #2960 0::int uf_7) +#3026 := (uf_110 uf_273 #3016) +decl uf_285 :: int +#3025 := uf_285 +#3027 := (= uf_285 #3026) +#13968 := (and #3027 #13405) +#13973 := (not #13968) +decl uf_292 :: int +#3052 := uf_292 +#13902 := (* -1::int uf_292) +#13956 := (+ uf_78 #13902) +#13955 := (>= #13956 0::int) +#13952 := (>= uf_292 0::int) +#13959 := (and #13952 #13955) +#13962 := (not #13959) +decl uf_293 :: int +#3056 := uf_293 +#13873 := (* -1::int uf_293) +#13943 := (+ uf_76 #13873) +#13942 := (>= #13943 0::int) +#13409 := (>= uf_293 0::int) +#13946 := (and #13409 #13942) +#13949 := (not #13946) +#13433 := (* -1::int uf_294) +#13434 := (+ uf_76 #13433) +#13933 := (>= #13434 0::int) +#13930 := (>= uf_294 0::int) +#13936 := (and #13930 #13933) +#13939 := (not #13936) +#13696 := (+ uf_272 #13433) +#13924 := (>= #13696 0::int) +#13927 := (not #13924) +#13903 := (+ #3045 #13902) +#13904 := (<= #13903 0::int) +#13891 := (+ #161 #13433) +#13890 := (>= #13891 0::int) +#4377 := (* -1::int uf_76) +#4378 := (+ #161 #4377) +#4379 := (<= #4378 0::int) +#4386 := (and #4065 #4379) +#5601 := (not #4386) +#13913 := (or #5601 #13890 #13904) +#13918 := (forall (vars (?x775 int)) #13913) +#13921 := (not #13918) +#13874 := (+ uf_272 #13873) +#13875 := (<= #13874 0::int) +#13876 := (not #13875) +#3073 := (uf_66 #2960 uf_293 uf_7) +#3074 := (uf_110 uf_273 #3073) +#11992 := (= uf_292 #3074) +#13882 := (and #11992 #13876) +#13887 := (not #13882) +#13338 := (* -1::int uf_272) +#13726 := (+ #161 #13338) +#13725 := (>= #13726 0::int) +#13727 := (not #13725) +decl uf_299 :: int +#3088 := uf_299 +#3095 := (= #3045 uf_299) +#13765 := (and #3095 #4065 #4379 #13727) +#13770 := (exists (vars (?x778 int)) #13765) +#13737 := (* -1::int uf_299) +#13738 := (+ #3045 #13737) +#13739 := (<= #13738 0::int) +#13748 := (or #5601 #13725 #13739) +#13753 := (forall (vars (?x776 int)) #13748) +#13773 := (not #13753) +#13779 := (or #13773 #13770) +#13784 := (and #13753 #13779) +#13414 := (and #13409 #13412) +#13417 := (not #13414) +#12020 := (= uf_292 uf_299) +#12069 := (not #12020) +decl uf_298 :: int +#3086 := uf_298 +#12017 := (= uf_293 uf_298) +#12078 := (not #12017) +decl uf_297 :: int +#3084 := uf_297 +#12014 := (= uf_294 uf_297) +#12087 := (not #12014) +decl uf_296 :: int +#3082 := uf_296 +#12011 := (= uf_292 uf_296) +#12096 := (not #12011) decl up_216 :: bool -#2477 := up_216 -#12719 := (not up_216) -#14036 := (or #12719 #12671 #12662 #12653 #12644 #13955 #14009) -#14041 := (and up_216 #14036) -#14070 := (or #13955 #14041 #14049) -#14376 := (and #14070 #14371) +#2482 := up_216 +#12144 := (not up_216) +#13811 := (or #12144 #12096 #12087 #12078 #12069 #13417 #13784) +#13816 := (and up_216 #13811) +#13697 := (<= #13696 0::int) +#13698 := (not #13697) +#13841 := (or #13417 #13698 #13816) +#13605 := (* -1::int #3189) +#13606 := (+ uf_292 #13605) +#13604 := (>= #13606 0::int) +#13603 := (not #13604) +#13432 := (>= #13434 1::int) +#13521 := (and #13430 #13432) +#13524 := (not #13521) +#13491 := (not #13490) +#13497 := (and #12404 #13491) +#13477 := (or #5601 #13454 #13468) +#13482 := (forall (vars (?x786 int)) #13477) +#13485 := (not #13482) +#13502 := (or #13485 #13497) +#13505 := (and #13482 #13502) +#13508 := (or #13451 #13505) +#13511 := (and #13446 #13508) +#13440 := (and #13421 #13438) +#13443 := (not #13440) +#13539 := (or #12469 #13443 #13511 #13518 #13524) +#13547 := (and #13430 #13432 #13539) +#13423 := (and #13412 #13421) +#13426 := (not #13423) +#12647 := (= uf_293 uf_304) +#12653 := (not #12647) +#12644 := (= uf_292 uf_303) +#12662 := (not #12644) +#13658 := (or #12662 #12653 #13417 #13426 #13547 #13603) +#12357 := (and #12345 #12354) +#12576 := (not #12357) +#13579 := (or #12558 #12549 #12576 #12567 #12524 #12515 #13552 #13426 #13547) +#13587 := (and #12345 #12354 #13579) +#12351 := (and #12345 #12348) +#12588 := (not #12351) +#13592 := (or #12588 #13587) +#13598 := (and #12345 #12348 #13592) +#13628 := (or #13417 #13598 #13604) +#13663 := (and #13628 #13658) +#13672 := (or #12576 #13417 #13663) +#13680 := (and #12345 #12354 #13672) +#13685 := (or #12588 #13680) +#13691 := (and #12345 #12348 #13685) +#13720 := (or #13417 #13691 #13697) +#13846 := (and #13720 #13841) +decl uf_178 :: (-> T4 T4 T2) +#3161 := (uf_178 uf_273 uf_273) +#12306 := (= uf_9 #3161) +#13870 := (not #12306) +decl uf_202 :: (-> T1 T4 T2) +decl uf_295 :: T1 +#3078 := uf_295 +#3079 := (uf_202 uf_295 uf_273) +#12000 := (= uf_9 #3079) decl uf_55 :: (-> T4 T2) -#2978 := (uf_55 uf_273) -#12332 := (= uf_9 #2978) -#12556 := (and #12332 #12553) -#13376 := (not #12556) -#3127 := (and #3125 #3126) -#13331 := (not #3127) -#14573 := (or #13367 #13358 #13349 #13340 #13331 #14399 #13376 #13955 #14376 #14416 #14450 #14456 #14468 #14483 #14496 #14507) -#14581 := (and #3022 #13943 #14573) -#13931 := (or #5739 #13910 #13921) -#13936 := (forall (vars (?x773 int)) #13931) -#13939 := (not #13936) -#14586 := (or #13939 #14581) -#14589 := (and #13936 #14586) -#14592 := (or #13906 #14589) -#14595 := (and #13903 #14592) -#12373 := (and #12361 #12367) -#13681 := (not #12373) -#14616 := (or #13672 #13663 #13654 #13645 #13681 #14595) -#14624 := (and #12361 #12367 #14616) -#12364 := (and #12358 #12361) -#13698 := (not #12364) -#14629 := (or #13698 #14624) -#14635 := (and #12358 #12361 #14629) -#14640 := (or #13715 #14635) -#14643 := (and #12355 #14640) -#13877 := (>= uf_272 0::int) -#13874 := (+ uf_76 #13873) -#13872 := (>= #13874 0::int) -#13880 := (and #13872 #13877) -#13883 := (not #13880) +#2986 := (uf_55 uf_273) +#11908 := (= uf_9 #2986) +#12006 := (and #11908 #12000) +#12179 := (not #12006) +decl up_280 :: (-> T4 T1 T1 T5 T3 bool) +decl uf_6 :: (-> T3 T3) +#11 := (uf_6 uf_7) +decl uf_279 :: T1 +#2995 := uf_279 +#3174 := (up_280 uf_273 uf_295 uf_279 #2960 #11) +#3173 := (up_278 uf_273 uf_295 uf_279 #2996 #11) +#3175 := (and #3173 #3174) +#12878 := (not #3175) +decl uf_281 :: T1 +#3000 := uf_281 +#3172 := (up_278 uf_273 uf_295 uf_281 uf_272 uf_4) +#12887 := (not #3172) +#3171 := (up_278 uf_273 uf_295 uf_287 uf_292 uf_7) +#12896 := (not #3171) +#3170 := (up_278 uf_273 uf_295 uf_289 uf_293 uf_4) +#12905 := (not #3170) +#3169 := (up_278 uf_273 uf_295 uf_291 uf_294 uf_4) +#12914 := (not #3169) +#14039 := (or #12914 #12905 #12896 #12887 #12878 #12179 #13870 #13417 #13846 #13887 #13921 #13927 #13939 #13949 #13962 #13973) +#14047 := (and #3027 #13405 #14039) +#13384 := (* -1::int #3045) +#13385 := (+ uf_285 #13384) +#13383 := (>= #13385 0::int) +#13371 := (>= #161 1::int) +#13393 := (or #5601 #13371 #13383) +#13398 := (forall (vars (?x773 int)) #13393) +#13401 := (not #13398) +#14052 := (or #13401 #14047) +#14055 := (and #13398 #14052) +#13365 := (>= uf_272 1::int) +#13368 := (not #13365) +#14058 := (or #13368 #14055) +#14061 := (and #13365 #14058) +#3022 := (uf_68 uf_273 #3016) +#11940 := (= uf_9 #3022) +#3017 := (uf_48 #3016 uf_7) +#11931 := (= uf_9 #3017) +#11943 := (and #11931 #11940) +#13158 := (not #11943) +decl uf_290 :: T1 +#3034 := uf_290 +#3036 := (up_278 uf_273 uf_290 uf_291 1::int uf_4) +#13122 := (not #3036) +decl uf_288 :: T1 +#3031 := uf_288 +#3033 := (up_278 uf_273 uf_288 uf_289 0::int uf_4) +#13131 := (not #3033) +decl uf_286 :: T1 +#3028 := uf_286 +#3030 := (up_278 uf_273 uf_286 uf_287 uf_285 uf_7) +#13140 := (not #3030) +#13149 := (not #3027) +#14082 := (or #13149 #13140 #13131 #13122 #13158 #14061) +#14090 := (and #11931 #11940 #14082) +#3019 := (uf_24 uf_273 #3016) +#11934 := (= uf_9 #3019) +#11937 := (and #11931 #11934) +#13170 := (not #11937) +#14095 := (or #13170 #14090) +#14101 := (and #11931 #11934 #14095) +decl uf_200 :: (-> T4 T5 T5 T16 T2) +decl uf_284 :: T16 +#3013 := uf_284 +#3014 := (uf_200 uf_273 #2962 #2962 uf_284) +#11928 := (= uf_9 #3014) +#13182 := (not #11928) +#14106 := (or #13182 #14101) +#14109 := (and #11928 #14106) +#13339 := (+ uf_76 #13338) +#13337 := (>= #13339 0::int) +#13335 := (>= uf_272 0::int) +#13342 := (and #13335 #13337) +#13345 := (not #13342) decl uf_283 :: (-> int T5 T2) -#26 := (:var 0 T5) decl uf_282 :: int -#2997 := uf_282 -#3000 := (uf_283 uf_282 #26) -#3001 := (pattern #3000) -#12341 := (= uf_9 #3000) -#12347 := (not #12341) -#12352 := (forall (vars (?x772 T5)) (:pat #3001) #12347) -#13741 := (not #12352) +#3002 := uf_282 +#3005 := (uf_283 uf_282 #23) +#3006 := (pattern #3005) +#11914 := (= uf_9 #3005) +#11920 := (not #11914) +#11925 := (forall (vars (?x772 T5)) (:pat #3006) #11920) +#13203 := (not #11925) decl uf_275 :: T1 -#2980 := uf_275 -#2981 := (uf_202 uf_275 uf_273) -#12335 := (= uf_9 #2981) -#12338 := (and #12332 #12335) -#13786 := (not #12338) +#2983 := uf_275 +#2984 := (uf_202 uf_275 uf_273) +#11905 := (= uf_9 #2984) +#11911 := (and #11905 #11908) +#13248 := (not #11911) decl uf_203 :: (-> T4 T2) -#2976 := (uf_203 uf_273) -#12329 := (= uf_9 #2976) -#13795 := (not #12329) -decl uf_171 :: (-> T4 int) -#2998 := (uf_171 uf_273) -#2999 := (= uf_282 #2998) -#13750 := (not #2999) -#2996 := (up_280 uf_273 uf_275 uf_281 uf_272 uf_4) -#13759 := (not #2996) -#2993 := (up_280 uf_273 uf_275 uf_279 #2992 #11) -#2991 := (up_278 uf_273 uf_275 uf_279 #2960 #11) -#2994 := (and #2991 #2993) -#13768 := (not #2994) -#14766 := (or #13768 #13759 #13750 #13795 #13786 #13741 #13883 #13942 #14643 #14655 #14661 #14671 #14693 #14707 #14721) -#14771 := (not #14766) -#3010 := (= #3009 uf_9) -#3015 := (= #3014 uf_9) -#3013 := (= #3012 uf_9) -#3016 := (and #3013 #3015) -#3018 := (= #3017 uf_9) -#3019 := (and #3018 #3015) -#3037 := (<= 1::int uf_272) -#3041 := (<= #3040 uf_285) -#3038 := (< #161 1::int) -#3042 := (implies #3038 #3041) +#2981 := (uf_203 uf_273) +#11902 := (= uf_9 #2981) +#13257 := (not #11902) +decl uf_173 :: (-> T4 int) +#3003 := (uf_173 uf_273) +#3004 := (= uf_282 #3003) +#13212 := (not #3004) +#3001 := (up_278 uf_273 uf_275 uf_281 uf_272 uf_4) +#13221 := (not #3001) +#2998 := (up_280 uf_273 uf_275 uf_279 #2960 #11) +#2997 := (up_278 uf_273 uf_275 uf_279 #2996 #11) +#2999 := (and #2997 #2998) +#13230 := (not #2999) +#14232 := (or #13230 #13221 #13212 #13257 #13248 #13203 #13345 #13404 #14109 #14121 #14127 #14137 #14159 #14173 #14187) +#14237 := (not #14232) +#1 := true +#3090 := (< #161 uf_272) +#3096 := (and #3090 #3095) +#411 := (<= #161 uf_76) +#3097 := (and #411 #3096) #285 := (<= 0::int #161) -#410 := (<= #161 uf_76) -#645 := (and #410 #285) -#3043 := (implies #645 #3042) -#3044 := (forall (vars (?x773 int)) #3043) -#2951 := (< 0::int uf_272) -#3045 := (= #3021 uf_285) -#3046 := (and #3045 #2951) -#3141 := (<= #3040 uf_299) -#3140 := (< #161 uf_272) -#3142 := (implies #3140 #3141) -#3143 := (implies #645 #3142) -#3144 := (forall (vars (?x781 int)) #3143) -#3146 := (and #3140 #645) -#3147 := (and #3145 #3146) -#3148 := (exists (vars (?x782 int)) #3147) -#1 := true -#3149 := (implies #3148 true) -#3150 := (and #3149 #3148) -#3151 := (implies #3144 #3150) -#3152 := (and #3151 #3144) -#3153 := (implies true #3152) -#3139 := (= uf_299 uf_292) -#3154 := (implies #3139 #3153) -#3137 := (= uf_298 uf_293) -#3155 := (implies #3137 #3154) -#3135 := (= uf_297 uf_294) -#3156 := (implies #3135 #3155) -#3133 := (= uf_296 uf_292) -#3157 := (implies #3133 #3156) -#3158 := (implies true #3157) -#3059 := (<= 1::int uf_294) -#3053 := (<= 0::int uf_293) -#3060 := (and #3053 #3059) -#3159 := (implies #3060 #3158) -#3160 := (implies #3060 #3159) -#3161 := (implies true #3160) -#3162 := (implies #3060 #3161) -#3163 := (implies up_216 #3162) -#3164 := (and #3163 up_216) -#3165 := (implies #3060 #3164) -#3166 := (implies true #3165) -#3167 := (implies #3060 #3166) -#3119 := (= #3118 uf_9) -#2979 := (= #2978 uf_9) -#3120 := (and #2979 #3119) -#3295 := (implies #3120 #3167) -#3296 := (implies #3060 #3295) -#3297 := (implies true #3296) -#3298 := (implies #3060 #3297) -#3294 := (not true) -#3299 := (implies #3294 #3298) -#3300 := (implies #3060 #3299) -#3301 := (implies true #3300) -#3179 := (= #3178 uf_9) -#3177 := (= #3176 uf_9) -#3180 := (and #3177 #3179) +#3098 := (and #285 #3097) +#3099 := (exists (vars (?x778 int)) #3098) +#3100 := (implies #3099 true) +#3101 := (and #3099 #3100) +#3091 := (<= #3045 uf_299) +#3092 := (implies #3090 #3091) +#412 := (and #285 #411) +#3093 := (implies #412 #3092) +#3094 := (forall (vars (?x776 int)) #3093) +#3102 := (implies #3094 #3101) +#3103 := (and #3094 #3102) +#3104 := (implies true #3103) +#3089 := (= uf_299 uf_292) +#3105 := (implies #3089 #3104) +#3087 := (= uf_298 uf_293) +#3106 := (implies #3087 #3105) +#3085 := (= uf_297 uf_294) +#3107 := (implies #3085 #3106) +#3083 := (= uf_296 uf_292) +#3108 := (implies #3083 #3107) +#3109 := (implies true #3108) +#3057 := (<= 0::int uf_293) +#3064 := (<= 1::int uf_294) +#3065 := (and #3064 #3057) +#3110 := (implies #3065 #3109) +#3111 := (implies #3065 #3110) +#3112 := (implies true #3111) +#3113 := (implies #3065 #3112) +#3114 := (implies up_216 #3113) +#3115 := (and up_216 #3114) +#3116 := (implies #3065 #3115) +#3117 := (implies true #3116) +#3118 := (implies #3065 #3117) +#3283 := (implies #3065 #3118) +#3284 := (implies true #3283) +#3285 := (implies #3065 #3284) +#3282 := (<= uf_272 uf_294) +#3286 := (implies #3282 #3285) +#3287 := (implies #3065 #3286) +#3288 := (implies true #3287) +#3225 := (implies false true) +#3223 := (= #3222 uf_303) +#3220 := (< uf_304 uf_272) +#3224 := (and #3220 #3223) +#3226 := (implies #3224 #3225) +#3227 := (and #3224 #3226) +#3216 := (<= #3045 uf_303) +#3215 := (< #161 uf_305) +#3217 := (implies #3215 #3216) +#3218 := (implies #412 #3217) +#3219 := (forall (vars (?x786 int)) #3218) +#3228 := (implies #3219 #3227) +#3229 := (and #3219 #3228) +#3214 := (<= uf_305 uf_272) +#3230 := (implies #3214 #3229) +#3231 := (and #3214 #3230) +#3232 := (implies true #3231) +#3202 := (<= 0::int uf_304) +#3212 := (<= 2::int uf_305) +#3213 := (and #3212 #3202) +#3233 := (implies #3213 #3232) +#3234 := (implies #3211 #3233) +#3204 := (+ uf_294 1::int) +#3209 := (= uf_305 #3204) +#3235 := (implies #3209 #3234) +#3206 := (<= #3204 uf_76) +#3205 := (<= 0::int #3204) +#3207 := (and #3205 #3206) +#3236 := (implies #3207 #3235) +#3237 := (and #3207 #3236) +#3203 := (and #3064 #3202) +#3238 := (implies #3203 #3237) +#3239 := (implies true #3238) +#3259 := (= uf_304 uf_293) +#3260 := (implies #3259 #3239) +#3258 := (= uf_303 uf_292) +#3261 := (implies #3258 #3260) +#3262 := (implies true #3261) +#3263 := (implies #3065 #3262) +#3264 := (implies #3065 #3263) +#3265 := (implies true #3264) +#3266 := (implies #3065 #3265) +#3257 := (<= #3189 uf_292) +#3267 := (implies #3257 #3266) +#3268 := (implies #3065 #3267) +#3269 := (implies true #3268) +#3201 := (= uf_304 uf_294) +#3240 := (implies #3201 #3239) +#3199 := (= uf_303 uf_300) +#3241 := (implies #3199 #3240) +#3242 := (implies true #3241) +#3197 := (and #3064 #3064) +#3243 := (implies #3197 #3242) +#3244 := (implies #3196 #3243) +#3245 := (implies #3194 #3244) +#3192 := (= uf_300 #3189) +#3246 := (implies #3192 #3245) +#3187 := (= #3186 uf_9) #3182 := (= #3181 uf_9) -#3183 := (and #3182 #3179) -#3192 := (+ uf_294 1::int) -#3194 := (<= 0::int #3192) -#3193 := (<= #3192 uf_76) -#3195 := (and #3193 #3194) -#3202 := (<= uf_302 uf_272) -#3204 := (<= #3040 uf_300) -#3203 := (< #161 uf_302) -#3205 := (implies #3203 #3204) -#3206 := (implies #645 #3205) -#3207 := (forall (vars (?x785 int)) #3206) -#3211 := (< uf_301 uf_272) -#3210 := (= #3209 uf_300) -#3212 := (and #3210 #3211) -#3213 := (implies false true) -#3214 := (implies #3212 #3213) -#3215 := (and #3214 #3212) -#3216 := (implies #3207 #3215) -#3217 := (and #3216 #3207) -#3218 := (implies #3202 #3217) -#3219 := (and #3218 #3202) -#3220 := (implies true #3219) -#3200 := (<= 2::int uf_302) -#3190 := (<= 0::int uf_301) -#3201 := (and #3190 #3200) -#3221 := (implies #3201 #3220) -#3222 := (implies #3199 #3221) -#3197 := (= uf_302 #3192) -#3223 := (implies #3197 #3222) -#3224 := (implies #3195 #3223) -#3225 := (and #3224 #3195) -#3191 := (and #3190 #3059) -#3226 := (implies #3191 #3225) -#3227 := (implies true #3226) -#3247 := (= uf_301 uf_294) -#3248 := (implies #3247 #3227) -#3249 := (implies #3246 #3248) -#3250 := (implies true #3249) -#3245 := (and #3059 #3059) -#3251 := (implies #3245 #3250) -#3252 := (implies #3244 #3251) -#3253 := (implies #3242 #3252) -#3240 := (= uf_304 #3184) -#3254 := (implies #3240 #3253) -#3255 := (implies #3183 #3254) -#3256 := (and #3255 #3183) -#3257 := (implies #3180 #3256) -#3258 := (and #3257 #3180) -#3259 := (implies #3060 #3258) -#3260 := (implies true #3259) -#3261 := (implies #3060 #3260) -#3238 := (< uf_292 #3184) -#3262 := (implies #3238 #3261) -#3263 := (implies #3060 #3262) -#3264 := (implies true #3263) -#3189 := (= uf_301 uf_293) -#3228 := (implies #3189 #3227) -#3187 := (= uf_300 uf_292) -#3229 := (implies #3187 #3228) -#3230 := (implies true #3229) -#3231 := (implies #3060 #3230) -#3232 := (implies #3060 #3231) -#3233 := (implies true #3232) -#3234 := (implies #3060 #3233) -#3185 := (<= #3184 uf_292) -#3235 := (implies #3185 #3234) -#3236 := (implies #3060 #3235) -#3237 := (implies true #3236) -#3265 := (and #3237 #3264) -#3266 := (implies #3060 #3265) -#3267 := (implies #3183 #3266) -#3268 := (and #3267 #3183) -#3269 := (implies #3180 #3268) -#3270 := (and #3269 #3180) -#3271 := (implies #3060 #3270) -#3272 := (implies true #3271) -#3273 := (implies #3060 #3272) -#3174 := (< uf_294 uf_272) -#3274 := (implies #3174 #3273) -#3275 := (implies #3060 #3274) -#3276 := (implies true #3275) -#3168 := (implies #3060 #3167) -#3169 := (implies true #3168) -#3170 := (implies #3060 #3169) -#3131 := (<= uf_272 uf_294) -#3171 := (implies #3131 #3170) -#3172 := (implies #3060 #3171) -#3173 := (implies true #3172) -#3277 := (and #3173 #3276) -#3278 := (implies #3060 #3277) -decl uf_59 :: (-> T4 T13) -#3079 := (uf_59 uf_273) -#3129 := (= #3079 #3079) +#3188 := (and #3182 #3187) +#3247 := (implies #3188 #3246) +#3248 := (and #3188 #3247) +#3184 := (= #3183 uf_9) +#3185 := (and #3182 #3184) +#3249 := (implies #3185 #3248) +#3250 := (and #3185 #3249) +#3251 := (implies #3065 #3250) +#3252 := (implies true #3251) +#3253 := (implies #3065 #3252) +#3190 := (< uf_292 #3189) +#3254 := (implies #3190 #3253) +#3255 := (implies #3065 #3254) +#3256 := (implies true #3255) +#3270 := (and #3256 #3269) +#3271 := (implies #3065 #3270) +#3272 := (implies #3188 #3271) +#3273 := (and #3188 #3272) +#3274 := (implies #3185 #3273) +#3275 := (and #3185 #3274) +#3276 := (implies #3065 #3275) +#3277 := (implies true #3276) +#3278 := (implies #3065 #3277) +#3179 := (< uf_294 uf_272) +#3279 := (implies #3179 #3278) +#3280 := (implies #3065 #3279) +#3281 := (implies true #3280) +#3289 := (and #3281 #3288) +#3290 := (implies #3065 #3289) decl uf_41 :: (-> T4 T12) -#3088 := (uf_41 uf_273) -#3128 := (= #3088 #3088) -#3130 := (and #3128 #3129) -#3279 := (implies #3130 #3278) -#3280 := (implies #3127 #3279) -#3281 := (implies #3124 #3280) -#3282 := (implies #3123 #3281) -#3283 := (implies #3122 #3282) -#3284 := (implies #3121 #3283) -#3285 := (implies #3120 #3284) -#3078 := (<= #2998 #2998) -decl uf_170 :: (-> T4 T5 int) -#3074 := (uf_170 uf_273 #26) -#3075 := (pattern #3074) -#3076 := (<= #3074 #3074) -#3077 := (forall (vars (?x775 T5)) (:pat #3075) #3076) -#3115 := (and #3077 #3078) -#3073 := (= #3072 uf_9) -#3116 := (and #3073 #3115) -#3286 := (implies #3116 #3285) +#3126 := (uf_41 uf_273) +#3177 := (= #3126 #3126) +#3176 := (= #3149 #3149) +#3178 := (and #3176 #3177) +#3291 := (implies #3178 #3290) +#3292 := (implies #3175 #3291) +#3293 := (implies #3172 #3292) +#3294 := (implies #3171 #3293) +#3295 := (implies #3170 #3294) +#3296 := (implies #3169 #3295) +#2987 := (= #2986 uf_9) +#3080 := (= #3079 uf_9) +#3081 := (and #3080 #2987) +#3297 := (implies #3081 #3296) +#3162 := (= #3161 uf_9) +decl uf_172 :: (-> T4 T5 int) +#3157 := (uf_172 uf_273 #23) +#3158 := (pattern #3157) +#3159 := (<= #3157 #3157) +#3160 := (forall (vars (?x784 T5)) (:pat #3158) #3159) +#3163 := (and #3160 #3162) +#3156 := (<= #3003 #3003) +#3164 := (and #3156 #3163) +#3298 := (implies #3164 #3297) +#3150 := (uf_58 #3149 #23) +#3151 := (pattern #3150) +#3139 := (uf_68 uf_273 #23) +#3140 := (= #3139 uf_9) +#3152 := (= #3150 #3150) +#3153 := (and #3152 #3140) +#3154 := (implies #3140 #3153) +#3155 := (forall (vars (?x783 T5)) (:pat #3151) #3154) +#3165 := (and #3155 #3164) decl uf_40 :: (-> T12 T5 T11) -#3089 := (uf_40 #3088 #26) -#3090 := (pattern #3089) -decl uf_261 :: T8 -#2832 := uf_261 -#3102 := (uf_25 uf_273 #26) -#3103 := (uf_15 #3102) -#3104 := (uf_14 #3103) -#3105 := (= #3104 uf_261) -#3106 := (not #3105) -#3107 := (implies #3106 #3106) -#3108 := (forall (vars (?x779 T5)) (:pat #3090) #3107) +#3127 := (uf_40 #3126 #23) +#3128 := (pattern #3127) +#3145 := (= #3127 #3127) +#3146 := (and #3145 #3140) +#3147 := (implies #3140 #3146) +#3148 := (forall (vars (?x782 T5)) (:pat #3128) #3147) +#3166 := (and #3148 #3165) decl uf_19 :: (-> T9 T5 int) decl uf_20 :: (-> T4 T9) -#3095 := (uf_20 uf_273) -#3096 := (uf_19 #3095 #26) -#3097 := (pattern #3096) -#3098 := (= #3096 #3096) -#3082 := (uf_67 uf_273 #26) -#3083 := (= #3082 uf_9) -#3099 := (and #3083 #3098) -#3100 := (implies #3083 #3099) -#3101 := (forall (vars (?x778 T5)) (:pat #3097) #3100) -#3109 := (and #3101 #3108) -#3091 := (= #3089 #3089) -#3092 := (and #3083 #3091) -#3093 := (implies #3083 #3092) -#3094 := (forall (vars (?x777 T5)) (:pat #3090) #3093) -#3110 := (and #3094 #3109) -decl uf_58 :: (-> T13 T5 T14) -#3080 := (uf_58 #3079 #26) -#3081 := (pattern #3080) -#3084 := (= #3080 #3080) -#3085 := (and #3083 #3084) -#3086 := (implies #3083 #3085) -#3087 := (forall (vars (?x776 T5)) (:pat #3081) #3086) -#3111 := (and #3087 #3110) -#3112 := (and #3078 #3111) -#3113 := (and #3077 #3112) -#3114 := (and #3073 #3113) -#3287 := (implies #3114 #3286) -#3288 := (implies #3060 #3287) -#3289 := (implies true #3288) -#3290 := (implies #3060 #3289) -#3291 := (implies true #3290) -#3292 := (implies #3060 #3291) -#3293 := (implies true #3292) -#3302 := (and #3293 #3301) -#3303 := (implies #3060 #3302) -#3070 := (< uf_293 uf_272) -#3069 := (= #3068 uf_292) -#3071 := (and #3069 #3070) -#3304 := (implies #3071 #3303) -#3063 := (<= #3040 uf_292) -#3062 := (< #161 uf_294) -#3064 := (implies #3062 #3063) -#3065 := (implies #645 #3064) -#3066 := (forall (vars (?x774 int)) #3065) -#3305 := (implies #3066 #3304) -#3061 := (<= uf_294 uf_272) -#3306 := (implies #3061 #3305) -#3307 := (implies #3060 #3306) -#3057 := (<= 0::int uf_294) -#3056 := (<= uf_294 uf_76) -#3058 := (and #3056 #3057) -#3308 := (implies #3058 #3307) -#3052 := (<= uf_293 uf_76) -#3054 := (and #3052 #3053) -#3309 := (implies #3054 #3308) -#3049 := (<= 0::int uf_292) -#3048 := (<= uf_292 uf_78) -#3050 := (and #3048 #3049) -#3310 := (implies #3050 #3309) -#3311 := (implies true #3310) -#3312 := (implies #3046 #3311) -#3313 := (and #3312 #3046) -#3314 := (implies #3044 #3313) -#3315 := (and #3314 #3044) -#3316 := (implies #3037 #3315) -#3317 := (and #3316 #3037) -#3033 := (<= 1::int 1::int) -#3034 := (and #3033 #3033) -#3032 := (<= 0::int 0::int) -#3035 := (and #3032 #3034) -#3036 := (and #3032 #3035) -#3318 := (implies #3036 #3317) -#3319 := (implies #3031 #3318) -#3320 := (implies #3028 #3319) -#3321 := (implies #3025 #3320) -#3322 := (implies #3022 #3321) -#3323 := (implies #3019 #3322) -#3324 := (and #3323 #3019) -#3325 := (implies #3016 #3324) -#3326 := (and #3325 #3016) -#3327 := (implies #3010 #3326) -#3328 := (and #3327 #3010) -#3006 := (<= 0::int uf_272) -#3005 := (<= uf_272 uf_76) -#3007 := (and #3005 #3006) -#3329 := (implies #3007 #3328) -#3002 := (= #3000 uf_9) -#3003 := (iff #3002 false) -#3004 := (forall (vars (?x772 T5)) (:pat #3001) #3003) -#3330 := (implies #3004 #3329) -#3331 := (implies #2999 #3330) -#3332 := (implies #2996 #3331) -#3333 := (implies #2994 #3332) -#2988 := (< #2985 uf_277) -#2989 := (forall (vars (?x771 T19)) (:pat #2986) #2988) -#3334 := (implies #2989 #3333) +#3136 := (uf_20 uf_273) +#3137 := (uf_19 #3136 #23) +#3138 := (pattern #3137) +#3141 := (= #3137 #3137) +#3142 := (and #3141 #3140) +#3143 := (implies #3140 #3142) +#3144 := (forall (vars (?x781 T5)) (:pat #3138) #3143) +#3167 := (and #3144 #3166) +decl uf_261 :: T8 +#2837 := uf_261 +#3129 := (uf_25 uf_273 #23) +#3130 := (uf_13 #3129) +#3131 := (uf_12 #3130) +#3132 := (= #3131 uf_261) +#3133 := (not #3132) +#3134 := (implies #3133 #3133) +#3135 := (forall (vars (?x780 T5)) (:pat #3128) #3134) +#3168 := (and #3135 #3167) +#3299 := (implies #3168 #3298) +#3300 := (implies #3065 #3299) +#3301 := (implies true #3300) +#3302 := (implies #3065 #3301) +#3303 := (implies true #3302) +#3304 := (implies #3065 #3303) +#3305 := (implies true #3304) +#3119 := (implies #3081 #3118) +#3120 := (implies #3065 #3119) +#3121 := (implies true #3120) +#3122 := (implies #3065 #3121) +#3077 := (not true) +#3123 := (implies #3077 #3122) +#3124 := (implies #3065 #3123) +#3125 := (implies true #3124) +#3306 := (and #3125 #3305) +#3307 := (implies #3065 #3306) +#3075 := (= #3074 uf_292) +#3072 := (< uf_293 uf_272) +#3076 := (and #3072 #3075) +#3308 := (implies #3076 #3307) +#3068 := (<= #3045 uf_292) +#3067 := (< #161 uf_294) +#3069 := (implies #3067 #3068) +#3070 := (implies #412 #3069) +#3071 := (forall (vars (?x775 int)) #3070) +#3309 := (implies #3071 #3308) +#3066 := (<= uf_294 uf_272) +#3310 := (implies #3066 #3309) +#3311 := (implies #3065 #3310) +#3062 := (<= uf_294 uf_76) +#3061 := (<= 0::int uf_294) +#3063 := (and #3061 #3062) +#3312 := (implies #3063 #3311) +#3058 := (<= uf_293 uf_76) +#3059 := (and #3057 #3058) +#3313 := (implies #3059 #3312) +#3054 := (<= uf_292 uf_78) +#3053 := (<= 0::int uf_292) +#3055 := (and #3053 #3054) +#3314 := (implies #3055 #3313) +#3315 := (implies true #3314) +#3050 := (= #3026 uf_285) +#2956 := (< 0::int uf_272) +#3051 := (and #2956 #3050) +#3316 := (implies #3051 #3315) +#3317 := (and #3051 #3316) +#3046 := (<= #3045 uf_285) +#3043 := (< #161 1::int) +#3047 := (implies #3043 #3046) +#3048 := (implies #412 #3047) +#3049 := (forall (vars (?x773 int)) #3048) +#3318 := (implies #3049 #3317) +#3319 := (and #3049 #3318) +#3042 := (<= 1::int uf_272) +#3320 := (implies #3042 #3319) +#3321 := (and #3042 #3320) +#3038 := (<= 0::int 0::int) +#3039 := (and #3038 #3038) +#3037 := (<= 1::int 1::int) +#3040 := (and #3037 #3039) +#3041 := (and #3037 #3040) +#3322 := (implies #3041 #3321) +#3323 := (implies #3036 #3322) +#3324 := (implies #3033 #3323) +#3325 := (implies #3030 #3324) +#3326 := (implies #3027 #3325) +#3023 := (= #3022 uf_9) +#3018 := (= #3017 uf_9) +#3024 := (and #3018 #3023) +#3327 := (implies #3024 #3326) +#3328 := (and #3024 #3327) +#3020 := (= #3019 uf_9) +#3021 := (and #3018 #3020) +#3329 := (implies #3021 #3328) +#3330 := (and #3021 #3329) +#3015 := (= #3014 uf_9) +#3331 := (implies #3015 #3330) +#3332 := (and #3015 #3331) +#3011 := (<= uf_272 uf_76) +#3010 := (<= 0::int uf_272) +#3012 := (and #3010 #3011) +#3333 := (implies #3012 #3332) +#3007 := (= #3005 uf_9) +#3008 := (iff #3007 false) +#3009 := (forall (vars (?x772 T5)) (:pat #3006) #3008) +#3334 := (implies #3009 #3333) +#3335 := (implies #3004 #3334) +#3336 := (implies #3001 #3335) +#3337 := (implies #2999 #3336) +#2993 := (< #2990 uf_277) +#2994 := (forall (vars (?x771 T19)) (:pat #2991) #2993) +#3338 := (implies #2994 #3337) +#2985 := (= #2984 uf_9) +#2988 := (and #2985 #2987) +#3339 := (implies #2988 #3338) #2982 := (= #2981 uf_9) -#2983 := (and #2979 #2982) -#3335 := (implies #2983 #3334) -#2977 := (= #2976 uf_9) -#3336 := (implies #2977 #3335) -#3337 := (implies true #3336) +#3340 := (implies #2982 #3339) +#3341 := (implies true #3340) +#2975 := (= #2974 uf_9) +#2972 := (= #2971 uf_14) +#2973 := (not #2972) +#2976 := (and #2973 #2975) #2970 := (= #2969 uf_9) -#2968 := (= #2967 uf_26) -#2971 := (and #2968 #2970) -#2966 := (= #2965 uf_9) -#2972 := (and #2966 #2971) +#2977 := (and #2970 #2976) +#2968 := (= #2967 uf_9) +#2978 := (and #2968 #2977) +#2966 := (= #2965 uf_26) +#2979 := (and #2966 #2978) #2964 := (= #2963 uf_9) -#2973 := (and #2964 #2972) -#2956 := (= #2955 uf_16) -#2957 := (not #2956) -#2974 := (and #2957 #2973) -#2954 := (= #2953 uf_9) -#2975 := (and #2954 #2974) -#3338 := (implies #2975 #3337) -#3339 := (implies #2951 #3338) -#2950 := (< uf_272 1099511627776::int) -#3340 := (implies #2950 #3339) -#2947 := (<= 0::int uf_271) -#2946 := (<= uf_271 uf_76) -#2948 := (and #2946 #2947) -#3341 := (implies #2948 #3340) -#2943 := (<= 0::int uf_270) -#2942 := (<= uf_270 uf_76) -#2944 := (and #2942 #2943) -#3342 := (implies #2944 #3341) -#2939 := (<= 0::int uf_269) -#2938 := (<= uf_269 uf_78) -#2940 := (and #2938 #2939) -#3343 := (implies #2940 #3342) -#3344 := (implies true #3343) -#3345 := (not #3344) -#14774 := (iff #3345 #14771) -#12868 := (and #3211 #12862) -#12847 := (not #3203) -#12848 := (or #12847 #3204) -#5718 := (and #285 #410) -#5727 := (not #5718) -#12854 := (or #5727 #12848) -#12859 := (forall (vars (?x785 int)) #12854) -#12892 := (not #12859) -#12893 := (or #12892 #12868) -#12901 := (and #12859 #12893) -#12909 := (not #3202) -#12910 := (or #12909 #12901) -#12918 := (and #3202 #12910) -#12933 := (not #3201) -#12934 := (or #12933 #12918) -#12943 := (or #12942 #12934) -#12832 := (+ 1::int uf_294) -#12844 := (= uf_302 #12832) -#12951 := (not #12844) -#12952 := (or #12951 #12943) -#12838 := (<= 0::int #12832) -#12835 := (<= #12832 uf_76) -#12841 := (and #12835 #12838) -#12960 := (not #12841) -#12961 := (or #12960 #12952) -#12969 := (and #12841 #12961) -#12829 := (and #3059 #3190) -#12977 := (not #12829) -#12978 := (or #12977 #12969) -#13082 := (or #12978 #13081) -#13091 := (or #13090 #13082) -#13106 := (not #3059) -#13107 := (or #13106 #13091) -#13116 := (or #13115 #13107) -#13125 := (or #13124 #13116) -#13134 := (or #13133 #13125) -#13143 := (or #13142 #13134) -#13151 := (and #12818 #13143) -#13160 := (or #13159 #13151) -#13168 := (and #12809 #13160) -#12687 := (not #3060) -#13176 := (or #12687 #13168) -#13191 := (or #12687 #13176) -#13199 := (not #3238) -#13200 := (or #13199 #13191) -#13208 := (or #12687 #13200) -#12994 := (or #12993 #12978) -#13003 := (or #13002 #12994) -#13018 := (or #12687 #13003) -#13026 := (or #12687 #13018) -#13041 := (or #12687 #13026) -#13049 := (not #3185) -#13050 := (or #13049 #13041) -#13058 := (or #12687 #13050) -#13220 := (and #13058 #13208) -#13226 := (or #12687 #13220) -#13234 := (or #13142 #13226) -#13242 := (and #12818 #13234) -#13250 := (or #13159 #13242) -#13258 := (and #12809 #13250) -#13266 := (or #12687 #13258) -#13281 := (or #12687 #13266) -#13289 := (not #3174) -#13290 := (or #13289 #13281) -#13298 := (or #12687 #13290) -#12594 := (and #3140 #5718) -#12597 := (and #3145 #12594) -#12600 := (exists (vars (?x782 int)) #12597) -#12579 := (not #3140) -#12580 := (or #12579 #3141) -#12586 := (or #5727 #12580) -#12591 := (forall (vars (?x781 int)) #12586) -#12620 := (not #12591) -#12621 := (or #12620 #12600) -#12629 := (and #12591 #12621) -#12645 := (or #12644 #12629) -#12654 := (or #12653 #12645) +#2980 := (and #2964 #2979) +#3342 := (implies #2980 #3341) +#3343 := (implies #2956 #3342) +#2955 := (< uf_272 1099511627776::int) +#3344 := (implies #2955 #3343) +#2952 := (<= uf_271 uf_76) +#2951 := (<= 0::int uf_271) +#2953 := (and #2951 #2952) +#3345 := (implies #2953 #3344) +#2948 := (<= uf_270 uf_76) +#2947 := (<= 0::int uf_270) +#2949 := (and #2947 #2948) +#3346 := (implies #2949 #3345) +#2944 := (<= uf_269 uf_78) +#2943 := (<= 0::int uf_269) +#2945 := (and #2943 #2944) +#3347 := (implies #2945 #3346) +#3348 := (implies true #3347) +#3349 := (not #3348) +#14240 := (iff #3349 #14237) +#12023 := (not #3090) +#12024 := (or #12023 #3091) +#5592 := (not #412) +#12030 := (or #5592 #12024) +#12035 := (forall (vars (?x776 int)) #12030) +#12050 := (not #12035) +#12051 := (or #3099 #12050) +#12056 := (and #12035 #12051) +#12070 := (or #12069 #12056) +#12079 := (or #12078 #12070) +#12088 := (or #12087 #12079) +#12097 := (or #12096 #12088) +#11974 := (and #3057 #3064) +#12112 := (not #11974) +#12113 := (or #12112 #12097) +#12121 := (or #12112 #12113) +#12136 := (or #12112 #12121) +#12145 := (or #12144 #12136) +#12150 := (and up_216 #12145) +#12156 := (or #12112 #12150) +#12171 := (or #12112 #12156) +#12813 := (or #12112 #12171) +#12828 := (or #12112 #12813) +#12836 := (not #3282) +#12837 := (or #12836 #12828) +#12845 := (or #12112 #12837) +#12407 := (and #3220 #12404) +#12389 := (not #3215) +#12390 := (or #12389 #3216) +#12396 := (or #5592 #12390) +#12401 := (forall (vars (?x786 int)) #12396) +#12429 := (not #12401) +#12430 := (or #12429 #12407) +#12435 := (and #12401 #12430) +#12441 := (not #3214) +#12442 := (or #12441 #12435) +#12447 := (and #3214 #12442) +#12386 := (and #3202 #3212) +#12460 := (not #12386) +#12461 := (or #12460 #12447) +#12470 := (or #12469 #12461) +#12371 := (+ 1::int uf_294) +#12383 := (= uf_305 #12371) +#12478 := (not #12383) +#12479 := (or #12478 #12470) +#12377 := (<= #12371 uf_76) +#12374 := (<= 0::int #12371) +#12380 := (and #12374 #12377) +#12487 := (not #12380) +#12488 := (or #12487 #12479) +#12493 := (and #12380 #12488) +#12499 := (not #3203) +#12500 := (or #12499 #12493) +#12654 := (or #12500 #12653) #12663 := (or #12662 #12654) -#12672 := (or #12671 #12663) -#12688 := (or #12687 #12672) -#12696 := (or #12687 #12688) -#12711 := (or #12687 #12696) -#12720 := (or #12719 #12711) -#12728 := (and up_216 #12720) -#12736 := (or #12687 #12728) -#12751 := (or #12687 #12736) -#12759 := (or #12687 #12751) -#12774 := (or #12687 #12759) -#12782 := (not #3131) -#12783 := (or #12782 #12774) -#12791 := (or #12687 #12783) -#13310 := (and #12791 #13298) -#13316 := (or #12687 #13310) -#13332 := (or #13331 #13316) -#13341 := (or #13340 #13332) -#13350 := (or #13349 #13341) -#13359 := (or #13358 #13350) -#13368 := (or #13367 #13359) -#13377 := (or #13376 #13368) -#12544 := (and #3115 #12437) -#13385 := (not #12544) -#13386 := (or #13385 #13377) -#13394 := (or #13385 #13386) -#13402 := (or #12687 #13394) -#13417 := (or #12687 #13402) -#13432 := (or #12687 #13417) -#13508 := (or #12687 #13432) -#12432 := (and #3070 #12426) -#13516 := (not #12432) -#13517 := (or #13516 #13508) -#12411 := (not #3062) -#12412 := (or #12411 #3063) -#12418 := (or #5727 #12412) -#12423 := (forall (vars (?x774 int)) #12418) -#13525 := (not #12423) -#13526 := (or #13525 #13517) -#13534 := (not #3061) -#13535 := (or #13534 #13526) -#13543 := (or #12687 #13535) -#13551 := (not #3058) -#13552 := (or #13551 #13543) -#13560 := (not #3054) -#13561 := (or #13560 #13552) -#13569 := (not #3050) -#13570 := (or #13569 #13561) -#12406 := (and #2951 #3022) -#13585 := (not #12406) -#13586 := (or #13585 #13570) -#13594 := (and #12406 #13586) -#12386 := (not #3038) -#12387 := (or #12386 #3041) -#12393 := (or #5727 #12387) -#12398 := (forall (vars (?x773 int)) #12393) -#13602 := (not #12398) -#13603 := (or #13602 #13594) -#13611 := (and #12398 #13603) -#13619 := (not #3037) -#13620 := (or #13619 #13611) -#13628 := (and #3037 #13620) -#12380 := (and #3032 #3033) -#12383 := (and #3032 #12380) -#13636 := (not #12383) -#13637 := (or #13636 #13628) -#13646 := (or #13645 #13637) -#13655 := (or #13654 #13646) -#13664 := (or #13663 #13655) -#13673 := (or #13672 #13664) -#13682 := (or #13681 #13673) -#13690 := (and #12373 #13682) -#13699 := (or #13698 #13690) -#13707 := (and #12364 #13699) -#13716 := (or #13715 #13707) -#13724 := (and #12355 #13716) -#13732 := (not #3007) -#13733 := (or #13732 #13724) -#13742 := (or #13741 #13733) -#13751 := (or #13750 #13742) -#13760 := (or #13759 #13751) -#13769 := (or #13768 #13760) -#13777 := (not #2989) -#13778 := (or #13777 #13769) -#13787 := (or #13786 #13778) -#13796 := (or #13795 #13787) -#12314 := (and #12308 #12311) -#12317 := (and #12305 #12314) -#12320 := (and #12302 #12317) -#12323 := (and #12299 #12320) -#12326 := (and #12293 #12323) -#13811 := (not #12326) -#13812 := (or #13811 #13796) -#13820 := (not #2951) -#13821 := (or #13820 #13812) -#13829 := (not #2950) -#13830 := (or #13829 #13821) -#13838 := (not #2948) -#13839 := (or #13838 #13830) -#13847 := (not #2944) -#13848 := (or #13847 #13839) -#13856 := (not #2940) -#13857 := (or #13856 #13848) -#13869 := (not #13857) -#14772 := (iff #13869 #14771) -#14769 := (iff #13857 #14766) -#14724 := (or #13883 #14643) -#14727 := (or #13741 #14724) -#14730 := (or #13750 #14727) -#14733 := (or #13759 #14730) -#14736 := (or #13768 #14733) -#14739 := (or #14655 #14736) -#14742 := (or #13786 #14739) -#14745 := (or #13795 #14742) -#14748 := (or #14661 #14745) -#14751 := (or #13942 #14748) -#14754 := (or #14671 #14751) -#14757 := (or #14693 #14754) -#14760 := (or #14707 #14757) -#14763 := (or #14721 #14760) -#14767 := (iff #14763 #14766) -#14768 := [rewrite]: #14767 -#14764 := (iff #13857 #14763) -#14761 := (iff #13848 #14760) -#14758 := (iff #13839 #14757) -#14755 := (iff #13830 #14754) -#14752 := (iff #13821 #14751) -#14749 := (iff #13812 #14748) -#14746 := (iff #13796 #14745) -#14743 := (iff #13787 #14742) -#14740 := (iff #13778 #14739) -#14737 := (iff #13769 #14736) -#14734 := (iff #13760 #14733) -#14731 := (iff #13751 #14730) -#14728 := (iff #13742 #14727) -#14725 := (iff #13733 #14724) -#14644 := (iff #13724 #14643) -#14641 := (iff #13716 #14640) -#14638 := (iff #13707 #14635) -#14632 := (and #12364 #14629) -#14636 := (iff #14632 #14635) -#14637 := [rewrite]: #14636 -#14633 := (iff #13707 #14632) -#14630 := (iff #13699 #14629) -#14627 := (iff #13690 #14624) -#14621 := (and #12373 #14616) -#14625 := (iff #14621 #14624) -#14626 := [rewrite]: #14625 -#14622 := (iff #13690 #14621) -#14619 := (iff #13682 #14616) -#14598 := (or false #14595) -#14601 := (or #13645 #14598) -#14604 := (or #13654 #14601) -#14607 := (or #13663 #14604) -#14610 := (or #13672 #14607) -#14613 := (or #13681 #14610) -#14617 := (iff #14613 #14616) -#14618 := [rewrite]: #14617 -#14614 := (iff #13682 #14613) -#14611 := (iff #13673 #14610) -#14608 := (iff #13664 #14607) -#14605 := (iff #13655 #14604) -#14602 := (iff #13646 #14601) -#14599 := (iff #13637 #14598) -#14596 := (iff #13628 #14595) -#14593 := (iff #13620 #14592) -#14590 := (iff #13611 #14589) -#14587 := (iff #13603 #14586) -#14584 := (iff #13594 #14581) -#14499 := (and #13943 #3022) -#14578 := (and #14499 #14573) -#14582 := (iff #14578 #14581) -#14583 := [rewrite]: #14582 -#14579 := (iff #13594 #14578) -#14576 := (iff #13586 #14573) -#14510 := (or #13955 #14376) -#14513 := (or #13331 #14510) -#14516 := (or #13340 #14513) -#14519 := (or #13349 #14516) -#14522 := (or #13358 #14519) -#14525 := (or #13367 #14522) -#14528 := (or #13376 #14525) -#14531 := (or #14399 #14528) -#14534 := (or #14399 #14531) -#14537 := (or #13955 #14534) -#14540 := (or #13955 #14537) -#14543 := (or #13955 #14540) -#14546 := (or #13955 #14543) -#14549 := (or #14416 #14546) -#14552 := (or #14450 #14549) -#14555 := (or #14456 #14552) -#14558 := (or #13955 #14555) -#14561 := (or #14468 #14558) -#14564 := (or #14483 #14561) -#14567 := (or #14496 #14564) -#14570 := (or #14507 #14567) -#14574 := (iff #14570 #14573) -#14575 := [rewrite]: #14574 -#14571 := (iff #13586 #14570) -#14568 := (iff #13570 #14567) -#14565 := (iff #13561 #14564) -#14562 := (iff #13552 #14561) -#14559 := (iff #13543 #14558) -#14556 := (iff #13535 #14555) -#14553 := (iff #13526 #14552) -#14550 := (iff #13517 #14549) -#14547 := (iff #13508 #14546) -#14544 := (iff #13432 #14543) -#14541 := (iff #13417 #14540) -#14538 := (iff #13402 #14537) -#14535 := (iff #13394 #14534) -#14532 := (iff #13386 #14531) -#14529 := (iff #13377 #14528) -#14526 := (iff #13368 #14525) -#14523 := (iff #13359 #14522) -#14520 := (iff #13350 #14519) -#14517 := (iff #13341 #14516) -#14514 := (iff #13332 #14513) -#14511 := (iff #13316 #14510) -#14377 := (iff #13310 #14376) -#14374 := (iff #13298 #14371) -#14359 := (or #13955 #14345) -#14362 := (or #13955 #14359) -#14365 := (or #14046 #14362) -#14368 := (or #13955 #14365) -#14372 := (iff #14368 #14371) -#14373 := [rewrite]: #14372 -#14369 := (iff #13298 #14368) -#14366 := (iff #13290 #14365) -#14363 := (iff #13281 #14362) -#14360 := (iff #13266 #14359) -#14348 := (iff #13258 #14345) -#14342 := (and #12809 #14339) -#14346 := (iff #14342 #14345) -#14347 := [rewrite]: #14346 -#14343 := (iff #13258 #14342) -#14340 := (iff #13250 #14339) -#14337 := (iff #13242 #14334) -#14331 := (and #12818 #14326) -#14335 := (iff #14331 #14334) -#14336 := [rewrite]: #14335 -#14332 := (iff #13242 #14331) -#14329 := (iff #13234 #14326) -#14320 := (or #13955 #14317) -#14323 := (or #13142 #14320) -#14327 := (iff #14323 #14326) -#14328 := [rewrite]: #14327 -#14324 := (iff #13234 #14323) -#14321 := (iff #13226 #14320) -#14318 := (iff #13220 #14317) -#14315 := (iff #13208 #14312) -#14300 := (or #13955 #14286) -#14303 := (or #13955 #14300) -#14306 := (or #14207 #14303) -#14309 := (or #13955 #14306) -#14313 := (iff #14309 #14312) -#14314 := [rewrite]: #14313 -#14310 := (iff #13208 #14309) -#14307 := (iff #13200 #14306) -#14304 := (iff #13191 #14303) -#14301 := (iff #13176 #14300) -#14289 := (iff #13168 #14286) -#14283 := (and #12809 #14280) -#14287 := (iff #14283 #14286) -#14288 := [rewrite]: #14287 -#14284 := (iff #13168 #14283) -#14281 := (iff #13160 #14280) -#14278 := (iff #13151 #14275) -#14272 := (and #12818 #14267) -#14276 := (iff #14272 #14275) -#14277 := [rewrite]: #14276 -#14273 := (iff #13151 #14272) -#14270 := (iff #13143 #14267) -#14214 := (or #14081 #14201) -#14246 := (or #14214 #13081) -#14249 := (or #13090 #14246) -#14252 := (or #14243 #14249) -#14255 := (or #13115 #14252) -#14258 := (or #13124 #14255) -#14261 := (or #13133 #14258) -#14264 := (or #13142 #14261) -#14268 := (iff #14264 #14267) -#14269 := [rewrite]: #14268 -#14265 := (iff #13143 #14264) -#14262 := (iff #13134 #14261) -#14259 := (iff #13125 #14258) -#14256 := (iff #13116 #14255) -#14253 := (iff #13107 #14252) -#14250 := (iff #13091 #14249) -#14247 := (iff #13082 #14246) -#14215 := (iff #12978 #14214) -#14204 := (iff #12969 #14201) -#14198 := (and #14175 #14193) -#14202 := (iff #14198 #14201) -#14203 := [rewrite]: #14202 -#14199 := (iff #12969 #14198) -#14196 := (iff #12961 #14193) -#14181 := (or #14097 #14165) -#14184 := (or #12942 #14181) -#14187 := (or #14172 #14184) -#14190 := (or #14178 #14187) -#14194 := (iff #14190 #14193) -#14195 := [rewrite]: #14194 -#14191 := (iff #12961 #14190) -#14188 := (iff #12952 #14187) -#14185 := (iff #12943 #14184) -#14182 := (iff #12934 #14181) -#14166 := (iff #12918 #14165) -#14163 := (iff #12910 #14162) -#14160 := (iff #12901 #14159) -#14157 := (iff #12893 #14156) -#14154 := (iff #12868 #14151) -#14148 := (and #14145 #12862) -#14152 := (iff #14148 #14151) -#14153 := [rewrite]: #14152 -#14149 := (iff #12868 #14148) -#14146 := (iff #3211 #14145) -#14147 := [rewrite]: #14146 -#14150 := [monotonicity #14147]: #14149 -#14155 := [trans #14150 #14153]: #14154 -#14140 := (iff #12892 #14139) -#14137 := (iff #12859 #14136) -#14134 := (iff #12854 #14131) -#14125 := (or #14109 #14122) -#14128 := (or #5739 #14125) -#14132 := (iff #14128 #14131) -#14133 := [rewrite]: #14132 -#14129 := (iff #12854 #14128) -#14126 := (iff #12848 #14125) -#14123 := (iff #3204 #14122) -#14124 := [rewrite]: #14123 -#14118 := (iff #12847 #14109) -#14108 := (not #14109) -#14113 := (not #14108) -#14116 := (iff #14113 #14109) -#14117 := [rewrite]: #14116 -#14114 := (iff #12847 #14113) -#14111 := (iff #3203 #14108) -#14112 := [rewrite]: #14111 -#14115 := [monotonicity #14112]: #14114 -#14119 := [trans #14115 #14117]: #14118 -#14127 := [monotonicity #14119 #14124]: #14126 -#5740 := (iff #5727 #5739) -#5737 := (iff #5718 #5736) -#4420 := (iff #410 #4419) -#4421 := [rewrite]: #4420 -#4083 := (iff #285 #4084) -#4085 := [rewrite]: #4083 -#5738 := [monotonicity #4085 #4421]: #5737 -#5741 := [monotonicity #5738]: #5740 -#14130 := [monotonicity #5741 #14127]: #14129 -#14135 := [trans #14130 #14133]: #14134 -#14138 := [quant-intro #14135]: #14137 -#14141 := [monotonicity #14138]: #14140 -#14158 := [monotonicity #14141 #14155]: #14157 -#14161 := [monotonicity #14138 #14158]: #14160 -#14106 := (iff #12909 #14105) -#14103 := (iff #3202 #14100) -#14104 := [rewrite]: #14103 -#14107 := [monotonicity #14104]: #14106 -#14164 := [monotonicity #14107 #14161]: #14163 -#14167 := [monotonicity #14104 #14164]: #14166 -#14098 := (iff #12933 #14097) -#14095 := (iff #3201 #14094) -#14091 := (iff #3200 #14092) -#14093 := [rewrite]: #14091 -#14075 := (iff #3190 #14076) -#14077 := [rewrite]: #14075 -#14096 := [monotonicity #14077 #14093]: #14095 -#14099 := [monotonicity #14096]: #14098 -#14183 := [monotonicity #14099 #14167]: #14182 -#14186 := [monotonicity #14183]: #14185 -#14173 := (iff #12951 #14172) -#14170 := (iff #12844 #14168) -#14171 := [rewrite]: #14170 -#14174 := [monotonicity #14171]: #14173 -#14189 := [monotonicity #14174 #14186]: #14188 -#14179 := (iff #12960 #14178) -#14176 := (iff #12841 #14175) -#14089 := (iff #12838 #14088) -#14090 := [rewrite]: #14089 -#14086 := (iff #12835 #14084) -#14087 := [rewrite]: #14086 -#14177 := [monotonicity #14087 #14090]: #14176 -#14180 := [monotonicity #14177]: #14179 -#14192 := [monotonicity #14180 #14189]: #14191 -#14197 := [trans #14192 #14195]: #14196 -#14200 := [monotonicity #14177 #14197]: #14199 -#14205 := [trans #14200 #14203]: #14204 -#14082 := (iff #12977 #14081) -#14079 := (iff #12829 #14078) -#13949 := (iff #3059 #13950) -#13951 := [rewrite]: #13949 -#14080 := [monotonicity #13951 #14077]: #14079 -#14083 := [monotonicity #14080]: #14082 -#14216 := [monotonicity #14083 #14205]: #14215 -#14248 := [monotonicity #14216]: #14247 -#14251 := [monotonicity #14248]: #14250 -#14244 := (iff #13106 #14243) -#14245 := [monotonicity #13951]: #14244 -#14254 := [monotonicity #14245 #14251]: #14253 -#14257 := [monotonicity #14254]: #14256 -#14260 := [monotonicity #14257]: #14259 -#14263 := [monotonicity #14260]: #14262 -#14266 := [monotonicity #14263]: #14265 -#14271 := [trans #14266 #14269]: #14270 -#14274 := [monotonicity #14271]: #14273 -#14279 := [trans #14274 #14277]: #14278 -#14282 := [monotonicity #14279]: #14281 -#14285 := [monotonicity #14282]: #14284 -#14290 := [trans #14285 #14288]: #14289 -#13956 := (iff #12687 #13955) -#13953 := (iff #3060 #13952) -#13946 := (iff #3053 #13947) -#13948 := [rewrite]: #13946 -#13954 := [monotonicity #13948 #13951]: #13953 -#13957 := [monotonicity #13954]: #13956 -#14302 := [monotonicity #13957 #14290]: #14301 -#14305 := [monotonicity #13957 #14302]: #14304 -#14298 := (iff #13199 #14207) -#14293 := (not #14211) -#14296 := (iff #14293 #14207) -#14297 := [rewrite]: #14296 -#14294 := (iff #13199 #14293) -#14291 := (iff #3238 #14211) -#14292 := [rewrite]: #14291 -#14295 := [monotonicity #14292]: #14294 -#14299 := [trans #14295 #14297]: #14298 -#14308 := [monotonicity #14299 #14305]: #14307 -#14311 := [monotonicity #13957 #14308]: #14310 -#14316 := [trans #14311 #14314]: #14315 -#14241 := (iff #13058 #14238) -#14217 := (or #12993 #14214) -#14220 := (or #13002 #14217) -#14223 := (or #13955 #14220) -#14226 := (or #13955 #14223) -#14229 := (or #13955 #14226) -#14232 := (or #14211 #14229) -#14235 := (or #13955 #14232) -#14239 := (iff #14235 #14238) -#14240 := [rewrite]: #14239 -#14236 := (iff #13058 #14235) -#14233 := (iff #13050 #14232) -#14230 := (iff #13041 #14229) -#14227 := (iff #13026 #14226) -#14224 := (iff #13018 #14223) -#14221 := (iff #13003 #14220) -#14218 := (iff #12994 #14217) -#14219 := [monotonicity #14216]: #14218 -#14222 := [monotonicity #14219]: #14221 -#14225 := [monotonicity #13957 #14222]: #14224 -#14228 := [monotonicity #13957 #14225]: #14227 -#14231 := [monotonicity #13957 #14228]: #14230 -#14212 := (iff #13049 #14211) -#14206 := (iff #3185 #14207) -#14210 := [rewrite]: #14206 -#14213 := [monotonicity #14210]: #14212 -#14234 := [monotonicity #14213 #14231]: #14233 -#14237 := [monotonicity #13957 #14234]: #14236 -#14242 := [trans #14237 #14240]: #14241 -#14319 := [monotonicity #14242 #14316]: #14318 -#14322 := [monotonicity #13957 #14319]: #14321 -#14325 := [monotonicity #14322]: #14324 -#14330 := [trans #14325 #14328]: #14329 -#14333 := [monotonicity #14330]: #14332 -#14338 := [trans #14333 #14336]: #14337 -#14341 := [monotonicity #14338]: #14340 -#14344 := [monotonicity #14341]: #14343 -#14349 := [trans #14344 #14347]: #14348 -#14361 := [monotonicity #13957 #14349]: #14360 -#14364 := [monotonicity #13957 #14361]: #14363 -#14357 := (iff #13289 #14046) -#14352 := (not #14049) -#14355 := (iff #14352 #14046) -#14356 := [rewrite]: #14355 -#14353 := (iff #13289 #14352) -#14350 := (iff #3174 #14049) -#14351 := [rewrite]: #14350 -#14354 := [monotonicity #14351]: #14353 -#14358 := [trans #14354 #14356]: #14357 -#14367 := [monotonicity #14358 #14364]: #14366 -#14370 := [monotonicity #13957 #14367]: #14369 -#14375 := [trans #14370 #14373]: #14374 -#14073 := (iff #12791 #14070) -#14052 := (or #13955 #14041) -#14055 := (or #13955 #14052) -#14058 := (or #13955 #14055) -#14061 := (or #13955 #14058) -#14064 := (or #14049 #14061) -#14067 := (or #13955 #14064) -#14071 := (iff #14067 #14070) -#14072 := [rewrite]: #14071 -#14068 := (iff #12791 #14067) -#14065 := (iff #12783 #14064) -#14062 := (iff #12774 #14061) -#14059 := (iff #12759 #14058) -#14056 := (iff #12751 #14055) -#14053 := (iff #12736 #14052) -#14042 := (iff #12728 #14041) -#14039 := (iff #12720 #14036) -#14012 := (or #12644 #14009) -#14015 := (or #12653 #14012) -#14018 := (or #12662 #14015) -#14021 := (or #12671 #14018) -#14024 := (or #13955 #14021) -#14027 := (or #13955 #14024) -#14030 := (or #13955 #14027) -#14033 := (or #12719 #14030) -#14037 := (iff #14033 #14036) -#14038 := [rewrite]: #14037 -#14034 := (iff #12720 #14033) -#14031 := (iff #12711 #14030) -#14028 := (iff #12696 #14027) -#14025 := (iff #12688 #14024) -#14022 := (iff #12672 #14021) -#14019 := (iff #12663 #14018) -#14016 := (iff #12654 #14015) -#14013 := (iff #12645 #14012) -#14010 := (iff #12629 #14009) -#14007 := (iff #12621 #14006) -#14004 := (iff #12600 #14003) -#14001 := (iff #12597 #13998) -#13992 := (and #13958 #5736) -#13995 := (and #3145 #13992) -#13999 := (iff #13995 #13998) -#14000 := [rewrite]: #13999 -#13996 := (iff #12597 #13995) -#13993 := (iff #12594 #13992) -#13961 := (iff #3140 #13958) -#13962 := [rewrite]: #13961 -#13994 := [monotonicity #13962 #5738]: #13993 -#13997 := [monotonicity #13994]: #13996 -#14002 := [trans #13997 #14000]: #14001 -#14005 := [quant-intro #14002]: #14004 -#13990 := (iff #12620 #13989) -#13987 := (iff #12591 #13986) -#13984 := (iff #12586 #13981) -#13975 := (or #13959 #13972) -#13978 := (or #5739 #13975) -#13982 := (iff #13978 #13981) -#13983 := [rewrite]: #13982 -#13979 := (iff #12586 #13978) -#13976 := (iff #12580 #13975) -#13973 := (iff #3141 #13972) -#13974 := [rewrite]: #13973 -#13968 := (iff #12579 #13959) -#13963 := (not #13958) -#13966 := (iff #13963 #13959) -#13967 := [rewrite]: #13966 -#13964 := (iff #12579 #13963) -#13965 := [monotonicity #13962]: #13964 -#13969 := [trans #13965 #13967]: #13968 -#13977 := [monotonicity #13969 #13974]: #13976 -#13980 := [monotonicity #5741 #13977]: #13979 -#13985 := [trans #13980 #13983]: #13984 -#13988 := [quant-intro #13985]: #13987 -#13991 := [monotonicity #13988]: #13990 -#14008 := [monotonicity #13991 #14005]: #14007 -#14011 := [monotonicity #13988 #14008]: #14010 -#14014 := [monotonicity #14011]: #14013 -#14017 := [monotonicity #14014]: #14016 -#14020 := [monotonicity #14017]: #14019 -#14023 := [monotonicity #14020]: #14022 -#14026 := [monotonicity #13957 #14023]: #14025 -#14029 := [monotonicity #13957 #14026]: #14028 -#14032 := [monotonicity #13957 #14029]: #14031 -#14035 := [monotonicity #14032]: #14034 -#14040 := [trans #14035 #14038]: #14039 -#14043 := [monotonicity #14040]: #14042 -#14054 := [monotonicity #13957 #14043]: #14053 -#14057 := [monotonicity #13957 #14054]: #14056 -#14060 := [monotonicity #13957 #14057]: #14059 -#14063 := [monotonicity #13957 #14060]: #14062 -#14050 := (iff #12782 #14049) -#14047 := (iff #3131 #14046) -#14048 := [rewrite]: #14047 -#14051 := [monotonicity #14048]: #14050 -#14066 := [monotonicity #14051 #14063]: #14065 -#14069 := [monotonicity #13957 #14066]: #14068 -#14074 := [trans #14069 #14072]: #14073 -#14378 := [monotonicity #14074 #14375]: #14377 -#14512 := [monotonicity #13957 #14378]: #14511 -#14515 := [monotonicity #14512]: #14514 -#14518 := [monotonicity #14515]: #14517 -#14521 := [monotonicity #14518]: #14520 -#14524 := [monotonicity #14521]: #14523 -#14527 := [monotonicity #14524]: #14526 -#14530 := [monotonicity #14527]: #14529 -#14400 := (iff #13385 #14399) -#14397 := (iff #12544 #12437) -#12517 := (and true true) -#14392 := (and #12517 #12437) -#14395 := (iff #14392 #12437) -#14396 := [rewrite]: #14395 -#14393 := (iff #12544 #14392) -#14390 := (iff #3115 #12517) -#14388 := (iff #3078 true) -#14389 := [rewrite]: #14388 -#14386 := (iff #3077 true) -#14381 := (forall (vars (?x775 T5)) (:pat #3075) true) -#14384 := (iff #14381 true) -#14385 := [elim-unused]: #14384 -#14382 := (iff #3077 #14381) -#14379 := (iff #3076 true) -#14380 := [rewrite]: #14379 -#14383 := [quant-intro #14380]: #14382 -#14387 := [trans #14383 #14385]: #14386 -#14391 := [monotonicity #14387 #14389]: #14390 -#14394 := [monotonicity #14391]: #14393 -#14398 := [trans #14394 #14396]: #14397 -#14401 := [monotonicity #14398]: #14400 -#14533 := [monotonicity #14401 #14530]: #14532 -#14536 := [monotonicity #14401 #14533]: #14535 -#14539 := [monotonicity #13957 #14536]: #14538 -#14542 := [monotonicity #13957 #14539]: #14541 -#14545 := [monotonicity #13957 #14542]: #14544 -#14548 := [monotonicity #13957 #14545]: #14547 -#14417 := (iff #13516 #14416) -#14414 := (iff #12432 #14411) -#14408 := (and #14405 #12426) -#14412 := (iff #14408 #14411) -#14413 := [rewrite]: #14412 -#14409 := (iff #12432 #14408) -#14406 := (iff #3070 #14405) -#14407 := [rewrite]: #14406 -#14410 := [monotonicity #14407]: #14409 -#14415 := [trans #14410 #14413]: #14414 -#14418 := [monotonicity #14415]: #14417 -#14551 := [monotonicity #14418 #14548]: #14550 -#14451 := (iff #13525 #14450) -#14448 := (iff #12423 #14447) -#14445 := (iff #12418 #14442) -#14436 := (or #14420 #14433) -#14439 := (or #5739 #14436) -#14443 := (iff #14439 #14442) -#14444 := [rewrite]: #14443 -#14440 := (iff #12418 #14439) -#14437 := (iff #12412 #14436) -#14434 := (iff #3063 #14433) -#14435 := [rewrite]: #14434 -#14429 := (iff #12411 #14420) -#14419 := (not #14420) -#14424 := (not #14419) -#14427 := (iff #14424 #14420) -#14428 := [rewrite]: #14427 -#14425 := (iff #12411 #14424) -#14422 := (iff #3062 #14419) -#14423 := [rewrite]: #14422 -#14426 := [monotonicity #14423]: #14425 -#14430 := [trans #14426 #14428]: #14429 -#14438 := [monotonicity #14430 #14435]: #14437 -#14441 := [monotonicity #5741 #14438]: #14440 -#14446 := [trans #14441 #14444]: #14445 -#14449 := [quant-intro #14446]: #14448 -#14452 := [monotonicity #14449]: #14451 -#14554 := [monotonicity #14452 #14551]: #14553 -#14457 := (iff #13534 #14456) -#14454 := (iff #3061 #14453) -#14455 := [rewrite]: #14454 -#14458 := [monotonicity #14455]: #14457 -#14557 := [monotonicity #14458 #14554]: #14556 -#14560 := [monotonicity #13957 #14557]: #14559 -#14469 := (iff #13551 #14468) -#14466 := (iff #3058 #14465) -#14463 := (iff #3057 #14462) -#14464 := [rewrite]: #14463 -#14460 := (iff #3056 #14459) -#14461 := [rewrite]: #14460 -#14467 := [monotonicity #14461 #14464]: #14466 -#14470 := [monotonicity #14467]: #14469 -#14563 := [monotonicity #14470 #14560]: #14562 -#14484 := (iff #13560 #14483) -#14481 := (iff #3054 #14478) -#14475 := (and #14471 #13947) -#14479 := (iff #14475 #14478) -#14480 := [rewrite]: #14479 -#14476 := (iff #3054 #14475) -#14473 := (iff #3052 #14471) -#14474 := [rewrite]: #14473 -#14477 := [monotonicity #14474 #13948]: #14476 -#14482 := [trans #14477 #14480]: #14481 -#14485 := [monotonicity #14482]: #14484 -#14566 := [monotonicity #14485 #14563]: #14565 -#14497 := (iff #13569 #14496) -#14494 := (iff #3050 #14493) -#14491 := (iff #3049 #14490) -#14492 := [rewrite]: #14491 -#14488 := (iff #3048 #14486) -#14489 := [rewrite]: #14488 -#14495 := [monotonicity #14489 #14492]: #14494 -#14498 := [monotonicity #14495]: #14497 -#14569 := [monotonicity #14498 #14566]: #14568 -#14508 := (iff #13585 #14507) -#14505 := (iff #12406 #14502) -#14503 := (iff #14499 #14502) -#14504 := [rewrite]: #14503 -#14500 := (iff #12406 #14499) -#13944 := (iff #2951 #13943) -#13945 := [rewrite]: #13944 -#14501 := [monotonicity #13945]: #14500 -#14506 := [trans #14501 #14504]: #14505 -#14509 := [monotonicity #14506]: #14508 -#14572 := [monotonicity #14509 #14569]: #14571 -#14577 := [trans #14572 #14575]: #14576 -#14580 := [monotonicity #14501 #14577]: #14579 -#14585 := [trans #14580 #14583]: #14584 -#13940 := (iff #13602 #13939) -#13937 := (iff #12398 #13936) -#13934 := (iff #12393 #13931) -#13925 := (or #13910 #13921) -#13928 := (or #5739 #13925) -#13932 := (iff #13928 #13931) -#13933 := [rewrite]: #13932 -#13929 := (iff #12393 #13928) -#13926 := (iff #12387 #13925) -#13920 := (iff #3041 #13921) -#13924 := [rewrite]: #13920 -#13918 := (iff #12386 #13910) -#13909 := (not #13910) -#13913 := (not #13909) -#13916 := (iff #13913 #13910) -#13917 := [rewrite]: #13916 -#13914 := (iff #12386 #13913) -#13911 := (iff #3038 #13909) -#13912 := [rewrite]: #13911 -#13915 := [monotonicity #13912]: #13914 -#13919 := [trans #13915 #13917]: #13918 -#13927 := [monotonicity #13919 #13924]: #13926 -#13930 := [monotonicity #5741 #13927]: #13929 -#13935 := [trans #13930 #13933]: #13934 -#13938 := [quant-intro #13935]: #13937 -#13941 := [monotonicity #13938]: #13940 -#14588 := [monotonicity #13941 #14585]: #14587 -#14591 := [monotonicity #13938 #14588]: #14590 -#13907 := (iff #13619 #13906) -#13904 := (iff #3037 #13903) -#13905 := [rewrite]: #13904 -#13908 := [monotonicity #13905]: #13907 -#14594 := [monotonicity #13908 #14591]: #14593 -#14597 := [monotonicity #13905 #14594]: #14596 -#13901 := (iff #13636 false) -#13444 := (iff #3294 false) -#13445 := [rewrite]: #13444 -#13899 := (iff #13636 #3294) -#13897 := (iff #12383 true) -#13892 := (and true #12517) -#13895 := (iff #13892 true) -#13896 := [rewrite]: #13895 -#13893 := (iff #12383 #13892) -#13890 := (iff #12380 #12517) -#13888 := (iff #3033 true) -#13889 := [rewrite]: #13888 -#13886 := (iff #3032 true) -#13887 := [rewrite]: #13886 -#13891 := [monotonicity #13887 #13889]: #13890 -#13894 := [monotonicity #13887 #13891]: #13893 -#13898 := [trans #13894 #13896]: #13897 -#13900 := [monotonicity #13898]: #13899 -#13902 := [trans #13900 #13445]: #13901 -#14600 := [monotonicity #13902 #14597]: #14599 -#14603 := [monotonicity #14600]: #14602 -#14606 := [monotonicity #14603]: #14605 -#14609 := [monotonicity #14606]: #14608 -#14612 := [monotonicity #14609]: #14611 -#14615 := [monotonicity #14612]: #14614 -#14620 := [trans #14615 #14618]: #14619 -#14623 := [monotonicity #14620]: #14622 -#14628 := [trans #14623 #14626]: #14627 -#14631 := [monotonicity #14628]: #14630 -#14634 := [monotonicity #14631]: #14633 -#14639 := [trans #14634 #14637]: #14638 -#14642 := [monotonicity #14639]: #14641 -#14645 := [monotonicity #14642]: #14644 -#13884 := (iff #13732 #13883) -#13881 := (iff #3007 #13880) -#13878 := (iff #3006 #13877) -#13879 := [rewrite]: #13878 -#13875 := (iff #3005 #13872) -#13876 := [rewrite]: #13875 -#13882 := [monotonicity #13876 #13879]: #13881 -#13885 := [monotonicity #13882]: #13884 -#14726 := [monotonicity #13885 #14645]: #14725 -#14729 := [monotonicity #14726]: #14728 -#14732 := [monotonicity #14729]: #14731 -#14735 := [monotonicity #14732]: #14734 -#14738 := [monotonicity #14735]: #14737 -#14656 := (iff #13777 #14655) -#14653 := (iff #2989 #14652) -#14650 := (iff #2988 #14646) -#14651 := [rewrite]: #14650 -#14654 := [quant-intro #14651]: #14653 -#14657 := [monotonicity #14654]: #14656 -#14741 := [monotonicity #14657 #14738]: #14740 -#14744 := [monotonicity #14741]: #14743 -#14747 := [monotonicity #14744]: #14746 -#14662 := (iff #13811 #14661) -#14659 := (iff #12326 #14658) -#14660 := [rewrite]: #14659 -#14663 := [monotonicity #14660]: #14662 -#14750 := [monotonicity #14663 #14747]: #14749 -#14669 := (iff #13820 #13942) -#14664 := (not #13943) -#14667 := (iff #14664 #13942) -#14668 := [rewrite]: #14667 -#14665 := (iff #13820 #14664) -#14666 := [monotonicity #13945]: #14665 -#14670 := [trans #14666 #14668]: #14669 -#14753 := [monotonicity #14670 #14750]: #14752 -#14680 := (iff #13829 #14671) -#14672 := (not #14671) -#14675 := (not #14672) -#14678 := (iff #14675 #14671) -#14679 := [rewrite]: #14678 -#14676 := (iff #13829 #14675) -#14673 := (iff #2950 #14672) -#14674 := [rewrite]: #14673 -#14677 := [monotonicity #14674]: #14676 -#14681 := [trans #14677 #14679]: #14680 -#14756 := [monotonicity #14681 #14753]: #14755 -#14694 := (iff #13838 #14693) -#14691 := (iff #2948 #14690) -#14688 := (iff #2947 #14687) -#14689 := [rewrite]: #14688 -#14685 := (iff #2946 #14682) -#14686 := [rewrite]: #14685 -#14692 := [monotonicity #14686 #14689]: #14691 -#14695 := [monotonicity #14692]: #14694 -#14759 := [monotonicity #14695 #14756]: #14758 -#14708 := (iff #13847 #14707) -#14705 := (iff #2944 #14704) -#14702 := (iff #2943 #14701) -#14703 := [rewrite]: #14702 -#14699 := (iff #2942 #14696) -#14700 := [rewrite]: #14699 -#14706 := [monotonicity #14700 #14703]: #14705 -#14709 := [monotonicity #14706]: #14708 -#14762 := [monotonicity #14709 #14759]: #14761 -#14722 := (iff #13856 #14721) -#14719 := (iff #2940 #14718) -#14716 := (iff #2939 #14715) -#14717 := [rewrite]: #14716 -#14713 := (iff #2938 #14710) -#14714 := [rewrite]: #14713 -#14720 := [monotonicity #14714 #14717]: #14719 -#14723 := [monotonicity #14720]: #14722 -#14765 := [monotonicity #14723 #14762]: #14764 -#14770 := [trans #14765 #14768]: #14769 -#14773 := [monotonicity #14770]: #14772 -#13870 := (iff #3345 #13869) -#13867 := (iff #3344 #13857) -#13862 := (implies true #13857) -#13865 := (iff #13862 #13857) -#13866 := [rewrite]: #13865 -#13863 := (iff #3344 #13862) -#13860 := (iff #3343 #13857) -#13853 := (implies #2940 #13848) -#13858 := (iff #13853 #13857) -#13859 := [rewrite]: #13858 -#13854 := (iff #3343 #13853) -#13851 := (iff #3342 #13848) -#13844 := (implies #2944 #13839) -#13849 := (iff #13844 #13848) -#13850 := [rewrite]: #13849 -#13845 := (iff #3342 #13844) -#13842 := (iff #3341 #13839) -#13835 := (implies #2948 #13830) -#13840 := (iff #13835 #13839) -#13841 := [rewrite]: #13840 -#13836 := (iff #3341 #13835) -#13833 := (iff #3340 #13830) -#13826 := (implies #2950 #13821) -#13831 := (iff #13826 #13830) -#13832 := [rewrite]: #13831 -#13827 := (iff #3340 #13826) -#13824 := (iff #3339 #13821) -#13817 := (implies #2951 #13812) -#13822 := (iff #13817 #13821) -#13823 := [rewrite]: #13822 -#13818 := (iff #3339 #13817) -#13815 := (iff #3338 #13812) -#13808 := (implies #12326 #13796) -#13813 := (iff #13808 #13812) -#13814 := [rewrite]: #13813 -#13809 := (iff #3338 #13808) -#13806 := (iff #3337 #13796) -#13801 := (implies true #13796) -#13804 := (iff #13801 #13796) -#13805 := [rewrite]: #13804 -#13802 := (iff #3337 #13801) -#13799 := (iff #3336 #13796) -#13792 := (implies #12329 #13787) -#13797 := (iff #13792 #13796) -#13798 := [rewrite]: #13797 -#13793 := (iff #3336 #13792) -#13790 := (iff #3335 #13787) -#13783 := (implies #12338 #13778) -#13788 := (iff #13783 #13787) -#13789 := [rewrite]: #13788 -#13784 := (iff #3335 #13783) -#13781 := (iff #3334 #13778) -#13774 := (implies #2989 #13769) -#13779 := (iff #13774 #13778) -#13780 := [rewrite]: #13779 -#13775 := (iff #3334 #13774) -#13772 := (iff #3333 #13769) -#13765 := (implies #2994 #13760) -#13770 := (iff #13765 #13769) -#13771 := [rewrite]: #13770 -#13766 := (iff #3333 #13765) -#13763 := (iff #3332 #13760) -#13756 := (implies #2996 #13751) -#13761 := (iff #13756 #13760) -#13762 := [rewrite]: #13761 -#13757 := (iff #3332 #13756) -#13754 := (iff #3331 #13751) -#13747 := (implies #2999 #13742) -#13752 := (iff #13747 #13751) -#13753 := [rewrite]: #13752 -#13748 := (iff #3331 #13747) -#13745 := (iff #3330 #13742) -#13738 := (implies #12352 #13733) -#13743 := (iff #13738 #13742) -#13744 := [rewrite]: #13743 -#13739 := (iff #3330 #13738) -#13736 := (iff #3329 #13733) -#13729 := (implies #3007 #13724) -#13734 := (iff #13729 #13733) -#13735 := [rewrite]: #13734 -#13730 := (iff #3329 #13729) -#13727 := (iff #3328 #13724) -#13721 := (and #13716 #12355) -#13725 := (iff #13721 #13724) -#13726 := [rewrite]: #13725 -#13722 := (iff #3328 #13721) -#12356 := (iff #3010 #12355) -#12357 := [rewrite]: #12356 -#13719 := (iff #3327 #13716) -#13712 := (implies #12355 #13707) -#13717 := (iff #13712 #13716) -#13718 := [rewrite]: #13717 -#13713 := (iff #3327 #13712) -#13710 := (iff #3326 #13707) -#13704 := (and #13699 #12364) -#13708 := (iff #13704 #13707) -#13709 := [rewrite]: #13708 -#13705 := (iff #3326 #13704) -#12365 := (iff #3016 #12364) -#12362 := (iff #3015 #12361) -#12363 := [rewrite]: #12362 -#12359 := (iff #3013 #12358) -#12360 := [rewrite]: #12359 -#12366 := [monotonicity #12360 #12363]: #12365 -#13702 := (iff #3325 #13699) -#13695 := (implies #12364 #13690) -#13700 := (iff #13695 #13699) -#13701 := [rewrite]: #13700 -#13696 := (iff #3325 #13695) -#13693 := (iff #3324 #13690) -#13687 := (and #13682 #12373) -#13691 := (iff #13687 #13690) -#13692 := [rewrite]: #13691 -#13688 := (iff #3324 #13687) -#12376 := (iff #3019 #12373) -#12370 := (and #12367 #12361) -#12374 := (iff #12370 #12373) -#12375 := [rewrite]: #12374 -#12371 := (iff #3019 #12370) -#12368 := (iff #3018 #12367) -#12369 := [rewrite]: #12368 -#12372 := [monotonicity #12369 #12363]: #12371 -#12377 := [trans #12372 #12375]: #12376 -#13685 := (iff #3323 #13682) -#13678 := (implies #12373 #13673) -#13683 := (iff #13678 #13682) -#13684 := [rewrite]: #13683 -#13679 := (iff #3323 #13678) -#13676 := (iff #3322 #13673) -#13669 := (implies #3022 #13664) -#13674 := (iff #13669 #13673) -#13675 := [rewrite]: #13674 -#13670 := (iff #3322 #13669) -#13667 := (iff #3321 #13664) -#13660 := (implies #3025 #13655) -#13665 := (iff #13660 #13664) -#13666 := [rewrite]: #13665 -#13661 := (iff #3321 #13660) -#13658 := (iff #3320 #13655) -#13651 := (implies #3028 #13646) -#13656 := (iff #13651 #13655) -#13657 := [rewrite]: #13656 -#13652 := (iff #3320 #13651) -#13649 := (iff #3319 #13646) -#13642 := (implies #3031 #13637) -#13647 := (iff #13642 #13646) -#13648 := [rewrite]: #13647 -#13643 := (iff #3319 #13642) -#13640 := (iff #3318 #13637) -#13633 := (implies #12383 #13628) -#13638 := (iff #13633 #13637) -#13639 := [rewrite]: #13638 -#13634 := (iff #3318 #13633) -#13631 := (iff #3317 #13628) -#13625 := (and #13620 #3037) +#12678 := (or #12112 #12663) +#12686 := (or #12112 #12678) +#12701 := (or #12112 #12686) +#12709 := (not #3257) +#12710 := (or #12709 #12701) +#12718 := (or #12112 #12710) +#12516 := (or #12515 #12500) +#12525 := (or #12524 #12516) +#12540 := (not #3064) +#12541 := (or #12540 #12525) +#12550 := (or #12549 #12541) +#12559 := (or #12558 #12550) +#12568 := (or #12567 #12559) +#12577 := (or #12576 #12568) +#12582 := (and #12357 #12577) +#12589 := (or #12588 #12582) +#12594 := (and #12351 #12589) +#12600 := (or #12112 #12594) +#12615 := (or #12112 #12600) +#12623 := (not #3190) +#12624 := (or #12623 #12615) +#12632 := (or #12112 #12624) +#12730 := (and #12632 #12718) +#12736 := (or #12112 #12730) +#12744 := (or #12576 #12736) +#12749 := (and #12357 #12744) +#12755 := (or #12588 #12749) +#12760 := (and #12351 #12755) +#12766 := (or #12112 #12760) +#12781 := (or #12112 #12766) +#12789 := (not #3179) +#12790 := (or #12789 #12781) +#12798 := (or #12112 #12790) +#12857 := (and #12798 #12845) +#12863 := (or #12112 #12857) +#12879 := (or #12878 #12863) +#12888 := (or #12887 #12879) +#12897 := (or #12896 #12888) +#12906 := (or #12905 #12897) +#12915 := (or #12914 #12906) +#12923 := (or #12179 #12915) +#12309 := (and #3160 #12306) +#12312 := (and #3156 #12309) +#12931 := (not #12312) +#12932 := (or #12931 #12923) +#12940 := (or #12931 #12932) +#12948 := (or #12112 #12940) +#12963 := (or #12112 #12948) +#12978 := (or #12112 #12963) +#13000 := (or #12112 #12978) +#11995 := (and #3072 #11992) +#13008 := (not #11995) +#13009 := (or #13008 #13000) +#11977 := (not #3067) +#11978 := (or #11977 #3068) +#11984 := (or #5592 #11978) +#11989 := (forall (vars (?x775 int)) #11984) +#13017 := (not #11989) +#13018 := (or #13017 #13009) +#13026 := (not #3066) +#13027 := (or #13026 #13018) +#13035 := (or #12112 #13027) +#13043 := (not #3063) +#13044 := (or #13043 #13035) +#13052 := (not #3059) +#13053 := (or #13052 #13044) +#13061 := (not #3055) +#13062 := (or #13061 #13053) +#11971 := (and #2956 #3027) +#13077 := (not #11971) +#13078 := (or #13077 #13062) +#13083 := (and #11971 #13078) +#11954 := (not #3043) +#11955 := (or #11954 #3046) +#11961 := (or #5592 #11955) +#11966 := (forall (vars (?x773 int)) #11961) +#13089 := (not #11966) +#13090 := (or #13089 #13083) +#13095 := (and #11966 #13090) +#13101 := (not #3042) +#13102 := (or #13101 #13095) +#13107 := (and #3042 #13102) +#11948 := (and #3037 #3038) +#11951 := (and #3037 #11948) +#13113 := (not #11951) +#13114 := (or #13113 #13107) +#13123 := (or #13122 #13114) +#13132 := (or #13131 #13123) +#13141 := (or #13140 #13132) +#13150 := (or #13149 #13141) +#13159 := (or #13158 #13150) +#13164 := (and #11943 #13159) +#13171 := (or #13170 #13164) +#13176 := (and #11937 #13171) +#13183 := (or #13182 #13176) +#13188 := (and #11928 #13183) +#13194 := (not #3012) +#13195 := (or #13194 #13188) +#13204 := (or #13203 #13195) +#13213 := (or #13212 #13204) +#13222 := (or #13221 #13213) +#13231 := (or #13230 #13222) +#13239 := (not #2994) +#13240 := (or #13239 #13231) +#13249 := (or #13248 #13240) +#13258 := (or #13257 #13249) +#11887 := (and #11881 #11884) +#11890 := (and #11875 #11887) +#11893 := (and #11872 #11890) +#11896 := (and #11869 #11893) +#11899 := (and #11866 #11896) +#13273 := (not #11899) +#13274 := (or #13273 #13258) +#13282 := (not #2956) +#13283 := (or #13282 #13274) +#13291 := (not #2955) +#13292 := (or #13291 #13283) +#13300 := (not #2953) +#13301 := (or #13300 #13292) +#13309 := (not #2949) +#13310 := (or #13309 #13301) +#13318 := (not #2945) +#13319 := (or #13318 #13310) +#13331 := (not #13319) +#14238 := (iff #13331 #14237) +#14235 := (iff #13319 #14232) +#14190 := (or #13345 #14109) +#14193 := (or #13203 #14190) +#14196 := (or #13212 #14193) +#14199 := (or #13221 #14196) +#14202 := (or #13230 #14199) +#14205 := (or #14121 #14202) +#14208 := (or #13248 #14205) +#14211 := (or #13257 #14208) +#14214 := (or #14127 #14211) +#14217 := (or #13404 #14214) +#14220 := (or #14137 #14217) +#14223 := (or #14159 #14220) +#14226 := (or #14173 #14223) +#14229 := (or #14187 #14226) +#14233 := (iff #14229 #14232) +#14234 := [rewrite]: #14233 +#14230 := (iff #13319 #14229) +#14227 := (iff #13310 #14226) +#14224 := (iff #13301 #14223) +#14221 := (iff #13292 #14220) +#14218 := (iff #13283 #14217) +#14215 := (iff #13274 #14214) +#14212 := (iff #13258 #14211) +#14209 := (iff #13249 #14208) +#14206 := (iff #13240 #14205) +#14203 := (iff #13231 #14202) +#14200 := (iff #13222 #14199) +#14197 := (iff #13213 #14196) +#14194 := (iff #13204 #14193) +#14191 := (iff #13195 #14190) +#14110 := (iff #13188 #14109) +#14107 := (iff #13183 #14106) +#14104 := (iff #13176 #14101) +#14098 := (and #11937 #14095) +#14102 := (iff #14098 #14101) +#14103 := [rewrite]: #14102 +#14099 := (iff #13176 #14098) +#14096 := (iff #13171 #14095) +#14093 := (iff #13164 #14090) +#14087 := (and #11943 #14082) +#14091 := (iff #14087 #14090) +#14092 := [rewrite]: #14091 +#14088 := (iff #13164 #14087) +#14085 := (iff #13159 #14082) +#14064 := (or false #14061) +#14067 := (or #13122 #14064) +#14070 := (or #13131 #14067) +#14073 := (or #13140 #14070) +#14076 := (or #13149 #14073) +#14079 := (or #13158 #14076) +#14083 := (iff #14079 #14082) +#14084 := [rewrite]: #14083 +#14080 := (iff #13159 #14079) +#14077 := (iff #13150 #14076) +#14074 := (iff #13141 #14073) +#14071 := (iff #13132 #14070) +#14068 := (iff #13123 #14067) +#14065 := (iff #13114 #14064) +#14062 := (iff #13107 #14061) +#14059 := (iff #13102 #14058) +#14056 := (iff #13095 #14055) +#14053 := (iff #13090 #14052) +#14050 := (iff #13083 #14047) +#13965 := (and #13405 #3027) +#14044 := (and #13965 #14039) +#14048 := (iff #14044 #14047) +#14049 := [rewrite]: #14048 +#14045 := (iff #13083 #14044) +#14042 := (iff #13078 #14039) +#13976 := (or #13417 #13846) +#13979 := (or #12878 #13976) +#13982 := (or #12887 #13979) +#13985 := (or #12896 #13982) +#13988 := (or #12905 #13985) +#13991 := (or #12914 #13988) +#13994 := (or #12179 #13991) +#13997 := (or #13870 #13994) +#14000 := (or #13870 #13997) +#14003 := (or #13417 #14000) +#14006 := (or #13417 #14003) +#14009 := (or #13417 #14006) +#14012 := (or #13417 #14009) +#14015 := (or #13887 #14012) +#14018 := (or #13921 #14015) +#14021 := (or #13927 #14018) +#14024 := (or #13417 #14021) +#14027 := (or #13939 #14024) +#14030 := (or #13949 #14027) +#14033 := (or #13962 #14030) +#14036 := (or #13973 #14033) +#14040 := (iff #14036 #14039) +#14041 := [rewrite]: #14040 +#14037 := (iff #13078 #14036) +#14034 := (iff #13062 #14033) +#14031 := (iff #13053 #14030) +#14028 := (iff #13044 #14027) +#14025 := (iff #13035 #14024) +#14022 := (iff #13027 #14021) +#14019 := (iff #13018 #14018) +#14016 := (iff #13009 #14015) +#14013 := (iff #13000 #14012) +#14010 := (iff #12978 #14009) +#14007 := (iff #12963 #14006) +#14004 := (iff #12948 #14003) +#14001 := (iff #12940 #14000) +#13998 := (iff #12932 #13997) +#13995 := (iff #12923 #13994) +#13992 := (iff #12915 #13991) +#13989 := (iff #12906 #13988) +#13986 := (iff #12897 #13985) +#13983 := (iff #12888 #13982) +#13980 := (iff #12879 #13979) +#13977 := (iff #12863 #13976) +#13847 := (iff #12857 #13846) +#13844 := (iff #12845 #13841) +#13823 := (or #13417 #13816) +#13826 := (or #13417 #13823) +#13829 := (or #13417 #13826) +#13832 := (or #13417 #13829) +#13835 := (or #13698 #13832) +#13838 := (or #13417 #13835) +#13842 := (iff #13838 #13841) +#13843 := [rewrite]: #13842 +#13839 := (iff #12845 #13838) +#13836 := (iff #12837 #13835) +#13833 := (iff #12828 #13832) +#13830 := (iff #12813 #13829) +#13827 := (iff #12171 #13826) +#13824 := (iff #12156 #13823) +#13817 := (iff #12150 #13816) +#13814 := (iff #12145 #13811) +#13787 := (or #12069 #13784) +#13790 := (or #12078 #13787) +#13793 := (or #12087 #13790) +#13796 := (or #12096 #13793) +#13799 := (or #13417 #13796) +#13802 := (or #13417 #13799) +#13805 := (or #13417 #13802) +#13808 := (or #12144 #13805) +#13812 := (iff #13808 #13811) +#13813 := [rewrite]: #13812 +#13809 := (iff #12145 #13808) +#13806 := (iff #12136 #13805) +#13803 := (iff #12121 #13802) +#13800 := (iff #12113 #13799) +#13797 := (iff #12097 #13796) +#13794 := (iff #12088 #13793) +#13791 := (iff #12079 #13790) +#13788 := (iff #12070 #13787) +#13785 := (iff #12056 #13784) +#13782 := (iff #12051 #13779) +#13776 := (or #13770 #13773) +#13780 := (iff #13776 #13779) +#13781 := [rewrite]: #13780 +#13777 := (iff #12051 #13776) +#13774 := (iff #12050 #13773) +#13754 := (iff #12035 #13753) +#13751 := (iff #12030 #13748) +#13742 := (or #13725 #13739) +#13745 := (or #5601 #13742) +#13749 := (iff #13745 #13748) +#13750 := [rewrite]: #13749 +#13746 := (iff #12030 #13745) +#13743 := (iff #12024 #13742) +#13740 := (iff #3091 #13739) +#13741 := [rewrite]: #13740 +#13735 := (iff #12023 #13725) +#13730 := (not #13727) +#13733 := (iff #13730 #13725) +#13734 := [rewrite]: #13733 +#13731 := (iff #12023 #13730) +#13728 := (iff #3090 #13727) +#13729 := [rewrite]: #13728 +#13732 := [monotonicity #13729]: #13731 +#13736 := [trans #13732 #13734]: #13735 +#13744 := [monotonicity #13736 #13741]: #13743 +#5602 := (iff #5592 #5601) +#4387 := (iff #412 #4386) +#4380 := (iff #411 #4379) +#4381 := [rewrite]: #4380 +#4063 := (iff #285 #4065) +#4064 := [rewrite]: #4063 +#4388 := [monotonicity #4064 #4381]: #4387 +#5603 := [monotonicity #4388]: #5602 +#13747 := [monotonicity #5603 #13744]: #13746 +#13752 := [trans #13747 #13750]: #13751 +#13755 := [quant-intro #13752]: #13754 +#13775 := [monotonicity #13755]: #13774 +#13771 := (iff #3099 #13770) +#13768 := (iff #3098 #13765) +#13756 := (and #13727 #3095) +#13759 := (and #4379 #13756) +#13762 := (and #4065 #13759) +#13766 := (iff #13762 #13765) +#13767 := [rewrite]: #13766 +#13763 := (iff #3098 #13762) +#13760 := (iff #3097 #13759) +#13757 := (iff #3096 #13756) +#13758 := [monotonicity #13729]: #13757 +#13761 := [monotonicity #4381 #13758]: #13760 +#13764 := [monotonicity #4064 #13761]: #13763 +#13769 := [trans #13764 #13767]: #13768 +#13772 := [quant-intro #13769]: #13771 +#13778 := [monotonicity #13772 #13775]: #13777 +#13783 := [trans #13778 #13781]: #13782 +#13786 := [monotonicity #13755 #13783]: #13785 +#13789 := [monotonicity #13786]: #13788 +#13792 := [monotonicity #13789]: #13791 +#13795 := [monotonicity #13792]: #13794 +#13798 := [monotonicity #13795]: #13797 +#13418 := (iff #12112 #13417) +#13415 := (iff #11974 #13414) +#13411 := (iff #3064 #13412) +#13413 := [rewrite]: #13411 +#13408 := (iff #3057 #13409) +#13410 := [rewrite]: #13408 +#13416 := [monotonicity #13410 #13413]: #13415 +#13419 := [monotonicity #13416]: #13418 +#13801 := [monotonicity #13419 #13798]: #13800 +#13804 := [monotonicity #13419 #13801]: #13803 +#13807 := [monotonicity #13419 #13804]: #13806 +#13810 := [monotonicity #13807]: #13809 +#13815 := [trans #13810 #13813]: #13814 +#13818 := [monotonicity #13815]: #13817 +#13825 := [monotonicity #13419 #13818]: #13824 +#13828 := [monotonicity #13419 #13825]: #13827 +#13831 := [monotonicity #13419 #13828]: #13830 +#13834 := [monotonicity #13419 #13831]: #13833 +#13821 := (iff #12836 #13698) +#13819 := (iff #3282 #13697) +#13820 := [rewrite]: #13819 +#13822 := [monotonicity #13820]: #13821 +#13837 := [monotonicity #13822 #13834]: #13836 +#13840 := [monotonicity #13419 #13837]: #13839 +#13845 := [trans #13840 #13843]: #13844 +#13723 := (iff #12798 #13720) +#13708 := (or #13417 #13691) +#13711 := (or #13417 #13708) +#13714 := (or #13697 #13711) +#13717 := (or #13417 #13714) +#13721 := (iff #13717 #13720) +#13722 := [rewrite]: #13721 +#13718 := (iff #12798 #13717) +#13715 := (iff #12790 #13714) +#13712 := (iff #12781 #13711) +#13709 := (iff #12766 #13708) +#13694 := (iff #12760 #13691) +#13688 := (and #12351 #13685) +#13692 := (iff #13688 #13691) +#13693 := [rewrite]: #13692 +#13689 := (iff #12760 #13688) +#13686 := (iff #12755 #13685) +#13683 := (iff #12749 #13680) +#13677 := (and #12357 #13672) +#13681 := (iff #13677 #13680) +#13682 := [rewrite]: #13681 +#13678 := (iff #12749 #13677) +#13675 := (iff #12744 #13672) +#13666 := (or #13417 #13663) +#13669 := (or #12576 #13666) +#13673 := (iff #13669 #13672) +#13674 := [rewrite]: #13673 +#13670 := (iff #12744 #13669) +#13667 := (iff #12736 #13666) +#13664 := (iff #12730 #13663) +#13661 := (iff #12718 #13658) +#13555 := (or #13426 #13547) +#13637 := (or #13555 #12653) +#13640 := (or #12662 #13637) +#13643 := (or #13417 #13640) +#13646 := (or #13417 #13643) +#13649 := (or #13417 #13646) +#13652 := (or #13603 #13649) +#13655 := (or #13417 #13652) +#13659 := (iff #13655 #13658) +#13660 := [rewrite]: #13659 +#13656 := (iff #12718 #13655) +#13653 := (iff #12710 #13652) +#13650 := (iff #12701 #13649) +#13647 := (iff #12686 #13646) +#13644 := (iff #12678 #13643) +#13641 := (iff #12663 #13640) +#13638 := (iff #12654 #13637) +#13556 := (iff #12500 #13555) +#13550 := (iff #12493 #13547) +#13544 := (and #13521 #13539) +#13548 := (iff #13544 #13547) +#13549 := [rewrite]: #13548 +#13545 := (iff #12493 #13544) +#13542 := (iff #12488 #13539) +#13527 := (or #13443 #13511) +#13530 := (or #12469 #13527) +#13533 := (or #13518 #13530) +#13536 := (or #13524 #13533) +#13540 := (iff #13536 #13539) +#13541 := [rewrite]: #13540 +#13537 := (iff #12488 #13536) +#13534 := (iff #12479 #13533) +#13531 := (iff #12470 #13530) +#13528 := (iff #12461 #13527) +#13512 := (iff #12447 #13511) +#13509 := (iff #12442 #13508) +#13506 := (iff #12435 #13505) +#13503 := (iff #12430 #13502) +#13500 := (iff #12407 #13497) +#13494 := (and #13491 #12404) +#13498 := (iff #13494 #13497) +#13499 := [rewrite]: #13498 +#13495 := (iff #12407 #13494) +#13492 := (iff #3220 #13491) +#13493 := [rewrite]: #13492 +#13496 := [monotonicity #13493]: #13495 +#13501 := [trans #13496 #13499]: #13500 +#13486 := (iff #12429 #13485) +#13483 := (iff #12401 #13482) +#13480 := (iff #12396 #13477) +#13471 := (or #13454 #13468) +#13474 := (or #5601 #13471) +#13478 := (iff #13474 #13477) +#13479 := [rewrite]: #13478 +#13475 := (iff #12396 #13474) +#13472 := (iff #12390 #13471) +#13469 := (iff #3216 #13468) +#13470 := [rewrite]: #13469 +#13464 := (iff #12389 #13454) +#13456 := (not #13454) +#13459 := (not #13456) +#13462 := (iff #13459 #13454) +#13463 := [rewrite]: #13462 +#13460 := (iff #12389 #13459) +#13457 := (iff #3215 #13456) +#13458 := [rewrite]: #13457 +#13461 := [monotonicity #13458]: #13460 +#13465 := [trans #13461 #13463]: #13464 +#13473 := [monotonicity #13465 #13470]: #13472 +#13476 := [monotonicity #5603 #13473]: #13475 +#13481 := [trans #13476 #13479]: #13480 +#13484 := [quant-intro #13481]: #13483 +#13487 := [monotonicity #13484]: #13486 +#13504 := [monotonicity #13487 #13501]: #13503 +#13507 := [monotonicity #13484 #13504]: #13506 +#13452 := (iff #12441 #13451) +#13449 := (iff #3214 #13446) +#13450 := [rewrite]: #13449 +#13453 := [monotonicity #13450]: #13452 +#13510 := [monotonicity #13453 #13507]: #13509 +#13513 := [monotonicity #13450 #13510]: #13512 +#13444 := (iff #12460 #13443) +#13441 := (iff #12386 #13440) +#13437 := (iff #3212 #13438) +#13439 := [rewrite]: #13437 +#13420 := (iff #3202 #13421) +#13422 := [rewrite]: #13420 +#13442 := [monotonicity #13422 #13439]: #13441 +#13445 := [monotonicity #13442]: #13444 +#13529 := [monotonicity #13445 #13513]: #13528 +#13532 := [monotonicity #13529]: #13531 +#13519 := (iff #12478 #13518) +#13516 := (iff #12383 #13514) +#13517 := [rewrite]: #13516 +#13520 := [monotonicity #13517]: #13519 +#13535 := [monotonicity #13520 #13532]: #13534 +#13525 := (iff #12487 #13524) +#13522 := (iff #12380 #13521) +#13435 := (iff #12377 #13432) +#13436 := [rewrite]: #13435 +#13429 := (iff #12374 #13430) +#13431 := [rewrite]: #13429 +#13523 := [monotonicity #13431 #13436]: #13522 +#13526 := [monotonicity #13523]: #13525 +#13538 := [monotonicity #13526 #13535]: #13537 +#13543 := [trans #13538 #13541]: #13542 +#13546 := [monotonicity #13523 #13543]: #13545 +#13551 := [trans #13546 #13549]: #13550 +#13427 := (iff #12499 #13426) +#13424 := (iff #3203 #13423) +#13425 := [monotonicity #13413 #13422]: #13424 +#13428 := [monotonicity #13425]: #13427 +#13557 := [monotonicity #13428 #13551]: #13556 +#13639 := [monotonicity #13557]: #13638 +#13642 := [monotonicity #13639]: #13641 +#13645 := [monotonicity #13419 #13642]: #13644 +#13648 := [monotonicity #13419 #13645]: #13647 +#13651 := [monotonicity #13419 #13648]: #13650 +#13635 := (iff #12709 #13603) +#13633 := (iff #3257 #13604) +#13634 := [rewrite]: #13633 +#13636 := [monotonicity #13634]: #13635 +#13654 := [monotonicity #13636 #13651]: #13653 +#13657 := [monotonicity #13419 #13654]: #13656 +#13662 := [trans #13657 #13660]: #13661 +#13631 := (iff #12632 #13628) +#13616 := (or #13417 #13598) +#13619 := (or #13417 #13616) +#13622 := (or #13604 #13619) +#13625 := (or #13417 #13622) #13629 := (iff #13625 #13628) #13630 := [rewrite]: #13629 -#13626 := (iff #3317 #13625) -#13623 := (iff #3316 #13620) -#13616 := (implies #3037 #13611) -#13621 := (iff #13616 #13620) -#13622 := [rewrite]: #13621 -#13617 := (iff #3316 #13616) -#13614 := (iff #3315 #13611) -#13608 := (and #13603 #12398) -#13612 := (iff #13608 #13611) +#13626 := (iff #12632 #13625) +#13623 := (iff #12624 #13622) +#13620 := (iff #12615 #13619) +#13617 := (iff #12600 #13616) +#13601 := (iff #12594 #13598) +#13595 := (and #12351 #13592) +#13599 := (iff #13595 #13598) +#13600 := [rewrite]: #13599 +#13596 := (iff #12594 #13595) +#13593 := (iff #12589 #13592) +#13590 := (iff #12582 #13587) +#13584 := (and #12357 #13579) +#13588 := (iff #13584 #13587) +#13589 := [rewrite]: #13588 +#13585 := (iff #12582 #13584) +#13582 := (iff #12577 #13579) +#13558 := (or #12515 #13555) +#13561 := (or #12524 #13558) +#13564 := (or #13552 #13561) +#13567 := (or #12549 #13564) +#13570 := (or #12558 #13567) +#13573 := (or #12567 #13570) +#13576 := (or #12576 #13573) +#13580 := (iff #13576 #13579) +#13581 := [rewrite]: #13580 +#13577 := (iff #12577 #13576) +#13574 := (iff #12568 #13573) +#13571 := (iff #12559 #13570) +#13568 := (iff #12550 #13567) +#13565 := (iff #12541 #13564) +#13562 := (iff #12525 #13561) +#13559 := (iff #12516 #13558) +#13560 := [monotonicity #13557]: #13559 +#13563 := [monotonicity #13560]: #13562 +#13553 := (iff #12540 #13552) +#13554 := [monotonicity #13413]: #13553 +#13566 := [monotonicity #13554 #13563]: #13565 +#13569 := [monotonicity #13566]: #13568 +#13572 := [monotonicity #13569]: #13571 +#13575 := [monotonicity #13572]: #13574 +#13578 := [monotonicity #13575]: #13577 +#13583 := [trans #13578 #13581]: #13582 +#13586 := [monotonicity #13583]: #13585 +#13591 := [trans #13586 #13589]: #13590 +#13594 := [monotonicity #13591]: #13593 +#13597 := [monotonicity #13594]: #13596 +#13602 := [trans #13597 #13600]: #13601 +#13618 := [monotonicity #13419 #13602]: #13617 +#13621 := [monotonicity #13419 #13618]: #13620 +#13614 := (iff #12623 #13604) +#13609 := (not #13603) +#13612 := (iff #13609 #13604) #13613 := [rewrite]: #13612 -#13609 := (iff #3315 #13608) -#12399 := (iff #3044 #12398) -#12396 := (iff #3043 #12393) -#12390 := (implies #5718 #12387) -#12394 := (iff #12390 #12393) -#12395 := [rewrite]: #12394 -#12391 := (iff #3043 #12390) -#12388 := (iff #3042 #12387) -#12389 := [rewrite]: #12388 -#5719 := (iff #645 #5718) -#5720 := [rewrite]: #5719 -#12392 := [monotonicity #5720 #12389]: #12391 -#12397 := [trans #12392 #12395]: #12396 -#12400 := [quant-intro #12397]: #12399 -#13606 := (iff #3314 #13603) -#13599 := (implies #12398 #13594) -#13604 := (iff #13599 #13603) -#13605 := [rewrite]: #13604 -#13600 := (iff #3314 #13599) -#13597 := (iff #3313 #13594) -#13591 := (and #13586 #12406) -#13595 := (iff #13591 #13594) -#13596 := [rewrite]: #13595 -#13592 := (iff #3313 #13591) -#12409 := (iff #3046 #12406) -#12403 := (and #3022 #2951) -#12407 := (iff #12403 #12406) -#12408 := [rewrite]: #12407 -#12404 := (iff #3046 #12403) -#12401 := (iff #3045 #3022) -#12402 := [rewrite]: #12401 -#12405 := [monotonicity #12402]: #12404 -#12410 := [trans #12405 #12408]: #12409 -#13589 := (iff #3312 #13586) -#13582 := (implies #12406 #13570) -#13587 := (iff #13582 #13586) -#13588 := [rewrite]: #13587 -#13583 := (iff #3312 #13582) -#13580 := (iff #3311 #13570) -#13575 := (implies true #13570) -#13578 := (iff #13575 #13570) -#13579 := [rewrite]: #13578 -#13576 := (iff #3311 #13575) -#13573 := (iff #3310 #13570) -#13566 := (implies #3050 #13561) -#13571 := (iff #13566 #13570) -#13572 := [rewrite]: #13571 -#13567 := (iff #3310 #13566) -#13564 := (iff #3309 #13561) -#13557 := (implies #3054 #13552) -#13562 := (iff #13557 #13561) -#13563 := [rewrite]: #13562 -#13558 := (iff #3309 #13557) -#13555 := (iff #3308 #13552) -#13548 := (implies #3058 #13543) -#13553 := (iff #13548 #13552) -#13554 := [rewrite]: #13553 -#13549 := (iff #3308 #13548) -#13546 := (iff #3307 #13543) -#13540 := (implies #3060 #13535) -#13544 := (iff #13540 #13543) -#13545 := [rewrite]: #13544 -#13541 := (iff #3307 #13540) -#13538 := (iff #3306 #13535) -#13531 := (implies #3061 #13526) -#13536 := (iff #13531 #13535) -#13537 := [rewrite]: #13536 -#13532 := (iff #3306 #13531) -#13529 := (iff #3305 #13526) -#13522 := (implies #12423 #13517) -#13527 := (iff #13522 #13526) -#13528 := [rewrite]: #13527 -#13523 := (iff #3305 #13522) -#13520 := (iff #3304 #13517) -#13513 := (implies #12432 #13508) -#13518 := (iff #13513 #13517) -#13519 := [rewrite]: #13518 -#13514 := (iff #3304 #13513) -#13511 := (iff #3303 #13508) -#13505 := (implies #3060 #13432) -#13509 := (iff #13505 #13508) -#13510 := [rewrite]: #13509 -#13506 := (iff #3303 #13505) -#13503 := (iff #3302 #13432) -#13498 := (and #13432 true) -#13501 := (iff #13498 #13432) -#13502 := [rewrite]: #13501 -#13499 := (iff #3302 #13498) -#13496 := (iff #3301 true) -#13491 := (implies true true) -#13494 := (iff #13491 true) -#13495 := [rewrite]: #13494 -#13492 := (iff #3301 #13491) -#13489 := (iff #3300 true) -#13484 := (implies #3060 true) -#13487 := (iff #13484 true) -#13488 := [rewrite]: #13487 -#13485 := (iff #3300 #13484) -#13482 := (iff #3299 true) -#13449 := (or #13376 #12751) -#13457 := (or #12687 #13449) -#13472 := (or #12687 #13457) -#13477 := (implies false #13472) -#13480 := (iff #13477 true) -#13481 := [rewrite]: #13480 -#13478 := (iff #3299 #13477) -#13475 := (iff #3298 #13472) -#13469 := (implies #3060 #13457) -#13473 := (iff #13469 #13472) -#13474 := [rewrite]: #13473 -#13470 := (iff #3298 #13469) -#13467 := (iff #3297 #13457) -#13462 := (implies true #13457) -#13465 := (iff #13462 #13457) -#13466 := [rewrite]: #13465 -#13463 := (iff #3297 #13462) -#13460 := (iff #3296 #13457) -#13454 := (implies #3060 #13449) -#13458 := (iff #13454 #13457) -#13459 := [rewrite]: #13458 -#13455 := (iff #3296 #13454) -#13452 := (iff #3295 #13449) -#13446 := (implies #12556 #12751) -#13450 := (iff #13446 #13449) -#13451 := [rewrite]: #13450 -#13447 := (iff #3295 #13446) -#12754 := (iff #3167 #12751) -#12748 := (implies #3060 #12736) -#12752 := (iff #12748 #12751) -#12753 := [rewrite]: #12752 -#12749 := (iff #3167 #12748) -#12746 := (iff #3166 #12736) -#12741 := (implies true #12736) -#12744 := (iff #12741 #12736) -#12745 := [rewrite]: #12744 -#12742 := (iff #3166 #12741) -#12739 := (iff #3165 #12736) -#12733 := (implies #3060 #12728) -#12737 := (iff #12733 #12736) -#12738 := [rewrite]: #12737 -#12734 := (iff #3165 #12733) -#12731 := (iff #3164 #12728) -#12725 := (and #12720 up_216) -#12729 := (iff #12725 #12728) -#12730 := [rewrite]: #12729 -#12726 := (iff #3164 #12725) -#12723 := (iff #3163 #12720) -#12716 := (implies up_216 #12711) -#12721 := (iff #12716 #12720) -#12722 := [rewrite]: #12721 -#12717 := (iff #3163 #12716) -#12714 := (iff #3162 #12711) -#12708 := (implies #3060 #12696) -#12712 := (iff #12708 #12711) -#12713 := [rewrite]: #12712 -#12709 := (iff #3162 #12708) -#12706 := (iff #3161 #12696) -#12701 := (implies true #12696) -#12704 := (iff #12701 #12696) -#12705 := [rewrite]: #12704 -#12702 := (iff #3161 #12701) -#12699 := (iff #3160 #12696) -#12693 := (implies #3060 #12688) -#12697 := (iff #12693 #12696) -#12698 := [rewrite]: #12697 -#12694 := (iff #3160 #12693) -#12691 := (iff #3159 #12688) -#12684 := (implies #3060 #12672) -#12689 := (iff #12684 #12688) -#12690 := [rewrite]: #12689 -#12685 := (iff #3159 #12684) -#12682 := (iff #3158 #12672) -#12677 := (implies true #12672) -#12680 := (iff #12677 #12672) -#12681 := [rewrite]: #12680 -#12678 := (iff #3158 #12677) -#12675 := (iff #3157 #12672) -#12668 := (implies #12567 #12663) -#12673 := (iff #12668 #12672) -#12674 := [rewrite]: #12673 -#12669 := (iff #3157 #12668) -#12666 := (iff #3156 #12663) -#12659 := (implies #12570 #12654) -#12664 := (iff #12659 #12663) -#12665 := [rewrite]: #12664 -#12660 := (iff #3156 #12659) -#12657 := (iff #3155 #12654) -#12650 := (implies #12573 #12645) -#12655 := (iff #12650 #12654) -#12656 := [rewrite]: #12655 -#12651 := (iff #3155 #12650) -#12648 := (iff #3154 #12645) -#12641 := (implies #12576 #12629) -#12646 := (iff #12641 #12645) -#12647 := [rewrite]: #12646 -#12642 := (iff #3154 #12641) -#12639 := (iff #3153 #12629) -#12634 := (implies true #12629) -#12637 := (iff #12634 #12629) -#12638 := [rewrite]: #12637 -#12635 := (iff #3153 #12634) -#12632 := (iff #3152 #12629) -#12626 := (and #12621 #12591) -#12630 := (iff #12626 #12629) -#12631 := [rewrite]: #12630 -#12627 := (iff #3152 #12626) -#12592 := (iff #3144 #12591) -#12589 := (iff #3143 #12586) -#12583 := (implies #5718 #12580) -#12587 := (iff #12583 #12586) -#12588 := [rewrite]: #12587 -#12584 := (iff #3143 #12583) -#12581 := (iff #3142 #12580) -#12582 := [rewrite]: #12581 -#12585 := [monotonicity #5720 #12582]: #12584 -#12590 := [trans #12585 #12588]: #12589 -#12593 := [quant-intro #12590]: #12592 -#12624 := (iff #3151 #12621) -#12617 := (implies #12591 #12600) -#12622 := (iff #12617 #12621) -#12623 := [rewrite]: #12622 -#12618 := (iff #3151 #12617) -#12615 := (iff #3150 #12600) -#12610 := (and true #12600) -#12613 := (iff #12610 #12600) -#12614 := [rewrite]: #12613 -#12611 := (iff #3150 #12610) -#12601 := (iff #3148 #12600) -#12598 := (iff #3147 #12597) -#12595 := (iff #3146 #12594) -#12596 := [monotonicity #5720]: #12595 -#12599 := [monotonicity #12596]: #12598 -#12602 := [quant-intro #12599]: #12601 -#12608 := (iff #3149 true) -#12603 := (implies #12600 true) -#12606 := (iff #12603 true) -#12607 := [rewrite]: #12606 -#12604 := (iff #3149 #12603) -#12605 := [monotonicity #12602]: #12604 -#12609 := [trans #12605 #12607]: #12608 -#12612 := [monotonicity #12609 #12602]: #12611 -#12616 := [trans #12612 #12614]: #12615 -#12619 := [monotonicity #12593 #12616]: #12618 -#12625 := [trans #12619 #12623]: #12624 -#12628 := [monotonicity #12625 #12593]: #12627 -#12633 := [trans #12628 #12631]: #12632 -#12636 := [monotonicity #12633]: #12635 -#12640 := [trans #12636 #12638]: #12639 -#12577 := (iff #3139 #12576) -#12578 := [rewrite]: #12577 -#12643 := [monotonicity #12578 #12640]: #12642 -#12649 := [trans #12643 #12647]: #12648 -#12574 := (iff #3137 #12573) -#12575 := [rewrite]: #12574 -#12652 := [monotonicity #12575 #12649]: #12651 -#12658 := [trans #12652 #12656]: #12657 -#12571 := (iff #3135 #12570) -#12572 := [rewrite]: #12571 -#12661 := [monotonicity #12572 #12658]: #12660 -#12667 := [trans #12661 #12665]: #12666 -#12568 := (iff #3133 #12567) -#12569 := [rewrite]: #12568 -#12670 := [monotonicity #12569 #12667]: #12669 -#12676 := [trans #12670 #12674]: #12675 -#12679 := [monotonicity #12676]: #12678 -#12683 := [trans #12679 #12681]: #12682 -#12686 := [monotonicity #12683]: #12685 -#12692 := [trans #12686 #12690]: #12691 -#12695 := [monotonicity #12692]: #12694 -#12700 := [trans #12695 #12698]: #12699 -#12703 := [monotonicity #12700]: #12702 -#12707 := [trans #12703 #12705]: #12706 -#12710 := [monotonicity #12707]: #12709 -#12715 := [trans #12710 #12713]: #12714 -#12718 := [monotonicity #12715]: #12717 -#12724 := [trans #12718 #12722]: #12723 -#12727 := [monotonicity #12724]: #12726 -#12732 := [trans #12727 #12730]: #12731 -#12735 := [monotonicity #12732]: #12734 -#12740 := [trans #12735 #12738]: #12739 -#12743 := [monotonicity #12740]: #12742 -#12747 := [trans #12743 #12745]: #12746 -#12750 := [monotonicity #12747]: #12749 -#12755 := [trans #12750 #12753]: #12754 -#12557 := (iff #3120 #12556) -#12554 := (iff #3119 #12553) -#12555 := [rewrite]: #12554 -#12333 := (iff #2979 #12332) -#12334 := [rewrite]: #12333 -#12558 := [monotonicity #12334 #12555]: #12557 -#13448 := [monotonicity #12558 #12755]: #13447 -#13453 := [trans #13448 #13451]: #13452 -#13456 := [monotonicity #13453]: #13455 -#13461 := [trans #13456 #13459]: #13460 -#13464 := [monotonicity #13461]: #13463 -#13468 := [trans #13464 #13466]: #13467 -#13471 := [monotonicity #13468]: #13470 -#13476 := [trans #13471 #13474]: #13475 -#13479 := [monotonicity #13445 #13476]: #13478 -#13483 := [trans #13479 #13481]: #13482 -#13486 := [monotonicity #13483]: #13485 -#13490 := [trans #13486 #13488]: #13489 -#13493 := [monotonicity #13490]: #13492 -#13497 := [trans #13493 #13495]: #13496 -#13442 := (iff #3293 #13432) -#13437 := (implies true #13432) -#13440 := (iff #13437 #13432) -#13441 := [rewrite]: #13440 -#13438 := (iff #3293 #13437) -#13435 := (iff #3292 #13432) -#13429 := (implies #3060 #13417) -#13433 := (iff #13429 #13432) -#13434 := [rewrite]: #13433 -#13430 := (iff #3292 #13429) -#13427 := (iff #3291 #13417) -#13422 := (implies true #13417) -#13425 := (iff #13422 #13417) -#13426 := [rewrite]: #13425 -#13423 := (iff #3291 #13422) -#13420 := (iff #3290 #13417) -#13414 := (implies #3060 #13402) -#13418 := (iff #13414 #13417) -#13419 := [rewrite]: #13418 -#13415 := (iff #3290 #13414) -#13412 := (iff #3289 #13402) -#13407 := (implies true #13402) -#13410 := (iff #13407 #13402) -#13411 := [rewrite]: #13410 -#13408 := (iff #3289 #13407) -#13405 := (iff #3288 #13402) -#13399 := (implies #3060 #13394) -#13403 := (iff #13399 #13402) -#13404 := [rewrite]: #13403 -#13400 := (iff #3288 #13399) -#13397 := (iff #3287 #13394) -#13391 := (implies #12544 #13386) -#13395 := (iff #13391 #13394) -#13396 := [rewrite]: #13395 -#13392 := (iff #3287 #13391) -#13389 := (iff #3286 #13386) -#13382 := (implies #12544 #13377) -#13387 := (iff #13382 #13386) -#13388 := [rewrite]: #13387 -#13383 := (iff #3286 #13382) -#13380 := (iff #3285 #13377) -#13373 := (implies #12556 #13368) -#13378 := (iff #13373 #13377) +#13610 := (iff #12623 #13609) +#13607 := (iff #3190 #13603) +#13608 := [rewrite]: #13607 +#13611 := [monotonicity #13608]: #13610 +#13615 := [trans #13611 #13613]: #13614 +#13624 := [monotonicity #13615 #13621]: #13623 +#13627 := [monotonicity #13419 #13624]: #13626 +#13632 := [trans #13627 #13630]: #13631 +#13665 := [monotonicity #13632 #13662]: #13664 +#13668 := [monotonicity #13419 #13665]: #13667 +#13671 := [monotonicity #13668]: #13670 +#13676 := [trans #13671 #13674]: #13675 +#13679 := [monotonicity #13676]: #13678 +#13684 := [trans #13679 #13682]: #13683 +#13687 := [monotonicity #13684]: #13686 +#13690 := [monotonicity #13687]: #13689 +#13695 := [trans #13690 #13693]: #13694 +#13710 := [monotonicity #13419 #13695]: #13709 +#13713 := [monotonicity #13419 #13710]: #13712 +#13706 := (iff #12789 #13697) +#13701 := (not #13698) +#13704 := (iff #13701 #13697) +#13705 := [rewrite]: #13704 +#13702 := (iff #12789 #13701) +#13699 := (iff #3179 #13698) +#13700 := [rewrite]: #13699 +#13703 := [monotonicity #13700]: #13702 +#13707 := [trans #13703 #13705]: #13706 +#13716 := [monotonicity #13707 #13713]: #13715 +#13719 := [monotonicity #13419 #13716]: #13718 +#13724 := [trans #13719 #13722]: #13723 +#13848 := [monotonicity #13724 #13845]: #13847 +#13978 := [monotonicity #13419 #13848]: #13977 +#13981 := [monotonicity #13978]: #13980 +#13984 := [monotonicity #13981]: #13983 +#13987 := [monotonicity #13984]: #13986 +#13990 := [monotonicity #13987]: #13989 +#13993 := [monotonicity #13990]: #13992 +#13996 := [monotonicity #13993]: #13995 +#13871 := (iff #12931 #13870) +#13868 := (iff #12312 #12306) +#13860 := (and true #12306) +#13863 := (and true #13860) +#13866 := (iff #13863 #12306) +#13867 := [rewrite]: #13866 +#13864 := (iff #12312 #13863) +#13861 := (iff #12309 #13860) +#13856 := (iff #3160 true) +#13851 := (forall (vars (?x784 T5)) (:pat #3158) true) +#13854 := (iff #13851 true) +#13855 := [elim-unused]: #13854 +#13852 := (iff #3160 #13851) +#13849 := (iff #3159 true) +#13850 := [rewrite]: #13849 +#13853 := [quant-intro #13850]: #13852 +#13857 := [trans #13853 #13855]: #13856 +#13862 := [monotonicity #13857]: #13861 +#13858 := (iff #3156 true) +#13859 := [rewrite]: #13858 +#13865 := [monotonicity #13859 #13862]: #13864 +#13869 := [trans #13865 #13867]: #13868 +#13872 := [monotonicity #13869]: #13871 +#13999 := [monotonicity #13872 #13996]: #13998 +#14002 := [monotonicity #13872 #13999]: #14001 +#14005 := [monotonicity #13419 #14002]: #14004 +#14008 := [monotonicity #13419 #14005]: #14007 +#14011 := [monotonicity #13419 #14008]: #14010 +#14014 := [monotonicity #13419 #14011]: #14013 +#13888 := (iff #13008 #13887) +#13885 := (iff #11995 #13882) +#13879 := (and #13876 #11992) +#13883 := (iff #13879 #13882) +#13884 := [rewrite]: #13883 +#13880 := (iff #11995 #13879) +#13877 := (iff #3072 #13876) +#13878 := [rewrite]: #13877 +#13881 := [monotonicity #13878]: #13880 +#13886 := [trans #13881 #13884]: #13885 +#13889 := [monotonicity #13886]: #13888 +#14017 := [monotonicity #13889 #14014]: #14016 +#13922 := (iff #13017 #13921) +#13919 := (iff #11989 #13918) +#13916 := (iff #11984 #13913) +#13907 := (or #13890 #13904) +#13910 := (or #5601 #13907) +#13914 := (iff #13910 #13913) +#13915 := [rewrite]: #13914 +#13911 := (iff #11984 #13910) +#13908 := (iff #11978 #13907) +#13905 := (iff #3068 #13904) +#13906 := [rewrite]: #13905 +#13900 := (iff #11977 #13890) +#13892 := (not #13890) +#13895 := (not #13892) +#13898 := (iff #13895 #13890) +#13899 := [rewrite]: #13898 +#13896 := (iff #11977 #13895) +#13893 := (iff #3067 #13892) +#13894 := [rewrite]: #13893 +#13897 := [monotonicity #13894]: #13896 +#13901 := [trans #13897 #13899]: #13900 +#13909 := [monotonicity #13901 #13906]: #13908 +#13912 := [monotonicity #5603 #13909]: #13911 +#13917 := [trans #13912 #13915]: #13916 +#13920 := [quant-intro #13917]: #13919 +#13923 := [monotonicity #13920]: #13922 +#14020 := [monotonicity #13923 #14017]: #14019 +#13928 := (iff #13026 #13927) +#13925 := (iff #3066 #13924) +#13926 := [rewrite]: #13925 +#13929 := [monotonicity #13926]: #13928 +#14023 := [monotonicity #13929 #14020]: #14022 +#14026 := [monotonicity #13419 #14023]: #14025 +#13940 := (iff #13043 #13939) +#13937 := (iff #3063 #13936) +#13934 := (iff #3062 #13933) +#13935 := [rewrite]: #13934 +#13931 := (iff #3061 #13930) +#13932 := [rewrite]: #13931 +#13938 := [monotonicity #13932 #13935]: #13937 +#13941 := [monotonicity #13938]: #13940 +#14029 := [monotonicity #13941 #14026]: #14028 +#13950 := (iff #13052 #13949) +#13947 := (iff #3059 #13946) +#13944 := (iff #3058 #13942) +#13945 := [rewrite]: #13944 +#13948 := [monotonicity #13410 #13945]: #13947 +#13951 := [monotonicity #13948]: #13950 +#14032 := [monotonicity #13951 #14029]: #14031 +#13963 := (iff #13061 #13962) +#13960 := (iff #3055 #13959) +#13957 := (iff #3054 #13955) +#13958 := [rewrite]: #13957 +#13953 := (iff #3053 #13952) +#13954 := [rewrite]: #13953 +#13961 := [monotonicity #13954 #13958]: #13960 +#13964 := [monotonicity #13961]: #13963 +#14035 := [monotonicity #13964 #14032]: #14034 +#13974 := (iff #13077 #13973) +#13971 := (iff #11971 #13968) +#13969 := (iff #13965 #13968) +#13970 := [rewrite]: #13969 +#13966 := (iff #11971 #13965) +#13406 := (iff #2956 #13405) +#13407 := [rewrite]: #13406 +#13967 := [monotonicity #13407]: #13966 +#13972 := [trans #13967 #13970]: #13971 +#13975 := [monotonicity #13972]: #13974 +#14038 := [monotonicity #13975 #14035]: #14037 +#14043 := [trans #14038 #14041]: #14042 +#14046 := [monotonicity #13967 #14043]: #14045 +#14051 := [trans #14046 #14049]: #14050 +#13402 := (iff #13089 #13401) +#13399 := (iff #11966 #13398) +#13396 := (iff #11961 #13393) +#13387 := (or #13371 #13383) +#13390 := (or #5601 #13387) +#13394 := (iff #13390 #13393) +#13395 := [rewrite]: #13394 +#13391 := (iff #11961 #13390) +#13388 := (iff #11955 #13387) +#13382 := (iff #3046 #13383) +#13386 := [rewrite]: #13382 +#13380 := (iff #11954 #13371) +#13372 := (not #13371) +#13375 := (not #13372) +#13378 := (iff #13375 #13371) #13379 := [rewrite]: #13378 -#13374 := (iff #3285 #13373) -#13371 := (iff #3284 #13368) -#13364 := (implies #3121 #13359) -#13369 := (iff #13364 #13368) -#13370 := [rewrite]: #13369 -#13365 := (iff #3284 #13364) -#13362 := (iff #3283 #13359) -#13355 := (implies #3122 #13350) -#13360 := (iff #13355 #13359) -#13361 := [rewrite]: #13360 -#13356 := (iff #3283 #13355) -#13353 := (iff #3282 #13350) -#13346 := (implies #3123 #13341) -#13351 := (iff #13346 #13350) -#13352 := [rewrite]: #13351 -#13347 := (iff #3282 #13346) -#13344 := (iff #3281 #13341) -#13337 := (implies #3124 #13332) -#13342 := (iff #13337 #13341) -#13343 := [rewrite]: #13342 -#13338 := (iff #3281 #13337) -#13335 := (iff #3280 #13332) -#13328 := (implies #3127 #13316) -#13333 := (iff #13328 #13332) -#13334 := [rewrite]: #13333 -#13329 := (iff #3280 #13328) -#13326 := (iff #3279 #13316) -#13321 := (implies true #13316) -#13324 := (iff #13321 #13316) -#13325 := [rewrite]: #13324 -#13322 := (iff #3279 #13321) -#13319 := (iff #3278 #13316) -#13313 := (implies #3060 #13310) -#13317 := (iff #13313 #13316) -#13318 := [rewrite]: #13317 -#13314 := (iff #3278 #13313) -#13311 := (iff #3277 #13310) -#13308 := (iff #3276 #13298) -#13303 := (implies true #13298) -#13306 := (iff #13303 #13298) -#13307 := [rewrite]: #13306 -#13304 := (iff #3276 #13303) -#13301 := (iff #3275 #13298) -#13295 := (implies #3060 #13290) -#13299 := (iff #13295 #13298) -#13300 := [rewrite]: #13299 -#13296 := (iff #3275 #13295) -#13293 := (iff #3274 #13290) -#13286 := (implies #3174 #13281) -#13291 := (iff #13286 #13290) -#13292 := [rewrite]: #13291 -#13287 := (iff #3274 #13286) -#13284 := (iff #3273 #13281) -#13278 := (implies #3060 #13266) -#13282 := (iff #13278 #13281) -#13283 := [rewrite]: #13282 -#13279 := (iff #3273 #13278) -#13276 := (iff #3272 #13266) -#13271 := (implies true #13266) -#13274 := (iff #13271 #13266) -#13275 := [rewrite]: #13274 -#13272 := (iff #3272 #13271) -#13269 := (iff #3271 #13266) -#13263 := (implies #3060 #13258) -#13267 := (iff #13263 #13266) -#13268 := [rewrite]: #13267 -#13264 := (iff #3271 #13263) -#13261 := (iff #3270 #13258) -#13255 := (and #13250 #12809) -#13259 := (iff #13255 #13258) +#13376 := (iff #11954 #13375) +#13373 := (iff #3043 #13372) +#13374 := [rewrite]: #13373 +#13377 := [monotonicity #13374]: #13376 +#13381 := [trans #13377 #13379]: #13380 +#13389 := [monotonicity #13381 #13386]: #13388 +#13392 := [monotonicity #5603 #13389]: #13391 +#13397 := [trans #13392 #13395]: #13396 +#13400 := [quant-intro #13397]: #13399 +#13403 := [monotonicity #13400]: #13402 +#14054 := [monotonicity #13403 #14051]: #14053 +#14057 := [monotonicity #13400 #14054]: #14056 +#13369 := (iff #13101 #13368) +#13366 := (iff #3042 #13365) +#13367 := [rewrite]: #13366 +#13370 := [monotonicity #13367]: #13369 +#14060 := [monotonicity #13370 #14057]: #14059 +#14063 := [monotonicity #13367 #14060]: #14062 +#13363 := (iff #13113 false) +#11998 := (iff #3077 false) +#11999 := [rewrite]: #11998 +#13361 := (iff #13113 #3077) +#13359 := (iff #11951 true) +#12338 := (and true true) +#13354 := (and true #12338) +#13357 := (iff #13354 true) +#13358 := [rewrite]: #13357 +#13355 := (iff #11951 #13354) +#13352 := (iff #11948 #12338) +#13350 := (iff #3038 true) +#13351 := [rewrite]: #13350 +#13348 := (iff #3037 true) +#13349 := [rewrite]: #13348 +#13353 := [monotonicity #13349 #13351]: #13352 +#13356 := [monotonicity #13349 #13353]: #13355 +#13360 := [trans #13356 #13358]: #13359 +#13362 := [monotonicity #13360]: #13361 +#13364 := [trans #13362 #11999]: #13363 +#14066 := [monotonicity #13364 #14063]: #14065 +#14069 := [monotonicity #14066]: #14068 +#14072 := [monotonicity #14069]: #14071 +#14075 := [monotonicity #14072]: #14074 +#14078 := [monotonicity #14075]: #14077 +#14081 := [monotonicity #14078]: #14080 +#14086 := [trans #14081 #14084]: #14085 +#14089 := [monotonicity #14086]: #14088 +#14094 := [trans #14089 #14092]: #14093 +#14097 := [monotonicity #14094]: #14096 +#14100 := [monotonicity #14097]: #14099 +#14105 := [trans #14100 #14103]: #14104 +#14108 := [monotonicity #14105]: #14107 +#14111 := [monotonicity #14108]: #14110 +#13346 := (iff #13194 #13345) +#13343 := (iff #3012 #13342) +#13340 := (iff #3011 #13337) +#13341 := [rewrite]: #13340 +#13334 := (iff #3010 #13335) +#13336 := [rewrite]: #13334 +#13344 := [monotonicity #13336 #13341]: #13343 +#13347 := [monotonicity #13344]: #13346 +#14192 := [monotonicity #13347 #14111]: #14191 +#14195 := [monotonicity #14192]: #14194 +#14198 := [monotonicity #14195]: #14197 +#14201 := [monotonicity #14198]: #14200 +#14204 := [monotonicity #14201]: #14203 +#14122 := (iff #13239 #14121) +#14119 := (iff #2994 #14118) +#14116 := (iff #2993 #14112) +#14117 := [rewrite]: #14116 +#14120 := [quant-intro #14117]: #14119 +#14123 := [monotonicity #14120]: #14122 +#14207 := [monotonicity #14123 #14204]: #14206 +#14210 := [monotonicity #14207]: #14209 +#14213 := [monotonicity #14210]: #14212 +#14128 := (iff #13273 #14127) +#14125 := (iff #11899 #14124) +#14126 := [rewrite]: #14125 +#14129 := [monotonicity #14126]: #14128 +#14216 := [monotonicity #14129 #14213]: #14215 +#14135 := (iff #13282 #13404) +#14130 := (not #13405) +#14133 := (iff #14130 #13404) +#14134 := [rewrite]: #14133 +#14131 := (iff #13282 #14130) +#14132 := [monotonicity #13407]: #14131 +#14136 := [trans #14132 #14134]: #14135 +#14219 := [monotonicity #14136 #14216]: #14218 +#14146 := (iff #13291 #14137) +#14138 := (not #14137) +#14141 := (not #14138) +#14144 := (iff #14141 #14137) +#14145 := [rewrite]: #14144 +#14142 := (iff #13291 #14141) +#14139 := (iff #2955 #14138) +#14140 := [rewrite]: #14139 +#14143 := [monotonicity #14140]: #14142 +#14147 := [trans #14143 #14145]: #14146 +#14222 := [monotonicity #14147 #14219]: #14221 +#14160 := (iff #13300 #14159) +#14157 := (iff #2953 #14156) +#14154 := (iff #2952 #14151) +#14155 := [rewrite]: #14154 +#14148 := (iff #2951 #14149) +#14150 := [rewrite]: #14148 +#14158 := [monotonicity #14150 #14155]: #14157 +#14161 := [monotonicity #14158]: #14160 +#14225 := [monotonicity #14161 #14222]: #14224 +#14174 := (iff #13309 #14173) +#14171 := (iff #2949 #14170) +#14168 := (iff #2948 #14165) +#14169 := [rewrite]: #14168 +#14162 := (iff #2947 #14163) +#14164 := [rewrite]: #14162 +#14172 := [monotonicity #14164 #14169]: #14171 +#14175 := [monotonicity #14172]: #14174 +#14228 := [monotonicity #14175 #14225]: #14227 +#14188 := (iff #13318 #14187) +#14185 := (iff #2945 #14184) +#14182 := (iff #2944 #14179) +#14183 := [rewrite]: #14182 +#14176 := (iff #2943 #14177) +#14178 := [rewrite]: #14176 +#14186 := [monotonicity #14178 #14183]: #14185 +#14189 := [monotonicity #14186]: #14188 +#14231 := [monotonicity #14189 #14228]: #14230 +#14236 := [trans #14231 #14234]: #14235 +#14239 := [monotonicity #14236]: #14238 +#13332 := (iff #3349 #13331) +#13329 := (iff #3348 #13319) +#13324 := (implies true #13319) +#13327 := (iff #13324 #13319) +#13328 := [rewrite]: #13327 +#13325 := (iff #3348 #13324) +#13322 := (iff #3347 #13319) +#13315 := (implies #2945 #13310) +#13320 := (iff #13315 #13319) +#13321 := [rewrite]: #13320 +#13316 := (iff #3347 #13315) +#13313 := (iff #3346 #13310) +#13306 := (implies #2949 #13301) +#13311 := (iff #13306 #13310) +#13312 := [rewrite]: #13311 +#13307 := (iff #3346 #13306) +#13304 := (iff #3345 #13301) +#13297 := (implies #2953 #13292) +#13302 := (iff #13297 #13301) +#13303 := [rewrite]: #13302 +#13298 := (iff #3345 #13297) +#13295 := (iff #3344 #13292) +#13288 := (implies #2955 #13283) +#13293 := (iff #13288 #13292) +#13294 := [rewrite]: #13293 +#13289 := (iff #3344 #13288) +#13286 := (iff #3343 #13283) +#13279 := (implies #2956 #13274) +#13284 := (iff #13279 #13283) +#13285 := [rewrite]: #13284 +#13280 := (iff #3343 #13279) +#13277 := (iff #3342 #13274) +#13270 := (implies #11899 #13258) +#13275 := (iff #13270 #13274) +#13276 := [rewrite]: #13275 +#13271 := (iff #3342 #13270) +#13268 := (iff #3341 #13258) +#13263 := (implies true #13258) +#13266 := (iff #13263 #13258) +#13267 := [rewrite]: #13266 +#13264 := (iff #3341 #13263) +#13261 := (iff #3340 #13258) +#13254 := (implies #11902 #13249) +#13259 := (iff #13254 #13258) #13260 := [rewrite]: #13259 -#13256 := (iff #3270 #13255) -#12810 := (iff #3180 #12809) -#12807 := (iff #3179 #12806) -#12808 := [rewrite]: #12807 -#12804 := (iff #3177 #12803) -#12805 := [rewrite]: #12804 -#12811 := [monotonicity #12805 #12808]: #12810 -#13253 := (iff #3269 #13250) -#13247 := (implies #12809 #13242) -#13251 := (iff #13247 #13250) -#13252 := [rewrite]: #13251 -#13248 := (iff #3269 #13247) -#13245 := (iff #3268 #13242) -#13239 := (and #13234 #12818) -#13243 := (iff #13239 #13242) -#13244 := [rewrite]: #13243 -#13240 := (iff #3268 #13239) -#12821 := (iff #3183 #12818) -#12815 := (and #12812 #12806) -#12819 := (iff #12815 #12818) -#12820 := [rewrite]: #12819 -#12816 := (iff #3183 #12815) -#12813 := (iff #3182 #12812) -#12814 := [rewrite]: #12813 -#12817 := [monotonicity #12814 #12808]: #12816 -#12822 := [trans #12817 #12820]: #12821 -#13237 := (iff #3267 #13234) -#13231 := (implies #12818 #13226) -#13235 := (iff #13231 #13234) -#13236 := [rewrite]: #13235 -#13232 := (iff #3267 #13231) -#13229 := (iff #3266 #13226) -#13223 := (implies #3060 #13220) -#13227 := (iff #13223 #13226) -#13228 := [rewrite]: #13227 -#13224 := (iff #3266 #13223) -#13221 := (iff #3265 #13220) -#13218 := (iff #3264 #13208) -#13213 := (implies true #13208) -#13216 := (iff #13213 #13208) -#13217 := [rewrite]: #13216 -#13214 := (iff #3264 #13213) -#13211 := (iff #3263 #13208) -#13205 := (implies #3060 #13200) -#13209 := (iff #13205 #13208) -#13210 := [rewrite]: #13209 -#13206 := (iff #3263 #13205) -#13203 := (iff #3262 #13200) -#13196 := (implies #3238 #13191) -#13201 := (iff #13196 #13200) -#13202 := [rewrite]: #13201 -#13197 := (iff #3262 #13196) -#13194 := (iff #3261 #13191) -#13188 := (implies #3060 #13176) -#13192 := (iff #13188 #13191) -#13193 := [rewrite]: #13192 -#13189 := (iff #3261 #13188) -#13186 := (iff #3260 #13176) -#13181 := (implies true #13176) -#13184 := (iff #13181 #13176) +#13255 := (iff #3340 #13254) +#13252 := (iff #3339 #13249) +#13245 := (implies #11911 #13240) +#13250 := (iff #13245 #13249) +#13251 := [rewrite]: #13250 +#13246 := (iff #3339 #13245) +#13243 := (iff #3338 #13240) +#13236 := (implies #2994 #13231) +#13241 := (iff #13236 #13240) +#13242 := [rewrite]: #13241 +#13237 := (iff #3338 #13236) +#13234 := (iff #3337 #13231) +#13227 := (implies #2999 #13222) +#13232 := (iff #13227 #13231) +#13233 := [rewrite]: #13232 +#13228 := (iff #3337 #13227) +#13225 := (iff #3336 #13222) +#13218 := (implies #3001 #13213) +#13223 := (iff #13218 #13222) +#13224 := [rewrite]: #13223 +#13219 := (iff #3336 #13218) +#13216 := (iff #3335 #13213) +#13209 := (implies #3004 #13204) +#13214 := (iff #13209 #13213) +#13215 := [rewrite]: #13214 +#13210 := (iff #3335 #13209) +#13207 := (iff #3334 #13204) +#13200 := (implies #11925 #13195) +#13205 := (iff #13200 #13204) +#13206 := [rewrite]: #13205 +#13201 := (iff #3334 #13200) +#13198 := (iff #3333 #13195) +#13191 := (implies #3012 #13188) +#13196 := (iff #13191 #13195) +#13197 := [rewrite]: #13196 +#13192 := (iff #3333 #13191) +#13189 := (iff #3332 #13188) +#13186 := (iff #3331 #13183) +#13179 := (implies #11928 #13176) +#13184 := (iff #13179 #13183) #13185 := [rewrite]: #13184 -#13182 := (iff #3260 #13181) -#13179 := (iff #3259 #13176) -#13173 := (implies #3060 #13168) -#13177 := (iff #13173 #13176) -#13178 := [rewrite]: #13177 -#13174 := (iff #3259 #13173) -#13171 := (iff #3258 #13168) -#13165 := (and #13160 #12809) -#13169 := (iff #13165 #13168) -#13170 := [rewrite]: #13169 -#13166 := (iff #3258 #13165) -#13163 := (iff #3257 #13160) -#13156 := (implies #12809 #13151) -#13161 := (iff #13156 #13160) -#13162 := [rewrite]: #13161 -#13157 := (iff #3257 #13156) -#13154 := (iff #3256 #13151) -#13148 := (and #13143 #12818) -#13152 := (iff #13148 #13151) -#13153 := [rewrite]: #13152 -#13149 := (iff #3256 #13148) -#13146 := (iff #3255 #13143) -#13139 := (implies #12818 #13134) -#13144 := (iff #13139 #13143) -#13145 := [rewrite]: #13144 -#13140 := (iff #3255 #13139) -#13137 := (iff #3254 #13134) -#13130 := (implies #13070 #13125) -#13135 := (iff #13130 #13134) -#13136 := [rewrite]: #13135 -#13131 := (iff #3254 #13130) -#13128 := (iff #3253 #13125) -#13121 := (implies #3242 #13116) -#13126 := (iff #13121 #13125) -#13127 := [rewrite]: #13126 -#13122 := (iff #3253 #13121) -#13119 := (iff #3252 #13116) -#13112 := (implies #3244 #13107) -#13117 := (iff #13112 #13116) -#13118 := [rewrite]: #13117 -#13113 := (iff #3252 #13112) -#13110 := (iff #3251 #13107) -#13103 := (implies #3059 #13091) -#13108 := (iff #13103 #13107) -#13109 := [rewrite]: #13108 -#13104 := (iff #3251 #13103) -#13101 := (iff #3250 #13091) -#13096 := (implies true #13091) -#13099 := (iff #13096 #13091) -#13100 := [rewrite]: #13099 -#13097 := (iff #3250 #13096) -#13094 := (iff #3249 #13091) -#13087 := (implies #3246 #13082) -#13092 := (iff #13087 #13091) -#13093 := [rewrite]: #13092 -#13088 := (iff #3249 #13087) -#13085 := (iff #3248 #13082) -#13078 := (implies #13075 #12978) -#13083 := (iff #13078 #13082) -#13084 := [rewrite]: #13083 -#13079 := (iff #3248 #13078) -#12988 := (iff #3227 #12978) +#13180 := (iff #3331 #13179) +#13177 := (iff #3330 #13176) +#13174 := (iff #3329 #13171) +#13167 := (implies #11937 #13164) +#13172 := (iff #13167 #13171) +#13173 := [rewrite]: #13172 +#13168 := (iff #3329 #13167) +#13165 := (iff #3328 #13164) +#13162 := (iff #3327 #13159) +#13155 := (implies #11943 #13150) +#13160 := (iff #13155 #13159) +#13161 := [rewrite]: #13160 +#13156 := (iff #3327 #13155) +#13153 := (iff #3326 #13150) +#13146 := (implies #3027 #13141) +#13151 := (iff #13146 #13150) +#13152 := [rewrite]: #13151 +#13147 := (iff #3326 #13146) +#13144 := (iff #3325 #13141) +#13137 := (implies #3030 #13132) +#13142 := (iff #13137 #13141) +#13143 := [rewrite]: #13142 +#13138 := (iff #3325 #13137) +#13135 := (iff #3324 #13132) +#13128 := (implies #3033 #13123) +#13133 := (iff #13128 #13132) +#13134 := [rewrite]: #13133 +#13129 := (iff #3324 #13128) +#13126 := (iff #3323 #13123) +#13119 := (implies #3036 #13114) +#13124 := (iff #13119 #13123) +#13125 := [rewrite]: #13124 +#13120 := (iff #3323 #13119) +#13117 := (iff #3322 #13114) +#13110 := (implies #11951 #13107) +#13115 := (iff #13110 #13114) +#13116 := [rewrite]: #13115 +#13111 := (iff #3322 #13110) +#13108 := (iff #3321 #13107) +#13105 := (iff #3320 #13102) +#13098 := (implies #3042 #13095) +#13103 := (iff #13098 #13102) +#13104 := [rewrite]: #13103 +#13099 := (iff #3320 #13098) +#13096 := (iff #3319 #13095) +#13093 := (iff #3318 #13090) +#13086 := (implies #11966 #13083) +#13091 := (iff #13086 #13090) +#13092 := [rewrite]: #13091 +#13087 := (iff #3318 #13086) +#13084 := (iff #3317 #13083) +#13081 := (iff #3316 #13078) +#13074 := (implies #11971 #13062) +#13079 := (iff #13074 #13078) +#13080 := [rewrite]: #13079 +#13075 := (iff #3316 #13074) +#13072 := (iff #3315 #13062) +#13067 := (implies true #13062) +#13070 := (iff #13067 #13062) +#13071 := [rewrite]: #13070 +#13068 := (iff #3315 #13067) +#13065 := (iff #3314 #13062) +#13058 := (implies #3055 #13053) +#13063 := (iff #13058 #13062) +#13064 := [rewrite]: #13063 +#13059 := (iff #3314 #13058) +#13056 := (iff #3313 #13053) +#13049 := (implies #3059 #13044) +#13054 := (iff #13049 #13053) +#13055 := [rewrite]: #13054 +#13050 := (iff #3313 #13049) +#13047 := (iff #3312 #13044) +#13040 := (implies #3063 #13035) +#13045 := (iff #13040 #13044) +#13046 := [rewrite]: #13045 +#13041 := (iff #3312 #13040) +#13038 := (iff #3311 #13035) +#13032 := (implies #11974 #13027) +#13036 := (iff #13032 #13035) +#13037 := [rewrite]: #13036 +#13033 := (iff #3311 #13032) +#13030 := (iff #3310 #13027) +#13023 := (implies #3066 #13018) +#13028 := (iff #13023 #13027) +#13029 := [rewrite]: #13028 +#13024 := (iff #3310 #13023) +#13021 := (iff #3309 #13018) +#13014 := (implies #11989 #13009) +#13019 := (iff #13014 #13018) +#13020 := [rewrite]: #13019 +#13015 := (iff #3309 #13014) +#13012 := (iff #3308 #13009) +#13005 := (implies #11995 #13000) +#13010 := (iff #13005 #13009) +#13011 := [rewrite]: #13010 +#13006 := (iff #3308 #13005) +#13003 := (iff #3307 #13000) +#12997 := (implies #11974 #12978) +#13001 := (iff #12997 #13000) +#13002 := [rewrite]: #13001 +#12998 := (iff #3307 #12997) +#12995 := (iff #3306 #12978) +#12990 := (and true #12978) +#12993 := (iff #12990 #12978) +#12994 := [rewrite]: #12993 +#12991 := (iff #3306 #12990) +#12988 := (iff #3305 #12978) #12983 := (implies true #12978) #12986 := (iff #12983 #12978) #12987 := [rewrite]: #12986 -#12984 := (iff #3227 #12983) -#12981 := (iff #3226 #12978) -#12974 := (implies #12829 #12969) -#12979 := (iff #12974 #12978) +#12984 := (iff #3305 #12983) +#12981 := (iff #3304 #12978) +#12975 := (implies #11974 #12963) +#12979 := (iff #12975 #12978) #12980 := [rewrite]: #12979 -#12975 := (iff #3226 #12974) -#12972 := (iff #3225 #12969) -#12966 := (and #12961 #12841) -#12970 := (iff #12966 #12969) -#12971 := [rewrite]: #12970 -#12967 := (iff #3225 #12966) -#12842 := (iff #3195 #12841) -#12839 := (iff #3194 #12838) -#12833 := (= #3192 #12832) -#12834 := [rewrite]: #12833 -#12840 := [monotonicity #12834]: #12839 -#12836 := (iff #3193 #12835) -#12837 := [monotonicity #12834]: #12836 -#12843 := [monotonicity #12837 #12840]: #12842 -#12964 := (iff #3224 #12961) -#12957 := (implies #12841 #12952) -#12962 := (iff #12957 #12961) -#12963 := [rewrite]: #12962 -#12958 := (iff #3224 #12957) -#12955 := (iff #3223 #12952) -#12948 := (implies #12844 #12943) -#12953 := (iff #12948 #12952) -#12954 := [rewrite]: #12953 -#12949 := (iff #3223 #12948) -#12946 := (iff #3222 #12943) -#12939 := (implies #3199 #12934) -#12944 := (iff #12939 #12943) -#12945 := [rewrite]: #12944 -#12940 := (iff #3222 #12939) -#12937 := (iff #3221 #12934) -#12930 := (implies #3201 #12918) -#12935 := (iff #12930 #12934) -#12936 := [rewrite]: #12935 -#12931 := (iff #3221 #12930) -#12928 := (iff #3220 #12918) -#12923 := (implies true #12918) -#12926 := (iff #12923 #12918) -#12927 := [rewrite]: #12926 -#12924 := (iff #3220 #12923) -#12921 := (iff #3219 #12918) -#12915 := (and #12910 #3202) -#12919 := (iff #12915 #12918) -#12920 := [rewrite]: #12919 -#12916 := (iff #3219 #12915) -#12913 := (iff #3218 #12910) -#12906 := (implies #3202 #12901) -#12911 := (iff #12906 #12910) -#12912 := [rewrite]: #12911 -#12907 := (iff #3218 #12906) -#12904 := (iff #3217 #12901) -#12898 := (and #12893 #12859) -#12902 := (iff #12898 #12901) -#12903 := [rewrite]: #12902 -#12899 := (iff #3217 #12898) -#12860 := (iff #3207 #12859) -#12857 := (iff #3206 #12854) -#12851 := (implies #5718 #12848) -#12855 := (iff #12851 #12854) -#12856 := [rewrite]: #12855 -#12852 := (iff #3206 #12851) -#12849 := (iff #3205 #12848) -#12850 := [rewrite]: #12849 -#12853 := [monotonicity #5720 #12850]: #12852 -#12858 := [trans #12853 #12856]: #12857 -#12861 := [quant-intro #12858]: #12860 -#12896 := (iff #3216 #12893) -#12889 := (implies #12859 #12868) -#12894 := (iff #12889 #12893) -#12895 := [rewrite]: #12894 -#12890 := (iff #3216 #12889) -#12887 := (iff #3215 #12868) -#12882 := (and true #12868) -#12885 := (iff #12882 #12868) -#12886 := [rewrite]: #12885 -#12883 := (iff #3215 #12882) -#12871 := (iff #3212 #12868) -#12865 := (and #12862 #3211) -#12869 := (iff #12865 #12868) -#12870 := [rewrite]: #12869 -#12866 := (iff #3212 #12865) -#12863 := (iff #3210 #12862) -#12864 := [rewrite]: #12863 -#12867 := [monotonicity #12864]: #12866 -#12872 := [trans #12867 #12870]: #12871 -#12880 := (iff #3214 true) -#12875 := (implies #12868 true) -#12878 := (iff #12875 true) -#12879 := [rewrite]: #12878 -#12876 := (iff #3214 #12875) -#12873 := (iff #3213 true) -#12874 := [rewrite]: #12873 -#12877 := [monotonicity #12872 #12874]: #12876 -#12881 := [trans #12877 #12879]: #12880 -#12884 := [monotonicity #12881 #12872]: #12883 -#12888 := [trans #12884 #12886]: #12887 -#12891 := [monotonicity #12861 #12888]: #12890 -#12897 := [trans #12891 #12895]: #12896 -#12900 := [monotonicity #12897 #12861]: #12899 -#12905 := [trans #12900 #12903]: #12904 -#12908 := [monotonicity #12905]: #12907 -#12914 := [trans #12908 #12912]: #12913 -#12917 := [monotonicity #12914]: #12916 -#12922 := [trans #12917 #12920]: #12921 -#12925 := [monotonicity #12922]: #12924 -#12929 := [trans #12925 #12927]: #12928 -#12932 := [monotonicity #12929]: #12931 -#12938 := [trans #12932 #12936]: #12937 -#12941 := [monotonicity #12938]: #12940 -#12947 := [trans #12941 #12945]: #12946 -#12845 := (iff #3197 #12844) -#12846 := [monotonicity #12834]: #12845 -#12950 := [monotonicity #12846 #12947]: #12949 -#12956 := [trans #12950 #12954]: #12955 -#12959 := [monotonicity #12843 #12956]: #12958 -#12965 := [trans #12959 #12963]: #12964 -#12968 := [monotonicity #12965 #12843]: #12967 -#12973 := [trans #12968 #12971]: #12972 -#12830 := (iff #3191 #12829) -#12831 := [rewrite]: #12830 -#12976 := [monotonicity #12831 #12973]: #12975 -#12982 := [trans #12976 #12980]: #12981 +#12976 := (iff #3304 #12975) +#12973 := (iff #3303 #12963) +#12968 := (implies true #12963) +#12971 := (iff #12968 #12963) +#12972 := [rewrite]: #12971 +#12969 := (iff #3303 #12968) +#12966 := (iff #3302 #12963) +#12960 := (implies #11974 #12948) +#12964 := (iff #12960 #12963) +#12965 := [rewrite]: #12964 +#12961 := (iff #3302 #12960) +#12958 := (iff #3301 #12948) +#12953 := (implies true #12948) +#12956 := (iff #12953 #12948) +#12957 := [rewrite]: #12956 +#12954 := (iff #3301 #12953) +#12951 := (iff #3300 #12948) +#12945 := (implies #11974 #12940) +#12949 := (iff #12945 #12948) +#12950 := [rewrite]: #12949 +#12946 := (iff #3300 #12945) +#12943 := (iff #3299 #12940) +#12937 := (implies #12312 #12932) +#12941 := (iff #12937 #12940) +#12942 := [rewrite]: #12941 +#12938 := (iff #3299 #12937) +#12935 := (iff #3298 #12932) +#12928 := (implies #12312 #12923) +#12933 := (iff #12928 #12932) +#12934 := [rewrite]: #12933 +#12929 := (iff #3298 #12928) +#12926 := (iff #3297 #12923) +#12920 := (implies #12006 #12915) +#12924 := (iff #12920 #12923) +#12925 := [rewrite]: #12924 +#12921 := (iff #3297 #12920) +#12918 := (iff #3296 #12915) +#12911 := (implies #3169 #12906) +#12916 := (iff #12911 #12915) +#12917 := [rewrite]: #12916 +#12912 := (iff #3296 #12911) +#12909 := (iff #3295 #12906) +#12902 := (implies #3170 #12897) +#12907 := (iff #12902 #12906) +#12908 := [rewrite]: #12907 +#12903 := (iff #3295 #12902) +#12900 := (iff #3294 #12897) +#12893 := (implies #3171 #12888) +#12898 := (iff #12893 #12897) +#12899 := [rewrite]: #12898 +#12894 := (iff #3294 #12893) +#12891 := (iff #3293 #12888) +#12884 := (implies #3172 #12879) +#12889 := (iff #12884 #12888) +#12890 := [rewrite]: #12889 +#12885 := (iff #3293 #12884) +#12882 := (iff #3292 #12879) +#12875 := (implies #3175 #12863) +#12880 := (iff #12875 #12879) +#12881 := [rewrite]: #12880 +#12876 := (iff #3292 #12875) +#12873 := (iff #3291 #12863) +#12868 := (implies true #12863) +#12871 := (iff #12868 #12863) +#12872 := [rewrite]: #12871 +#12869 := (iff #3291 #12868) +#12866 := (iff #3290 #12863) +#12860 := (implies #11974 #12857) +#12864 := (iff #12860 #12863) +#12865 := [rewrite]: #12864 +#12861 := (iff #3290 #12860) +#12858 := (iff #3289 #12857) +#12855 := (iff #3288 #12845) +#12850 := (implies true #12845) +#12853 := (iff #12850 #12845) +#12854 := [rewrite]: #12853 +#12851 := (iff #3288 #12850) +#12848 := (iff #3287 #12845) +#12842 := (implies #11974 #12837) +#12846 := (iff #12842 #12845) +#12847 := [rewrite]: #12846 +#12843 := (iff #3287 #12842) +#12840 := (iff #3286 #12837) +#12833 := (implies #3282 #12828) +#12838 := (iff #12833 #12837) +#12839 := [rewrite]: #12838 +#12834 := (iff #3286 #12833) +#12831 := (iff #3285 #12828) +#12825 := (implies #11974 #12813) +#12829 := (iff #12825 #12828) +#12830 := [rewrite]: #12829 +#12826 := (iff #3285 #12825) +#12823 := (iff #3284 #12813) +#12818 := (implies true #12813) +#12821 := (iff #12818 #12813) +#12822 := [rewrite]: #12821 +#12819 := (iff #3284 #12818) +#12816 := (iff #3283 #12813) +#12810 := (implies #11974 #12171) +#12814 := (iff #12810 #12813) +#12815 := [rewrite]: #12814 +#12811 := (iff #3283 #12810) +#12174 := (iff #3118 #12171) +#12168 := (implies #11974 #12156) +#12172 := (iff #12168 #12171) +#12173 := [rewrite]: #12172 +#12169 := (iff #3118 #12168) +#12166 := (iff #3117 #12156) +#12161 := (implies true #12156) +#12164 := (iff #12161 #12156) +#12165 := [rewrite]: #12164 +#12162 := (iff #3117 #12161) +#12159 := (iff #3116 #12156) +#12153 := (implies #11974 #12150) +#12157 := (iff #12153 #12156) +#12158 := [rewrite]: #12157 +#12154 := (iff #3116 #12153) +#12151 := (iff #3115 #12150) +#12148 := (iff #3114 #12145) +#12141 := (implies up_216 #12136) +#12146 := (iff #12141 #12145) +#12147 := [rewrite]: #12146 +#12142 := (iff #3114 #12141) +#12139 := (iff #3113 #12136) +#12133 := (implies #11974 #12121) +#12137 := (iff #12133 #12136) +#12138 := [rewrite]: #12137 +#12134 := (iff #3113 #12133) +#12131 := (iff #3112 #12121) +#12126 := (implies true #12121) +#12129 := (iff #12126 #12121) +#12130 := [rewrite]: #12129 +#12127 := (iff #3112 #12126) +#12124 := (iff #3111 #12121) +#12118 := (implies #11974 #12113) +#12122 := (iff #12118 #12121) +#12123 := [rewrite]: #12122 +#12119 := (iff #3111 #12118) +#12116 := (iff #3110 #12113) +#12109 := (implies #11974 #12097) +#12114 := (iff #12109 #12113) +#12115 := [rewrite]: #12114 +#12110 := (iff #3110 #12109) +#12107 := (iff #3109 #12097) +#12102 := (implies true #12097) +#12105 := (iff #12102 #12097) +#12106 := [rewrite]: #12105 +#12103 := (iff #3109 #12102) +#12100 := (iff #3108 #12097) +#12093 := (implies #12011 #12088) +#12098 := (iff #12093 #12097) +#12099 := [rewrite]: #12098 +#12094 := (iff #3108 #12093) +#12091 := (iff #3107 #12088) +#12084 := (implies #12014 #12079) +#12089 := (iff #12084 #12088) +#12090 := [rewrite]: #12089 +#12085 := (iff #3107 #12084) +#12082 := (iff #3106 #12079) +#12075 := (implies #12017 #12070) +#12080 := (iff #12075 #12079) +#12081 := [rewrite]: #12080 +#12076 := (iff #3106 #12075) +#12073 := (iff #3105 #12070) +#12066 := (implies #12020 #12056) +#12071 := (iff #12066 #12070) +#12072 := [rewrite]: #12071 +#12067 := (iff #3105 #12066) +#12064 := (iff #3104 #12056) +#12059 := (implies true #12056) +#12062 := (iff #12059 #12056) +#12063 := [rewrite]: #12062 +#12060 := (iff #3104 #12059) +#12057 := (iff #3103 #12056) +#12054 := (iff #3102 #12051) +#12047 := (implies #12035 #3099) +#12052 := (iff #12047 #12051) +#12053 := [rewrite]: #12052 +#12048 := (iff #3102 #12047) +#12045 := (iff #3101 #3099) +#12040 := (and #3099 true) +#12043 := (iff #12040 #3099) +#12044 := [rewrite]: #12043 +#12041 := (iff #3101 #12040) +#12038 := (iff #3100 true) +#12039 := [rewrite]: #12038 +#12042 := [monotonicity #12039]: #12041 +#12046 := [trans #12042 #12044]: #12045 +#12036 := (iff #3094 #12035) +#12033 := (iff #3093 #12030) +#12027 := (implies #412 #12024) +#12031 := (iff #12027 #12030) +#12032 := [rewrite]: #12031 +#12028 := (iff #3093 #12027) +#12025 := (iff #3092 #12024) +#12026 := [rewrite]: #12025 +#12029 := [monotonicity #12026]: #12028 +#12034 := [trans #12029 #12032]: #12033 +#12037 := [quant-intro #12034]: #12036 +#12049 := [monotonicity #12037 #12046]: #12048 +#12055 := [trans #12049 #12053]: #12054 +#12058 := [monotonicity #12037 #12055]: #12057 +#12061 := [monotonicity #12058]: #12060 +#12065 := [trans #12061 #12063]: #12064 +#12021 := (iff #3089 #12020) +#12022 := [rewrite]: #12021 +#12068 := [monotonicity #12022 #12065]: #12067 +#12074 := [trans #12068 #12072]: #12073 +#12018 := (iff #3087 #12017) +#12019 := [rewrite]: #12018 +#12077 := [monotonicity #12019 #12074]: #12076 +#12083 := [trans #12077 #12081]: #12082 +#12015 := (iff #3085 #12014) +#12016 := [rewrite]: #12015 +#12086 := [monotonicity #12016 #12083]: #12085 +#12092 := [trans #12086 #12090]: #12091 +#12012 := (iff #3083 #12011) +#12013 := [rewrite]: #12012 +#12095 := [monotonicity #12013 #12092]: #12094 +#12101 := [trans #12095 #12099]: #12100 +#12104 := [monotonicity #12101]: #12103 +#12108 := [trans #12104 #12106]: #12107 +#11975 := (iff #3065 #11974) +#11976 := [rewrite]: #11975 +#12111 := [monotonicity #11976 #12108]: #12110 +#12117 := [trans #12111 #12115]: #12116 +#12120 := [monotonicity #11976 #12117]: #12119 +#12125 := [trans #12120 #12123]: #12124 +#12128 := [monotonicity #12125]: #12127 +#12132 := [trans #12128 #12130]: #12131 +#12135 := [monotonicity #11976 #12132]: #12134 +#12140 := [trans #12135 #12138]: #12139 +#12143 := [monotonicity #12140]: #12142 +#12149 := [trans #12143 #12147]: #12148 +#12152 := [monotonicity #12149]: #12151 +#12155 := [monotonicity #11976 #12152]: #12154 +#12160 := [trans #12155 #12158]: #12159 +#12163 := [monotonicity #12160]: #12162 +#12167 := [trans #12163 #12165]: #12166 +#12170 := [monotonicity #11976 #12167]: #12169 +#12175 := [trans #12170 #12173]: #12174 +#12812 := [monotonicity #11976 #12175]: #12811 +#12817 := [trans #12812 #12815]: #12816 +#12820 := [monotonicity #12817]: #12819 +#12824 := [trans #12820 #12822]: #12823 +#12827 := [monotonicity #11976 #12824]: #12826 +#12832 := [trans #12827 #12830]: #12831 +#12835 := [monotonicity #12832]: #12834 +#12841 := [trans #12835 #12839]: #12840 +#12844 := [monotonicity #11976 #12841]: #12843 +#12849 := [trans #12844 #12847]: #12848 +#12852 := [monotonicity #12849]: #12851 +#12856 := [trans #12852 #12854]: #12855 +#12808 := (iff #3281 #12798) +#12803 := (implies true #12798) +#12806 := (iff #12803 #12798) +#12807 := [rewrite]: #12806 +#12804 := (iff #3281 #12803) +#12801 := (iff #3280 #12798) +#12795 := (implies #11974 #12790) +#12799 := (iff #12795 #12798) +#12800 := [rewrite]: #12799 +#12796 := (iff #3280 #12795) +#12793 := (iff #3279 #12790) +#12786 := (implies #3179 #12781) +#12791 := (iff #12786 #12790) +#12792 := [rewrite]: #12791 +#12787 := (iff #3279 #12786) +#12784 := (iff #3278 #12781) +#12778 := (implies #11974 #12766) +#12782 := (iff #12778 #12781) +#12783 := [rewrite]: #12782 +#12779 := (iff #3278 #12778) +#12776 := (iff #3277 #12766) +#12771 := (implies true #12766) +#12774 := (iff #12771 #12766) +#12775 := [rewrite]: #12774 +#12772 := (iff #3277 #12771) +#12769 := (iff #3276 #12766) +#12763 := (implies #11974 #12760) +#12767 := (iff #12763 #12766) +#12768 := [rewrite]: #12767 +#12764 := (iff #3276 #12763) +#12761 := (iff #3275 #12760) +#12758 := (iff #3274 #12755) +#12752 := (implies #12351 #12749) +#12756 := (iff #12752 #12755) +#12757 := [rewrite]: #12756 +#12753 := (iff #3274 #12752) +#12750 := (iff #3273 #12749) +#12747 := (iff #3272 #12744) +#12741 := (implies #12357 #12736) +#12745 := (iff #12741 #12744) +#12746 := [rewrite]: #12745 +#12742 := (iff #3272 #12741) +#12739 := (iff #3271 #12736) +#12733 := (implies #11974 #12730) +#12737 := (iff #12733 #12736) +#12738 := [rewrite]: #12737 +#12734 := (iff #3271 #12733) +#12731 := (iff #3270 #12730) +#12728 := (iff #3269 #12718) +#12723 := (implies true #12718) +#12726 := (iff #12723 #12718) +#12727 := [rewrite]: #12726 +#12724 := (iff #3269 #12723) +#12721 := (iff #3268 #12718) +#12715 := (implies #11974 #12710) +#12719 := (iff #12715 #12718) +#12720 := [rewrite]: #12719 +#12716 := (iff #3268 #12715) +#12713 := (iff #3267 #12710) +#12706 := (implies #3257 #12701) +#12711 := (iff #12706 #12710) +#12712 := [rewrite]: #12711 +#12707 := (iff #3267 #12706) +#12704 := (iff #3266 #12701) +#12698 := (implies #11974 #12686) +#12702 := (iff #12698 #12701) +#12703 := [rewrite]: #12702 +#12699 := (iff #3266 #12698) +#12696 := (iff #3265 #12686) +#12691 := (implies true #12686) +#12694 := (iff #12691 #12686) +#12695 := [rewrite]: #12694 +#12692 := (iff #3265 #12691) +#12689 := (iff #3264 #12686) +#12683 := (implies #11974 #12678) +#12687 := (iff #12683 #12686) +#12688 := [rewrite]: #12687 +#12684 := (iff #3264 #12683) +#12681 := (iff #3263 #12678) +#12675 := (implies #11974 #12663) +#12679 := (iff #12675 #12678) +#12680 := [rewrite]: #12679 +#12676 := (iff #3263 #12675) +#12673 := (iff #3262 #12663) +#12668 := (implies true #12663) +#12671 := (iff #12668 #12663) +#12672 := [rewrite]: #12671 +#12669 := (iff #3262 #12668) +#12666 := (iff #3261 #12663) +#12659 := (implies #12644 #12654) +#12664 := (iff #12659 #12663) +#12665 := [rewrite]: #12664 +#12660 := (iff #3261 #12659) +#12657 := (iff #3260 #12654) +#12650 := (implies #12647 #12500) +#12655 := (iff #12650 #12654) +#12656 := [rewrite]: #12655 +#12651 := (iff #3260 #12650) +#12510 := (iff #3239 #12500) +#12505 := (implies true #12500) +#12508 := (iff #12505 #12500) +#12509 := [rewrite]: #12508 +#12506 := (iff #3239 #12505) +#12503 := (iff #3238 #12500) +#12496 := (implies #3203 #12493) +#12501 := (iff #12496 #12500) +#12502 := [rewrite]: #12501 +#12497 := (iff #3238 #12496) +#12494 := (iff #3237 #12493) +#12491 := (iff #3236 #12488) +#12484 := (implies #12380 #12479) +#12489 := (iff #12484 #12488) +#12490 := [rewrite]: #12489 +#12485 := (iff #3236 #12484) +#12482 := (iff #3235 #12479) +#12475 := (implies #12383 #12470) +#12480 := (iff #12475 #12479) +#12481 := [rewrite]: #12480 +#12476 := (iff #3235 #12475) +#12473 := (iff #3234 #12470) +#12466 := (implies #3211 #12461) +#12471 := (iff #12466 #12470) +#12472 := [rewrite]: #12471 +#12467 := (iff #3234 #12466) +#12464 := (iff #3233 #12461) +#12457 := (implies #12386 #12447) +#12462 := (iff #12457 #12461) +#12463 := [rewrite]: #12462 +#12458 := (iff #3233 #12457) +#12455 := (iff #3232 #12447) +#12450 := (implies true #12447) +#12453 := (iff #12450 #12447) +#12454 := [rewrite]: #12453 +#12451 := (iff #3232 #12450) +#12448 := (iff #3231 #12447) +#12445 := (iff #3230 #12442) +#12438 := (implies #3214 #12435) +#12443 := (iff #12438 #12442) +#12444 := [rewrite]: #12443 +#12439 := (iff #3230 #12438) +#12436 := (iff #3229 #12435) +#12433 := (iff #3228 #12430) +#12426 := (implies #12401 #12407) +#12431 := (iff #12426 #12430) +#12432 := [rewrite]: #12431 +#12427 := (iff #3228 #12426) +#12424 := (iff #3227 #12407) +#12419 := (and #12407 true) +#12422 := (iff #12419 #12407) +#12423 := [rewrite]: #12422 +#12420 := (iff #3227 #12419) +#12417 := (iff #3226 true) +#12412 := (implies #12407 true) +#12415 := (iff #12412 true) +#12416 := [rewrite]: #12415 +#12413 := (iff #3226 #12412) +#12410 := (iff #3225 true) +#12411 := [rewrite]: #12410 +#12408 := (iff #3224 #12407) +#12405 := (iff #3223 #12404) +#12406 := [rewrite]: #12405 +#12409 := [monotonicity #12406]: #12408 +#12414 := [monotonicity #12409 #12411]: #12413 +#12418 := [trans #12414 #12416]: #12417 +#12421 := [monotonicity #12409 #12418]: #12420 +#12425 := [trans #12421 #12423]: #12424 +#12402 := (iff #3219 #12401) +#12399 := (iff #3218 #12396) +#12393 := (implies #412 #12390) +#12397 := (iff #12393 #12396) +#12398 := [rewrite]: #12397 +#12394 := (iff #3218 #12393) +#12391 := (iff #3217 #12390) +#12392 := [rewrite]: #12391 +#12395 := [monotonicity #12392]: #12394 +#12400 := [trans #12395 #12398]: #12399 +#12403 := [quant-intro #12400]: #12402 +#12428 := [monotonicity #12403 #12425]: #12427 +#12434 := [trans #12428 #12432]: #12433 +#12437 := [monotonicity #12403 #12434]: #12436 +#12440 := [monotonicity #12437]: #12439 +#12446 := [trans #12440 #12444]: #12445 +#12449 := [monotonicity #12446]: #12448 +#12452 := [monotonicity #12449]: #12451 +#12456 := [trans #12452 #12454]: #12455 +#12387 := (iff #3213 #12386) +#12388 := [rewrite]: #12387 +#12459 := [monotonicity #12388 #12456]: #12458 +#12465 := [trans #12459 #12463]: #12464 +#12468 := [monotonicity #12465]: #12467 +#12474 := [trans #12468 #12472]: #12473 +#12384 := (iff #3209 #12383) +#12372 := (= #3204 #12371) +#12373 := [rewrite]: #12372 +#12385 := [monotonicity #12373]: #12384 +#12477 := [monotonicity #12385 #12474]: #12476 +#12483 := [trans #12477 #12481]: #12482 +#12381 := (iff #3207 #12380) +#12378 := (iff #3206 #12377) +#12379 := [monotonicity #12373]: #12378 +#12375 := (iff #3205 #12374) +#12376 := [monotonicity #12373]: #12375 +#12382 := [monotonicity #12376 #12379]: #12381 +#12486 := [monotonicity #12382 #12483]: #12485 +#12492 := [trans #12486 #12490]: #12491 +#12495 := [monotonicity #12382 #12492]: #12494 +#12498 := [monotonicity #12495]: #12497 +#12504 := [trans #12498 #12502]: #12503 +#12507 := [monotonicity #12504]: #12506 +#12511 := [trans #12507 #12509]: #12510 +#12648 := (iff #3259 #12647) +#12649 := [rewrite]: #12648 +#12652 := [monotonicity #12649 #12511]: #12651 +#12658 := [trans #12652 #12656]: #12657 +#12645 := (iff #3258 #12644) +#12646 := [rewrite]: #12645 +#12661 := [monotonicity #12646 #12658]: #12660 +#12667 := [trans #12661 #12665]: #12666 +#12670 := [monotonicity #12667]: #12669 +#12674 := [trans #12670 #12672]: #12673 +#12677 := [monotonicity #11976 #12674]: #12676 +#12682 := [trans #12677 #12680]: #12681 +#12685 := [monotonicity #11976 #12682]: #12684 +#12690 := [trans #12685 #12688]: #12689 +#12693 := [monotonicity #12690]: #12692 +#12697 := [trans #12693 #12695]: #12696 +#12700 := [monotonicity #11976 #12697]: #12699 +#12705 := [trans #12700 #12703]: #12704 +#12708 := [monotonicity #12705]: #12707 +#12714 := [trans #12708 #12712]: #12713 +#12717 := [monotonicity #11976 #12714]: #12716 +#12722 := [trans #12717 #12720]: #12721 +#12725 := [monotonicity #12722]: #12724 +#12729 := [trans #12725 #12727]: #12728 +#12642 := (iff #3256 #12632) +#12637 := (implies true #12632) +#12640 := (iff #12637 #12632) +#12641 := [rewrite]: #12640 +#12638 := (iff #3256 #12637) +#12635 := (iff #3255 #12632) +#12629 := (implies #11974 #12624) +#12633 := (iff #12629 #12632) +#12634 := [rewrite]: #12633 +#12630 := (iff #3255 #12629) +#12627 := (iff #3254 #12624) +#12620 := (implies #3190 #12615) +#12625 := (iff #12620 #12624) +#12626 := [rewrite]: #12625 +#12621 := (iff #3254 #12620) +#12618 := (iff #3253 #12615) +#12612 := (implies #11974 #12600) +#12616 := (iff #12612 #12615) +#12617 := [rewrite]: #12616 +#12613 := (iff #3253 #12612) +#12610 := (iff #3252 #12600) +#12605 := (implies true #12600) +#12608 := (iff #12605 #12600) +#12609 := [rewrite]: #12608 +#12606 := (iff #3252 #12605) +#12603 := (iff #3251 #12600) +#12597 := (implies #11974 #12594) +#12601 := (iff #12597 #12600) +#12602 := [rewrite]: #12601 +#12598 := (iff #3251 #12597) +#12595 := (iff #3250 #12594) +#12592 := (iff #3249 #12589) +#12585 := (implies #12351 #12582) +#12590 := (iff #12585 #12589) +#12591 := [rewrite]: #12590 +#12586 := (iff #3249 #12585) +#12583 := (iff #3248 #12582) +#12580 := (iff #3247 #12577) +#12573 := (implies #12357 #12568) +#12578 := (iff #12573 #12577) +#12579 := [rewrite]: #12578 +#12574 := (iff #3247 #12573) +#12571 := (iff #3246 #12568) +#12564 := (implies #12360 #12559) +#12569 := (iff #12564 #12568) +#12570 := [rewrite]: #12569 +#12565 := (iff #3246 #12564) +#12562 := (iff #3245 #12559) +#12555 := (implies #3194 #12550) +#12560 := (iff #12555 #12559) +#12561 := [rewrite]: #12560 +#12556 := (iff #3245 #12555) +#12553 := (iff #3244 #12550) +#12546 := (implies #3196 #12541) +#12551 := (iff #12546 #12550) +#12552 := [rewrite]: #12551 +#12547 := (iff #3244 #12546) +#12544 := (iff #3243 #12541) +#12537 := (implies #3064 #12525) +#12542 := (iff #12537 #12541) +#12543 := [rewrite]: #12542 +#12538 := (iff #3243 #12537) +#12535 := (iff #3242 #12525) +#12530 := (implies true #12525) +#12533 := (iff #12530 #12525) +#12534 := [rewrite]: #12533 +#12531 := (iff #3242 #12530) +#12528 := (iff #3241 #12525) +#12521 := (implies #12365 #12516) +#12526 := (iff #12521 #12525) +#12527 := [rewrite]: #12526 +#12522 := (iff #3241 #12521) +#12519 := (iff #3240 #12516) +#12512 := (implies #12368 #12500) +#12517 := (iff #12512 #12516) +#12518 := [rewrite]: #12517 +#12513 := (iff #3240 #12512) +#12369 := (iff #3201 #12368) +#12370 := [rewrite]: #12369 +#12514 := [monotonicity #12370 #12511]: #12513 +#12520 := [trans #12514 #12518]: #12519 +#12366 := (iff #3199 #12365) +#12367 := [rewrite]: #12366 +#12523 := [monotonicity #12367 #12520]: #12522 +#12529 := [trans #12523 #12527]: #12528 +#12532 := [monotonicity #12529]: #12531 +#12536 := [trans #12532 #12534]: #12535 +#12363 := (iff #3197 #3064) +#12364 := [rewrite]: #12363 +#12539 := [monotonicity #12364 #12536]: #12538 +#12545 := [trans #12539 #12543]: #12544 +#12548 := [monotonicity #12545]: #12547 +#12554 := [trans #12548 #12552]: #12553 +#12557 := [monotonicity #12554]: #12556 +#12563 := [trans #12557 #12561]: #12562 +#12361 := (iff #3192 #12360) +#12362 := [rewrite]: #12361 +#12566 := [monotonicity #12362 #12563]: #12565 +#12572 := [trans #12566 #12570]: #12571 +#12358 := (iff #3188 #12357) +#12355 := (iff #3187 #12354) +#12356 := [rewrite]: #12355 +#12346 := (iff #3182 #12345) +#12347 := [rewrite]: #12346 +#12359 := [monotonicity #12347 #12356]: #12358 +#12575 := [monotonicity #12359 #12572]: #12574 +#12581 := [trans #12575 #12579]: #12580 +#12584 := [monotonicity #12359 #12581]: #12583 +#12352 := (iff #3185 #12351) +#12349 := (iff #3184 #12348) +#12350 := [rewrite]: #12349 +#12353 := [monotonicity #12347 #12350]: #12352 +#12587 := [monotonicity #12353 #12584]: #12586 +#12593 := [trans #12587 #12591]: #12592 +#12596 := [monotonicity #12353 #12593]: #12595 +#12599 := [monotonicity #11976 #12596]: #12598 +#12604 := [trans #12599 #12602]: #12603 +#12607 := [monotonicity #12604]: #12606 +#12611 := [trans #12607 #12609]: #12610 +#12614 := [monotonicity #11976 #12611]: #12613 +#12619 := [trans #12614 #12617]: #12618 +#12622 := [monotonicity #12619]: #12621 +#12628 := [trans #12622 #12626]: #12627 +#12631 := [monotonicity #11976 #12628]: #12630 +#12636 := [trans #12631 #12634]: #12635 +#12639 := [monotonicity #12636]: #12638 +#12643 := [trans #12639 #12641]: #12642 +#12732 := [monotonicity #12643 #12729]: #12731 +#12735 := [monotonicity #11976 #12732]: #12734 +#12740 := [trans #12735 #12738]: #12739 +#12743 := [monotonicity #12359 #12740]: #12742 +#12748 := [trans #12743 #12746]: #12747 +#12751 := [monotonicity #12359 #12748]: #12750 +#12754 := [monotonicity #12353 #12751]: #12753 +#12759 := [trans #12754 #12757]: #12758 +#12762 := [monotonicity #12353 #12759]: #12761 +#12765 := [monotonicity #11976 #12762]: #12764 +#12770 := [trans #12765 #12768]: #12769 +#12773 := [monotonicity #12770]: #12772 +#12777 := [trans #12773 #12775]: #12776 +#12780 := [monotonicity #11976 #12777]: #12779 +#12785 := [trans #12780 #12783]: #12784 +#12788 := [monotonicity #12785]: #12787 +#12794 := [trans #12788 #12792]: #12793 +#12797 := [monotonicity #11976 #12794]: #12796 +#12802 := [trans #12797 #12800]: #12801 +#12805 := [monotonicity #12802]: #12804 +#12809 := [trans #12805 #12807]: #12808 +#12859 := [monotonicity #12809 #12856]: #12858 +#12862 := [monotonicity #11976 #12859]: #12861 +#12867 := [trans #12862 #12865]: #12866 +#12343 := (iff #3178 true) +#12341 := (iff #12338 true) +#12342 := [rewrite]: #12341 +#12339 := (iff #3178 #12338) +#12336 := (iff #3177 true) +#12337 := [rewrite]: #12336 +#12334 := (iff #3176 true) +#12335 := [rewrite]: #12334 +#12340 := [monotonicity #12335 #12337]: #12339 +#12344 := [trans #12340 #12342]: #12343 +#12870 := [monotonicity #12344 #12867]: #12869 +#12874 := [trans #12870 #12872]: #12873 +#12877 := [monotonicity #12874]: #12876 +#12883 := [trans #12877 #12881]: #12882 +#12886 := [monotonicity #12883]: #12885 +#12892 := [trans #12886 #12890]: #12891 +#12895 := [monotonicity #12892]: #12894 +#12901 := [trans #12895 #12899]: #12900 +#12904 := [monotonicity #12901]: #12903 +#12910 := [trans #12904 #12908]: #12909 +#12913 := [monotonicity #12910]: #12912 +#12919 := [trans #12913 #12917]: #12918 +#12009 := (iff #3081 #12006) +#12003 := (and #12000 #11908) +#12007 := (iff #12003 #12006) +#12008 := [rewrite]: #12007 +#12004 := (iff #3081 #12003) +#11909 := (iff #2987 #11908) +#11910 := [rewrite]: #11909 +#12001 := (iff #3080 #12000) +#12002 := [rewrite]: #12001 +#12005 := [monotonicity #12002 #11910]: #12004 +#12010 := [trans #12005 #12008]: #12009 +#12922 := [monotonicity #12010 #12919]: #12921 +#12927 := [trans #12922 #12925]: #12926 +#12313 := (iff #3164 #12312) +#12310 := (iff #3163 #12309) +#12307 := (iff #3162 #12306) +#12308 := [rewrite]: #12307 +#12311 := [monotonicity #12308]: #12310 +#12314 := [monotonicity #12311]: #12313 +#12930 := [monotonicity #12314 #12927]: #12929 +#12936 := [trans #12930 #12934]: #12935 +#12332 := (iff #3168 #12312) +#12315 := (and true #12312) +#12318 := (iff #12315 #12312) +#12319 := [rewrite]: #12318 +#12330 := (iff #3168 #12315) +#12328 := (iff #3167 #12312) +#12326 := (iff #3167 #12315) +#12324 := (iff #3166 #12312) +#12322 := (iff #3166 #12315) +#12320 := (iff #3165 #12312) +#12316 := (iff #3165 #12315) +#12304 := (iff #3155 true) +#12299 := (forall (vars (?x783 T5)) (:pat #3151) true) +#12302 := (iff #12299 true) +#12303 := [elim-unused]: #12302 +#12300 := (iff #3155 #12299) +#12297 := (iff #3154 true) +#12249 := (= uf_9 #3139) +#12261 := (implies #12249 #12249) +#12264 := (iff #12261 true) +#12265 := [rewrite]: #12264 +#12295 := (iff #3154 #12261) +#12293 := (iff #3153 #12249) +#12254 := (and true #12249) +#12257 := (iff #12254 #12249) +#12258 := [rewrite]: #12257 +#12291 := (iff #3153 #12254) +#12250 := (iff #3140 #12249) +#12251 := [rewrite]: #12250 +#12289 := (iff #3152 true) +#12290 := [rewrite]: #12289 +#12292 := [monotonicity #12290 #12251]: #12291 +#12294 := [trans #12292 #12258]: #12293 +#12296 := [monotonicity #12251 #12294]: #12295 +#12298 := [trans #12296 #12265]: #12297 +#12301 := [quant-intro #12298]: #12300 +#12305 := [trans #12301 #12303]: #12304 +#12317 := [monotonicity #12305 #12314]: #12316 +#12321 := [trans #12317 #12319]: #12320 +#12287 := (iff #3148 true) +#12242 := (forall (vars (?x780 T5)) (:pat #3128) true) +#12245 := (iff #12242 true) +#12246 := [elim-unused]: #12245 +#12285 := (iff #3148 #12242) +#12283 := (iff #3147 true) +#12281 := (iff #3147 #12261) +#12279 := (iff #3146 #12249) +#12277 := (iff #3146 #12254) +#12275 := (iff #3145 true) +#12276 := [rewrite]: #12275 +#12278 := [monotonicity #12276 #12251]: #12277 +#12280 := [trans #12278 #12258]: #12279 +#12282 := [monotonicity #12251 #12280]: #12281 +#12284 := [trans #12282 #12265]: #12283 +#12286 := [quant-intro #12284]: #12285 +#12288 := [trans #12286 #12246]: #12287 +#12323 := [monotonicity #12288 #12321]: #12322 +#12325 := [trans #12323 #12319]: #12324 +#12273 := (iff #3144 true) +#12268 := (forall (vars (?x781 T5)) (:pat #3138) true) +#12271 := (iff #12268 true) +#12272 := [elim-unused]: #12271 +#12269 := (iff #3144 #12268) +#12266 := (iff #3143 true) +#12262 := (iff #3143 #12261) +#12259 := (iff #3142 #12249) +#12255 := (iff #3142 #12254) +#12252 := (iff #3141 true) +#12253 := [rewrite]: #12252 +#12256 := [monotonicity #12253 #12251]: #12255 +#12260 := [trans #12256 #12258]: #12259 +#12263 := [monotonicity #12251 #12260]: #12262 +#12267 := [trans #12263 #12265]: #12266 +#12270 := [quant-intro #12267]: #12269 +#12274 := [trans #12270 #12272]: #12273 +#12327 := [monotonicity #12274 #12325]: #12326 +#12329 := [trans #12327 #12319]: #12328 +#12247 := (iff #3135 true) +#12243 := (iff #3135 #12242) +#12240 := (iff #3134 true) +#12229 := (= uf_261 #3131) +#12232 := (not #12229) +#12235 := (implies #12232 #12232) +#12238 := (iff #12235 true) +#12239 := [rewrite]: #12238 +#12236 := (iff #3134 #12235) +#12233 := (iff #3133 #12232) +#12230 := (iff #3132 #12229) +#12231 := [rewrite]: #12230 +#12234 := [monotonicity #12231]: #12233 +#12237 := [monotonicity #12234 #12234]: #12236 +#12241 := [trans #12237 #12239]: #12240 +#12244 := [quant-intro #12241]: #12243 +#12248 := [trans #12244 #12246]: #12247 +#12331 := [monotonicity #12248 #12329]: #12330 +#12333 := [trans #12331 #12319]: #12332 +#12939 := [monotonicity #12333 #12936]: #12938 +#12944 := [trans #12939 #12942]: #12943 +#12947 := [monotonicity #11976 #12944]: #12946 +#12952 := [trans #12947 #12950]: #12951 +#12955 := [monotonicity #12952]: #12954 +#12959 := [trans #12955 #12957]: #12958 +#12962 := [monotonicity #11976 #12959]: #12961 +#12967 := [trans #12962 #12965]: #12966 +#12970 := [monotonicity #12967]: #12969 +#12974 := [trans #12970 #12972]: #12973 +#12977 := [monotonicity #11976 #12974]: #12976 +#12982 := [trans #12977 #12980]: #12981 #12985 := [monotonicity #12982]: #12984 #12989 := [trans #12985 #12987]: #12988 -#13076 := (iff #3247 #13075) -#13077 := [rewrite]: #13076 -#13080 := [monotonicity #13077 #12989]: #13079 -#13086 := [trans #13080 #13084]: #13085 -#13089 := [monotonicity #13086]: #13088 -#13095 := [trans #13089 #13093]: #13094 -#13098 := [monotonicity #13095]: #13097 -#13102 := [trans #13098 #13100]: #13101 -#13073 := (iff #3245 #3059) -#13074 := [rewrite]: #13073 -#13105 := [monotonicity #13074 #13102]: #13104 -#13111 := [trans #13105 #13109]: #13110 -#13114 := [monotonicity #13111]: #13113 -#13120 := [trans #13114 #13118]: #13119 -#13123 := [monotonicity #13120]: #13122 -#13129 := [trans #13123 #13127]: #13128 -#13071 := (iff #3240 #13070) -#13072 := [rewrite]: #13071 -#13132 := [monotonicity #13072 #13129]: #13131 -#13138 := [trans #13132 #13136]: #13137 -#13141 := [monotonicity #12822 #13138]: #13140 -#13147 := [trans #13141 #13145]: #13146 -#13150 := [monotonicity #13147 #12822]: #13149 -#13155 := [trans #13150 #13153]: #13154 -#13158 := [monotonicity #12811 #13155]: #13157 -#13164 := [trans #13158 #13162]: #13163 -#13167 := [monotonicity #13164 #12811]: #13166 -#13172 := [trans #13167 #13170]: #13171 -#13175 := [monotonicity #13172]: #13174 -#13180 := [trans #13175 #13178]: #13179 -#13183 := [monotonicity #13180]: #13182 -#13187 := [trans #13183 #13185]: #13186 -#13190 := [monotonicity #13187]: #13189 -#13195 := [trans #13190 #13193]: #13194 -#13198 := [monotonicity #13195]: #13197 -#13204 := [trans #13198 #13202]: #13203 -#13207 := [monotonicity #13204]: #13206 -#13212 := [trans #13207 #13210]: #13211 -#13215 := [monotonicity #13212]: #13214 -#13219 := [trans #13215 #13217]: #13218 -#13068 := (iff #3237 #13058) -#13063 := (implies true #13058) -#13066 := (iff #13063 #13058) -#13067 := [rewrite]: #13066 -#13064 := (iff #3237 #13063) -#13061 := (iff #3236 #13058) -#13055 := (implies #3060 #13050) -#13059 := (iff #13055 #13058) -#13060 := [rewrite]: #13059 -#13056 := (iff #3236 #13055) -#13053 := (iff #3235 #13050) -#13046 := (implies #3185 #13041) -#13051 := (iff #13046 #13050) -#13052 := [rewrite]: #13051 -#13047 := (iff #3235 #13046) -#13044 := (iff #3234 #13041) -#13038 := (implies #3060 #13026) -#13042 := (iff #13038 #13041) -#13043 := [rewrite]: #13042 -#13039 := (iff #3234 #13038) -#13036 := (iff #3233 #13026) -#13031 := (implies true #13026) -#13034 := (iff #13031 #13026) -#13035 := [rewrite]: #13034 -#13032 := (iff #3233 #13031) -#13029 := (iff #3232 #13026) -#13023 := (implies #3060 #13018) -#13027 := (iff #13023 #13026) -#13028 := [rewrite]: #13027 -#13024 := (iff #3232 #13023) -#13021 := (iff #3231 #13018) -#13015 := (implies #3060 #13003) -#13019 := (iff #13015 #13018) -#13020 := [rewrite]: #13019 -#13016 := (iff #3231 #13015) -#13013 := (iff #3230 #13003) -#13008 := (implies true #13003) -#13011 := (iff #13008 #13003) -#13012 := [rewrite]: #13011 -#13009 := (iff #3230 #13008) -#13006 := (iff #3229 #13003) -#12999 := (implies #12823 #12994) -#13004 := (iff #12999 #13003) -#13005 := [rewrite]: #13004 -#13000 := (iff #3229 #12999) -#12997 := (iff #3228 #12994) -#12990 := (implies #12826 #12978) -#12995 := (iff #12990 #12994) -#12996 := [rewrite]: #12995 -#12991 := (iff #3228 #12990) -#12827 := (iff #3189 #12826) -#12828 := [rewrite]: #12827 -#12992 := [monotonicity #12828 #12989]: #12991 -#12998 := [trans #12992 #12996]: #12997 -#12824 := (iff #3187 #12823) -#12825 := [rewrite]: #12824 -#13001 := [monotonicity #12825 #12998]: #13000 -#13007 := [trans #13001 #13005]: #13006 -#13010 := [monotonicity #13007]: #13009 -#13014 := [trans #13010 #13012]: #13013 -#13017 := [monotonicity #13014]: #13016 -#13022 := [trans #13017 #13020]: #13021 +#12227 := (iff #3125 true) +#12222 := (implies true true) +#12225 := (iff #12222 true) +#12226 := [rewrite]: #12225 +#12223 := (iff #3125 #12222) +#12220 := (iff #3124 true) +#12215 := (implies #11974 true) +#12218 := (iff #12215 true) +#12219 := [rewrite]: #12218 +#12216 := (iff #3124 #12215) +#12213 := (iff #3123 true) +#12180 := (or #12179 #12171) +#12188 := (or #12112 #12180) +#12203 := (or #12112 #12188) +#12208 := (implies false #12203) +#12211 := (iff #12208 true) +#12212 := [rewrite]: #12211 +#12209 := (iff #3123 #12208) +#12206 := (iff #3122 #12203) +#12200 := (implies #11974 #12188) +#12204 := (iff #12200 #12203) +#12205 := [rewrite]: #12204 +#12201 := (iff #3122 #12200) +#12198 := (iff #3121 #12188) +#12193 := (implies true #12188) +#12196 := (iff #12193 #12188) +#12197 := [rewrite]: #12196 +#12194 := (iff #3121 #12193) +#12191 := (iff #3120 #12188) +#12185 := (implies #11974 #12180) +#12189 := (iff #12185 #12188) +#12190 := [rewrite]: #12189 +#12186 := (iff #3120 #12185) +#12183 := (iff #3119 #12180) +#12176 := (implies #12006 #12171) +#12181 := (iff #12176 #12180) +#12182 := [rewrite]: #12181 +#12177 := (iff #3119 #12176) +#12178 := [monotonicity #12010 #12175]: #12177 +#12184 := [trans #12178 #12182]: #12183 +#12187 := [monotonicity #11976 #12184]: #12186 +#12192 := [trans #12187 #12190]: #12191 +#12195 := [monotonicity #12192]: #12194 +#12199 := [trans #12195 #12197]: #12198 +#12202 := [monotonicity #11976 #12199]: #12201 +#12207 := [trans #12202 #12205]: #12206 +#12210 := [monotonicity #11999 #12207]: #12209 +#12214 := [trans #12210 #12212]: #12213 +#12217 := [monotonicity #11976 #12214]: #12216 +#12221 := [trans #12217 #12219]: #12220 +#12224 := [monotonicity #12221]: #12223 +#12228 := [trans #12224 #12226]: #12227 +#12992 := [monotonicity #12228 #12989]: #12991 +#12996 := [trans #12992 #12994]: #12995 +#12999 := [monotonicity #11976 #12996]: #12998 +#13004 := [trans #12999 #13002]: #13003 +#11996 := (iff #3076 #11995) +#11993 := (iff #3075 #11992) +#11994 := [rewrite]: #11993 +#11997 := [monotonicity #11994]: #11996 +#13007 := [monotonicity #11997 #13004]: #13006 +#13013 := [trans #13007 #13011]: #13012 +#11990 := (iff #3071 #11989) +#11987 := (iff #3070 #11984) +#11981 := (implies #412 #11978) +#11985 := (iff #11981 #11984) +#11986 := [rewrite]: #11985 +#11982 := (iff #3070 #11981) +#11979 := (iff #3069 #11978) +#11980 := [rewrite]: #11979 +#11983 := [monotonicity #11980]: #11982 +#11988 := [trans #11983 #11986]: #11987 +#11991 := [quant-intro #11988]: #11990 +#13016 := [monotonicity #11991 #13013]: #13015 +#13022 := [trans #13016 #13020]: #13021 #13025 := [monotonicity #13022]: #13024 -#13030 := [trans #13025 #13028]: #13029 -#13033 := [monotonicity #13030]: #13032 -#13037 := [trans #13033 #13035]: #13036 -#13040 := [monotonicity #13037]: #13039 -#13045 := [trans #13040 #13043]: #13044 -#13048 := [monotonicity #13045]: #13047 -#13054 := [trans #13048 #13052]: #13053 -#13057 := [monotonicity #13054]: #13056 -#13062 := [trans #13057 #13060]: #13061 -#13065 := [monotonicity #13062]: #13064 -#13069 := [trans #13065 #13067]: #13068 -#13222 := [monotonicity #13069 #13219]: #13221 -#13225 := [monotonicity #13222]: #13224 -#13230 := [trans #13225 #13228]: #13229 -#13233 := [monotonicity #12822 #13230]: #13232 -#13238 := [trans #13233 #13236]: #13237 -#13241 := [monotonicity #13238 #12822]: #13240 -#13246 := [trans #13241 #13244]: #13245 -#13249 := [monotonicity #12811 #13246]: #13248 -#13254 := [trans #13249 #13252]: #13253 -#13257 := [monotonicity #13254 #12811]: #13256 -#13262 := [trans #13257 #13260]: #13261 +#13031 := [trans #13025 #13029]: #13030 +#13034 := [monotonicity #11976 #13031]: #13033 +#13039 := [trans #13034 #13037]: #13038 +#13042 := [monotonicity #13039]: #13041 +#13048 := [trans #13042 #13046]: #13047 +#13051 := [monotonicity #13048]: #13050 +#13057 := [trans #13051 #13055]: #13056 +#13060 := [monotonicity #13057]: #13059 +#13066 := [trans #13060 #13064]: #13065 +#13069 := [monotonicity #13066]: #13068 +#13073 := [trans #13069 #13071]: #13072 +#11972 := (iff #3051 #11971) +#11969 := (iff #3050 #3027) +#11970 := [rewrite]: #11969 +#11973 := [monotonicity #11970]: #11972 +#13076 := [monotonicity #11973 #13073]: #13075 +#13082 := [trans #13076 #13080]: #13081 +#13085 := [monotonicity #11973 #13082]: #13084 +#11967 := (iff #3049 #11966) +#11964 := (iff #3048 #11961) +#11958 := (implies #412 #11955) +#11962 := (iff #11958 #11961) +#11963 := [rewrite]: #11962 +#11959 := (iff #3048 #11958) +#11956 := (iff #3047 #11955) +#11957 := [rewrite]: #11956 +#11960 := [monotonicity #11957]: #11959 +#11965 := [trans #11960 #11963]: #11964 +#11968 := [quant-intro #11965]: #11967 +#13088 := [monotonicity #11968 #13085]: #13087 +#13094 := [trans #13088 #13092]: #13093 +#13097 := [monotonicity #11968 #13094]: #13096 +#13100 := [monotonicity #13097]: #13099 +#13106 := [trans #13100 #13104]: #13105 +#13109 := [monotonicity #13106]: #13108 +#11952 := (iff #3041 #11951) +#11949 := (iff #3040 #11948) +#11946 := (iff #3039 #3038) +#11947 := [rewrite]: #11946 +#11950 := [monotonicity #11947]: #11949 +#11953 := [monotonicity #11950]: #11952 +#13112 := [monotonicity #11953 #13109]: #13111 +#13118 := [trans #13112 #13116]: #13117 +#13121 := [monotonicity #13118]: #13120 +#13127 := [trans #13121 #13125]: #13126 +#13130 := [monotonicity #13127]: #13129 +#13136 := [trans #13130 #13134]: #13135 +#13139 := [monotonicity #13136]: #13138 +#13145 := [trans #13139 #13143]: #13144 +#13148 := [monotonicity #13145]: #13147 +#13154 := [trans #13148 #13152]: #13153 +#11944 := (iff #3024 #11943) +#11941 := (iff #3023 #11940) +#11942 := [rewrite]: #11941 +#11932 := (iff #3018 #11931) +#11933 := [rewrite]: #11932 +#11945 := [monotonicity #11933 #11942]: #11944 +#13157 := [monotonicity #11945 #13154]: #13156 +#13163 := [trans #13157 #13161]: #13162 +#13166 := [monotonicity #11945 #13163]: #13165 +#11938 := (iff #3021 #11937) +#11935 := (iff #3020 #11934) +#11936 := [rewrite]: #11935 +#11939 := [monotonicity #11933 #11936]: #11938 +#13169 := [monotonicity #11939 #13166]: #13168 +#13175 := [trans #13169 #13173]: #13174 +#13178 := [monotonicity #11939 #13175]: #13177 +#11929 := (iff #3015 #11928) +#11930 := [rewrite]: #11929 +#13181 := [monotonicity #11930 #13178]: #13180 +#13187 := [trans #13181 #13185]: #13186 +#13190 := [monotonicity #11930 #13187]: #13189 +#13193 := [monotonicity #13190]: #13192 +#13199 := [trans #13193 #13197]: #13198 +#11926 := (iff #3009 #11925) +#11923 := (iff #3008 #11920) +#11917 := (iff #11914 false) +#11921 := (iff #11917 #11920) +#11922 := [rewrite]: #11921 +#11918 := (iff #3008 #11917) +#11915 := (iff #3007 #11914) +#11916 := [rewrite]: #11915 +#11919 := [monotonicity #11916]: #11918 +#11924 := [trans #11919 #11922]: #11923 +#11927 := [quant-intro #11924]: #11926 +#13202 := [monotonicity #11927 #13199]: #13201 +#13208 := [trans #13202 #13206]: #13207 +#13211 := [monotonicity #13208]: #13210 +#13217 := [trans #13211 #13215]: #13216 +#13220 := [monotonicity #13217]: #13219 +#13226 := [trans #13220 #13224]: #13225 +#13229 := [monotonicity #13226]: #13228 +#13235 := [trans #13229 #13233]: #13234 +#13238 := [monotonicity #13235]: #13237 +#13244 := [trans #13238 #13242]: #13243 +#11912 := (iff #2988 #11911) +#11906 := (iff #2985 #11905) +#11907 := [rewrite]: #11906 +#11913 := [monotonicity #11907 #11910]: #11912 +#13247 := [monotonicity #11913 #13244]: #13246 +#13253 := [trans #13247 #13251]: #13252 +#11903 := (iff #2982 #11902) +#11904 := [rewrite]: #11903 +#13256 := [monotonicity #11904 #13253]: #13255 +#13262 := [trans #13256 #13260]: #13261 #13265 := [monotonicity #13262]: #13264 -#13270 := [trans #13265 #13268]: #13269 -#13273 := [monotonicity #13270]: #13272 -#13277 := [trans #13273 #13275]: #13276 -#13280 := [monotonicity #13277]: #13279 -#13285 := [trans #13280 #13283]: #13284 -#13288 := [monotonicity #13285]: #13287 -#13294 := [trans #13288 #13292]: #13293 -#13297 := [monotonicity #13294]: #13296 -#13302 := [trans #13297 #13300]: #13301 -#13305 := [monotonicity #13302]: #13304 -#13309 := [trans #13305 #13307]: #13308 -#12801 := (iff #3173 #12791) -#12796 := (implies true #12791) -#12799 := (iff #12796 #12791) -#12800 := [rewrite]: #12799 -#12797 := (iff #3173 #12796) -#12794 := (iff #3172 #12791) -#12788 := (implies #3060 #12783) -#12792 := (iff #12788 #12791) -#12793 := [rewrite]: #12792 -#12789 := (iff #3172 #12788) -#12786 := (iff #3171 #12783) -#12779 := (implies #3131 #12774) -#12784 := (iff #12779 #12783) -#12785 := [rewrite]: #12784 -#12780 := (iff #3171 #12779) -#12777 := (iff #3170 #12774) -#12771 := (implies #3060 #12759) -#12775 := (iff #12771 #12774) -#12776 := [rewrite]: #12775 -#12772 := (iff #3170 #12771) -#12769 := (iff #3169 #12759) -#12764 := (implies true #12759) -#12767 := (iff #12764 #12759) -#12768 := [rewrite]: #12767 -#12765 := (iff #3169 #12764) -#12762 := (iff #3168 #12759) -#12756 := (implies #3060 #12751) -#12760 := (iff #12756 #12759) -#12761 := [rewrite]: #12760 -#12757 := (iff #3168 #12756) -#12758 := [monotonicity #12755]: #12757 -#12763 := [trans #12758 #12761]: #12762 -#12766 := [monotonicity #12763]: #12765 -#12770 := [trans #12766 #12768]: #12769 -#12773 := [monotonicity #12770]: #12772 -#12778 := [trans #12773 #12776]: #12777 -#12781 := [monotonicity #12778]: #12780 -#12787 := [trans #12781 #12785]: #12786 -#12790 := [monotonicity #12787]: #12789 -#12795 := [trans #12790 #12793]: #12794 -#12798 := [monotonicity #12795]: #12797 -#12802 := [trans #12798 #12800]: #12801 -#13312 := [monotonicity #12802 #13309]: #13311 -#13315 := [monotonicity #13312]: #13314 -#13320 := [trans #13315 #13318]: #13319 -#12565 := (iff #3130 true) -#12520 := (iff #12517 true) -#12521 := [rewrite]: #12520 -#12563 := (iff #3130 #12517) -#12561 := (iff #3129 true) -#12562 := [rewrite]: #12561 -#12559 := (iff #3128 true) -#12560 := [rewrite]: #12559 -#12564 := [monotonicity #12560 #12562]: #12563 -#12566 := [trans #12564 #12521]: #12565 -#13323 := [monotonicity #12566 #13320]: #13322 -#13327 := [trans #13323 #13325]: #13326 -#13330 := [monotonicity #13327]: #13329 -#13336 := [trans #13330 #13334]: #13335 -#13339 := [monotonicity #13336]: #13338 -#13345 := [trans #13339 #13343]: #13344 -#13348 := [monotonicity #13345]: #13347 -#13354 := [trans #13348 #13352]: #13353 -#13357 := [monotonicity #13354]: #13356 -#13363 := [trans #13357 #13361]: #13362 -#13366 := [monotonicity #13363]: #13365 -#13372 := [trans #13366 #13370]: #13371 -#13375 := [monotonicity #12558 #13372]: #13374 -#13381 := [trans #13375 #13379]: #13380 -#12551 := (iff #3116 #12544) -#12541 := (and #12437 #3115) -#12545 := (iff #12541 #12544) -#12546 := [rewrite]: #12545 -#12549 := (iff #3116 #12541) -#12438 := (iff #3073 #12437) -#12439 := [rewrite]: #12438 -#12550 := [monotonicity #12439]: #12549 -#12552 := [trans #12550 #12546]: #12551 -#13384 := [monotonicity #12552 #13381]: #13383 -#13390 := [trans #13384 #13388]: #13389 -#12547 := (iff #3114 #12544) -#12542 := (iff #3114 #12541) -#12539 := (iff #3113 #3115) -#12537 := (iff #3112 #3078) -#12532 := (and #3078 true) -#12535 := (iff #12532 #3078) -#12536 := [rewrite]: #12535 -#12533 := (iff #3112 #12532) -#12530 := (iff #3111 true) -#12528 := (iff #3111 #12517) -#12526 := (iff #3110 true) -#12524 := (iff #3110 #12517) -#12522 := (iff #3109 true) -#12518 := (iff #3109 #12517) -#12515 := (iff #3108 true) -#12476 := (forall (vars (?x777 T5)) (:pat #3090) true) -#12479 := (iff #12476 true) -#12480 := [elim-unused]: #12479 -#12513 := (iff #3108 #12476) -#12511 := (iff #3107 true) -#12500 := (= uf_261 #3104) -#12503 := (not #12500) -#12506 := (implies #12503 #12503) -#12509 := (iff #12506 true) -#12510 := [rewrite]: #12509 -#12507 := (iff #3107 #12506) -#12504 := (iff #3106 #12503) -#12501 := (iff #3105 #12500) -#12502 := [rewrite]: #12501 -#12505 := [monotonicity #12502]: #12504 -#12508 := [monotonicity #12505 #12505]: #12507 -#12512 := [trans #12508 #12510]: #12511 -#12514 := [quant-intro #12512]: #12513 -#12516 := [trans #12514 #12480]: #12515 -#12498 := (iff #3101 true) -#12493 := (forall (vars (?x778 T5)) (:pat #3097) true) -#12496 := (iff #12493 true) -#12497 := [elim-unused]: #12496 -#12494 := (iff #3101 #12493) -#12491 := (iff #3100 true) -#12440 := (= uf_9 #3082) -#12452 := (implies #12440 #12440) -#12455 := (iff #12452 true) -#12456 := [rewrite]: #12455 -#12489 := (iff #3100 #12452) -#12487 := (iff #3099 #12440) -#12445 := (and #12440 true) -#12448 := (iff #12445 #12440) -#12449 := [rewrite]: #12448 -#12485 := (iff #3099 #12445) -#12483 := (iff #3098 true) -#12484 := [rewrite]: #12483 -#12441 := (iff #3083 #12440) -#12442 := [rewrite]: #12441 -#12486 := [monotonicity #12442 #12484]: #12485 -#12488 := [trans #12486 #12449]: #12487 -#12490 := [monotonicity #12442 #12488]: #12489 -#12492 := [trans #12490 #12456]: #12491 -#12495 := [quant-intro #12492]: #12494 -#12499 := [trans #12495 #12497]: #12498 -#12519 := [monotonicity #12499 #12516]: #12518 -#12523 := [trans #12519 #12521]: #12522 -#12481 := (iff #3094 true) -#12477 := (iff #3094 #12476) -#12474 := (iff #3093 true) -#12472 := (iff #3093 #12452) -#12470 := (iff #3092 #12440) -#12468 := (iff #3092 #12445) -#12466 := (iff #3091 true) -#12467 := [rewrite]: #12466 -#12469 := [monotonicity #12442 #12467]: #12468 -#12471 := [trans #12469 #12449]: #12470 -#12473 := [monotonicity #12442 #12471]: #12472 -#12475 := [trans #12473 #12456]: #12474 -#12478 := [quant-intro #12475]: #12477 -#12482 := [trans #12478 #12480]: #12481 -#12525 := [monotonicity #12482 #12523]: #12524 -#12527 := [trans #12525 #12521]: #12526 -#12464 := (iff #3087 true) -#12459 := (forall (vars (?x776 T5)) (:pat #3081) true) -#12462 := (iff #12459 true) -#12463 := [elim-unused]: #12462 -#12460 := (iff #3087 #12459) -#12457 := (iff #3086 true) -#12453 := (iff #3086 #12452) -#12450 := (iff #3085 #12440) -#12446 := (iff #3085 #12445) -#12443 := (iff #3084 true) -#12444 := [rewrite]: #12443 -#12447 := [monotonicity #12442 #12444]: #12446 -#12451 := [trans #12447 #12449]: #12450 -#12454 := [monotonicity #12442 #12451]: #12453 -#12458 := [trans #12454 #12456]: #12457 -#12461 := [quant-intro #12458]: #12460 -#12465 := [trans #12461 #12463]: #12464 -#12529 := [monotonicity #12465 #12527]: #12528 -#12531 := [trans #12529 #12521]: #12530 -#12534 := [monotonicity #12531]: #12533 -#12538 := [trans #12534 #12536]: #12537 -#12540 := [monotonicity #12538]: #12539 -#12543 := [monotonicity #12439 #12540]: #12542 -#12548 := [trans #12543 #12546]: #12547 -#13393 := [monotonicity #12548 #13390]: #13392 -#13398 := [trans #13393 #13396]: #13397 -#13401 := [monotonicity #13398]: #13400 -#13406 := [trans #13401 #13404]: #13405 -#13409 := [monotonicity #13406]: #13408 -#13413 := [trans #13409 #13411]: #13412 -#13416 := [monotonicity #13413]: #13415 -#13421 := [trans #13416 #13419]: #13420 -#13424 := [monotonicity #13421]: #13423 -#13428 := [trans #13424 #13426]: #13427 -#13431 := [monotonicity #13428]: #13430 -#13436 := [trans #13431 #13434]: #13435 -#13439 := [monotonicity #13436]: #13438 -#13443 := [trans #13439 #13441]: #13442 -#13500 := [monotonicity #13443 #13497]: #13499 -#13504 := [trans #13500 #13502]: #13503 -#13507 := [monotonicity #13504]: #13506 -#13512 := [trans #13507 #13510]: #13511 -#12435 := (iff #3071 #12432) -#12429 := (and #12426 #3070) -#12433 := (iff #12429 #12432) -#12434 := [rewrite]: #12433 -#12430 := (iff #3071 #12429) -#12427 := (iff #3069 #12426) -#12428 := [rewrite]: #12427 -#12431 := [monotonicity #12428]: #12430 -#12436 := [trans #12431 #12434]: #12435 -#13515 := [monotonicity #12436 #13512]: #13514 -#13521 := [trans #13515 #13519]: #13520 -#12424 := (iff #3066 #12423) -#12421 := (iff #3065 #12418) -#12415 := (implies #5718 #12412) -#12419 := (iff #12415 #12418) -#12420 := [rewrite]: #12419 -#12416 := (iff #3065 #12415) -#12413 := (iff #3064 #12412) -#12414 := [rewrite]: #12413 -#12417 := [monotonicity #5720 #12414]: #12416 -#12422 := [trans #12417 #12420]: #12421 -#12425 := [quant-intro #12422]: #12424 -#13524 := [monotonicity #12425 #13521]: #13523 -#13530 := [trans #13524 #13528]: #13529 -#13533 := [monotonicity #13530]: #13532 -#13539 := [trans #13533 #13537]: #13538 -#13542 := [monotonicity #13539]: #13541 -#13547 := [trans #13542 #13545]: #13546 -#13550 := [monotonicity #13547]: #13549 -#13556 := [trans #13550 #13554]: #13555 -#13559 := [monotonicity #13556]: #13558 -#13565 := [trans #13559 #13563]: #13564 -#13568 := [monotonicity #13565]: #13567 -#13574 := [trans #13568 #13572]: #13573 -#13577 := [monotonicity #13574]: #13576 -#13581 := [trans #13577 #13579]: #13580 -#13584 := [monotonicity #12410 #13581]: #13583 -#13590 := [trans #13584 #13588]: #13589 -#13593 := [monotonicity #13590 #12410]: #13592 -#13598 := [trans #13593 #13596]: #13597 -#13601 := [monotonicity #12400 #13598]: #13600 -#13607 := [trans #13601 #13605]: #13606 -#13610 := [monotonicity #13607 #12400]: #13609 -#13615 := [trans #13610 #13613]: #13614 -#13618 := [monotonicity #13615]: #13617 -#13624 := [trans #13618 #13622]: #13623 -#13627 := [monotonicity #13624]: #13626 -#13632 := [trans #13627 #13630]: #13631 -#12384 := (iff #3036 #12383) -#12381 := (iff #3035 #12380) -#12378 := (iff #3034 #3033) -#12379 := [rewrite]: #12378 -#12382 := [monotonicity #12379]: #12381 -#12385 := [monotonicity #12382]: #12384 -#13635 := [monotonicity #12385 #13632]: #13634 -#13641 := [trans #13635 #13639]: #13640 -#13644 := [monotonicity #13641]: #13643 -#13650 := [trans #13644 #13648]: #13649 -#13653 := [monotonicity #13650]: #13652 -#13659 := [trans #13653 #13657]: #13658 -#13662 := [monotonicity #13659]: #13661 -#13668 := [trans #13662 #13666]: #13667 -#13671 := [monotonicity #13668]: #13670 -#13677 := [trans #13671 #13675]: #13676 -#13680 := [monotonicity #12377 #13677]: #13679 -#13686 := [trans #13680 #13684]: #13685 -#13689 := [monotonicity #13686 #12377]: #13688 -#13694 := [trans #13689 #13692]: #13693 -#13697 := [monotonicity #12366 #13694]: #13696 -#13703 := [trans #13697 #13701]: #13702 -#13706 := [monotonicity #13703 #12366]: #13705 -#13711 := [trans #13706 #13709]: #13710 -#13714 := [monotonicity #12357 #13711]: #13713 -#13720 := [trans #13714 #13718]: #13719 -#13723 := [monotonicity #13720 #12357]: #13722 -#13728 := [trans #13723 #13726]: #13727 -#13731 := [monotonicity #13728]: #13730 -#13737 := [trans #13731 #13735]: #13736 -#12353 := (iff #3004 #12352) -#12350 := (iff #3003 #12347) -#12344 := (iff #12341 false) -#12348 := (iff #12344 #12347) -#12349 := [rewrite]: #12348 -#12345 := (iff #3003 #12344) -#12342 := (iff #3002 #12341) -#12343 := [rewrite]: #12342 -#12346 := [monotonicity #12343]: #12345 -#12351 := [trans #12346 #12349]: #12350 -#12354 := [quant-intro #12351]: #12353 -#13740 := [monotonicity #12354 #13737]: #13739 -#13746 := [trans #13740 #13744]: #13745 -#13749 := [monotonicity #13746]: #13748 -#13755 := [trans #13749 #13753]: #13754 -#13758 := [monotonicity #13755]: #13757 -#13764 := [trans #13758 #13762]: #13763 -#13767 := [monotonicity #13764]: #13766 -#13773 := [trans #13767 #13771]: #13772 -#13776 := [monotonicity #13773]: #13775 -#13782 := [trans #13776 #13780]: #13781 -#12339 := (iff #2983 #12338) -#12336 := (iff #2982 #12335) -#12337 := [rewrite]: #12336 -#12340 := [monotonicity #12334 #12337]: #12339 -#13785 := [monotonicity #12340 #13782]: #13784 -#13791 := [trans #13785 #13789]: #13790 -#12330 := (iff #2977 #12329) -#12331 := [rewrite]: #12330 -#13794 := [monotonicity #12331 #13791]: #13793 -#13800 := [trans #13794 #13798]: #13799 -#13803 := [monotonicity #13800]: #13802 -#13807 := [trans #13803 #13805]: #13806 -#12327 := (iff #2975 #12326) -#12324 := (iff #2974 #12323) -#12321 := (iff #2973 #12320) -#12318 := (iff #2972 #12317) -#12315 := (iff #2971 #12314) -#12312 := (iff #2970 #12311) -#12313 := [rewrite]: #12312 -#12309 := (iff #2968 #12308) -#12310 := [rewrite]: #12309 -#12316 := [monotonicity #12310 #12313]: #12315 -#12306 := (iff #2966 #12305) -#12307 := [rewrite]: #12306 -#12319 := [monotonicity #12307 #12316]: #12318 -#12303 := (iff #2964 #12302) -#12304 := [rewrite]: #12303 -#12322 := [monotonicity #12304 #12319]: #12321 -#12300 := (iff #2957 #12299) -#12297 := (iff #2956 #12296) -#12298 := [rewrite]: #12297 -#12301 := [monotonicity #12298]: #12300 -#12325 := [monotonicity #12301 #12322]: #12324 -#12294 := (iff #2954 #12293) -#12295 := [rewrite]: #12294 -#12328 := [monotonicity #12295 #12325]: #12327 -#13810 := [monotonicity #12328 #13807]: #13809 -#13816 := [trans #13810 #13814]: #13815 -#13819 := [monotonicity #13816]: #13818 -#13825 := [trans #13819 #13823]: #13824 -#13828 := [monotonicity #13825]: #13827 -#13834 := [trans #13828 #13832]: #13833 -#13837 := [monotonicity #13834]: #13836 -#13843 := [trans #13837 #13841]: #13842 -#13846 := [monotonicity #13843]: #13845 -#13852 := [trans #13846 #13850]: #13851 -#13855 := [monotonicity #13852]: #13854 -#13861 := [trans #13855 #13859]: #13860 -#13864 := [monotonicity #13861]: #13863 -#13868 := [trans #13864 #13866]: #13867 -#13871 := [monotonicity #13868]: #13870 -#14775 := [trans #13871 #14773]: #14774 -#12292 := [asserted]: #3345 -#14776 := [mp #12292 #14775]: #14771 -#14794 := [not-or-elim #14776]: #14658 -#14798 := [and-elim #14794]: #12305 +#13269 := [trans #13265 #13267]: #13268 +#11900 := (iff #2980 #11899) +#11897 := (iff #2979 #11896) +#11894 := (iff #2978 #11893) +#11891 := (iff #2977 #11890) +#11888 := (iff #2976 #11887) +#11885 := (iff #2975 #11884) +#11886 := [rewrite]: #11885 +#11882 := (iff #2973 #11881) +#11879 := (iff #2972 #11878) +#11880 := [rewrite]: #11879 +#11883 := [monotonicity #11880]: #11882 +#11889 := [monotonicity #11883 #11886]: #11888 +#11876 := (iff #2970 #11875) +#11877 := [rewrite]: #11876 +#11892 := [monotonicity #11877 #11889]: #11891 +#11873 := (iff #2968 #11872) +#11874 := [rewrite]: #11873 +#11895 := [monotonicity #11874 #11892]: #11894 +#11870 := (iff #2966 #11869) +#11871 := [rewrite]: #11870 +#11898 := [monotonicity #11871 #11895]: #11897 +#11867 := (iff #2964 #11866) +#11868 := [rewrite]: #11867 +#11901 := [monotonicity #11868 #11898]: #11900 +#13272 := [monotonicity #11901 #13269]: #13271 +#13278 := [trans #13272 #13276]: #13277 +#13281 := [monotonicity #13278]: #13280 +#13287 := [trans #13281 #13285]: #13286 +#13290 := [monotonicity #13287]: #13289 +#13296 := [trans #13290 #13294]: #13295 +#13299 := [monotonicity #13296]: #13298 +#13305 := [trans #13299 #13303]: #13304 +#13308 := [monotonicity #13305]: #13307 +#13314 := [trans #13308 #13312]: #13313 +#13317 := [monotonicity #13314]: #13316 +#13323 := [trans #13317 #13321]: #13322 +#13326 := [monotonicity #13323]: #13325 +#13330 := [trans #13326 #13328]: #13329 +#13333 := [monotonicity #13330]: #13332 +#14241 := [trans #13333 #14239]: #14240 +#11865 := [asserted]: #3349 +#14242 := [mp #11865 #14241]: #14237 +#14260 := [not-or-elim #14242]: #14124 +#14264 := [and-elim #14260]: #11875 +#27446 := [trans #14264 #27445]: #25896 +#27208 := (>= #27025 0::int) +decl ?x776!15 :: int +#18607 := ?x776!15 +#18612 := (uf_66 #2960 ?x776!15 uf_7) +#18613 := (uf_110 uf_273 #18612) +#18958 := (* -1::int #18613) +#18959 := (+ uf_299 #18958) +#18960 := (>= #18959 0::int) +#18945 := (* -1::int ?x776!15) +#18946 := (+ uf_272 #18945) +#18947 := (<= #18946 0::int) +#18609 := (>= ?x776!15 0::int) +#22442 := (not #18609) +#18608 := (<= ?x776!15 4294967295::int) +#22441 := (not #18608) +#22457 := (or #22441 #22442 #18947 #18960) +#22462 := (not #22457) +#22415 := (not #3095) +#22416 := (or #22415 #4987 #13725 #19482) +#23335 := (forall (vars (?x778 int)) (:pat #23194) #22416) +#23340 := (not #23335) +#22407 := (or #4987 #13725 #13739 #19482) +#23327 := (forall (vars (?x776 int)) (:pat #23194) #22407) +#23332 := (not #23327) +#23343 := (or #23332 #23340) +#23346 := (not #23343) +#23349 := (or #23346 #22462) +#23352 := (not #23349) +#22348 := (not #13409) +#23355 := (or #12096 #12087 #12078 #12069 #22348 #13552 #13698 #23352) +#23358 := (not #23355) +#23285 := (or #12662 #12653 #22348 #13552 #22307 #13603 #23252) +#23288 := (not #23285) +#23261 := (or #18425 #18434 #23258) +#23264 := (not #23261) +#23267 := (or #18425 #18428 #23264) +#23270 := (not #23267) +#23273 := (or #18425 #18428 #23270) +#23276 := (not #23273) +#23279 := (or #22348 #13552 #13604 #23276) +#23282 := (not #23279) +#23291 := (or #23282 #23288) +#23294 := (not #23291) +#23297 := (or #18425 #18434 #22348 #13552 #23294) +#23300 := (not #23297) +#23303 := (or #18425 #18434 #23300) +#23306 := (not #23303) +#23309 := (or #18425 #18428 #23306) +#23312 := (not #23309) +#23315 := (or #18425 #18428 #23312) +#23318 := (not #23315) +#23321 := (or #22348 #13552 #13697 #23318) +#23324 := (not #23321) +#23361 := (or #23324 #23358) +#23364 := (not #23361) +#22223 := (or #4987 #13890 #13904 #19482) +#23203 := (forall (vars (?x775 int)) (:pat #23194) #22223) +#23208 := (not #23203) +#1331 := 255::int +#15781 := (<= uf_292 255::int) +#22492 := (not #15781) +#15764 := (<= uf_293 4294967295::int) +#22491 := (not #15764) +#15747 := (<= uf_294 4294967295::int) +#22490 := (not #15747) +#22488 := (not #13952) +#22487 := (not #13930) +#15177 := (not #12000) +#22486 := (not #11992) +#22485 := (not #3174) +#22484 := (not #3173) +#23367 := (or #13149 #12914 #12905 #12896 #12887 #22484 #22485 #22486 #15177 #13870 #13404 #22348 #13552 #13875 #13927 #22487 #22488 #22490 #22491 #22492 #23208 #23364) +#23370 := (not #23367) +#23373 := (or #13149 #13404 #23370) +#23376 := (not #23373) +#22212 := (or #4987 #13371 #13383 #19482) +#23195 := (forall (vars (?x773 int)) (:pat #23194) #22212) +#23200 := (not #23195) +#23379 := (or #23200 #23376) +#23382 := (not #23379) +decl ?x773!13 :: int +#18346 := ?x773!13 +#18356 := (>= ?x773!13 1::int) +#18351 := (uf_66 #2960 ?x773!13 uf_7) +#18352 := (uf_110 uf_273 #18351) +#18353 := (* -1::int #18352) +#18354 := (+ uf_285 #18353) +#18355 := (>= #18354 0::int) +#18348 := (>= ?x773!13 0::int) +#22186 := (not #18348) +#18347 := (<= ?x773!13 4294967295::int) +#22185 := (not #18347) +#22201 := (or #22185 #22186 #18355 #18356) +#22206 := (not #22201) +#23385 := (or #22206 #23382) +#23388 := (not #23385) +#23391 := (or #13368 #23388) +#23394 := (not #23391) +#23397 := (or #13368 #23394) +#23400 := (not #23397) +#18323 := (not #11940) +#18314 := (not #11931) +#23403 := (or #13149 #13140 #13131 #13122 #18314 #18323 #23400) +#23406 := (not #23403) +#23409 := (or #18314 #18323 #23406) +#23412 := (not #23409) +#18317 := (not #11934) +#23415 := (or #18314 #18317 #23412) +#23418 := (not #23415) +#23421 := (or #18314 #18317 #23418) +#23424 := (not #23421) +#23427 := (or #13182 #23424) +#23430 := (not #23427) +#23468 := (uf_116 #2962) +#24856 := (uf_43 #2958 #23468) +#25434 := (uf_200 uf_273 #24856 #24856 uf_284) +#25872 := (= #25434 #3014) +#25946 := (= #3014 #25434) +#24863 := (= #2962 #24856) +#14263 := [and-elim #14260]: #11872 #233 := (:var 0 T3) #15 := (:var 1 T5) -#2661 := (uf_48 #15 #233) -#2662 := (pattern #2661) -#11594 := (= uf_9 #2661) -#11601 := (not #11594) -#1250 := (uf_116 #15) -#2664 := (uf_43 #233 #1250) -#2665 := (= #15 #2664) -#11602 := (or #2665 #11601) -#11607 := (forall (vars (?x710 T5) (?x711 T3)) (:pat #2662) #11602) -#18734 := (~ #11607 #11607) -#18732 := (~ #11602 #11602) -#18733 := [refl]: #18732 -#18735 := [nnf-pos #18733]: #18734 -#2663 := (= #2661 uf_9) -#2666 := (implies #2663 #2665) -#2667 := (forall (vars (?x710 T5) (?x711 T3)) (:pat #2662) #2666) -#11608 := (iff #2667 #11607) -#11605 := (iff #2666 #11602) -#11598 := (implies #11594 #2665) -#11603 := (iff #11598 #11602) -#11604 := [rewrite]: #11603 -#11599 := (iff #2666 #11598) -#11596 := (iff #2663 #11594) -#11597 := [rewrite]: #11596 -#11600 := [monotonicity #11597]: #11599 -#11606 := [trans #11600 #11604]: #11605 -#11609 := [quant-intro #11606]: #11608 -#11593 := [asserted]: #2667 -#11612 := [mp #11593 #11609]: #11607 -#18736 := [mp~ #11612 #18735]: #11607 -#25403 := (not #12305) -#25416 := (not #11607) -#25417 := (or #25416 #25403 #25411) -#25412 := (or #25411 #25403) -#25418 := (or #25416 #25412) -#25425 := (iff #25418 #25417) -#25413 := (or #25403 #25411) -#25420 := (or #25416 #25413) -#25423 := (iff #25420 #25417) -#25424 := [rewrite]: #25423 -#25421 := (iff #25418 #25420) -#25414 := (iff #25412 #25413) -#25415 := [rewrite]: #25414 -#25422 := [monotonicity #25415]: #25421 -#25426 := [trans #25422 #25424]: #25425 -#25419 := [quant-inst]: #25418 -#25427 := [mp #25419 #25426]: #25417 -#27939 := [unit-resolution #25427 #18736 #14798]: #25411 -#27941 := [symm #27939]: #27940 -#26337 := [monotonicity #27941]: #26336 -#26339 := [trans #26337 #28359]: #26338 -#26341 := [monotonicity #26339]: #26340 -#26306 := [monotonicity #26341]: #26342 -#26296 := [symm #26306]: #26293 -#26299 := [monotonicity #26296]: #26298 -#14796 := [and-elim #14794]: #12299 -#26307 := [mp #14796 #26299]: #26297 +#2666 := (uf_48 #15 #233) +#2667 := (pattern #2666) +#11162 := (= uf_9 #2666) +#11169 := (not #11162) +#1259 := (uf_116 #15) +#2669 := (uf_43 #233 #1259) +#2670 := (= #15 #2669) +#11170 := (or #2670 #11169) +#11175 := (forall (vars (?x710 T5) (?x711 T3)) (:pat #2667) #11170) +#18151 := (~ #11175 #11175) +#18149 := (~ #11170 #11170) +#18150 := [refl]: #18149 +#18152 := [nnf-pos #18150]: #18151 +#2668 := (= #2666 uf_9) +#2671 := (implies #2668 #2670) +#2672 := (forall (vars (?x710 T5) (?x711 T3)) (:pat #2667) #2671) +#11176 := (iff #2672 #11175) +#11173 := (iff #2671 #11170) +#11166 := (implies #11162 #2670) +#11171 := (iff #11166 #11170) +#11172 := [rewrite]: #11171 +#11167 := (iff #2671 #11166) +#11164 := (iff #2668 #11162) +#11165 := [rewrite]: #11164 +#11168 := [monotonicity #11165]: #11167 +#11174 := [trans #11168 #11172]: #11173 +#11177 := [quant-intro #11174]: #11176 +#11161 := [asserted]: #2672 +#11180 := [mp #11161 #11177]: #11175 +#18153 := [mp~ #11180 #18152]: #11175 +#24855 := (not #11872) +#24868 := (not #11175) +#24869 := (or #24868 #24855 #24863) +#24864 := (or #24863 #24855) +#24870 := (or #24868 #24864) +#24877 := (iff #24870 #24869) +#24865 := (or #24855 #24863) +#24872 := (or #24868 #24865) +#24875 := (iff #24872 #24869) +#24876 := [rewrite]: #24875 +#24873 := (iff #24870 #24872) +#24866 := (iff #24864 #24865) +#24867 := [rewrite]: #24866 +#24874 := [monotonicity #24867]: #24873 +#24878 := [trans #24874 #24876]: #24877 +#24871 := [quant-inst]: #24870 +#24879 := [mp #24871 #24878]: #24869 +#25847 := [unit-resolution #24879 #18153 #14263]: #24863 +#25533 := [monotonicity #25847 #25847]: #25946 +#25593 := [symm #25533]: #25872 +#25435 := (= uf_9 #25434) decl uf_196 :: (-> T4 T5 T5 T2) -#25980 := (uf_196 uf_273 #25404 #25404) -#25981 := (= uf_9 #25980) -#26002 := (not #25981) -#25982 := (uf_200 uf_273 #25404 #25404 uf_284) -#25983 := (= uf_9 #25982) -#25985 := (iff #25981 #25983) -#2240 := (:var 0 T16) -#24 := (:var 2 T5) +#25432 := (uf_196 uf_273 #24856 #24856) +#25433 := (= uf_9 #25432) +#25437 := (iff #25433 #25435) +#2245 := (:var 0 T16) +#21 := (:var 2 T5) #13 := (:var 3 T4) -#2251 := (uf_200 #13 #24 #15 #2240) -#2252 := (pattern #2251) -#2254 := (uf_196 #13 #24 #15) -#10555 := (= uf_9 #2254) -#10551 := (= uf_9 #2251) -#10558 := (iff #10551 #10555) -#10561 := (forall (vars (?x586 T4) (?x587 T5) (?x588 T5) (?x589 T16)) (:pat #2252) #10558) -#18376 := (~ #10561 #10561) -#18374 := (~ #10558 #10558) -#18375 := [refl]: #18374 -#18377 := [nnf-pos #18375]: #18376 -#2255 := (= #2254 uf_9) -#2253 := (= #2251 uf_9) -#2256 := (iff #2253 #2255) -#2257 := (forall (vars (?x586 T4) (?x587 T5) (?x588 T5) (?x589 T16)) (:pat #2252) #2256) -#10562 := (iff #2257 #10561) -#10559 := (iff #2256 #10558) -#10556 := (iff #2255 #10555) -#10557 := [rewrite]: #10556 -#10553 := (iff #2253 #10551) -#10554 := [rewrite]: #10553 -#10560 := [monotonicity #10554 #10557]: #10559 -#10563 := [quant-intro #10560]: #10562 -#10550 := [asserted]: #2257 -#10566 := [mp #10550 #10563]: #10561 -#18378 := [mp~ #10566 #18377]: #10561 -#25995 := (not #10561) -#25996 := (or #25995 #25985) -#25984 := (iff #25983 #25981) -#25997 := (or #25995 #25984) -#26025 := (iff #25997 #25996) -#26077 := (iff #25996 #25996) -#26078 := [rewrite]: #26077 -#25986 := (iff #25984 #25985) -#25987 := [rewrite]: #25986 -#26076 := [monotonicity #25987]: #26025 -#26079 := [trans #26076 #26078]: #26025 -#26023 := [quant-inst]: #25997 -#26015 := [mp #26023 #26079]: #25996 -#27937 := [unit-resolution #26015 #18378]: #25985 -#25999 := (not #25983) -#26332 := (iff #13715 #25999) -#26334 := (iff #12355 #25983) -#26301 := (iff #25983 #12355) -#27942 := (= #25982 #3009) -#27943 := [monotonicity #27941 #27941]: #27942 -#26333 := [monotonicity #27943]: #26301 -#26335 := [symm #26333]: #26334 -#26349 := [monotonicity #26335]: #26332 -#26300 := [hypothesis]: #13715 -#26350 := [mp #26300 #26349]: #25999 -#26022 := (not #25985) -#25991 := (or #26022 #26002 #25983) -#25989 := [def-axiom]: #25991 -#26348 := [unit-resolution #25989 #26350 #27937]: #26002 -#26086 := (uf_48 #25404 #25815) -#26087 := (= uf_9 #26086) -#26398 := (= #2965 #26086) -#26351 := (= #26086 #2965) -#26352 := [monotonicity #27941 #26339]: #26351 -#26399 := [symm #26352]: #26398 -#26400 := [trans #14798 #26399]: #26087 -#26089 := (uf_27 uf_273 #25404) -#26090 := (= uf_9 #26089) -#26324 := (= #2963 #26089) -#26323 := (= #26089 #2963) -#26325 := [monotonicity #27941]: #26323 -#26327 := [symm #26325]: #26324 -#14797 := [and-elim #14794]: #12302 -#26322 := [trans #14797 #26327]: #26090 -#26091 := (not #26090) -#26088 := (not #26087) -#26490 := (or #25981 #26088 #26091 #26095) -#25827 := (uf_25 uf_273 #25404) -#26084 := (= uf_26 #25827) -#26331 := (= #2967 #25827) -#26328 := (= #25827 #2967) -#26329 := [monotonicity #27941]: #26328 -#26401 := [symm #26329]: #26331 -#14799 := [and-elim #14794]: #12308 -#26403 := [trans #14799 #26401]: #26084 -#25853 := (uf_24 uf_273 #25404) -#25854 := (= uf_9 #25853) -#26391 := (= #2969 #25853) -#26388 := (= #25853 #2969) -#26402 := [monotonicity #27941]: #26388 -#26389 := [symm #26402]: #26391 -#14800 := [and-elim #14794]: #12311 -#26392 := [trans #14800 #26389]: #25854 -#25816 := (uf_22 #25815) -#25823 := (= uf_9 #25816) -#26413 := (= #2953 #25816) -#26393 := (= #25816 #2953) -#26394 := [monotonicity #26339]: #26393 -#26414 := [symm #26394]: #26413 -#14795 := [and-elim #14794]: #12293 -#26488 := [trans #14795 #26414]: #25823 -#14783 := [not-or-elim #14776]: #12338 -#14784 := [and-elim #14783]: #12332 -#47 := (:var 1 T4) -#2213 := (uf_196 #47 #26 #26) -#2214 := (pattern #2213) -#10431 := (= uf_9 #2213) +#2256 := (uf_200 #13 #21 #15 #2245) +#2257 := (pattern #2256) +#2259 := (uf_196 #13 #21 #15) +#10125 := (= uf_9 #2259) +#10121 := (= uf_9 #2256) +#10128 := (iff #10121 #10125) +#10131 := (forall (vars (?x586 T4) (?x587 T5) (?x588 T5) (?x589 T16)) (:pat #2257) #10128) +#17793 := (~ #10131 #10131) +#17791 := (~ #10128 #10128) +#17792 := [refl]: #17791 +#17794 := [nnf-pos #17792]: #17793 +#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) +#10132 := (iff #2262 #10131) +#10129 := (iff #2261 #10128) +#10126 := (iff #2260 #10125) +#10127 := [rewrite]: #10126 +#10123 := (iff #2258 #10121) +#10124 := [rewrite]: #10123 +#10130 := [monotonicity #10124 #10127]: #10129 +#10133 := [quant-intro #10130]: #10132 +#10120 := [asserted]: #2262 +#10136 := [mp #10120 #10133]: #10131 +#17795 := [mp~ #10136 #17794]: #10131 +#25449 := (not #10131) +#25475 := (or #25449 #25437) +#25436 := (iff #25435 #25433) +#25448 := (or #25449 #25436) +#25528 := (iff #25448 #25475) +#25530 := (iff #25475 #25475) +#25531 := [rewrite]: #25530 +#25438 := (iff #25436 #25437) +#25439 := [rewrite]: #25438 +#25529 := [monotonicity #25439]: #25528 +#25467 := [trans #25529 #25531]: #25528 +#25477 := [quant-inst]: #25448 +#25474 := [mp #25477 #25467]: #25475 +#25951 := [unit-resolution #25474 #17795]: #25437 +#25473 := (not #25437) +#25516 := (or #25473 #25435) +#25267 := (uf_13 #24856) +#25544 := (uf_12 #25267) +#25547 := (= uf_14 #25544) +#25790 := (not #25547) +#25792 := (iff #11881 #25790) +#25854 := (iff #11878 #25547) +#25852 := (iff #25547 #11878) +#25851 := (= #25544 #2971) +#25787 := (= #25267 #2958) +#23686 := (uf_13 #2962) +#25785 := (= #23686 #2958) +#23689 := (= #2958 #23686) +#2697 := (uf_43 #326 #161) +#23131 := (pattern #2697) +#2701 := (uf_13 #2697) +#11240 := (= #326 #2701) +#23138 := (forall (vars (?x720 T3) (?x721 int)) (:pat #23131) #11240) +#11244 := (forall (vars (?x720 T3) (?x721 int)) #11240) +#23141 := (iff #11244 #23138) +#23139 := (iff #11240 #11240) +#23140 := [refl]: #23139 +#23142 := [quant-intro #23140]: #23141 +#18176 := (~ #11244 #11244) +#18174 := (~ #11240 #11240) +#18175 := [refl]: #18174 +#18177 := [nnf-pos #18175]: #18176 +#2702 := (= #2701 #326) +#2703 := (forall (vars (?x720 T3) (?x721 int)) #2702) +#11245 := (iff #2703 #11244) +#11242 := (iff #2702 #11240) +#11243 := [rewrite]: #11242 +#11246 := [quant-intro #11243]: #11245 +#11239 := [asserted]: #2703 +#11249 := [mp #11239 #11246]: #11244 +#18178 := [mp~ #11249 #18177]: #11244 +#23143 := [mp #18178 #23142]: #23138 +#23633 := (not #23138) +#23694 := (or #23633 #23689) +#23695 := [quant-inst]: #23694 +#25772 := [unit-resolution #23695 #23143]: #23689 +#25786 := [symm #25772]: #25785 +#25773 := (= #25267 #23686) +#25870 := (= #24856 #2962) +#25871 := [symm #25847]: #25870 +#25789 := [monotonicity #25871]: #25773 +#25788 := [trans #25789 #25786]: #25787 +#25784 := [monotonicity #25788]: #25851 +#25853 := [monotonicity #25784]: #25852 +#25855 := [symm #25853]: #25854 +#25794 := [monotonicity #25855]: #25792 +#14265 := [and-elim #14260]: #11881 +#25795 := [mp #14265 #25794]: #25790 +#25536 := (uf_24 uf_273 #24856) +#25537 := (= uf_9 #25536) +#25793 := (= #2969 #25536) +#25796 := (= #25536 #2969) +#25791 := [monotonicity #25871]: #25796 +#25798 := [symm #25791]: #25793 +#25781 := [trans #14264 #25798]: #25537 +#25539 := (uf_48 #24856 #25267) +#25540 := (= uf_9 #25539) +#25841 := (= #2967 #25539) +#25782 := (= #25539 #2967) +#25780 := [monotonicity #25871 #25788]: #25782 +#26004 := [symm #25780]: #25841 +#26005 := [trans #14263 #26004]: #25540 +#25541 := (not #25540) +#25538 := (not #25537) +#26027 := (or #25538 #25541 #25547) +#25279 := (uf_25 uf_273 #24856) +#25542 := (= uf_26 #25279) +#25938 := (= #2965 #25279) +#26006 := (= #25279 #2965) +#25934 := [monotonicity #25871]: #26006 +#25939 := [symm #25934]: #25938 +#14262 := [and-elim #14260]: #11869 +#25940 := [trans #14262 #25939]: #25542 +#25454 := (not #25433) +#25935 := [hypothesis]: #25454 +#25305 := (uf_27 uf_273 #24856) +#25306 := (= uf_9 #25305) +#25943 := (= #2963 #25305) +#25936 := (= #25305 #2963) +#25941 := [monotonicity #25871]: #25936 +#25944 := [symm #25941]: #25943 +#14261 := [and-elim #14260]: #11866 +#25945 := [trans #14261 #25944]: #25306 +#25268 := (uf_23 #25267) +#25275 := (= uf_9 #25268) +#26035 := (= #2974 #25268) +#26030 := (= #25268 #2974) +#26031 := [monotonicity #25788]: #26030 +#26036 := [symm #26031]: #26035 +#14266 := [and-elim #14260]: #11884 +#26026 := [trans #14266 #26036]: #25275 +#14249 := [not-or-elim #14242]: #11911 +#14251 := [and-elim #14249]: #11908 +#2217 := (uf_196 #47 #23 #23) +#2218 := (pattern #2217) +#10006 := (= uf_9 #2217) #227 := (uf_55 #47) -#3939 := (= uf_9 #227) -#19933 := (not #3939) -#150 := (uf_25 #47 #26) -#3656 := (= uf_26 #150) -#19808 := (not #3656) -#33 := (uf_15 #26) -#148 := (uf_48 #26 #33) -#3653 := (= uf_9 #148) -#19807 := (not #3653) -#146 := (uf_27 #47 #26) -#3650 := (= uf_9 #146) -#11522 := (not #3650) -#135 := (uf_24 #47 #26) -#3635 := (= uf_9 #135) -#11145 := (not #3635) -#69 := (uf_22 #33) -#3470 := (= uf_9 #69) -#11200 := (not #3470) -#34 := (uf_14 #33) -#36 := (= #34 uf_16) -#22334 := (or #36 #11200 #11145 #11522 #19807 #19808 #19933 #10431) -#22339 := (forall (vars (?x572 T4) (?x573 T5)) (:pat #2214) #22334) -#52 := (not #36) -#10446 := (and #52 #3470 #3635 #3650 #3653 #3656 #3939) -#10449 := (not #10446) -#10455 := (or #10431 #10449) -#10460 := (forall (vars (?x572 T4) (?x573 T5)) (:pat #2214) #10455) -#22340 := (iff #10460 #22339) -#22337 := (iff #10455 #22334) -#22320 := (or #36 #11200 #11145 #11522 #19807 #19808 #19933) -#22331 := (or #10431 #22320) -#22335 := (iff #22331 #22334) -#22336 := [rewrite]: #22335 -#22332 := (iff #10455 #22331) -#22329 := (iff #10449 #22320) -#22321 := (not #22320) -#22324 := (not #22321) -#22327 := (iff #22324 #22320) -#22328 := [rewrite]: #22327 -#22325 := (iff #10449 #22324) -#22322 := (iff #10446 #22321) -#22323 := [rewrite]: #22322 -#22326 := [monotonicity #22323]: #22325 -#22330 := [trans #22326 #22328]: #22329 -#22333 := [monotonicity #22330]: #22332 -#22338 := [trans #22333 #22336]: #22337 -#22341 := [quant-intro #22338]: #22340 -#18344 := (~ #10460 #10460) -#18342 := (~ #10455 #10455) -#18343 := [refl]: #18342 -#18345 := [nnf-pos #18343]: #18344 -#2220 := (= #2213 uf_9) -#229 := (= #227 uf_9) +#3921 := (= uf_9 #227) +#19350 := (not #3921) +#144 := (uf_48 #23 #26) +#3642 := (= uf_9 #144) +#19225 := (not #3642) +#19224 := (not #3639) +#135 := (uf_27 #47 #23) +#3624 := (= uf_9 #135) +#10715 := (not #3624) +#71 := (uf_23 #26) +#3477 := (= uf_9 #71) +#10770 := (not #3477) +#21783 := (or #29 #10770 #10715 #19224 #19225 #11090 #19350 #10006) +#21788 := (forall (vars (?x572 T4) (?x573 T5)) (:pat #2218) #21783) +#10021 := (and #52 #3477 #3624 #3639 #3642 #3645 #3921) +#10024 := (not #10021) +#10030 := (or #10006 #10024) +#10035 := (forall (vars (?x572 T4) (?x573 T5)) (:pat #2218) #10030) +#21789 := (iff #10035 #21788) +#21786 := (iff #10030 #21783) +#21769 := (or #29 #10770 #10715 #19224 #19225 #11090 #19350) +#21780 := (or #10006 #21769) +#21784 := (iff #21780 #21783) +#21785 := [rewrite]: #21784 +#21781 := (iff #10030 #21780) +#21778 := (iff #10024 #21769) +#21770 := (not #21769) +#21773 := (not #21770) +#21776 := (iff #21773 #21769) +#21777 := [rewrite]: #21776 +#21774 := (iff #10024 #21773) +#21771 := (iff #10021 #21770) +#21772 := [rewrite]: #21771 +#21775 := [monotonicity #21772]: #21774 +#21779 := [trans #21775 #21777]: #21778 +#21782 := [monotonicity #21779]: #21781 +#21787 := [trans #21782 #21785]: #21786 +#21790 := [quant-intro #21787]: #21789 +#17761 := (~ #10035 #10035) +#17759 := (~ #10030 #10030) +#17760 := [refl]: #17759 +#17762 := [nnf-pos #17760]: #17761 +#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) -#230 := (and #136 #229) -#151 := (= #150 uf_26) -#2215 := (and #151 #230) -#149 := (= #148 uf_9) -#2216 := (and #149 #2215) -#147 := (= #146 uf_9) -#2217 := (and #147 #2216) -#2218 := (and #52 #2217) -#70 := (= #69 uf_9) -#2219 := (and #70 #2218) -#2221 := (implies #2219 #2220) -#2222 := (forall (vars (?x572 T4) (?x573 T5)) (:pat #2214) #2221) -#10463 := (iff #2222 #10460) -#3943 := (and #3635 #3939) -#10415 := (and #3656 #3943) -#10419 := (and #3653 #10415) -#10422 := (and #3650 #10419) -#10425 := (and #52 #10422) -#10428 := (and #3470 #10425) -#10437 := (not #10428) -#10438 := (or #10437 #10431) -#10443 := (forall (vars (?x572 T4) (?x573 T5)) (:pat #2214) #10438) -#10461 := (iff #10443 #10460) -#10458 := (iff #10438 #10455) -#10452 := (or #10449 #10431) -#10456 := (iff #10452 #10455) -#10457 := [rewrite]: #10456 -#10453 := (iff #10438 #10452) -#10450 := (iff #10437 #10449) -#10447 := (iff #10428 #10446) -#10448 := [rewrite]: #10447 -#10451 := [monotonicity #10448]: #10450 -#10454 := [monotonicity #10451]: #10453 -#10459 := [trans #10454 #10457]: #10458 -#10462 := [quant-intro #10459]: #10461 -#10444 := (iff #2222 #10443) -#10441 := (iff #2221 #10438) -#10434 := (implies #10428 #10431) -#10439 := (iff #10434 #10438) -#10440 := [rewrite]: #10439 -#10435 := (iff #2221 #10434) -#10432 := (iff #2220 #10431) -#10433 := [rewrite]: #10432 -#10429 := (iff #2219 #10428) -#10426 := (iff #2218 #10425) -#10423 := (iff #2217 #10422) -#10420 := (iff #2216 #10419) -#10417 := (iff #2215 #10415) -#3944 := (iff #230 #3943) -#3941 := (iff #229 #3939) -#3942 := [rewrite]: #3941 -#3637 := (iff #136 #3635) -#3638 := [rewrite]: #3637 -#3945 := [monotonicity #3638 #3942]: #3944 -#3657 := (iff #151 #3656) -#3658 := [rewrite]: #3657 -#10418 := [monotonicity #3658 #3945]: #10417 -#3654 := (iff #149 #3653) -#3655 := [rewrite]: #3654 -#10421 := [monotonicity #3655 #10418]: #10420 -#3651 := (iff #147 #3650) -#3652 := [rewrite]: #3651 -#10424 := [monotonicity #3652 #10421]: #10423 -#10427 := [monotonicity #10424]: #10426 -#3471 := (iff #70 #3470) -#3472 := [rewrite]: #3471 -#10430 := [monotonicity #3472 #10427]: #10429 -#10436 := [monotonicity #10430 #10433]: #10435 -#10442 := [trans #10436 #10440]: #10441 -#10445 := [quant-intro #10442]: #10444 -#10464 := [trans #10445 #10462]: #10463 -#10414 := [asserted]: #2222 -#10465 := [mp #10414 #10464]: #10460 -#18346 := [mp~ #10465 #18345]: #10460 -#22342 := [mp #18346 #22341]: #22339 -#26085 := (not #26084) -#25880 := (not #25854) -#25824 := (not #25823) -#23209 := (not #12332) -#26081 := (not #22339) -#26110 := (or #26081 #23209 #25824 #25880 #25981 #26085 #26088 #26091 #26095) -#26093 := (= #26092 uf_16) -#26094 := (or #26093 #25824 #25880 #26091 #26088 #26085 #23209 #25981) -#26111 := (or #26081 #26094) -#26219 := (iff #26111 #26110) -#26101 := (or #23209 #25824 #25880 #25981 #26085 #26088 #26091 #26095) -#26107 := (or #26081 #26101) -#26215 := (iff #26107 #26110) -#26218 := [rewrite]: #26215 -#26113 := (iff #26111 #26107) -#26104 := (iff #26094 #26101) -#26098 := (or #26095 #25824 #25880 #26091 #26088 #26085 #23209 #25981) -#26102 := (iff #26098 #26101) -#26103 := [rewrite]: #26102 -#26099 := (iff #26094 #26098) -#26096 := (iff #26093 #26095) -#26097 := [rewrite]: #26096 -#26100 := [monotonicity #26097]: #26099 -#26105 := [trans #26100 #26103]: #26104 -#26150 := [monotonicity #26105]: #26113 -#26181 := [trans #26150 #26218]: #26219 -#26112 := [quant-inst]: #26111 -#26165 := [mp #26112 #26181]: #26110 -#26491 := [unit-resolution #26165 #22342 #14784 #26488 #26392 #26403]: #26490 -#26493 := [unit-resolution #26491 #26322 #26400 #26348 #26307]: false -#26494 := [lemma #26493]: #12355 -#23984 := (or #13715 #23981) -#22978 := (forall (vars (?x782 int)) #22967) -#22985 := (not #22978) -#22963 := (forall (vars (?x781 int)) #22958) -#22984 := (not #22963) -#22986 := (or #22984 #22985) -#22987 := (not #22986) -#23016 := (or #22987 #23013) -#23022 := (not #23016) -#23023 := (or #12671 #12662 #12653 #12644 #22873 #14243 #14049 #23022) -#23024 := (not #23023) -#22802 := (forall (vars (?x785 int)) #22797) -#22808 := (not #22802) -#22809 := (or #22784 #22808) -#22810 := (not #22809) -#22839 := (or #22810 #22836) -#22845 := (not #22839) -#22846 := (or #14105 #22845) -#22847 := (not #22846) -#22852 := (or #14105 #22847) -#22860 := (not #22852) -#22861 := (or #12942 #22858 #19034 #22859 #14172 #19037 #22860) -#22862 := (not #22861) -#22867 := (or #19034 #19037 #22862) -#22874 := (not #22867) -#22884 := (or #13124 #13115 #13090 #19011 #19017 #13133 #13081 #14243 #22858 #22874) -#22885 := (not #22884) -#22890 := (or #19011 #19017 #22885) -#22896 := (not #22890) -#22897 := (or #19008 #19011 #22896) -#22898 := (not #22897) -#22903 := (or #19008 #19011 #22898) -#22909 := (not #22903) -#22910 := (or #22873 #14243 #14207 #22909) -#22911 := (not #22910) -#22875 := (or #13002 #12993 #22873 #14243 #22858 #14211 #22874) -#22876 := (not #22875) -#22916 := (or #22876 #22911) -#22922 := (not #22916) -#22923 := (or #19011 #19017 #22873 #14243 #22922) -#22924 := (not #22923) -#22929 := (or #19011 #19017 #22924) -#22935 := (not #22929) -#22936 := (or #19008 #19011 #22935) -#22937 := (not #22936) -#22942 := (or #19008 #19011 #22937) -#22948 := (not #22942) -#22949 := (or #22873 #14243 #14046 #22948) -#22950 := (not #22949) -#23029 := (or #22950 #23024) -#23044 := (not #23029) -#22779 := (forall (vars (?x774 int)) #22774) -#23040 := (not #22779) -#23045 := (or #13672 #13367 #13358 #13349 #13340 #23035 #23036 #23037 #14399 #15709 #13942 #22873 #14243 #14404 #14456 #23038 #23039 #23041 #23042 #23043 #23040 #23044) -#23046 := (not #23045) -#23051 := (or #13672 #13942 #23046) -#23058 := (not #23051) -#22768 := (forall (vars (?x773 int)) #22763) -#23057 := (not #22768) -#23059 := (or #23057 #23058) -#23060 := (not #23059) -#23065 := (or #22757 #23060) -#23071 := (not #23065) -#23072 := (or #13906 #23071) -#23073 := (not #23072) -#23078 := (or #13906 #23073) -#23084 := (not #23078) -#23085 := (or #13672 #13663 #13654 #13645 #18900 #18906 #23084) -#23086 := (not #23085) -#23091 := (or #18900 #18906 #23086) -#23097 := (not #23091) -#23098 := (or #18897 #18900 #23097) -#23099 := (not #23098) -#23104 := (or #18897 #18900 #23099) -#23110 := (not #23104) -#23111 := (or #13715 #23110) -#23112 := (not #23111) -#23117 := (or #13715 #23112) -#23985 := (iff #23117 #23984) -#23982 := (iff #23112 #23981) -#23979 := (iff #23111 #23978) -#23976 := (iff #23110 #23975) -#23973 := (iff #23104 #23972) -#23970 := (iff #23099 #23969) -#23967 := (iff #23098 #23966) -#23964 := (iff #23097 #23963) -#23961 := (iff #23091 #23960) -#23958 := (iff #23086 #23957) -#23955 := (iff #23085 #23954) -#23952 := (iff #23084 #23951) -#23949 := (iff #23078 #23948) -#23946 := (iff #23073 #23945) -#23943 := (iff #23072 #23942) -#23940 := (iff #23071 #23939) -#23937 := (iff #23065 #23936) -#23934 := (iff #23060 #23933) -#23931 := (iff #23059 #23930) -#23928 := (iff #23058 #23927) -#23925 := (iff #23051 #23924) -#23922 := (iff #23046 #23921) -#23919 := (iff #23045 #23918) -#23916 := (iff #23044 #23915) -#23913 := (iff #23029 #23912) -#23910 := (iff #23024 #23909) -#23907 := (iff #23023 #23906) -#23904 := (iff #23022 #23903) -#23901 := (iff #23016 #23900) -#23898 := (iff #22987 #23897) -#23895 := (iff #22986 #23894) -#23892 := (iff #22985 #23891) -#23889 := (iff #22978 #23886) -#23887 := (iff #22967 #22967) -#23888 := [refl]: #23887 -#23890 := [quant-intro #23888]: #23889 -#23893 := [monotonicity #23890]: #23892 -#23884 := (iff #22984 #23883) -#23881 := (iff #22963 #23878) -#23879 := (iff #22958 #22958) -#23880 := [refl]: #23879 -#23882 := [quant-intro #23880]: #23881 -#23885 := [monotonicity #23882]: #23884 -#23896 := [monotonicity #23885 #23893]: #23895 -#23899 := [monotonicity #23896]: #23898 -#23902 := [monotonicity #23899]: #23901 -#23905 := [monotonicity #23902]: #23904 -#23908 := [monotonicity #23905]: #23907 -#23911 := [monotonicity #23908]: #23910 -#23876 := (iff #22950 #23875) -#23873 := (iff #22949 #23872) -#23870 := (iff #22948 #23869) -#23867 := (iff #22942 #23866) -#23864 := (iff #22937 #23863) -#23861 := (iff #22936 #23860) -#23858 := (iff #22935 #23857) -#23855 := (iff #22929 #23854) -#23852 := (iff #22924 #23851) -#23849 := (iff #22923 #23848) -#23846 := (iff #22922 #23845) -#23843 := (iff #22916 #23842) -#23840 := (iff #22911 #23839) -#23837 := (iff #22910 #23836) -#23834 := (iff #22909 #23833) -#23831 := (iff #22903 #23830) -#23828 := (iff #22898 #23827) -#23825 := (iff #22897 #23824) -#23822 := (iff #22896 #23821) -#23819 := (iff #22890 #23818) -#23816 := (iff #22885 #23815) -#23813 := (iff #22884 #23812) -#23804 := (iff #22874 #23803) -#23801 := (iff #22867 #23800) -#23798 := (iff #22862 #23797) -#23795 := (iff #22861 #23794) -#23792 := (iff #22860 #23791) -#23789 := (iff #22852 #23788) -#23786 := (iff #22847 #23785) -#23783 := (iff #22846 #23782) -#23780 := (iff #22845 #23779) -#23777 := (iff #22839 #23776) -#23774 := (iff #22810 #23773) -#23771 := (iff #22809 #23770) -#23768 := (iff #22808 #23767) -#23765 := (iff #22802 #23762) -#23763 := (iff #22797 #22797) -#23764 := [refl]: #23763 -#23766 := [quant-intro #23764]: #23765 -#23769 := [monotonicity #23766]: #23768 -#23772 := [monotonicity #23769]: #23771 -#23775 := [monotonicity #23772]: #23774 -#23778 := [monotonicity #23775]: #23777 -#23781 := [monotonicity #23778]: #23780 -#23784 := [monotonicity #23781]: #23783 -#23787 := [monotonicity #23784]: #23786 -#23790 := [monotonicity #23787]: #23789 -#23793 := [monotonicity #23790]: #23792 -#23796 := [monotonicity #23793]: #23795 -#23799 := [monotonicity #23796]: #23798 -#23802 := [monotonicity #23799]: #23801 -#23805 := [monotonicity #23802]: #23804 -#23814 := [monotonicity #23805]: #23813 -#23817 := [monotonicity #23814]: #23816 -#23820 := [monotonicity #23817]: #23819 -#23823 := [monotonicity #23820]: #23822 -#23826 := [monotonicity #23823]: #23825 -#23829 := [monotonicity #23826]: #23828 -#23832 := [monotonicity #23829]: #23831 -#23835 := [monotonicity #23832]: #23834 -#23838 := [monotonicity #23835]: #23837 -#23841 := [monotonicity #23838]: #23840 -#23810 := (iff #22876 #23809) -#23807 := (iff #22875 #23806) -#23808 := [monotonicity #23805]: #23807 -#23811 := [monotonicity #23808]: #23810 -#23844 := [monotonicity #23811 #23841]: #23843 -#23847 := [monotonicity #23844]: #23846 -#23850 := [monotonicity #23847]: #23849 -#23853 := [monotonicity #23850]: #23852 -#23856 := [monotonicity #23853]: #23855 -#23859 := [monotonicity #23856]: #23858 -#23862 := [monotonicity #23859]: #23861 -#23865 := [monotonicity #23862]: #23864 -#23868 := [monotonicity #23865]: #23867 -#23871 := [monotonicity #23868]: #23870 -#23874 := [monotonicity #23871]: #23873 -#23877 := [monotonicity #23874]: #23876 -#23914 := [monotonicity #23877 #23911]: #23913 -#23917 := [monotonicity #23914]: #23916 -#23760 := (iff #23040 #23759) -#23757 := (iff #22779 #23754) -#23755 := (iff #22774 #22774) -#23756 := [refl]: #23755 -#23758 := [quant-intro #23756]: #23757 -#23761 := [monotonicity #23758]: #23760 -#23920 := [monotonicity #23761 #23917]: #23919 -#23923 := [monotonicity #23920]: #23922 -#23926 := [monotonicity #23923]: #23925 -#23929 := [monotonicity #23926]: #23928 -#23752 := (iff #23057 #23751) -#23749 := (iff #22768 #23746) -#23747 := (iff #22763 #22763) -#23748 := [refl]: #23747 -#23750 := [quant-intro #23748]: #23749 -#23753 := [monotonicity #23750]: #23752 -#23932 := [monotonicity #23753 #23929]: #23931 -#23935 := [monotonicity #23932]: #23934 -#23938 := [monotonicity #23935]: #23937 -#23941 := [monotonicity #23938]: #23940 -#23944 := [monotonicity #23941]: #23943 -#23947 := [monotonicity #23944]: #23946 -#23950 := [monotonicity #23947]: #23949 -#23953 := [monotonicity #23950]: #23952 -#23956 := [monotonicity #23953]: #23955 -#23959 := [monotonicity #23956]: #23958 -#23962 := [monotonicity #23959]: #23961 -#23965 := [monotonicity #23962]: #23964 -#23968 := [monotonicity #23965]: #23967 -#23971 := [monotonicity #23968]: #23970 -#23974 := [monotonicity #23971]: #23973 -#23977 := [monotonicity #23974]: #23976 -#23980 := [monotonicity #23977]: #23979 -#23983 := [monotonicity #23980]: #23982 -#23986 := [monotonicity #23983]: #23985 -#19548 := (and #19191 #19192) -#19551 := (not #19548) -#19554 := (or #19530 #19543 #19551) -#19557 := (not #19554) -#16489 := (and #3145 #4084 #13958 #15606) -#19214 := (not #16489) -#19217 := (forall (vars (?x782 int)) #19214) -#14858 := (and #4084 #15606) -#14857 := (not #14858) -#16475 := (or #13959 #13972 #14857) -#16480 := (forall (vars (?x781 int)) #16475) -#19221 := (and #16480 #19217) -#19563 := (or #19221 #19557) -#19571 := (and #12567 #12570 #12573 #12576 #13947 #13950 #14046 #19563) -#19392 := (and #19055 #19056) -#19395 := (not #19392) -#19398 := (or #19374 #19387 #19395) -#19401 := (not #19398) -#16376 := (or #14109 #14122 #14857) -#16381 := (forall (vars (?x785 int)) #16376) -#19071 := (not #14151) -#19081 := (and #19071 #16381) -#19407 := (or #19081 #19401) -#19412 := (and #14100 #19407) -#19415 := (or #14105 #19412) -#19423 := (and #3199 #14076 #14088 #14092 #14168 #16368 #19415) -#19428 := (or #19034 #19037 #19423) -#19454 := (and #3242 #3244 #3246 #12806 #12812 #13070 #13075 #13950 #14076 #19428) -#19459 := (or #19011 #19017 #19454) -#19465 := (and #12803 #12806 #19459) -#19470 := (or #19008 #19011 #19465) -#19476 := (and #13947 #13950 #14211 #19470) -#19434 := (and #12823 #12826 #13947 #13950 #14076 #14207 #19428) -#19481 := (or #19434 #19476) -#19487 := (and #12806 #12812 #13947 #13950 #19481) -#19492 := (or #19011 #19017 #19487) -#19498 := (and #12803 #12806 #19492) -#19503 := (or #19008 #19011 #19498) -#19509 := (and #13947 #13950 #14049 #19503) -#19576 := (or #19509 #19571) -#16293 := (or #14420 #14433 #14857) -#16298 := (forall (vars (?x774 int)) #16293) -#19582 := (and #3022 #3121 #3122 #3123 #3124 #3125 #3126 #12426 #12437 #12553 #13943 #13947 #13950 #14405 #14453 #14462 #14490 #16298 #16310 #16332 #16349 #19576) -#19587 := (or #13672 #13942 #19582) -#16279 := (or #13910 #13921 #14857) -#16284 := (forall (vars (?x773 int)) #16279) -#19590 := (and #16284 #19587) -#19303 := (and #18930 #18931) -#19306 := (not #19303) -#19312 := (or #18938 #18939 #19306) -#19317 := (not #19312) -#19593 := (or #19317 #19590) -#19596 := (and #13903 #19593) -#19599 := (or #13906 #19596) -#19605 := (and #3022 #3025 #3028 #3031 #12361 #12367 #19599) -#19610 := (or #18900 #18906 #19605) -#19616 := (and #12358 #12361 #19610) -#19621 := (or #18897 #18900 #19616) -#19624 := (and #12355 #19621) -#19627 := (or #13715 #19624) -#23118 := (iff #19627 #23117) -#23115 := (iff #19624 #23112) -#23107 := (and #12355 #23104) -#23113 := (iff #23107 #23112) -#23114 := [rewrite]: #23113 -#23108 := (iff #19624 #23107) -#23105 := (iff #19621 #23104) -#23102 := (iff #19616 #23099) -#23094 := (and #12358 #12361 #23091) -#23100 := (iff #23094 #23099) -#23101 := [rewrite]: #23100 -#23095 := (iff #19616 #23094) -#23092 := (iff #19610 #23091) -#23089 := (iff #19605 #23086) -#23081 := (and #3022 #3025 #3028 #3031 #12361 #12367 #23078) -#23087 := (iff #23081 #23086) -#23088 := [rewrite]: #23087 -#23082 := (iff #19605 #23081) -#23079 := (iff #19599 #23078) -#23076 := (iff #19596 #23073) -#23068 := (and #13903 #23065) -#23074 := (iff #23068 #23073) -#23075 := [rewrite]: #23074 -#23069 := (iff #19596 #23068) -#23066 := (iff #19593 #23065) -#23063 := (iff #19590 #23060) -#23054 := (and #22768 #23051) -#23061 := (iff #23054 #23060) -#23062 := [rewrite]: #23061 -#23055 := (iff #19590 #23054) -#23052 := (iff #19587 #23051) -#23049 := (iff #19582 #23046) -#23032 := (and #3022 #3121 #3122 #3123 #3124 #3125 #3126 #12426 #12437 #12553 #13943 #13947 #13950 #14405 #14453 #14462 #14490 #22779 #16310 #16332 #16349 #23029) -#23047 := (iff #23032 #23046) -#23048 := [rewrite]: #23047 -#23033 := (iff #19582 #23032) -#23030 := (iff #19576 #23029) -#23027 := (iff #19571 #23024) -#23019 := (and #12567 #12570 #12573 #12576 #13947 #13950 #14046 #23016) -#23025 := (iff #23019 #23024) -#23026 := [rewrite]: #23025 -#23020 := (iff #19571 #23019) -#23017 := (iff #19563 #23016) -#23014 := (iff #19557 #23013) -#23011 := (iff #19554 #23008) -#22994 := (or #22992 #22993) -#23005 := (or #19530 #19543 #22994) -#23009 := (iff #23005 #23008) -#23010 := [rewrite]: #23009 -#23006 := (iff #19554 #23005) -#23003 := (iff #19551 #22994) -#22995 := (not #22994) -#22998 := (not #22995) -#23001 := (iff #22998 #22994) -#23002 := [rewrite]: #23001 -#22999 := (iff #19551 #22998) -#22996 := (iff #19548 #22995) -#22997 := [rewrite]: #22996 -#23000 := [monotonicity #22997]: #22999 -#23004 := [trans #23000 #23002]: #23003 -#23007 := [monotonicity #23004]: #23006 -#23012 := [trans #23007 #23010]: #23011 -#23015 := [monotonicity #23012]: #23014 -#22990 := (iff #19221 #22987) -#22981 := (and #22963 #22978) -#22988 := (iff #22981 #22987) -#22989 := [rewrite]: #22988 -#22982 := (iff #19221 #22981) -#22979 := (iff #19217 #22978) -#22976 := (iff #19214 #22967) -#22968 := (not #22967) -#22971 := (not #22968) -#22974 := (iff #22971 #22967) -#22975 := [rewrite]: #22974 -#22972 := (iff #19214 #22971) -#22969 := (iff #16489 #22968) -#22970 := [rewrite]: #22969 -#22973 := [monotonicity #22970]: #22972 -#22977 := [trans #22973 #22975]: #22976 -#22980 := [quant-intro #22977]: #22979 -#22964 := (iff #16480 #22963) -#22961 := (iff #16475 #22958) -#20695 := (or #5113 #20064) -#22955 := (or #13959 #13972 #20695) -#22959 := (iff #22955 #22958) -#22960 := [rewrite]: #22959 -#22956 := (iff #16475 #22955) -#20704 := (iff #14857 #20695) -#20696 := (not #20695) -#20699 := (not #20696) -#20702 := (iff #20699 #20695) -#20703 := [rewrite]: #20702 -#20700 := (iff #14857 #20699) -#20697 := (iff #14858 #20696) -#20698 := [rewrite]: #20697 -#20701 := [monotonicity #20698]: #20700 -#20705 := [trans #20701 #20703]: #20704 -#22957 := [monotonicity #20705]: #22956 -#22962 := [trans #22957 #22960]: #22961 -#22965 := [quant-intro #22962]: #22964 -#22983 := [monotonicity #22965 #22980]: #22982 -#22991 := [trans #22983 #22989]: #22990 -#23018 := [monotonicity #22991 #23015]: #23017 -#23021 := [monotonicity #23018]: #23020 -#23028 := [trans #23021 #23026]: #23027 -#22953 := (iff #19509 #22950) -#22945 := (and #13947 #13950 #14049 #22942) -#22951 := (iff #22945 #22950) -#22952 := [rewrite]: #22951 -#22946 := (iff #19509 #22945) -#22943 := (iff #19503 #22942) -#22940 := (iff #19498 #22937) -#22932 := (and #12803 #12806 #22929) -#22938 := (iff #22932 #22937) -#22939 := [rewrite]: #22938 -#22933 := (iff #19498 #22932) -#22930 := (iff #19492 #22929) -#22927 := (iff #19487 #22924) -#22919 := (and #12806 #12812 #13947 #13950 #22916) -#22925 := (iff #22919 #22924) -#22926 := [rewrite]: #22925 -#22920 := (iff #19487 #22919) -#22917 := (iff #19481 #22916) -#22914 := (iff #19476 #22911) -#22906 := (and #13947 #13950 #14211 #22903) -#22912 := (iff #22906 #22911) -#22913 := [rewrite]: #22912 -#22907 := (iff #19476 #22906) -#22904 := (iff #19470 #22903) -#22901 := (iff #19465 #22898) -#22893 := (and #12803 #12806 #22890) -#22899 := (iff #22893 #22898) -#22900 := [rewrite]: #22899 -#22894 := (iff #19465 #22893) -#22891 := (iff #19459 #22890) -#22888 := (iff #19454 #22885) -#22881 := (and #3242 #3244 #3246 #12806 #12812 #13070 #13075 #13950 #14076 #22867) -#22886 := (iff #22881 #22885) -#22887 := [rewrite]: #22886 -#22882 := (iff #19454 #22881) -#22868 := (iff #19428 #22867) -#22865 := (iff #19423 #22862) -#22855 := (and #3199 #14076 #14088 #14092 #14168 #16368 #22852) -#22863 := (iff #22855 #22862) -#22864 := [rewrite]: #22863 -#22856 := (iff #19423 #22855) -#22853 := (iff #19415 #22852) -#22850 := (iff #19412 #22847) -#22842 := (and #14100 #22839) -#22848 := (iff #22842 #22847) -#22849 := [rewrite]: #22848 -#22843 := (iff #19412 #22842) -#22840 := (iff #19407 #22839) -#22837 := (iff #19401 #22836) -#22834 := (iff #19398 #22831) -#22817 := (or #22815 #22816) -#22828 := (or #19374 #19387 #22817) -#22832 := (iff #22828 #22831) -#22833 := [rewrite]: #22832 -#22829 := (iff #19398 #22828) -#22826 := (iff #19395 #22817) -#22818 := (not #22817) -#22821 := (not #22818) -#22824 := (iff #22821 #22817) -#22825 := [rewrite]: #22824 -#22822 := (iff #19395 #22821) -#22819 := (iff #19392 #22818) -#22820 := [rewrite]: #22819 -#22823 := [monotonicity #22820]: #22822 -#22827 := [trans #22823 #22825]: #22826 -#22830 := [monotonicity #22827]: #22829 -#22835 := [trans #22830 #22833]: #22834 -#22838 := [monotonicity #22835]: #22837 -#22813 := (iff #19081 #22810) -#22805 := (and #22783 #22802) -#22811 := (iff #22805 #22810) -#22812 := [rewrite]: #22811 -#22806 := (iff #19081 #22805) -#22803 := (iff #16381 #22802) -#22800 := (iff #16376 #22797) -#22794 := (or #14109 #14122 #20695) -#22798 := (iff #22794 #22797) -#22799 := [rewrite]: #22798 -#22795 := (iff #16376 #22794) -#22796 := [monotonicity #20705]: #22795 -#22801 := [trans #22796 #22799]: #22800 -#22804 := [quant-intro #22801]: #22803 -#22792 := (iff #19071 #22783) -#22787 := (not #22784) -#22790 := (iff #22787 #22783) -#22791 := [rewrite]: #22790 -#22788 := (iff #19071 #22787) -#22785 := (iff #14151 #22784) -#22786 := [rewrite]: #22785 -#22789 := [monotonicity #22786]: #22788 -#22793 := [trans #22789 #22791]: #22792 -#22807 := [monotonicity #22793 #22804]: #22806 -#22814 := [trans #22807 #22812]: #22813 -#22841 := [monotonicity #22814 #22838]: #22840 -#22844 := [monotonicity #22841]: #22843 -#22851 := [trans #22844 #22849]: #22850 -#22854 := [monotonicity #22851]: #22853 -#22857 := [monotonicity #22854]: #22856 -#22866 := [trans #22857 #22864]: #22865 -#22869 := [monotonicity #22866]: #22868 -#22883 := [monotonicity #22869]: #22882 -#22889 := [trans #22883 #22887]: #22888 -#22892 := [monotonicity #22889]: #22891 -#22895 := [monotonicity #22892]: #22894 -#22902 := [trans #22895 #22900]: #22901 -#22905 := [monotonicity #22902]: #22904 -#22908 := [monotonicity #22905]: #22907 -#22915 := [trans #22908 #22913]: #22914 -#22879 := (iff #19434 #22876) -#22870 := (and #12823 #12826 #13947 #13950 #14076 #14207 #22867) -#22877 := (iff #22870 #22876) -#22878 := [rewrite]: #22877 -#22871 := (iff #19434 #22870) -#22872 := [monotonicity #22869]: #22871 -#22880 := [trans #22872 #22878]: #22879 -#22918 := [monotonicity #22880 #22915]: #22917 -#22921 := [monotonicity #22918]: #22920 -#22928 := [trans #22921 #22926]: #22927 -#22931 := [monotonicity #22928]: #22930 -#22934 := [monotonicity #22931]: #22933 -#22941 := [trans #22934 #22939]: #22940 -#22944 := [monotonicity #22941]: #22943 -#22947 := [monotonicity #22944]: #22946 -#22954 := [trans #22947 #22952]: #22953 -#23031 := [monotonicity #22954 #23028]: #23030 -#22780 := (iff #16298 #22779) -#22777 := (iff #16293 #22774) -#22771 := (or #14420 #14433 #20695) -#22775 := (iff #22771 #22774) -#22776 := [rewrite]: #22775 -#22772 := (iff #16293 #22771) -#22773 := [monotonicity #20705]: #22772 -#22778 := [trans #22773 #22776]: #22777 -#22781 := [quant-intro #22778]: #22780 -#23034 := [monotonicity #22781 #23031]: #23033 -#23050 := [trans #23034 #23048]: #23049 -#23053 := [monotonicity #23050]: #23052 -#22769 := (iff #16284 #22768) -#22766 := (iff #16279 #22763) -#22760 := (or #13910 #13921 #20695) -#22764 := (iff #22760 #22763) -#22765 := [rewrite]: #22764 -#22761 := (iff #16279 #22760) -#22762 := [monotonicity #20705]: #22761 -#22767 := [trans #22762 #22765]: #22766 -#22770 := [quant-intro #22767]: #22769 -#23056 := [monotonicity #22770 #23053]: #23055 -#23064 := [trans #23056 #23062]: #23063 -#22758 := (iff #19317 #22757) -#22755 := (iff #19312 #22752) -#22738 := (or #22736 #22737) -#22749 := (or #18938 #18939 #22738) -#22753 := (iff #22749 #22752) -#22754 := [rewrite]: #22753 -#22750 := (iff #19312 #22749) -#22747 := (iff #19306 #22738) -#22739 := (not #22738) -#22742 := (not #22739) -#22745 := (iff #22742 #22738) -#22746 := [rewrite]: #22745 -#22743 := (iff #19306 #22742) -#22740 := (iff #19303 #22739) -#22741 := [rewrite]: #22740 -#22744 := [monotonicity #22741]: #22743 -#22748 := [trans #22744 #22746]: #22747 -#22751 := [monotonicity #22748]: #22750 -#22756 := [trans #22751 #22754]: #22755 -#22759 := [monotonicity #22756]: #22758 -#23067 := [monotonicity #22759 #23064]: #23066 -#23070 := [monotonicity #23067]: #23069 -#23077 := [trans #23070 #23075]: #23076 -#23080 := [monotonicity #23077]: #23079 -#23083 := [monotonicity #23080]: #23082 -#23090 := [trans #23083 #23088]: #23089 -#23093 := [monotonicity #23090]: #23092 -#23096 := [monotonicity #23093]: #23095 -#23103 := [trans #23096 #23101]: #23102 -#23106 := [monotonicity #23103]: #23105 -#23109 := [monotonicity #23106]: #23108 -#23116 := [trans #23109 #23114]: #23115 -#23119 := [monotonicity #23116]: #23118 -#19193 := (and #19192 #19191) -#19194 := (not #19193) -#19197 := (+ #19196 #13970) -#19198 := (<= #19197 0::int) -#19199 := (+ ?x781!15 #13873) -#19200 := (>= #19199 0::int) -#19201 := (or #19200 #19198 #19194) -#19202 := (not #19201) -#19225 := (or #19202 #19221) -#18978 := (not #13955) -#19185 := (not #12644) -#19182 := (not #12653) -#19179 := (not #12662) -#19176 := (not #12671) -#19229 := (and #19176 #19179 #19182 #19185 #18978 #14352 #19225) -#16407 := (and #14088 #16368) -#16412 := (not #16407) -#19097 := (not #16412) -#19057 := (and #19056 #19055) -#19058 := (not #19057) -#19061 := (+ #19060 #14120) -#19062 := (<= #19061 0::int) -#19063 := (+ ?x785!14 #14101) -#19064 := (>= #19063 0::int) -#19065 := (or #19064 #19062 #19058) -#19066 := (not #19065) -#19085 := (or #19066 #19081) -#19051 := (not #14105) -#19089 := (and #19051 #19085) -#19093 := (or #14105 #19089) -#19046 := (not #14172) -#19043 := (not #14097) -#19040 := (not #12942) -#19100 := (and #19040 #19043 #19046 #19093 #19097) -#19104 := (or #19034 #19037 #19100) -#19029 := (not #14081) -#19129 := (not #14243) -#19126 := (not #13081) -#19123 := (not #13133) -#19020 := (not #13142) -#19120 := (not #13090) -#19117 := (not #13115) -#19114 := (not #13124) -#19132 := (and #19114 #19117 #19120 #19020 #19123 #19126 #19129 #19029 #19104) -#19136 := (or #19011 #19017 #19132) -#19014 := (not #13159) -#19140 := (and #19014 #19136) -#19144 := (or #19008 #19011 #19140) -#19148 := (and #18978 #14211 #19144) -#19026 := (not #12993) -#19023 := (not #13002) -#19108 := (and #19023 #19026 #18978 #19029 #14293 #19104) -#19152 := (or #19108 #19148) -#19156 := (and #19020 #18978 #19152) -#19160 := (or #19011 #19017 #19156) -#19164 := (and #19014 #19160) -#19168 := (or #19008 #19011 #19164) -#19172 := (and #18978 #14049 #19168) -#19233 := (or #19172 #19229) -#16357 := (and #14490 #16349) -#16362 := (not #16357) -#19003 := (not #16362) -#16337 := (and #13947 #16332) -#16340 := (not #16337) -#19000 := (not #16340) -#16318 := (and #14462 #16310) -#16323 := (not #16318) -#18997 := (not #16323) -#18987 := (not #14507) -#18984 := (not #14456) -#18981 := (not #14416) -#18975 := (not #15709) -#18972 := (not #14399) -#18969 := (not #13331) -#18966 := (not #13340) -#18963 := (not #13349) -#18960 := (not #13358) -#18957 := (not #13367) -#19237 := (and #18957 #18960 #18963 #18966 #18969 #18972 #18975 #18978 #18981 #18984 #18987 #16298 #18997 #19000 #19003 #19233) -#19241 := (or #13672 #14664 #19237) -#19245 := (and #16284 #19241) -#18932 := (and #18931 #18930) -#18933 := (not #18932) -#18940 := (or #18939 #18938 #18933) -#18941 := (not #18940) -#19249 := (or #18941 #19245) -#18926 := (not #13906) -#19253 := (and #18926 #19249) -#19257 := (or #13906 #19253) -#18921 := (not #13681) -#18918 := (not #13645) -#18915 := (not #13654) -#18912 := (not #13663) -#18909 := (not #13672) -#19261 := (and #18909 #18912 #18915 #18918 #18921 #19257) -#19265 := (or #18900 #18906 #19261) -#18903 := (not #13698) -#19269 := (and #18903 #19265) -#19273 := (or #18897 #18900 #19269) -#18894 := (not #13715) -#19277 := (and #18894 #19273) -#19281 := (or #13715 #19277) -#19628 := (iff #19281 #19627) -#19625 := (iff #19277 #19624) -#19622 := (iff #19273 #19621) -#19619 := (iff #19269 #19616) -#19613 := (and #12364 #19610) -#19617 := (iff #19613 #19616) -#19618 := [rewrite]: #19617 -#19614 := (iff #19269 #19613) -#19611 := (iff #19265 #19610) -#19608 := (iff #19261 #19605) -#19602 := (and #3022 #3025 #3028 #3031 #12373 #19599) -#19606 := (iff #19602 #19605) -#19607 := [rewrite]: #19606 -#19603 := (iff #19261 #19602) -#19600 := (iff #19257 #19599) -#19597 := (iff #19253 #19596) -#19594 := (iff #19249 #19593) -#19591 := (iff #19245 #19590) -#19588 := (iff #19241 #19587) -#19585 := (iff #19237 #19582) -#19579 := (and #3121 #3122 #3123 #3124 #3127 #12437 #12553 #13952 #14411 #14453 #14502 #16298 #16318 #16337 #16357 #19576) -#19583 := (iff #19579 #19582) -#19584 := [rewrite]: #19583 -#19580 := (iff #19237 #19579) -#19577 := (iff #19233 #19576) -#19574 := (iff #19229 #19571) -#19568 := (and #12567 #12570 #12573 #12576 #13952 #14046 #19563) -#19572 := (iff #19568 #19571) -#19573 := [rewrite]: #19572 -#19569 := (iff #19229 #19568) -#19566 := (iff #19225 #19563) -#19560 := (or #19557 #19221) -#19564 := (iff #19560 #19563) -#19565 := [rewrite]: #19564 -#19561 := (iff #19225 #19560) -#19558 := (iff #19202 #19557) -#19555 := (iff #19201 #19554) -#19552 := (iff #19194 #19551) -#19549 := (iff #19193 #19548) -#19550 := [rewrite]: #19549 -#19553 := [monotonicity #19550]: #19552 -#19546 := (iff #19198 #19543) -#19535 := (+ #13970 #19196) -#19538 := (<= #19535 0::int) -#19544 := (iff #19538 #19543) -#19545 := [rewrite]: #19544 -#19539 := (iff #19198 #19538) -#19536 := (= #19197 #19535) -#19537 := [rewrite]: #19536 -#19540 := [monotonicity #19537]: #19539 -#19547 := [trans #19540 #19545]: #19546 -#19533 := (iff #19200 #19530) -#19522 := (+ #13873 ?x781!15) -#19525 := (>= #19522 0::int) -#19531 := (iff #19525 #19530) -#19532 := [rewrite]: #19531 -#19526 := (iff #19200 #19525) -#19523 := (= #19199 #19522) -#19524 := [rewrite]: #19523 -#19527 := [monotonicity #19524]: #19526 -#19534 := [trans #19527 #19532]: #19533 -#19556 := [monotonicity #19534 #19547 #19553]: #19555 -#19559 := [monotonicity #19556]: #19558 -#19562 := [monotonicity #19559]: #19561 -#19567 := [trans #19562 #19565]: #19566 -#19334 := (iff #18978 #13952) -#19335 := [rewrite]: #19334 -#19520 := (iff #19185 #12576) -#19521 := [rewrite]: #19520 -#19518 := (iff #19182 #12573) -#19519 := [rewrite]: #19518 -#19516 := (iff #19179 #12570) -#19517 := [rewrite]: #19516 -#19514 := (iff #19176 #12567) -#19515 := [rewrite]: #19514 -#19570 := [monotonicity #19515 #19517 #19519 #19521 #19335 #14356 #19567]: #19569 -#19575 := [trans #19570 #19573]: #19574 -#19512 := (iff #19172 #19509) -#19506 := (and #13952 #14049 #19503) -#19510 := (iff #19506 #19509) -#19511 := [rewrite]: #19510 -#19507 := (iff #19172 #19506) -#19504 := (iff #19168 #19503) -#19501 := (iff #19164 #19498) -#19495 := (and #12809 #19492) -#19499 := (iff #19495 #19498) -#19500 := [rewrite]: #19499 -#19496 := (iff #19164 #19495) -#19493 := (iff #19160 #19492) -#19490 := (iff #19156 #19487) -#19484 := (and #12818 #13952 #19481) -#19488 := (iff #19484 #19487) -#19489 := [rewrite]: #19488 -#19485 := (iff #19156 #19484) -#19482 := (iff #19152 #19481) -#19479 := (iff #19148 #19476) -#19473 := (and #13952 #14211 #19470) -#19477 := (iff #19473 #19476) -#19478 := [rewrite]: #19477 -#19474 := (iff #19148 #19473) -#19471 := (iff #19144 #19470) -#19468 := (iff #19140 #19465) -#19462 := (and #12809 #19459) -#19466 := (iff #19462 #19465) -#19467 := [rewrite]: #19466 -#19463 := (iff #19140 #19462) -#19460 := (iff #19136 #19459) -#19457 := (iff #19132 #19454) -#19451 := (and #3242 #3244 #3246 #12818 #13070 #13075 #13950 #14078 #19428) -#19455 := (iff #19451 #19454) -#19456 := [rewrite]: #19455 -#19452 := (iff #19132 #19451) -#19429 := (iff #19104 #19428) -#19426 := (iff #19100 #19423) -#19420 := (and #3199 #14094 #14168 #19415 #16407) -#19424 := (iff #19420 #19423) -#19425 := [rewrite]: #19424 -#19421 := (iff #19100 #19420) -#19418 := (iff #19097 #16407) -#19419 := [rewrite]: #19418 -#19416 := (iff #19093 #19415) -#19413 := (iff #19089 #19412) -#19410 := (iff #19085 #19407) -#19404 := (or #19401 #19081) -#19408 := (iff #19404 #19407) -#19409 := [rewrite]: #19408 -#19405 := (iff #19085 #19404) -#19402 := (iff #19066 #19401) -#19399 := (iff #19065 #19398) -#19396 := (iff #19058 #19395) -#19393 := (iff #19057 #19392) -#19394 := [rewrite]: #19393 -#19397 := [monotonicity #19394]: #19396 -#19390 := (iff #19062 #19387) -#19379 := (+ #14120 #19060) -#19382 := (<= #19379 0::int) -#19388 := (iff #19382 #19387) -#19389 := [rewrite]: #19388 -#19383 := (iff #19062 #19382) -#19380 := (= #19061 #19379) -#19381 := [rewrite]: #19380 -#19384 := [monotonicity #19381]: #19383 -#19391 := [trans #19384 #19389]: #19390 -#19377 := (iff #19064 #19374) -#19366 := (+ #14101 ?x785!14) -#19369 := (>= #19366 0::int) -#19375 := (iff #19369 #19374) -#19376 := [rewrite]: #19375 -#19370 := (iff #19064 #19369) -#19367 := (= #19063 #19366) -#19368 := [rewrite]: #19367 -#19371 := [monotonicity #19368]: #19370 -#19378 := [trans #19371 #19376]: #19377 -#19400 := [monotonicity #19378 #19391 #19397]: #19399 -#19403 := [monotonicity #19400]: #19402 -#19406 := [monotonicity #19403]: #19405 -#19411 := [trans #19406 #19409]: #19410 -#19364 := (iff #19051 #14100) -#19365 := [rewrite]: #19364 -#19414 := [monotonicity #19365 #19411]: #19413 -#19417 := [monotonicity #19414]: #19416 -#19362 := (iff #19046 #14168) -#19363 := [rewrite]: #19362 -#19360 := (iff #19043 #14094) -#19361 := [rewrite]: #19360 -#19358 := (iff #19040 #3199) -#19359 := [rewrite]: #19358 -#19422 := [monotonicity #19359 #19361 #19363 #19417 #19419]: #19421 -#19427 := [trans #19422 #19425]: #19426 -#19430 := [monotonicity #19427]: #19429 -#19356 := (iff #19029 #14078) -#19357 := [rewrite]: #19356 -#19449 := (iff #19129 #13950) -#19450 := [rewrite]: #19449 -#19447 := (iff #19126 #13075) -#19448 := [rewrite]: #19447 -#19445 := (iff #19123 #13070) -#19446 := [rewrite]: #19445 -#19350 := (iff #19020 #12818) -#19351 := [rewrite]: #19350 -#19443 := (iff #19120 #3246) -#19444 := [rewrite]: #19443 -#19441 := (iff #19117 #3244) -#19442 := [rewrite]: #19441 -#19439 := (iff #19114 #3242) -#19440 := [rewrite]: #19439 -#19453 := [monotonicity #19440 #19442 #19444 #19351 #19446 #19448 #19450 #19357 #19430]: #19452 -#19458 := [trans #19453 #19456]: #19457 -#19461 := [monotonicity #19458]: #19460 -#19348 := (iff #19014 #12809) -#19349 := [rewrite]: #19348 -#19464 := [monotonicity #19349 #19461]: #19463 -#19469 := [trans #19464 #19467]: #19468 -#19472 := [monotonicity #19469]: #19471 -#19475 := [monotonicity #19335 #19472]: #19474 -#19480 := [trans #19475 #19478]: #19479 -#19437 := (iff #19108 #19434) -#19431 := (and #12823 #12826 #13952 #14078 #14207 #19428) -#19435 := (iff #19431 #19434) -#19436 := [rewrite]: #19435 -#19432 := (iff #19108 #19431) -#19354 := (iff #19026 #12826) -#19355 := [rewrite]: #19354 -#19352 := (iff #19023 #12823) -#19353 := [rewrite]: #19352 -#19433 := [monotonicity #19353 #19355 #19335 #19357 #14297 #19430]: #19432 -#19438 := [trans #19433 #19436]: #19437 -#19483 := [monotonicity #19438 #19480]: #19482 -#19486 := [monotonicity #19351 #19335 #19483]: #19485 -#19491 := [trans #19486 #19489]: #19490 -#19494 := [monotonicity #19491]: #19493 -#19497 := [monotonicity #19349 #19494]: #19496 -#19502 := [trans #19497 #19500]: #19501 -#19505 := [monotonicity #19502]: #19504 -#19508 := [monotonicity #19335 #19505]: #19507 -#19513 := [trans #19508 #19511]: #19512 -#19578 := [monotonicity #19513 #19575]: #19577 -#19346 := (iff #19003 #16357) -#19347 := [rewrite]: #19346 -#19344 := (iff #19000 #16337) -#19345 := [rewrite]: #19344 -#19342 := (iff #18997 #16318) -#19343 := [rewrite]: #19342 -#19340 := (iff #18987 #14502) -#19341 := [rewrite]: #19340 -#19338 := (iff #18984 #14453) -#19339 := [rewrite]: #19338 -#19336 := (iff #18981 #14411) -#19337 := [rewrite]: #19336 -#19332 := (iff #18975 #12553) -#19333 := [rewrite]: #19332 -#19330 := (iff #18972 #12437) -#19331 := [rewrite]: #19330 -#19328 := (iff #18969 #3127) -#19329 := [rewrite]: #19328 -#19326 := (iff #18966 #3124) -#19327 := [rewrite]: #19326 -#19324 := (iff #18963 #3123) -#19325 := [rewrite]: #19324 -#19322 := (iff #18960 #3122) -#19323 := [rewrite]: #19322 -#19320 := (iff #18957 #3121) -#19321 := [rewrite]: #19320 -#19581 := [monotonicity #19321 #19323 #19325 #19327 #19329 #19331 #19333 #19335 #19337 #19339 #19341 #19343 #19345 #19347 #19578]: #19580 -#19586 := [trans #19581 #19584]: #19585 -#19589 := [monotonicity #14668 #19586]: #19588 -#19592 := [monotonicity #19589]: #19591 -#19318 := (iff #18941 #19317) -#19315 := (iff #18940 #19312) -#19309 := (or #18939 #18938 #19306) -#19313 := (iff #19309 #19312) -#19314 := [rewrite]: #19313 -#19310 := (iff #18940 #19309) -#19307 := (iff #18933 #19306) -#19304 := (iff #18932 #19303) -#19305 := [rewrite]: #19304 -#19308 := [monotonicity #19305]: #19307 -#19311 := [monotonicity #19308]: #19310 -#19316 := [trans #19311 #19314]: #19315 -#19319 := [monotonicity #19316]: #19318 -#19595 := [monotonicity #19319 #19592]: #19594 -#19301 := (iff #18926 #13903) -#19302 := [rewrite]: #19301 -#19598 := [monotonicity #19302 #19595]: #19597 -#19601 := [monotonicity #19598]: #19600 -#19299 := (iff #18921 #12373) -#19300 := [rewrite]: #19299 -#19297 := (iff #18918 #3031) -#19298 := [rewrite]: #19297 -#19295 := (iff #18915 #3028) -#19296 := [rewrite]: #19295 -#19293 := (iff #18912 #3025) -#19294 := [rewrite]: #19293 -#19291 := (iff #18909 #3022) -#19292 := [rewrite]: #19291 -#19604 := [monotonicity #19292 #19294 #19296 #19298 #19300 #19601]: #19603 -#19609 := [trans #19604 #19607]: #19608 -#19612 := [monotonicity #19609]: #19611 -#19289 := (iff #18903 #12364) -#19290 := [rewrite]: #19289 -#19615 := [monotonicity #19290 #19612]: #19614 -#19620 := [trans #19615 #19618]: #19619 -#19623 := [monotonicity #19620]: #19622 -#19287 := (iff #18894 #12355) -#19288 := [rewrite]: #19287 -#19626 := [monotonicity #19288 #19623]: #19625 -#19629 := [monotonicity #19626]: #19628 -#16494 := (exists (vars (?x782 int)) #16489) -#16483 := (not #16480) -#16497 := (or #16483 #16494) -#16500 := (and #16480 #16497) -#16506 := (or #12671 #12662 #12653 #12644 #13955 #14049 #16500) -#16384 := (not #16381) -#16390 := (or #14151 #16384) -#16395 := (and #16381 #16390) -#16398 := (or #14105 #16395) -#16401 := (and #14100 #16398) -#16418 := (or #12942 #14097 #14172 #16401 #16412) -#16426 := (and #14088 #16368 #16418) -#16439 := (or #13124 #13115 #13090 #13142 #13133 #13081 #14243 #14081 #16426) -#16442 := (and #12806 #12812 #16439) -#16445 := (or #13159 #16442) -#16448 := (and #12803 #12806 #16445) -#16451 := (or #13955 #14207 #16448) -#16434 := (or #13002 #12993 #13955 #14081 #14211 #16426) -#16454 := (and #16434 #16451) -#16457 := (or #13142 #13955 #16454) -#16460 := (and #12806 #12812 #16457) -#16463 := (or #13159 #16460) -#16466 := (and #12803 #12806 #16463) -#16469 := (or #13955 #14046 #16466) -#16511 := (and #16469 #16506) -#16301 := (not #16298) -#16517 := (or #13367 #13358 #13349 #13340 #13331 #14399 #15709 #13955 #14416 #14456 #14507 #16301 #16323 #16340 #16362 #16511) -#16522 := (and #3022 #13943 #16517) -#16287 := (not #16284) -#16525 := (or #16287 #16522) -#16528 := (and #16284 #16525) -#16531 := (or #13906 #16528) -#16534 := (and #13903 #16531) -#16537 := (or #13672 #13663 #13654 #13645 #13681 #16534) -#16540 := (and #12361 #12367 #16537) -#16543 := (or #13698 #16540) -#16546 := (and #12358 #12361 #16543) -#16549 := (or #13715 #16546) -#16552 := (and #12355 #16549) -#16555 := (not #16552) -#19282 := (~ #16555 #19281) -#19278 := (not #16549) -#19279 := (~ #19278 #19277) -#19274 := (not #16546) -#19275 := (~ #19274 #19273) -#19270 := (not #16543) -#19271 := (~ #19270 #19269) -#19266 := (not #16540) -#19267 := (~ #19266 #19265) -#19262 := (not #16537) -#19263 := (~ #19262 #19261) -#19258 := (not #16534) -#19259 := (~ #19258 #19257) -#19254 := (not #16531) -#19255 := (~ #19254 #19253) -#19250 := (not #16528) -#19251 := (~ #19250 #19249) -#19246 := (not #16525) -#19247 := (~ #19246 #19245) -#19242 := (not #16522) -#19243 := (~ #19242 #19241) -#19238 := (not #16517) -#19239 := (~ #19238 #19237) -#19234 := (not #16511) -#19235 := (~ #19234 #19233) -#19230 := (not #16506) -#19231 := (~ #19230 #19229) -#19226 := (not #16500) -#19227 := (~ #19226 #19225) -#19222 := (not #16497) -#19223 := (~ #19222 #19221) -#19218 := (not #16494) -#19219 := (~ #19218 #19217) -#19215 := (~ #19214 #19214) -#19216 := [refl]: #19215 -#19220 := [nnf-neg #19216]: #19219 -#19211 := (not #16483) -#19212 := (~ #19211 #16480) -#19209 := (~ #16480 #16480) -#19207 := (~ #16475 #16475) -#19208 := [refl]: #19207 -#19210 := [nnf-pos #19208]: #19209 -#19213 := [nnf-neg #19210]: #19212 -#19224 := [nnf-neg #19213 #19220]: #19223 -#19203 := (~ #16483 #19202) -#19204 := [sk]: #19203 -#19228 := [nnf-neg #19204 #19224]: #19227 -#19188 := (~ #14352 #14352) -#19189 := [refl]: #19188 -#18979 := (~ #18978 #18978) -#18980 := [refl]: #18979 -#19186 := (~ #19185 #19185) -#19187 := [refl]: #19186 -#19183 := (~ #19182 #19182) -#19184 := [refl]: #19183 -#19180 := (~ #19179 #19179) -#19181 := [refl]: #19180 -#19177 := (~ #19176 #19176) -#19178 := [refl]: #19177 -#19232 := [nnf-neg #19178 #19181 #19184 #19187 #18980 #19189 #19228]: #19231 -#19173 := (not #16469) -#19174 := (~ #19173 #19172) -#19169 := (not #16466) -#19170 := (~ #19169 #19168) -#19165 := (not #16463) -#19166 := (~ #19165 #19164) -#19161 := (not #16460) -#19162 := (~ #19161 #19160) -#19157 := (not #16457) -#19158 := (~ #19157 #19156) -#19153 := (not #16454) -#19154 := (~ #19153 #19152) -#19149 := (not #16451) -#19150 := (~ #19149 #19148) -#19145 := (not #16448) -#19146 := (~ #19145 #19144) -#19141 := (not #16445) -#19142 := (~ #19141 #19140) -#19137 := (not #16442) -#19138 := (~ #19137 #19136) -#19133 := (not #16439) -#19134 := (~ #19133 #19132) -#19105 := (not #16426) -#19106 := (~ #19105 #19104) -#19101 := (not #16418) -#19102 := (~ #19101 #19100) -#19098 := (~ #19097 #19097) -#19099 := [refl]: #19098 -#19094 := (not #16401) -#19095 := (~ #19094 #19093) -#19090 := (not #16398) -#19091 := (~ #19090 #19089) -#19086 := (not #16395) -#19087 := (~ #19086 #19085) -#19082 := (not #16390) -#19083 := (~ #19082 #19081) -#19078 := (not #16384) -#19079 := (~ #19078 #16381) -#19076 := (~ #16381 #16381) -#19074 := (~ #16376 #16376) -#19075 := [refl]: #19074 -#19077 := [nnf-pos #19075]: #19076 -#19080 := [nnf-neg #19077]: #19079 -#19072 := (~ #19071 #19071) -#19073 := [refl]: #19072 -#19084 := [nnf-neg #19073 #19080]: #19083 -#19067 := (~ #16384 #19066) -#19068 := [sk]: #19067 -#19088 := [nnf-neg #19068 #19084]: #19087 -#19052 := (~ #19051 #19051) -#19053 := [refl]: #19052 -#19092 := [nnf-neg #19053 #19088]: #19091 -#19049 := (~ #14105 #14105) -#19050 := [refl]: #19049 -#19096 := [nnf-neg #19050 #19092]: #19095 -#19047 := (~ #19046 #19046) -#19048 := [refl]: #19047 -#19044 := (~ #19043 #19043) -#19045 := [refl]: #19044 -#19041 := (~ #19040 #19040) -#19042 := [refl]: #19041 -#19103 := [nnf-neg #19042 #19045 #19048 #19096 #19099]: #19102 -#19038 := (~ #19037 #19037) -#19039 := [refl]: #19038 -#19035 := (~ #19034 #19034) -#19036 := [refl]: #19035 -#19107 := [nnf-neg #19036 #19039 #19103]: #19106 -#19030 := (~ #19029 #19029) -#19031 := [refl]: #19030 -#19130 := (~ #19129 #19129) -#19131 := [refl]: #19130 -#19127 := (~ #19126 #19126) -#19128 := [refl]: #19127 -#19124 := (~ #19123 #19123) -#19125 := [refl]: #19124 -#19021 := (~ #19020 #19020) -#19022 := [refl]: #19021 -#19121 := (~ #19120 #19120) -#19122 := [refl]: #19121 -#19118 := (~ #19117 #19117) -#19119 := [refl]: #19118 -#19115 := (~ #19114 #19114) -#19116 := [refl]: #19115 -#19135 := [nnf-neg #19116 #19119 #19122 #19022 #19125 #19128 #19131 #19031 #19107]: #19134 -#19018 := (~ #19017 #19017) -#19019 := [refl]: #19018 -#19012 := (~ #19011 #19011) -#19013 := [refl]: #19012 -#19139 := [nnf-neg #19013 #19019 #19135]: #19138 -#19015 := (~ #19014 #19014) -#19016 := [refl]: #19015 -#19143 := [nnf-neg #19016 #19139]: #19142 -#19009 := (~ #19008 #19008) -#19010 := [refl]: #19009 -#19147 := [nnf-neg #19010 #19013 #19143]: #19146 -#19112 := (~ #14211 #14211) -#19113 := [refl]: #19112 -#19151 := [nnf-neg #18980 #19113 #19147]: #19150 -#19109 := (not #16434) -#19110 := (~ #19109 #19108) -#19032 := (~ #14293 #14293) -#19033 := [refl]: #19032 -#19027 := (~ #19026 #19026) -#19028 := [refl]: #19027 -#19024 := (~ #19023 #19023) -#19025 := [refl]: #19024 -#19111 := [nnf-neg #19025 #19028 #18980 #19031 #19033 #19107]: #19110 -#19155 := [nnf-neg #19111 #19151]: #19154 -#19159 := [nnf-neg #19022 #18980 #19155]: #19158 -#19163 := [nnf-neg #19013 #19019 #19159]: #19162 -#19167 := [nnf-neg #19016 #19163]: #19166 -#19171 := [nnf-neg #19010 #19013 #19167]: #19170 -#19006 := (~ #14049 #14049) -#19007 := [refl]: #19006 -#19175 := [nnf-neg #18980 #19007 #19171]: #19174 -#19236 := [nnf-neg #19175 #19232]: #19235 -#19004 := (~ #19003 #19003) -#19005 := [refl]: #19004 -#19001 := (~ #19000 #19000) -#19002 := [refl]: #19001 -#18998 := (~ #18997 #18997) -#18999 := [refl]: #18998 -#18994 := (not #16301) -#18995 := (~ #18994 #16298) -#18992 := (~ #16298 #16298) -#18990 := (~ #16293 #16293) -#18991 := [refl]: #18990 -#18993 := [nnf-pos #18991]: #18992 -#18996 := [nnf-neg #18993]: #18995 -#18988 := (~ #18987 #18987) -#18989 := [refl]: #18988 -#18985 := (~ #18984 #18984) -#18986 := [refl]: #18985 -#18982 := (~ #18981 #18981) -#18983 := [refl]: #18982 -#18976 := (~ #18975 #18975) -#18977 := [refl]: #18976 -#18973 := (~ #18972 #18972) -#18974 := [refl]: #18973 -#18970 := (~ #18969 #18969) -#18971 := [refl]: #18970 -#18967 := (~ #18966 #18966) -#18968 := [refl]: #18967 -#18964 := (~ #18963 #18963) -#18965 := [refl]: #18964 -#18961 := (~ #18960 #18960) -#18962 := [refl]: #18961 -#18958 := (~ #18957 #18957) -#18959 := [refl]: #18958 -#19240 := [nnf-neg #18959 #18962 #18965 #18968 #18971 #18974 #18977 #18980 #18983 #18986 #18989 #18996 #18999 #19002 #19005 #19236]: #19239 -#18955 := (~ #14664 #14664) -#18956 := [refl]: #18955 -#18953 := (~ #13672 #13672) -#18954 := [refl]: #18953 -#19244 := [nnf-neg #18954 #18956 #19240]: #19243 -#18950 := (not #16287) -#18951 := (~ #18950 #16284) -#18948 := (~ #16284 #16284) -#18946 := (~ #16279 #16279) -#18947 := [refl]: #18946 -#18949 := [nnf-pos #18947]: #18948 -#18952 := [nnf-neg #18949]: #18951 -#19248 := [nnf-neg #18952 #19244]: #19247 -#18942 := (~ #16287 #18941) -#18943 := [sk]: #18942 -#19252 := [nnf-neg #18943 #19248]: #19251 -#18927 := (~ #18926 #18926) -#18928 := [refl]: #18927 -#19256 := [nnf-neg #18928 #19252]: #19255 -#18924 := (~ #13906 #13906) -#18925 := [refl]: #18924 -#19260 := [nnf-neg #18925 #19256]: #19259 -#18922 := (~ #18921 #18921) -#18923 := [refl]: #18922 -#18919 := (~ #18918 #18918) -#18920 := [refl]: #18919 -#18916 := (~ #18915 #18915) -#18917 := [refl]: #18916 -#18913 := (~ #18912 #18912) -#18914 := [refl]: #18913 -#18910 := (~ #18909 #18909) -#18911 := [refl]: #18910 -#19264 := [nnf-neg #18911 #18914 #18917 #18920 #18923 #19260]: #19263 -#18907 := (~ #18906 #18906) -#18908 := [refl]: #18907 -#18901 := (~ #18900 #18900) -#18902 := [refl]: #18901 -#19268 := [nnf-neg #18902 #18908 #19264]: #19267 -#18904 := (~ #18903 #18903) -#18905 := [refl]: #18904 -#19272 := [nnf-neg #18905 #19268]: #19271 -#18898 := (~ #18897 #18897) -#18899 := [refl]: #18898 -#19276 := [nnf-neg #18899 #18902 #19272]: #19275 -#18895 := (~ #18894 #18894) -#18896 := [refl]: #18895 -#19280 := [nnf-neg #18896 #19276]: #19279 -#18892 := (~ #13715 #13715) -#18893 := [refl]: #18892 -#19283 := [nnf-neg #18893 #19280]: #19282 -#15734 := (or #12671 #12662 #12653 #12644 #13955 #14009 #14049) -#15742 := (and #14371 #15734) -#15750 := (or #13367 #13358 #13349 #13340 #13331 #14399 #15709 #13955 #14416 #14450 #14456 #14468 #14483 #14496 #14507 #15742) -#15755 := (and #3022 #13943 #15750) -#15758 := (or #13939 #15755) -#15761 := (and #13936 #15758) -#15764 := (or #13906 #15761) -#15767 := (and #13903 #15764) -#15770 := (or #13672 #13663 #13654 #13645 #13681 #15767) -#15773 := (and #12361 #12367 #15770) -#15776 := (or #13698 #15773) -#15779 := (and #12358 #12361 #15776) -#15782 := (or #13715 #15779) -#15785 := (and #12355 #15782) -#15788 := (not #15785) -#16556 := (iff #15788 #16555) -#16553 := (iff #15785 #16552) -#16550 := (iff #15782 #16549) -#16547 := (iff #15779 #16546) -#16544 := (iff #15776 #16543) -#16541 := (iff #15773 #16540) -#16538 := (iff #15770 #16537) -#16535 := (iff #15767 #16534) -#16532 := (iff #15764 #16531) -#16529 := (iff #15761 #16528) -#16526 := (iff #15758 #16525) -#16523 := (iff #15755 #16522) -#16520 := (iff #15750 #16517) -#16514 := (or #13367 #13358 #13349 #13340 #13331 #14399 #15709 #13955 #14416 #16301 #14456 #16323 #16340 #16362 #14507 #16511) -#16518 := (iff #16514 #16517) -#16519 := [rewrite]: #16518 -#16515 := (iff #15750 #16514) -#16512 := (iff #15742 #16511) -#16509 := (iff #15734 #16506) -#16503 := (or #12671 #12662 #12653 #12644 #13955 #16500 #14049) -#16507 := (iff #16503 #16506) -#16508 := [rewrite]: #16507 -#16504 := (iff #15734 #16503) -#16501 := (iff #14009 #16500) -#16498 := (iff #14006 #16497) -#16495 := (iff #14003 #16494) -#16492 := (iff #13998 #16489) -#16486 := (and #3145 #4084 #15606 #13958) -#16490 := (iff #16486 #16489) -#16491 := [rewrite]: #16490 -#16487 := (iff #13998 #16486) -#15605 := (iff #4419 #15606) -#15638 := -131073::int -#15614 := (+ -131073::int #161) -#15611 := (<= #15614 0::int) -#15607 := (iff #15611 #15606) -#15604 := [rewrite]: #15607 -#15608 := (iff #4419 #15611) -#15613 := (= #4418 #15614) -#15619 := (+ #161 -131073::int) -#15615 := (= #15619 #15614) -#15612 := [rewrite]: #15615 -#15616 := (= #4418 #15619) -#15637 := (= #4413 -131073::int) -#15643 := (* -1::int 131073::int) -#15639 := (= #15643 -131073::int) -#15636 := [rewrite]: #15639 -#15640 := (= #4413 #15643) -#7883 := (= uf_76 131073::int) -#1070 := 65536::int -#1313 := (+ 65536::int 65536::int) -#1318 := (+ #1313 1::int) -#1319 := (= uf_76 #1318) -#7884 := (iff #1319 #7883) -#7881 := (= #1318 131073::int) -#7874 := (+ 131072::int 1::int) -#7879 := (= #7874 131073::int) -#7880 := [rewrite]: #7879 -#7876 := (= #1318 #7874) -#7845 := (= #1313 131072::int) -#7846 := [rewrite]: #7845 -#7877 := [monotonicity #7846]: #7876 -#7882 := [trans #7877 #7880]: #7881 -#7885 := [monotonicity #7882]: #7884 -#7873 := [asserted]: #1319 -#7888 := [mp #7873 #7885]: #7883 -#15641 := [monotonicity #7888]: #15640 -#15634 := [trans #15641 #15636]: #15637 -#15617 := [monotonicity #15634]: #15616 -#15610 := [trans #15617 #15612]: #15613 -#15609 := [monotonicity #15610]: #15608 -#15602 := [trans #15609 #15604]: #15605 -#16488 := [monotonicity #15602]: #16487 -#16493 := [trans #16488 #16491]: #16492 -#16496 := [quant-intro #16493]: #16495 -#16484 := (iff #13989 #16483) -#16481 := (iff #13986 #16480) -#16478 := (iff #13981 #16475) -#16472 := (or #14857 #13959 #13972) -#16476 := (iff #16472 #16475) -#16477 := [rewrite]: #16476 -#16473 := (iff #13981 #16472) -#14854 := (iff #5739 #14857) -#14859 := (iff #5736 #14858) -#14856 := [monotonicity #15602]: #14859 -#14855 := [monotonicity #14856]: #14854 -#16474 := [monotonicity #14855]: #16473 -#16479 := [trans #16474 #16477]: #16478 -#16482 := [quant-intro #16479]: #16481 -#16485 := [monotonicity #16482]: #16484 -#16499 := [monotonicity #16485 #16496]: #16498 -#16502 := [monotonicity #16482 #16499]: #16501 -#16505 := [monotonicity #16502]: #16504 -#16510 := [trans #16505 #16508]: #16509 -#16470 := (iff #14371 #16469) -#16467 := (iff #14345 #16466) -#16464 := (iff #14339 #16463) -#16461 := (iff #14334 #16460) -#16458 := (iff #14326 #16457) -#16455 := (iff #14317 #16454) -#16452 := (iff #14312 #16451) -#16449 := (iff #14286 #16448) -#16446 := (iff #14280 #16445) -#16443 := (iff #14275 #16442) -#16440 := (iff #14267 #16439) -#16429 := (iff #14201 #16426) -#16423 := (and #16368 #14088 #16418) -#16427 := (iff #16423 #16426) -#16428 := [rewrite]: #16427 -#16424 := (iff #14201 #16423) -#16421 := (iff #14193 #16418) -#16415 := (or #12942 #14097 #16401 #14172 #16412) -#16419 := (iff #16415 #16418) -#16420 := [rewrite]: #16419 -#16416 := (iff #14193 #16415) -#16413 := (iff #14178 #16412) -#16410 := (iff #14175 #16407) -#16404 := (and #16368 #14088) -#16408 := (iff #16404 #16407) -#16409 := [rewrite]: #16408 -#16405 := (iff #14175 #16404) -#16371 := (iff #14084 #16368) -#16304 := (+ 131073::int #14044) -#16365 := (>= #16304 1::int) -#16369 := (iff #16365 #16368) -#16370 := [rewrite]: #16369 -#16366 := (iff #14084 #16365) -#16305 := (= #14085 #16304) -#16306 := [monotonicity #7888]: #16305 -#16367 := [monotonicity #16306]: #16366 -#16372 := [trans #16367 #16370]: #16371 -#16406 := [monotonicity #16372]: #16405 -#16411 := [trans #16406 #16409]: #16410 -#16414 := [monotonicity #16411]: #16413 -#16402 := (iff #14165 #16401) -#16399 := (iff #14162 #16398) -#16396 := (iff #14159 #16395) -#16393 := (iff #14156 #16390) -#16387 := (or #16384 #14151) -#16391 := (iff #16387 #16390) -#16392 := [rewrite]: #16391 -#16388 := (iff #14156 #16387) -#16385 := (iff #14139 #16384) -#16382 := (iff #14136 #16381) -#16379 := (iff #14131 #16376) -#16373 := (or #14857 #14109 #14122) -#16377 := (iff #16373 #16376) -#16378 := [rewrite]: #16377 -#16374 := (iff #14131 #16373) -#16375 := [monotonicity #14855]: #16374 -#16380 := [trans #16375 #16378]: #16379 -#16383 := [quant-intro #16380]: #16382 -#16386 := [monotonicity #16383]: #16385 -#16389 := [monotonicity #16386]: #16388 -#16394 := [trans #16389 #16392]: #16393 -#16397 := [monotonicity #16383 #16394]: #16396 -#16400 := [monotonicity #16397]: #16399 -#16403 := [monotonicity #16400]: #16402 -#16417 := [monotonicity #16403 #16414]: #16416 -#16422 := [trans #16417 #16420]: #16421 -#16425 := [monotonicity #16372 #16422]: #16424 -#16430 := [trans #16425 #16428]: #16429 -#16441 := [monotonicity #16430]: #16440 -#16444 := [monotonicity #16441]: #16443 -#16447 := [monotonicity #16444]: #16446 -#16450 := [monotonicity #16447]: #16449 -#16453 := [monotonicity #16450]: #16452 -#16437 := (iff #14238 #16434) -#16431 := (or #13002 #12993 #13955 #14081 #16426 #14211) -#16435 := (iff #16431 #16434) -#16436 := [rewrite]: #16435 -#16432 := (iff #14238 #16431) -#16433 := [monotonicity #16430]: #16432 -#16438 := [trans #16433 #16436]: #16437 -#16456 := [monotonicity #16438 #16453]: #16455 -#16459 := [monotonicity #16456]: #16458 -#16462 := [monotonicity #16459]: #16461 -#16465 := [monotonicity #16462]: #16464 -#16468 := [monotonicity #16465]: #16467 -#16471 := [monotonicity #16468]: #16470 -#16513 := [monotonicity #16471 #16510]: #16512 -#16363 := (iff #14496 #16362) -#16360 := (iff #14493 #16357) -#16354 := (and #16349 #14490) -#16358 := (iff #16354 #16357) -#16359 := [rewrite]: #16358 -#16355 := (iff #14493 #16354) -#16352 := (iff #14486 #16349) -#16343 := (+ 255::int #14431) -#16346 := (>= #16343 0::int) -#16350 := (iff #16346 #16349) -#16351 := [rewrite]: #16350 -#16347 := (iff #14486 #16346) -#16344 := (= #14487 #16343) -#1323 := (= uf_78 255::int) -#7887 := [asserted]: #1323 -#16345 := [monotonicity #7887]: #16344 -#16348 := [monotonicity #16345]: #16347 -#16353 := [trans #16348 #16351]: #16352 -#16356 := [monotonicity #16353]: #16355 -#16361 := [trans #16356 #16359]: #16360 -#16364 := [monotonicity #16361]: #16363 -#16341 := (iff #14483 #16340) -#16338 := (iff #14478 #16337) -#16335 := (iff #14471 #16332) -#16326 := (+ 131073::int #14402) -#16329 := (>= #16326 0::int) -#16333 := (iff #16329 #16332) -#16334 := [rewrite]: #16333 -#16330 := (iff #14471 #16329) -#16327 := (= #14472 #16326) -#16328 := [monotonicity #7888]: #16327 -#16331 := [monotonicity #16328]: #16330 -#16336 := [trans #16331 #16334]: #16335 -#16339 := [monotonicity #16336]: #16338 -#16342 := [monotonicity #16339]: #16341 -#16324 := (iff #14468 #16323) -#16321 := (iff #14465 #16318) -#16315 := (and #16310 #14462) -#16319 := (iff #16315 #16318) -#16320 := [rewrite]: #16319 -#16316 := (iff #14465 #16315) -#16313 := (iff #14459 #16310) -#16307 := (>= #16304 0::int) -#16311 := (iff #16307 #16310) -#16312 := [rewrite]: #16311 -#16308 := (iff #14459 #16307) -#16309 := [monotonicity #16306]: #16308 -#16314 := [trans #16309 #16312]: #16313 -#16317 := [monotonicity #16314]: #16316 -#16322 := [trans #16317 #16320]: #16321 -#16325 := [monotonicity #16322]: #16324 -#16302 := (iff #14450 #16301) -#16299 := (iff #14447 #16298) -#16296 := (iff #14442 #16293) -#16290 := (or #14857 #14420 #14433) -#16294 := (iff #16290 #16293) -#16295 := [rewrite]: #16294 -#16291 := (iff #14442 #16290) -#16292 := [monotonicity #14855]: #16291 -#16297 := [trans #16292 #16295]: #16296 -#16300 := [quant-intro #16297]: #16299 -#16303 := [monotonicity #16300]: #16302 -#16516 := [monotonicity #16303 #16325 #16342 #16364 #16513]: #16515 -#16521 := [trans #16516 #16519]: #16520 -#16524 := [monotonicity #16521]: #16523 -#16288 := (iff #13939 #16287) -#16285 := (iff #13936 #16284) -#16282 := (iff #13931 #16279) -#16276 := (or #14857 #13910 #13921) -#16280 := (iff #16276 #16279) -#16281 := [rewrite]: #16280 -#16277 := (iff #13931 #16276) -#16278 := [monotonicity #14855]: #16277 -#16283 := [trans #16278 #16281]: #16282 -#16286 := [quant-intro #16283]: #16285 -#16289 := [monotonicity #16286]: #16288 -#16527 := [monotonicity #16289 #16524]: #16526 -#16530 := [monotonicity #16286 #16527]: #16529 -#16533 := [monotonicity #16530]: #16532 -#16536 := [monotonicity #16533]: #16535 -#16539 := [monotonicity #16536]: #16538 -#16542 := [monotonicity #16539]: #16541 -#16545 := [monotonicity #16542]: #16544 -#16548 := [monotonicity #16545]: #16547 -#16551 := [monotonicity #16548]: #16550 -#16554 := [monotonicity #16551]: #16553 -#16557 := [monotonicity #16554]: #16556 -#14791 := (not #14643) -#15789 := (iff #14791 #15788) -#15786 := (iff #14643 #15785) -#15783 := (iff #14640 #15782) -#15780 := (iff #14635 #15779) -#15777 := (iff #14629 #15776) -#15774 := (iff #14624 #15773) -#15771 := (iff #14616 #15770) -#15768 := (iff #14595 #15767) -#15765 := (iff #14592 #15764) -#15762 := (iff #14589 #15761) -#15759 := (iff #14586 #15758) -#15756 := (iff #14581 #15755) -#15753 := (iff #14573 #15750) -#15747 := (or #13367 #13358 #13349 #13340 #13331 #14399 #15709 #13955 #15742 #14416 #14450 #14456 #14468 #14483 #14496 #14507) -#15751 := (iff #15747 #15750) -#15752 := [rewrite]: #15751 -#15748 := (iff #14573 #15747) -#15745 := (iff #14376 #15742) -#15739 := (and #15734 #14371) -#15743 := (iff #15739 #15742) -#15744 := [rewrite]: #15743 -#15740 := (iff #14376 #15739) -#15737 := (iff #14070 #15734) -#15719 := (or #12671 #12662 #12653 #12644 #13955 #14009) -#15731 := (or #13955 #15719 #14049) -#15735 := (iff #15731 #15734) -#15736 := [rewrite]: #15735 -#15732 := (iff #14070 #15731) -#15729 := (iff #14041 #15719) -#15724 := (and true #15719) -#15727 := (iff #15724 #15719) -#15728 := [rewrite]: #15727 -#15725 := (iff #14041 #15724) -#15722 := (iff #14036 #15719) -#15716 := (or false #12671 #12662 #12653 #12644 #13955 #14009) -#15720 := (iff #15716 #15719) -#15721 := [rewrite]: #15720 -#15717 := (iff #14036 #15716) -#15714 := (iff #12719 false) -#15712 := (iff #12719 #3294) -#15456 := (iff up_216 true) -#11194 := [asserted]: up_216 -#15457 := [iff-true #11194]: #15456 -#15713 := [monotonicity #15457]: #15712 -#15715 := [trans #15713 #13445]: #15714 -#15718 := [monotonicity #15715]: #15717 -#15723 := [trans #15718 #15721]: #15722 -#15726 := [monotonicity #15457 #15723]: #15725 -#15730 := [trans #15726 #15728]: #15729 -#15733 := [monotonicity #15730]: #15732 -#15738 := [trans #15733 #15736]: #15737 -#15741 := [monotonicity #15738]: #15740 -#15746 := [trans #15741 #15744]: #15745 -#15710 := (iff #13376 #15709) -#15707 := (iff #12556 #12553) -#15702 := (and true #12553) -#15705 := (iff #15702 #12553) -#15706 := [rewrite]: #15705 -#15703 := (iff #12556 #15702) -#15690 := (iff #12332 true) -#15691 := [iff-true #14784]: #15690 -#15704 := [monotonicity #15691]: #15703 -#15708 := [trans #15704 #15706]: #15707 -#15711 := [monotonicity #15708]: #15710 -#15749 := [monotonicity #15711 #15746]: #15748 -#15754 := [trans #15749 #15752]: #15753 +#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) +#10038 := (iff #2227 #10035) +#9987 := (and #52 #3477) +#9991 := (and #3645 #9987) +#9994 := (and #3642 #9991) +#9997 := (and #3639 #9994) +#10000 := (and #3624 #9997) +#10003 := (and #3921 #10000) +#10012 := (not #10003) +#10013 := (or #10012 #10006) +#10018 := (forall (vars (?x572 T4) (?x573 T5)) (:pat #2218) #10013) +#10036 := (iff #10018 #10035) +#10033 := (iff #10013 #10030) +#10027 := (or #10024 #10006) +#10031 := (iff #10027 #10030) +#10032 := [rewrite]: #10031 +#10028 := (iff #10013 #10027) +#10025 := (iff #10012 #10024) +#10022 := (iff #10003 #10021) +#10023 := [rewrite]: #10022 +#10026 := [monotonicity #10023]: #10025 +#10029 := [monotonicity #10026]: #10028 +#10034 := [trans #10029 #10032]: #10033 +#10037 := [quant-intro #10034]: #10036 +#10019 := (iff #2227 #10018) +#10016 := (iff #2226 #10013) +#10009 := (implies #10003 #10006) +#10014 := (iff #10009 #10013) +#10015 := [rewrite]: #10014 +#10010 := (iff #2226 #10009) +#10007 := (iff #2225 #10006) +#10008 := [rewrite]: #10007 +#10004 := (iff #2224 #10003) +#10001 := (iff #2223 #10000) +#9998 := (iff #2222 #9997) +#9995 := (iff #2221 #9994) +#9992 := (iff #2220 #9991) +#9989 := (iff #2219 #9987) +#3478 := (iff #72 #3477) +#3479 := [rewrite]: #3478 +#9990 := [monotonicity #3479]: #9989 +#9993 := [monotonicity #3647 #9990]: #9992 +#3643 := (iff #145 #3642) +#3644 := [rewrite]: #3643 +#9996 := [monotonicity #3644 #9993]: #9995 +#9999 := [monotonicity #3641 #9996]: #9998 +#3626 := (iff #136 #3624) +#3627 := [rewrite]: #3626 +#10002 := [monotonicity #3627 #9999]: #10001 +#3923 := (iff #229 #3921) +#3924 := [rewrite]: #3923 +#10005 := [monotonicity #3924 #10002]: #10004 +#10011 := [monotonicity #10005 #10008]: #10010 +#10017 := [trans #10011 #10015]: #10016 +#10020 := [quant-intro #10017]: #10019 +#10039 := [trans #10020 #10037]: #10038 +#9986 := [asserted]: #2227 +#10040 := [mp #9986 #10039]: #10035 +#17763 := [mp~ #10040 #17762]: #10035 +#21791 := [mp #17763 #21790]: #21788 +#25543 := (not #25542) +#25332 := (not #25306) +#25276 := (not #25275) +#22661 := (not #11908) +#25535 := (not #21788) +#25562 := (or #25535 #22661 #25276 #25332 #25433 #25538 #25541 #25543 #25547) +#25545 := (= #25544 uf_14) +#25546 := (or #25545 #25276 #25332 #25543 #25541 #25538 #22661 #25433) +#25563 := (or #25535 #25546) +#25625 := (iff #25563 #25562) +#25553 := (or #22661 #25276 #25332 #25433 #25538 #25541 #25543 #25547) +#25559 := (or #25535 #25553) +#25618 := (iff #25559 #25562) +#25624 := [rewrite]: #25618 +#25565 := (iff #25563 #25559) +#25556 := (iff #25546 #25553) +#25550 := (or #25547 #25276 #25332 #25543 #25541 #25538 #22661 #25433) +#25554 := (iff #25550 #25553) +#25555 := [rewrite]: #25554 +#25551 := (iff #25546 #25550) +#25548 := (iff #25545 #25547) +#25549 := [rewrite]: #25548 +#25552 := [monotonicity #25549]: #25551 +#25557 := [trans #25552 #25555]: #25556 +#25623 := [monotonicity #25557]: #25565 +#25601 := [trans #25623 #25624]: #25625 +#25564 := [quant-inst]: #25563 +#25626 := [mp #25564 #25601]: #25562 +#26028 := [unit-resolution #25626 #21791 #14251 #26026 #25945 #25935 #25940]: #26027 +#26029 := [unit-resolution #26028 #26005 #25781 #25795]: false +#26048 := [lemma #26029]: #25433 +#25441 := (or #25473 #25454 #25435) +#25442 := [def-axiom]: #25441 +#25561 := [unit-resolution #25442 #26048]: #25516 +#25532 := [unit-resolution #25561 #25951]: #25435 +#25581 := [trans #25532 #25593]: #11928 +#25476 := [hypothesis]: #13182 +#25585 := [unit-resolution #25476 #25581]: false +#25596 := [lemma #25585]: #11928 +#23433 := (or #13182 #23430) +#22427 := (forall (vars (?x778 int)) #22416) +#22434 := (not #22427) +#22412 := (forall (vars (?x776 int)) #22407) +#22433 := (not #22412) +#22435 := (or #22433 #22434) +#22436 := (not #22435) +#22465 := (or #22436 #22462) +#22471 := (not #22465) +#22472 := (or #12096 #12087 #12078 #12069 #22348 #13552 #13698 #22471) +#22473 := (not #22472) +#22251 := (forall (vars (?x786 int)) #22246) +#22257 := (not #22251) +#22258 := (or #22233 #22257) +#22259 := (not #22258) +#22288 := (or #22259 #22285) +#22294 := (not #22288) +#22295 := (or #13451 #22294) +#22296 := (not #22295) +#22301 := (or #13451 #22296) +#22309 := (not #22301) +#22310 := (or #12469 #22307 #18463 #22308 #13518 #18466 #22309) +#22311 := (not #22310) +#22316 := (or #18463 #18466 #22311) +#22322 := (not #22316) +#22359 := (or #12662 #12653 #22348 #13552 #22307 #13603 #22322) +#22360 := (not #22359) +#22323 := (or #12558 #12549 #18425 #18434 #12567 #12524 #12515 #13552 #22307 #22322) +#22324 := (not #22323) +#22329 := (or #18425 #18434 #22324) +#22335 := (not #22329) +#22336 := (or #18425 #18428 #22335) +#22337 := (not #22336) +#22342 := (or #18425 #18428 #22337) +#22349 := (not #22342) +#22350 := (or #22348 #13552 #13604 #22349) +#22351 := (not #22350) +#22365 := (or #22351 #22360) +#22371 := (not #22365) +#22372 := (or #18425 #18434 #22348 #13552 #22371) +#22373 := (not #22372) +#22378 := (or #18425 #18434 #22373) +#22384 := (not #22378) +#22385 := (or #18425 #18428 #22384) +#22386 := (not #22385) +#22391 := (or #18425 #18428 #22386) +#22397 := (not #22391) +#22398 := (or #22348 #13552 #13697 #22397) +#22399 := (not #22398) +#22478 := (or #22399 #22473) +#22493 := (not #22478) +#22228 := (forall (vars (?x775 int)) #22223) +#22489 := (not #22228) +#22494 := (or #13149 #12914 #12905 #12896 #12887 #22484 #22485 #22486 #15177 #13870 #13404 #22348 #13552 #13875 #13927 #22487 #22488 #22490 #22491 #22492 #22489 #22493) +#22495 := (not #22494) +#22500 := (or #13149 #13404 #22495) +#22507 := (not #22500) +#22217 := (forall (vars (?x773 int)) #22212) +#22506 := (not #22217) +#22508 := (or #22506 #22507) +#22509 := (not #22508) +#22514 := (or #22206 #22509) +#22520 := (not #22514) +#22521 := (or #13368 #22520) +#22522 := (not #22521) +#22527 := (or #13368 #22522) +#22533 := (not #22527) +#22534 := (or #13149 #13140 #13131 #13122 #18314 #18323 #22533) +#22535 := (not #22534) +#22540 := (or #18314 #18323 #22535) +#22546 := (not #22540) +#22547 := (or #18314 #18317 #22546) +#22548 := (not #22547) +#22553 := (or #18314 #18317 #22548) +#22559 := (not #22553) +#22560 := (or #13182 #22559) +#22561 := (not #22560) +#22566 := (or #13182 #22561) +#23434 := (iff #22566 #23433) +#23431 := (iff #22561 #23430) +#23428 := (iff #22560 #23427) +#23425 := (iff #22559 #23424) +#23422 := (iff #22553 #23421) +#23419 := (iff #22548 #23418) +#23416 := (iff #22547 #23415) +#23413 := (iff #22546 #23412) +#23410 := (iff #22540 #23409) +#23407 := (iff #22535 #23406) +#23404 := (iff #22534 #23403) +#23401 := (iff #22533 #23400) +#23398 := (iff #22527 #23397) +#23395 := (iff #22522 #23394) +#23392 := (iff #22521 #23391) +#23389 := (iff #22520 #23388) +#23386 := (iff #22514 #23385) +#23383 := (iff #22509 #23382) +#23380 := (iff #22508 #23379) +#23377 := (iff #22507 #23376) +#23374 := (iff #22500 #23373) +#23371 := (iff #22495 #23370) +#23368 := (iff #22494 #23367) +#23365 := (iff #22493 #23364) +#23362 := (iff #22478 #23361) +#23359 := (iff #22473 #23358) +#23356 := (iff #22472 #23355) +#23353 := (iff #22471 #23352) +#23350 := (iff #22465 #23349) +#23347 := (iff #22436 #23346) +#23344 := (iff #22435 #23343) +#23341 := (iff #22434 #23340) +#23338 := (iff #22427 #23335) +#23336 := (iff #22416 #22416) +#23337 := [refl]: #23336 +#23339 := [quant-intro #23337]: #23338 +#23342 := [monotonicity #23339]: #23341 +#23333 := (iff #22433 #23332) +#23330 := (iff #22412 #23327) +#23328 := (iff #22407 #22407) +#23329 := [refl]: #23328 +#23331 := [quant-intro #23329]: #23330 +#23334 := [monotonicity #23331]: #23333 +#23345 := [monotonicity #23334 #23342]: #23344 +#23348 := [monotonicity #23345]: #23347 +#23351 := [monotonicity #23348]: #23350 +#23354 := [monotonicity #23351]: #23353 +#23357 := [monotonicity #23354]: #23356 +#23360 := [monotonicity #23357]: #23359 +#23325 := (iff #22399 #23324) +#23322 := (iff #22398 #23321) +#23319 := (iff #22397 #23318) +#23316 := (iff #22391 #23315) +#23313 := (iff #22386 #23312) +#23310 := (iff #22385 #23309) +#23307 := (iff #22384 #23306) +#23304 := (iff #22378 #23303) +#23301 := (iff #22373 #23300) +#23298 := (iff #22372 #23297) +#23295 := (iff #22371 #23294) +#23292 := (iff #22365 #23291) +#23289 := (iff #22360 #23288) +#23286 := (iff #22359 #23285) +#23253 := (iff #22322 #23252) +#23250 := (iff #22316 #23249) +#23247 := (iff #22311 #23246) +#23244 := (iff #22310 #23243) +#23241 := (iff #22309 #23240) +#23238 := (iff #22301 #23237) +#23235 := (iff #22296 #23234) +#23232 := (iff #22295 #23231) +#23229 := (iff #22294 #23228) +#23226 := (iff #22288 #23225) +#23223 := (iff #22259 #23222) +#23220 := (iff #22258 #23219) +#23217 := (iff #22257 #23216) +#23214 := (iff #22251 #23211) +#23212 := (iff #22246 #22246) +#23213 := [refl]: #23212 +#23215 := [quant-intro #23213]: #23214 +#23218 := [monotonicity #23215]: #23217 +#23221 := [monotonicity #23218]: #23220 +#23224 := [monotonicity #23221]: #23223 +#23227 := [monotonicity #23224]: #23226 +#23230 := [monotonicity #23227]: #23229 +#23233 := [monotonicity #23230]: #23232 +#23236 := [monotonicity #23233]: #23235 +#23239 := [monotonicity #23236]: #23238 +#23242 := [monotonicity #23239]: #23241 +#23245 := [monotonicity #23242]: #23244 +#23248 := [monotonicity #23245]: #23247 +#23251 := [monotonicity #23248]: #23250 +#23254 := [monotonicity #23251]: #23253 +#23287 := [monotonicity #23254]: #23286 +#23290 := [monotonicity #23287]: #23289 +#23283 := (iff #22351 #23282) +#23280 := (iff #22350 #23279) +#23277 := (iff #22349 #23276) +#23274 := (iff #22342 #23273) +#23271 := (iff #22337 #23270) +#23268 := (iff #22336 #23267) +#23265 := (iff #22335 #23264) +#23262 := (iff #22329 #23261) +#23259 := (iff #22324 #23258) +#23256 := (iff #22323 #23255) +#23257 := [monotonicity #23254]: #23256 +#23260 := [monotonicity #23257]: #23259 +#23263 := [monotonicity #23260]: #23262 +#23266 := [monotonicity #23263]: #23265 +#23269 := [monotonicity #23266]: #23268 +#23272 := [monotonicity #23269]: #23271 +#23275 := [monotonicity #23272]: #23274 +#23278 := [monotonicity #23275]: #23277 +#23281 := [monotonicity #23278]: #23280 +#23284 := [monotonicity #23281]: #23283 +#23293 := [monotonicity #23284 #23290]: #23292 +#23296 := [monotonicity #23293]: #23295 +#23299 := [monotonicity #23296]: #23298 +#23302 := [monotonicity #23299]: #23301 +#23305 := [monotonicity #23302]: #23304 +#23308 := [monotonicity #23305]: #23307 +#23311 := [monotonicity #23308]: #23310 +#23314 := [monotonicity #23311]: #23313 +#23317 := [monotonicity #23314]: #23316 +#23320 := [monotonicity #23317]: #23319 +#23323 := [monotonicity #23320]: #23322 +#23326 := [monotonicity #23323]: #23325 +#23363 := [monotonicity #23326 #23360]: #23362 +#23366 := [monotonicity #23363]: #23365 +#23209 := (iff #22489 #23208) +#23206 := (iff #22228 #23203) +#23204 := (iff #22223 #22223) +#23205 := [refl]: #23204 +#23207 := [quant-intro #23205]: #23206 +#23210 := [monotonicity #23207]: #23209 +#23369 := [monotonicity #23210 #23366]: #23368 +#23372 := [monotonicity #23369]: #23371 +#23375 := [monotonicity #23372]: #23374 +#23378 := [monotonicity #23375]: #23377 +#23201 := (iff #22506 #23200) +#23198 := (iff #22217 #23195) +#23196 := (iff #22212 #22212) +#23197 := [refl]: #23196 +#23199 := [quant-intro #23197]: #23198 +#23202 := [monotonicity #23199]: #23201 +#23381 := [monotonicity #23202 #23378]: #23380 +#23384 := [monotonicity #23381]: #23383 +#23387 := [monotonicity #23384]: #23386 +#23390 := [monotonicity #23387]: #23389 +#23393 := [monotonicity #23390]: #23392 +#23396 := [monotonicity #23393]: #23395 +#23399 := [monotonicity #23396]: #23398 +#23402 := [monotonicity #23399]: #23401 +#23405 := [monotonicity #23402]: #23404 +#23408 := [monotonicity #23405]: #23407 +#23411 := [monotonicity #23408]: #23410 +#23414 := [monotonicity #23411]: #23413 +#23417 := [monotonicity #23414]: #23416 +#23420 := [monotonicity #23417]: #23419 +#23423 := [monotonicity #23420]: #23422 +#23426 := [monotonicity #23423]: #23425 +#23429 := [monotonicity #23426]: #23428 +#23432 := [monotonicity #23429]: #23431 +#23435 := [monotonicity #23432]: #23434 +#18965 := (and #18608 #18609) +#18968 := (not #18965) +#18971 := (or #18947 #18960 #18968) +#18974 := (not #18971) +#15917 := (and #3095 #4065 #13727 #15097) +#18631 := (not #15917) +#18634 := (forall (vars (?x778 int)) #18631) +#14340 := (and #4065 #15097) +#14339 := (not #14340) +#15903 := (or #13725 #13739 #14339) +#15908 := (forall (vars (?x776 int)) #15903) +#18638 := (and #15908 #18634) +#18980 := (or #18638 #18974) +#18988 := (and #12011 #12014 #12017 #12020 #13409 #13412 #13697 #18980) +#18817 := (and #18484 #18485) +#18820 := (not #18817) +#18823 := (or #18799 #18812 #18820) +#18826 := (not #18823) +#15804 := (or #13454 #13468 #14339) +#15809 := (forall (vars (?x786 int)) #15804) +#18500 := (not #13497) +#18510 := (and #18500 #15809) +#18832 := (or #18510 #18826) +#18837 := (and #13446 #18832) +#18840 := (or #13451 #18837) +#18848 := (and #3211 #13421 #13430 #13438 #13514 #15796 #18840) +#18853 := (or #18463 #18466 #18848) +#18893 := (and #12644 #12647 #13409 #13412 #13421 #13604 #18853) +#18859 := (and #3194 #3196 #12345 #12354 #12360 #12365 #12368 #13412 #13421 #18853) +#18864 := (or #18425 #18434 #18859) +#18870 := (and #12345 #12348 #18864) +#18875 := (or #18425 #18428 #18870) +#18881 := (and #13409 #13412 #13603 #18875) +#18898 := (or #18881 #18893) +#18904 := (and #12345 #12354 #13409 #13412 #18898) +#18909 := (or #18425 #18434 #18904) +#18915 := (and #12345 #12348 #18909) +#18920 := (or #18425 #18428 #18915) +#18926 := (and #13409 #13412 #13698 #18920) +#18993 := (or #18926 #18988) +#15730 := (or #13890 #13904 #14339) +#15735 := (forall (vars (?x775 int)) #15730) +#18999 := (and #3027 #3169 #3170 #3171 #3172 #3173 #3174 #11992 #12000 #12306 #13405 #13409 #13412 #13876 #13924 #13930 #13952 #15735 #15747 #15764 #15781 #18993) +#19004 := (or #13149 #13404 #18999) +#15716 := (or #13371 #13383 #14339) +#15721 := (forall (vars (?x773 int)) #15716) +#19007 := (and #15721 #19004) +#18720 := (and #18347 #18348) +#18723 := (not #18720) +#18729 := (or #18355 #18356 #18723) +#18734 := (not #18729) +#19010 := (or #18734 #19007) +#19013 := (and #13365 #19010) +#19016 := (or #13368 #19013) +#19022 := (and #3027 #3030 #3033 #3036 #11931 #11940 #19016) +#19027 := (or #18314 #18323 #19022) +#19033 := (and #11931 #11934 #19027) +#19038 := (or #18314 #18317 #19033) +#19041 := (and #11928 #19038) +#19044 := (or #13182 #19041) +#22567 := (iff #19044 #22566) +#22564 := (iff #19041 #22561) +#22556 := (and #11928 #22553) +#22562 := (iff #22556 #22561) +#22563 := [rewrite]: #22562 +#22557 := (iff #19041 #22556) +#22554 := (iff #19038 #22553) +#22551 := (iff #19033 #22548) +#22543 := (and #11931 #11934 #22540) +#22549 := (iff #22543 #22548) +#22550 := [rewrite]: #22549 +#22544 := (iff #19033 #22543) +#22541 := (iff #19027 #22540) +#22538 := (iff #19022 #22535) +#22530 := (and #3027 #3030 #3033 #3036 #11931 #11940 #22527) +#22536 := (iff #22530 #22535) +#22537 := [rewrite]: #22536 +#22531 := (iff #19022 #22530) +#22528 := (iff #19016 #22527) +#22525 := (iff #19013 #22522) +#22517 := (and #13365 #22514) +#22523 := (iff #22517 #22522) +#22524 := [rewrite]: #22523 +#22518 := (iff #19013 #22517) +#22515 := (iff #19010 #22514) +#22512 := (iff #19007 #22509) +#22503 := (and #22217 #22500) +#22510 := (iff #22503 #22509) +#22511 := [rewrite]: #22510 +#22504 := (iff #19007 #22503) +#22501 := (iff #19004 #22500) +#22498 := (iff #18999 #22495) +#22481 := (and #3027 #3169 #3170 #3171 #3172 #3173 #3174 #11992 #12000 #12306 #13405 #13409 #13412 #13876 #13924 #13930 #13952 #22228 #15747 #15764 #15781 #22478) +#22496 := (iff #22481 #22495) +#22497 := [rewrite]: #22496 +#22482 := (iff #18999 #22481) +#22479 := (iff #18993 #22478) +#22476 := (iff #18988 #22473) +#22468 := (and #12011 #12014 #12017 #12020 #13409 #13412 #13697 #22465) +#22474 := (iff #22468 #22473) +#22475 := [rewrite]: #22474 +#22469 := (iff #18988 #22468) +#22466 := (iff #18980 #22465) +#22463 := (iff #18974 #22462) +#22460 := (iff #18971 #22457) +#22443 := (or #22441 #22442) +#22454 := (or #18947 #18960 #22443) +#22458 := (iff #22454 #22457) +#22459 := [rewrite]: #22458 +#22455 := (iff #18971 #22454) +#22452 := (iff #18968 #22443) +#22444 := (not #22443) +#22447 := (not #22444) +#22450 := (iff #22447 #22443) +#22451 := [rewrite]: #22450 +#22448 := (iff #18968 #22447) +#22445 := (iff #18965 #22444) +#22446 := [rewrite]: #22445 +#22449 := [monotonicity #22446]: #22448 +#22453 := [trans #22449 #22451]: #22452 +#22456 := [monotonicity #22453]: #22455 +#22461 := [trans #22456 #22459]: #22460 +#22464 := [monotonicity #22461]: #22463 +#22439 := (iff #18638 #22436) +#22430 := (and #22412 #22427) +#22437 := (iff #22430 #22436) +#22438 := [rewrite]: #22437 +#22431 := (iff #18638 #22430) +#22428 := (iff #18634 #22427) +#22425 := (iff #18631 #22416) +#22417 := (not #22416) +#22420 := (not #22417) +#22423 := (iff #22420 #22416) +#22424 := [rewrite]: #22423 +#22421 := (iff #18631 #22420) +#22418 := (iff #15917 #22417) +#22419 := [rewrite]: #22418 +#22422 := [monotonicity #22419]: #22421 +#22426 := [trans #22422 #22424]: #22425 +#22429 := [quant-intro #22426]: #22428 +#22413 := (iff #15908 #22412) +#22410 := (iff #15903 #22407) +#20120 := (or #4987 #19482) +#22404 := (or #13725 #13739 #20120) +#22408 := (iff #22404 #22407) +#22409 := [rewrite]: #22408 +#22405 := (iff #15903 #22404) +#20129 := (iff #14339 #20120) +#20121 := (not #20120) +#20124 := (not #20121) +#20127 := (iff #20124 #20120) +#20128 := [rewrite]: #20127 +#20125 := (iff #14339 #20124) +#20122 := (iff #14340 #20121) +#20123 := [rewrite]: #20122 +#20126 := [monotonicity #20123]: #20125 +#20130 := [trans #20126 #20128]: #20129 +#22406 := [monotonicity #20130]: #22405 +#22411 := [trans #22406 #22409]: #22410 +#22414 := [quant-intro #22411]: #22413 +#22432 := [monotonicity #22414 #22429]: #22431 +#22440 := [trans #22432 #22438]: #22439 +#22467 := [monotonicity #22440 #22464]: #22466 +#22470 := [monotonicity #22467]: #22469 +#22477 := [trans #22470 #22475]: #22476 +#22402 := (iff #18926 #22399) +#22394 := (and #13409 #13412 #13698 #22391) +#22400 := (iff #22394 #22399) +#22401 := [rewrite]: #22400 +#22395 := (iff #18926 #22394) +#22392 := (iff #18920 #22391) +#22389 := (iff #18915 #22386) +#22381 := (and #12345 #12348 #22378) +#22387 := (iff #22381 #22386) +#22388 := [rewrite]: #22387 +#22382 := (iff #18915 #22381) +#22379 := (iff #18909 #22378) +#22376 := (iff #18904 #22373) +#22368 := (and #12345 #12354 #13409 #13412 #22365) +#22374 := (iff #22368 #22373) +#22375 := [rewrite]: #22374 +#22369 := (iff #18904 #22368) +#22366 := (iff #18898 #22365) +#22363 := (iff #18893 #22360) +#22356 := (and #12644 #12647 #13409 #13412 #13421 #13604 #22316) +#22361 := (iff #22356 #22360) +#22362 := [rewrite]: #22361 +#22357 := (iff #18893 #22356) +#22317 := (iff #18853 #22316) +#22314 := (iff #18848 #22311) +#22304 := (and #3211 #13421 #13430 #13438 #13514 #15796 #22301) +#22312 := (iff #22304 #22311) +#22313 := [rewrite]: #22312 +#22305 := (iff #18848 #22304) +#22302 := (iff #18840 #22301) +#22299 := (iff #18837 #22296) +#22291 := (and #13446 #22288) +#22297 := (iff #22291 #22296) +#22298 := [rewrite]: #22297 +#22292 := (iff #18837 #22291) +#22289 := (iff #18832 #22288) +#22286 := (iff #18826 #22285) +#22283 := (iff #18823 #22280) +#22266 := (or #22264 #22265) +#22277 := (or #18799 #18812 #22266) +#22281 := (iff #22277 #22280) +#22282 := [rewrite]: #22281 +#22278 := (iff #18823 #22277) +#22275 := (iff #18820 #22266) +#22267 := (not #22266) +#22270 := (not #22267) +#22273 := (iff #22270 #22266) +#22274 := [rewrite]: #22273 +#22271 := (iff #18820 #22270) +#22268 := (iff #18817 #22267) +#22269 := [rewrite]: #22268 +#22272 := [monotonicity #22269]: #22271 +#22276 := [trans #22272 #22274]: #22275 +#22279 := [monotonicity #22276]: #22278 +#22284 := [trans #22279 #22282]: #22283 +#22287 := [monotonicity #22284]: #22286 +#22262 := (iff #18510 #22259) +#22254 := (and #22232 #22251) +#22260 := (iff #22254 #22259) +#22261 := [rewrite]: #22260 +#22255 := (iff #18510 #22254) +#22252 := (iff #15809 #22251) +#22249 := (iff #15804 #22246) +#22243 := (or #13454 #13468 #20120) +#22247 := (iff #22243 #22246) +#22248 := [rewrite]: #22247 +#22244 := (iff #15804 #22243) +#22245 := [monotonicity #20130]: #22244 +#22250 := [trans #22245 #22248]: #22249 +#22253 := [quant-intro #22250]: #22252 +#22241 := (iff #18500 #22232) +#22236 := (not #22233) +#22239 := (iff #22236 #22232) +#22240 := [rewrite]: #22239 +#22237 := (iff #18500 #22236) +#22234 := (iff #13497 #22233) +#22235 := [rewrite]: #22234 +#22238 := [monotonicity #22235]: #22237 +#22242 := [trans #22238 #22240]: #22241 +#22256 := [monotonicity #22242 #22253]: #22255 +#22263 := [trans #22256 #22261]: #22262 +#22290 := [monotonicity #22263 #22287]: #22289 +#22293 := [monotonicity #22290]: #22292 +#22300 := [trans #22293 #22298]: #22299 +#22303 := [monotonicity #22300]: #22302 +#22306 := [monotonicity #22303]: #22305 +#22315 := [trans #22306 #22313]: #22314 +#22318 := [monotonicity #22315]: #22317 +#22358 := [monotonicity #22318]: #22357 +#22364 := [trans #22358 #22362]: #22363 +#22354 := (iff #18881 #22351) +#22345 := (and #13409 #13412 #13603 #22342) +#22352 := (iff #22345 #22351) +#22353 := [rewrite]: #22352 +#22346 := (iff #18881 #22345) +#22343 := (iff #18875 #22342) +#22340 := (iff #18870 #22337) +#22332 := (and #12345 #12348 #22329) +#22338 := (iff #22332 #22337) +#22339 := [rewrite]: #22338 +#22333 := (iff #18870 #22332) +#22330 := (iff #18864 #22329) +#22327 := (iff #18859 #22324) +#22319 := (and #3194 #3196 #12345 #12354 #12360 #12365 #12368 #13412 #13421 #22316) +#22325 := (iff #22319 #22324) +#22326 := [rewrite]: #22325 +#22320 := (iff #18859 #22319) +#22321 := [monotonicity #22318]: #22320 +#22328 := [trans #22321 #22326]: #22327 +#22331 := [monotonicity #22328]: #22330 +#22334 := [monotonicity #22331]: #22333 +#22341 := [trans #22334 #22339]: #22340 +#22344 := [monotonicity #22341]: #22343 +#22347 := [monotonicity #22344]: #22346 +#22355 := [trans #22347 #22353]: #22354 +#22367 := [monotonicity #22355 #22364]: #22366 +#22370 := [monotonicity #22367]: #22369 +#22377 := [trans #22370 #22375]: #22376 +#22380 := [monotonicity #22377]: #22379 +#22383 := [monotonicity #22380]: #22382 +#22390 := [trans #22383 #22388]: #22389 +#22393 := [monotonicity #22390]: #22392 +#22396 := [monotonicity #22393]: #22395 +#22403 := [trans #22396 #22401]: #22402 +#22480 := [monotonicity #22403 #22477]: #22479 +#22229 := (iff #15735 #22228) +#22226 := (iff #15730 #22223) +#22220 := (or #13890 #13904 #20120) +#22224 := (iff #22220 #22223) +#22225 := [rewrite]: #22224 +#22221 := (iff #15730 #22220) +#22222 := [monotonicity #20130]: #22221 +#22227 := [trans #22222 #22225]: #22226 +#22230 := [quant-intro #22227]: #22229 +#22483 := [monotonicity #22230 #22480]: #22482 +#22499 := [trans #22483 #22497]: #22498 +#22502 := [monotonicity #22499]: #22501 +#22218 := (iff #15721 #22217) +#22215 := (iff #15716 #22212) +#22209 := (or #13371 #13383 #20120) +#22213 := (iff #22209 #22212) +#22214 := [rewrite]: #22213 +#22210 := (iff #15716 #22209) +#22211 := [monotonicity #20130]: #22210 +#22216 := [trans #22211 #22214]: #22215 +#22219 := [quant-intro #22216]: #22218 +#22505 := [monotonicity #22219 #22502]: #22504 +#22513 := [trans #22505 #22511]: #22512 +#22207 := (iff #18734 #22206) +#22204 := (iff #18729 #22201) +#22187 := (or #22185 #22186) +#22198 := (or #18355 #18356 #22187) +#22202 := (iff #22198 #22201) +#22203 := [rewrite]: #22202 +#22199 := (iff #18729 #22198) +#22196 := (iff #18723 #22187) +#22188 := (not #22187) +#22191 := (not #22188) +#22194 := (iff #22191 #22187) +#22195 := [rewrite]: #22194 +#22192 := (iff #18723 #22191) +#22189 := (iff #18720 #22188) +#22190 := [rewrite]: #22189 +#22193 := [monotonicity #22190]: #22192 +#22197 := [trans #22193 #22195]: #22196 +#22200 := [monotonicity #22197]: #22199 +#22205 := [trans #22200 #22203]: #22204 +#22208 := [monotonicity #22205]: #22207 +#22516 := [monotonicity #22208 #22513]: #22515 +#22519 := [monotonicity #22516]: #22518 +#22526 := [trans #22519 #22524]: #22525 +#22529 := [monotonicity #22526]: #22528 +#22532 := [monotonicity #22529]: #22531 +#22539 := [trans #22532 #22537]: #22538 +#22542 := [monotonicity #22539]: #22541 +#22545 := [monotonicity #22542]: #22544 +#22552 := [trans #22545 #22550]: #22551 +#22555 := [monotonicity #22552]: #22554 +#22558 := [monotonicity #22555]: #22557 +#22565 := [trans #22558 #22563]: #22564 +#22568 := [monotonicity #22565]: #22567 +#18610 := (and #18609 #18608) +#18611 := (not #18610) +#18614 := (+ #18613 #13737) +#18615 := (<= #18614 0::int) +#18616 := (+ ?x776!15 #13338) +#18617 := (>= #18616 0::int) +#18618 := (or #18617 #18615 #18611) +#18619 := (not #18618) +#18642 := (or #18619 #18638) +#18395 := (not #13417) +#18602 := (not #12069) +#18599 := (not #12078) +#18596 := (not #12087) +#18593 := (not #12096) +#18646 := (and #18593 #18596 #18599 #18602 #18395 #13701 #18642) +#15832 := (and #13430 #15796) +#15835 := (not #15832) +#18526 := (not #15835) +#18486 := (and #18485 #18484) +#18487 := (not #18486) +#18490 := (+ #18489 #13466) +#18491 := (<= #18490 0::int) +#18492 := (+ ?x786!14 #13447) +#18493 := (>= #18492 0::int) +#18494 := (or #18493 #18491 #18487) +#18495 := (not #18494) +#18514 := (or #18495 #18510) +#18480 := (not #13451) +#18518 := (and #18480 #18514) +#18522 := (or #13451 #18518) +#18475 := (not #13518) +#18472 := (not #13443) +#18469 := (not #12469) +#18529 := (and #18469 #18472 #18475 #18522 #18526) +#18533 := (or #18463 #18466 #18529) +#18460 := (not #13426) +#18560 := (not #12653) +#18557 := (not #12662) +#18565 := (and #18557 #18560 #18395 #18460 #13609 #18533) +#18457 := (not #13552) +#18454 := (not #12515) +#18451 := (not #12524) +#18448 := (not #12567) +#18437 := (not #12576) +#18445 := (not #12549) +#18442 := (not #12558) +#18537 := (and #18442 #18445 #18437 #18448 #18451 #18454 #18457 #18460 #18533) +#18541 := (or #18425 #18434 #18537) +#18431 := (not #12588) +#18545 := (and #18431 #18541) +#18549 := (or #18425 #18428 #18545) +#18553 := (and #18395 #13603 #18549) +#18569 := (or #18553 #18565) +#18573 := (and #18437 #18395 #18569) +#18577 := (or #18425 #18434 #18573) +#18581 := (and #18431 #18577) +#18585 := (or #18425 #18428 #18581) +#18589 := (and #18395 #13698 #18585) +#18650 := (or #18589 #18646) +#15786 := (and #13952 #15781) +#15789 := (not #15786) +#18420 := (not #15789) +#15769 := (and #13409 #15764) +#15772 := (not #15769) +#18417 := (not #15772) +#15752 := (and #13930 #15747) +#15755 := (not #15752) +#18414 := (not #15755) +#18404 := (not #13973) +#18401 := (not #13927) +#18398 := (not #13887) +#18392 := (not #13870) +#18389 := (not #15177) +#18386 := (not #12878) +#18383 := (not #12887) +#18380 := (not #12896) +#18377 := (not #12905) +#18374 := (not #12914) +#18654 := (and #18374 #18377 #18380 #18383 #18386 #18389 #18392 #18395 #18398 #18401 #18404 #15735 #18414 #18417 #18420 #18650) +#18658 := (or #13149 #14130 #18654) +#18662 := (and #15721 #18658) +#18349 := (and #18348 #18347) +#18350 := (not #18349) +#18357 := (or #18356 #18355 #18350) +#18358 := (not #18357) +#18666 := (or #18358 #18662) +#18343 := (not #13368) +#18670 := (and #18343 #18666) +#18674 := (or #13368 #18670) +#18338 := (not #13158) +#18335 := (not #13122) +#18332 := (not #13131) +#18329 := (not #13140) +#18326 := (not #13149) +#18678 := (and #18326 #18329 #18332 #18335 #18338 #18674) +#18682 := (or #18314 #18323 #18678) +#18320 := (not #13170) +#18686 := (and #18320 #18682) +#18690 := (or #18314 #18317 #18686) +#18311 := (not #13182) +#18694 := (and #18311 #18690) +#18698 := (or #13182 #18694) +#19045 := (iff #18698 #19044) +#19042 := (iff #18694 #19041) +#19039 := (iff #18690 #19038) +#19036 := (iff #18686 #19033) +#19030 := (and #11937 #19027) +#19034 := (iff #19030 #19033) +#19035 := [rewrite]: #19034 +#19031 := (iff #18686 #19030) +#19028 := (iff #18682 #19027) +#19025 := (iff #18678 #19022) +#19019 := (and #3027 #3030 #3033 #3036 #11943 #19016) +#19023 := (iff #19019 #19022) +#19024 := [rewrite]: #19023 +#19020 := (iff #18678 #19019) +#19017 := (iff #18674 #19016) +#19014 := (iff #18670 #19013) +#19011 := (iff #18666 #19010) +#19008 := (iff #18662 #19007) +#19005 := (iff #18658 #19004) +#19002 := (iff #18654 #18999) +#18996 := (and #3169 #3170 #3171 #3172 #3175 #12000 #12306 #13414 #13882 #13924 #13968 #15735 #15752 #15769 #15786 #18993) +#19000 := (iff #18996 #18999) +#19001 := [rewrite]: #19000 +#18997 := (iff #18654 #18996) +#18994 := (iff #18650 #18993) +#18991 := (iff #18646 #18988) +#18985 := (and #12011 #12014 #12017 #12020 #13414 #13697 #18980) +#18989 := (iff #18985 #18988) +#18990 := [rewrite]: #18989 +#18986 := (iff #18646 #18985) +#18983 := (iff #18642 #18980) +#18977 := (or #18974 #18638) +#18981 := (iff #18977 #18980) +#18982 := [rewrite]: #18981 +#18978 := (iff #18642 #18977) +#18975 := (iff #18619 #18974) +#18972 := (iff #18618 #18971) +#18969 := (iff #18611 #18968) +#18966 := (iff #18610 #18965) +#18967 := [rewrite]: #18966 +#18970 := [monotonicity #18967]: #18969 +#18963 := (iff #18615 #18960) +#18952 := (+ #13737 #18613) +#18955 := (<= #18952 0::int) +#18961 := (iff #18955 #18960) +#18962 := [rewrite]: #18961 +#18956 := (iff #18615 #18955) +#18953 := (= #18614 #18952) +#18954 := [rewrite]: #18953 +#18957 := [monotonicity #18954]: #18956 +#18964 := [trans #18957 #18962]: #18963 +#18950 := (iff #18617 #18947) +#18939 := (+ #13338 ?x776!15) +#18942 := (>= #18939 0::int) +#18948 := (iff #18942 #18947) +#18949 := [rewrite]: #18948 +#18943 := (iff #18617 #18942) +#18940 := (= #18616 #18939) +#18941 := [rewrite]: #18940 +#18944 := [monotonicity #18941]: #18943 +#18951 := [trans #18944 #18949]: #18950 +#18973 := [monotonicity #18951 #18964 #18970]: #18972 +#18976 := [monotonicity #18973]: #18975 +#18979 := [monotonicity #18976]: #18978 +#18984 := [trans #18979 #18982]: #18983 +#18751 := (iff #18395 #13414) +#18752 := [rewrite]: #18751 +#18937 := (iff #18602 #12020) +#18938 := [rewrite]: #18937 +#18935 := (iff #18599 #12017) +#18936 := [rewrite]: #18935 +#18933 := (iff #18596 #12014) +#18934 := [rewrite]: #18933 +#18931 := (iff #18593 #12011) +#18932 := [rewrite]: #18931 +#18987 := [monotonicity #18932 #18934 #18936 #18938 #18752 #13705 #18984]: #18986 +#18992 := [trans #18987 #18990]: #18991 +#18929 := (iff #18589 #18926) +#18923 := (and #13414 #13698 #18920) +#18927 := (iff #18923 #18926) +#18928 := [rewrite]: #18927 +#18924 := (iff #18589 #18923) +#18921 := (iff #18585 #18920) +#18918 := (iff #18581 #18915) +#18912 := (and #12351 #18909) +#18916 := (iff #18912 #18915) +#18917 := [rewrite]: #18916 +#18913 := (iff #18581 #18912) +#18910 := (iff #18577 #18909) +#18907 := (iff #18573 #18904) +#18901 := (and #12357 #13414 #18898) +#18905 := (iff #18901 #18904) +#18906 := [rewrite]: #18905 +#18902 := (iff #18573 #18901) +#18899 := (iff #18569 #18898) +#18896 := (iff #18565 #18893) +#18890 := (and #12644 #12647 #13414 #13423 #13604 #18853) +#18894 := (iff #18890 #18893) +#18895 := [rewrite]: #18894 +#18891 := (iff #18565 #18890) +#18854 := (iff #18533 #18853) +#18851 := (iff #18529 #18848) +#18845 := (and #3211 #13440 #13514 #18840 #15832) +#18849 := (iff #18845 #18848) +#18850 := [rewrite]: #18849 +#18846 := (iff #18529 #18845) +#18843 := (iff #18526 #15832) +#18844 := [rewrite]: #18843 +#18841 := (iff #18522 #18840) +#18838 := (iff #18518 #18837) +#18835 := (iff #18514 #18832) +#18829 := (or #18826 #18510) +#18833 := (iff #18829 #18832) +#18834 := [rewrite]: #18833 +#18830 := (iff #18514 #18829) +#18827 := (iff #18495 #18826) +#18824 := (iff #18494 #18823) +#18821 := (iff #18487 #18820) +#18818 := (iff #18486 #18817) +#18819 := [rewrite]: #18818 +#18822 := [monotonicity #18819]: #18821 +#18815 := (iff #18491 #18812) +#18804 := (+ #13466 #18489) +#18807 := (<= #18804 0::int) +#18813 := (iff #18807 #18812) +#18814 := [rewrite]: #18813 +#18808 := (iff #18491 #18807) +#18805 := (= #18490 #18804) +#18806 := [rewrite]: #18805 +#18809 := [monotonicity #18806]: #18808 +#18816 := [trans #18809 #18814]: #18815 +#18802 := (iff #18493 #18799) +#18791 := (+ #13447 ?x786!14) +#18794 := (>= #18791 0::int) +#18800 := (iff #18794 #18799) +#18801 := [rewrite]: #18800 +#18795 := (iff #18493 #18794) +#18792 := (= #18492 #18791) +#18793 := [rewrite]: #18792 +#18796 := [monotonicity #18793]: #18795 +#18803 := [trans #18796 #18801]: #18802 +#18825 := [monotonicity #18803 #18816 #18822]: #18824 +#18828 := [monotonicity #18825]: #18827 +#18831 := [monotonicity #18828]: #18830 +#18836 := [trans #18831 #18834]: #18835 +#18789 := (iff #18480 #13446) +#18790 := [rewrite]: #18789 +#18839 := [monotonicity #18790 #18836]: #18838 +#18842 := [monotonicity #18839]: #18841 +#18787 := (iff #18475 #13514) +#18788 := [rewrite]: #18787 +#18785 := (iff #18472 #13440) +#18786 := [rewrite]: #18785 +#18783 := (iff #18469 #3211) +#18784 := [rewrite]: #18783 +#18847 := [monotonicity #18784 #18786 #18788 #18842 #18844]: #18846 +#18852 := [trans #18847 #18850]: #18851 +#18855 := [monotonicity #18852]: #18854 +#18781 := (iff #18460 #13423) +#18782 := [rewrite]: #18781 +#18888 := (iff #18560 #12647) +#18889 := [rewrite]: #18888 +#18886 := (iff #18557 #12644) +#18887 := [rewrite]: #18886 +#18892 := [monotonicity #18887 #18889 #18752 #18782 #13613 #18855]: #18891 +#18897 := [trans #18892 #18895]: #18896 +#18884 := (iff #18553 #18881) +#18878 := (and #13414 #13603 #18875) +#18882 := (iff #18878 #18881) +#18883 := [rewrite]: #18882 +#18879 := (iff #18553 #18878) +#18876 := (iff #18549 #18875) +#18873 := (iff #18545 #18870) +#18867 := (and #12351 #18864) +#18871 := (iff #18867 #18870) +#18872 := [rewrite]: #18871 +#18868 := (iff #18545 #18867) +#18865 := (iff #18541 #18864) +#18862 := (iff #18537 #18859) +#18856 := (and #3194 #3196 #12357 #12360 #12365 #12368 #13412 #13423 #18853) +#18860 := (iff #18856 #18859) +#18861 := [rewrite]: #18860 +#18857 := (iff #18537 #18856) +#18779 := (iff #18457 #13412) +#18780 := [rewrite]: #18779 +#18777 := (iff #18454 #12368) +#18778 := [rewrite]: #18777 +#18775 := (iff #18451 #12365) +#18776 := [rewrite]: #18775 +#18773 := (iff #18448 #12360) +#18774 := [rewrite]: #18773 +#18767 := (iff #18437 #12357) +#18768 := [rewrite]: #18767 +#18771 := (iff #18445 #3196) +#18772 := [rewrite]: #18771 +#18769 := (iff #18442 #3194) +#18770 := [rewrite]: #18769 +#18858 := [monotonicity #18770 #18772 #18768 #18774 #18776 #18778 #18780 #18782 #18855]: #18857 +#18863 := [trans #18858 #18861]: #18862 +#18866 := [monotonicity #18863]: #18865 +#18765 := (iff #18431 #12351) +#18766 := [rewrite]: #18765 +#18869 := [monotonicity #18766 #18866]: #18868 +#18874 := [trans #18869 #18872]: #18873 +#18877 := [monotonicity #18874]: #18876 +#18880 := [monotonicity #18752 #18877]: #18879 +#18885 := [trans #18880 #18883]: #18884 +#18900 := [monotonicity #18885 #18897]: #18899 +#18903 := [monotonicity #18768 #18752 #18900]: #18902 +#18908 := [trans #18903 #18906]: #18907 +#18911 := [monotonicity #18908]: #18910 +#18914 := [monotonicity #18766 #18911]: #18913 +#18919 := [trans #18914 #18917]: #18918 +#18922 := [monotonicity #18919]: #18921 +#18925 := [monotonicity #18752 #18922]: #18924 +#18930 := [trans #18925 #18928]: #18929 +#18995 := [monotonicity #18930 #18992]: #18994 +#18763 := (iff #18420 #15786) +#18764 := [rewrite]: #18763 +#18761 := (iff #18417 #15769) +#18762 := [rewrite]: #18761 +#18759 := (iff #18414 #15752) +#18760 := [rewrite]: #18759 +#18757 := (iff #18404 #13968) +#18758 := [rewrite]: #18757 +#18755 := (iff #18401 #13924) +#18756 := [rewrite]: #18755 +#18753 := (iff #18398 #13882) +#18754 := [rewrite]: #18753 +#18749 := (iff #18392 #12306) +#18750 := [rewrite]: #18749 +#18747 := (iff #18389 #12000) +#18748 := [rewrite]: #18747 +#18745 := (iff #18386 #3175) +#18746 := [rewrite]: #18745 +#18743 := (iff #18383 #3172) +#18744 := [rewrite]: #18743 +#18741 := (iff #18380 #3171) +#18742 := [rewrite]: #18741 +#18739 := (iff #18377 #3170) +#18740 := [rewrite]: #18739 +#18737 := (iff #18374 #3169) +#18738 := [rewrite]: #18737 +#18998 := [monotonicity #18738 #18740 #18742 #18744 #18746 #18748 #18750 #18752 #18754 #18756 #18758 #18760 #18762 #18764 #18995]: #18997 +#19003 := [trans #18998 #19001]: #19002 +#19006 := [monotonicity #14134 #19003]: #19005 +#19009 := [monotonicity #19006]: #19008 +#18735 := (iff #18358 #18734) +#18732 := (iff #18357 #18729) +#18726 := (or #18356 #18355 #18723) +#18730 := (iff #18726 #18729) +#18731 := [rewrite]: #18730 +#18727 := (iff #18357 #18726) +#18724 := (iff #18350 #18723) +#18721 := (iff #18349 #18720) +#18722 := [rewrite]: #18721 +#18725 := [monotonicity #18722]: #18724 +#18728 := [monotonicity #18725]: #18727 +#18733 := [trans #18728 #18731]: #18732 +#18736 := [monotonicity #18733]: #18735 +#19012 := [monotonicity #18736 #19009]: #19011 +#18718 := (iff #18343 #13365) +#18719 := [rewrite]: #18718 +#19015 := [monotonicity #18719 #19012]: #19014 +#19018 := [monotonicity #19015]: #19017 +#18716 := (iff #18338 #11943) +#18717 := [rewrite]: #18716 +#18714 := (iff #18335 #3036) +#18715 := [rewrite]: #18714 +#18712 := (iff #18332 #3033) +#18713 := [rewrite]: #18712 +#18710 := (iff #18329 #3030) +#18711 := [rewrite]: #18710 +#18708 := (iff #18326 #3027) +#18709 := [rewrite]: #18708 +#19021 := [monotonicity #18709 #18711 #18713 #18715 #18717 #19018]: #19020 +#19026 := [trans #19021 #19024]: #19025 +#19029 := [monotonicity #19026]: #19028 +#18706 := (iff #18320 #11937) +#18707 := [rewrite]: #18706 +#19032 := [monotonicity #18707 #19029]: #19031 +#19037 := [trans #19032 #19035]: #19036 +#19040 := [monotonicity #19037]: #19039 +#18704 := (iff #18311 #11928) +#18705 := [rewrite]: #18704 +#19043 := [monotonicity #18705 #19040]: #19042 +#19046 := [monotonicity #19043]: #19045 +#15922 := (exists (vars (?x778 int)) #15917) +#15911 := (not #15908) +#15925 := (or #15911 #15922) +#15928 := (and #15908 #15925) +#15931 := (or #12096 #12087 #12078 #12069 #13417 #13698 #15928) +#15812 := (not #15809) +#15818 := (or #13497 #15812) +#15823 := (and #15809 #15818) +#15826 := (or #13451 #15823) +#15829 := (and #13446 #15826) +#15841 := (or #12469 #13443 #13518 #15829 #15835) +#15846 := (and #13430 #15796 #15841) +#15872 := (or #12662 #12653 #13417 #13426 #13603 #15846) +#15849 := (or #12558 #12549 #12576 #12567 #12524 #12515 #13552 #13426 #15846) +#15852 := (and #12345 #12354 #15849) +#15855 := (or #12588 #15852) +#15858 := (and #12345 #12348 #15855) +#15864 := (or #13417 #13604 #15858) +#15877 := (and #15864 #15872) +#15880 := (or #12576 #13417 #15877) +#15883 := (and #12345 #12354 #15880) +#15886 := (or #12588 #15883) +#15889 := (and #12345 #12348 #15886) +#15895 := (or #13417 #13697 #15889) +#15934 := (and #15895 #15931) +#15738 := (not #15735) +#15940 := (or #12914 #12905 #12896 #12887 #12878 #15177 #13870 #13417 #13887 #13927 #13973 #15738 #15755 #15772 #15789 #15934) +#15945 := (and #3027 #13405 #15940) +#15724 := (not #15721) +#15948 := (or #15724 #15945) +#15951 := (and #15721 #15948) +#15954 := (or #13368 #15951) +#15957 := (and #13365 #15954) +#15960 := (or #13149 #13140 #13131 #13122 #13158 #15957) +#15963 := (and #11931 #11940 #15960) +#15966 := (or #13170 #15963) +#15969 := (and #11931 #11934 #15966) +#15972 := (or #13182 #15969) +#15975 := (and #11928 #15972) +#15978 := (not #15975) +#18699 := (~ #15978 #18698) +#18695 := (not #15972) +#18696 := (~ #18695 #18694) +#18691 := (not #15969) +#18692 := (~ #18691 #18690) +#18687 := (not #15966) +#18688 := (~ #18687 #18686) +#18683 := (not #15963) +#18684 := (~ #18683 #18682) +#18679 := (not #15960) +#18680 := (~ #18679 #18678) +#18675 := (not #15957) +#18676 := (~ #18675 #18674) +#18671 := (not #15954) +#18672 := (~ #18671 #18670) +#18667 := (not #15951) +#18668 := (~ #18667 #18666) +#18663 := (not #15948) +#18664 := (~ #18663 #18662) +#18659 := (not #15945) +#18660 := (~ #18659 #18658) +#18655 := (not #15940) +#18656 := (~ #18655 #18654) +#18651 := (not #15934) +#18652 := (~ #18651 #18650) +#18647 := (not #15931) +#18648 := (~ #18647 #18646) +#18643 := (not #15928) +#18644 := (~ #18643 #18642) +#18639 := (not #15925) +#18640 := (~ #18639 #18638) +#18635 := (not #15922) +#18636 := (~ #18635 #18634) +#18632 := (~ #18631 #18631) +#18633 := [refl]: #18632 +#18637 := [nnf-neg #18633]: #18636 +#18628 := (not #15911) +#18629 := (~ #18628 #15908) +#18626 := (~ #15908 #15908) +#18624 := (~ #15903 #15903) +#18625 := [refl]: #18624 +#18627 := [nnf-pos #18625]: #18626 +#18630 := [nnf-neg #18627]: #18629 +#18641 := [nnf-neg #18630 #18637]: #18640 +#18620 := (~ #15911 #18619) +#18621 := [sk]: #18620 +#18645 := [nnf-neg #18621 #18641]: #18644 +#18605 := (~ #13701 #13701) +#18606 := [refl]: #18605 +#18396 := (~ #18395 #18395) +#18397 := [refl]: #18396 +#18603 := (~ #18602 #18602) +#18604 := [refl]: #18603 +#18600 := (~ #18599 #18599) +#18601 := [refl]: #18600 +#18597 := (~ #18596 #18596) +#18598 := [refl]: #18597 +#18594 := (~ #18593 #18593) +#18595 := [refl]: #18594 +#18649 := [nnf-neg #18595 #18598 #18601 #18604 #18397 #18606 #18645]: #18648 +#18590 := (not #15895) +#18591 := (~ #18590 #18589) +#18586 := (not #15889) +#18587 := (~ #18586 #18585) +#18582 := (not #15886) +#18583 := (~ #18582 #18581) +#18578 := (not #15883) +#18579 := (~ #18578 #18577) +#18574 := (not #15880) +#18575 := (~ #18574 #18573) +#18570 := (not #15877) +#18571 := (~ #18570 #18569) +#18566 := (not #15872) +#18567 := (~ #18566 #18565) +#18534 := (not #15846) +#18535 := (~ #18534 #18533) +#18530 := (not #15841) +#18531 := (~ #18530 #18529) +#18527 := (~ #18526 #18526) +#18528 := [refl]: #18527 +#18523 := (not #15829) +#18524 := (~ #18523 #18522) +#18519 := (not #15826) +#18520 := (~ #18519 #18518) +#18515 := (not #15823) +#18516 := (~ #18515 #18514) +#18511 := (not #15818) +#18512 := (~ #18511 #18510) +#18507 := (not #15812) +#18508 := (~ #18507 #15809) +#18505 := (~ #15809 #15809) +#18503 := (~ #15804 #15804) +#18504 := [refl]: #18503 +#18506 := [nnf-pos #18504]: #18505 +#18509 := [nnf-neg #18506]: #18508 +#18501 := (~ #18500 #18500) +#18502 := [refl]: #18501 +#18513 := [nnf-neg #18502 #18509]: #18512 +#18496 := (~ #15812 #18495) +#18497 := [sk]: #18496 +#18517 := [nnf-neg #18497 #18513]: #18516 +#18481 := (~ #18480 #18480) +#18482 := [refl]: #18481 +#18521 := [nnf-neg #18482 #18517]: #18520 +#18478 := (~ #13451 #13451) +#18479 := [refl]: #18478 +#18525 := [nnf-neg #18479 #18521]: #18524 +#18476 := (~ #18475 #18475) +#18477 := [refl]: #18476 +#18473 := (~ #18472 #18472) +#18474 := [refl]: #18473 +#18470 := (~ #18469 #18469) +#18471 := [refl]: #18470 +#18532 := [nnf-neg #18471 #18474 #18477 #18525 #18528]: #18531 +#18467 := (~ #18466 #18466) +#18468 := [refl]: #18467 +#18464 := (~ #18463 #18463) +#18465 := [refl]: #18464 +#18536 := [nnf-neg #18465 #18468 #18532]: #18535 +#18563 := (~ #13609 #13609) +#18564 := [refl]: #18563 +#18461 := (~ #18460 #18460) +#18462 := [refl]: #18461 +#18561 := (~ #18560 #18560) +#18562 := [refl]: #18561 +#18558 := (~ #18557 #18557) +#18559 := [refl]: #18558 +#18568 := [nnf-neg #18559 #18562 #18397 #18462 #18564 #18536]: #18567 +#18554 := (not #15864) +#18555 := (~ #18554 #18553) +#18550 := (not #15858) +#18551 := (~ #18550 #18549) +#18546 := (not #15855) +#18547 := (~ #18546 #18545) +#18542 := (not #15852) +#18543 := (~ #18542 #18541) +#18538 := (not #15849) +#18539 := (~ #18538 #18537) +#18458 := (~ #18457 #18457) +#18459 := [refl]: #18458 +#18455 := (~ #18454 #18454) +#18456 := [refl]: #18455 +#18452 := (~ #18451 #18451) +#18453 := [refl]: #18452 +#18449 := (~ #18448 #18448) +#18450 := [refl]: #18449 +#18438 := (~ #18437 #18437) +#18439 := [refl]: #18438 +#18446 := (~ #18445 #18445) +#18447 := [refl]: #18446 +#18443 := (~ #18442 #18442) +#18444 := [refl]: #18443 +#18540 := [nnf-neg #18444 #18447 #18439 #18450 #18453 #18456 #18459 #18462 #18536]: #18539 +#18435 := (~ #18434 #18434) +#18436 := [refl]: #18435 +#18426 := (~ #18425 #18425) +#18427 := [refl]: #18426 +#18544 := [nnf-neg #18427 #18436 #18540]: #18543 +#18432 := (~ #18431 #18431) +#18433 := [refl]: #18432 +#18548 := [nnf-neg #18433 #18544]: #18547 +#18429 := (~ #18428 #18428) +#18430 := [refl]: #18429 +#18552 := [nnf-neg #18427 #18430 #18548]: #18551 +#18440 := (~ #13603 #13603) +#18441 := [refl]: #18440 +#18556 := [nnf-neg #18397 #18441 #18552]: #18555 +#18572 := [nnf-neg #18556 #18568]: #18571 +#18576 := [nnf-neg #18439 #18397 #18572]: #18575 +#18580 := [nnf-neg #18427 #18436 #18576]: #18579 +#18584 := [nnf-neg #18433 #18580]: #18583 +#18588 := [nnf-neg #18427 #18430 #18584]: #18587 +#18423 := (~ #13698 #13698) +#18424 := [refl]: #18423 +#18592 := [nnf-neg #18397 #18424 #18588]: #18591 +#18653 := [nnf-neg #18592 #18649]: #18652 +#18421 := (~ #18420 #18420) +#18422 := [refl]: #18421 +#18418 := (~ #18417 #18417) +#18419 := [refl]: #18418 +#18415 := (~ #18414 #18414) +#18416 := [refl]: #18415 +#18411 := (not #15738) +#18412 := (~ #18411 #15735) +#18409 := (~ #15735 #15735) +#18407 := (~ #15730 #15730) +#18408 := [refl]: #18407 +#18410 := [nnf-pos #18408]: #18409 +#18413 := [nnf-neg #18410]: #18412 +#18405 := (~ #18404 #18404) +#18406 := [refl]: #18405 +#18402 := (~ #18401 #18401) +#18403 := [refl]: #18402 +#18399 := (~ #18398 #18398) +#18400 := [refl]: #18399 +#18393 := (~ #18392 #18392) +#18394 := [refl]: #18393 +#18390 := (~ #18389 #18389) +#18391 := [refl]: #18390 +#18387 := (~ #18386 #18386) +#18388 := [refl]: #18387 +#18384 := (~ #18383 #18383) +#18385 := [refl]: #18384 +#18381 := (~ #18380 #18380) +#18382 := [refl]: #18381 +#18378 := (~ #18377 #18377) +#18379 := [refl]: #18378 +#18375 := (~ #18374 #18374) +#18376 := [refl]: #18375 +#18657 := [nnf-neg #18376 #18379 #18382 #18385 #18388 #18391 #18394 #18397 #18400 #18403 #18406 #18413 #18416 #18419 #18422 #18653]: #18656 +#18372 := (~ #14130 #14130) +#18373 := [refl]: #18372 +#18370 := (~ #13149 #13149) +#18371 := [refl]: #18370 +#18661 := [nnf-neg #18371 #18373 #18657]: #18660 +#18367 := (not #15724) +#18368 := (~ #18367 #15721) +#18365 := (~ #15721 #15721) +#18363 := (~ #15716 #15716) +#18364 := [refl]: #18363 +#18366 := [nnf-pos #18364]: #18365 +#18369 := [nnf-neg #18366]: #18368 +#18665 := [nnf-neg #18369 #18661]: #18664 +#18359 := (~ #15724 #18358) +#18360 := [sk]: #18359 +#18669 := [nnf-neg #18360 #18665]: #18668 +#18344 := (~ #18343 #18343) +#18345 := [refl]: #18344 +#18673 := [nnf-neg #18345 #18669]: #18672 +#18341 := (~ #13368 #13368) +#18342 := [refl]: #18341 +#18677 := [nnf-neg #18342 #18673]: #18676 +#18339 := (~ #18338 #18338) +#18340 := [refl]: #18339 +#18336 := (~ #18335 #18335) +#18337 := [refl]: #18336 +#18333 := (~ #18332 #18332) +#18334 := [refl]: #18333 +#18330 := (~ #18329 #18329) +#18331 := [refl]: #18330 +#18327 := (~ #18326 #18326) +#18328 := [refl]: #18327 +#18681 := [nnf-neg #18328 #18331 #18334 #18337 #18340 #18677]: #18680 +#18324 := (~ #18323 #18323) +#18325 := [refl]: #18324 +#18315 := (~ #18314 #18314) +#18316 := [refl]: #18315 +#18685 := [nnf-neg #18316 #18325 #18681]: #18684 +#18321 := (~ #18320 #18320) +#18322 := [refl]: #18321 +#18689 := [nnf-neg #18322 #18685]: #18688 +#18318 := (~ #18317 #18317) +#18319 := [refl]: #18318 +#18693 := [nnf-neg #18316 #18319 #18689]: #18692 +#18312 := (~ #18311 #18311) +#18313 := [refl]: #18312 +#18697 := [nnf-neg #18313 #18693]: #18696 +#18309 := (~ #13182 #13182) +#18310 := [refl]: #18309 +#18700 := [nnf-neg #18310 #18697]: #18699 +#15202 := (or #12096 #12087 #12078 #12069 #13417 #13698 #13784) +#15207 := (and #13720 #15202) +#15213 := (or #12914 #12905 #12896 #12887 #12878 #15177 #13870 #13417 #13887 #13921 #13927 #13939 #13949 #13962 #13973 #15207) +#15218 := (and #3027 #13405 #15213) +#15221 := (or #13401 #15218) +#15224 := (and #13398 #15221) +#15227 := (or #13368 #15224) +#15230 := (and #13365 #15227) +#15233 := (or #13149 #13140 #13131 #13122 #13158 #15230) +#15236 := (and #11931 #11940 #15233) +#15239 := (or #13170 #15236) +#15242 := (and #11931 #11934 #15239) +#15245 := (or #13182 #15242) +#15248 := (and #11928 #15245) +#15251 := (not #15248) +#15979 := (iff #15251 #15978) +#15976 := (iff #15248 #15975) +#15973 := (iff #15245 #15972) +#15970 := (iff #15242 #15969) +#15967 := (iff #15239 #15966) +#15964 := (iff #15236 #15963) +#15961 := (iff #15233 #15960) +#15958 := (iff #15230 #15957) +#15955 := (iff #15227 #15954) +#15952 := (iff #15224 #15951) +#15949 := (iff #15221 #15948) +#15946 := (iff #15218 #15945) +#15943 := (iff #15213 #15940) +#15937 := (or #12914 #12905 #12896 #12887 #12878 #15177 #13870 #13417 #13887 #15738 #13927 #15755 #15772 #15789 #13973 #15934) +#15941 := (iff #15937 #15940) +#15942 := [rewrite]: #15941 +#15938 := (iff #15213 #15937) +#15935 := (iff #15207 #15934) +#15932 := (iff #15202 #15931) +#15929 := (iff #13784 #15928) +#15926 := (iff #13779 #15925) +#15923 := (iff #13770 #15922) +#15920 := (iff #13765 #15917) +#15914 := (and #3095 #4065 #15097 #13727) +#15918 := (iff #15914 #15917) +#15919 := [rewrite]: #15918 +#15915 := (iff #13765 #15914) +#15092 := (iff #4379 #15097) +#15113 := -4294967295::int +#15105 := (+ -4294967295::int #161) +#15098 := (<= #15105 0::int) +#15094 := (iff #15098 #15097) +#15095 := [rewrite]: #15094 +#15099 := (iff #4379 #15098) +#15100 := (= #4378 #15105) +#15106 := (+ #161 -4294967295::int) +#15102 := (= #15106 #15105) +#15103 := [rewrite]: #15102 +#15107 := (= #4378 #15106) +#15108 := (= #4377 -4294967295::int) +#15114 := (* -1::int 4294967295::int) +#15110 := (= #15114 -4294967295::int) +#15111 := [rewrite]: #15110 +#15115 := (= #4377 #15114) +#7500 := (= uf_76 4294967295::int) +#947 := 65536::int +#1322 := (* 65536::int 65536::int) +#1327 := (- #1322 1::int) +#1328 := (= uf_76 #1327) +#7501 := (iff #1328 #7500) +#7498 := (= #1327 4294967295::int) +#1010 := 4294967296::int +#7491 := (- 4294967296::int 1::int) +#7496 := (= #7491 4294967295::int) +#7497 := [rewrite]: #7496 +#7493 := (= #1327 #7491) +#7462 := (= #1322 4294967296::int) +#7463 := [rewrite]: #7462 +#7494 := [monotonicity #7463]: #7493 +#7499 := [trans #7494 #7497]: #7498 +#7502 := [monotonicity #7499]: #7501 +#7490 := [asserted]: #1328 +#7505 := [mp #7490 #7502]: #7500 +#15112 := [monotonicity #7505]: #15115 +#15109 := [trans #15112 #15111]: #15108 +#15104 := [monotonicity #15109]: #15107 +#15101 := [trans #15104 #15103]: #15100 +#15096 := [monotonicity #15101]: #15099 +#15093 := [trans #15096 #15095]: #15092 +#15916 := [monotonicity #15093]: #15915 +#15921 := [trans #15916 #15919]: #15920 +#15924 := [quant-intro #15921]: #15923 +#15912 := (iff #13773 #15911) +#15909 := (iff #13753 #15908) +#15906 := (iff #13748 #15903) +#15900 := (or #14339 #13725 #13739) +#15904 := (iff #15900 #15903) +#15905 := [rewrite]: #15904 +#15901 := (iff #13748 #15900) +#14336 := (iff #5601 #14339) +#14341 := (iff #4386 #14340) +#14338 := [monotonicity #15093]: #14341 +#14337 := [monotonicity #14338]: #14336 +#15902 := [monotonicity #14337]: #15901 +#15907 := [trans #15902 #15905]: #15906 +#15910 := [quant-intro #15907]: #15909 +#15913 := [monotonicity #15910]: #15912 +#15927 := [monotonicity #15913 #15924]: #15926 +#15930 := [monotonicity #15910 #15927]: #15929 +#15933 := [monotonicity #15930]: #15932 +#15898 := (iff #13720 #15895) +#15892 := (or #13417 #15889 #13697) +#15896 := (iff #15892 #15895) +#15897 := [rewrite]: #15896 +#15893 := (iff #13720 #15892) +#15890 := (iff #13691 #15889) +#15887 := (iff #13685 #15886) +#15884 := (iff #13680 #15883) +#15881 := (iff #13672 #15880) +#15878 := (iff #13663 #15877) +#15875 := (iff #13658 #15872) +#15869 := (or #12662 #12653 #13417 #13426 #15846 #13603) +#15873 := (iff #15869 #15872) +#15874 := [rewrite]: #15873 +#15870 := (iff #13658 #15869) +#15847 := (iff #13547 #15846) +#15844 := (iff #13539 #15841) +#15838 := (or #12469 #13443 #15829 #13518 #15835) +#15842 := (iff #15838 #15841) +#15843 := [rewrite]: #15842 +#15839 := (iff #13539 #15838) +#15836 := (iff #13524 #15835) +#15833 := (iff #13521 #15832) +#15799 := (iff #13432 #15796) +#15741 := (+ 4294967295::int #13433) +#15792 := (>= #15741 1::int) +#15797 := (iff #15792 #15796) +#15798 := [rewrite]: #15797 +#15793 := (iff #13432 #15792) +#15742 := (= #13434 #15741) +#15743 := [monotonicity #7505]: #15742 +#15794 := [monotonicity #15743]: #15793 +#15800 := [trans #15794 #15798]: #15799 +#15834 := [monotonicity #15800]: #15833 +#15837 := [monotonicity #15834]: #15836 +#15830 := (iff #13511 #15829) +#15827 := (iff #13508 #15826) +#15824 := (iff #13505 #15823) +#15821 := (iff #13502 #15818) +#15815 := (or #15812 #13497) +#15819 := (iff #15815 #15818) +#15820 := [rewrite]: #15819 +#15816 := (iff #13502 #15815) +#15813 := (iff #13485 #15812) +#15810 := (iff #13482 #15809) +#15807 := (iff #13477 #15804) +#15801 := (or #14339 #13454 #13468) +#15805 := (iff #15801 #15804) +#15806 := [rewrite]: #15805 +#15802 := (iff #13477 #15801) +#15803 := [monotonicity #14337]: #15802 +#15808 := [trans #15803 #15806]: #15807 +#15811 := [quant-intro #15808]: #15810 +#15814 := [monotonicity #15811]: #15813 +#15817 := [monotonicity #15814]: #15816 +#15822 := [trans #15817 #15820]: #15821 +#15825 := [monotonicity #15811 #15822]: #15824 +#15828 := [monotonicity #15825]: #15827 +#15831 := [monotonicity #15828]: #15830 +#15840 := [monotonicity #15831 #15837]: #15839 +#15845 := [trans #15840 #15843]: #15844 +#15848 := [monotonicity #15800 #15845]: #15847 +#15871 := [monotonicity #15848]: #15870 +#15876 := [trans #15871 #15874]: #15875 +#15867 := (iff #13628 #15864) +#15861 := (or #13417 #15858 #13604) +#15865 := (iff #15861 #15864) +#15866 := [rewrite]: #15865 +#15862 := (iff #13628 #15861) +#15859 := (iff #13598 #15858) +#15856 := (iff #13592 #15855) +#15853 := (iff #13587 #15852) +#15850 := (iff #13579 #15849) +#15851 := [monotonicity #15848]: #15850 +#15854 := [monotonicity #15851]: #15853 +#15857 := [monotonicity #15854]: #15856 +#15860 := [monotonicity #15857]: #15859 +#15863 := [monotonicity #15860]: #15862 +#15868 := [trans #15863 #15866]: #15867 +#15879 := [monotonicity #15868 #15876]: #15878 +#15882 := [monotonicity #15879]: #15881 +#15885 := [monotonicity #15882]: #15884 +#15888 := [monotonicity #15885]: #15887 +#15891 := [monotonicity #15888]: #15890 +#15894 := [monotonicity #15891]: #15893 +#15899 := [trans #15894 #15897]: #15898 +#15936 := [monotonicity #15899 #15933]: #15935 +#15790 := (iff #13962 #15789) +#15787 := (iff #13959 #15786) +#15784 := (iff #13955 #15781) +#15775 := (+ 255::int #13902) +#15778 := (>= #15775 0::int) +#15782 := (iff #15778 #15781) +#15783 := [rewrite]: #15782 +#15779 := (iff #13955 #15778) +#15776 := (= #13956 #15775) +#1332 := (= uf_78 255::int) +#7504 := [asserted]: #1332 +#15777 := [monotonicity #7504]: #15776 +#15780 := [monotonicity #15777]: #15779 +#15785 := [trans #15780 #15783]: #15784 +#15788 := [monotonicity #15785]: #15787 +#15791 := [monotonicity #15788]: #15790 +#15773 := (iff #13949 #15772) +#15770 := (iff #13946 #15769) +#15767 := (iff #13942 #15764) +#15758 := (+ 4294967295::int #13873) +#15761 := (>= #15758 0::int) +#15765 := (iff #15761 #15764) +#15766 := [rewrite]: #15765 +#15762 := (iff #13942 #15761) +#15759 := (= #13943 #15758) +#15760 := [monotonicity #7505]: #15759 +#15763 := [monotonicity #15760]: #15762 +#15768 := [trans #15763 #15766]: #15767 +#15771 := [monotonicity #15768]: #15770 +#15774 := [monotonicity #15771]: #15773 +#15756 := (iff #13939 #15755) +#15753 := (iff #13936 #15752) +#15750 := (iff #13933 #15747) +#15744 := (>= #15741 0::int) +#15748 := (iff #15744 #15747) +#15749 := [rewrite]: #15748 +#15745 := (iff #13933 #15744) +#15746 := [monotonicity #15743]: #15745 +#15751 := [trans #15746 #15749]: #15750 +#15754 := [monotonicity #15751]: #15753 #15757 := [monotonicity #15754]: #15756 -#15760 := [monotonicity #15757]: #15759 -#15763 := [monotonicity #15760]: #15762 -#15766 := [monotonicity #15763]: #15765 -#15769 := [monotonicity #15766]: #15768 -#15772 := [monotonicity #15769]: #15771 -#15775 := [monotonicity #15772]: #15774 -#15778 := [monotonicity #15775]: #15777 -#15781 := [monotonicity #15778]: #15780 -#15784 := [monotonicity #15781]: #15783 -#15787 := [monotonicity #15784]: #15786 -#15790 := [monotonicity #15787]: #15789 -#14792 := [not-or-elim #14776]: #14791 -#15791 := [mp #14792 #15790]: #15788 -#16558 := [mp #15791 #16557]: #16555 -#19284 := [mp~ #16558 #19283]: #19281 -#19285 := [mp #19284 #19629]: #19627 -#23120 := [mp #19285 #23119]: #23117 -#23987 := [mp #23120 #23986]: #23984 -#28241 := [unit-resolution #23987 #26494]: #23981 -#28348 := (or #23978 #23957) -decl uf_136 :: (-> T14 T5) -#26312 := (uf_58 #3079 #3011) -#26553 := (uf_136 #26312) -#26565 := (uf_24 uf_273 #26553) -#26566 := (= uf_9 #26565) -#26600 := (not #26566) -decl uf_135 :: (-> T14 T2) -#26546 := (uf_135 #26312) -#26551 := (= uf_9 #26546) -#26552 := (not #26551) -#26788 := (or #26552 #26600) -#26791 := (not #26788) -decl uf_210 :: (-> T4 T5 T2) -#26631 := (uf_210 uf_273 #26553) -#26632 := (= uf_9 #26631) -#26630 := (uf_25 uf_273 #26553) -#26610 := (= uf_26 #26630) -#26753 := (or #26610 #26632) -#26766 := (not #26753) -#26287 := (uf_15 #3011) -#26634 := (uf_14 #26287) -#26726 := (= uf_16 #26634) -#26750 := (not #26726) -#26608 := (uf_15 #26553) -#26609 := (uf_14 #26608) -#26629 := (= uf_16 #26609) -#26796 := (or #26629 #26750 #26766 #26791) -#26807 := (not #26796) -#26557 := (uf_25 uf_273 #3011) -#26558 := (= uf_26 #26557) -#26555 := (uf_210 uf_273 #3011) -#26556 := (= uf_9 #26555) -#26756 := (or #26556 #26558) -#26759 := (not #26756) -#26745 := (or #26726 #26759) -#26748 := (not #26745) -#26809 := (or #26748 #26807) -#26812 := (not #26809) -#26819 := (or #18897 #26812) -#26823 := (not #26819) -#26851 := (iff #12367 #26823) -#2376 := (uf_67 #47 #26) -#2377 := (pattern #2376) -#281 := (uf_59 #47) -#2383 := (uf_58 #281 #26) -#2397 := (uf_135 #2383) -#10938 := (= uf_9 #2397) -#10941 := (not #10938) -#2384 := (uf_136 #2383) -#2394 := (uf_24 #47 #2384) -#10932 := (= uf_9 #2394) -#10935 := (not #10932) -#10944 := (or #10935 #10941) -#22490 := (not #10944) -#2390 := (uf_15 #2384) -#2391 := (uf_14 #2390) -#10926 := (= uf_16 #2391) -#2387 := (uf_25 #47 #2384) -#10920 := (= uf_26 #2387) -#2385 := (uf_210 #47 #2384) -#10917 := (= uf_9 #2385) -#10923 := (or #10917 #10920) -#22489 := (not #10923) -#22491 := (or #52 #22489 #10926 #22490) -#22492 := (not #22491) -#2379 := (uf_210 #47 #26) -#10898 := (= uf_9 #2379) -#10904 := (or #3656 #10898) -#22484 := (not #10904) -#22485 := (or #36 #22484) -#22486 := (not #22485) -#22495 := (or #22486 #22492) -#22501 := (not #22495) -#22502 := (or #11522 #22501) -#22503 := (not #22502) -#10894 := (= uf_9 #2376) -#22508 := (iff #10894 #22503) -#22511 := (forall (vars (?x632 T4) (?x633 T5)) (:pat #2377) #22508) -#10929 := (not #10926) -#10978 := (and #36 #10923 #10929 #10944) -#10912 := (and #52 #10904) -#10981 := (or #10912 #10978) -#10984 := (and #3650 #10981) -#10987 := (iff #10894 #10984) -#10990 := (forall (vars (?x632 T4) (?x633 T5)) (:pat #2377) #10987) -#22512 := (iff #10990 #22511) -#22509 := (iff #10987 #22508) -#22506 := (iff #10984 #22503) -#22498 := (and #3650 #22495) -#22504 := (iff #22498 #22503) -#22505 := [rewrite]: #22504 -#22499 := (iff #10984 #22498) -#22496 := (iff #10981 #22495) -#22493 := (iff #10978 #22492) -#22494 := [rewrite]: #22493 -#22487 := (iff #10912 #22486) -#22488 := [rewrite]: #22487 -#22497 := [monotonicity #22488 #22494]: #22496 -#22500 := [monotonicity #22497]: #22499 -#22507 := [trans #22500 #22505]: #22506 -#22510 := [monotonicity #22507]: #22509 -#22513 := [quant-intro #22510]: #22512 -#18466 := (~ #10990 #10990) -#18464 := (~ #10987 #10987) -#18465 := [refl]: #18464 -#18467 := [nnf-pos #18465]: #18466 -#2398 := (= #2397 uf_9) -#2399 := (not #2398) -#2395 := (= #2394 uf_9) -#2396 := (not #2395) -#2400 := (or #2396 #2399) -#2401 := (and #2400 #36) -#2392 := (= #2391 uf_16) -#2393 := (not #2392) -#2402 := (and #2393 #2401) -#2388 := (= #2387 uf_26) -#2386 := (= #2385 uf_9) -#2389 := (or #2386 #2388) -#2403 := (and #2389 #2402) -#2380 := (= #2379 uf_9) -#2381 := (or #2380 #151) -#2382 := (and #2381 #52) -#2404 := (or #2382 #2403) -#2405 := (and #2404 #147) -#2378 := (= #2376 uf_9) -#2406 := (iff #2378 #2405) -#2407 := (forall (vars (?x632 T4) (?x633 T5)) (:pat #2377) #2406) -#10993 := (iff #2407 #10990) -#10950 := (and #36 #10944) -#10955 := (and #10929 #10950) -#10958 := (and #10923 #10955) -#10961 := (or #10912 #10958) -#10967 := (and #3650 #10961) -#10972 := (iff #10894 #10967) -#10975 := (forall (vars (?x632 T4) (?x633 T5)) (:pat #2377) #10972) -#10991 := (iff #10975 #10990) -#10988 := (iff #10972 #10987) -#10985 := (iff #10967 #10984) -#10982 := (iff #10961 #10981) -#10979 := (iff #10958 #10978) -#10980 := [rewrite]: #10979 -#10983 := [monotonicity #10980]: #10982 -#10986 := [monotonicity #10983]: #10985 -#10989 := [monotonicity #10986]: #10988 -#10992 := [quant-intro #10989]: #10991 -#10976 := (iff #2407 #10975) -#10973 := (iff #2406 #10972) -#10970 := (iff #2405 #10967) -#10964 := (and #10961 #3650) -#10968 := (iff #10964 #10967) -#10969 := [rewrite]: #10968 -#10965 := (iff #2405 #10964) -#10962 := (iff #2404 #10961) -#10959 := (iff #2403 #10958) -#10956 := (iff #2402 #10955) -#10953 := (iff #2401 #10950) -#10947 := (and #10944 #36) -#10951 := (iff #10947 #10950) -#10952 := [rewrite]: #10951 -#10948 := (iff #2401 #10947) -#10945 := (iff #2400 #10944) -#10942 := (iff #2399 #10941) -#10939 := (iff #2398 #10938) -#10940 := [rewrite]: #10939 -#10943 := [monotonicity #10940]: #10942 -#10936 := (iff #2396 #10935) -#10933 := (iff #2395 #10932) -#10934 := [rewrite]: #10933 -#10937 := [monotonicity #10934]: #10936 -#10946 := [monotonicity #10937 #10943]: #10945 -#10949 := [monotonicity #10946]: #10948 -#10954 := [trans #10949 #10952]: #10953 -#10930 := (iff #2393 #10929) -#10927 := (iff #2392 #10926) -#10928 := [rewrite]: #10927 -#10931 := [monotonicity #10928]: #10930 -#10957 := [monotonicity #10931 #10954]: #10956 -#10924 := (iff #2389 #10923) -#10921 := (iff #2388 #10920) -#10922 := [rewrite]: #10921 -#10918 := (iff #2386 #10917) -#10919 := [rewrite]: #10918 -#10925 := [monotonicity #10919 #10922]: #10924 -#10960 := [monotonicity #10925 #10957]: #10959 -#10915 := (iff #2382 #10912) -#10909 := (and #10904 #52) -#10913 := (iff #10909 #10912) -#10914 := [rewrite]: #10913 -#10910 := (iff #2382 #10909) -#10907 := (iff #2381 #10904) -#10901 := (or #10898 #3656) -#10905 := (iff #10901 #10904) -#10906 := [rewrite]: #10905 -#10902 := (iff #2381 #10901) -#10899 := (iff #2380 #10898) -#10900 := [rewrite]: #10899 -#10903 := [monotonicity #10900 #3658]: #10902 -#10908 := [trans #10903 #10906]: #10907 -#10911 := [monotonicity #10908]: #10910 -#10916 := [trans #10911 #10914]: #10915 -#10963 := [monotonicity #10916 #10960]: #10962 -#10966 := [monotonicity #10963 #3652]: #10965 -#10971 := [trans #10966 #10969]: #10970 -#10896 := (iff #2378 #10894) -#10897 := [rewrite]: #10896 -#10974 := [monotonicity #10897 #10971]: #10973 -#10977 := [quant-intro #10974]: #10976 -#10994 := [trans #10977 #10992]: #10993 -#10893 := [asserted]: #2407 -#10995 := [mp #10893 #10994]: #10990 -#18468 := [mp~ #10995 #18467]: #10990 -#22514 := [mp #18468 #22513]: #22511 -#26854 := (not #22511) -#26855 := (or #26854 #26851) -#26606 := (or #26600 #26552) -#26607 := (not #26606) -#26633 := (or #26632 #26610) -#26628 := (not #26633) -#26635 := (= #26634 uf_16) -#26684 := (not #26635) -#26685 := (or #26684 #26628 #26629 #26607) -#26554 := (not #26685) -#26559 := (or #26558 #26556) -#26560 := (not #26559) -#26544 := (or #26635 #26560) -#26636 := (not #26544) -#26637 := (or #26636 #26554) -#26681 := (not #26637) -#26713 := (or #18897 #26681) -#26714 := (not #26713) -#26725 := (iff #12367 #26714) -#26840 := (or #26854 #26725) -#26842 := (iff #26840 #26855) -#26844 := (iff #26855 #26855) -#26839 := [rewrite]: #26844 -#26852 := (iff #26725 #26851) -#26824 := (iff #26714 #26823) -#26820 := (iff #26713 #26819) -#26813 := (iff #26681 #26812) -#26810 := (iff #26637 #26809) -#26808 := (iff #26554 #26807) -#26805 := (iff #26685 #26796) -#26793 := (or #26750 #26766 #26629 #26791) -#26802 := (iff #26793 #26796) -#26804 := [rewrite]: #26802 -#26794 := (iff #26685 #26793) -#26786 := (iff #26607 #26791) -#26789 := (iff #26606 #26788) -#26790 := [rewrite]: #26789 -#26792 := [monotonicity #26790]: #26786 -#26785 := (iff #26628 #26766) -#26764 := (iff #26633 #26753) -#26765 := [rewrite]: #26764 -#26787 := [monotonicity #26765]: #26785 -#26751 := (iff #26684 #26750) -#26754 := (iff #26635 #26726) -#26755 := [rewrite]: #26754 -#26752 := [monotonicity #26755]: #26751 -#26795 := [monotonicity #26752 #26787 #26792]: #26794 -#26806 := [trans #26795 #26804]: #26805 -#26803 := [monotonicity #26806]: #26808 -#26743 := (iff #26636 #26748) -#26746 := (iff #26544 #26745) -#26742 := (iff #26560 #26759) -#26757 := (iff #26559 #26756) -#26758 := [rewrite]: #26757 -#26744 := [monotonicity #26758]: #26742 -#26747 := [monotonicity #26755 #26744]: #26746 -#26749 := [monotonicity #26747]: #26743 -#26811 := [monotonicity #26749 #26803]: #26810 -#26818 := [monotonicity #26811]: #26813 -#26822 := [monotonicity #26818]: #26820 -#26850 := [monotonicity #26822]: #26824 -#26853 := [monotonicity #26850]: #26852 -#26843 := [monotonicity #26853]: #26842 -#26845 := [trans #26843 #26839]: #26842 -#26841 := [quant-inst]: #26840 -#26846 := [mp #26841 #26845]: #26855 -#27857 := [unit-resolution #26846 #22514]: #26851 -#27023 := (not #26851) -#27960 := (or #27023 #26819) -#27858 := [hypothesis]: #23954 -decl uf_144 :: (-> T3 T3) -#24114 := (uf_144 #2952) -#26288 := (= #24114 #26287) -#26263 := (uf_48 #3011 #24114) -#26264 := (= uf_9 #26263) -#26290 := (iff #26264 #26288) -#26074 := (not #26290) -#26175 := [hypothesis]: #26074 -#1381 := (uf_15 #15) -#9506 := (= #233 #1381) -#11615 := (iff #9506 #11594) -#23676 := (forall (vars (?x712 T5) (?x713 T3)) (:pat #2662) #11615) -#11620 := (forall (vars (?x712 T5) (?x713 T3)) #11615) -#23679 := (iff #11620 #23676) -#23677 := (iff #11615 #11615) -#23678 := [refl]: #23677 -#23680 := [quant-intro #23678]: #23679 -#18739 := (~ #11620 #11620) -#18737 := (~ #11615 #11615) -#18738 := [refl]: #18737 -#18740 := [nnf-pos #18738]: #18739 -#1882 := (= #1381 #233) -#2668 := (iff #2663 #1882) -#2669 := (forall (vars (?x712 T5) (?x713 T3)) #2668) -#11621 := (iff #2669 #11620) -#11618 := (iff #2668 #11615) -#11611 := (iff #11594 #9506) -#11616 := (iff #11611 #11615) -#11617 := [rewrite]: #11616 -#11613 := (iff #2668 #11611) -#9507 := (iff #1882 #9506) -#9508 := [rewrite]: #9507 -#11614 := [monotonicity #11597 #9508]: #11613 -#11619 := [trans #11614 #11617]: #11618 -#11622 := [quant-intro #11619]: #11621 -#11610 := [asserted]: #2669 -#11625 := [mp #11610 #11622]: #11620 -#18741 := [mp~ #11625 #18740]: #11620 -#23681 := [mp #18741 #23680]: #23676 -#25432 := (not #23676) -#26067 := (or #25432 #26290) -#26289 := (iff #26288 #26264) -#26068 := (or #25432 #26289) -#26069 := (iff #26068 #26067) -#26065 := (iff #26067 #26067) -#26071 := [rewrite]: #26065 -#26291 := (iff #26289 #26290) -#26292 := [rewrite]: #26291 -#26070 := [monotonicity #26292]: #26069 -#26072 := [trans #26070 #26071]: #26069 -#26066 := [quant-inst]: #26068 -#26073 := [mp #26066 #26072]: #26067 -#26176 := [unit-resolution #26073 #23681 #26175]: false -#26214 := [lemma #26176]: #26290 -#26294 := (or #26074 #12361) -#26357 := (uf_116 #23223) -decl uf_138 :: (-> T3 int) -#26356 := (uf_138 #24114) -#26365 := (+ #26356 #26357) -#26368 := (uf_43 #24114 #26365) -#26561 := (uf_15 #26368) -#26308 := (= #26561 #26287) -#26304 := (= #26287 #26561) -#26302 := (= #3011 #26368) -#26346 := (uf_66 #23223 0::int #24114) -#26371 := (= #26346 #26368) -#26374 := (not #26371) -decl uf_139 :: (-> T5 T5 T2) -#26347 := (uf_139 #26346 #23223) -#26354 := (= uf_9 #26347) -#26355 := (not #26354) -#26380 := (or #26355 #26374) -#26385 := (not #26380) +#15739 := (iff #13921 #15738) +#15736 := (iff #13918 #15735) +#15733 := (iff #13913 #15730) +#15727 := (or #14339 #13890 #13904) +#15731 := (iff #15727 #15730) +#15732 := [rewrite]: #15731 +#15728 := (iff #13913 #15727) +#15729 := [monotonicity #14337]: #15728 +#15734 := [trans #15729 #15732]: #15733 +#15737 := [quant-intro #15734]: #15736 +#15740 := [monotonicity #15737]: #15739 +#15939 := [monotonicity #15740 #15757 #15774 #15791 #15936]: #15938 +#15944 := [trans #15939 #15942]: #15943 +#15947 := [monotonicity #15944]: #15946 +#15725 := (iff #13401 #15724) +#15722 := (iff #13398 #15721) +#15719 := (iff #13393 #15716) +#15713 := (or #14339 #13371 #13383) +#15717 := (iff #15713 #15716) +#15718 := [rewrite]: #15717 +#15714 := (iff #13393 #15713) +#15715 := [monotonicity #14337]: #15714 +#15720 := [trans #15715 #15718]: #15719 +#15723 := [quant-intro #15720]: #15722 +#15726 := [monotonicity #15723]: #15725 +#15950 := [monotonicity #15726 #15947]: #15949 +#15953 := [monotonicity #15723 #15950]: #15952 +#15956 := [monotonicity #15953]: #15955 +#15959 := [monotonicity #15956]: #15958 +#15962 := [monotonicity #15959]: #15961 +#15965 := [monotonicity #15962]: #15964 +#15968 := [monotonicity #15965]: #15967 +#15971 := [monotonicity #15968]: #15970 +#15974 := [monotonicity #15971]: #15973 +#15977 := [monotonicity #15974]: #15976 +#15980 := [monotonicity #15977]: #15979 +#14257 := (not #14109) +#15252 := (iff #14257 #15251) +#15249 := (iff #14109 #15248) +#15246 := (iff #14106 #15245) +#15243 := (iff #14101 #15242) +#15240 := (iff #14095 #15239) +#15237 := (iff #14090 #15236) +#15234 := (iff #14082 #15233) +#15231 := (iff #14061 #15230) +#15228 := (iff #14058 #15227) +#15225 := (iff #14055 #15224) +#15222 := (iff #14052 #15221) +#15219 := (iff #14047 #15218) +#15216 := (iff #14039 #15213) +#15210 := (or #12914 #12905 #12896 #12887 #12878 #15177 #13870 #13417 #15207 #13887 #13921 #13927 #13939 #13949 #13962 #13973) +#15214 := (iff #15210 #15213) +#15215 := [rewrite]: #15214 +#15211 := (iff #14039 #15210) +#15208 := (iff #13846 #15207) +#15205 := (iff #13841 #15202) +#15187 := (or #12096 #12087 #12078 #12069 #13417 #13784) +#15199 := (or #13417 #13698 #15187) +#15203 := (iff #15199 #15202) +#15204 := [rewrite]: #15203 +#15200 := (iff #13841 #15199) +#15197 := (iff #13816 #15187) +#15192 := (and true #15187) +#15195 := (iff #15192 #15187) +#15196 := [rewrite]: #15195 +#15193 := (iff #13816 #15192) +#15190 := (iff #13811 #15187) +#15184 := (or false #12096 #12087 #12078 #12069 #13417 #13784) +#15188 := (iff #15184 #15187) +#15189 := [rewrite]: #15188 +#15185 := (iff #13811 #15184) +#15182 := (iff #12144 false) +#15180 := (iff #12144 #3077) +#14924 := (iff up_216 true) +#10764 := [asserted]: up_216 +#14925 := [iff-true #10764]: #14924 +#15181 := [monotonicity #14925]: #15180 +#15183 := [trans #15181 #11999]: #15182 +#15186 := [monotonicity #15183]: #15185 +#15191 := [trans #15186 #15189]: #15190 +#15194 := [monotonicity #14925 #15191]: #15193 +#15198 := [trans #15194 #15196]: #15197 +#15201 := [monotonicity #15198]: #15200 +#15206 := [trans #15201 #15204]: #15205 +#15209 := [monotonicity #15206]: #15208 +#15178 := (iff #12179 #15177) +#15175 := (iff #12006 #12000) +#15170 := (and true #12000) +#15173 := (iff #15170 #12000) +#15174 := [rewrite]: #15173 +#15171 := (iff #12006 #15170) +#15160 := (iff #11908 true) +#15161 := [iff-true #14251]: #15160 +#15172 := [monotonicity #15161]: #15171 +#15176 := [trans #15172 #15174]: #15175 +#15179 := [monotonicity #15176]: #15178 +#15212 := [monotonicity #15179 #15209]: #15211 +#15217 := [trans #15212 #15215]: #15216 +#15220 := [monotonicity #15217]: #15219 +#15223 := [monotonicity #15220]: #15222 +#15226 := [monotonicity #15223]: #15225 +#15229 := [monotonicity #15226]: #15228 +#15232 := [monotonicity #15229]: #15231 +#15235 := [monotonicity #15232]: #15234 +#15238 := [monotonicity #15235]: #15237 +#15241 := [monotonicity #15238]: #15240 +#15244 := [monotonicity #15241]: #15243 +#15247 := [monotonicity #15244]: #15246 +#15250 := [monotonicity #15247]: #15249 +#15253 := [monotonicity #15250]: #15252 +#14258 := [not-or-elim #14242]: #14257 +#15254 := [mp #14258 #15253]: #15251 +#15981 := [mp #15254 #15980]: #15978 +#18701 := [mp~ #15981 #18700]: #18698 +#18702 := [mp #18701 #19046]: #19044 +#22569 := [mp #18702 #22568]: #22566 +#23436 := [mp #22569 #23435]: #23433 +#27882 := [unit-resolution #23436 #25596]: #23430 +#22674 := (or #23427 #23421) +#22672 := [def-axiom]: #22674 +#27883 := [unit-resolution #22672 #27882]: #23421 +#25774 := (uf_13 #3016) +#25775 := (= #23566 #25774) +#25748 := (uf_48 #3016 #23566) +#25749 := (= uf_9 #25748) +#25777 := (iff #25749 #25775) +#25526 := (not #25777) +#25620 := [hypothesis]: #25526 +#1390 := (uf_13 #15) +#9063 := (= #233 #1390) +#11183 := (iff #9063 #11162) +#23125 := (forall (vars (?x712 T5) (?x713 T3)) (:pat #2667) #11183) +#11188 := (forall (vars (?x712 T5) (?x713 T3)) #11183) +#23128 := (iff #11188 #23125) +#23126 := (iff #11183 #11183) +#23127 := [refl]: #23126 +#23129 := [quant-intro #23127]: #23128 +#18156 := (~ #11188 #11188) +#18154 := (~ #11183 #11183) +#18155 := [refl]: #18154 +#18157 := [nnf-pos #18155]: #18156 +#1890 := (= #1390 #233) +#2673 := (iff #2668 #1890) +#2674 := (forall (vars (?x712 T5) (?x713 T3)) #2673) +#11189 := (iff #2674 #11188) +#11186 := (iff #2673 #11183) +#11179 := (iff #11162 #9063) +#11184 := (iff #11179 #11183) +#11185 := [rewrite]: #11184 +#11181 := (iff #2673 #11179) +#9064 := (iff #1890 #9063) +#9065 := [rewrite]: #9064 +#11182 := [monotonicity #11165 #9065]: #11181 +#11187 := [trans #11182 #11185]: #11186 +#11190 := [quant-intro #11187]: #11189 +#11178 := [asserted]: #2674 +#11193 := [mp #11178 #11190]: #11188 +#18158 := [mp~ #11193 #18157]: #11188 +#23130 := [mp #18158 #23129]: #23125 +#24884 := (not #23125) +#25519 := (or #24884 #25777) +#25776 := (iff #25775 #25749) +#25520 := (or #24884 #25776) +#25521 := (iff #25520 #25519) +#25517 := (iff #25519 #25519) +#25523 := [rewrite]: #25517 +#25778 := (iff #25776 #25777) +#25779 := [rewrite]: #25778 +#25522 := [monotonicity #25779]: #25521 +#25524 := [trans #25522 #25523]: #25521 +#25518 := [quant-inst]: #25520 +#25525 := [mp #25518 #25524]: #25519 +#25621 := [unit-resolution #25525 #23130 #25620]: false +#25622 := [lemma #25621]: #25777 +#25723 := (or #25526 #11931) +#25686 := [hypothesis]: #25777 +#25756 := (not #25749) +#25677 := (iff #18314 #25756) +#25676 := (iff #11931 #25749) +#25673 := (iff #25749 #11931) +#25691 := (= #25748 #3017) +#25692 := [monotonicity #25690]: #25691 +#25675 := [monotonicity #25692]: #25673 +#25674 := [symm #25675]: #25676 +#25678 := [monotonicity #25674]: #25677 +#25687 := [hypothesis]: #18314 +#25679 := [mp #25687 #25678]: #25756 +#23681 := (uf_13 #2960) +#25743 := (= #23681 #25774) +#25739 := (= #25774 #23681) +#25737 := (= #3016 #2960) +#25714 := (= #25821 #2960) +#25703 := (= #25805 uf_274) +#25701 := (= #2961 uf_274) +#23685 := (= uf_274 #2961) +#2698 := (uf_116 #2697) +#11232 := (= #161 #2698) +#23132 := (forall (vars (?x718 T3) (?x719 int)) (:pat #23131) #11232) +#11236 := (forall (vars (?x718 T3) (?x719 int)) #11232) +#23135 := (iff #11236 #23132) +#23133 := (iff #11232 #11232) +#23134 := [refl]: #23133 +#23136 := [quant-intro #23134]: #23135 +#18171 := (~ #11236 #11236) +#18169 := (~ #11232 #11232) +#18170 := [refl]: #18169 +#18172 := [nnf-pos #18170]: #18171 +#2699 := (= #2698 #161) +#2700 := (forall (vars (?x718 T3) (?x719 int)) #2699) +#11237 := (iff #2700 #11236) +#11234 := (iff #2699 #11232) +#11235 := [rewrite]: #11234 +#11238 := [quant-intro #11235]: #11237 +#11231 := [asserted]: #2700 +#11241 := [mp #11231 #11238]: #11236 +#18173 := [mp~ #11241 #18172]: #11236 +#23137 := [mp #18173 #23136]: #23132 +#23639 := (not #23132) +#23690 := (or #23639 #23685) +#23691 := [quant-inst]: #23690 +#25681 := [unit-resolution #23691 #23137]: #23685 +#25702 := [symm #25681]: #25701 +#25700 := (= #25805 #2961) +#25697 := (= #22665 #2960) +#22666 := (= #2960 #22665) +#93 := (uf_29 #23) +#23042 := (pattern #93) +#94 := (uf_28 #93) +#3564 := (= #23 #94) +#23043 := (forall (vars (?x14 T5)) (:pat #23042) #3564) +#3567 := (forall (vars (?x14 T5)) #3564) +#23044 := (iff #3567 #23043) +#23046 := (iff #23043 #23043) +#23047 := [rewrite]: #23046 +#23045 := [rewrite]: #23044 +#23048 := [trans #23045 #23047]: #23044 +#16213 := (~ #3567 #3567) +#16203 := (~ #3564 #3564) +#16204 := [refl]: #16203 +#16269 := [nnf-pos #16204]: #16213 +#95 := (= #94 #23) +#96 := (forall (vars (?x14 T5)) #95) +#3568 := (iff #96 #3567) +#3565 := (iff #95 #3564) +#3566 := [rewrite]: #3565 +#3569 := [quant-intro #3566]: #3568 +#3563 := [asserted]: #96 +#3572 := [mp #3563 #3569]: #3567 +#16270 := [mp~ #3572 #16269]: #3567 +#23049 := [mp #16270 #23048]: #23043 +#22670 := (not #23043) +#22645 := (or #22670 #22666) +#22649 := [quant-inst]: #22645 +#25682 := [unit-resolution #22649 #23049]: #22666 +#25699 := [symm #25682]: #25697 +#25698 := [monotonicity #25699]: #25700 +#25704 := [trans #25698 #25702]: #25703 +#25734 := [monotonicity #25690 #25704]: #25714 +#25735 := (= #3016 #25821) +#25799 := (uf_66 #22665 0::int #23566) +#25824 := (= #25799 #25821) +#25827 := (not #25824) +decl uf_138 :: (-> T5 T5 T2) +#25800 := (uf_138 #25799 #22665) +#25801 := (= uf_9 #25800) +#25802 := (not #25801) +#25833 := (or #25802 #25827) +#25838 := (not #25833) #247 := (:var 1 int) -#1568 := (uf_66 #24 #247 #233) -#1569 := (pattern #1568) -#1576 := (uf_139 #1568 #24) -#8688 := (= uf_9 #1576) -#21652 := (not #8688) -#1571 := (uf_138 #233) -#1570 := (uf_116 #24) -#8678 := (+ #1570 #1571) -#8679 := (+ #247 #8678) -#8682 := (uf_43 #233 #8679) -#8685 := (= #1568 #8682) -#21651 := (not #8685) -#21653 := (or #21651 #21652) -#21654 := (not #21653) -#21657 := (forall (vars (?x375 T5) (?x376 int) (?x377 T3)) (:pat #1569) #21654) -#8691 := (and #8685 #8688) -#8694 := (forall (vars (?x375 T5) (?x376 int) (?x377 T3)) (:pat #1569) #8691) -#21658 := (iff #8694 #21657) -#21655 := (iff #8691 #21654) -#21656 := [rewrite]: #21655 -#21659 := [quant-intro #21656]: #21658 -#17817 := (~ #8694 #8694) -#17815 := (~ #8691 #8691) -#17816 := [refl]: #17815 -#17818 := [nnf-pos #17816]: #17817 -#1577 := (= #1576 uf_9) -#1572 := (+ #247 #1571) -#1573 := (+ #1570 #1572) -#1574 := (uf_43 #233 #1573) -#1575 := (= #1568 #1574) -#1578 := (and #1575 #1577) -#1579 := (forall (vars (?x375 T5) (?x376 int) (?x377 T3)) (:pat #1569) #1578) -#8695 := (iff #1579 #8694) -#8692 := (iff #1578 #8691) -#8689 := (iff #1577 #8688) -#8690 := [rewrite]: #8689 -#8686 := (iff #1575 #8685) -#8683 := (= #1574 #8682) -#8680 := (= #1573 #8679) -#8681 := [rewrite]: #8680 -#8684 := [monotonicity #8681]: #8683 -#8687 := [monotonicity #8684]: #8686 -#8693 := [monotonicity #8687 #8690]: #8692 -#8696 := [quant-intro #8693]: #8695 -#8677 := [asserted]: #1579 -#8699 := [mp #8677 #8696]: #8694 -#17819 := [mp~ #8699 #17818]: #8694 -#21660 := [mp #17819 #21659]: #21657 -#26114 := (not #21657) -#26115 := (or #26114 #26385) -#26358 := (+ #26357 #26356) -#26359 := (+ 0::int #26358) -#26360 := (uf_43 #24114 #26359) -#26361 := (= #26346 #26360) -#26362 := (not #26361) -#26363 := (or #26362 #26355) -#26364 := (not #26363) -#26116 := (or #26114 #26364) -#26122 := (iff #26116 #26115) -#26125 := (iff #26115 #26115) -#26126 := [rewrite]: #26125 -#26386 := (iff #26364 #26385) -#26383 := (iff #26363 #26380) -#26377 := (or #26374 #26355) -#26381 := (iff #26377 #26380) -#26382 := [rewrite]: #26381 -#26378 := (iff #26363 #26377) -#26375 := (iff #26362 #26374) -#26372 := (iff #26361 #26371) -#26369 := (= #26360 #26368) -#26366 := (= #26359 #26365) -#26367 := [rewrite]: #26366 -#26370 := [monotonicity #26367]: #26369 -#26373 := [monotonicity #26370]: #26372 -#26376 := [monotonicity #26373]: #26375 -#26379 := [monotonicity #26376]: #26378 -#26384 := [trans #26379 #26382]: #26383 -#26387 := [monotonicity #26384]: #26386 -#26124 := [monotonicity #26387]: #26122 -#26127 := [trans #26124 #26126]: #26122 -#26117 := [quant-inst]: #26116 -#26128 := [mp #26117 #26127]: #26115 -#26282 := [unit-resolution #26128 #21660]: #26385 -#26130 := (or #26380 #26371) -#26131 := [def-axiom]: #26130 -#26283 := [unit-resolution #26131 #26282]: #26371 -#26285 := (= #3011 #26346) -#24115 := (= uf_7 #24114) -#1349 := (uf_124 #326 #161) -#1584 := (pattern #1349) -#1597 := (uf_144 #1349) -#8734 := (= #326 #1597) -#8738 := (forall (vars (?x388 T3) (?x389 int)) (:pat #1584) #8734) -#17847 := (~ #8738 #8738) -#17845 := (~ #8734 #8734) -#17846 := [refl]: #17845 -#17848 := [nnf-pos #17846]: #17847 -#1598 := (= #1597 #326) -#1599 := (forall (vars (?x388 T3) (?x389 int)) (:pat #1584) #1598) -#8739 := (iff #1599 #8738) -#8736 := (iff #1598 #8734) -#8737 := [rewrite]: #8736 -#8740 := [quant-intro #8737]: #8739 -#8733 := [asserted]: #1599 -#8743 := [mp #8733 #8740]: #8738 -#17849 := [mp~ #8743 #17848]: #8738 -#24118 := (not #8738) -#24119 := (or #24118 #24115) -#24120 := [quant-inst]: #24119 -#27681 := [unit-resolution #24120 #17849]: #24115 -#23226 := (= #2960 #23223) -#93 := (uf_29 #26) -#23593 := (pattern #93) -#94 := (uf_28 #93) -#3575 := (= #26 #94) -#23594 := (forall (vars (?x14 T5)) (:pat #23593) #3575) -#3578 := (forall (vars (?x14 T5)) #3575) -#23595 := (iff #3578 #23594) -#23597 := (iff #23594 #23594) -#23598 := [rewrite]: #23597 -#23596 := [rewrite]: #23595 -#23599 := [trans #23596 #23598]: #23595 -#16790 := (~ #3578 #3578) -#16780 := (~ #3575 #3575) -#16781 := [refl]: #16780 -#16851 := [nnf-pos #16781]: #16790 -#95 := (= #94 #26) -#96 := (forall (vars (?x14 T5)) #95) -#3579 := (iff #96 #3578) -#3576 := (iff #95 #3575) -#3577 := [rewrite]: #3576 -#3580 := [quant-intro #3577]: #3579 -#3574 := [asserted]: #96 -#3583 := [mp #3574 #3580]: #3578 -#16852 := [mp~ #3583 #16851]: #3578 -#23600 := [mp #16852 #23599]: #23594 -#23217 := (not #23594) -#23220 := (or #23217 #23226) -#23215 := [quant-inst]: #23220 -#26284 := [unit-resolution #23215 #23600]: #23226 -#26286 := [monotonicity #26284 #27681]: #26285 -#26303 := [trans #26286 #26283]: #26302 -#26305 := [monotonicity #26303]: #26304 -#26309 := [symm #26305]: #26308 -#26562 := (= #24114 #26561) -#26231 := (or #24181 #26562) -#26232 := [quant-inst]: #26231 -#26276 := [unit-resolution #26232 #23694]: #26562 -#26310 := [trans #26276 #26309]: #26288 -#26075 := (not #26288) -#26256 := [hypothesis]: #26290 -#26268 := (not #26264) -#26278 := (iff #18900 #26268) -#26267 := (iff #12361 #26264) -#26265 := (iff #26264 #12361) -#26258 := (= #26263 #3014) -#27682 := (= #24114 uf_7) -#27683 := [symm #27681]: #27682 -#26259 := [monotonicity #27683]: #26258 -#26266 := [monotonicity #26259]: #26265 -#26277 := [symm #26266]: #26267 -#26279 := [monotonicity #26277]: #26278 -#26257 := [hypothesis]: #18900 -#26280 := [mp #26257 #26279]: #26268 -#26106 := (or #26074 #26264 #26075) -#26108 := [def-axiom]: #26106 -#26281 := [unit-resolution #26108 #26280 #26256]: #26075 -#26311 := [unit-resolution #26281 #26310]: false -#26295 := [lemma #26311]: #26294 -#27925 := [unit-resolution #26295 #26214]: #12361 -#27926 := [hypothesis]: #23981 -#23241 := (or #23978 #23972) -#23222 := [def-axiom]: #23241 -#27936 := [unit-resolution #23222 #27926]: #23972 -decl uf_13 :: (-> T5 T6 T2) -decl uf_10 :: (-> T4 T5 T6) -#26039 := (uf_10 uf_273 #25404) -decl uf_143 :: (-> T3 int) -#24116 := (uf_143 #2952) -#26431 := (uf_124 #24114 #24116) -#26432 := (uf_43 #26431 #2961) -#26521 := (uf_13 #26432 #26039) -#26522 := (= uf_9 #26521) -#26040 := (uf_13 #25404 #26039) -#27955 := (= #26040 #26521) -#27949 := (= #26521 #26040) -#27947 := (= #26432 #25404) -#27934 := (= #26432 #2962) -#27932 := (= #26431 #2952) -#27923 := (= #24116 uf_272) -#24117 := (= uf_272 #24116) -#1594 := (uf_143 #1349) -#8727 := (= #161 #1594) -#8730 := (forall (vars (?x386 T3) (?x387 int)) (:pat #1584) #8727) -#17842 := (~ #8730 #8730) -#17840 := (~ #8727 #8727) -#17841 := [refl]: #17840 -#17843 := [nnf-pos #17841]: #17842 -#1595 := (= #1594 #161) -#1596 := (forall (vars (?x386 T3) (?x387 int)) (:pat #1584) #1595) -#8731 := (iff #1596 #8730) -#8728 := (iff #1595 #8727) -#8729 := [rewrite]: #8728 -#8732 := [quant-intro #8729]: #8731 -#8726 := [asserted]: #1596 -#8735 := [mp #8726 #8732]: #8730 -#17844 := [mp~ #8735 #17843]: #8730 -#24123 := (not #8730) -#24124 := (or #24123 #24117) -#24125 := [quant-inst]: #24124 -#27703 := [unit-resolution #24125 #17844]: #24117 -#27931 := [symm #27703]: #27923 -#27933 := [monotonicity #27683 #27931]: #27932 -#27935 := [monotonicity #27933]: #27934 -#27948 := [trans #27935 #27939]: #27947 -#27950 := [monotonicity #27948]: #27949 -#27953 := [symm #27950]: #27955 -#26041 := (= uf_9 #26040) +#1576 := (uf_66 #21 #247 #233) +#1577 := (pattern #1576) +#1578 := (uf_138 #1576 #21) +#8244 := (= uf_9 #1578) +#21101 := (not #8244) +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) +#21100 := (not #1585) +#21102 := (or #21100 #21101) +#21103 := (not #21102) +#21106 := (forall (vars (?x375 T5) (?x376 int) (?x377 T3)) (:pat #1577) #21103) +#8250 := (and #1585 #8244) +#8255 := (forall (vars (?x375 T5) (?x376 int) (?x377 T3)) (:pat #1577) #8250) +#21107 := (iff #8255 #21106) +#21104 := (iff #8250 #21103) +#21105 := [rewrite]: #21104 +#21108 := [quant-intro #21105]: #21107 +#17234 := (~ #8255 #8255) +#17232 := (~ #8250 #8250) +#17233 := [refl]: #17232 +#17235 := [nnf-pos #17233]: #17234 +#1579 := (= #1578 uf_9) +#1586 := (and #1579 #1585) +#1587 := (forall (vars (?x375 T5) (?x376 int) (?x377 T3)) (:pat #1577) #1586) +#8256 := (iff #1587 #8255) +#8253 := (iff #1586 #8250) +#8247 := (and #8244 #1585) +#8251 := (iff #8247 #8250) +#8252 := [rewrite]: #8251 +#8248 := (iff #1586 #8247) +#8245 := (iff #1579 #8244) +#8246 := [rewrite]: #8245 +#8249 := [monotonicity #8246]: #8248 +#8254 := [trans #8249 #8252]: #8253 +#8257 := [quant-intro #8254]: #8256 +#8243 := [asserted]: #1587 +#8260 := [mp #8243 #8257]: #8255 +#17236 := [mp~ #8260 #17235]: #8255 +#21109 := [mp #17236 #21108]: #21106 +#25566 := (not #21106) +#25567 := (or #25566 #25838) +#25803 := (uf_139 #23566) +#25804 := (* 0::int #25803) +#25806 := (+ #25805 #25804) +#25807 := (uf_43 #23566 #25806) +#25808 := (= #25799 #25807) +#25809 := (not #25808) +#25810 := (or #25809 #25802) +#25811 := (not #25810) +#25568 := (or #25566 #25811) +#25574 := (iff #25568 #25567) +#25577 := (iff #25567 #25567) +#25578 := [rewrite]: #25577 +#25839 := (iff #25811 #25838) +#25836 := (iff #25810 #25833) +#25830 := (or #25827 #25802) +#25834 := (iff #25830 #25833) +#25835 := [rewrite]: #25834 +#25831 := (iff #25810 #25830) +#25828 := (iff #25809 #25827) +#25825 := (iff #25808 #25824) +#25822 := (= #25807 #25821) +#25819 := (= #25806 #25805) +#25814 := (+ #25805 0::int) +#25817 := (= #25814 #25805) +#25818 := [rewrite]: #25817 +#25815 := (= #25806 #25814) +#25812 := (= #25804 0::int) +#25813 := [rewrite]: #25812 +#25816 := [monotonicity #25813]: #25815 +#25820 := [trans #25816 #25818]: #25819 +#25823 := [monotonicity #25820]: #25822 +#25826 := [monotonicity #25823]: #25825 +#25829 := [monotonicity #25826]: #25828 +#25832 := [monotonicity #25829]: #25831 +#25837 := [trans #25832 #25835]: #25836 +#25840 := [monotonicity #25837]: #25839 +#25576 := [monotonicity #25840]: #25574 +#25579 := [trans #25576 #25578]: #25574 +#25569 := [quant-inst]: #25568 +#25580 := [mp #25569 #25579]: #25567 +#25705 := [unit-resolution #25580 #21109]: #25838 +#25582 := (or #25833 #25824) +#25583 := [def-axiom]: #25582 +#25706 := [unit-resolution #25583 #25705]: #25824 +#25712 := (= #3016 #25799) +#25713 := [monotonicity #25682 #25688]: #25712 +#25736 := [trans #25713 #25706]: #25735 +#25738 := [trans #25736 #25734]: #25737 +#25740 := [monotonicity #25738]: #25739 +#25724 := [symm #25740]: #25743 +#25725 := (= #23566 #23681) +#23682 := (= uf_7 #23681) +#23687 := (or #23633 #23682) +#23688 := [quant-inst]: #23687 +#25680 := [unit-resolution #23688 #23143]: #23682 +#25726 := [trans #25690 #25680]: #25725 +#25727 := [trans #25726 #25724]: #25775 +#25527 := (not #25775) +#25558 := (or #25526 #25749 #25527) +#25560 := [def-axiom]: #25558 +#25728 := [unit-resolution #25560 #25727 #25679 #25686]: false +#25729 := [lemma #25728]: #25723 +#27884 := [unit-resolution #25729 #25622]: #11931 +decl uf_15 :: (-> T5 T6 T2) +decl uf_16 :: (-> T4 T5 T6) +#25491 := (uf_16 uf_273 #24856) +#25967 := (uf_15 #25880 #25491) +#25968 := (= uf_9 #25967) +#25492 := (uf_15 #24856 #25491) +#25493 := (= uf_9 #25492) +#26002 := (or #13182 #25493) +#25873 := [monotonicity #25871 #25871]: #25872 +#25947 := [symm #25873]: #25946 +#25952 := [hypothesis]: #11928 +#25948 := [trans #25952 #25947]: #25435 decl uf_53 :: (-> T4 T5 T6) -#26030 := (uf_53 uf_273 #25404) -#26031 := (uf_13 #26 #26030) -#26036 := (pattern #26031) +#25482 := (uf_53 uf_273 #24856) +#25483 := (uf_15 #23 #25482) +#25488 := (pattern #25483) decl up_197 :: (-> T3 bool) -#26034 := (up_197 #25815) -#26032 := (= uf_9 #26031) -#26033 := (not #26032) +#25486 := (up_197 #25267) +#25484 := (= uf_9 #25483) +#25485 := (not #25484) decl uf_147 :: (-> T5 T6 T2) decl uf_192 :: (-> T7 T6) -decl uf_12 :: (-> T4 T5 T7) -#26026 := (uf_12 uf_273 #25404) -#26027 := (uf_192 #26026) -#26028 := (uf_147 #26 #26027) -#26029 := (= uf_9 #26028) -#26046 := (or #26029 #26033 #26034) -#26049 := (forall (vars (?x577 T5)) (:pat #26036) #26046) -#26052 := (not #26049) -#26042 := (not #26041) -#26055 := (or #25880 #26042 #26052) -#26058 := (not #26055) -#27945 := (= #3009 #25982) -#27946 := [symm #27943]: #27945 -#23240 := (or #23978 #12355) -#23229 := [def-axiom]: #23240 -#27938 := [unit-resolution #23229 #27926]: #12355 -#27924 := [trans #27938 #27946]: #25983 -#25988 := (or #26022 #25981 #25999) -#26021 := [def-axiom]: #25988 -#27927 := [unit-resolution #26021 #27924 #27937]: #25981 -#26061 := (or #26002 #26058) +decl uf_11 :: (-> T4 T5 T7) +#25478 := (uf_11 uf_273 #24856) +#25479 := (uf_192 #25478) +#25480 := (uf_147 #23 #25479) +#25481 := (= uf_9 #25480) +#25498 := (or #25481 #25485 #25486) +#25501 := (forall (vars (?x577 T5)) (:pat #25488) #25498) +#25504 := (not #25501) +#25494 := (not #25493) +#25507 := (or #25332 #25494 #25504) +#25949 := [hypothesis]: #25494 +#25658 := (or #25507 #25493) +#25685 := [def-axiom]: #25658 +#25950 := [unit-resolution #25685 #25949]: #25507 #14 := (:var 2 T4) -#2162 := (uf_196 #14 #15 #26) -#2223 := (pattern #2162) -#2224 := (uf_53 #13 #24) -#2225 := (uf_13 #26 #2224) -#2226 := (pattern #2225) -#2154 := (uf_12 #13 #15) -#2231 := (uf_192 #2154) -#2232 := (uf_147 #26 #2231) -#10478 := (= uf_9 #2232) -#10467 := (= uf_9 #2225) -#22343 := (not #10467) -#1373 := (uf_15 #24) -#2228 := (up_197 #1373) -#22358 := (or #2228 #22343 #10478) -#22363 := (forall (vars (?x577 T5)) (:pat #2226) #22358) -#22369 := (not #22363) -#2140 := (uf_10 #14 #26) -#2141 := (uf_13 #15 #2140) -#10170 := (= uf_9 #2141) -#22177 := (not #10170) -#180 := (uf_24 #14 #15) -#3758 := (= uf_9 #180) -#10821 := (not #3758) -#22370 := (or #10821 #22177 #22369) -#22371 := (not #22370) -#10219 := (= uf_9 #2162) -#10502 := (not #10219) -#22376 := (or #10502 #22371) -#22379 := (forall (vars (?x574 T4) (?x575 T5) (?x576 T5)) (:pat #2223) #22376) -#2229 := (not #2228) -#10473 := (and #2229 #10467) -#10484 := (not #10473) -#10485 := (or #10484 #10478) -#10490 := (forall (vars (?x577 T5)) (:pat #2226) #10485) -#10511 := (and #3758 #10170 #10490) -#10514 := (or #10502 #10511) -#10517 := (forall (vars (?x574 T4) (?x575 T5) (?x576 T5)) (:pat #2223) #10514) -#22380 := (iff #10517 #22379) -#22377 := (iff #10514 #22376) -#22374 := (iff #10511 #22371) -#22366 := (and #3758 #10170 #22363) -#22372 := (iff #22366 #22371) -#22373 := [rewrite]: #22372 -#22367 := (iff #10511 #22366) -#22364 := (iff #10490 #22363) -#22361 := (iff #10485 #22358) -#22344 := (or #2228 #22343) -#22355 := (or #22344 #10478) -#22359 := (iff #22355 #22358) -#22360 := [rewrite]: #22359 -#22356 := (iff #10485 #22355) -#22353 := (iff #10484 #22344) -#22345 := (not #22344) -#22348 := (not #22345) -#22351 := (iff #22348 #22344) -#22352 := [rewrite]: #22351 -#22349 := (iff #10484 #22348) -#22346 := (iff #10473 #22345) -#22347 := [rewrite]: #22346 -#22350 := [monotonicity #22347]: #22349 -#22354 := [trans #22350 #22352]: #22353 -#22357 := [monotonicity #22354]: #22356 -#22362 := [trans #22357 #22360]: #22361 -#22365 := [quant-intro #22362]: #22364 -#22368 := [monotonicity #22365]: #22367 -#22375 := [trans #22368 #22373]: #22374 -#22378 := [monotonicity #22375]: #22377 -#22381 := [quant-intro #22378]: #22380 -#18361 := (~ #10517 #10517) -#18359 := (~ #10514 #10514) -#18357 := (~ #10511 #10511) -#18355 := (~ #10490 #10490) -#18353 := (~ #10485 #10485) -#18354 := [refl]: #18353 -#18356 := [nnf-pos #18354]: #18355 -#18351 := (~ #10170 #10170) -#18352 := [refl]: #18351 -#18349 := (~ #3758 #3758) -#18350 := [refl]: #18349 -#18358 := [monotonicity #18350 #18352 #18356]: #18357 -#18347 := (~ #10502 #10502) -#18348 := [refl]: #18347 -#18360 := [monotonicity #18348 #18358]: #18359 -#18362 := [nnf-pos #18360]: #18361 -#2145 := (= #2141 uf_9) +#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) +#10048 := (= uf_9 #2237) +#10042 := (= uf_9 #2230) +#21792 := (not #10042) +#1382 := (uf_13 #21) +#2232 := (up_197 #1382) +#21807 := (or #2232 #21792 #10048) +#21812 := (forall (vars (?x577 T5)) (:pat #2231) #21807) +#21818 := (not #21812) +#2145 := (uf_16 #14 #23) +#2146 := (uf_15 #15 #2145) +#9748 := (= uf_9 #2146) +#21627 := (not #9748) +#180 := (uf_27 #14 #15) +#3742 := (= uf_9 #180) +#10385 := (not #3742) +#21819 := (or #10385 #21627 #21818) +#21820 := (not #21819) +#9796 := (= uf_9 #2166) +#10072 := (not #9796) +#21825 := (or #10072 #21820) +#21828 := (forall (vars (?x574 T4) (?x575 T5) (?x576 T5)) (:pat #2228) #21825) +#2233 := (not #2232) +#10045 := (and #2233 #10042) +#10054 := (not #10045) +#10055 := (or #10054 #10048) +#10060 := (forall (vars (?x577 T5)) (:pat #2231) #10055) +#10081 := (and #3742 #9748 #10060) +#10084 := (or #10072 #10081) +#10087 := (forall (vars (?x574 T4) (?x575 T5) (?x576 T5)) (:pat #2228) #10084) +#21829 := (iff #10087 #21828) +#21826 := (iff #10084 #21825) +#21823 := (iff #10081 #21820) +#21815 := (and #3742 #9748 #21812) +#21821 := (iff #21815 #21820) +#21822 := [rewrite]: #21821 +#21816 := (iff #10081 #21815) +#21813 := (iff #10060 #21812) +#21810 := (iff #10055 #21807) +#21793 := (or #2232 #21792) +#21804 := (or #21793 #10048) +#21808 := (iff #21804 #21807) +#21809 := [rewrite]: #21808 +#21805 := (iff #10055 #21804) +#21802 := (iff #10054 #21793) +#21794 := (not #21793) +#21797 := (not #21794) +#21800 := (iff #21797 #21793) +#21801 := [rewrite]: #21800 +#21798 := (iff #10054 #21797) +#21795 := (iff #10045 #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 +#21817 := [monotonicity #21814]: #21816 +#21824 := [trans #21817 #21822]: #21823 +#21827 := [monotonicity #21824]: #21826 +#21830 := [quant-intro #21827]: #21829 +#17778 := (~ #10087 #10087) +#17776 := (~ #10084 #10084) +#17774 := (~ #10081 #10081) +#17772 := (~ #10060 #10060) +#17770 := (~ #10055 #10055) +#17771 := [refl]: #17770 +#17773 := [nnf-pos #17771]: #17772 +#17768 := (~ #9748 #9748) +#17769 := [refl]: #17768 +#17766 := (~ #3742 #3742) +#17767 := [refl]: #17766 +#17775 := [monotonicity #17767 #17769 #17773]: #17774 +#17764 := (~ #10072 #10072) +#17765 := [refl]: #17764 +#17777 := [monotonicity #17765 #17775]: #17776 +#17779 := [nnf-pos #17777]: #17778 +#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) -#2236 := (and #184 #2145) -#2233 := (= #2232 uf_9) -#2227 := (= #2225 uf_9) -#2230 := (and #2227 #2229) -#2234 := (implies #2230 #2233) -#2235 := (forall (vars (?x577 T5)) (:pat #2226) #2234) -#2237 := (and #2235 #2236) -#2163 := (= #2162 uf_9) -#2238 := (implies #2163 #2237) -#2239 := (forall (vars (?x574 T4) (?x575 T5) (?x576 T5)) (:pat #2223) #2238) -#10520 := (iff #2239 #10517) -#10493 := (and #3758 #10170) -#10496 := (and #10490 #10493) -#10503 := (or #10502 #10496) -#10508 := (forall (vars (?x574 T4) (?x575 T5) (?x576 T5)) (:pat #2223) #10503) -#10518 := (iff #10508 #10517) -#10515 := (iff #10503 #10514) -#10512 := (iff #10496 #10511) -#10513 := [rewrite]: #10512 -#10516 := [monotonicity #10513]: #10515 -#10519 := [quant-intro #10516]: #10518 -#10509 := (iff #2239 #10508) -#10506 := (iff #2238 #10503) -#10499 := (implies #10219 #10496) -#10504 := (iff #10499 #10503) -#10505 := [rewrite]: #10504 -#10500 := (iff #2238 #10499) -#10497 := (iff #2237 #10496) -#10494 := (iff #2236 #10493) -#10171 := (iff #2145 #10170) -#10172 := [rewrite]: #10171 -#3759 := (iff #184 #3758) -#3760 := [rewrite]: #3759 -#10495 := [monotonicity #3760 #10172]: #10494 -#10491 := (iff #2235 #10490) -#10488 := (iff #2234 #10485) -#10481 := (implies #10473 #10478) -#10486 := (iff #10481 #10485) -#10487 := [rewrite]: #10486 -#10482 := (iff #2234 #10481) -#10479 := (iff #2233 #10478) -#10480 := [rewrite]: #10479 -#10476 := (iff #2230 #10473) -#10470 := (and #10467 #2229) -#10474 := (iff #10470 #10473) -#10475 := [rewrite]: #10474 -#10471 := (iff #2230 #10470) -#10468 := (iff #2227 #10467) -#10469 := [rewrite]: #10468 -#10472 := [monotonicity #10469]: #10471 -#10477 := [trans #10472 #10475]: #10476 -#10483 := [monotonicity #10477 #10480]: #10482 -#10489 := [trans #10483 #10487]: #10488 -#10492 := [quant-intro #10489]: #10491 -#10498 := [monotonicity #10492 #10495]: #10497 -#10220 := (iff #2163 #10219) -#10221 := [rewrite]: #10220 -#10501 := [monotonicity #10221 #10498]: #10500 -#10507 := [trans #10501 #10505]: #10506 -#10510 := [quant-intro #10507]: #10509 -#10521 := [trans #10510 #10519]: #10520 -#10466 := [asserted]: #2239 -#10522 := [mp #10466 #10521]: #10517 -#18363 := [mp~ #10522 #18362]: #10517 -#22382 := [mp #18363 #22381]: #22379 -#26123 := (not #22379) -#26129 := (or #26123 #26002 #26058) -#26035 := (or #26034 #26033 #26029) -#26037 := (forall (vars (?x577 T5)) (:pat #26036) #26035) -#26038 := (not #26037) -#26043 := (or #25880 #26042 #26038) -#26044 := (not #26043) -#26045 := (or #26002 #26044) -#26132 := (or #26123 #26045) -#26147 := (iff #26132 #26129) -#26144 := (or #26123 #26061) -#26145 := (iff #26144 #26129) -#26146 := [rewrite]: #26145 -#26142 := (iff #26132 #26144) -#26062 := (iff #26045 #26061) -#26059 := (iff #26044 #26058) -#26056 := (iff #26043 #26055) -#26053 := (iff #26038 #26052) -#26050 := (iff #26037 #26049) -#26047 := (iff #26035 #26046) -#26048 := [rewrite]: #26047 -#26051 := [quant-intro #26048]: #26050 -#26054 := [monotonicity #26051]: #26053 -#26057 := [monotonicity #26054]: #26056 -#26060 := [monotonicity #26057]: #26059 -#26063 := [monotonicity #26060]: #26062 -#26143 := [monotonicity #26063]: #26142 -#26148 := [trans #26143 #26146]: #26147 -#26133 := [quant-inst]: #26132 -#26149 := [mp #26133 #26148]: #26129 -#27928 := [unit-resolution #26149 #22382]: #26061 -#27929 := [unit-resolution #27928 #27927]: #26058 -#26216 := (or #26055 #26041) -#26217 := [def-axiom]: #26216 -#27930 := [unit-resolution #26217 #27929]: #26041 -#27956 := [trans #27930 #27953]: #26522 -#26523 := (not #26522) -#26711 := (or #12358 #26523) -#26511 := (uf_43 #24114 #2961) -#26512 := (uf_66 #26511 0::int #24114) -#26513 := (uf_27 uf_273 #26512) -#26514 := (= uf_9 #26513) -#26515 := (not #26514) -#26678 := (iff #18897 #26515) -#26674 := (iff #12358 #26514) -#26675 := (iff #26514 #12358) -#26687 := (= #26513 #3012) -#26683 := (= #26512 #3011) -#27689 := (= #26511 #2960) -#27687 := (= #2961 uf_274) -#24233 := (= uf_274 #2961) -#2693 := (uf_116 #2692) -#11669 := (= #161 #2693) -#23683 := (forall (vars (?x718 T3) (?x719 int)) (:pat #23682) #11669) -#11673 := (forall (vars (?x718 T3) (?x719 int)) #11669) -#23686 := (iff #11673 #23683) -#23684 := (iff #11669 #11669) -#23685 := [refl]: #23684 -#23687 := [quant-intro #23685]: #23686 -#18754 := (~ #11673 #11673) -#18752 := (~ #11669 #11669) -#18753 := [refl]: #18752 -#18755 := [nnf-pos #18753]: #18754 -#2694 := (= #2693 #161) -#2695 := (forall (vars (?x718 T3) (?x719 int)) #2694) -#11674 := (iff #2695 #11673) -#11671 := (iff #2694 #11669) -#11672 := [rewrite]: #11671 -#11675 := [quant-intro #11672]: #11674 -#11668 := [asserted]: #2695 -#11678 := [mp #11668 #11675]: #11673 -#18756 := [mp~ #11678 #18755]: #11673 -#23688 := [mp #18756 #23687]: #23683 -#24187 := (not #23683) -#24238 := (or #24187 #24233) -#24239 := [quant-inst]: #24238 -#27686 := [unit-resolution #24239 #23688]: #24233 -#27688 := [symm #27686]: #27687 -#27690 := [monotonicity #27683 #27688]: #27689 -#26686 := [monotonicity #27690 #27683]: #26683 -#26688 := [monotonicity #26686]: #26687 -#26676 := [monotonicity #26688]: #26675 -#26677 := [symm #26676]: #26674 -#26679 := [monotonicity #26677]: #26678 -#26638 := [hypothesis]: #18897 -#26680 := [mp #26638 #26679]: #26515 -#26516 := (uf_58 #3079 #26512) -#26517 := (uf_135 #26516) -#26518 := (= uf_9 #26517) -#26528 := (or #26515 #26518) -#26531 := (not #26528) -decl uf_23 :: (-> T3 T2) -#26524 := (uf_23 #24114) -#26525 := (= uf_9 #26524) -#2778 := (uf_23 uf_7) -#27721 := (= #2778 #26524) -#27718 := (= #26524 #2778) -#27719 := [monotonicity #27683]: #27718 -#27722 := [symm #27719]: #27721 -#11835 := (= uf_9 #2778) -#2779 := (= #2778 uf_9) -#11837 := (iff #2779 #11835) -#11838 := [rewrite]: #11837 -#11834 := [asserted]: #2779 -#11841 := [mp #11834 #11838]: #11835 -#27723 := [trans #11841 #27722]: #26525 -#26526 := (not #26525) -#26708 := (or #26526 #26531) -#27724 := [hypothesis]: #26522 -#26469 := (<= #24116 0::int) -#26682 := (not #26469) -#14790 := [not-or-elim #14776]: #13943 -#26452 := (* -1::int #24116) -#26584 := (+ uf_272 #26452) -#26585 := (<= #26584 0::int) -#27704 := (not #24117) -#27705 := (or #27704 #26585) -#27706 := [th-lemma]: #27705 -#27707 := [unit-resolution #27706 #27703]: #26585 -#27713 := (not #26585) -#26698 := (or #26682 #13942 #27713) -#26699 := [th-lemma]: #26698 -#26707 := [unit-resolution #26699 #27707 #14790]: #26682 -#237 := (uf_23 #233) -#758 := (:var 4 int) -#2062 := (uf_43 #233 #758) -#2063 := (uf_66 #2062 #247 #233) -#1364 := (:var 5 T4) -#2080 := (uf_25 #1364 #2063) -#1356 := (:var 3 T5) -#2060 := (uf_10 #1364 #1356) +#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) +#10090 := (iff #2244 #10087) +#10063 := (and #3742 #10060) +#10066 := (and #9748 #10063) +#10073 := (or #10072 #10066) +#10078 := (forall (vars (?x574 T4) (?x575 T5) (?x576 T5)) (:pat #2228) #10073) +#10088 := (iff #10078 #10087) +#10085 := (iff #10073 #10084) +#10082 := (iff #10066 #10081) +#10083 := [rewrite]: #10082 +#10086 := [monotonicity #10083]: #10085 +#10089 := [quant-intro #10086]: #10088 +#10079 := (iff #2244 #10078) +#10076 := (iff #2243 #10073) +#10069 := (implies #9796 #10066) +#10074 := (iff #10069 #10073) +#10075 := [rewrite]: #10074 +#10070 := (iff #2243 #10069) +#10067 := (iff #2242 #10066) +#10064 := (iff #2241 #10063) +#10061 := (iff #2240 #10060) +#10058 := (iff #2239 #10055) +#10051 := (implies #10045 #10048) +#10056 := (iff #10051 #10055) +#10057 := [rewrite]: #10056 +#10052 := (iff #2239 #10051) +#10049 := (iff #2238 #10048) +#10050 := [rewrite]: #10049 +#10046 := (iff #2235 #10045) +#10043 := (iff #2234 #10042) +#10044 := [rewrite]: #10043 +#10047 := [monotonicity #10044]: #10046 +#10053 := [monotonicity #10047 #10050]: #10052 +#10059 := [trans #10053 #10057]: #10058 +#10062 := [quant-intro #10059]: #10061 +#3743 := (iff #184 #3742) +#3744 := [rewrite]: #3743 +#10065 := [monotonicity #3744 #10062]: #10064 +#9749 := (iff #2151 #9748) +#9750 := [rewrite]: #9749 +#10068 := [monotonicity #9750 #10065]: #10067 +#9797 := (iff #2167 #9796) +#9798 := [rewrite]: #9797 +#10071 := [monotonicity #9798 #10068]: #10070 +#10077 := [trans #10071 #10075]: #10076 +#10080 := [quant-intro #10077]: #10079 +#10091 := [trans #10080 #10089]: #10090 +#10041 := [asserted]: #2244 +#10092 := [mp #10041 #10091]: #10087 +#17780 := [mp~ #10092 #17779]: #10087 +#21831 := [mp #17780 #21830]: #21828 +#25510 := (not #25507) +#25649 := (not #21828) +#25651 := (or #25649 #25454 #25510) +#25487 := (or #25486 #25485 #25481) +#25489 := (forall (vars (?x577 T5)) (:pat #25488) #25487) +#25490 := (not #25489) +#25495 := (or #25332 #25494 #25490) +#25496 := (not #25495) +#25497 := (or #25454 #25496) +#25665 := (or #25649 #25497) +#25654 := (iff #25665 #25651) +#25513 := (or #25454 #25510) +#25653 := (or #25649 #25513) +#25656 := (iff #25653 #25651) +#25657 := [rewrite]: #25656 +#25655 := (iff #25665 #25653) +#25514 := (iff #25497 #25513) +#25511 := (iff #25496 #25510) +#25508 := (iff #25495 #25507) +#25505 := (iff #25490 #25504) +#25502 := (iff #25489 #25501) +#25499 := (iff #25487 #25498) +#25500 := [rewrite]: #25499 +#25503 := [quant-intro #25500]: #25502 +#25506 := [monotonicity #25503]: #25505 +#25509 := [monotonicity #25506]: #25508 +#25512 := [monotonicity #25509]: #25511 +#25515 := [monotonicity #25512]: #25514 +#25650 := [monotonicity #25515]: #25655 +#25659 := [trans #25650 #25657]: #25654 +#25652 := [quant-inst]: #25665 +#25660 := [mp #25652 #25659]: #25651 +#26000 := [unit-resolution #25660 #21831 #25950]: #25454 +#25451 := (not #25435) +#25443 := (or #25473 #25433 #25451) +#25440 := [def-axiom]: #25443 +#26001 := [unit-resolution #25440 #26000 #25948 #25951]: false +#26003 := [lemma #26001]: #26002 +#27340 := [unit-resolution #26003 #25596]: #25493 +#25856 := (or #25494 #25968) +#25953 := (= #25492 #25967) +#25844 := (= #25967 #25492) +#25868 := (= #25880 #24856) +#25843 := [trans #25865 #25847]: #25868 +#25842 := [monotonicity #25843]: #25844 +#25954 := [symm #25842]: #25953 +#25846 := [hypothesis]: #25493 +#25955 := [trans #25846 #25954]: #25968 +#25969 := (not #25968) +#25845 := [hypothesis]: #25969 +#25956 := [unit-resolution #25845 #25955]: false +#25869 := [lemma #25956]: #25856 +#27341 := [unit-resolution #25869 #27340]: #25968 +#26259 := (or #25969 #11934) +#25957 := (uf_43 #23566 #2961) +#25958 := (uf_66 #25957 0::int #23566) +#25962 := (uf_24 uf_273 #25958) +#25963 := (= uf_9 #25962) +#25964 := (not #25963) +#26241 := (iff #18317 #25964) +#26214 := (iff #11934 #25963) +#26211 := (iff #25963 #11934) +#26209 := (= #25962 #3019) +#26203 := (= #25958 #3016) +#26201 := (= #25957 #2960) +#26202 := [monotonicity #25690 #25702]: #26201 +#26204 := [monotonicity #26202 #25690]: #26203 +#26210 := [monotonicity #26204]: #26209 +#26213 := [monotonicity #26210]: #26211 +#26215 := [symm #26213]: #26214 +#26242 := [monotonicity #26215]: #26241 +#26200 := [hypothesis]: #18317 +#26243 := [mp #26200 #26242]: #25964 +#25959 := (uf_58 #3149 #25958) +#25960 := (uf_136 #25959) +#25961 := (= uf_9 #25960) +#25974 := (or #25961 #25964) +#25977 := (not #25974) +#26244 := [hypothesis]: #25968 +decl uf_22 :: (-> T3 T2) +#25970 := (uf_22 #23566) +#25971 := (= uf_9 #25970) +#2783 := (uf_22 uf_7) +#26232 := (= #2783 #25970) +#26245 := (= #25970 #2783) +#26246 := [monotonicity #25690]: #26245 +#26233 := [symm #26246]: #26232 +#11408 := (= uf_9 #2783) +#2784 := (= #2783 uf_9) +#11410 := (iff #2784 #11408) +#11411 := [rewrite]: #11410 +#11407 := [asserted]: #2784 +#11414 := [mp #11407 #11411]: #11408 +#26234 := [trans #11414 #26233]: #25971 +#25972 := (not #25971) +#26261 := (or #25969 #25972 #25977) +#25915 := (<= #23568 0::int) +#26239 := (not #25915) +#14256 := [not-or-elim #14242]: #13405 +#25898 := (* -1::int #23568) +#26007 := (+ uf_272 #25898) +#26008 := (<= #26007 0::int) +#26230 := (not #23569) +#26236 := (or #26230 #26008) +#26237 := [th-lemma]: #26236 +#26238 := [unit-resolution #26237 #26235]: #26008 +#26240 := (not #26008) +#26251 := (or #26239 #13404 #26240) +#26252 := [th-lemma]: #26251 +#26260 := [unit-resolution #26252 #26238 #14256]: #26239 +#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) -#2058 := (uf_124 #233 #268) -#2059 := (uf_43 #2058 #758) -#2061 := (uf_13 #2059 #2060) -#2081 := (pattern #2061 #2080 #237) -#1535 := (uf_59 #1364) -#2078 := (uf_58 #1535 #2063) -#2079 := (pattern #2061 #2078 #237) -#2085 := (uf_27 #1364 #2063) -#9989 := (= uf_9 #2085) -#22088 := (not #9989) -#2082 := (uf_135 #2078) -#9983 := (= uf_9 #2082) -#22089 := (or #9983 #22088) -#22090 := (not #22089) -#2067 := (uf_55 #1364) -#9932 := (= uf_9 #2067) -#22064 := (not #9932) -#9929 := (= uf_9 #2061) -#22063 := (not #9929) -#4079 := (* -1::int #268) -#6249 := (+ #247 #4079) -#6838 := (>= #6249 0::int) -#4346 := (>= #247 0::int) -#20033 := (not #4346) -#3963 := (= uf_9 #237) -#10698 := (not #3963) -#22096 := (or #10698 #20033 #6838 #22063 #22064 #22090) -#22101 := (forall (vars (?x526 T4) (?x527 int) (?x528 T5) (?x529 int) (?x530 int) (?x531 T3)) (:pat #2079 #2081) #22096) -#9986 := (not #9983) -#9992 := (and #9986 #9989) -#8189 := (not #6838) -#9965 := (and #3963 #4346 #8189 #9929 #9932) -#9970 := (not #9965) -#10006 := (or #9970 #9992) -#10009 := (forall (vars (?x526 T4) (?x527 int) (?x528 T5) (?x529 int) (?x530 int) (?x531 T3)) (:pat #2079 #2081) #10006) -#22102 := (iff #10009 #22101) -#22099 := (iff #10006 #22096) -#22065 := (or #10698 #20033 #6838 #22063 #22064) -#22093 := (or #22065 #22090) -#22097 := (iff #22093 #22096) -#22098 := [rewrite]: #22097 -#22094 := (iff #10006 #22093) -#22091 := (iff #9992 #22090) -#22092 := [rewrite]: #22091 -#22074 := (iff #9970 #22065) -#22066 := (not #22065) -#22069 := (not #22066) -#22072 := (iff #22069 #22065) -#22073 := [rewrite]: #22072 -#22070 := (iff #9970 #22069) -#22067 := (iff #9965 #22066) -#22068 := [rewrite]: #22067 -#22071 := [monotonicity #22068]: #22070 -#22075 := [trans #22071 #22073]: #22074 -#22095 := [monotonicity #22075 #22092]: #22094 -#22100 := [trans #22095 #22098]: #22099 -#22103 := [quant-intro #22100]: #22102 -#18227 := (~ #10009 #10009) -#18225 := (~ #10006 #10006) -#18226 := [refl]: #18225 -#18228 := [nnf-pos #18226]: #18227 -#2086 := (= #2085 uf_9) -#2083 := (= #2082 uf_9) -#2084 := (not #2083) -#2087 := (and #2084 #2086) -#2068 := (= #2067 uf_9) +#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) +#9556 := (= uf_9 #2090) +#2088 := (uf_24 #1373 #2070) +#9553 := (= uf_9 #2088) +#21537 := (not #9553) +#21538 := (or #21537 #9556) +#21539 := (not #21538) +#9497 := (= uf_9 #2068) +#21513 := (not #9497) +#2073 := (uf_55 #1373) +#9494 := (= uf_9 #2073) +#21512 := (not #9494) +#4069 := (* -1::int #268) +#6133 := (+ #247 #4069) +#6730 := (>= #6133 0::int) +#4331 := (>= #247 0::int) +#19450 := (not #4331) +#3950 := (= uf_9 #237) +#10268 := (not #3950) +#21545 := (or #10268 #19450 #6730 #21512 #21513 #21539) +#21550 := (forall (vars (?x526 T4) (?x527 int) (?x528 T5) (?x529 int) (?x530 int) (?x531 T3)) (:pat #2085 #2087) #21545) +#9559 := (not #9556) +#9562 := (and #9553 #9559) +#7797 := (not #6730) +#9535 := (and #3950 #4331 #7797 #9494 #9497) +#9540 := (not #9535) +#9576 := (or #9540 #9562) +#9579 := (forall (vars (?x526 T4) (?x527 int) (?x528 T5) (?x529 int) (?x530 int) (?x531 T3)) (:pat #2085 #2087) #9576) +#21551 := (iff #9579 #21550) +#21548 := (iff #9576 #21545) +#21514 := (or #10268 #19450 #6730 #21512 #21513) +#21542 := (or #21514 #21539) +#21546 := (iff #21542 #21545) +#21547 := [rewrite]: #21546 +#21543 := (iff #9576 #21542) +#21540 := (iff #9562 #21539) +#21541 := [rewrite]: #21540 +#21523 := (iff #9540 #21514) +#21515 := (not #21514) +#21518 := (not #21515) +#21521 := (iff #21518 #21514) +#21522 := [rewrite]: #21521 +#21519 := (iff #9540 #21518) +#21516 := (iff #9535 #21515) +#21517 := [rewrite]: #21516 +#21520 := [monotonicity #21517]: #21519 +#21524 := [trans #21520 #21522]: #21523 +#21544 := [monotonicity #21524 #21541]: #21543 +#21549 := [trans #21544 #21547]: #21548 +#21552 := [quant-intro #21549]: #21551 +#17644 := (~ #9579 #9579) +#17642 := (~ #9576 #9576) +#17643 := [refl]: #17642 +#17645 := [nnf-pos #17643]: #17644 +#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) -#2069 := (and #238 #2068) -#2066 := (= #2061 uf_9) -#2070 := (and #2066 #2069) -#400 := (<= 0::int #247) -#2071 := (and #400 #2070) -#1425 := (< #247 #268) -#2072 := (and #1425 #2071) -#2088 := (implies #2072 #2087) -#2089 := (forall (vars (?x526 T4) (?x527 int) (?x528 T5) (?x529 int) (?x530 int) (?x531 T3)) (:pat #2079 #2081) #2088) -#10012 := (iff #2089 #10009) -#9935 := (and #3963 #9932) -#9938 := (and #9929 #9935) -#9941 := (and #400 #9938) -#9944 := (and #1425 #9941) -#9950 := (not #9944) -#9998 := (or #9950 #9992) -#10003 := (forall (vars (?x526 T4) (?x527 int) (?x528 T5) (?x529 int) (?x530 int) (?x531 T3)) (:pat #2079 #2081) #9998) -#10010 := (iff #10003 #10009) -#10007 := (iff #9998 #10006) -#9971 := (iff #9950 #9970) -#9968 := (iff #9944 #9965) -#9959 := (and #4346 #9938) -#9962 := (and #8189 #9959) -#9966 := (iff #9962 #9965) -#9967 := [rewrite]: #9966 -#9963 := (iff #9944 #9962) -#9960 := (iff #9941 #9959) -#4345 := (iff #400 #4346) -#4347 := [rewrite]: #4345 -#9961 := [monotonicity #4347]: #9960 -#8190 := (iff #1425 #8189) -#8191 := [rewrite]: #8190 -#9964 := [monotonicity #8191 #9961]: #9963 -#9969 := [trans #9964 #9967]: #9968 -#9972 := [monotonicity #9969]: #9971 -#10008 := [monotonicity #9972]: #10007 -#10011 := [quant-intro #10008]: #10010 -#10004 := (iff #2089 #10003) -#10001 := (iff #2088 #9998) -#9995 := (implies #9944 #9992) -#9999 := (iff #9995 #9998) -#10000 := [rewrite]: #9999 -#9996 := (iff #2088 #9995) -#9993 := (iff #2087 #9992) -#9990 := (iff #2086 #9989) -#9991 := [rewrite]: #9990 -#9987 := (iff #2084 #9986) -#9984 := (iff #2083 #9983) -#9985 := [rewrite]: #9984 -#9988 := [monotonicity #9985]: #9987 -#9994 := [monotonicity #9988 #9991]: #9993 -#9945 := (iff #2072 #9944) -#9942 := (iff #2071 #9941) -#9939 := (iff #2070 #9938) -#9936 := (iff #2069 #9935) -#9933 := (iff #2068 #9932) -#9934 := [rewrite]: #9933 -#3964 := (iff #238 #3963) -#3965 := [rewrite]: #3964 -#9937 := [monotonicity #3965 #9934]: #9936 -#9930 := (iff #2066 #9929) -#9931 := [rewrite]: #9930 -#9940 := [monotonicity #9931 #9937]: #9939 -#9943 := [monotonicity #9940]: #9942 -#9946 := [monotonicity #9943]: #9945 -#9997 := [monotonicity #9946 #9994]: #9996 -#10002 := [trans #9997 #10000]: #10001 -#10005 := [quant-intro #10002]: #10004 -#10013 := [trans #10005 #10011]: #10012 -#9982 := [asserted]: #2089 -#10014 := [mp #9982 #10013]: #10009 -#18229 := [mp~ #10014 #18228]: #10009 -#22104 := [mp #18229 #22103]: #22101 -#26542 := (not #22101) -#26613 := (or #26542 #23209 #26469 #26523 #26526 #26531) -#26519 := (or #26518 #26515) -#26520 := (not #26519) -#26453 := (+ 0::int #26452) -#26454 := (>= #26453 0::int) -#26455 := (>= 0::int 0::int) -#26456 := (not #26455) -#26527 := (or #26526 #26456 #26454 #26523 #23209 #26520) -#26614 := (or #26542 #26527) -#26601 := (iff #26614 #26613) -#26537 := (or #23209 #26469 #26523 #26526 #26531) -#26616 := (or #26542 #26537) -#26619 := (iff #26616 #26613) -#26620 := [rewrite]: #26619 -#26617 := (iff #26614 #26616) -#26540 := (iff #26527 #26537) -#26534 := (or #26526 false #26469 #26523 #23209 #26531) -#26538 := (iff #26534 #26537) -#26539 := [rewrite]: #26538 -#26535 := (iff #26527 #26534) -#26532 := (iff #26520 #26531) -#26529 := (iff #26519 #26528) -#26530 := [rewrite]: #26529 -#26533 := [monotonicity #26530]: #26532 -#26472 := (iff #26454 #26469) -#26466 := (>= #26452 0::int) -#26470 := (iff #26466 #26469) -#26471 := [rewrite]: #26470 -#26467 := (iff #26454 #26466) -#26464 := (= #26453 #26452) -#26465 := [rewrite]: #26464 -#26468 := [monotonicity #26465]: #26467 -#26473 := [trans #26468 #26471]: #26472 -#26462 := (iff #26456 false) -#26460 := (iff #26456 #3294) -#26458 := (iff #26455 true) -#26459 := [rewrite]: #26458 -#26461 := [monotonicity #26459]: #26460 -#26463 := [trans #26461 #13445]: #26462 -#26536 := [monotonicity #26463 #26473 #26533]: #26535 -#26541 := [trans #26536 #26539]: #26540 -#26618 := [monotonicity #26541]: #26617 -#26602 := [trans #26618 #26620]: #26601 -#26615 := [quant-inst]: #26614 -#26603 := [mp #26615 #26602]: #26613 -#26706 := [unit-resolution #26603 #22104 #14784 #26707 #27724]: #26708 -#26709 := [unit-resolution #26706 #27723]: #26531 -#26604 := (or #26528 #26514) -#26605 := [def-axiom]: #26604 -#26710 := [unit-resolution #26605 #26709 #26680]: false -#26712 := [lemma #26710]: #26711 -#27952 := [unit-resolution #26712 #27956]: #12358 -#23238 := (or #23975 #18897 #18900 #23969) -#23239 := [def-axiom]: #23238 -#27957 := [unit-resolution #23239 #27952 #27925 #27936]: #23969 -#23252 := (or #23966 #23960) -#23263 := [def-axiom]: #23252 -#27958 := [unit-resolution #23263 #27957]: #23960 -#23245 := (or #23963 #18900 #18906 #23957) -#23258 := [def-axiom]: #23245 -#27959 := [unit-resolution #23258 #27958 #27925 #27858]: #18906 -#27024 := (or #27023 #12367 #26819) -#27025 := [def-axiom]: #27024 -#27961 := [unit-resolution #27025 #27959]: #27960 -#27962 := [unit-resolution #27961 #27857]: #26819 -#27902 := (or #26823 #26812) -#26997 := (or #26823 #18897 #26812) -#26998 := [def-axiom]: #26997 -#27904 := [unit-resolution #26998 #27952]: #27902 -#27905 := [unit-resolution #27904 #27962]: #26812 -#26956 := (or #26809 #26796) -#26991 := [def-axiom]: #26956 -#27903 := [unit-resolution #26991 #27905]: #26796 -#27585 := (not #26518) -#27980 := (iff #27585 #26552) -#27976 := (iff #26518 #26551) -#27987 := (= #26517 #26546) -#27910 := (= #26516 #26312) -#27911 := [monotonicity #26686]: #27910 -#27988 := [monotonicity #27911]: #27987 -#27979 := [monotonicity #27988]: #27976 -#27981 := [monotonicity #27979]: #27980 -#27907 := [unit-resolution #26603 #22104 #14784 #26707 #27956]: #26708 -#27908 := [unit-resolution #27907 #27723]: #26531 -#27597 := (or #26528 #27585) -#27598 := [def-axiom]: #27597 -#27909 := [unit-resolution #27598 #27908]: #27585 -#27982 := [mp #27909 #27981]: #26552 -#26910 := (or #26788 #26551) -#26911 := [def-axiom]: #26910 -#27983 := [unit-resolution #26911 #27982]: #26788 -#24653 := (uf_14 uf_7) -#27977 := (= #24653 #26634) -#27985 := (= #26634 #24653) -#27991 := (= #26287 uf_7) -#27989 := (= #26287 #24114) -#28002 := [mp #27925 #26277]: #26264 -#26014 := (or #26074 #26268 #26288) -#26016 := [def-axiom]: #26014 -#27986 := [unit-resolution #26016 #28002 #26214]: #26288 -#27990 := [symm #27986]: #27989 -#27992 := [trans #27990 #27683]: #27991 -#27993 := [monotonicity #27992]: #27985 -#27978 := [symm #27993]: #27977 -#24654 := (= uf_16 #24653) -#24661 := (iff #11835 #24654) -#2303 := (pattern #237) -#2831 := (uf_14 #233) -#12008 := (= uf_16 #2831) -#12012 := (iff #3963 #12008) -#12015 := (forall (vars (?x761 T3)) (:pat #2303) #12012) -#18854 := (~ #12015 #12015) -#18852 := (~ #12012 #12012) -#18853 := [refl]: #18852 -#18855 := [nnf-pos #18853]: #18854 -#2844 := (= #2831 uf_16) -#2845 := (iff #238 #2844) -#2846 := (forall (vars (?x761 T3)) (:pat #2303) #2845) -#12016 := (iff #2846 #12015) -#12013 := (iff #2845 #12012) -#12010 := (iff #2844 #12008) -#12011 := [rewrite]: #12010 -#12014 := [monotonicity #3965 #12011]: #12013 -#12017 := [quant-intro #12014]: #12016 -#12007 := [asserted]: #2846 -#12020 := [mp #12007 #12017]: #12015 -#18856 := [mp~ #12020 #18855]: #12015 -#24285 := (not #12015) -#24664 := (or #24285 #24661) -#24665 := [quant-inst]: #24664 -#27984 := [unit-resolution #24665 #18856]: #24661 -#24666 := (not #24661) -#28001 := (or #24666 #24654) -#24670 := (not #11835) -#24671 := (or #24666 #24670 #24654) -#24672 := [def-axiom]: #24671 -#28003 := [unit-resolution #24672 #11841]: #28001 -#28004 := [unit-resolution #28003 #27984]: #24654 -#28005 := [trans #28004 #27978]: #26726 -#26958 := (not #26629) -#28390 := (iff #12299 #26958) -#28388 := (iff #12296 #26629) -#28355 := (iff #26629 #12296) -#28362 := (= #26609 #2955) -#28360 := (= #26608 #2952) -#28357 := (= #26608 #24234) -#28329 := (= #26553 #2962) -#28302 := (= #26553 #26432) -#26435 := (uf_66 #26432 0::int #24114) -#26436 := (uf_58 #3079 #26435) -#26437 := (uf_136 #26436) -#28300 := (= #26437 #26432) -#26438 := (= #26432 #26437) -decl up_68 :: (-> T14 bool) -#26445 := (up_68 #26436) -#26446 := (not #26445) -#26442 := (uf_27 uf_273 #26435) -#26443 := (= uf_9 #26442) -#26444 := (not #26443) -#26440 := (uf_135 #26436) -#26441 := (= uf_9 #26440) -#26439 := (not #26438) -#26474 := (or #26439 #26441 #26444 #26446) -#26477 := (not #26474) -#26449 := (uf_27 uf_273 #26432) -#26450 := (= uf_9 #26449) -#28032 := (= #2963 #26449) -#28007 := (= #26449 #2963) -#28013 := [monotonicity #27935]: #28007 -#28033 := [symm #28013]: #28032 -#28031 := [trans #14797 #28033]: #26450 -#26451 := (not #26450) -#28034 := (or #26451 #26477) -#276 := (:var 3 int) +#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) +#9582 := (iff #2095 #9579) +#9503 := (and #1435 #9497) +#9508 := (and #3950 #9503) +#9511 := (and #9494 #9508) +#9517 := (not #9511) +#9568 := (or #9517 #9562) +#9573 := (forall (vars (?x526 T4) (?x527 int) (?x528 T5) (?x529 int) (?x530 int) (?x531 T3)) (:pat #2085 #2087) #9568) +#9580 := (iff #9573 #9579) +#9577 := (iff #9568 #9576) +#9541 := (iff #9517 #9540) +#9538 := (iff #9511 #9535) +#7800 := (and #4331 #7797) +#9526 := (and #7800 #9497) +#9529 := (and #3950 #9526) +#9532 := (and #9494 #9529) +#9536 := (iff #9532 #9535) +#9537 := [rewrite]: #9536 +#9533 := (iff #9511 #9532) +#9530 := (iff #9508 #9529) +#9527 := (iff #9503 #9526) +#7801 := (iff #1435 #7800) +#7798 := (iff #1434 #7797) +#7799 := [rewrite]: #7798 +#4330 := (iff #397 #4331) +#4332 := [rewrite]: #4330 +#7802 := [monotonicity #4332 #7799]: #7801 +#9528 := [monotonicity #7802]: #9527 +#9531 := [monotonicity #9528]: #9530 +#9534 := [monotonicity #9531]: #9533 +#9539 := [trans #9534 #9537]: #9538 +#9542 := [monotonicity #9539]: #9541 +#9578 := [monotonicity #9542]: #9577 +#9581 := [quant-intro #9578]: #9580 +#9574 := (iff #2095 #9573) +#9571 := (iff #2094 #9568) +#9565 := (implies #9511 #9562) +#9569 := (iff #9565 #9568) +#9570 := [rewrite]: #9569 +#9566 := (iff #2094 #9565) +#9563 := (iff #2093 #9562) +#9560 := (iff #2092 #9559) +#9557 := (iff #2091 #9556) +#9558 := [rewrite]: #9557 +#9561 := [monotonicity #9558]: #9560 +#9554 := (iff #2089 #9553) +#9555 := [rewrite]: #9554 +#9564 := [monotonicity #9555 #9561]: #9563 +#9512 := (iff #2078 #9511) +#9509 := (iff #2077 #9508) +#9506 := (iff #2076 #9503) +#9500 := (and #9497 #1435) +#9504 := (iff #9500 #9503) +#9505 := [rewrite]: #9504 +#9501 := (iff #2076 #9500) +#9498 := (iff #2075 #9497) +#9499 := [rewrite]: #9498 +#9502 := [monotonicity #9499]: #9501 +#9507 := [trans #9502 #9505]: #9506 +#3951 := (iff #238 #3950) +#3952 := [rewrite]: #3951 +#9510 := [monotonicity #3952 #9507]: #9509 +#9495 := (iff #2074 #9494) +#9496 := [rewrite]: #9495 +#9513 := [monotonicity #9496 #9510]: #9512 +#9567 := [monotonicity #9513 #9564]: #9566 +#9572 := [trans #9567 #9570]: #9571 +#9575 := [quant-intro #9572]: #9574 +#9583 := [trans #9575 #9581]: #9582 +#9552 := [asserted]: #2095 +#9584 := [mp #9552 #9583]: #9579 +#17646 := [mp~ #9584 #17645]: #9579 +#21553 := [mp #17646 #21552]: #21550 +#25988 := (not #21550) +#25990 := (or #25988 #22661 #25915 #25969 #25972 #25977) +#25965 := (or #25964 #25961) +#25966 := (not #25965) +#25899 := (+ 0::int #25898) +#25900 := (>= #25899 0::int) +#25901 := (>= 0::int 0::int) +#25902 := (not #25901) +#25973 := (or #25972 #25902 #25900 #22661 #25969 #25966) +#25991 := (or #25988 #25973) +#25997 := (iff #25991 #25990) +#25983 := (or #22661 #25915 #25969 #25972 #25977) +#25993 := (or #25988 #25983) +#25995 := (iff #25993 #25990) +#25996 := [rewrite]: #25995 +#25994 := (iff #25991 #25993) +#25986 := (iff #25973 #25983) +#25980 := (or #25972 false #25915 #22661 #25969 #25977) +#25984 := (iff #25980 #25983) +#25985 := [rewrite]: #25984 +#25981 := (iff #25973 #25980) +#25978 := (iff #25966 #25977) +#25975 := (iff #25965 #25974) +#25976 := [rewrite]: #25975 +#25979 := [monotonicity #25976]: #25978 +#25918 := (iff #25900 #25915) +#25912 := (>= #25898 0::int) +#25916 := (iff #25912 #25915) +#25917 := [rewrite]: #25916 +#25913 := (iff #25900 #25912) +#25910 := (= #25899 #25898) +#25911 := [rewrite]: #25910 +#25914 := [monotonicity #25911]: #25913 +#25919 := [trans #25914 #25917]: #25918 +#25908 := (iff #25902 false) +#25906 := (iff #25902 #3077) +#25904 := (iff #25901 true) +#25905 := [rewrite]: #25904 +#25907 := [monotonicity #25905]: #25906 +#25909 := [trans #25907 #11999]: #25908 +#25982 := [monotonicity #25909 #25919 #25979]: #25981 +#25987 := [trans #25982 #25985]: #25986 +#25989 := [monotonicity #25987]: #25994 +#25998 := [trans #25989 #25996]: #25997 +#25992 := [quant-inst]: #25991 +#25999 := [mp #25992 #25998]: #25990 +#26262 := [unit-resolution #25999 #21553 #14251 #26260]: #26261 +#26263 := [unit-resolution #26262 #26234 #26244]: #25977 +#26033 := (or #25974 #25963) +#26034 := [def-axiom]: #26033 +#26264 := [unit-resolution #26034 #26263 #26243]: false +#26265 := [lemma #26264]: #26259 +#27342 := [unit-resolution #26265 #27341]: #11934 +#22678 := (or #23424 #18314 #18317 #23418) +#22690 := [def-axiom]: #22678 +#27885 := [unit-resolution #22690 #27342 #27884 #27883]: #23418 +#22697 := (or #23415 #23409) +#22698 := [def-axiom]: #22697 +#27886 := [unit-resolution #22698 #27885]: #23409 +#26720 := (uf_58 #3149 #22665) +#26901 := (uf_136 #26720) +#26917 := (= uf_9 #26901) +#26918 := (not #26917) +#26721 := (uf_135 #26720) +#26916 := (uf_27 uf_273 #26721) +#26911 := (= uf_9 #26916) +#26899 := (not #26911) +#27016 := (or #26899 #26918) +#27020 := (not #27016) +#25857 := (uf_13 #22665) +#26921 := (uf_12 #25857) +#26953 := (= uf_14 #26921) +#27014 := (not #26953) +#26914 := (uf_13 #26721) +#26902 := (uf_12 #26914) +#26915 := (= uf_14 #26902) +#26838 := (uf_210 uf_273 #26721) +#26903 := (= uf_9 #26838) +#26912 := (uf_25 uf_273 #26721) +#26898 := (= uf_26 #26912) +#26913 := (or #26898 #26903) +#26900 := (not #26913) +#27156 := (or #26900 #26915 #27014 #27020) +#27161 := (not #27156) +#26942 := (uf_25 uf_273 #22665) +#26943 := (= uf_26 #26942) +#26940 := (uf_210 uf_273 #22665) +#26941 := (= uf_9 #26940) +#26959 := (or #26941 #26943) +#26962 := (not #26959) +#27022 := (or #26953 #26962) +#27011 := (not #27022) +#27164 := (or #27011 #27161) +#27167 := (not #27164) +#26931 := (uf_24 uf_273 #22665) +#26932 := (= uf_9 #26931) +#27346 := (= #3019 #26931) +#27380 := (= #26931 #3019) +#27364 := (= #22665 #3016) +#27294 := (= #25799 #3016) +#27306 := [symm #25713]: #27294 +#27381 := (= #22665 #25799) +#27339 := (= #25821 #25799) +#27344 := [symm #25706]: #27339 +#27307 := (= #22665 #25821) +#27343 := (= #2960 #25821) +#27338 := [symm #25734]: #27343 +#27310 := [trans #25699 #27338]: #27307 +#27362 := [trans #27310 #27344]: #27381 +#27365 := [trans #27362 #27306]: #27364 +#27363 := [monotonicity #27365]: #27380 +#27347 := [symm #27363]: #27346 +#27348 := [trans #27342 #27347]: #26932 +#26933 := (not #26932) +#27170 := (or #26933 #27167) +#27173 := (not #27170) +#26956 := (uf_68 uf_273 #22665) +#26957 := (= uf_9 #26956) +#27176 := (iff #26957 #27173) +#27180 := (or #27179 #27176) +#26919 := (or #26918 #26899) +#26920 := (not #26919) +#26936 := (= #26921 uf_14) +#26938 := (not #26936) +#26939 := (or #26938 #26920 #26915 #26900) +#26937 := (not #26939) +#26944 := (or #26943 #26941) +#26945 := (not #26944) +#26929 := (or #26936 #26945) +#26930 := (not #26929) +#26946 := (or #26930 #26937) +#26928 := (not #26946) +#26934 := (or #26933 #26928) +#26935 := (not #26934) +#26958 := (iff #26957 #26935) +#27181 := (or #27179 #26958) +#27183 := (iff #27181 #27180) +#27185 := (iff #27180 #27180) +#27186 := [rewrite]: #27185 +#27177 := (iff #26958 #27176) +#27174 := (iff #26935 #27173) +#27171 := (iff #26934 #27170) +#27168 := (iff #26928 #27167) +#27165 := (iff #26946 #27164) +#27162 := (iff #26937 #27161) +#27159 := (iff #26939 #27156) +#27153 := (or #27014 #27020 #26915 #26900) +#27157 := (iff #27153 #27156) +#27158 := [rewrite]: #27157 +#27154 := (iff #26939 #27153) +#27018 := (iff #26920 #27020) +#27017 := (iff #26919 #27016) +#27019 := [rewrite]: #27017 +#27152 := [monotonicity #27019]: #27018 +#27012 := (iff #26938 #27014) +#26954 := (iff #26936 #26953) +#26955 := [rewrite]: #26954 +#27015 := [monotonicity #26955]: #27012 +#27155 := [monotonicity #27015 #27152]: #27154 +#27160 := [trans #27155 #27158]: #27159 +#27163 := [monotonicity #27160]: #27162 +#27009 := (iff #26930 #27011) +#27023 := (iff #26929 #27022) +#26963 := (iff #26945 #26962) +#26960 := (iff #26944 #26959) +#26961 := [rewrite]: #26960 +#27021 := [monotonicity #26961]: #26963 +#27010 := [monotonicity #26955 #27021]: #27023 +#27013 := [monotonicity #27010]: #27009 +#27166 := [monotonicity #27013 #27163]: #27165 +#27169 := [monotonicity #27166]: #27168 +#27172 := [monotonicity #27169]: #27171 +#27175 := [monotonicity #27172]: #27174 +#27178 := [monotonicity #27175]: #27177 +#27184 := [monotonicity #27178]: #27183 +#27187 := [trans #27184 #27186]: #27183 +#27182 := [quant-inst]: #27181 +#27188 := [mp #27182 #27187]: #27180 +#27372 := [unit-resolution #27188 #21963]: #27176 +#27311 := (not #26957) +#27376 := (iff #18323 #27311) +#27371 := (iff #11940 #26957) +#27378 := (iff #26957 #11940) +#27345 := (= #26956 #3022) +#27377 := [monotonicity #27365]: #27345 +#27370 := [monotonicity #27377]: #27378 +#27388 := [symm #27370]: #27371 +#27390 := [monotonicity #27388]: #27376 +#27373 := [hypothesis]: #18323 +#27369 := [mp #27373 #27390]: #27311 +#27247 := (not #27176) +#27248 := (or #27247 #26957 #27170) +#27238 := [def-axiom]: #27248 +#27391 := [unit-resolution #27238 #27369 #27372]: #27170 +#27245 := (or #27173 #26933 #27167) +#27246 := [def-axiom]: #27245 +#27368 := [unit-resolution #27246 #27391 #27348]: #27167 +#27681 := (= #2965 #26912) +#27679 := (= #26912 #2965) +#27687 := (= #26721 #2962) +#27496 := (= #26721 #25880) +#25881 := (uf_66 #25880 0::int #23566) +#25885 := (uf_58 #3149 #25881) +#25888 := (uf_135 #25885) +#27471 := (= #25888 #25880) +#25889 := (= #25880 #25888) +#25891 := (up_67 #25885) +#25892 := (not #25891) +#25890 := (not #25889) +#25886 := (uf_136 #25885) +#25887 := (= uf_9 #25886) +#25882 := (uf_24 uf_273 #25881) +#25883 := (= uf_9 #25882) +#25884 := (not #25883) +#25920 := (or #25884 #25887 #25890 #25892) +#25923 := (not #25920) +#25897 := (not #25896) +#27447 := (or #25897 #25923) +#277 := (:var 3 int) #310 := (:var 2 T3) -#1463 := (uf_124 #310 #247) -#1464 := (uf_43 #1463 #276) -#1460 := (uf_43 #310 #276) -#1461 := (uf_66 #1460 #161 #310) -#38 := (:var 4 T4) -#1466 := (uf_59 #38) -#1467 := (uf_58 #1466 #1461) -#1468 := (pattern #1467 #1464) -#1459 := (uf_41 #38) -#1462 := (uf_40 #1459 #1461) -#1465 := (pattern #1462 #1464) -#1471 := (uf_66 #1464 #161 #310) -#1474 := (uf_58 #1466 #1471) -#1479 := (uf_136 #1474) -#8354 := (= #1464 #1479) -#21428 := (not #8354) -#1476 := (uf_135 #1474) -#8348 := (= uf_9 #1476) -#1472 := (uf_27 #38 #1471) -#8345 := (= uf_9 #1472) -#21427 := (not #8345) -#1475 := (up_68 #1474) -#21426 := (not #1475) -#21429 := (or #21426 #21427 #8348 #21428) -#21430 := (not #21429) -#1469 := (uf_27 #38 #1464) -#8342 := (= uf_9 #1469) -#8377 := (not #8342) -#5373 := (* -1::int #247) -#6256 := (+ #161 #5373) -#6255 := (>= #6256 0::int) -#21436 := (or #5113 #6255 #8377 #21430) -#21441 := (forall (vars (?x346 T4) (?x347 int) (?x348 T3) (?x349 int) (?x350 int)) (:pat #1465 #1468) #21436) -#8351 := (not #8348) -#8386 := (and #1475 #8345 #8351 #8354) -#8026 := (not #6255) -#8029 := (and #4084 #8026) -#8032 := (not #8029) -#8395 := (or #8032 #8377 #8386) -#8400 := (forall (vars (?x346 T4) (?x347 int) (?x348 T3) (?x349 int) (?x350 int)) (:pat #1465 #1468) #8395) -#21442 := (iff #8400 #21441) -#21439 := (iff #8395 #21436) -#21311 := (or #5113 #6255) -#21433 := (or #21311 #8377 #21430) -#21437 := (iff #21433 #21436) -#21438 := [rewrite]: #21437 -#21434 := (iff #8395 #21433) -#21431 := (iff #8386 #21430) -#21432 := [rewrite]: #21431 -#21320 := (iff #8032 #21311) -#21312 := (not #21311) -#21315 := (not #21312) -#21318 := (iff #21315 #21311) -#21319 := [rewrite]: #21318 -#21316 := (iff #8032 #21315) -#21313 := (iff #8029 #21312) -#21314 := [rewrite]: #21313 -#21317 := [monotonicity #21314]: #21316 -#21321 := [trans #21317 #21319]: #21320 -#21435 := [monotonicity #21321 #21432]: #21434 -#21440 := [trans #21435 #21438]: #21439 -#21443 := [quant-intro #21440]: #21442 -#17588 := (~ #8400 #8400) -#17586 := (~ #8395 #8395) -#17587 := [refl]: #17586 -#17589 := [nnf-pos #17587]: #17588 -#1480 := (= #1479 #1464) +#1470 := (uf_124 #310 #247) +#1471 := (uf_43 #1470 #277) +#1467 := (uf_43 #310 #277) +#1468 := (uf_66 #1467 #161 #310) +#35 := (:var 4 T4) +#1473 := (uf_59 #35) +#1474 := (uf_58 #1473 #1468) +#1475 := (pattern #1474 #1471) +#1466 := (uf_41 #35) +#1469 := (uf_40 #1466 #1468) +#1472 := (pattern #1469 #1471) +#1478 := (uf_66 #1471 #161 #310) +#1486 := (uf_24 #35 #1478) +#7955 := (= uf_9 #1486) +#20877 := (not #7955) +#1479 := (uf_58 #1473 #1478) +#1482 := (uf_136 #1479) +#7949 := (= uf_9 #1482) +#1480 := (uf_135 #1479) +#7946 := (= #1471 #1480) +#20876 := (not #7946) +#1485 := (up_67 #1479) +#20875 := (not #1485) +#20878 := (or #20875 #20876 #7949 #20877) +#20879 := (not #20878) +#1476 := (uf_24 #35 #1471) +#7943 := (= uf_9 #1476) +#7978 := (not #7943) +#5258 := (* -1::int #247) +#6138 := (+ #161 #5258) +#6139 := (>= #6138 0::int) +#20885 := (or #4987 #6139 #7978 #20879) +#20890 := (forall (vars (?x346 T4) (?x347 int) (?x348 T3) (?x349 int) (?x350 int)) (:pat #1472 #1475) #20885) +#7952 := (not #7949) +#7987 := (and #1485 #7946 #7952 #7955) +#7642 := (not #6139) +#7645 := (and #4065 #7642) +#7648 := (not #7645) +#7996 := (or #7648 #7978 #7987) +#8001 := (forall (vars (?x346 T4) (?x347 int) (?x348 T3) (?x349 int) (?x350 int)) (:pat #1472 #1475) #7996) +#20891 := (iff #8001 #20890) +#20888 := (iff #7996 #20885) +#20760 := (or #4987 #6139) +#20882 := (or #20760 #7978 #20879) +#20886 := (iff #20882 #20885) +#20887 := [rewrite]: #20886 +#20883 := (iff #7996 #20882) +#20880 := (iff #7987 #20879) +#20881 := [rewrite]: #20880 +#20769 := (iff #7648 #20760) +#20761 := (not #20760) +#20764 := (not #20761) +#20767 := (iff #20764 #20760) +#20768 := [rewrite]: #20767 +#20765 := (iff #7648 #20764) +#20762 := (iff #7645 #20761) +#20763 := [rewrite]: #20762 +#20766 := [monotonicity #20763]: #20765 +#20770 := [trans #20766 #20768]: #20769 +#20884 := [monotonicity #20770 #20881]: #20883 +#20889 := [trans #20884 #20887]: #20888 +#20892 := [quant-intro #20889]: #20891 +#17011 := (~ #8001 #8001) +#17009 := (~ #7996 #7996) +#17010 := [refl]: #17009 +#17012 := [nnf-pos #17010]: #17011 +#1487 := (= #1486 uf_9) +#1488 := (and #1485 #1487) +#1483 := (= #1482 uf_9) +#1484 := (not #1483) +#1489 := (and #1484 #1488) +#1481 := (= #1480 #1471) +#1490 := (and #1481 #1489) +#1371 := (< #161 #247) +#1372 := (and #285 #1371) +#1491 := (implies #1372 #1490) #1477 := (= #1476 uf_9) -#1478 := (not #1477) -#1481 := (and #1478 #1480) -#1482 := (and #1475 #1481) -#1473 := (= #1472 uf_9) -#1483 := (and #1473 #1482) -#1362 := (< #161 #247) -#1363 := (and #1362 #285) -#1484 := (implies #1363 #1483) -#1470 := (= #1469 uf_9) -#1485 := (implies #1470 #1484) -#1486 := (forall (vars (?x346 T4) (?x347 int) (?x348 T3) (?x349 int) (?x350 int)) (:pat #1465 #1468) #1485) -#8403 := (iff #1486 #8400) -#8357 := (and #8351 #8354) -#8360 := (and #1475 #8357) -#8363 := (and #8345 #8360) -#7987 := (and #285 #1362) -#7996 := (not #7987) -#8369 := (or #7996 #8363) -#8378 := (or #8377 #8369) -#8383 := (forall (vars (?x346 T4) (?x347 int) (?x348 T3) (?x349 int) (?x350 int)) (:pat #1465 #1468) #8378) -#8401 := (iff #8383 #8400) -#8398 := (iff #8378 #8395) -#8389 := (or #8032 #8386) -#8392 := (or #8377 #8389) -#8396 := (iff #8392 #8395) -#8397 := [rewrite]: #8396 -#8393 := (iff #8378 #8392) -#8390 := (iff #8369 #8389) -#8387 := (iff #8363 #8386) -#8388 := [rewrite]: #8387 -#8033 := (iff #7996 #8032) -#8030 := (iff #7987 #8029) -#8027 := (iff #1362 #8026) -#8028 := [rewrite]: #8027 -#8031 := [monotonicity #4085 #8028]: #8030 -#8034 := [monotonicity #8031]: #8033 -#8391 := [monotonicity #8034 #8388]: #8390 -#8394 := [monotonicity #8391]: #8393 -#8399 := [trans #8394 #8397]: #8398 -#8402 := [quant-intro #8399]: #8401 -#8384 := (iff #1486 #8383) -#8381 := (iff #1485 #8378) -#8374 := (implies #8342 #8369) -#8379 := (iff #8374 #8378) -#8380 := [rewrite]: #8379 -#8375 := (iff #1485 #8374) -#8372 := (iff #1484 #8369) -#8366 := (implies #7987 #8363) -#8370 := (iff #8366 #8369) -#8371 := [rewrite]: #8370 -#8367 := (iff #1484 #8366) -#8364 := (iff #1483 #8363) -#8361 := (iff #1482 #8360) -#8358 := (iff #1481 #8357) -#8355 := (iff #1480 #8354) -#8356 := [rewrite]: #8355 -#8352 := (iff #1478 #8351) -#8349 := (iff #1477 #8348) -#8350 := [rewrite]: #8349 -#8353 := [monotonicity #8350]: #8352 -#8359 := [monotonicity #8353 #8356]: #8358 -#8362 := [monotonicity #8359]: #8361 -#8346 := (iff #1473 #8345) -#8347 := [rewrite]: #8346 -#8365 := [monotonicity #8347 #8362]: #8364 -#7988 := (iff #1363 #7987) +#1492 := (implies #1477 #1491) +#1493 := (forall (vars (?x346 T4) (?x347 int) (?x348 T3) (?x349 int) (?x350 int)) (:pat #1472 #1475) #1492) +#8004 := (iff #1493 #8001) +#7958 := (and #1485 #7955) +#7961 := (and #7952 #7958) +#7964 := (and #7946 #7961) +#7612 := (not #1372) +#7970 := (or #7612 #7964) +#7979 := (or #7978 #7970) +#7984 := (forall (vars (?x346 T4) (?x347 int) (?x348 T3) (?x349 int) (?x350 int)) (:pat #1472 #1475) #7979) +#8002 := (iff #7984 #8001) +#7999 := (iff #7979 #7996) +#7990 := (or #7648 #7987) +#7993 := (or #7978 #7990) +#7997 := (iff #7993 #7996) +#7998 := [rewrite]: #7997 +#7994 := (iff #7979 #7993) +#7991 := (iff #7970 #7990) +#7988 := (iff #7964 #7987) #7989 := [rewrite]: #7988 -#8368 := [monotonicity #7989 #8365]: #8367 -#8373 := [trans #8368 #8371]: #8372 -#8343 := (iff #1470 #8342) -#8344 := [rewrite]: #8343 -#8376 := [monotonicity #8344 #8373]: #8375 -#8382 := [trans #8376 #8380]: #8381 -#8385 := [quant-intro #8382]: #8384 -#8404 := [trans #8385 #8402]: #8403 -#8341 := [asserted]: #1486 -#8405 := [mp #8341 #8404]: #8400 -#17590 := [mp~ #8405 #17589]: #8400 -#21444 := [mp #17590 #21443]: #21441 -#27098 := (not #21441) -#27099 := (or #27098 #26451 #26469 #26477) -#26447 := (or #26446 #26444 #26441 #26439) -#26448 := (not #26447) -#26457 := (or #26456 #26454 #26451 #26448) -#27111 := (or #27098 #26457) -#27522 := (iff #27111 #27099) -#26483 := (or #26451 #26469 #26477) -#27450 := (or #27098 #26483) -#27435 := (iff #27450 #27099) -#27448 := [rewrite]: #27435 -#27451 := (iff #27111 #27450) -#26486 := (iff #26457 #26483) -#26480 := (or false #26469 #26451 #26477) -#26484 := (iff #26480 #26483) -#26485 := [rewrite]: #26484 -#26481 := (iff #26457 #26480) -#26478 := (iff #26448 #26477) -#26475 := (iff #26447 #26474) -#26476 := [rewrite]: #26475 -#26479 := [monotonicity #26476]: #26478 -#26482 := [monotonicity #26463 #26473 #26479]: #26481 -#26487 := [trans #26482 #26485]: #26486 -#27428 := [monotonicity #26487]: #27451 -#27523 := [trans #27428 #27448]: #27522 -#27449 := [quant-inst]: #27111 -#27524 := [mp #27449 #27523]: #27099 -#28015 := [unit-resolution #27524 #21444 #26707]: #28034 -#28035 := [unit-resolution #28015 #28031]: #26477 -#27525 := (or #26474 #26438) -#27526 := [def-axiom]: #27525 -#28036 := [unit-resolution #27526 #28035]: #26438 -#28301 := [symm #28036]: #28300 -#28299 := (= #26553 #26437) -#28298 := (= #26312 #26436) -#28308 := (= #26436 #26312) -#28325 := (= #26435 #3011) -#26269 := (uf_116 #3011) -#26270 := (uf_43 #24114 #26269) -#28320 := (= #26270 #3011) -#26271 := (= #3011 #26270) -#26500 := (or #25416 #26268 #26271) -#26272 := (or #26271 #26268) -#26501 := (or #25416 #26272) -#26508 := (iff #26501 #26500) -#26273 := (or #26268 #26271) -#26503 := (or #25416 #26273) -#26506 := (iff #26503 #26500) -#26507 := [rewrite]: #26506 -#26504 := (iff #26501 #26503) -#26274 := (iff #26272 #26273) -#26275 := [rewrite]: #26274 -#26505 := [monotonicity #26275]: #26504 -#26509 := [trans #26505 #26507]: #26508 -#26502 := [quant-inst]: #26501 -#26396 := [mp #26502 #26509]: #26500 -#28037 := [unit-resolution #26396 #18736 #28002]: #26271 -#28321 := [symm #28037]: #28320 -#28324 := (= #26435 #26270) -#26643 := (uf_116 #25404) -#26651 := (+ #26356 #26643) -#26654 := (uf_43 #24114 #26651) -#28304 := (= #26654 #26270) -#28237 := (= #26651 #26269) -#26563 := (uf_116 #26368) -#28283 := (= #26563 #26269) -#28058 := (= #26368 #3011) -#28056 := (= #26346 #3011) -#28038 := (= #23223 #2960) -#28039 := [symm #26284]: #28038 -#28057 := [monotonicity #28039 #27683]: #28056 -#28040 := (= #26368 #26346) -#28050 := [symm #26283]: #28040 -#28059 := [trans #28050 #28057]: #28058 -#28284 := [monotonicity #28059]: #28283 -#28282 := (= #26651 #26563) -#28272 := (= #26563 #26651) -#27071 := (* -1::int #26357) -#27072 := (+ #24016 #27071) -#27073 := (<= #27072 0::int) -#27070 := (= #24016 #26357) -#28067 := (= #2961 #26357) -#28087 := (= #26357 #2961) -#28088 := [monotonicity #28039]: #28087 -#28068 := [symm #28088]: #28067 -#28085 := (= #24016 #2961) -#24240 := (= #2961 #24016) -#24245 := (or #24187 #24240) -#24246 := [quant-inst]: #24245 -#28060 := [unit-resolution #24246 #23688]: #24240 -#28086 := [symm #28060]: #28085 -#28069 := [trans #28086 #28068]: #27070 -#28070 := (not #27070) -#28049 := (or #28070 #27073) -#28066 := [th-lemma]: #28049 -#28051 := [unit-resolution #28066 #28069]: #27073 -#27068 := (>= #27072 0::int) -#28052 := (or #28070 #27068) -#28053 := [th-lemma]: #28052 -#28054 := [unit-resolution #28053 #28069]: #27068 -#26567 := (* -1::int #26563) -#26568 := (+ #26357 #26567) -#26569 := (+ #26356 #26568) -#27092 := (<= #26569 0::int) -#26570 := (= #26569 0::int) -#27074 := (or #24187 #26570) -#26564 := (= #26365 #26563) -#27075 := (or #24187 #26564) -#27077 := (iff #27075 #27074) -#27083 := (iff #27074 #27074) -#27084 := [rewrite]: #27083 -#26571 := (iff #26564 #26570) -#26572 := [rewrite]: #26571 -#27078 := [monotonicity #26572]: #27077 -#27093 := [trans #27078 #27084]: #27077 -#27076 := [quant-inst]: #27075 -#27094 := [mp #27076 #27093]: #27074 -#28055 := [unit-resolution #27094 #23688]: #26570 -#28076 := (not #26570) -#28079 := (or #28076 #27092) -#28080 := [th-lemma]: #28079 -#28081 := [unit-resolution #28080 #28055]: #27092 -#27095 := (>= #26569 0::int) -#28082 := (or #28076 #27095) -#28078 := [th-lemma]: #28082 -#28083 := [unit-resolution #28078 #28055]: #27095 -#27032 := (<= #26356 1::int) -#27031 := (= #26356 1::int) -#2927 := (uf_138 uf_7) -#2928 := (= #2927 1::int) -#12262 := [asserted]: #2928 -#28084 := (= #26356 #2927) -#28103 := [monotonicity #27683]: #28084 -#28105 := [trans #28103 #12262]: #27031 -#28106 := (not #27031) -#28261 := (or #28106 #27032) -#28262 := [th-lemma]: #28261 -#28263 := [unit-resolution #28262 #28105]: #27032 -#27069 := (>= #26356 1::int) -#28264 := (or #28106 #27069) -#28265 := [th-lemma]: #28264 -#28266 := [unit-resolution #28265 #28105]: #27069 -#27890 := (* -1::int #26643) -#27891 := (+ #24016 #27890) -#27892 := (<= #27891 0::int) -#27887 := (= #24016 #26643) -#28253 := (= #26643 #24016) -#28254 := [monotonicity #27941]: #28253 -#28252 := [symm #28254]: #27887 -#28255 := (not #27887) -#28256 := (or #28255 #27892) -#28257 := [th-lemma]: #28256 -#28258 := [unit-resolution #28257 #28252]: #27892 -#27893 := (>= #27891 0::int) -#28259 := (or #28255 #27893) -#28260 := [th-lemma]: #28259 -#28271 := [unit-resolution #28260 #28252]: #27893 -#28281 := [th-lemma #28266 #28263 #28271 #28258 #28266 #28263 #28083 #28081 #28054 #28051]: #28272 -#28280 := [symm #28281]: #28282 -#28239 := [trans #28280 #28284]: #28237 -#28305 := [monotonicity #28239]: #28304 -#28322 := (= #26435 #26654) -#26639 := (uf_66 #25404 0::int #24114) -#26657 := (= #26639 #26654) -#26660 := (not #26657) -#26640 := (uf_139 #26639 #25404) -#26641 := (= uf_9 #26640) -#26642 := (not #26641) -#26666 := (or #26642 #26660) -#26671 := (not #26666) -#27881 := (or #26114 #26671) -#26644 := (+ #26643 #26356) -#26645 := (+ 0::int #26644) -#26646 := (uf_43 #24114 #26645) -#26647 := (= #26639 #26646) -#26648 := (not #26647) -#26649 := (or #26648 #26642) -#26650 := (not #26649) -#27869 := (or #26114 #26650) -#27883 := (iff #27869 #27881) -#27885 := (iff #27881 #27881) -#27886 := [rewrite]: #27885 -#26672 := (iff #26650 #26671) -#26669 := (iff #26649 #26666) -#26663 := (or #26660 #26642) -#26667 := (iff #26663 #26666) -#26668 := [rewrite]: #26667 -#26664 := (iff #26649 #26663) -#26661 := (iff #26648 #26660) -#26658 := (iff #26647 #26657) -#26655 := (= #26646 #26654) -#26652 := (= #26645 #26651) -#26653 := [rewrite]: #26652 -#26656 := [monotonicity #26653]: #26655 -#26659 := [monotonicity #26656]: #26658 -#26662 := [monotonicity #26659]: #26661 -#26665 := [monotonicity #26662]: #26664 -#26670 := [trans #26665 #26668]: #26669 -#26673 := [monotonicity #26670]: #26672 -#27884 := [monotonicity #26673]: #27883 -#27896 := [trans #27884 #27886]: #27883 -#27882 := [quant-inst]: #27869 -#27897 := [mp #27882 #27896]: #27881 -#28240 := [unit-resolution #27897 #21660]: #26671 -#27900 := (or #26666 #26657) -#27901 := [def-axiom]: #27900 -#28238 := [unit-resolution #27901 #28240]: #26657 -#28310 := (= #26435 #26639) -#28311 := [monotonicity #27948]: #28310 -#28323 := [trans #28311 #28238]: #28322 -#28319 := [trans #28323 #28305]: #28324 -#28326 := [trans #28319 #28321]: #28325 -#28296 := [monotonicity #28326]: #28308 -#28309 := [symm #28296]: #28298 -#28297 := [monotonicity #28309]: #28299 -#28303 := [trans #28297 #28301]: #28302 -#28335 := [trans #28303 #27935]: #28329 -#28334 := [monotonicity #28335]: #28357 -#28361 := [trans #28334 #28359]: #28360 -#28363 := [monotonicity #28361]: #28362 -#28356 := [monotonicity #28363]: #28355 -#28389 := [symm #28356]: #28388 -#28391 := [monotonicity #28389]: #28390 -#28392 := [mp #14796 #28391]: #26958 -#28395 := (= #2967 #26630) -#28387 := (= #26630 #2967) -#28393 := [monotonicity #28335]: #28387 -#28328 := [symm #28393]: #28395 -#28349 := [trans #14799 #28328]: #26610 -#26877 := (not #26610) -#26878 := (or #26753 #26877) -#26905 := [def-axiom]: #26878 -#28327 := [unit-resolution #26905 #28349]: #26753 -#26952 := (or #26807 #26629 #26750 #26766 #26791) -#26953 := [def-axiom]: #26952 -#28350 := [unit-resolution #26953 #28327 #28392 #28005 #27983 #27903]: false -#28351 := [lemma #28350]: #28348 -#28242 := [unit-resolution #28351 #28241]: #23957 -#23303 := (or #23954 #3022) -#23302 := [def-axiom]: #23303 -#28243 := [unit-resolution #23302 #28242]: #3022 -#28633 := (+ #3021 #18936) -#26421 := (>= #28633 0::int) -#28632 := (= #3021 #18935) -#27102 := (= #18935 #3021) -#26769 := (= #18934 #3011) -#26767 := (= ?x773!13 0::int) -#23266 := (not #18939) -#26720 := [hypothesis]: #22757 -#23257 := (or #22752 #23266) -#23268 := [def-axiom]: #23257 -#26762 := [unit-resolution #23268 #26720]: #23266 -#23178 := (or #22752 #18931) -#23264 := [def-axiom]: #23178 -#26763 := [unit-resolution #23264 #26720]: #18931 -#26768 := [th-lemma #26763 #26762]: #26767 -#27101 := [monotonicity #26768]: #26769 -#27157 := [monotonicity #27101]: #27102 -#28041 := [symm #27157]: #28632 -#28023 := (not #28632) -#28024 := (or #28023 #26421) -#28022 := [th-lemma]: #28024 -#28025 := [unit-resolution #28022 #28041]: #26421 -#23179 := (not #18938) -#23265 := (or #22752 #23179) -#23180 := [def-axiom]: #23265 -#28026 := [unit-resolution #23180 #26720]: #23179 -#26970 := (* -1::int #3021) -#26971 := (+ uf_285 #26970) -#26972 := (>= #26971 0::int) -#28244 := (or #13672 #26972) -#28245 := [th-lemma]: #28244 -#28246 := [unit-resolution #28245 #28243]: #26972 -#28641 := [th-lemma #28246 #28026 #28025]: false -#28642 := [lemma #28641]: #22752 -#23280 := (or #23954 #23948) -#23281 := [def-axiom]: #23280 -#29203 := [unit-resolution #23281 #28242]: #23948 -#28560 := [hypothesis]: #13906 -#28561 := [th-lemma #14790 #28560]: false -#28562 := [lemma #28561]: #13903 -#23300 := (or #23951 #13906 #23945) -#23301 := [def-axiom]: #23300 -#29204 := [unit-resolution #23301 #28562 #29203]: #23945 -#23309 := (or #23942 #23936) -#23310 := [def-axiom]: #23309 -#29207 := [unit-resolution #23310 #29204]: #23936 -#23328 := (or #23939 #22757 #23933) -#23305 := [def-axiom]: #23328 -#29208 := [unit-resolution #23305 #29207 #28642]: #23933 -#23321 := (or #23930 #23924) -#23322 := [def-axiom]: #23321 -#29209 := [unit-resolution #23322 #29208]: #23924 -#29210 := (or #23927 #13672 #23921) -#23317 := (or #23927 #13672 #13942 #23921) -#23318 := [def-axiom]: #23317 -#29211 := [unit-resolution #23318 #14790]: #29210 -#29212 := [unit-resolution #29211 #29209 #28243]: #23921 -#23351 := (or #23918 #13947) -#23355 := [def-axiom]: #23351 -#29213 := [unit-resolution #23355 #29212]: #13947 -#27053 := (* -1::int #26964) -#27103 := (+ uf_293 #27053) -#27104 := (<= #27103 0::int) -#26965 := (= uf_293 #26964) -#1382 := (uf_66 #15 #161 #1381) -#1383 := (pattern #1382) -#1384 := (uf_125 #1382 #15) -#8071 := (= #161 #1384) -#8075 := (forall (vars (?x319 T5) (?x320 int)) (:pat #1383) #8071) -#17553 := (~ #8075 #8075) -#17551 := (~ #8071 #8071) -#17552 := [refl]: #17551 -#17554 := [nnf-pos #17552]: #17553 -#1385 := (= #1384 #161) -#1386 := (forall (vars (?x319 T5) (?x320 int)) (:pat #1383) #1385) -#8076 := (iff #1386 #8075) -#8073 := (iff #1385 #8071) -#8074 := [rewrite]: #8073 -#8077 := [quant-intro #8074]: #8076 -#8070 := [asserted]: #1386 -#8080 := [mp #8070 #8077]: #8075 -#17555 := [mp~ #8080 #17554]: #8075 -#26411 := (not #8075) -#26968 := (or #26411 #26965) -#26969 := [quant-inst]: #26968 -#27438 := [unit-resolution #26969 #17555]: #26965 -#27439 := (not #26965) -#29214 := (or #27439 #27104) -#29215 := [th-lemma]: #29214 -#29216 := [unit-resolution #29215 #27438]: #27104 -#29217 := (not #27104) -#29218 := (or #27037 #22873 #29217) -#29219 := [th-lemma]: #29218 -#29220 := [unit-resolution #29219 #29216 #29213]: #27037 -#23345 := (or #23918 #23754) -#23338 := [def-axiom]: #23345 -#29221 := [unit-resolution #23338 #29212]: #23754 -#23365 := (or #23918 #12426) -#23366 := [def-axiom]: #23365 -#29222 := [unit-resolution #23366 #29212]: #12426 -#27373 := (+ uf_272 #27053) -#27374 := (<= #27373 0::int) -#27445 := (not #27374) -#23356 := (or #23918 #14405) -#23359 := [def-axiom]: #23356 -#29223 := [unit-resolution #23359 #29212]: #14405 -#27446 := (or #27445 #14404) -#27437 := [hypothesis]: #14405 -#27105 := (>= #27103 0::int) -#27440 := (or #27439 #27105) -#27441 := [th-lemma]: #27440 -#27442 := [unit-resolution #27441 #27438]: #27105 -#27443 := [hypothesis]: #27374 -#27444 := [th-lemma #27443 #27442 #27437]: false -#27447 := [lemma #27444]: #27446 -#29224 := [unit-resolution #27447 #29223]: #27445 -#23346 := (or #23918 #23912) -#23339 := [def-axiom]: #23346 -#29225 := [unit-resolution #23339 #29212]: #23912 -#27311 := (<= #26964 131073::int) -#23336 := (or #23918 #16332) -#23337 := [def-axiom]: #23336 -#29226 := [unit-resolution #23337 #29212]: #16332 -#29227 := (not #27105) -#29228 := (or #27311 #23042 #29227) -#29229 := [th-lemma]: #29228 -#29230 := [unit-resolution #29229 #27442 #29226]: #27311 -#27312 := (not #27311) -#27038 := (not #27037) -#27757 := (or #14049 #27038 #27312 #27374 #23037 #23759 #23915) -#27327 := (uf_66 #2960 #26964 uf_7) -#27328 := (uf_110 uf_273 #27327) -#27331 := (= uf_299 #27328) -#27162 := (= #3068 #27328) -#27733 := (= #27328 #3068) -#27638 := (= #27327 #3067) -#27584 := (= #26964 uf_293) -#27589 := [symm #27438]: #27584 -#27639 := [monotonicity #27589]: #27638 -#27734 := [monotonicity #27639]: #27733 -#27665 := [symm #27734]: #27162 -#27735 := (= uf_299 #3068) -#27640 := [hypothesis]: #12426 -#27641 := [hypothesis]: #23912 -#27352 := [hypothesis]: #14046 -#23335 := (or #23872 #14049) -#23446 := [def-axiom]: #23335 -#27720 := [unit-resolution #23446 #27352]: #23872 -#23378 := (or #23915 #23875 #23909) -#23380 := [def-axiom]: #23378 -#27731 := [unit-resolution #23380 #27720 #27641]: #23909 -#23397 := (or #23906 #12576) -#23398 := [def-axiom]: #23397 -#27666 := [unit-resolution #23398 #27731]: #12576 -#27732 := [symm #27666]: #3139 -#27736 := [trans #27732 #27640]: #27735 -#27737 := [trans #27736 #27665]: #27331 -#27738 := [hypothesis]: #27445 -#27675 := [hypothesis]: #27311 -#27739 := [hypothesis]: #27037 -#23405 := (or #23906 #23900) -#23406 := [def-axiom]: #23405 -#27740 := [unit-resolution #23406 #27731]: #23900 -#27363 := [hypothesis]: #23754 -#27108 := (+ uf_292 #13970) -#27109 := (<= #27108 0::int) -#27741 := (or #12644 #27109) -#27742 := [th-lemma]: #27741 -#27743 := [unit-resolution #27742 #27666]: #27109 -#27349 := (not #27109) -#27367 := (or #23008 #23759 #27349 #14049) -#27179 := (+ uf_294 #19528) -#27180 := (<= #27179 0::int) -#27355 := (not #27180) -#23419 := (not #19530) -#27353 := [hypothesis]: #23013 -#23443 := (or #23008 #23419) -#23444 := [def-axiom]: #23443 -#27354 := [unit-resolution #23444 #27353]: #23419 -#27356 := (or #27355 #14049 #19530) -#27357 := [th-lemma]: #27356 -#27358 := [unit-resolution #27357 #27354 #27352]: #27355 -#27191 := (+ uf_292 #19541) -#27192 := (>= #27191 0::int) -#27348 := (not #27192) -#27342 := [hypothesis]: #27109 -#23439 := (not #19543) -#23445 := (or #23008 #23439) -#23413 := [def-axiom]: #23445 -#27359 := [unit-resolution #23413 #27353]: #23439 -#27350 := (or #27348 #19543 #27349) -#27343 := [hypothesis]: #23439 -#27346 := [hypothesis]: #27192 -#27347 := [th-lemma #27346 #27343 #27342]: false -#27351 := [lemma #27347]: #27350 -#27360 := [unit-resolution #27351 #27359 #27342]: #27348 -#27364 := (or #27180 #27192) -#23383 := (or #23008 #19192) -#23438 := [def-axiom]: #23383 -#27361 := [unit-resolution #23438 #27353]: #19192 -#23457 := (or #23008 #19191) -#23437 := [def-axiom]: #23457 -#27362 := [unit-resolution #23437 #27353]: #19191 -#27205 := (or #23759 #22992 #22993 #27180 #27192) -#27168 := (+ #19196 #14431) -#27169 := (<= #27168 0::int) -#27170 := (+ ?x781!15 #14044) -#27171 := (>= #27170 0::int) -#27172 := (or #22993 #27171 #27169 #22992) -#27206 := (or #23759 #27172) -#27213 := (iff #27206 #27205) -#27200 := (or #22992 #22993 #27180 #27192) -#27208 := (or #23759 #27200) -#27211 := (iff #27208 #27205) -#27212 := [rewrite]: #27211 -#27209 := (iff #27206 #27208) -#27203 := (iff #27172 #27200) -#27197 := (or #22993 #27180 #27192 #22992) -#27201 := (iff #27197 #27200) -#27202 := [rewrite]: #27201 -#27198 := (iff #27172 #27197) -#27195 := (iff #27169 #27192) -#27185 := (+ #14431 #19196) -#27188 := (<= #27185 0::int) -#27193 := (iff #27188 #27192) -#27194 := [rewrite]: #27193 -#27189 := (iff #27169 #27188) -#27186 := (= #27168 #27185) -#27187 := [rewrite]: #27186 -#27190 := [monotonicity #27187]: #27189 -#27196 := [trans #27190 #27194]: #27195 -#27183 := (iff #27171 #27180) -#27173 := (+ #14044 ?x781!15) -#27176 := (>= #27173 0::int) -#27181 := (iff #27176 #27180) -#27182 := [rewrite]: #27181 -#27177 := (iff #27171 #27176) -#27174 := (= #27170 #27173) -#27175 := [rewrite]: #27174 -#27178 := [monotonicity #27175]: #27177 -#27184 := [trans #27178 #27182]: #27183 -#27199 := [monotonicity #27184 #27196]: #27198 -#27204 := [trans #27199 #27202]: #27203 -#27210 := [monotonicity #27204]: #27209 -#27214 := [trans #27210 #27212]: #27213 -#27207 := [quant-inst]: #27206 -#27215 := [mp #27207 #27214]: #27205 -#27365 := [unit-resolution #27215 #27363 #27362 #27361]: #27364 -#27366 := [unit-resolution #27365 #27360 #27358]: false -#27368 := [lemma #27366]: #27367 -#27753 := [unit-resolution #27368 #27743 #27352 #27363]: #23008 -#23423 := (or #23903 #23897 #23013) -#23424 := [def-axiom]: #23423 -#27754 := [unit-resolution #23424 #27753 #27740]: #23897 -#23454 := (or #23894 #23886) -#23455 := [def-axiom]: #23454 -#27755 := [unit-resolution #23455 #27754]: #23886 -#27334 := (not #27331) -#27520 := (or #23891 #27038 #27312 #27334 #27374) -#27317 := (+ #26964 #13873) -#27318 := (>= #27317 0::int) -#27326 := (= #27328 uf_299) -#27329 := (not #27326) -#27330 := (or #27329 #27038 #27318 #27312) -#27518 := (or #23891 #27330) -#27642 := (iff #27518 #27520) -#27382 := (or #27038 #27312 #27334 #27374) -#27590 := (or #23891 #27382) -#27593 := (iff #27590 #27520) -#27594 := [rewrite]: #27593 -#27591 := (iff #27518 #27590) -#27385 := (iff #27330 #27382) -#27379 := (or #27334 #27038 #27374 #27312) -#27383 := (iff #27379 #27382) -#27384 := [rewrite]: #27383 -#27380 := (iff #27330 #27379) -#27377 := (iff #27318 #27374) -#27335 := (+ #13873 #26964) -#27370 := (>= #27335 0::int) -#27375 := (iff #27370 #27374) -#27376 := [rewrite]: #27375 -#27371 := (iff #27318 #27370) -#27336 := (= #27317 #27335) -#27369 := [rewrite]: #27336 -#27372 := [monotonicity #27369]: #27371 -#27378 := [trans #27372 #27376]: #27377 -#27344 := (iff #27329 #27334) -#27332 := (iff #27326 #27331) -#27333 := [rewrite]: #27332 -#27345 := [monotonicity #27333]: #27344 -#27381 := [monotonicity #27345 #27378]: #27380 -#27386 := [trans #27381 #27384]: #27385 -#27592 := [monotonicity #27386]: #27591 -#27647 := [trans #27592 #27594]: #27642 -#27521 := [quant-inst]: #27518 -#27648 := [mp #27521 #27647]: #27520 -#27756 := [unit-resolution #27648 #27755 #27739 #27675 #27738 #27737]: false -#27758 := [lemma #27756]: #27757 -#29231 := [unit-resolution #27758 #29230 #29225 #29224 #29222 #29221 #29220]: #14049 -#23541 := (+ uf_294 #14142) -#23536 := (>= #23541 0::int) -#27163 := (uf_58 #3079 #3175) -#27762 := (uf_136 #27163) -#27763 := (uf_24 uf_273 #27762) -#27764 := (= uf_9 #27763) -#27765 := (not #27764) -#27759 := (uf_135 #27163) -#27760 := (= uf_9 #27759) -#27761 := (not #27760) -#27819 := (or #27761 #27765) -#27822 := (not #27819) -#27773 := (uf_210 uf_273 #27762) -#27774 := (= uf_9 #27773) -#27771 := (uf_25 uf_273 #27762) -#27772 := (= uf_26 #27771) -#27813 := (or #27772 #27774) -#27816 := (not #27813) -#27527 := (uf_15 #3175) -#27777 := (uf_14 #27527) -#27795 := (= uf_16 #27777) -#27810 := (not #27795) -#27768 := (uf_15 #27762) -#27769 := (uf_14 #27768) -#27770 := (= uf_16 #27769) -#27828 := (or #27770 #27810 #27816 #27822) -#27833 := (not #27828) -#27784 := (uf_25 uf_273 #3175) -#27785 := (= uf_26 #27784) -#27782 := (uf_210 uf_273 #3175) -#27783 := (= uf_9 #27782) -#27798 := (or #27783 #27785) -#27801 := (not #27798) -#27804 := (or #27795 #27801) -#27807 := (not #27804) -#27836 := (or #27807 #27833) -#27839 := (not #27836) -#27842 := (or #19008 #27839) -#27845 := (not #27842) -#27848 := (iff #12812 #27845) -#29445 := (or #26854 #27848) -#27766 := (or #27765 #27761) -#27767 := (not #27766) -#27775 := (or #27774 #27772) -#27776 := (not #27775) -#27778 := (= #27777 uf_16) -#27779 := (not #27778) -#27780 := (or #27779 #27776 #27770 #27767) -#27781 := (not #27780) -#27786 := (or #27785 #27783) -#27787 := (not #27786) -#27788 := (or #27778 #27787) -#27789 := (not #27788) -#27790 := (or #27789 #27781) -#27791 := (not #27790) -#27792 := (or #19008 #27791) -#27793 := (not #27792) -#27794 := (iff #12812 #27793) -#29439 := (or #26854 #27794) -#29438 := (iff #29439 #29445) -#29456 := (iff #29445 #29445) -#29454 := [rewrite]: #29456 -#27849 := (iff #27794 #27848) -#27846 := (iff #27793 #27845) -#27843 := (iff #27792 #27842) -#27840 := (iff #27791 #27839) -#27837 := (iff #27790 #27836) -#27834 := (iff #27781 #27833) -#27831 := (iff #27780 #27828) -#27825 := (or #27810 #27816 #27770 #27822) -#27829 := (iff #27825 #27828) -#27830 := [rewrite]: #27829 -#27826 := (iff #27780 #27825) -#27823 := (iff #27767 #27822) -#27820 := (iff #27766 #27819) -#27821 := [rewrite]: #27820 -#27824 := [monotonicity #27821]: #27823 -#27817 := (iff #27776 #27816) -#27814 := (iff #27775 #27813) -#27815 := [rewrite]: #27814 -#27818 := [monotonicity #27815]: #27817 -#27811 := (iff #27779 #27810) -#27796 := (iff #27778 #27795) -#27797 := [rewrite]: #27796 -#27812 := [monotonicity #27797]: #27811 -#27827 := [monotonicity #27812 #27818 #27824]: #27826 -#27832 := [trans #27827 #27830]: #27831 -#27835 := [monotonicity #27832]: #27834 -#27808 := (iff #27789 #27807) -#27805 := (iff #27788 #27804) -#27802 := (iff #27787 #27801) -#27799 := (iff #27786 #27798) -#27800 := [rewrite]: #27799 -#27803 := [monotonicity #27800]: #27802 -#27806 := [monotonicity #27797 #27803]: #27805 -#27809 := [monotonicity #27806]: #27808 -#27838 := [monotonicity #27809 #27835]: #27837 -#27841 := [monotonicity #27838]: #27840 -#27844 := [monotonicity #27841]: #27843 -#27847 := [monotonicity #27844]: #27846 -#27850 := [monotonicity #27847]: #27849 -#29455 := [monotonicity #27850]: #29438 -#29457 := [trans #29455 #29454]: #29438 -#29446 := [quant-inst]: #29439 -#29459 := [mp #29446 #29457]: #29445 -#29640 := [unit-resolution #29459 #22514]: #27848 -#29381 := (not #27848) -#29642 := (or #29381 #27842) -#29641 := [hypothesis]: #19017 -#29377 := (or #29381 #12812 #27842) -#29382 := [def-axiom]: #29377 -#28866 := [unit-resolution #29382 #29641]: #29642 -#29632 := [unit-resolution #28866 #29640]: #27842 -#29635 := (or #27845 #27839) -#23357 := (or #23918 #13950) -#23358 := [def-axiom]: #23357 -#28992 := [unit-resolution #23358 #29212]: #13950 -#28994 := [trans #26494 #27946]: #25983 -#28995 := [unit-resolution #26021 #28994 #27937]: #25981 -#28996 := [unit-resolution #27928 #28995]: #26058 -#28997 := [unit-resolution #26217 #28996]: #26041 -#29000 := [trans #28997 #27953]: #26522 -#27729 := (or #12803 #14243 #26523 #14046) -#27672 := [hypothesis]: #13950 -#27528 := (uf_66 #23223 uf_294 #26404) -#27529 := (uf_125 #27528 #23223) -#27558 := (* -1::int #27529) -#27667 := (+ uf_294 #27558) -#27668 := (<= #27667 0::int) -#27530 := (= uf_294 #27529) -#27533 := (or #26411 #27530) -#27534 := [quant-inst]: #27533 -#27673 := [unit-resolution #27534 #17555]: #27530 -#27676 := (not #27530) -#27677 := (or #27676 #27668) -#27678 := [th-lemma]: #27677 -#27679 := [unit-resolution #27678 #27673]: #27668 -#27549 := (>= #27529 0::int) -#27550 := (not #27549) -#27601 := (uf_66 #26511 #27529 #24114) -#27605 := (uf_58 #3079 #27601) -#27606 := (uf_135 #27605) -#27607 := (= uf_9 #27606) -#27602 := (uf_27 uf_273 #27601) -#27603 := (= uf_9 #27602) -#27604 := (not #27603) -#27611 := (or #27604 #27607) -#27699 := (iff #19008 #27604) -#27697 := (iff #12803 #27603) -#27695 := (iff #27603 #12803) -#27693 := (= #27602 #3176) -#27691 := (= #27601 #3175) -#27684 := (= #27529 uf_294) -#27685 := [symm #27673]: #27684 -#27692 := [monotonicity #27690 #27685 #27683]: #27691 -#27694 := [monotonicity #27692]: #27693 -#27696 := [monotonicity #27694]: #27695 -#27698 := [symm #27696]: #27697 -#27700 := [monotonicity #27698]: #27699 -#27680 := [hypothesis]: #19008 -#27701 := [mp #27680 #27700]: #27604 -#27636 := (or #27611 #27603) -#27637 := [def-axiom]: #27636 -#27702 := [unit-resolution #27637 #27701]: #27611 -#27559 := (+ #24116 #27558) -#27560 := (<= #27559 0::int) -#27712 := (not #27560) -#27708 := [hypothesis]: #14049 -#27669 := (>= #27667 0::int) -#27709 := (or #27676 #27669) -#27710 := [th-lemma]: #27709 -#27711 := [unit-resolution #27710 #27673]: #27669 -#27714 := (not #27669) -#27715 := (or #27712 #27713 #27714 #14046) -#27716 := [th-lemma]: #27715 -#27717 := [unit-resolution #27716 #27711 #27708 #27707]: #27712 -#27614 := (not #27611) -#27725 := (or #27550 #27560 #27614) -#27625 := (or #26542 #23209 #26523 #26526 #27550 #27560 #27614) -#27608 := (or #27607 #27604) -#27609 := (not #27608) -#27547 := (+ #27529 #26452) -#27548 := (>= #27547 0::int) -#27610 := (or #26526 #27550 #27548 #26523 #23209 #27609) -#27626 := (or #26542 #27610) -#27633 := (iff #27626 #27625) -#27620 := (or #23209 #26523 #26526 #27550 #27560 #27614) -#27628 := (or #26542 #27620) -#27631 := (iff #27628 #27625) -#27632 := [rewrite]: #27631 -#27629 := (iff #27626 #27628) -#27623 := (iff #27610 #27620) -#27617 := (or #26526 #27550 #27560 #26523 #23209 #27614) -#27621 := (iff #27617 #27620) -#27622 := [rewrite]: #27621 -#27618 := (iff #27610 #27617) -#27615 := (iff #27609 #27614) -#27612 := (iff #27608 #27611) -#27613 := [rewrite]: #27612 -#27616 := [monotonicity #27613]: #27615 -#27563 := (iff #27548 #27560) -#27552 := (+ #26452 #27529) -#27555 := (>= #27552 0::int) -#27561 := (iff #27555 #27560) -#27562 := [rewrite]: #27561 -#27556 := (iff #27548 #27555) -#27553 := (= #27547 #27552) -#27554 := [rewrite]: #27553 -#27557 := [monotonicity #27554]: #27556 -#27564 := [trans #27557 #27562]: #27563 -#27619 := [monotonicity #27564 #27616]: #27618 -#27624 := [trans #27619 #27622]: #27623 -#27630 := [monotonicity #27624]: #27629 -#27634 := [trans #27630 #27632]: #27633 -#27627 := [quant-inst]: #27626 -#27635 := [mp #27627 #27634]: #27625 -#27726 := [unit-resolution #27635 #22104 #14784 #27724 #27723]: #27725 -#27727 := [unit-resolution #27726 #27717 #27702]: #27550 -#27728 := [th-lemma #27727 #27679 #27672]: false -#27730 := [lemma #27728]: #27729 -#29001 := [unit-resolution #27730 #29231 #29000 #28992]: #12803 -#29437 := (or #27845 #19008 #27839) -#29380 := [def-axiom]: #29437 -#29636 := [unit-resolution #29380 #29001]: #29635 -#29634 := [unit-resolution #29636 #29632]: #27839 -#29590 := (or #27836 #27828) -#29500 := [def-axiom]: #29590 -#29637 := [unit-resolution #29500 #29634]: #27828 -#29123 := (= #24653 #27777) -#29378 := (= #27777 #24653) -#29114 := (= #27527 uf_7) -#28862 := (= #27527 #24114) -#27514 := (= #24114 #27527) -#27302 := (uf_48 #3175 #24114) -#27308 := (= uf_9 #27302) -#27513 := (iff #27308 #27514) -#28999 := (or #25432 #27513) -#27515 := (iff #27514 #27308) -#28993 := (or #25432 #27515) -#29008 := (iff #28993 #28999) -#28998 := (iff #28999 #28999) -#29010 := [rewrite]: #28998 -#27516 := (iff #27515 #27513) -#27517 := [rewrite]: #27516 -#29009 := [monotonicity #27517]: #29008 -#29011 := [trans #29009 #29010]: #29008 -#29007 := [quant-inst]: #28993 -#29012 := [mp #29007 #29011]: #28999 -#29076 := [unit-resolution #29012 #23681]: #27513 -#29751 := (= #3178 #27302) -#29068 := (= #27302 #3178) -#29078 := [monotonicity #27683]: #29068 -#29752 := [symm #29078]: #29751 -#27490 := (+ uf_294 #26365) -#27493 := (uf_43 #24114 #27490) -#27643 := (uf_15 #27493) -#29135 := (= #27643 #27527) -#29116 := (= #27527 #27643) -#29048 := (= #3175 #27493) -#27480 := (uf_66 #23223 uf_294 #24114) -#27496 := (= #27480 #27493) -#27499 := (not #27496) -#27481 := (uf_139 #27480 #23223) -#27482 := (= uf_9 #27481) -#27483 := (not #27482) -#27505 := (or #27483 #27499) -#27510 := (not #27505) -#29033 := (or #26114 #27510) -#27484 := (+ uf_294 #26358) -#27485 := (uf_43 #24114 #27484) -#27486 := (= #27480 #27485) -#27487 := (not #27486) -#27488 := (or #27487 #27483) -#27489 := (not #27488) -#29034 := (or #26114 #27489) -#29030 := (iff #29034 #29033) -#29036 := (iff #29033 #29033) -#29037 := [rewrite]: #29036 -#27511 := (iff #27489 #27510) -#27508 := (iff #27488 #27505) -#27502 := (or #27499 #27483) -#27506 := (iff #27502 #27505) -#27507 := [rewrite]: #27506 -#27503 := (iff #27488 #27502) -#27500 := (iff #27487 #27499) -#27497 := (iff #27486 #27496) -#27494 := (= #27485 #27493) -#27491 := (= #27484 #27490) -#27492 := [rewrite]: #27491 -#27495 := [monotonicity #27492]: #27494 -#27498 := [monotonicity #27495]: #27497 -#27501 := [monotonicity #27498]: #27500 -#27504 := [monotonicity #27501]: #27503 -#27509 := [trans #27504 #27507]: #27508 -#27512 := [monotonicity #27509]: #27511 -#29035 := [monotonicity #27512]: #29030 -#29038 := [trans #29035 #29037]: #29030 -#29029 := [quant-inst]: #29034 -#29039 := [mp #29029 #29038]: #29033 -#29088 := [unit-resolution #29039 #21660]: #27510 -#28968 := (or #27505 #27496) -#29050 := [def-axiom]: #28968 -#29049 := [unit-resolution #29050 #29088]: #27496 -#29056 := (= #3175 #27480) -#29054 := (= #27480 #3175) -#29055 := [monotonicity #28039 #27683]: #29054 -#29057 := [symm #29055]: #29056 -#29082 := [trans #29057 #29049]: #29048 -#29117 := [monotonicity #29082]: #29116 -#29115 := [symm #29117]: #29135 -#27644 := (= #24114 #27643) -#29031 := (or #24181 #27644) -#29032 := [quant-inst]: #29031 -#29087 := [unit-resolution #29032 #23694]: #27644 -#29136 := [trans #29087 #29115]: #27514 -#28947 := (not #27514) -#27309 := (not #27308) -#29080 := (iff #19011 #27309) -#29071 := (iff #12806 #27308) -#29081 := (iff #27308 #12806) -#29066 := [monotonicity #29078]: #29081 -#29072 := [symm #29066]: #29071 -#29083 := [monotonicity #29072]: #29080 -#29077 := [hypothesis]: #19011 -#29079 := [mp #29077 #29083]: #27309 -#28946 := (not #27513) -#29042 := (or #28946 #27308 #28947) -#29043 := [def-axiom]: #29042 -#29084 := [unit-resolution #29043 #29079 #29076]: #28947 -#29137 := [unit-resolution #29084 #29136]: false -#29138 := [lemma #29137]: #12806 -#29753 := [trans #29138 #29752]: #27308 -#28964 := (or #28946 #27309 #27514) -#28951 := [def-axiom]: #28964 -#28867 := [unit-resolution #28951 #29753 #29076]: #27514 -#29113 := [symm #28867]: #28862 -#28864 := [trans #29113 #27683]: #29114 -#29141 := [monotonicity #28864]: #29378 -#29124 := [symm #29141]: #29123 -#29188 := [trans #28004 #29124]: #27795 -#29563 := (not #27770) -#29681 := (iff #12299 #29563) -#29679 := (iff #12296 #27770) -#29461 := (iff #27770 #12296) -#29267 := (= #27769 #2955) -#29265 := (= #27768 #2952) -#29264 := (= #27768 #24234) -#29783 := (= #27762 #2962) -#29781 := (= #27762 #26432) -#27531 := (uf_66 #26432 #27529 #24114) -#27532 := (uf_58 #3079 #27531) -#27535 := (uf_136 #27532) -#29779 := (= #27535 #26432) -#27536 := (= #26432 #27535) -#27543 := (up_68 #27532) -#27544 := (not #27543) -#27540 := (uf_27 uf_273 #27531) -#27541 := (= uf_9 #27540) -#27542 := (not #27541) -#27538 := (uf_135 #27532) -#27539 := (= uf_9 #27538) -#27537 := (not #27536) -#27565 := (or #27537 #27539 #27542 #27544) -#27568 := (not #27565) -#28420 := (or #27549 #14243) -#28416 := [hypothesis]: #27550 -#28417 := [th-lemma #28416 #27679 #27672]: false -#28421 := [lemma #28417]: #28420 -#29742 := [unit-resolution #28421 #28992]: #27549 -#29745 := (or #27712 #27714) -#29743 := (or #27712 #27714 #14046) -#29744 := [unit-resolution #27716 #27707]: #29743 -#29746 := [unit-resolution #29744 #29231]: #29745 -#29747 := [unit-resolution #29746 #27711]: #27712 -#29144 := (or #27098 #26451 #27550 #27560 #27568) -#27545 := (or #27544 #27542 #27539 #27537) -#27546 := (not #27545) -#27551 := (or #27550 #27548 #26451 #27546) -#29145 := (or #27098 #27551) -#29157 := (iff #29145 #29144) -#27574 := (or #26451 #27550 #27560 #27568) -#29168 := (or #27098 #27574) -#29156 := (iff #29168 #29144) -#29154 := [rewrite]: #29156 -#29155 := (iff #29145 #29168) -#27577 := (iff #27551 #27574) -#27571 := (or #27550 #27560 #26451 #27568) -#27575 := (iff #27571 #27574) -#27576 := [rewrite]: #27575 -#27572 := (iff #27551 #27571) -#27569 := (iff #27546 #27568) -#27566 := (iff #27545 #27565) -#27567 := [rewrite]: #27566 -#27570 := [monotonicity #27567]: #27569 -#27573 := [monotonicity #27564 #27570]: #27572 -#27578 := [trans #27573 #27576]: #27577 -#29164 := [monotonicity #27578]: #29155 -#29158 := [trans #29164 #29154]: #29157 -#29167 := [quant-inst]: #29145 -#29159 := [mp #29167 #29158]: #29144 -#29748 := [unit-resolution #29159 #21444 #29747 #29742 #28031]: #27568 -#29175 := (or #27565 #27536) -#29176 := [def-axiom]: #29175 -#29749 := [unit-resolution #29176 #29748]: #27536 -#29780 := [symm #29749]: #29779 -#29777 := (= #27762 #27535) -#29775 := (= #27163 #27532) -#29773 := (= #27532 #27163) -#29771 := (= #27531 #3175) -#27310 := (uf_116 #3175) -#27388 := (uf_43 #24114 #27310) -#29765 := (= #27388 #3175) -#27429 := (= #3175 #27388) -#27431 := (or #27309 #27429) -#29044 := (or #25416 #27309 #27429) -#27430 := (or #27429 #27309) -#29045 := (or #25416 #27430) -#28949 := (iff #29045 #29044) -#28962 := (or #25416 #27431) -#28965 := (iff #28962 #29044) -#28948 := [rewrite]: #28965 -#28960 := (iff #29045 #28962) -#27432 := (iff #27430 #27431) -#27433 := [rewrite]: #27432 -#28963 := [monotonicity #27433]: #28960 -#28945 := [trans #28963 #28948]: #28949 -#28961 := [quant-inst]: #29045 -#28950 := [mp #28961 #28945]: #29044 -#29754 := [unit-resolution #28950 #18736]: #27431 -#29755 := [unit-resolution #29754 #29753]: #27429 -#29766 := [symm #29755]: #29765 -#29769 := (= #27531 #27388) -#28122 := (+ #26643 #27529) -#28146 := (+ #26356 #28122) -#28149 := (uf_43 #24114 #28146) -#29763 := (= #28149 #27388) -#29757 := (= #28146 #27310) -#29735 := (= #27310 #28146) -#29736 := (* -1::int #28146) -#29737 := (+ #27310 #29736) -#29738 := (<= #29737 0::int) -#27645 := (uf_116 #27493) -#27649 := (* -1::int #27645) -#29118 := (+ #27310 #27649) -#29119 := (<= #29118 0::int) -#29091 := (= #27310 #27645) -#29592 := (= #27645 #27310) -#29585 := (= #27493 #3175) -#29612 := (= #27493 #27480) -#29613 := [symm #29049]: #29612 -#29591 := [trans #29613 #29055]: #29585 -#29588 := [monotonicity #29591]: #29592 -#29593 := [symm #29588]: #29091 -#29604 := (not #29091) -#29605 := (or #29604 #29119) -#29606 := [th-lemma]: #29605 -#29607 := [unit-resolution #29606 #29593]: #29119 -#27650 := (+ #26357 #27649) -#27651 := (+ #26356 #27650) -#27652 := (+ uf_294 #27651) -#29185 := (>= #27652 0::int) -#27653 := (= #27652 0::int) -#29089 := (or #24187 #27653) -#27646 := (= #27490 #27645) -#29085 := (or #24187 #27646) -#29108 := (iff #29085 #29089) -#29110 := (iff #29089 #29089) -#29111 := [rewrite]: #29110 -#27654 := (iff #27646 #27653) -#27655 := [rewrite]: #27654 -#29109 := [monotonicity #27655]: #29108 -#29112 := [trans #29109 #29111]: #29108 -#29086 := [quant-inst]: #29085 -#29107 := [mp #29086 #29112]: #29089 -#29608 := [unit-resolution #29107 #23688]: #27653 -#29595 := (not #27653) -#29596 := (or #29595 #29185) -#29597 := [th-lemma]: #29596 -#29598 := [unit-resolution #29597 #29608]: #29185 -#29601 := (not #27668) -#29600 := (not #27892) -#29599 := (not #27068) -#29594 := (not #29185) -#29586 := (not #29119) -#29602 := (or #29738 #29586 #29594 #29599 #29600 #29601) -#29603 := [th-lemma]: #29602 -#29625 := [unit-resolution #29603 #28258 #27679 #29598 #29607 #28054]: #29738 -#29739 := (>= #29737 0::int) -#29120 := (>= #29118 0::int) -#29626 := (or #29604 #29120) -#29616 := [th-lemma]: #29626 -#29614 := [unit-resolution #29616 #29593]: #29120 -#29196 := (<= #27652 0::int) -#29617 := (or #29595 #29196) -#29618 := [th-lemma]: #29617 -#29619 := [unit-resolution #29618 #29608]: #29196 -#29627 := (not #27893) -#29624 := (not #27073) -#29621 := (not #29196) -#29620 := (not #29120) -#29623 := (or #29739 #29620 #29621 #29624 #29627 #27714) -#29628 := [th-lemma]: #29623 -#29629 := [unit-resolution #29628 #28051 #28271 #29619 #29614 #27711]: #29739 -#29503 := (not #29739) -#29630 := (not #29738) -#29518 := (or #29735 #29630 #29503) -#29532 := [th-lemma]: #29518 -#29517 := [unit-resolution #29532 #29629 #29625]: #29735 -#29263 := [symm #29517]: #29757 -#29472 := [monotonicity #29263]: #29763 -#29767 := (= #27531 #28149) -#28104 := (uf_66 #25404 #27529 #24114) -#28136 := (= #28104 #28149) -#28137 := (not #28136) -#28107 := (uf_139 #28104 #25404) -#28108 := (= uf_9 #28107) -#28109 := (not #28108) -#28145 := (or #28109 #28137) -#28249 := (not #28145) -#29261 := (or #26114 #28249) -#28110 := (+ #27529 #26644) -#28111 := (uf_43 #24114 #28110) -#28112 := (= #28104 #28111) -#28117 := (not #28112) -#28118 := (or #28117 #28109) -#28121 := (not #28118) -#29262 := (or #26114 #28121) -#29295 := (iff #29262 #29261) -#29335 := (iff #29261 #29261) -#29336 := [rewrite]: #29335 -#28250 := (iff #28121 #28249) -#28247 := (iff #28118 #28145) -#28142 := (or #28137 #28109) -#28156 := (iff #28142 #28145) -#28157 := [rewrite]: #28156 -#28143 := (iff #28118 #28142) -#28140 := (iff #28117 #28137) -#28138 := (iff #28112 #28136) -#28150 := (= #28111 #28149) -#28147 := (= #28110 #28146) -#28148 := [rewrite]: #28147 -#28151 := [monotonicity #28148]: #28150 -#28139 := [monotonicity #28151]: #28138 -#28141 := [monotonicity #28139]: #28140 -#28144 := [monotonicity #28141]: #28143 -#28248 := [trans #28144 #28157]: #28247 -#28251 := [monotonicity #28248]: #28250 -#29334 := [monotonicity #28251]: #29295 -#29337 := [trans #29334 #29336]: #29295 -#29294 := [quant-inst]: #29262 -#29338 := [mp #29294 #29337]: #29261 -#29759 := [unit-resolution #29338 #21660]: #28249 -#29340 := (or #28145 #28136) -#29279 := [def-axiom]: #29340 -#29760 := [unit-resolution #29279 #29759]: #28136 -#29761 := (= #27531 #28104) -#29762 := [monotonicity #27948]: #29761 -#29768 := [trans #29762 #29760]: #29767 -#29473 := [trans #29768 #29472]: #29769 -#29499 := [trans #29473 #29766]: #29771 -#29656 := [monotonicity #29499]: #29773 -#29657 := [symm #29656]: #29775 -#29180 := [monotonicity #29657]: #29777 -#29677 := [trans #29180 #29780]: #29781 -#29199 := [trans #29677 #27935]: #29783 -#29181 := [monotonicity #29199]: #29264 -#29266 := [trans #29181 #28359]: #29265 -#29460 := [monotonicity #29266]: #29267 -#29462 := [monotonicity #29460]: #29461 -#29680 := [symm #29462]: #29679 -#29682 := [monotonicity #29680]: #29681 -#29683 := [mp #14796 #29682]: #29563 -#29170 := (not #27607) -#29696 := (iff #29170 #27761) -#29689 := (iff #27607 #27760) -#29693 := (iff #27760 #27607) -#29691 := (= #27759 #27606) -#29688 := (= #27163 #27605) -#29686 := (= #27605 #27163) -#29687 := [monotonicity #27692]: #29686 -#29690 := [symm #29687]: #29688 -#29692 := [monotonicity #29690]: #29691 -#29694 := [monotonicity #29692]: #29693 -#29695 := [symm #29694]: #29689 -#29697 := [monotonicity #29695]: #29696 -#29684 := [unit-resolution #27635 #22104 #14784 #29000 #29747 #29742 #27723]: #27614 -#29173 := (or #27611 #29170) -#29169 := [def-axiom]: #29173 -#29685 := [unit-resolution #29169 #29684]: #29170 -#29698 := [mp #29685 #29697]: #27761 -#29577 := (or #27819 #27760) -#29578 := [def-axiom]: #29577 -#29699 := [unit-resolution #29578 #29698]: #27819 -#29709 := (or #27833 #27770 #27810 #27822) -#29792 := (not #29735) -#29793 := (or #29792 #27772) -#29788 := (= #2967 #27771) -#29785 := (= #27771 #2967) -#29756 := [hypothesis]: #29735 -#29758 := [symm #29756]: #29757 -#29764 := [monotonicity #29758]: #29763 -#29770 := [trans #29768 #29764]: #29769 -#29772 := [trans #29770 #29766]: #29771 -#29774 := [monotonicity #29772]: #29773 -#29776 := [symm #29774]: #29775 -#29778 := [monotonicity #29776]: #29777 -#29782 := [trans #29778 #29780]: #29781 -#29784 := [trans #29782 #27935]: #29783 -#29786 := [monotonicity #29784]: #29785 -#29789 := [symm #29786]: #29788 -#29790 := [trans #14799 #29789]: #27772 -#29458 := (not #27772) -#29740 := [hypothesis]: #29458 -#29791 := [unit-resolution #29740 #29790]: false -#29794 := [lemma #29791]: #29793 -#29702 := [unit-resolution #29794 #29517]: #27772 -#29528 := (or #27813 #29458) -#29529 := [def-axiom]: #29528 -#29703 := [unit-resolution #29529 #29702]: #27813 -#29571 := (or #27833 #27770 #27810 #27816 #27822) -#29572 := [def-axiom]: #29571 -#29710 := [unit-resolution #29572 #29703]: #29709 -#29711 := [unit-resolution #29710 #29699 #29683 #29188 #29637]: false -#29712 := [lemma #29711]: #12812 -#23247 := (not #19374) -#29440 := [hypothesis]: #23803 -#23427 := (or #23812 #23800) -#23522 := [def-axiom]: #23427 -#29504 := [unit-resolution #23522 #29440]: #23812 -#23396 := (or #23806 #23800) -#23538 := [def-axiom]: #23396 -#29505 := [unit-resolution #23538 #29440]: #23806 -#29542 := (or #23818 #23809) -#23403 := (or #23906 #14046) -#23404 := [def-axiom]: #23403 -#29536 := [unit-resolution #23404 #29231]: #23906 -#29537 := [unit-resolution #23380 #29536 #29225]: #23875 -#23447 := (or #23872 #23866) -#23448 := [def-axiom]: #23447 -#29538 := [unit-resolution #23448 #29537]: #23866 -#27307 := (or #23818 #23809 #19008 #23869) -#27389 := [hypothesis]: #23821 -#23428 := (or #23818 #12812) -#23429 := [def-axiom]: #23428 -#27390 := [unit-resolution #23429 #27389]: #12812 -#23411 := (or #23818 #12806) -#23426 := [def-axiom]: #23411 -#27391 := [unit-resolution #23426 #27389]: #12806 -#27392 := [hypothesis]: #12803 -#27387 := [hypothesis]: #23866 -#23466 := (or #23869 #19008 #19011 #23863) -#23461 := [def-axiom]: #23466 -#27393 := [unit-resolution #23461 #27391 #27387 #27392]: #23863 -#23475 := (or #23860 #23854) -#23470 := [def-axiom]: #23475 -#27394 := [unit-resolution #23470 #27393]: #23854 -#23468 := (or #23857 #19011 #19017 #23851) -#23469 := [def-axiom]: #23468 -#27395 := [unit-resolution #23469 #27394 #27391 #27390]: #23851 -#27396 := [hypothesis]: #23806 -#23528 := (or #23824 #23818) -#23515 := [def-axiom]: #23528 -#27397 := [unit-resolution #23515 #27389]: #23824 -#23521 := (or #23833 #19008 #19011 #23827) -#23510 := [def-axiom]: #23521 -#27304 := [unit-resolution #23510 #27397 #27392 #27391]: #23833 -#23499 := (or #23836 #23830) -#23501 := [def-axiom]: #23499 -#27305 := [unit-resolution #23501 #27304]: #23836 -#23492 := (or #23845 #23809 #23839) -#23494 := [def-axiom]: #23492 -#27306 := [unit-resolution #23494 #27305 #27396]: #23845 -#23482 := (or #23848 #23842) -#23483 := [def-axiom]: #23482 -#27269 := [unit-resolution #23483 #27306 #27395]: false -#27303 := [lemma #27269]: #27307 -#29521 := [unit-resolution #27303 #29001 #29538]: #29542 -#29525 := [unit-resolution #29521 #29505]: #23818 -#29520 := (or #23821 #19017 #23815) -#23431 := (or #23821 #19011 #19017 #23815) -#23432 := [def-axiom]: #23431 -#29526 := [unit-resolution #23432 #29138]: #29520 -#29527 := [unit-resolution #29526 #29525 #29504 #29712]: false -#29530 := [lemma #29527]: #23800 -#29896 := (or #23803 #23797) -#16270 := (<= uf_272 131073::int) -#16273 := (iff #13872 #16270) -#16264 := (+ 131073::int #13873) -#16267 := (>= #16264 0::int) -#16271 := (iff #16267 #16270) -#16272 := [rewrite]: #16271 -#16268 := (iff #13872 #16267) -#16265 := (= #13874 #16264) -#16266 := [monotonicity #7888]: #16265 -#16269 := [monotonicity #16266]: #16268 -#16274 := [trans #16269 #16272]: #16273 -#14787 := [not-or-elim #14776]: #13880 -#14788 := [and-elim #14787]: #13872 -#16275 := [mp #14788 #16274]: #16270 -#29232 := [hypothesis]: #19037 -#29233 := [th-lemma #29232 #29231 #16275]: false -#29234 := [lemma #29233]: #16368 -#29894 := (or #23803 #19037 #23797) -#29891 := (or #14243 #14088) -#29892 := [th-lemma]: #29891 -#29893 := [unit-resolution #29892 #28992]: #14088 -#23558 := (or #23803 #19034 #19037 #23797) -#23555 := [def-axiom]: #23558 -#29895 := [unit-resolution #23555 #29893]: #29894 -#29897 := [unit-resolution #29895 #29234]: #29896 -#29898 := [unit-resolution #29897 #29530]: #23797 -#23561 := (or #23794 #23788) -#23565 := [def-axiom]: #23561 -#29899 := [unit-resolution #23565 #29898]: #23788 -#23271 := (>= #14169 -1::int) -#23285 := (or #23794 #14168) -#23286 := [def-axiom]: #23285 -#29900 := [unit-resolution #23286 #29898]: #14168 -#29901 := (or #14172 #23271) -#29902 := [th-lemma]: #29901 -#29903 := [unit-resolution #29902 #29900]: #23271 -#29238 := (not #23271) -#29239 := (or #14100 #29238) -#29201 := [hypothesis]: #23271 -#29202 := [hypothesis]: #14105 -#29237 := [th-lemma #29202 #29231 #29201]: false -#29240 := [lemma #29237]: #29239 -#29904 := [unit-resolution #29240 #29903]: #14100 -#23580 := (or #23791 #14105 #23785) -#23566 := [def-axiom]: #23580 -#29905 := [unit-resolution #23566 #29904 #29899]: #23785 -#23575 := (or #23782 #23776) -#23213 := [def-axiom]: #23575 -#29906 := [unit-resolution #23213 #29905]: #23776 -#29920 := (= #3068 #3209) -#29918 := (= #3209 #3068) -#29914 := (= #3208 #3067) -#29912 := (= #3208 #27327) -#29910 := (= uf_301 #26964) -#29907 := [hypothesis]: #23809 -#23549 := (or #23806 #12826) -#23550 := [def-axiom]: #23549 -#29908 := [unit-resolution #23550 #29907]: #12826 -#29909 := [symm #29908]: #3189 -#29911 := [trans #29909 #27438]: #29910 -#29913 := [monotonicity #29911]: #29912 -#29915 := [trans #29913 #27639]: #29914 -#29919 := [monotonicity #29915]: #29918 -#29921 := [symm #29919]: #29920 -#29922 := (= uf_300 #3068) -#23559 := (or #23806 #12823) -#23548 := [def-axiom]: #23559 -#29916 := [unit-resolution #23548 #29907]: #12823 -#29917 := [symm #29916]: #3187 -#29923 := [trans #29917 #29222]: #29922 -#29924 := [trans #29923 #29921]: #12862 -#28863 := (+ uf_293 #14142) -#28865 := (>= #28863 0::int) -#29925 := (or #12993 #28865) -#29926 := [th-lemma]: #29925 -#29927 := [unit-resolution #29926 #29908]: #28865 -#29483 := (not #28865) -#29930 := (or #14145 #29483) -#29928 := (or #14145 #14404 #29483) -#29929 := [th-lemma]: #29928 -#29931 := [unit-resolution #29929 #29223]: #29930 -#29932 := [unit-resolution #29931 #29927]: #14145 -#23374 := (or #22784 #22782 #14144) -#23581 := [def-axiom]: #23374 -#29933 := [unit-resolution #23581 #29932 #29924]: #22784 -#23255 := (or #23770 #22783) -#23256 := [def-axiom]: #23255 -#29934 := [unit-resolution #23256 #29933]: #23770 -#23572 := (or #23779 #23773 #22836) -#23573 := [def-axiom]: #23572 -#29935 := [unit-resolution #23573 #29934 #29906]: #22836 -#23583 := (or #22831 #23247) -#23243 := [def-axiom]: #23583 -#29936 := [unit-resolution #23243 #29935]: #23247 -#29307 := (+ uf_294 #19372) -#29860 := (>= #29307 0::int) -#29955 := (not #29860) -#29855 := (= uf_294 ?x785!14) -#29888 := (not #29855) -#29858 := (= #3184 #19060) -#29864 := (not #29858) -#29859 := (+ #3184 #19385) -#29861 := (>= #29859 0::int) -#29871 := (not #29861) -#23394 := (or #23806 #14207) -#23395 := [def-axiom]: #23394 -#29937 := [unit-resolution #23395 #29907]: #14207 -#29541 := (+ uf_292 #14120) -#29539 := (<= #29541 0::int) -#29938 := (or #13002 #29539) -#29939 := [th-lemma]: #29938 -#29940 := [unit-resolution #29939 #29916]: #29539 -#23227 := (or #22831 #23584) -#23568 := [def-axiom]: #23227 -#29941 := [unit-resolution #23568 #29935]: #23584 -#29872 := (not #29539) -#29873 := (or #29871 #19387 #29872 #14211) -#29866 := [hypothesis]: #14207 -#29867 := [hypothesis]: #29539 -#29868 := [hypothesis]: #23584 -#29869 := [hypothesis]: #29861 -#29870 := [th-lemma #29869 #29868 #29867 #29866]: false -#29874 := [lemma #29870]: #29873 -#29942 := [unit-resolution #29874 #29941 #29940 #29937]: #29871 -#29865 := (or #29864 #29861) -#29875 := [th-lemma]: #29865 -#29943 := [unit-resolution #29875 #29942]: #29864 -#29889 := (or #29888 #29858) -#29884 := (= #19060 #3184) -#29882 := (= #19059 #3175) -#29880 := (= ?x785!14 uf_294) -#29879 := [hypothesis]: #29855 -#29881 := [symm #29879]: #29880 -#29883 := [monotonicity #29881]: #29882 -#29885 := [monotonicity #29883]: #29884 -#29886 := [symm #29885]: #29858 -#29878 := [hypothesis]: #29864 -#29887 := [unit-resolution #29878 #29886]: false -#29890 := [lemma #29887]: #29889 -#29944 := [unit-resolution #29890 #29943]: #29888 -#29958 := (or #29855 #29955) -#29308 := (<= #29307 0::int) -#29319 := (+ uf_292 #19385) -#29320 := (>= #29319 0::int) -#29945 := (not #29320) -#29946 := (or #29945 #19387 #29872) -#29947 := [th-lemma]: #29946 -#29948 := [unit-resolution #29947 #29940 #29941]: #29945 -#29951 := (or #29308 #29320) -#23582 := (or #22831 #19056) -#23242 := [def-axiom]: #23582 -#29949 := [unit-resolution #23242 #29935]: #19056 -#23586 := (or #22831 #19055) -#23592 := [def-axiom]: #23586 -#29950 := [unit-resolution #23592 #29935]: #19055 -#29802 := (or #23759 #22815 #22816 #29308 #29320) -#29296 := (+ #19060 #14431) -#29297 := (<= #29296 0::int) -#29298 := (+ ?x785!14 #14044) -#29299 := (>= #29298 0::int) -#29300 := (or #22816 #29299 #29297 #22815) -#29803 := (or #23759 #29300) -#29810 := (iff #29803 #29802) -#29328 := (or #22815 #22816 #29308 #29320) -#29805 := (or #23759 #29328) -#29808 := (iff #29805 #29802) -#29809 := [rewrite]: #29808 -#29806 := (iff #29803 #29805) -#29331 := (iff #29300 #29328) -#29325 := (or #22816 #29308 #29320 #22815) -#29329 := (iff #29325 #29328) -#29330 := [rewrite]: #29329 -#29326 := (iff #29300 #29325) -#29323 := (iff #29297 #29320) -#29313 := (+ #14431 #19060) -#29316 := (<= #29313 0::int) -#29321 := (iff #29316 #29320) -#29322 := [rewrite]: #29321 -#29317 := (iff #29297 #29316) -#29314 := (= #29296 #29313) -#29315 := [rewrite]: #29314 -#29318 := [monotonicity #29315]: #29317 -#29324 := [trans #29318 #29322]: #29323 -#29311 := (iff #29299 #29308) -#29301 := (+ #14044 ?x785!14) -#29304 := (>= #29301 0::int) -#29309 := (iff #29304 #29308) -#29310 := [rewrite]: #29309 -#29305 := (iff #29299 #29304) -#29302 := (= #29298 #29301) -#29303 := [rewrite]: #29302 -#29306 := [monotonicity #29303]: #29305 -#29312 := [trans #29306 #29310]: #29311 -#29327 := [monotonicity #29312 #29324]: #29326 -#29332 := [trans #29327 #29330]: #29331 -#29807 := [monotonicity #29332]: #29806 -#29811 := [trans #29807 #29809]: #29810 -#29804 := [quant-inst]: #29803 -#29812 := [mp #29804 #29811]: #29802 -#29952 := [unit-resolution #29812 #29221 #29950 #29949]: #29951 -#29953 := [unit-resolution #29952 #29948]: #29308 -#29954 := (not #29308) -#29956 := (or #29855 #29954 #29955) -#29957 := [th-lemma]: #29956 -#29959 := [unit-resolution #29957 #29953]: #29958 -#29960 := [unit-resolution #29959 #29944]: #29955 -#29961 := [th-lemma #29903 #29960 #29936]: false -#29962 := [lemma #29961]: #23806 -#29633 := [unit-resolution #29521 #29962]: #23818 -#29615 := [unit-resolution #29526 #29633 #29712]: #23815 -#23534 := (or #23812 #13075) -#23416 := [def-axiom]: #23534 -#29713 := [unit-resolution #23416 #29615]: #13075 -#29142 := (or #13081 #23536) -#29708 := [th-lemma]: #29142 -#29714 := [unit-resolution #29708 #29713]: #23536 -#29715 := [hypothesis]: #14144 -#29716 := [th-lemma #29715 #29714 #29231]: false -#29717 := [lemma #29716]: #14145 -#29990 := (or #22784 #14144) -#29985 := (= #3184 #3209) -#29982 := (= #3209 #3184) -#29979 := (= #3208 #3175) -#29978 := [symm #29713]: #3247 -#29980 := [monotonicity #29978]: #29979 -#29983 := [monotonicity #29980]: #29982 -#29986 := [symm #29983]: #29985 -#29987 := (= uf_300 #3184) -#23533 := (or #23812 #13070) -#23531 := [def-axiom]: #23533 -#29977 := [unit-resolution #23531 #29615]: #13070 -#29984 := [symm #29977]: #3240 -#23373 := (or #23812 #3246) -#23375 := [def-axiom]: #23373 -#29981 := [unit-resolution #23375 #29615]: #3246 -#29988 := [trans #29981 #29984]: #29987 -#29989 := [trans #29988 #29986]: #12862 -#29991 := [unit-resolution #23581 #29989]: #29990 -#29992 := [unit-resolution #29991 #29717]: #22784 -#29993 := [unit-resolution #23256 #29992]: #23770 -#29994 := [unit-resolution #23573 #29906]: #23776 -#29995 := [unit-resolution #29994 #29993]: #22836 -#30004 := [unit-resolution #23568 #29995]: #23584 -#30026 := (or #29945 #19387) -#29570 := (+ #3184 #14120) -#29587 := (<= #29570 0::int) -#29569 := (= #3184 uf_300) -#30005 := (= uf_304 uf_300) -#30006 := [symm #29981]: #30005 -#30007 := [trans #29977 #30006]: #29569 -#30008 := (not #29569) -#30009 := (or #30008 #29587) -#30010 := [th-lemma]: #30009 -#30011 := [unit-resolution #30010 #30007]: #29587 -#30016 := (or #19017 #23851) -#30012 := (or #19011 #23863) -#30013 := [unit-resolution #23461 #29001 #29538]: #30012 -#30014 := [unit-resolution #30013 #29138]: #23863 -#30015 := [unit-resolution #23470 #30014]: #23854 -#30017 := [unit-resolution #23469 #29138 #30015]: #30016 -#30018 := [unit-resolution #30017 #29712]: #23851 -#30019 := [unit-resolution #23483 #30018]: #23842 -#30020 := [unit-resolution #23494 #29962 #30019]: #23839 -#23514 := (or #23836 #14211) -#23498 := [def-axiom]: #23514 -#30021 := [unit-resolution #23498 #30020]: #14211 -#30022 := (not #29587) -#30023 := (or #29539 #14207 #30022) -#30024 := [th-lemma]: #30023 -#30025 := [unit-resolution #30024 #30021 #30011]: #29539 -#30027 := [unit-resolution #29947 #30025]: #30026 -#30028 := [unit-resolution #30027 #30004]: #29945 -#30029 := [unit-resolution #23242 #29995]: #19056 -#30030 := [unit-resolution #23592 #29995]: #19055 -#30031 := [unit-resolution #29812 #29221 #30030 #30029 #30028]: #29308 -#29996 := [unit-resolution #23243 #29995]: #23247 -#29997 := [hypothesis]: #29955 -#29998 := [th-lemma #29903 #29997 #29996]: false -#29999 := [lemma #29998]: #29860 -#30032 := [unit-resolution #29957 #29999 #30031]: #29855 -#30033 := [unit-resolution #29890 #30032]: #29858 -#30034 := [unit-resolution #29875 #30033]: #29861 -[th-lemma #30011 #30034 #30004]: false +#7649 := (iff #7612 #7648) +#7646 := (iff #1372 #7645) +#7643 := (iff #1371 #7642) +#7644 := [rewrite]: #7643 +#7647 := [monotonicity #4064 #7644]: #7646 +#7650 := [monotonicity #7647]: #7649 +#7992 := [monotonicity #7650 #7989]: #7991 +#7995 := [monotonicity #7992]: #7994 +#8000 := [trans #7995 #7998]: #7999 +#8003 := [quant-intro #8000]: #8002 +#7985 := (iff #1493 #7984) +#7982 := (iff #1492 #7979) +#7975 := (implies #7943 #7970) +#7980 := (iff #7975 #7979) +#7981 := [rewrite]: #7980 +#7976 := (iff #1492 #7975) +#7973 := (iff #1491 #7970) +#7967 := (implies #1372 #7964) +#7971 := (iff #7967 #7970) +#7972 := [rewrite]: #7971 +#7968 := (iff #1491 #7967) +#7965 := (iff #1490 #7964) +#7962 := (iff #1489 #7961) +#7959 := (iff #1488 #7958) +#7956 := (iff #1487 #7955) +#7957 := [rewrite]: #7956 +#7960 := [monotonicity #7957]: #7959 +#7953 := (iff #1484 #7952) +#7950 := (iff #1483 #7949) +#7951 := [rewrite]: #7950 +#7954 := [monotonicity #7951]: #7953 +#7963 := [monotonicity #7954 #7960]: #7962 +#7947 := (iff #1481 #7946) +#7948 := [rewrite]: #7947 +#7966 := [monotonicity #7948 #7963]: #7965 +#7969 := [monotonicity #7966]: #7968 +#7974 := [trans #7969 #7972]: #7973 +#7944 := (iff #1477 #7943) +#7945 := [rewrite]: #7944 +#7977 := [monotonicity #7945 #7974]: #7976 +#7983 := [trans #7977 #7981]: #7982 +#7986 := [quant-intro #7983]: #7985 +#8005 := [trans #7986 #8003]: #8004 +#7942 := [asserted]: #1493 +#8006 := [mp #7942 #8005]: #8001 +#17013 := [mp~ #8006 #17012]: #8001 +#20893 := [mp #17013 #20892]: #20890 +#26056 := (not #20890) +#26054 := (or #26056 #25897 #25915 #25923) +#25893 := (or #25892 #25890 #25887 #25884) +#25894 := (not #25893) +#25903 := (or #25902 #25900 #25897 #25894) +#26057 := (or #26056 #25903) +#26096 := (iff #26057 #26054) +#25929 := (or #25897 #25915 #25923) +#26106 := (or #26056 #25929) +#26109 := (iff #26106 #26054) +#26098 := [rewrite]: #26109 +#26107 := (iff #26057 #26106) +#25932 := (iff #25903 #25929) +#25926 := (or false #25915 #25897 #25923) +#25930 := (iff #25926 #25929) +#25931 := [rewrite]: #25930 +#25927 := (iff #25903 #25926) +#25924 := (iff #25894 #25923) +#25921 := (iff #25893 #25920) +#25922 := [rewrite]: #25921 +#25925 := [monotonicity #25922]: #25924 +#25928 := [monotonicity #25909 #25919 #25925]: #25927 +#25933 := [trans #25928 #25931]: #25932 +#26108 := [monotonicity #25933]: #26107 +#26099 := [trans #26108 #26098]: #26096 +#26105 := [quant-inst]: #26057 +#26149 := [mp #26105 #26099]: #26054 +#27437 := [unit-resolution #26149 #20893 #26260]: #27447 +#27472 := [unit-resolution #27437 #27446]: #25923 +#26138 := (or #25920 #25889) +#26139 := [def-axiom]: #26138 +#27473 := [unit-resolution #26139 #27472]: #25889 +#27495 := [symm #27473]: #27471 +#27491 := (= #26721 #25888) +#27538 := (= #26720 #25885) +#25710 := (uf_58 #3149 #3016) +#27490 := (= #25710 #25885) +#27532 := (= #25885 #25710) +#27497 := (= #25881 #3016) +#27536 := (= #25881 #25799) +#27534 := (= #25881 #25821) +#27475 := (= #25881 #2960) +#26062 := (uf_116 #24856) +#26076 := (uf_43 #23566 #26062) +#27504 := (= #26076 #2960) +#27428 := (= #26062 uf_274) +#27413 := (= #26062 #2961) +#27409 := (= #23468 #2961) +#23692 := (= #2961 #23468) +#23697 := (or #23639 #23692) +#23698 := [quant-inst]: #23697 +#27474 := [unit-resolution #23698 #23137]: #23692 +#27412 := [symm #27474]: #27409 +#27410 := (= #26062 #23468) +#27411 := [monotonicity #25871]: #27410 +#27401 := [trans #27411 #27412]: #27413 +#27414 := [trans #27401 #25702]: #27428 +#27505 := [monotonicity #25690 #27414]: #27504 +#27508 := (= #25881 #26076) +#26058 := (uf_66 #24856 0::int #23566) +#26079 := (= #26058 #26076) +#26082 := (not #26079) +#26059 := (uf_138 #26058 #24856) +#26060 := (= uf_9 #26059) +#26061 := (not #26060) +#26088 := (or #26061 #26082) +#26093 := (not #26088) +#26525 := (or #25566 #26093) +#26063 := (+ #26062 #25804) +#26064 := (uf_43 #23566 #26063) +#26065 := (= #26058 #26064) +#26066 := (not #26065) +#26067 := (or #26066 #26061) +#26068 := (not #26067) +#26526 := (or #25566 #26068) +#26840 := (iff #26526 #26525) +#26836 := (iff #26525 #26525) +#26837 := [rewrite]: #26836 +#26094 := (iff #26068 #26093) +#26091 := (iff #26067 #26088) +#26085 := (or #26082 #26061) +#26089 := (iff #26085 #26088) +#26090 := [rewrite]: #26089 +#26086 := (iff #26067 #26085) +#26083 := (iff #26066 #26082) +#26080 := (iff #26065 #26079) +#26077 := (= #26064 #26076) +#26074 := (= #26063 #26062) +#26069 := (+ #26062 0::int) +#26072 := (= #26069 #26062) +#26073 := [rewrite]: #26072 +#26070 := (= #26063 #26069) +#26071 := [monotonicity #25813]: #26070 +#26075 := [trans #26071 #26073]: #26074 +#26078 := [monotonicity #26075]: #26077 +#26081 := [monotonicity #26078]: #26080 +#26084 := [monotonicity #26081]: #26083 +#26087 := [monotonicity #26084]: #26086 +#26092 := [trans #26087 #26090]: #26091 +#26095 := [monotonicity #26092]: #26094 +#26841 := [monotonicity #26095]: #26840 +#26842 := [trans #26841 #26837]: #26840 +#26839 := [quant-inst]: #26526 +#26843 := [mp #26839 #26842]: #26525 +#27415 := [unit-resolution #26843 #21109]: #26093 +#26714 := (or #26088 #26079) +#26716 := [def-axiom]: #26714 +#27416 := [unit-resolution #26716 #27415]: #26079 +#27417 := (= #25881 #26058) +#27476 := [monotonicity #25843]: #27417 +#27509 := [trans #27476 #27416]: #27508 +#27533 := [trans #27509 #27505]: #27475 +#27535 := [trans #27533 #27338]: #27534 +#27537 := [trans #27535 #27344]: #27536 +#27498 := [trans #27537 #27306]: #27497 +#27543 := [monotonicity #27498]: #27532 +#27492 := [symm #27543]: #27490 +#27499 := (= #26720 #25710) +#27531 := [monotonicity #27365]: #27499 +#27493 := [trans #27531 #27492]: #27538 +#27494 := [monotonicity #27493]: #27491 +#27686 := [trans #27494 #27495]: #27496 +#27639 := [trans #27686 #25865]: #27687 +#27638 := [monotonicity #27639]: #27679 +#27694 := [symm #27638]: #27681 +#27630 := [trans #14262 #27694]: #26898 +#27134 := (not #26898) +#27132 := (or #26913 #27134) +#27135 := [def-axiom]: #27132 +#27631 := [unit-resolution #27135 #27630]: #26913 +#24105 := (uf_12 uf_7) +#27737 := (= #24105 #26921) +#27761 := (= #26921 #24105) +#27779 := (= #25857 uf_7) +#27771 := (= #23681 uf_7) +#27778 := [symm #25680]: #27771 +#27769 := (= #25857 #23681) +#27770 := [monotonicity #25699]: #27769 +#27780 := [trans #27770 #27778]: #27779 +#27767 := [monotonicity #27780]: #27761 +#27722 := [symm #27767]: #27737 +#24106 := (= uf_14 #24105) +#24113 := (iff #11408 #24106) +#2308 := (pattern #237) +#2836 := (uf_12 #233) +#11581 := (= uf_14 #2836) +#11585 := (iff #3950 #11581) +#11588 := (forall (vars (?x761 T3)) (:pat #2308) #11585) +#18271 := (~ #11588 #11588) +#18269 := (~ #11585 #11585) +#18270 := [refl]: #18269 +#18272 := [nnf-pos #18270]: #18271 +#2849 := (= #2836 uf_14) +#2850 := (iff #238 #2849) +#2851 := (forall (vars (?x761 T3)) (:pat #2308) #2850) +#11589 := (iff #2851 #11588) +#11586 := (iff #2850 #11585) +#11583 := (iff #2849 #11581) +#11584 := [rewrite]: #11583 +#11587 := [monotonicity #3952 #11584]: #11586 +#11590 := [quant-intro #11587]: #11589 +#11580 := [asserted]: #2851 +#11593 := [mp #11580 #11590]: #11588 +#18273 := [mp~ #11593 #18272]: #11588 +#23737 := (not #11588) +#24116 := (or #23737 #24113) +#24117 := [quant-inst]: #24116 +#27629 := [unit-resolution #24117 #18273]: #24113 +#24118 := (not #24113) +#27723 := (or #24118 #24106) +#24122 := (not #11408) +#24123 := (or #24118 #24122 #24106) +#24124 := [def-axiom]: #24123 +#27724 := [unit-resolution #24124 #11414]: #27723 +#27766 := [unit-resolution #27724 #27629]: #24106 +#27709 := [trans #27766 #27722]: #26953 +#27149 := (not #26915) +#27478 := (iff #11881 #27149) +#27477 := (iff #11878 #26915) +#27290 := (iff #26915 #11878) +#27729 := (= #26902 #2971) +#27727 := (= #26914 #2958) +#27725 := (= #26914 #23686) +#27726 := [monotonicity #27639]: #27725 +#27728 := [trans #27726 #25786]: #27727 +#27765 := [monotonicity #27728]: #27729 +#27418 := [monotonicity #27765]: #27290 +#27402 := [symm #27418]: #27477 +#27479 := [monotonicity #27402]: #27478 +#27480 := [mp #14265 #27479]: #27149 +#26305 := (not #25961) +#27693 := (iff #26305 #26918) +#27698 := (iff #25961 #26917) +#27804 := (iff #26917 #25961) +#27802 := (= #26901 #25960) +#27354 := (= #26720 #25959) +#27352 := (= #25710 #25959) +#27327 := (= #25959 #25710) +#27328 := [monotonicity #26204]: #27327 +#27353 := [symm #27328]: #27352 +#27355 := [trans #27531 #27353]: #27354 +#27803 := [monotonicity #27355]: #27802 +#27787 := [monotonicity #27803]: #27804 +#27786 := [symm #27787]: #27698 +#27699 := [monotonicity #27786]: #27693 +#27696 := (or #25972 #25977) +#27697 := [unit-resolution #25999 #21553 #14251 #26260 #27341]: #27696 +#27781 := [unit-resolution #27697 #26234]: #25977 +#26350 := (or #25974 #26305) +#26351 := [def-axiom]: #26350 +#27782 := [unit-resolution #26351 #27781]: #26305 +#27730 := [mp #27782 #27699]: #26918 +#27144 := (or #27016 #26917) +#27142 := [def-axiom]: #27144 +#27762 := [unit-resolution #27142 #27730]: #27016 +#27292 := (or #27161 #26900 #26915 #27014 #27020) +#27293 := [def-axiom]: #27292 +#27763 := [unit-resolution #27293 #27762 #27480 #27709 #27631]: #27161 +#27325 := (or #27164 #27156) +#27326 := [def-axiom]: #27325 +#27764 := [unit-resolution #27326 #27763 #27368]: false +#27818 := [lemma #27764]: #11940 +#22709 := (or #23412 #18314 #18323 #23406) +#22710 := [def-axiom]: #22709 +#27887 := [unit-resolution #22710 #27818 #27884 #27886]: #23406 +#22742 := (or #23403 #3027) +#22743 := [def-axiom]: #22742 +#27931 := [unit-resolution #22743 #27887]: #3027 +#22731 := (or #23403 #23397) +#22732 := [def-axiom]: #22731 +#28328 := [unit-resolution #22732 #27887]: #23397 +#27821 := [hypothesis]: #13368 +#27822 := [th-lemma #14256 #27821]: false +#27823 := [lemma #27822]: #13365 +#22751 := (or #23400 #13368 #23394) +#22753 := [def-axiom]: #22751 +#28329 := [unit-resolution #22753 #27823 #28328]: #23394 +#22761 := (or #23391 #23385) +#22762 := [def-axiom]: #22761 +#28330 := [unit-resolution #22762 #28329]: #23385 +#26364 := (* -1::int #3026) +#26365 := (+ uf_285 #26364) +#26366 := (>= #26365 0::int) +#27932 := (or #13149 #26366) +#27933 := [th-lemma]: #27932 +#27934 := [unit-resolution #27933 #27931]: #26366 +#22628 := (not #18355) +#27935 := [hypothesis]: #22206 +#22714 := (or #22201 #22628) +#22629 := [def-axiom]: #22714 +#27936 := [unit-resolution #22629 #27935]: #22628 +#27904 := (+ #3026 #18353) +#27906 := (>= #27904 0::int) +#27903 := (= #3026 #18352) +#27939 := (= #18352 #3026) +#27937 := (= #18351 #3016) +#27930 := (= ?x773!13 0::int) +#22715 := (not #18356) +#22706 := (or #22201 #22715) +#22717 := [def-axiom]: #22706 +#27928 := [unit-resolution #22717 #27935]: #22715 +#22627 := (or #22201 #18348) +#22713 := [def-axiom]: #22627 +#27929 := [unit-resolution #22713 #27935]: #18348 +#27927 := [th-lemma #27929 #27928]: #27930 +#27938 := [monotonicity #27927]: #27937 +#27940 := [monotonicity #27938]: #27939 +#27926 := [symm #27940]: #27903 +#27941 := (not #27903) +#27942 := (or #27941 #27906) +#27943 := [th-lemma]: #27942 +#27944 := [unit-resolution #27943 #27926]: #27906 +#27945 := [th-lemma #27944 #27936 #27934]: false +#27948 := [lemma #27945]: #22201 +#22757 := (or #23388 #22206 #23382) +#22758 := [def-axiom]: #22757 +#28331 := [unit-resolution #22758 #27948 #28330]: #23382 +#22773 := (or #23379 #23373) +#22774 := [def-axiom]: #22773 +#28332 := [unit-resolution #22774 #28331]: #23373 +#28333 := (or #23376 #13149 #23370) +#22769 := (or #23376 #13149 #13404 #23370) +#22770 := [def-axiom]: #22769 +#28334 := [unit-resolution #22770 #14256]: #28333 +#28335 := [unit-resolution #28334 #28332 #27931]: #23370 +#22808 := (or #23367 #13412) +#22796 := [def-axiom]: #22808 +#28968 := [unit-resolution #22796 #28335]: #13412 +#27800 := (or #27208 #13552) +#27790 := [hypothesis]: #13412 +#27217 := (* -1::int #27025) +#27349 := (+ uf_294 #27217) +#27350 := (<= #27349 0::int) +#27026 := (= uf_294 #27025) +#1391 := (uf_66 #15 #161 #1390) +#1392 := (pattern #1391) +#1393 := (uf_125 #1391 #15) +#7687 := (= #161 #1393) +#7691 := (forall (vars (?x319 T5) (?x320 int)) (:pat #1392) #7687) +#16976 := (~ #7691 #7691) +#16974 := (~ #7687 #7687) +#16975 := [refl]: #16974 +#16977 := [nnf-pos #16975]: #16976 +#1394 := (= #1393 #161) +#1395 := (forall (vars (?x319 T5) (?x320 int)) (:pat #1392) #1394) +#7692 := (iff #1395 #7691) +#7689 := (iff #1394 #7687) +#7690 := [rewrite]: #7689 +#7693 := [quant-intro #7690]: #7692 +#7686 := [asserted]: #1395 +#7696 := [mp #7686 #7693]: #7691 +#16978 := [mp~ #7696 #16977]: #7691 +#25864 := (not #7691) +#27029 := (or #25864 #27026) +#27030 := [quant-inst]: #27029 +#27791 := [unit-resolution #27030 #16978]: #27026 +#27794 := (not #27026) +#27795 := (or #27794 #27350) +#27796 := [th-lemma]: #27795 +#27797 := [unit-resolution #27796 #27791]: #27350 +#27209 := (not #27208) +#27798 := [hypothesis]: #27209 +#27799 := [th-lemma #27798 #27797 #27790]: false +#27801 := [lemma #27799]: #27800 +#29177 := [unit-resolution #27801 #28968]: #27208 +#27218 := (+ #23568 #27217) +#27219 := (<= #27218 0::int) +#29199 := (not #27219) +#27351 := (>= #27349 0::int) +#28150 := (uf_66 #22665 uf_294 #25857) +#28151 := (uf_125 #28150 #22665) +#28181 := (* -1::int #28151) +#28307 := (+ uf_294 #28181) +#28309 := (>= #28307 0::int) +#28152 := (= uf_294 #28151) +#28155 := (or #25864 #28152) +#28156 := [quant-inst]: #28155 +#28311 := [unit-resolution #28156 #16978]: #28152 +#28337 := (not #28152) +#28381 := (or #28337 #28309) +#28382 := [th-lemma]: #28381 +#28383 := [unit-resolution #28382 #28311]: #28309 +#28385 := (not #28309) +#29196 := (or #27351 #28385) +#28099 := (+ #27025 #28181) +#28100 := (<= #28099 0::int) +#28098 := (= #27025 #28151) +#29186 := (= #28151 #27025) +#29184 := (= #28150 #27024) +#29182 := (= #3180 #27024) +#29180 := (= #27024 #3180) +#28355 := (= #26356 uf_7) +#28353 := (= #26356 #23681) +#28354 := [monotonicity #25734]: #28353 +#28356 := [trans #28354 #27778]: #28355 +#29181 := [monotonicity #25734 #28356]: #29180 +#29183 := [symm #29181]: #29182 +#29178 := (= #28150 #3180) +#29179 := [monotonicity #25699 #27780]: #29178 +#29185 := [trans #29179 #29183]: #29184 +#29187 := [monotonicity #29185 #27310]: #29186 +#29188 := [symm #29187]: #28098 +#29189 := (not #28098) +#29190 := (or #29189 #28100) +#29191 := [th-lemma]: #29190 +#29192 := [unit-resolution #29191 #29188]: #28100 +#29193 := (not #28100) +#29194 := (or #27351 #29193 #28385) +#29195 := [th-lemma]: #29194 +#29197 := [unit-resolution #29195 #29192]: #29196 +#29198 := [unit-resolution #29197 #28383]: #27351 +#29200 := (not #27351) +#29205 := (or #29199 #29200) +#22792 := (or #23367 #23361) +#22778 := [def-axiom]: #22792 +#28345 := [unit-resolution #22778 #28335]: #23361 +#22788 := (or #23367 #23203) +#22791 := [def-axiom]: #22788 +#28346 := [unit-resolution #22791 #28335]: #23203 +#22659 := (or #23367 #11992) +#22817 := [def-axiom]: #22659 +#28347 := [unit-resolution #22817 #28335]: #11992 +#22790 := (or #23367 #15764) +#22794 := [def-axiom]: #22790 +#28348 := [unit-resolution #22794 #28335]: #15764 +#22799 := (or #23367 #13876) +#22801 := [def-axiom]: #22799 +#28349 := [unit-resolution #22801 #28335]: #13876 +#26357 := (uf_66 #25821 uf_293 #26356) +#26358 := (uf_125 #26357 #25821) +#26439 := (>= #26358 0::int) +#22807 := (or #23367 #13409) +#22805 := [def-axiom]: #22807 +#28350 := [unit-resolution #22805 #28335]: #13409 +#26455 := (* -1::int #26358) +#26528 := (+ uf_293 #26455) +#26529 := (<= #26528 0::int) +#26359 := (= uf_293 #26358) +#27978 := (uf_66 #22665 uf_293 #25857) +#27979 := (uf_125 #27978 #22665) +#28368 := (= #27979 #26358) +#28366 := (= #26358 #27979) +#28351 := (= #25821 #22665) +#28352 := [trans #25734 #25682]: #28351 +#28363 := (= #26357 #27978) +#28361 := (= #3073 #27978) +#28359 := (= #27978 #3073) +#28360 := [monotonicity #25699 #27780]: #28359 +#28362 := [symm #28360]: #28361 +#28357 := (= #26357 #3073) +#28358 := [monotonicity #25734 #28356]: #28357 +#28364 := [trans #28358 #28362]: #28363 +#28367 := [monotonicity #28364 #28352]: #28366 +#28369 := [symm #28367]: #28368 +#27980 := (= uf_293 #27979) +#27982 := (or #25864 #27980) +#27983 := [quant-inst]: #27982 +#28365 := [unit-resolution #27983 #16978]: #27980 +#28370 := [trans #28365 #28369]: #26359 +#26867 := (not #26359) +#28371 := (or #26867 #26529) +#28372 := [th-lemma]: #28371 +#28373 := [unit-resolution #28372 #28370]: #26529 +#28374 := (not #26529) +#28375 := (or #26439 #22348 #28374) +#28376 := [th-lemma]: #28375 +#28377 := [unit-resolution #28376 #28373 #28350]: #26439 +#26440 := (not #26439) +#26892 := (or #23355 #26440 #13875 #22491 #22486 #23208) +#26737 := (uf_66 #2960 #26358 uf_7) +#26738 := (uf_110 uf_273 #26737) +#26741 := (= uf_299 #26738) +#26847 := (= #3074 #26738) +#26860 := (= #26738 #3074) +#26854 := (= #26737 #3073) +#26852 := (= #26358 uf_293) +#26362 := (or #25864 #26359) +#26363 := [quant-inst]: #26362 +#26851 := [unit-resolution #26363 #16978]: #26359 +#26853 := [symm #26851]: #26852 +#26855 := [monotonicity #26853]: #26854 +#26861 := [monotonicity #26855]: #26860 +#26862 := [symm #26861]: #26847 +#26863 := (= uf_299 #3074) +#26856 := [hypothesis]: #11992 +#26857 := [hypothesis]: #23358 +#22849 := (or #23355 #12020) +#22850 := [def-axiom]: #22849 +#26858 := [unit-resolution #22850 #26857]: #12020 +#26859 := [symm #26858]: #3089 +#26864 := [trans #26859 #26856]: #26863 +#26865 := [trans #26864 #26862]: #26741 +#26722 := (<= #26358 4294967295::int) +#26866 := [hypothesis]: #15764 +#26530 := (>= #26528 0::int) +#26868 := (or #26867 #26530) +#26869 := [th-lemma]: #26868 +#26870 := [unit-resolution #26869 #26851]: #26530 +#26871 := (not #26530) +#26872 := (or #26722 #22491 #26871) +#26873 := [th-lemma]: #26872 +#26874 := [unit-resolution #26873 #26870 #26866]: #26722 +#26764 := (+ uf_272 #26455) +#26765 := (<= #26764 0::int) +#26876 := (not #26765) +#26875 := [hypothesis]: #13876 +#26877 := (or #26876 #13875 #26871) +#26878 := [th-lemma]: #26877 +#26879 := [unit-resolution #26878 #26870 #26875]: #26876 +#26744 := (not #26741) +#26723 := (not #26722) +#26889 := (or #26723 #26744 #26765) +#26880 := [hypothesis]: #26439 +#22857 := (or #23355 #23349) +#22842 := [def-axiom]: #22857 +#26881 := [unit-resolution #22842 #26857]: #23349 +#26534 := (+ uf_292 #13737) +#26537 := (<= #26534 0::int) +#26882 := (or #12069 #26537) +#26883 := [th-lemma]: #26882 +#26884 := [unit-resolution #26883 #26858]: #26537 +#26788 := [hypothesis]: #23203 +#22855 := (or #23355 #13697) +#22856 := [def-axiom]: #22855 +#26885 := [unit-resolution #22856 #26857]: #13697 +#26779 := (not #26537) +#26794 := (or #22457 #13698 #23208 #26779) +#26782 := [hypothesis]: #13697 +#26605 := (+ uf_294 #18945) +#26606 := (<= #26605 0::int) +#26617 := (+ uf_292 #18958) +#26618 := (>= #26617 0::int) +#26778 := (not #26618) +#26772 := [hypothesis]: #26537 +#22869 := (not #18960) +#26783 := [hypothesis]: #22462 +#22828 := (or #22457 #22869) +#22859 := [def-axiom]: #22828 +#26784 := [unit-resolution #22859 #26783]: #22869 +#26780 := (or #26778 #18960 #26779) +#26773 := [hypothesis]: #22869 +#26776 := [hypothesis]: #26618 +#26777 := [th-lemma #26776 #26773 #26772]: false +#26781 := [lemma #26777]: #26780 +#26785 := [unit-resolution #26781 #26784 #26772]: #26778 +#26789 := (or #26606 #26618) +#22892 := (or #22457 #18609) +#22893 := [def-axiom]: #22892 +#26786 := [unit-resolution #22893 #26783]: #18609 +#22887 := (or #22457 #18608) +#22868 := [def-axiom]: #22887 +#26787 := [unit-resolution #22868 #26783]: #18608 +#26631 := (or #23208 #22441 #22442 #26606 #26618) +#26594 := (+ #18613 #13902) +#26595 := (<= #26594 0::int) +#26596 := (+ ?x776!15 #13433) +#26597 := (>= #26596 0::int) +#26598 := (or #22442 #26597 #26595 #22441) +#26632 := (or #23208 #26598) +#26639 := (iff #26632 #26631) +#26626 := (or #22441 #22442 #26606 #26618) +#26634 := (or #23208 #26626) +#26637 := (iff #26634 #26631) +#26638 := [rewrite]: #26637 +#26635 := (iff #26632 #26634) +#26629 := (iff #26598 #26626) +#26623 := (or #22442 #26606 #26618 #22441) +#26627 := (iff #26623 #26626) +#26628 := [rewrite]: #26627 +#26624 := (iff #26598 #26623) +#26621 := (iff #26595 #26618) +#26611 := (+ #13902 #18613) +#26614 := (<= #26611 0::int) +#26619 := (iff #26614 #26618) +#26620 := [rewrite]: #26619 +#26615 := (iff #26595 #26614) +#26612 := (= #26594 #26611) +#26613 := [rewrite]: #26612 +#26616 := [monotonicity #26613]: #26615 +#26622 := [trans #26616 #26620]: #26621 +#26609 := (iff #26597 #26606) +#26599 := (+ #13433 ?x776!15) +#26602 := (>= #26599 0::int) +#26607 := (iff #26602 #26606) +#26608 := [rewrite]: #26607 +#26603 := (iff #26597 #26602) +#26600 := (= #26596 #26599) +#26601 := [rewrite]: #26600 +#26604 := [monotonicity #26601]: #26603 +#26610 := [trans #26604 #26608]: #26609 +#26625 := [monotonicity #26610 #26622]: #26624 +#26630 := [trans #26625 #26628]: #26629 +#26636 := [monotonicity #26630]: #26635 +#26640 := [trans #26636 #26638]: #26639 +#26633 := [quant-inst]: #26632 +#26641 := [mp #26633 #26640]: #26631 +#26790 := [unit-resolution #26641 #26788 #26787 #26786]: #26789 +#26791 := [unit-resolution #26790 #26785]: #26606 +#22888 := (not #18947) +#22894 := (or #22457 #22888) +#22862 := [def-axiom]: #22894 +#26792 := [unit-resolution #22862 #26783]: #22888 +#26793 := [th-lemma #26792 #26791 #26782]: false +#26795 := [lemma #26793]: #26794 +#26886 := [unit-resolution #26795 #26885 #26788 #26884]: #22457 +#22884 := (or #23352 #23346 #22462) +#22864 := [def-axiom]: #22884 +#26887 := [unit-resolution #22864 #26886 #26881]: #23346 +#22901 := (or #23343 #23335) +#22906 := [def-axiom]: #22901 +#26888 := [unit-resolution #22906 #26887]: #23335 +#26801 := (or #23340 #26440 #26723 #26744 #26765) +#26728 := (+ #26358 #13338) +#26729 := (>= #26728 0::int) +#26736 := (= #26738 uf_299) +#26739 := (not #26736) +#26740 := (or #26739 #26440 #26729 #26723) +#26802 := (or #23340 #26740) +#26809 := (iff #26802 #26801) +#26796 := (or #26440 #26723 #26744 #26765) +#26804 := (or #23340 #26796) +#26807 := (iff #26804 #26801) +#26808 := [rewrite]: #26807 +#26805 := (iff #26802 #26804) +#26799 := (iff #26740 #26796) +#26745 := (or #26744 #26440 #26765 #26723) +#26797 := (iff #26745 #26796) +#26798 := [rewrite]: #26797 +#26746 := (iff #26740 #26745) +#26775 := (iff #26729 #26765) +#26759 := (+ #13338 #26358) +#26756 := (>= #26759 0::int) +#26766 := (iff #26756 #26765) +#26774 := [rewrite]: #26766 +#26762 := (iff #26729 #26756) +#26760 := (= #26728 #26759) +#26761 := [rewrite]: #26760 +#26763 := [monotonicity #26761]: #26762 +#26677 := [trans #26763 #26774]: #26775 +#26757 := (iff #26739 #26744) +#26742 := (iff #26736 #26741) +#26743 := [rewrite]: #26742 +#26758 := [monotonicity #26743]: #26757 +#26769 := [monotonicity #26758 #26677]: #26746 +#26800 := [trans #26769 #26798]: #26799 +#26806 := [monotonicity #26800]: #26805 +#26810 := [trans #26806 #26808]: #26809 +#26803 := [quant-inst]: #26802 +#26811 := [mp #26803 #26810]: #26801 +#26890 := [unit-resolution #26811 #26888 #26880]: #26889 +#26891 := [unit-resolution #26890 #26879 #26874 #26865]: false +#26893 := [lemma #26891]: #26892 +#28378 := [unit-resolution #26893 #28377 #28349 #28348 #28347 #28346]: #23355 +#22831 := (or #23364 #23324 #23358) +#22833 := [def-axiom]: #22831 +#28379 := [unit-resolution #22833 #28378 #28345]: #23324 +#22897 := (or #23321 #13698) +#22902 := [def-axiom]: #22897 +#28380 := [unit-resolution #22902 #28379]: #13698 +#29203 := (or #29199 #29200 #13697) +#29201 := (or #29199 #26240 #29200 #13697) +#29202 := [th-lemma]: #29201 +#29204 := [unit-resolution #29202 #26238]: #29203 +#29206 := [unit-resolution #29204 #28380]: #29205 +#29207 := [unit-resolution #29206 #29198]: #29199 +#28412 := (or #26056 #25897 #27209 #27219 #27227) +#27204 := (or #27203 #27201 #27198 #27195) +#27205 := (not #27204) +#27206 := (+ #27025 #25898) +#27207 := (>= #27206 0::int) +#27210 := (or #27209 #27207 #25897 #27205) +#28413 := (or #26056 #27210) +#28419 := (iff #28413 #28412) +#27233 := (or #25897 #27209 #27219 #27227) +#28409 := (or #26056 #27233) +#28417 := (iff #28409 #28412) +#28418 := [rewrite]: #28417 +#28415 := (iff #28413 #28409) +#27236 := (iff #27210 #27233) +#27230 := (or #27209 #27219 #25897 #27227) +#27234 := (iff #27230 #27233) +#27235 := [rewrite]: #27234 +#27231 := (iff #27210 #27230) +#27228 := (iff #27205 #27227) +#27225 := (iff #27204 #27224) +#27226 := [rewrite]: #27225 +#27229 := [monotonicity #27226]: #27228 +#27222 := (iff #27207 #27219) +#27211 := (+ #25898 #27025) +#27214 := (>= #27211 0::int) +#27220 := (iff #27214 #27219) +#27221 := [rewrite]: #27220 +#27215 := (iff #27207 #27214) +#27212 := (= #27206 #27211) +#27213 := [rewrite]: #27212 +#27216 := [monotonicity #27213]: #27215 +#27223 := [trans #27216 #27221]: #27222 +#27232 := [monotonicity #27223 #27229]: #27231 +#27237 := [trans #27232 #27235]: #27236 +#28416 := [monotonicity #27237]: #28415 +#28429 := [trans #28416 #28418]: #28419 +#28414 := [quant-inst]: #28413 +#28430 := [mp #28414 #28429]: #28412 +#29208 := [unit-resolution #28430 #20893 #29207 #29177 #27446]: #27227 +#28434 := (or #27224 #27200) +#28435 := [def-axiom]: #28434 +#29209 := [unit-resolution #28435 #29208]: #27200 +#29240 := [symm #29209]: #29239 +#29237 := (= #27027 #27199) +#29235 := (= #26583 #27196) +#29233 := (= #27196 #26583) +#29231 := (= #27192 #3180) +#26846 := (uf_116 #3180) +#26905 := (uf_43 #23566 #26846) +#29225 := (= #26905 #3180) +#26906 := (= #3180 #26905) +#26896 := (uf_48 #3180 #23566) +#26897 := (= uf_9 #26896) +#29211 := (= #3181 #26896) +#28282 := (= #26896 #3181) +#28283 := [monotonicity #25690]: #28282 +#29212 := [symm #28283]: #29211 +#26923 := (= #23566 #26922) +#26996 := (* uf_294 #25803) +#28123 := (+ #25805 #26996) +#28124 := (uf_43 #23566 #28123) +#28266 := (uf_13 #28124) +#28306 := (= #28266 #26922) +#28313 := (= #26922 #28266) +#28408 := (= #3180 #28124) +#28119 := (uf_66 #22665 uf_294 #23566) +#28125 := (= #28119 #28124) +#28126 := (not #28125) +#28120 := (uf_138 #28119 #22665) +#28121 := (= uf_9 #28120) +#28122 := (not #28121) +#28129 := (or #28122 #28126) +#28132 := (not #28129) +#28139 := (or #25566 #28132) +#28127 := (or #28126 #28122) +#28128 := (not #28127) +#28140 := (or #25566 #28128) +#28141 := (iff #28140 #28139) +#28143 := (iff #28139 #28139) +#28213 := [rewrite]: #28143 +#28133 := (iff #28128 #28132) +#28130 := (iff #28127 #28129) +#28131 := [rewrite]: #28130 +#28134 := [monotonicity #28131]: #28133 +#28142 := [monotonicity #28134]: #28141 +#28214 := [trans #28142 #28213]: #28141 +#28144 := [quant-inst]: #28140 +#28203 := [mp #28144 #28214]: #28139 +#28400 := [unit-resolution #28203 #21109]: #28132 +#28206 := (or #28129 #28125) +#28207 := [def-axiom]: #28206 +#28401 := [unit-resolution #28207 #28400]: #28125 +#28406 := (= #3180 #28119) +#28404 := (= #28119 #3180) +#28405 := [monotonicity #25699 #25690]: #28404 +#28407 := [symm #28405]: #28406 +#28312 := [trans #28407 #28401]: #28408 +#28153 := [monotonicity #28312]: #28313 +#28389 := [symm #28153]: #28306 +#28267 := (= #23566 #28266) +#28215 := (or #23633 #28267) +#28218 := [quant-inst]: #28215 +#28399 := [unit-resolution #28218 #23143]: #28267 +#28394 := [trans #28399 #28389]: #26923 +#28090 := (not #26923) +#26925 := (iff #26897 #26923) +#28110 := (or #24884 #26925) +#26924 := (iff #26923 #26897) +#28111 := (or #24884 #26924) +#28113 := (iff #28111 #28110) +#28114 := (iff #28110 #28110) +#28115 := [rewrite]: #28114 +#26926 := (iff #26924 #26925) +#26927 := [rewrite]: #26926 +#28108 := [monotonicity #26927]: #28113 +#28116 := [trans #28108 #28115]: #28113 +#28112 := [quant-inst]: #28111 +#28117 := [mp #28112 #28116]: #28110 +#28303 := [unit-resolution #28117 #23130]: #26925 +#26904 := (not #26897) +#28297 := (iff #18425 #26904) +#28285 := (iff #12345 #26897) +#28281 := (iff #26897 #12345) +#28284 := [monotonicity #28283]: #28281 +#28286 := [symm #28284]: #28285 +#28298 := [monotonicity #28286]: #28297 +#28304 := [hypothesis]: #18425 +#28299 := [mp #28304 #28298]: #26904 +#28118 := (not #26925) +#28091 := (or #28118 #26897 #28090) +#28092 := [def-axiom]: #28091 +#28398 := [unit-resolution #28092 #28299 #28303]: #28090 +#28395 := [unit-resolution #28398 #28394]: false +#28396 := [lemma #28395]: #12345 +#29213 := [trans #28396 #29212]: #26897 +#26908 := (or #26904 #26906) +#28216 := (or #24868 #26904 #26906) +#26907 := (or #26906 #26904) +#28217 := (or #24868 #26907) +#28279 := (iff #28217 #28216) +#28259 := (or #24868 #26908) +#28270 := (iff #28259 #28216) +#28271 := [rewrite]: #28270 +#28260 := (iff #28217 #28259) +#26909 := (iff #26907 #26908) +#26910 := [rewrite]: #26909 +#28212 := [monotonicity #26910]: #28260 +#28280 := [trans #28212 #28271]: #28279 +#28223 := [quant-inst]: #28217 +#28265 := [mp #28223 #28280]: #28216 +#29214 := [unit-resolution #28265 #18153]: #26908 +#29215 := [unit-resolution #29214 #29213]: #26906 +#29226 := [symm #29215]: #29225 +#29229 := (= #27192 #26905) +#27651 := (* #25803 #27025) +#27654 := (+ #26062 #27651) +#27657 := (uf_43 #23566 #27654) +#29223 := (= #27657 #26905) +#29217 := (= #27654 #26846) +#29170 := (= #26846 #27654) +#29171 := (* -1::int #27654) +#29172 := (+ #26846 #29171) +#29173 := (<= #29172 0::int) +#28024 := (* -1::int #25805) +#28025 := (+ #23468 #28024) +#28027 := (>= #28025 0::int) +#28017 := (= #23468 #25805) +#29165 := (= #2961 #25805) +#29247 := [symm #25698]: #29165 +#29210 := [trans #27412 #29247]: #28017 +#29176 := (not #28017) +#29255 := (or #29176 #28027) +#29256 := [th-lemma]: #29255 +#29257 := [unit-resolution #29256 #29210]: #28027 +#28574 := (* -1::int #26062) +#28575 := (+ #23468 #28574) +#28576 := (<= #28575 0::int) +#28573 := (= #23468 #26062) +#29258 := [symm #27411]: #28573 +#29259 := (not #28573) +#29260 := (or #29259 #28576) +#29261 := [th-lemma]: #29260 +#29262 := [unit-resolution #29261 #29258]: #28576 +#28268 := (uf_116 #28124) +#28272 := (* -1::int #28268) +#28498 := (+ #26846 #28272) +#28499 := (<= #28498 0::int) +#28488 := (= #26846 #28268) +#29267 := (= #28268 #26846) +#29265 := (= #28124 #3180) +#29263 := (= #28124 #28119) +#29264 := [symm #28401]: #29263 +#29266 := [trans #29264 #28405]: #29265 +#29268 := [monotonicity #29266]: #29267 +#29269 := [symm #29268]: #28488 +#29270 := (not #28488) +#29271 := (or #29270 #28499) +#29272 := [th-lemma]: #29271 +#29273 := [unit-resolution #29272 #29269]: #28499 +#28273 := (+ #26996 #28272) +#28274 := (+ #25805 #28273) +#28466 := (>= #28274 0::int) +#28275 := (= #28274 0::int) +#28475 := (or #23639 #28275) +#28269 := (= #28123 #28268) +#28476 := (or #23639 #28269) +#28460 := (iff #28476 #28475) +#28462 := (iff #28475 #28475) +#28457 := [rewrite]: #28462 +#28276 := (iff #28269 #28275) +#28277 := [rewrite]: #28276 +#28461 := [monotonicity #28277]: #28460 +#28463 := [trans #28461 #28457]: #28460 +#28459 := [quant-inst]: #28476 +#28464 := [mp #28459 #28463]: #28475 +#29274 := [unit-resolution #28464 #23137]: #28275 +#29275 := (not #28275) +#29276 := (or #29275 #28466) +#29277 := [th-lemma]: #29276 +#29278 := [unit-resolution #29277 #29274]: #28466 +#28101 := (>= #28099 0::int) +#29279 := (or #29189 #28101) +#29280 := [th-lemma]: #29279 +#29281 := [unit-resolution #29280 #29188]: #28101 +#28308 := (<= #28307 0::int) +#28338 := (or #28337 #28308) +#28339 := [th-lemma]: #28338 +#28340 := [unit-resolution #28339 #28311]: #28308 +#28029 := (>= #25803 1::int) +#28028 := (= #25803 1::int) +#2932 := (uf_139 uf_7) +#2933 := (= #2932 1::int) +#11835 := [asserted]: #2933 +#29282 := (= #25803 #2932) +#29283 := [monotonicity #25690]: #29282 +#29284 := [trans #29283 #11835]: #28028 +#29285 := (not #28028) +#29286 := (or #29285 #28029) +#29287 := [th-lemma]: #29286 +#29288 := [unit-resolution #29287 #29284]: #28029 +#28023 := (<= #25803 1::int) +#29289 := (or #29285 #28023) +#29290 := [th-lemma]: #29289 +#29291 := [unit-resolution #29290 #29284]: #28023 +#28341 := (not #28308) +#29298 := (not #28101) +#29297 := (not #28576) +#29296 := (not #28023) +#29295 := (not #28029) +#29294 := (not #28027) +#29293 := (not #28466) +#29292 := (not #28499) +#29299 := (or #29173 #29292 #29293 #29294 #29295 #29296 #29295 #29296 #29297 #29298 #28341) +#29300 := [th-lemma]: #29299 +#29301 := [unit-resolution #29300 #29291 #29288 #28340 #29281 #29278 #29273 #29262 #29257]: #29173 +#29174 := (>= #29172 0::int) +#28026 := (<= #28025 0::int) +#29302 := (or #29176 #28026) +#29303 := [th-lemma]: #29302 +#29304 := [unit-resolution #29303 #29210]: #28026 +#28577 := (>= #28575 0::int) +#29305 := (or #29259 #28577) +#29306 := [th-lemma]: #29305 +#29307 := [unit-resolution #29306 #29258]: #28577 +#28500 := (>= #28498 0::int) +#29308 := (or #29270 #28500) +#29309 := [th-lemma]: #29308 +#29310 := [unit-resolution #29309 #29269]: #28500 +#28465 := (<= #28274 0::int) +#29311 := (or #29275 #28465) +#29312 := [th-lemma]: #29311 +#29313 := [unit-resolution #29312 #29274]: #28465 +#29317 := (not #28577) +#29316 := (not #28026) +#29315 := (not #28465) +#29314 := (not #28500) +#29318 := (or #29174 #29314 #29315 #29316 #29295 #29296 #29295 #29296 #29317 #29193 #28385) +#29319 := [th-lemma]: #29318 +#29320 := [unit-resolution #29319 #29291 #29288 #28383 #29192 #29313 #29310 #29307 #29304]: #29174 +#29322 := (not #29174) +#29321 := (not #29173) +#29323 := (or #29170 #29321 #29322) +#29324 := [th-lemma]: #29323 +#29325 := [unit-resolution #29324 #29320 #29301]: #29170 +#28689 := [symm #29325]: #29217 +#28690 := [monotonicity #28689]: #29223 +#29227 := (= #27192 #27657) +#27640 := (uf_66 #24856 #27025 #23566) +#27660 := (= #27640 #27657) +#27663 := (not #27660) +#27641 := (uf_138 #27640 #24856) +#27642 := (= uf_9 #27641) +#27643 := (not #27642) +#27669 := (or #27643 #27663) +#27627 := (not #27669) +#28558 := (or #25566 #27627) +#27644 := (* #27025 #25803) +#27645 := (+ #26062 #27644) +#27646 := (uf_43 #23566 #27645) +#27647 := (= #27640 #27646) +#27648 := (not #27647) +#27649 := (or #27648 #27643) +#27650 := (not #27649) +#28559 := (or #25566 #27650) +#28561 := (iff #28559 #28558) +#28563 := (iff #28558 #28558) +#28564 := [rewrite]: #28563 +#27625 := (iff #27650 #27627) +#27624 := (iff #27649 #27669) +#27666 := (or #27663 #27643) +#27670 := (iff #27666 #27669) +#27671 := [rewrite]: #27670 +#27667 := (iff #27649 #27666) +#27664 := (iff #27648 #27663) +#27661 := (iff #27647 #27660) +#27658 := (= #27646 #27657) +#27655 := (= #27645 #27654) +#27652 := (= #27644 #27651) +#27653 := [rewrite]: #27652 +#27656 := [monotonicity #27653]: #27655 +#27659 := [monotonicity #27656]: #27658 +#27662 := [monotonicity #27659]: #27661 +#27665 := [monotonicity #27662]: #27664 +#27668 := [monotonicity #27665]: #27667 +#27626 := [trans #27668 #27671]: #27624 +#27628 := [monotonicity #27626]: #27625 +#28562 := [monotonicity #27628]: #28561 +#28565 := [trans #28562 #28564]: #28561 +#28560 := [quant-inst]: #28559 +#28566 := [mp #28560 #28565]: #28558 +#29219 := [unit-resolution #28566 #21109]: #27627 +#28569 := (or #27669 #27660) +#28570 := [def-axiom]: #28569 +#29220 := [unit-resolution #28570 #29219]: #27660 +#29221 := (= #27192 #27640) +#29222 := [monotonicity #25843]: #29221 +#29228 := [trans #29222 #29220]: #29227 +#28664 := [trans #29228 #28690]: #29229 +#28668 := [trans #28664 #29226]: #29231 +#28660 := [monotonicity #28668]: #29233 +#28667 := [symm #28660]: #29235 +#28669 := [monotonicity #28667]: #29237 +#28670 := [trans #28669 #29240]: #29241 +#28671 := [trans #28670 #25865]: #29243 +#28672 := [monotonicity #28671]: #28541 +#28692 := [symm #28672]: #28691 +#28693 := [trans #14262 #28692]: #27033 +#28617 := (not #27033) +#28618 := (or #27089 #28617) +#28619 := [def-axiom]: #28618 +#28661 := [unit-resolution #28619 #28693]: #27089 +#28731 := (= #24105 #27047) +#28702 := (= #27047 #24105) +#28698 := (= #26922 uf_7) +#28678 := (= #26922 #23566) +#28278 := (or #28118 #26904 #26923) +#28300 := [def-axiom]: #28278 +#28694 := [unit-resolution #28300 #29213 #28303]: #26923 +#28681 := [symm #28694]: #28678 +#28699 := [trans #28681 #25690]: #28698 +#28701 := [monotonicity #28699]: #28702 +#28732 := [symm #28701]: #28731 +#28733 := [trans #27766 #28732]: #27065 +#28597 := (not #27038) +#28679 := (iff #11881 #28597) +#28704 := (iff #11878 #27038) +#28703 := (iff #27038 #11878) +#28714 := (= #27037 #2971) +#28713 := (= #27036 #2958) +#28734 := (= #27036 #23686) +#28712 := [monotonicity #28671]: #28734 +#28709 := [trans #28712 #25786]: #28713 +#28688 := [monotonicity #28709]: #28714 +#28550 := [monotonicity #28688]: #28703 +#28705 := [symm #28550]: #28704 +#28677 := [monotonicity #28705]: #28679 +#28700 := [mp #14265 #28677]: #28597 +#27260 := (uf_66 #25957 #27025 #23566) +#27261 := (uf_58 #3149 #27260) +#27262 := (uf_136 #27261) +#27263 := (= uf_9 #27262) +#28468 := (not #27263) +#28716 := (iff #28468 #27044) +#28728 := (iff #27263 #27043) +#28720 := (iff #27043 #27263) +#28706 := (= #27042 #27262) +#28818 := (= #26583 #27261) +#28749 := (= #27261 #26583) +#28745 := (= #27260 #3180) +#28743 := (= #27025 uf_294) +#28314 := (= #28151 uf_294) +#28315 := [symm #28311]: #28314 +#28744 := [trans #29188 #28315]: #28743 +#28746 := [monotonicity #26202 #28744 #25690]: #28745 +#28826 := [monotonicity #28746]: #28749 +#28819 := [symm #28826]: #28818 +#28719 := [monotonicity #28819]: #28706 +#28727 := [monotonicity #28719]: #28720 +#28507 := [symm #28727]: #28728 +#28521 := [monotonicity #28507]: #28716 +#27264 := (uf_24 uf_273 #27260) +#27265 := (= uf_9 #27264) +#27266 := (not #27265) +#27270 := (or #27263 #27266) +#27273 := (not #27270) +#28443 := (or #25988 #22661 #25969 #25972 #27209 #27219 #27273) +#27267 := (or #27266 #27263) +#27268 := (not #27267) +#27269 := (or #25972 #27209 #27207 #22661 #25969 #27268) +#28444 := (or #25988 #27269) +#28450 := (iff #28444 #28443) +#27279 := (or #22661 #25969 #25972 #27209 #27219 #27273) +#28446 := (or #25988 #27279) +#28448 := (iff #28446 #28443) +#28449 := [rewrite]: #28448 +#28447 := (iff #28444 #28446) +#27282 := (iff #27269 #27279) +#27276 := (or #25972 #27209 #27219 #22661 #25969 #27273) +#27280 := (iff #27276 #27279) +#27281 := [rewrite]: #27280 +#27277 := (iff #27269 #27276) +#27274 := (iff #27268 #27273) +#27271 := (iff #27267 #27270) +#27272 := [rewrite]: #27271 +#27275 := [monotonicity #27272]: #27274 +#27278 := [monotonicity #27223 #27275]: #27277 +#27283 := [trans #27278 #27281]: #27282 +#28442 := [monotonicity #27283]: #28447 +#28451 := [trans #28442 #28449]: #28450 +#28445 := [quant-inst]: #28444 +#28452 := [mp #28445 #28451]: #28443 +#28742 := [unit-resolution #28452 #21553 #14251 #29207 #29177 #27341 #26234]: #27273 +#28469 := (or #27270 #28468) +#28470 := [def-axiom]: #28469 +#28737 := [unit-resolution #28470 #28742]: #28468 +#28739 := [mp #28737 #28521]: #27044 +#28612 := (or #27083 #27043) +#28613 := [def-axiom]: #28612 +#28740 := [unit-resolution #28613 #28739]: #27083 +#28608 := (or #27103 #27038 #27080 #27086 #27092) +#28544 := [def-axiom]: #28608 +#28738 := [unit-resolution #28544 #28740 #28700 #28733 #28661 #28658]: false +#28741 := [lemma #28738]: #27103 +#28624 := (or #27106 #27098) +#28625 := [def-axiom]: #28624 +#29039 := [unit-resolution #28625 #28741]: #27106 +#28224 := (uf_66 #25957 #28151 #23566) +#28228 := (uf_24 uf_273 #28224) +#28229 := (= uf_9 #28228) +#28230 := (not #28229) +#28225 := (uf_58 #3149 #28224) +#28226 := (uf_136 #28225) +#28227 := (= uf_9 #28226) +#28234 := (or #28227 #28230) +#28324 := (iff #18428 #28230) +#28322 := (iff #12348 #28229) +#28320 := (iff #28229 #12348) +#28318 := (= #28228 #3183) +#28316 := (= #28224 #3180) +#28317 := [monotonicity #26202 #28315 #25690]: #28316 +#28319 := [monotonicity #28317]: #28318 +#28321 := [monotonicity #28319]: #28320 +#28323 := [symm #28321]: #28322 +#28325 := [monotonicity #28323]: #28324 +#28310 := [hypothesis]: #18428 +#28326 := [mp #28310 #28325]: #28230 +#28262 := (or #28234 #28229) +#28263 := [def-axiom]: #28262 +#28327 := [unit-resolution #28263 #28326]: #28234 +#28172 := (>= #28151 0::int) +#22803 := (or #23367 #13930) +#22793 := [def-axiom]: #22803 +#28336 := [unit-resolution #22793 #28335]: #13930 +#28342 := (or #28172 #22487 #28341) +#28343 := [th-lemma]: #28342 +#28344 := [unit-resolution #28343 #28340 #28336]: #28172 +#28182 := (+ #23568 #28181) +#28183 := (<= #28182 0::int) +#28384 := (not #28183) +#28386 := (or #28384 #26240 #13697 #28385) +#28387 := [th-lemma]: #28386 +#28388 := [unit-resolution #28387 #28383 #28380 #26238]: #28384 +#28237 := (not #28234) +#28173 := (not #28172) +#28390 := (or #28173 #28183 #28237) +#28248 := (or #25988 #22661 #25969 #25972 #28173 #28183 #28237) +#28231 := (or #28230 #28227) +#28232 := (not #28231) +#28170 := (+ #28151 #25898) +#28171 := (>= #28170 0::int) +#28233 := (or #25972 #28173 #28171 #22661 #25969 #28232) +#28249 := (or #25988 #28233) +#28256 := (iff #28249 #28248) +#28243 := (or #22661 #25969 #25972 #28173 #28183 #28237) +#28251 := (or #25988 #28243) +#28254 := (iff #28251 #28248) +#28255 := [rewrite]: #28254 +#28252 := (iff #28249 #28251) +#28246 := (iff #28233 #28243) +#28240 := (or #25972 #28173 #28183 #22661 #25969 #28237) +#28244 := (iff #28240 #28243) +#28245 := [rewrite]: #28244 +#28241 := (iff #28233 #28240) +#28238 := (iff #28232 #28237) +#28235 := (iff #28231 #28234) +#28236 := [rewrite]: #28235 +#28239 := [monotonicity #28236]: #28238 +#28186 := (iff #28171 #28183) +#28175 := (+ #25898 #28151) +#28178 := (>= #28175 0::int) +#28184 := (iff #28178 #28183) +#28185 := [rewrite]: #28184 +#28179 := (iff #28171 #28178) +#28176 := (= #28170 #28175) +#28177 := [rewrite]: #28176 +#28180 := [monotonicity #28177]: #28179 +#28187 := [trans #28180 #28185]: #28186 +#28242 := [monotonicity #28187 #28239]: #28241 +#28247 := [trans #28242 #28245]: #28246 +#28253 := [monotonicity #28247]: #28252 +#28257 := [trans #28253 #28255]: #28256 +#28250 := [quant-inst]: #28249 +#28258 := [mp #28250 #28257]: #28248 +#28391 := [unit-resolution #28258 #21553 #14251 #27341 #26234]: #28390 +#28392 := [unit-resolution #28391 #28388 #28344 #28327]: false +#28393 := [lemma #28392]: #12348 +#28676 := (or #27115 #18428 #27109) +#28663 := [def-axiom]: #28676 +#29045 := [unit-resolution #28663 #28393 #29039]: #27115 +#28636 := (or #28635 #12354 #27112) +#28634 := [def-axiom]: #28636 +#29048 := [unit-resolution #28634 #29045]: #29058 +#29033 := [unit-resolution #29048 #28730]: #12354 +#29036 := (or #18434 #23258) +#22696 := (not #18799) +#22967 := (+ uf_293 #13488) +#22969 := (>= #22967 0::int) +#28951 := [hypothesis]: #23288 +#22959 := (or #23285 #12647) +#22960 := [def-axiom]: #22959 +#28952 := [unit-resolution #22960 #28951]: #12647 +#28947 := (or #12653 #22969) +#28953 := [th-lemma]: #28947 +#28949 := [unit-resolution #28953 #28952]: #22969 +#28954 := (not #22969) +#28935 := (or #13491 #28954) +#28955 := (or #13491 #13875 #28954) +#28956 := [th-lemma]: #28955 +#28936 := [unit-resolution #28956 #28349]: #28935 +#28943 := [unit-resolution #28936 #28949]: #13491 +#28962 := (= #3074 #3222) +#28959 := (= #3222 #3074) +#28940 := (= #3221 #3073) +#28944 := [symm #28952]: #3259 +#28945 := [monotonicity #28944]: #28940 +#28961 := [monotonicity #28945]: #28959 +#28960 := [symm #28961]: #28962 +#28963 := (= uf_303 #3074) +#22965 := (or #23285 #12644) +#22970 := [def-axiom]: #22965 +#28946 := [unit-resolution #22970 #28951]: #12644 +#28942 := [symm #28946]: #3258 +#28964 := [trans #28942 #28347]: #28963 +#28965 := [trans #28964 #28960]: #12404 +#22823 := (or #22233 #22231 #13490) +#23030 := [def-axiom]: #22823 +#28966 := [unit-resolution #23030 #28965 #28943]: #22233 +#22704 := (or #23219 #22232) +#22705 := [def-axiom]: #22704 +#28939 := [unit-resolution #22705 #28966]: #23219 +#22952 := (or #23285 #23249) +#22953 := [def-axiom]: #22952 +#28967 := [unit-resolution #22953 #28951]: #23249 +#28980 := (or #23252 #23246) +#27793 := (or #15796 #13697) +#15707 := (<= uf_272 4294967295::int) +#15710 := (iff #13337 #15707) +#15701 := (+ 4294967295::int #13338) +#15704 := (>= #15701 0::int) +#15708 := (iff #15704 #15707) +#15709 := [rewrite]: #15708 +#15705 := (iff #13337 #15704) +#15702 := (= #13339 #15701) +#15703 := [monotonicity #7505]: #15702 +#15706 := [monotonicity #15703]: #15705 +#15711 := [trans #15706 #15709]: #15710 +#14253 := [not-or-elim #14242]: #13342 +#14255 := [and-elim #14253]: #13337 +#15712 := [mp #14255 #15711]: #15707 +#27784 := [hypothesis]: #13698 +#27785 := [hypothesis]: #18466 +#27792 := [th-lemma #27785 #27784 #15712]: false +#27191 := [lemma #27792]: #27793 +#28941 := [unit-resolution #27191 #28380]: #15796 +#28976 := (or #13552 #13430) +#28977 := [th-lemma]: #28976 +#28978 := [unit-resolution #28977 #28968]: #13430 +#23007 := (or #23252 #18463 #18466 #23246) +#23004 := [def-axiom]: #23007 +#28982 := [unit-resolution #23004 #28978 #28941]: #28980 +#28983 := [unit-resolution #28982 #28967]: #23246 +#23010 := (or #23243 #23237) +#23014 := [def-axiom]: #23010 +#28984 := [unit-resolution #23014 #28983]: #23237 +#22720 := (>= #13515 -1::int) +#22734 := (or #23243 #13514) +#22735 := [def-axiom]: #22734 +#28979 := [unit-resolution #22735 #28983]: #13514 +#28985 := (or #13518 #22720) +#28981 := [th-lemma]: #28985 +#28986 := [unit-resolution #28981 #28979]: #22720 +#28682 := (not #22720) +#28683 := (or #13446 #28682) +#28696 := [hypothesis]: #22720 +#28697 := [hypothesis]: #13451 +#28596 := [th-lemma #28697 #28380 #28696]: false +#28684 := [lemma #28596]: #28683 +#28958 := [unit-resolution #28684 #28986]: #13446 +#23029 := (or #23240 #13451 #23234) +#23015 := [def-axiom]: #23029 +#28969 := [unit-resolution #23015 #28958 #28984]: #23234 +#23024 := (or #23231 #23225) +#22662 := [def-axiom]: #23024 +#28957 := [unit-resolution #22662 #28969]: #23225 +#23021 := (or #23228 #23222 #22285) +#23022 := [def-axiom]: #23021 +#29004 := [unit-resolution #23022 #28957 #28939]: #22285 +#23032 := (or #22280 #22696) +#22692 := [def-axiom]: #23032 +#28972 := [unit-resolution #22692 #29004]: #22696 +#28754 := (+ uf_294 #18797) +#28950 := (>= #28754 0::int) +#28970 := (not #28950) +#29109 := (= uf_294 ?x786!14) +#29144 := (not #29109) +#22950 := (or #23285 #13604) +#22951 := [def-axiom]: #22950 +#28973 := [unit-resolution #22951 #28951]: #13604 +#22958 := (+ uf_292 #13466) +#22955 := (<= #22958 0::int) +#28971 := (or #12662 #22955) +#28974 := [th-lemma]: #28971 +#28975 := [unit-resolution #28974 #28946]: #22955 +#23033 := (not #18812) +#22676 := (or #22280 #23033) +#23017 := [def-axiom]: #22676 +#29000 := [unit-resolution #23017 #29004]: #23033 +#29131 := (not #22955) +#29132 := (or #29130 #18812 #29131 #13603) +#29125 := [hypothesis]: #13604 +#29126 := [hypothesis]: #22955 +#29127 := [hypothesis]: #23033 +#29128 := [hypothesis]: #29120 +#29129 := [th-lemma #29128 #29127 #29126 #29125]: false +#29133 := [lemma #29129]: #29132 +#29001 := [unit-resolution #29133 #29000 #28975 #28973]: #29130 +#29121 := (or #29119 #29120) +#29122 := [th-lemma]: #29121 +#29002 := [unit-resolution #29122 #29001]: #29119 +#29145 := (or #29144 #29114) +#29140 := (= #18489 #3189) +#29138 := (= #18488 #3180) +#29136 := (= ?x786!14 uf_294) +#29135 := [hypothesis]: #29109 +#29137 := [symm #29135]: #29136 +#29139 := [monotonicity #29137]: #29138 +#29141 := [monotonicity #29139]: #29140 +#29142 := [symm #29141]: #29114 +#29134 := [hypothesis]: #29119 +#29143 := [unit-resolution #29134 #29142]: false +#29146 := [lemma #29143]: #29145 +#29003 := [unit-resolution #29146 #29002]: #29144 +#29019 := (or #29109 #28970) +#28748 := (<= #28754 0::int) +#28780 := (+ uf_292 #18810) +#28781 := (>= #28780 0::int) +#28989 := (not #28781) +#29005 := (or #28989 #18812 #29131) +#29007 := [th-lemma]: #29005 +#29008 := [unit-resolution #29007 #28975 #29000]: #28989 +#28988 := (or #28748 #28781) +#23031 := (or #22280 #18485) +#22691 := [def-axiom]: #23031 +#29006 := [unit-resolution #22691 #29004]: #18485 +#23035 := (or #22280 #18484) +#23041 := [def-axiom]: #23035 +#29009 := [unit-resolution #23041 #29004]: #18484 +#28794 := (or #23208 #22264 #22265 #28748 #28781) +#28750 := (+ #18489 #13902) +#28751 := (<= #28750 0::int) +#28752 := (+ ?x786!14 #13433) +#28761 := (>= #28752 0::int) +#28762 := (or #22265 #28761 #28751 #22264) +#28795 := (or #23208 #28762) +#28802 := (iff #28795 #28794) +#28789 := (or #22264 #22265 #28748 #28781) +#28797 := (or #23208 #28789) +#28800 := (iff #28797 #28794) +#28801 := [rewrite]: #28800 +#28798 := (iff #28795 #28797) +#28792 := (iff #28762 #28789) +#28786 := (or #22265 #28748 #28781 #22264) +#28790 := (iff #28786 #28789) +#28791 := [rewrite]: #28790 +#28787 := (iff #28762 #28786) +#28784 := (iff #28751 #28781) +#28759 := (+ #13902 #18489) +#28770 := (<= #28759 0::int) +#28782 := (iff #28770 #28781) +#28783 := [rewrite]: #28782 +#28778 := (iff #28751 #28770) +#28760 := (= #28750 #28759) +#28769 := [rewrite]: #28760 +#28779 := [monotonicity #28769]: #28778 +#28785 := [trans #28779 #28783]: #28784 +#28757 := (iff #28761 #28748) +#28763 := (+ #13433 ?x786!14) +#28766 := (>= #28763 0::int) +#28755 := (iff #28766 #28748) +#28756 := [rewrite]: #28755 +#28747 := (iff #28761 #28766) +#28764 := (= #28752 #28763) +#28765 := [rewrite]: #28764 +#28753 := [monotonicity #28765]: #28747 +#28758 := [trans #28753 #28756]: #28757 +#28788 := [monotonicity #28758 #28785]: #28787 +#28793 := [trans #28788 #28791]: #28792 +#28799 := [monotonicity #28793]: #28798 +#28803 := [trans #28799 #28801]: #28802 +#28796 := [quant-inst]: #28795 +#28804 := [mp #28796 #28803]: #28794 +#28990 := [unit-resolution #28804 #28346 #29009 #29006]: #28988 +#28991 := [unit-resolution #28990 #29008]: #28748 +#28992 := (not #28748) +#28995 := (or #29109 #28992 #28970) +#28999 := [th-lemma]: #28995 +#28998 := [unit-resolution #28999 #28991]: #29019 +#29020 := [unit-resolution #28998 #29003]: #28970 +#29021 := [th-lemma #29020 #28986 #28972]: false +#29022 := [lemma #29021]: #23285 +#28539 := (or #23261 #23288 #18425) +#28527 := [hypothesis]: #23285 +#28528 := [hypothesis]: #12345 +#28102 := [hypothesis]: #23306 +#22938 := (or #23303 #12345) +#22928 := [def-axiom]: #22938 +#28103 := [unit-resolution #22928 #28102]: #12345 +#22911 := (or #23309 #23303) +#22907 := [def-axiom]: #22911 +#28104 := [unit-resolution #22907 #28102]: #23309 +#22898 := (or #23321 #23315) +#22899 := [def-axiom]: #22898 +#28105 := [unit-resolution #22899 #28379]: #23315 +#22889 := (or #23318 #18425 #18428 #23312) +#22890 := [def-axiom]: #22889 +#28106 := [unit-resolution #22890 #28105]: #23315 +#28107 := [unit-resolution #28106 #28104 #28103 #28393]: false +#28088 := [lemma #28107]: #23303 +#28531 := [hypothesis]: #23264 +#22824 := (or #23261 #12354) +#22825 := [def-axiom]: #22824 +#28532 := [unit-resolution #22825 #28531]: #12354 +#22921 := (or #23306 #18425 #18434 #23300) +#22922 := [def-axiom]: #22921 +#28533 := [unit-resolution #22922 #28532 #28088 #28528]: #23300 +#22934 := (or #23297 #23291) +#22935 := [def-axiom]: #22934 +#28534 := [unit-resolution #22935 #28533]: #23291 +#22867 := (or #23267 #23261) +#22826 := [def-axiom]: #22867 +#28535 := [unit-resolution #22826 #28531]: #23267 +#22878 := (or #23276 #18425 #18428 #23270) +#22871 := [def-axiom]: #22878 +#28536 := [unit-resolution #22871 #28535 #28393 #28528]: #23276 +#22973 := (or #23279 #23273) +#22977 := [def-axiom]: #22973 +#28537 := [unit-resolution #22977 #28536]: #23279 +#22945 := (or #23294 #23282 #23288) +#22942 := [def-axiom]: #22945 +#28538 := [unit-resolution #22942 #28537 #28534 #28527]: false +#28540 := [lemma #28538]: #28539 +#29035 := [unit-resolution #28540 #29022 #28396]: #23261 +#22981 := (or #23264 #18425 #18434 #23258) +#22982 := [def-axiom]: #22981 +#28715 := [unit-resolution #22982 #28396 #29035]: #29036 +#28908 := [unit-resolution #28715 #29033]: #23258 +#22843 := (or #23255 #12365) +#22844 := [def-axiom]: #22843 +#28948 := [unit-resolution #22844 #28908]: #12365 +#22996 := (or #23255 #12360) +#23002 := [def-axiom]: #22996 +#28816 := [unit-resolution #23002 #28908]: #12360 +#29012 := [trans #28816 #28948]: #28823 +#29010 := (not #28823) +#29013 := (or #29010 #28825) +#29032 := [th-lemma]: #29013 +#28994 := [unit-resolution #29032 #29012]: #28825 +#22985 := (or #23255 #23249) +#22991 := [def-axiom]: #22985 +#29023 := [unit-resolution #22991 #28908]: #23249 +#29062 := [unit-resolution #28982 #29023]: #23246 +#29059 := [unit-resolution #23014 #29062]: #23237 +#29060 := [unit-resolution #22735 #29062]: #13514 +#29063 := [unit-resolution #28981 #29060]: #22720 +#29075 := [unit-resolution #28684 #29063]: #13446 +#29077 := [unit-resolution #23015 #29075 #29059]: #23234 +#29049 := [unit-resolution #22662 #29077]: #23225 +#29069 := (= #3189 #3222) +#29067 := (= #3222 #3189) +#29051 := (= #3221 #3180) +#22845 := (or #23255 #12368) +#22987 := [def-axiom]: #22845 +#29052 := [unit-resolution #22987 #28908]: #12368 +#29061 := [symm #29052]: #3201 +#29066 := [monotonicity #29061]: #29051 +#28987 := [monotonicity #29066]: #29067 +#29070 := [symm #28987]: #29069 +#29073 := (= uf_303 #3189) +#29037 := [symm #28816]: #3192 +#29068 := [symm #28948]: #3199 +#29074 := [trans #29068 #29037]: #29073 +#29050 := [trans #29074 #29070]: #12404 +#28659 := (+ uf_294 #13488) +#28680 := (>= #28659 0::int) +#29079 := (or #12515 #28680) +#29076 := [th-lemma]: #29079 +#29099 := [unit-resolution #29076 #29052]: #28680 +#29105 := (not #28680) +#29106 := (or #13491 #29105 #13697) +#29100 := [th-lemma]: #29106 +#29101 := [unit-resolution #29100 #28380 #29099]: #13491 +#29102 := [unit-resolution #23030 #29101 #29050]: #22233 +#29103 := [unit-resolution #22705 #29102]: #23219 +#29078 := [unit-resolution #23022 #29103 #29049]: #22285 +#29080 := [unit-resolution #23017 #29078]: #23033 +#29072 := (not #28825) +#29081 := (or #29130 #18812 #29072) +#29082 := [th-lemma]: #29081 +#29083 := [unit-resolution #29082 #29080 #28994]: #29130 +#29086 := [unit-resolution #29122 #29083]: #29119 +#29087 := [unit-resolution #22692 #29078]: #22696 +#29104 := (or #28950 #18799 #28682) +#29112 := [th-lemma]: #29104 +#29071 := [unit-resolution #29112 #29087 #29063]: #28950 +#29108 := (or #18434 #23300) +#29107 := [unit-resolution #22922 #28088]: #23303 +#29088 := [unit-resolution #29107 #28396]: #29108 +#29111 := [unit-resolution #29088 #29033]: #23300 +#29084 := [unit-resolution #22935 #29111]: #23291 +#29089 := (or #23294 #23282) +#29090 := [unit-resolution #22942 #29022]: #29089 +#29085 := [unit-resolution #29090 #29084]: #23282 +#22883 := (or #23279 #13603) +#22976 := [def-axiom]: #22883 +#29091 := [unit-resolution #22976 #29085]: #13603 +#29110 := (or #22955 #29072 #13604) +#29113 := [th-lemma]: #29110 +#29056 := [unit-resolution #29113 #29091 #28994]: #22955 +#29030 := [unit-resolution #29007 #29080 #29056]: #28989 +#29115 := [unit-resolution #22691 #29078]: #18485 +#29092 := [unit-resolution #23041 #29078]: #18484 +#29147 := [unit-resolution #28804 #28346 #29092 #29115 #29030]: #28748 +#29148 := [unit-resolution #28999 #29147 #29071]: #29109 +[unit-resolution #29146 #29148 #29086]: false unsat diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Boogie/Tools/boogie_commands.ML --- a/src/HOL/Boogie/Tools/boogie_commands.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Boogie/Tools/boogie_commands.ML Sat Nov 14 09:40:27 2009 +0100 @@ -38,7 +38,7 @@ | SOME vc => let fun store thm = Boogie_VCs.discharge (name, thm) - fun after_qed [[thm]] = LocalTheory.theory (store thm) + fun after_qed [[thm]] = Local_Theory.theory (store thm) | after_qed _ = I in lthy diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Boogie/Tools/boogie_loader.ML --- a/src/HOL/Boogie/Tools/boogie_loader.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Boogie/Tools/boogie_loader.ML Sat Nov 14 09:40:27 2009 +0100 @@ -58,7 +58,6 @@ SOME type_name => ((type_name, false), thy) | NONE => let - val bname = Binding.name isa_name val args = Name.variant_list [] (replicate arity "'") val (T, thy') = ObjectLogic.typedecl (Binding.name isa_name, args, NoSyn) thy @@ -324,7 +323,7 @@ local fun mk_nary _ t [] = t - | mk_nary f _ (t :: ts) = fold f ts t + | mk_nary f _ ts = uncurry (fold_rev f) (split_last ts) fun mk_distinct T ts = Const (@{const_name distinct}, HOLogic.listT T --> @{typ bool}) $ @@ -376,7 +375,6 @@ 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} - val m = 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 @@ -451,8 +449,8 @@ binexp ">" (binop @{term "op < :: int => _"} o swap) || binexp ">=" (binop @{term "op <= :: int => _"} o swap) || binexp "+" (binop @{term "op + :: int => _"}) || - binexp "-" (binop @{term "op + :: int => _"}) || - binexp "*" (binop @{term "op + :: int => _"}) || + binexp "-" (binop @{term "op - :: int => _"}) || + binexp "*" (binop @{term "op * :: int => _"}) || binexp "/" (binop @{term boogie_div}) || binexp "%" (binop @{term boogie_mod}) || scan_line "select" num :|-- (fn arity => diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Boogie/Tools/boogie_split.ML --- a/src/HOL/Boogie/Tools/boogie_split.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Boogie/Tools/boogie_split.ML Sat Nov 14 09:40:27 2009 +0100 @@ -114,12 +114,15 @@ local val split_rules = [@{thm impI}, @{thm conjI}, @{thm TrueI}] - fun prep_tac ctxt args facts = + 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 (REPEAT_ALL_NEW (Tactic.resolve_tac split_rules)) THEN unique_case_names (ProofContext.theory_of ctxt) - THEN ALLGOALS (SUBGOAL (fn (t, i) => - TRY (sub_tactics_tac ctxt args (case_name_of t) i))) + 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})))) @@ -136,12 +139,14 @@ (ProofContext.theory_of ctxt, Thm.prop_of st) names) Tactical.all_tac st)) - val verbose_opt = Args.parens (Args.$$$ "verbose") >> K true + 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 verbose_opt false -- Scan.optional sub_tactics []) - >> 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 diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Deriv.thy --- a/src/HOL/Deriv.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Deriv.thy Sat Nov 14 09:40:27 2009 +0100 @@ -273,6 +273,11 @@ "DERIV f x :> D ==> DERIV (%x. c * f x) x :> c*D" by (drule DERIV_mult' [OF DERIV_const], simp) +lemma DERIV_cdivide: "DERIV f x :> D ==> DERIV (%x. f x / c) x :> D / c" + apply (subgoal_tac "DERIV (%x. (1 / c) * f x) x :> (1 / c) * D", force) + apply (erule DERIV_cmult) + done + text {* Standard version *} lemma DERIV_chain: "[| DERIV f (g x) :> Da; DERIV g x :> Db |] ==> DERIV (f o g) x :> Da * Db" by (drule (1) DERIV_chain', simp add: o_def real_scaleR_def mult_commute) @@ -702,14 +707,10 @@ apply safe apply simp_all apply (rename_tac x xa ya M Ma) -apply (cut_tac x = M and y = Ma in linorder_linear, safe) -apply (rule_tac x = Ma in exI, clarify) -apply (cut_tac x = xb and y = xa in linorder_linear, force) -apply (rule_tac x = M in exI, clarify) -apply (cut_tac x = xb and y = xa in linorder_linear, force) +apply (metis linorder_not_less order_le_less real_le_trans) apply (case_tac "a \ x & x \ b") -apply (rule_tac [2] x = 1 in exI) -prefer 2 apply force + prefer 2 + apply (rule_tac x = 1 in exI, force) apply (simp add: LIM_eq isCont_iff) apply (drule_tac x = x in spec, auto) apply (erule_tac V = "\M. \x. a \ x & x \ b & ~ f x \ M" in thin_rl) @@ -815,7 +816,7 @@ text{*If @{term "0 < f'(x)"} then @{term x} is Locally Strictly Increasing At The Right*} -lemma DERIV_left_inc: +lemma DERIV_pos_inc_right: fixes f :: "real => real" assumes der: "DERIV f x :> l" and l: "0 < l" @@ -845,7 +846,7 @@ qed qed -lemma DERIV_left_dec: +lemma DERIV_neg_dec_left: fixes f :: "real => real" assumes der: "DERIV f x :> l" and l: "l < 0" @@ -875,6 +876,21 @@ qed qed + +lemma DERIV_pos_inc_left: + fixes f :: "real => real" + shows "DERIV f x :> l \ 0 < l \ \d > 0. \h > 0. h < d --> f(x - h) < f(x)" + apply (rule DERIV_neg_dec_left [of "%x. - f x" x "-l", simplified]) + apply (auto simp add: DERIV_minus) + done + +lemma DERIV_neg_dec_right: + fixes f :: "real => real" + shows "DERIV f x :> l \ l < 0 \ \d > 0. \h > 0. h < d --> f(x) > f(x + h)" + apply (rule DERIV_pos_inc_right [of "%x. - f x" x "-l", simplified]) + apply (auto simp add: DERIV_minus) + done + lemma DERIV_local_max: fixes f :: "real => real" assumes der: "DERIV f x :> l" @@ -885,7 +901,7 @@ case equal thus ?thesis . next case less - from DERIV_left_dec [OF der less] + from DERIV_neg_dec_left [OF der less] obtain d' where d': "0 < d'" and lt: "\h > 0. h < d' \ f x < f (x-h)" by blast from real_lbound_gt_zero [OF d d'] @@ -894,7 +910,7 @@ show ?thesis by (auto simp add: abs_if) next case greater - from DERIV_left_inc [OF der greater] + from DERIV_pos_inc_right [OF der greater] obtain d' where d': "0 < d'" and lt: "\h > 0. h < d' \ f x < f (x + h)" by blast from real_lbound_gt_zero [OF d d'] @@ -1205,6 +1221,87 @@ ultimately show ?thesis using neq by (force simp add: add_commute) qed +(* A function with positive derivative is increasing. + A simple proof using the MVT, by Jeremy Avigad. And variants. +*) + +lemma DERIV_pos_imp_increasing: + fixes a::real and b::real and f::"real => real" + assumes "a < b" and "\x. a \ x & x \ b --> (EX y. DERIV f x :> y & y > 0)" + shows "f a < f b" +proof (rule ccontr) + assume "~ f a < f b" + from assms have "EX l z. a < z & z < b & DERIV f z :> l + & f b - f a = (b - a) * l" + by (metis MVT DERIV_isCont differentiableI real_less_def) + then obtain l z where "a < z" and "z < b" and "DERIV f z :> l" + and "f b - f a = (b - a) * l" + by auto + + from prems have "~(l > 0)" + by (metis assms(1) linorder_not_le mult_le_0_iff real_le_eq_diff) + with prems show False + by (metis DERIV_unique real_less_def) +qed + + +lemma DERIV_nonneg_imp_nonincreasing: + fixes a::real and b::real and f::"real => real" + assumes "a \ b" and + "\x. a \ x & x \ b --> (\y. DERIV f x :> y & y \ 0)" + shows "f a \ f b" +proof (rule ccontr, cases "a = b") + assume "~ f a \ f b" + assume "a = b" + with prems show False by auto + next assume "~ f a \ f b" + assume "a ~= b" + with assms have "EX l z. a < z & z < b & DERIV f z :> l + & f b - f a = (b - a) * l" + apply (intro MVT) + apply auto + apply (metis DERIV_isCont) + apply (metis differentiableI real_less_def) + done + then obtain l z where "a < z" and "z < b" and "DERIV f z :> l" + and "f b - f a = (b - a) * l" + by auto + from prems have "~(l >= 0)" + by (metis diff_self le_eqI le_iff_diff_le_0 real_le_antisym real_le_linear + split_mult_pos_le) + with prems show False + by (metis DERIV_unique order_less_imp_le) +qed + +lemma DERIV_neg_imp_decreasing: + fixes a::real and b::real and f::"real => real" + assumes "a < b" and + "\x. a \ x & x \ b --> (\y. DERIV f x :> y & y < 0)" + shows "f a > f b" +proof - + have "(%x. -f x) a < (%x. -f x) b" + apply (rule DERIV_pos_imp_increasing [of a b "%x. -f x"]) + apply (insert prems, auto) + apply (metis DERIV_minus neg_0_less_iff_less) + done + thus ?thesis + by simp +qed + +lemma DERIV_nonpos_imp_nonincreasing: + fixes a::real and b::real and f::"real => real" + assumes "a \ b" and + "\x. a \ x & x \ b --> (\y. DERIV f x :> y & y \ 0)" + shows "f a \ f b" +proof - + have "(%x. -f x) a \ (%x. -f x) b" + apply (rule DERIV_nonneg_imp_nonincreasing [of a b "%x. -f x"]) + apply (insert prems, auto) + apply (metis DERIV_minus neg_0_le_iff_le) + done + thus ?thesis + by simp +qed subsection {* Continuous injective functions *} diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Finite_Set.thy --- a/src/HOL/Finite_Set.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Finite_Set.thy Sat Nov 14 09:40:27 2009 +0100 @@ -2344,7 +2344,7 @@ lemma card_bij_eq: "[|inj_on f A; f ` A \ B; inj_on g B; g ` B \ A; finite A; finite B |] ==> card A = card B" -by (auto intro: le_anti_sym card_inj_on_le) +by (auto intro: le_antisym card_inj_on_le) subsubsection {* Cardinality of products *} diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/GCD.thy --- a/src/HOL/GCD.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/GCD.thy Sat Nov 14 09:40:27 2009 +0100 @@ -312,13 +312,13 @@ by (insert gcd_zero_int [of m n], insert gcd_ge_0_int [of m n], arith) lemma gcd_commute_nat: "gcd (m::nat) n = gcd n m" - by (rule dvd_anti_sym, auto) + by (rule dvd_antisym, auto) lemma gcd_commute_int: "gcd (m::int) n = gcd n m" by (auto simp add: gcd_int_def gcd_commute_nat) lemma gcd_assoc_nat: "gcd (gcd (k::nat) m) n = gcd k (gcd m n)" - apply (rule dvd_anti_sym) + apply (rule dvd_antisym) apply (blast intro: dvd_trans)+ done @@ -339,23 +339,18 @@ lemma gcd_unique_nat: "(d::nat) dvd a \ d dvd b \ (\e. e dvd a \ e dvd b \ e dvd d) \ d = gcd a b" apply auto - apply (rule dvd_anti_sym) + apply (rule dvd_antisym) apply (erule (1) gcd_greatest_nat) apply auto done lemma gcd_unique_int: "d >= 0 & (d::int) dvd a \ d dvd b \ (\e. e dvd a \ e dvd b \ e dvd d) \ d = gcd a b" - apply (case_tac "d = 0") - apply force - apply (rule iffI) - apply (rule zdvd_anti_sym) - apply arith - apply (subst gcd_pos_int) - apply clarsimp - apply (drule_tac x = "d + 1" in spec) - apply (frule zdvd_imp_le) - apply (auto intro: gcd_greatest_int) +apply (case_tac "d = 0") + apply simp +apply (rule iffI) + apply (rule zdvd_antisym_nonneg) + apply (auto intro: gcd_greatest_int) done lemma gcd_proj1_if_dvd_nat [simp]: "(x::nat) dvd y \ gcd x y = x" @@ -417,7 +412,7 @@ by (auto intro: coprime_dvd_mult_int) lemma gcd_mult_cancel_nat: "coprime k n \ gcd ((k::nat) * m) n = gcd m n" - apply (rule dvd_anti_sym) + apply (rule dvd_antisym) apply (rule gcd_greatest_nat) apply (rule_tac n = k in coprime_dvd_mult_nat) apply (simp add: gcd_assoc_nat) @@ -1381,11 +1376,11 @@ lemma lcm_unique_nat: "(a::nat) dvd d \ b dvd d \ (\e. a dvd e \ b dvd e \ d dvd e) \ d = lcm a b" - by (auto intro: dvd_anti_sym lcm_least_nat lcm_dvd1_nat lcm_dvd2_nat) + by (auto intro: dvd_antisym lcm_least_nat lcm_dvd1_nat lcm_dvd2_nat) lemma lcm_unique_int: "d >= 0 \ (a::int) dvd d \ b dvd d \ (\e. a dvd e \ b dvd e \ d dvd e) \ d = lcm a b" - by (auto intro: dvd_anti_sym [transferred] lcm_least_int) + by (auto intro: dvd_antisym [transferred] lcm_least_int) lemma lcm_proj2_if_dvd_nat [simp]: "(x::nat) dvd y \ lcm x y = y" apply (rule sym) diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Hoare/Arith2.thy --- a/src/HOL/Hoare/Arith2.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Hoare/Arith2.thy Sat Nov 14 09:40:27 2009 +0100 @@ -58,7 +58,7 @@ apply (frule cd_nnn) apply (rule some_equality [symmetric]) apply (blast dest: cd_le) - apply (blast intro: le_anti_sym dest: cd_le) + apply (blast intro: le_antisym dest: cd_le) done lemma gcd_swap: "gcd m n = gcd n m" diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Import/HOL/arithmetic.imp --- a/src/HOL/Import/HOL/arithmetic.imp Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Import/HOL/arithmetic.imp Sat Nov 14 09:40:27 2009 +0100 @@ -191,7 +191,7 @@ "LESS_EQ_ADD_SUB" > "Nat.add_diff_assoc" "LESS_EQ_ADD" > "Nat.le_add1" "LESS_EQ_0" > "Nat.le_0_eq" - "LESS_EQUAL_ANTISYM" > "Nat.le_anti_sym" + "LESS_EQUAL_ANTISYM" > "Nat.le_antisym" "LESS_EQUAL_ADD" > "HOL4Base.arithmetic.LESS_EQUAL_ADD" "LESS_EQ" > "Nat.le_simps_3" "LESS_DIV_EQ_ZERO" > "Divides.div_less" diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Import/HOL/divides.imp --- a/src/HOL/Import/HOL/divides.imp Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Import/HOL/divides.imp Sat Nov 14 09:40:27 2009 +0100 @@ -16,7 +16,7 @@ "DIVIDES_MULT" > "Divides.dvd_mult2" "DIVIDES_LE" > "Divides.dvd_imp_le" "DIVIDES_FACT" > "HOL4Base.divides.DIVIDES_FACT" - "DIVIDES_ANTISYM" > "Divides.dvd_anti_sym" + "DIVIDES_ANTISYM" > "Divides.dvd_antisym" "DIVIDES_ADD_2" > "HOL4Base.divides.DIVIDES_ADD_2" "DIVIDES_ADD_1" > "Ring_and_Field.dvd_add" "ALL_DIVIDES_0" > "Divides.dvd_0_right" diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Int.thy --- a/src/HOL/Int.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Int.thy Sat Nov 14 09:40:27 2009 +0100 @@ -1986,15 +1986,18 @@ subsection {* The divides relation *} -lemma zdvd_anti_sym: - "0 < m ==> 0 < n ==> m dvd n ==> n dvd m ==> m = (n::int)" +lemma zdvd_antisym_nonneg: + "0 <= m ==> 0 <= n ==> m dvd n ==> n dvd m ==> m = (n::int)" apply (simp add: dvd_def, auto) - apply (simp add: mult_assoc zero_less_mult_iff zmult_eq_1_iff) + apply (auto simp add: mult_assoc zero_le_mult_iff zmult_eq_1_iff) done -lemma zdvd_dvd_eq: assumes "a \ 0" and "(a::int) dvd b" and "b dvd a" +lemma zdvd_antisym_abs: assumes "(a::int) dvd b" and "b dvd a" shows "\a\ = \b\" -proof- +proof cases + assume "a = 0" with assms show ?thesis by simp +next + assume "a \ 0" from `a dvd b` obtain k where k:"b = a*k" unfolding dvd_def by blast from `b dvd a` obtain k' where k':"a = b*k'" unfolding dvd_def by blast from k k' have "a = a*k*k'" by simp @@ -2326,7 +2329,7 @@ lemmas zle_refl = order_refl [of "w::int", standard] lemmas zle_trans = order_trans [where 'a=int and x="i" and y="j" and z="k", standard] -lemmas zle_anti_sym = order_antisym [of "z::int" "w", standard] +lemmas zle_antisym = order_antisym [of "z::int" "w", standard] lemmas zle_linear = linorder_linear [of "z::int" "w", standard] lemmas zless_linear = linorder_less_linear [where 'a = int] diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Library/Abstract_Rat.thy --- a/src/HOL/Library/Abstract_Rat.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Library/Abstract_Rat.thy Sat Nov 14 09:40:27 2009 +0100 @@ -206,7 +206,7 @@ apply simp apply algebra done - from zdvd_dvd_eq[OF bz coprime_dvd_mult_int[OF gcd1(2) raw_dvd(2)] + from zdvd_antisym_abs[OF coprime_dvd_mult_int[OF gcd1(2) raw_dvd(2)] coprime_dvd_mult_int[OF gcd1(4) raw_dvd(4)]] have eq1: "b = b'" using pos by arith with eq have "a = a'" using pos by simp diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Ln.thy --- a/src/HOL/Ln.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Ln.thy Sat Nov 14 09:40:27 2009 +0100 @@ -342,9 +342,6 @@ apply auto done -lemma DERIV_ln: "0 < x ==> DERIV ln x :> 1 / x" - by (rule DERIV_ln[THEN DERIV_cong], simp, simp add: divide_inverse) - lemma ln_x_over_x_mono: "exp 1 <= x ==> x <= y ==> (ln y / y) <= (ln x / x)" proof - assume "exp 1 <= x" and "x <= y" diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Matrix/Matrix.thy --- a/src/HOL/Matrix/Matrix.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Matrix/Matrix.thy Sat Nov 14 09:40:27 2009 +0100 @@ -873,7 +873,7 @@ have th: "\ (\m. m \ j)" "\n. \ n \ i" by arith+ from th show ?thesis apply (auto) -apply (rule le_anti_sym) +apply (rule le_antisym) apply (subst nrows_le) apply (simp add: singleton_matrix_def, auto) apply (subst RepAbs_matrix) @@ -889,7 +889,7 @@ have th: "\ (\m. m \ j)" "\n. \ n \ i" by arith+ from th show ?thesis apply (auto) -apply (rule le_anti_sym) +apply (rule le_antisym) apply (subst ncols_le) apply (simp add: singleton_matrix_def, auto) apply (subst RepAbs_matrix) @@ -1004,7 +1004,7 @@ apply (subst foldseq_almostzero[of _ j]) apply (simp add: prems)+ apply (auto) - apply (metis comm_monoid_add.mult_1 le_anti_sym le_diff_eq not_neg_nat zero_le_imp_of_nat zle_int) + apply (metis comm_monoid_add.mult_1 le_antisym le_diff_eq not_neg_nat zero_le_imp_of_nat zle_int) done lemma mult_matrix_ext: diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Nat.thy --- a/src/HOL/Nat.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Nat.thy Sat Nov 14 09:40:27 2009 +0100 @@ -596,7 +596,7 @@ lemma le_trans: "[| i \ j; j \ k |] ==> i \ (k::nat)" by (rule order_trans) -lemma le_anti_sym: "[| m \ n; n \ m |] ==> m = (n::nat)" +lemma le_antisym: "[| m \ n; n \ m |] ==> m = (n::nat)" by (rule antisym) lemma nat_less_le: "((m::nat) < n) = (m \ n & m \ n)" @@ -1611,14 +1611,14 @@ lemma nat_dvd_1_iff_1 [simp]: "m dvd (1::nat) \ m = 1" by (simp add: dvd_def) -lemma dvd_anti_sym: "[| m dvd n; n dvd m |] ==> m = (n::nat)" +lemma dvd_antisym: "[| m dvd n; n dvd m |] ==> m = (n::nat)" unfolding dvd_def by (force dest: mult_eq_self_implies_10 simp add: mult_assoc mult_eq_1_iff) text {* @{term "op dvd"} is a partial order *} interpretation dvd: order "op dvd" "\n m \ nat. n dvd m \ \ m dvd n" - proof qed (auto intro: dvd_refl dvd_trans dvd_anti_sym) + proof qed (auto intro: dvd_refl dvd_trans dvd_antisym) lemma dvd_diff_nat[simp]: "[| k dvd m; k dvd n |] ==> k dvd (m-n :: nat)" unfolding dvd_def diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Nominal/nominal_datatype.ML --- a/src/HOL/Nominal/nominal_datatype.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Nominal/nominal_datatype.ML Sat Nov 14 09:40:27 2009 +0100 @@ -569,9 +569,8 @@ thy3 |> Sign.map_naming Name_Space.conceal |> Inductive.add_inductive_global (serial ()) - {quiet_mode = false, verbose = false, kind = "", - alt_name = Binding.name big_rep_name, coind = false, no_elim = true, no_ind = false, - skip_mono = true, fork_mono = false} + {quiet_mode = false, verbose = false, alt_name = Binding.name big_rep_name, + coind = false, no_elim = true, no_ind = false, skip_mono = true, fork_mono = false} (map (fn (s, T) => ((Binding.name s, T --> HOLogic.boolT), NoSyn)) (rep_set_names' ~~ recTs')) [] (map (fn x => (Attrib.empty_binding, x)) intr_ts) [] @@ -1513,9 +1512,8 @@ thy10 |> Sign.map_naming Name_Space.conceal |> Inductive.add_inductive_global (serial ()) - {quiet_mode = #quiet config, verbose = false, kind = "", - alt_name = Binding.name big_rec_name, coind = false, no_elim = false, no_ind = false, - skip_mono = true, fork_mono = false} + {quiet_mode = #quiet config, verbose = false, alt_name = Binding.name big_rec_name, + coind = false, no_elim = false, no_ind = false, skip_mono = true, fork_mono = false} (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts)) (map dest_Free rec_fns) (map (fn x => (Attrib.empty_binding, x)) rec_intr_ts) [] diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Nominal/nominal_inductive.ML --- a/src/HOL/Nominal/nominal_inductive.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Nominal/nominal_inductive.ML Sat Nov 14 09:40:27 2009 +0100 @@ -561,20 +561,19 @@ (strong_raw_induct, [ind_case_names, Rule_Cases.consumes 0]) else (strong_raw_induct RSN (2, rev_mp), [ind_case_names, Rule_Cases.consumes 1]); - val ((_, [strong_induct']), ctxt') = LocalTheory.note Thm.generatedK + val ((_, [strong_induct']), ctxt') = ctxt |> Local_Theory.note ((rec_qualified (Binding.name "strong_induct"), - map (Attrib.internal o K) (#2 strong_induct)), [#1 strong_induct]) - ctxt; + map (Attrib.internal o K) (#2 strong_induct)), [#1 strong_induct]); val strong_inducts = - Project_Rule.projects ctxt (1 upto length names) strong_induct' + Project_Rule.projects ctxt (1 upto length names) strong_induct'; in ctxt' |> - LocalTheory.note Thm.generatedK + Local_Theory.note ((rec_qualified (Binding.name "strong_inducts"), [Attrib.internal (K ind_case_names), Attrib.internal (K (Rule_Cases.consumes 1))]), strong_inducts) |> snd |> - LocalTheory.notes Thm.generatedK (map (fn ((name, elim), (_, cases)) => + Local_Theory.notes (map (fn ((name, elim), (_, cases)) => ((Binding.qualified_name (Long_Name.qualify (Long_Name.base_name name) "strong_cases"), [Attrib.internal (K (Rule_Cases.case_names (map snd cases))), Attrib.internal (K (Rule_Cases.consumes 1))]), [([elim], [])])) @@ -664,7 +663,7 @@ end) atoms in ctxt |> - LocalTheory.notes Thm.generatedK (map (fn (name, ths) => + Local_Theory.notes (map (fn (name, ths) => ((Binding.qualified_name (Long_Name.qualify (Long_Name.base_name name) "eqvt"), [Attrib.internal (K NominalThmDecls.eqvt_add)]), [(ths, [])])) (names ~~ transp thss)) |> snd diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Nominal/nominal_inductive2.ML --- a/src/HOL/Nominal/nominal_inductive2.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Nominal/nominal_inductive2.ML Sat Nov 14 09:40:27 2009 +0100 @@ -466,15 +466,14 @@ NONE => (rec_qualified (Binding.name "strong_induct"), rec_qualified (Binding.name "strong_inducts")) | SOME s => (Binding.name s, Binding.name (s ^ "s")); - val ((_, [strong_induct']), ctxt') = LocalTheory.note Thm.generatedK + val ((_, [strong_induct']), ctxt') = ctxt |> Local_Theory.note ((induct_name, - map (Attrib.internal o K) (#2 strong_induct)), [#1 strong_induct]) - ctxt; + map (Attrib.internal o K) (#2 strong_induct)), [#1 strong_induct]); val strong_inducts = Project_Rule.projects ctxt' (1 upto length names) strong_induct' in ctxt' |> - LocalTheory.note Thm.generatedK + Local_Theory.note ((inducts_name, [Attrib.internal (K ind_case_names), Attrib.internal (K (Rule_Cases.consumes 1))]), diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Nominal/nominal_primrec.ML --- a/src/HOL/Nominal/nominal_primrec.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Nominal/nominal_primrec.ML Sat Nov 14 09:40:27 2009 +0100 @@ -280,9 +280,8 @@ else primrec_err ("functions " ^ commas_quote names2 ^ "\nare not mutually recursive"); val (defs_thms, lthy') = lthy |> - set_group ? LocalTheory.set_group (serial ()) |> - fold_map (apfst (snd o snd) oo - LocalTheory.define Thm.definitionK o fst) defs'; + set_group ? Local_Theory.set_group (serial ()) |> + fold_map (apfst (snd o snd) oo Local_Theory.define Thm.definitionK o fst) defs'; val qualify = Binding.qualify false (space_implode "_" (map (Long_Name.base_name o #1) defs)); val names_atts' = map (apfst qualify) names_atts; @@ -367,11 +366,11 @@ (fn thss => fn goal_ctxt => let val simps = ProofContext.export goal_ctxt lthy' (flat thss); - val (simps', lthy'') = fold_map (LocalTheory.note Thm.generatedK) - (names_atts' ~~ map single simps) lthy' + val (simps', lthy'') = + fold_map Local_Theory.note (names_atts' ~~ map single simps) lthy'; in lthy'' - |> LocalTheory.note Thm.generatedK ((qualify (Binding.name "simps"), + |> Local_Theory.note ((qualify (Binding.name "simps"), map (Attrib.internal o K) [Simplifier.simp_add, Nitpick_Simps.add]), maps snd simps') |> snd diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Number_Theory/UniqueFactorization.thy --- a/src/HOL/Number_Theory/UniqueFactorization.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Number_Theory/UniqueFactorization.thy Sat Nov 14 09:40:27 2009 +0100 @@ -844,7 +844,7 @@ mult_eq [simp]: "!!p. prime p \ multiplicity p x = multiplicity p y" shows "x = y" - apply (rule dvd_anti_sym) + apply (rule dvd_antisym) apply (auto intro: multiplicity_dvd'_nat) done @@ -854,7 +854,7 @@ mult_eq [simp]: "!!p. prime p \ multiplicity p x = multiplicity p y" shows "x = y" - apply (rule dvd_anti_sym [transferred]) + apply (rule dvd_antisym [transferred]) apply (auto intro: multiplicity_dvd'_int) done diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Old_Number_Theory/IntPrimes.thy --- a/src/HOL/Old_Number_Theory/IntPrimes.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Old_Number_Theory/IntPrimes.thy Sat Nov 14 09:40:27 2009 +0100 @@ -204,7 +204,7 @@ apply (case_tac [2] "0 \ ka") apply (metis zdvd_mult_div_cancel dvd_refl dvd_mult_left zmult_commute zrelprime_zdvd_zmult) apply (metis abs_dvd_iff abs_of_nonneg zadd_0 zgcd_0_left zgcd_commute zgcd_zadd_zmult zgcd_zdvd_zgcd_zmult zgcd_zmult_distrib2_abs zmult_1_right zmult_commute) - apply (metis mult_le_0_iff zdvd_mono zdvd_mult_cancel dvd_triv_left zero_le_mult_iff zle_anti_sym zle_linear zle_refl zmult_commute zrelprime_zdvd_zmult) + apply (metis mult_le_0_iff zdvd_mono zdvd_mult_cancel dvd_triv_left zero_le_mult_iff zle_antisym zle_linear zle_refl zmult_commute zrelprime_zdvd_zmult) apply (metis dvd_triv_left) done diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Old_Number_Theory/Legacy_GCD.thy --- a/src/HOL/Old_Number_Theory/Legacy_GCD.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Old_Number_Theory/Legacy_GCD.thy Sat Nov 14 09:40:27 2009 +0100 @@ -23,7 +23,7 @@ text {* Uniqueness *} lemma is_gcd_unique: "is_gcd a b m \ is_gcd a b n \ m = n" - by (simp add: is_gcd_def) (blast intro: dvd_anti_sym) + by (simp add: is_gcd_def) (blast intro: dvd_antisym) text {* Connection to divides relation *} @@ -156,7 +156,7 @@ by (auto intro: relprime_dvd_mult dvd_mult2) lemma gcd_mult_cancel: "gcd k n = 1 ==> gcd (k * m) n = gcd m n" - apply (rule dvd_anti_sym) + apply (rule dvd_antisym) apply (rule gcd_greatest) apply (rule_tac n = k in relprime_dvd_mult) apply (simp add: gcd_assoc) @@ -223,7 +223,7 @@ assume H: "d dvd a" "d dvd b" "\e. e dvd a \ e dvd b \ e dvd d" from H(3)[rule_format] gcd_dvd1[of a b] gcd_dvd2[of a b] have th: "gcd a b dvd d" by blast - from dvd_anti_sym[OF th gcd_greatest[OF H(1,2)]] show "d = gcd a b" by blast + from dvd_antisym[OF th gcd_greatest[OF H(1,2)]] show "d = gcd a b" by blast qed lemma gcd_eq: assumes H: "\d. d dvd x \ d dvd y \ d dvd u \ d dvd v" diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Old_Number_Theory/Pocklington.thy --- a/src/HOL/Old_Number_Theory/Pocklington.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Old_Number_Theory/Pocklington.thy Sat Nov 14 09:40:27 2009 +0100 @@ -935,7 +935,7 @@ p: "prime p" "p dvd m" by blast from dvd_trans[OF p(2) m(1)] p(1) H have "p = n" by blast with p(2) have "n dvd m" by simp - hence "m=n" using dvd_anti_sym[OF m(1)] by simp } + hence "m=n" using dvd_antisym[OF m(1)] by simp } with n1 have "prime n" unfolding prime_def by auto } ultimately have ?thesis using n by blast} ultimately show ?thesis by auto diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Old_Number_Theory/Primes.thy --- a/src/HOL/Old_Number_Theory/Primes.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Old_Number_Theory/Primes.thy Sat Nov 14 09:40:27 2009 +0100 @@ -97,7 +97,7 @@ text {* Elementary theory of divisibility *} lemma divides_ge: "(a::nat) dvd b \ b = 0 \ a \ b" unfolding dvd_def by auto lemma divides_antisym: "(x::nat) dvd y \ y dvd x \ x = y" - using dvd_anti_sym[of x y] by auto + using dvd_antisym[of x y] by auto lemma divides_add_revr: assumes da: "(d::nat) dvd a" and dab:"d dvd (a + b)" shows "d dvd b" diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Probability/Borel.thy --- a/src/HOL/Probability/Borel.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Probability/Borel.thy Sat Nov 14 09:40:27 2009 +0100 @@ -73,7 +73,7 @@ with w have "real(Suc(natceiling(inverse(g w - f w)))) > inverse(g w - f w)" by (metis lessI order_le_less_trans real_natceiling_ge real_of_nat_less_iff) hence "inverse(real(Suc(natceiling(inverse(g w - f w))))) < inverse(inverse(g w - f w))" - by (metis less_iff_diff_less_0 less_imp_inverse_less linorder_neqE_ordered_idom nz positive_imp_inverse_positive real_le_anti_sym real_less_def w) + by (metis less_iff_diff_less_0 less_imp_inverse_less linorder_neqE_ordered_idom nz positive_imp_inverse_positive real_le_antisym real_less_def w) hence "inverse(real(Suc(natceiling(inverse(g w - f w))))) < g w - f w" by (metis inverse_inverse_eq order_less_le_trans real_le_refl) thus "\n. f w \ g w - inverse(real(Suc n))" using w diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Probability/Measure.thy --- a/src/HOL/Probability/Measure.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Probability/Measure.thy Sat Nov 14 09:40:27 2009 +0100 @@ -356,7 +356,7 @@ by (metis add_commute le_add_diff_inverse nat_less_le) thus ?thesis by (auto simp add: disjoint_family_def) - (metis insert_absorb insert_subset le_SucE le_anti_sym not_leE) + (metis insert_absorb insert_subset le_SucE le_antisym not_leE) qed diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/RealDef.thy --- a/src/HOL/RealDef.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/RealDef.thy Sat Nov 14 09:40:27 2009 +0100 @@ -321,7 +321,7 @@ apply (auto intro: real_le_lemma) done -lemma real_le_anti_sym: "[| z \ w; w \ z |] ==> z = (w::real)" +lemma real_le_antisym: "[| z \ w; w \ z |] ==> z = (w::real)" by (cases z, cases w, simp add: real_le) lemma real_trans_lemma: @@ -347,8 +347,8 @@ proof fix u v :: real show "u < v \ u \ v \ \ v \ u" - by (auto simp add: real_less_def intro: real_le_anti_sym) -qed (assumption | rule real_le_refl real_le_trans real_le_anti_sym)+ + by (auto simp add: real_less_def intro: real_le_antisym) +qed (assumption | rule real_le_refl real_le_trans real_le_antisym)+ (* Axiom 'linorder_linear' of class 'linorder': *) lemma real_le_linear: "(z::real) \ w | w \ z" diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Ring_and_Field.thy --- a/src/HOL/Ring_and_Field.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Ring_and_Field.thy Sat Nov 14 09:40:27 2009 +0100 @@ -1301,6 +1301,10 @@ lemma dvd_abs_iff [simp]: "m dvd (abs k) \ m dvd k" by (simp add: abs_if) +lemma dvd_if_abs_eq: + "abs l = abs (k) \ l dvd k" +by(subst abs_dvd_iff[symmetric]) simp + end class ordered_field = field + ordered_idom diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/SMT/Tools/cvc3_solver.ML --- a/src/HOL/SMT/Tools/cvc3_solver.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/SMT/Tools/cvc3_solver.ML Sat Nov 14 09:40:27 2009 +0100 @@ -23,11 +23,11 @@ fun raise_cex real = raise SMT_Solver.SMT_COUNTEREXAMPLE (real, []) -fun core_oracle ({context, output, recon, ...} : SMT_Solver.proof_data) = +fun core_oracle ({output, ...} : SMT_Solver.proof_data) = let val empty_line = (fn "" => true | _ => false) val split_first = (fn [] => ("", []) | l :: ls => (l, ls)) - val (l, ls) = split_first (dropwhile empty_line output) + val (l, _) = split_first (dropwhile empty_line output) in if is_unsat l then @{cprop False} else if is_sat l then raise_cex true diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/SMT/Tools/smt_normalize.ML --- a/src/HOL/SMT/Tools/smt_normalize.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/SMT/Tools/smt_normalize.ML Sat Nov 14 09:40:27 2009 +0100 @@ -285,14 +285,14 @@ fun add_sym t = if AList.defined (op =) defs t then insert (op =) t else I fun add_syms thms = fold (Term.fold_aterms add_sym o Thm.prop_of) thms [] - fun unfold_conv ctxt ct = + fun unfold_conv ct = (case AList.lookup (op =) defs (Term.head_of (Thm.term_of ct)) of SOME (_, eq) => Conv.rewr_conv eq | NONE => Conv.all_conv) ct in fun add_abs_min_max_rules ctxt thms = if Config.get ctxt unfold_defs - then map (Conv.fconv_rule (More_Conv.bottom_conv unfold_conv ctxt)) thms + then map (Conv.fconv_rule (More_Conv.bottom_conv (K unfold_conv) ctxt)) thms else map fst (map_filter (AList.lookup (op =) defs) (add_syms thms)) @ thms end @@ -319,13 +319,12 @@ let val cvs' = used_vars cvs ct val ct' = fold_rev Thm.cabs cvs' ct - val mk_repl = fold (fn ct => fn cu => Thm.capply cu ct) cvs' in (case Termtab.lookup defs (Thm.term_of ct') of SOME (_, eq) => (make_def cvs' eq, cx) | NONE => let - val {t, T, ...} = Thm.rep_cterm ct' + val {T, ...} = Thm.rep_cterm ct' val (n, nctxt') = fresh_name "" nctxt val eq = Thm.assume (mk_meta_eq (cert ctxt (Free (n, T))) ct') in (make_def cvs' eq, (nctxt', add_def ct' eq defs)) end) @@ -410,8 +409,8 @@ (case Term.strip_comb (Thm.term_of ct) of (Const (n, _), ts) => app_conv tb (const n) (length ts) ctxt | (Free (n, _), ts) => app_conv tb (free n) (length ts) ctxt - | (Abs _, ts) => nary_conv (abs_conv sub_conv tb ctxt) (sub_conv tb ctxt) - | (_, ts) => nary_conv Conv.all_conv (sub_conv tb ctxt)) ct + | (Abs _, _) => nary_conv (abs_conv sub_conv tb ctxt) (sub_conv tb ctxt) + | (_, _) => nary_conv Conv.all_conv (sub_conv tb ctxt)) ct and app_conv tb n i ctxt = (case Symtab.lookup tb n of NONE => nary_conv Conv.all_conv (sub_conv tb ctxt) diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/SMT/Tools/smt_solver.ML --- a/src/HOL/SMT/Tools/smt_solver.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/SMT/Tools/smt_solver.ML Sat Nov 14 09:40:27 2009 +0100 @@ -144,7 +144,7 @@ true) end -fun run_locally f ctxt env_var args ps = +fun run_locally f env_var args ps = if getenv env_var = "" then f ("Undefined Isabelle environment variable: " ^ quote env_var) else @@ -173,7 +173,7 @@ let val ps = [File.shell_path problem_path, ">", File.shell_path proof_path] in if use_certificate ctxt ps then () - else run_locally (run_remote remote_name args ps) ctxt env_var args ps + else run_locally (run_remote remote_name args ps) env_var args ps end in diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/SMT/Tools/smt_translate.ML --- a/src/HOL/SMT/Tools/smt_translate.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/SMT/Tools/smt_translate.ML Sat Nov 14 09:40:27 2009 +0100 @@ -142,7 +142,7 @@ SApp (SConst (@{const_name nat}, _), [SApp (SNum (i, _), _)]) => SOME i | _ => NONE) in -fun bv_rotate mk_name T ts = +fun bv_rotate mk_name _ ts = dest_nat (hd ts) |> Option.map (fn i => (mk_name i, tl ts)) fun bv_extend mk_name T ts = @@ -196,7 +196,7 @@ fun group_quant qname vs (t as Const (q, _) $ Abs (n, T, u)) = if q = qname then group_quant qname ((n, T) :: vs) u else (vs, t) - | group_quant qname vs t = (vs, t) + | group_quant _ vs t = (vs, t) fun dest_trigger (@{term trigger} $ tl $ t) = (HOLogic.dest_list tl, t) | dest_trigger t = ([], t) @@ -208,11 +208,11 @@ fun trans Ts t = (case Term.strip_comb t of - (t1 as Const (qn, qT), [t2 as Abs (n, T, t3)]) => + (Const (qn, _), [Abs (n, T, t1)]) => (case quantifier qn of SOME q => let - val (vs, u) = group_quant qn [(n, T)] t3 + val (vs, u) = group_quant qn [(n, T)] t1 val Us = map snd vs @ Ts val (ps, b) = dest_trigger u in SQuant (q, rev vs, map (pat (trans Us) []) ps, trans Us b) end @@ -277,7 +277,7 @@ fun sep loc t = (case t of - SVar i => pair t + SVar _ => pair t | SApp (c as SConst (@{const_name If}, _), u :: us) => mark sep false u ##>> fold_map (sep loc) us #>> app c o (op ::) | SApp (c, us) => @@ -429,7 +429,7 @@ let val (n, ns') = fresh_typ ns in (n, (vars, ns', add_typ (T, n) sgn)) end) - fun rep_var bs (n, T) (vars, ns, sgn) = + fun rep_var bs (_, T) (vars, ns, sgn) = let val (n', vars') = fresh_name vars in (vars', ns, sgn) |> rep_typ bs T |>> pair n' end @@ -458,9 +458,9 @@ fun sign loc t = (case t of SVar i => pair (SVar i) - | SApp (c as SConst (@{const_name term}, _), [u]) => + | SApp (SConst (@{const_name term}, _), [u]) => sign true u #>> app term_marker o single - | SApp (c as SConst (@{const_name formula}, _), [u]) => + | SApp (SConst (@{const_name formula}, _), [u]) => sign false u #>> app formula_marker o single | SApp (SConst (c as (_, T)), ts) => (case builtin_lookup (builtin_fun loc) thy c ts of diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/SMT/Tools/yices_solver.ML --- a/src/HOL/SMT/Tools/yices_solver.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/SMT/Tools/yices_solver.ML Sat Nov 14 09:40:27 2009 +0100 @@ -19,11 +19,11 @@ fun raise_cex real = raise SMT_Solver.SMT_COUNTEREXAMPLE (real, []) -fun core_oracle ({context, output, recon, ...} : SMT_Solver.proof_data) = +fun core_oracle ({output, ...} : SMT_Solver.proof_data) = let val empty_line = (fn "" => true | _ => false) val split_first = (fn [] => ("", []) | l :: ls => (l, ls)) - val (l, ls) = split_first (dropwhile empty_line output) + val (l, _) = split_first (dropwhile empty_line output) in if String.isPrefix "unsat" l then @{cprop False} else if String.isPrefix "sat" l then raise_cex true diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/SMT/Tools/z3_model.ML --- a/src/HOL/SMT/Tools/z3_model.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/SMT/Tools/z3_model.ML Sat Nov 14 09:40:27 2009 +0100 @@ -130,7 +130,8 @@ in (case (can HOLogic.dest_number t, cs) of (true, [c]) => trans 0 T c #>> (fn (_, u) => [mk_eq u ([], t)]) - | (_, (es, _) :: _) => fold_map (trans (length es) T) cs #>> map (mk_eq t)) + | (_, (es, _) :: _) => fold_map (trans (length es) T) cs #>> map (mk_eq t) + | _ => raise TERM ("translate: no cases", [t])) end diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/SMT/Tools/z3_proof_rules.ML --- a/src/HOL/SMT/Tools/z3_proof_rules.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/SMT/Tools/z3_proof_rules.ML Sat Nov 14 09:40:27 2009 +0100 @@ -97,7 +97,7 @@ | thm_of (Literals (thm, _)) = thm fun meta_eq_of (MetaEq thm) = thm - | meta_eq_of p = thm_of p COMP @{thm eq_reflection} + | meta_eq_of p = mk_meta_eq (thm_of p) datatype proof = Unproved of { @@ -222,7 +222,6 @@ and dest_disj3 = precompose (destd I) @{lemma "~(P | Q) ==> ~Q" by fast} and dest_disj4 = precompose (destd dn2) @{lemma "~(P | ~Q) ==> Q" by fast} - val is_neg = (fn @{term Not} $ _ => true | _ => false) fun dest_disj_rules t = (case dest_disj_term' is_neg t of SOME (true, true) => SOME (dest_disj2, dest_disj4) @@ -458,7 +457,7 @@ let val cv = mk_var ctxt idx (#T (Thm.rep_cterm ct)) in (cv, (ctxt, mk_var, idx + 1, gen, Ctermtab.update (ct, cv) tab)) end) -fun prove_abstraction tac ct (ctxt, _, _, gen, tab) = +fun prove_abstraction tac ct (_, _, _, gen, tab) = let val insts = map swap (Ctermtab.dest tab) val thm = Goal.prove_internal [] ct (fn _ => tac 1) @@ -1088,6 +1087,64 @@ in with_conv (all_distrib_conv ctxt) (prove_arith ctxt thms') ct end end +(** theory simpset **) +local + val antisym_le1 = mk_meta_eq @{thm order_class.antisym_conv} + val antisym_le2 = mk_meta_eq @{thm linorder_class.antisym_conv2} + val antisym_less1 = mk_meta_eq @{thm linorder_class.antisym_conv1} + val antisym_less2 = mk_meta_eq @{thm linorder_class.antisym_conv3} + + fun eq_prop t thm = HOLogic.mk_Trueprop t aconv Thm.prop_of thm + fun dest_binop ((c as Const _) $ t $ u) = (c, t, u) + | dest_binop t = raise TERM ("dest_binop", [t]) + + fun prove_antisym_le ss t = + let + val (le, r, s) = dest_binop t + val less = Const (@{const_name less}, Term.fastype_of le) + val prems = Simplifier.prems_of_ss ss + in + (case find_first (eq_prop (le $ s $ r)) prems of + NONE => + find_first (eq_prop (HOLogic.mk_not (less $ r $ s))) prems + |> Option.map (fn thm => thm RS antisym_less1) + | SOME thm => SOME (thm RS antisym_le1)) + end + handle THM _ => NONE + + fun prove_antisym_less ss t = + let + val (less, r, s) = dest_binop (HOLogic.dest_not t) + val le = Const (@{const_name less_eq}, Term.fastype_of less) + val prems = prems_of_ss ss + in + (case find_first (eq_prop (le $ r $ s)) prems of + NONE => + find_first (eq_prop (HOLogic.mk_not (less $ s $ r))) prems + |> Option.map (fn thm => thm RS antisym_less2) + | SOME thm => SOME (thm RS antisym_le2)) + end + handle THM _ => NONE +in +val z3_simpset = HOL_ss addsimps @{thms array_rules} + addsimps @{thms ring_distribs} addsimps @{thms field_eq_simps} + addsimps @{thms arith_special} addsimps @{thms less_bin_simps} + addsimps @{thms le_bin_simps} addsimps @{thms eq_bin_simps} + addsimps @{thms add_bin_simps} addsimps @{thms succ_bin_simps} + addsimps @{thms minus_bin_simps} addsimps @{thms pred_bin_simps} + addsimps @{thms mult_bin_simps} addsimps @{thms iszero_simps} + addsimprocs [ + Simplifier.simproc @{theory} "fast_int_arith" [ + "(m::int) < n", "(m::int) <= n", "(m::int) = n"] (K Lin_Arith.simproc), + Simplifier.simproc @{theory} "fast_real_arith" [ + "(m::real) < n", "(m::real) <= n", "(m::real) = n"] + (K Lin_Arith.simproc), + Simplifier.simproc @{theory} "antisym le" ["(x::'a::order) <= y"] + (K prove_antisym_le), + Simplifier.simproc @{theory} "antisym less" ["~ (x::'a::linorder) < y"] + (K prove_antisym_less)] +end + (** theory lemmas: linear arithmetic, arrays **) local val array_ss = HOL_ss addsimps @{thms array_rules} @@ -1098,15 +1155,20 @@ fun full_arith_tac ctxt thms = Tactic.cut_facts_tac thms THEN' Arith_Data.arith_tac ctxt + + fun simp_arith_tac ctxt thms = + Tactic.cut_facts_tac thms + THEN' Simplifier.asm_full_simp_tac z3_simpset + THEN' Arith_Data.arith_tac ctxt in fun th_lemma ctxt thms ct = Thm (try_apply ctxt "th-lemma" [ ("abstract arith", arith_lemma ctxt thms), ("array", by_tac' (array_tac thms)), - ("full arith", by_tac' (full_arith_tac ctxt thms))] (T.mk_prop ct)) + ("full arith", by_tac' (full_arith_tac ctxt thms)), + ("simp arith", by_tac' (simp_arith_tac ctxt thms))] (T.mk_prop ct)) end - (** rewriting: prove equalities: * ACI of conjunction/disjunction * contradiction, excluded middle @@ -1190,20 +1252,7 @@ Tactic.rtac iffI_rule THEN_ALL_NEW arith_tac ctxt ORELSE' arith_tac ctxt - val simpset = HOL_ss addsimps @{thms array_rules} - addsimps @{thms ring_distribs} addsimps @{thms field_eq_simps} - addsimps @{thms arith_special} addsimps @{thms less_bin_simps} - addsimps @{thms le_bin_simps} addsimps @{thms eq_bin_simps} - addsimps @{thms add_bin_simps} addsimps @{thms succ_bin_simps} - addsimps @{thms minus_bin_simps} addsimps @{thms pred_bin_simps} - addsimps @{thms mult_bin_simps} addsimps @{thms iszero_simps} - addsimprocs [ - Simplifier.simproc @{theory} "fast_int_arith" [ - "(m::int) < n", "(m::int) <= n", "(m::int) = n"] (K Lin_Arith.simproc), - Simplifier.simproc @{theory} "fast_real_arith" [ - "(m::real) < n", "(m::real) <= n", "(m::real) = n"] - (K Lin_Arith.simproc)] - val simp_tac = CHANGED o Simplifier.simp_tac simpset + fun simp_tac thms = CHANGED o Simplifier.simp_tac (z3_simpset addsimps thms) ORELSE' Classical.best_tac HOL_cs in fun rewrite ctxt thms ct = @@ -1215,7 +1264,7 @@ ("distinct", distinct), ("arith", by_tac' (arith_eq_tac ctxt)), ("classical", by_tac' (Classical.best_tac HOL_cs)), - ("simp", by_tac' simp_tac), + ("simp", by_tac' (simp_tac thms)), ("full arith", by_tac' (Arith_Data.arith_tac ctxt))] (T.mk_prop ct)) end end diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/SMT/Tools/z3_proof_terms.ML --- a/src/HOL/SMT/Tools/z3_proof_terms.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/SMT/Tools/z3_proof_terms.ML Sat Nov 14 09:40:27 2009 +0100 @@ -132,7 +132,7 @@ in mk_preterm (ct, [(i, ct)]) end local -fun mk_quant q thy (n, T) e = +fun mk_quant q thy (_, T) e = let val (ct, vs) = dest_preterm e val cv = diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Statespace/state_space.ML --- a/src/HOL/Statespace/state_space.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Statespace/state_space.ML Sat Nov 14 09:40:27 2009 +0100 @@ -155,13 +155,13 @@ thy |> Expression.add_locale (Binding.name name) (Binding.name name) expr elems |> snd - |> LocalTheory.exit; + |> Local_Theory.exit; fun add_locale_cmd name expr elems thy = thy |> Expression.add_locale_cmd (Binding.name name) Binding.empty expr elems |> snd - |> LocalTheory.exit; + |> Local_Theory.exit; type statespace_info = {args: (string * sort) list, (* type arguments *) @@ -350,8 +350,8 @@ fun add_declaration name decl thy = thy |> Theory_Target.init name - |> (fn lthy => LocalTheory.declaration false (decl lthy) lthy) - |> LocalTheory.exit_global; + |> (fn lthy => Local_Theory.declaration false (decl lthy) lthy) + |> Local_Theory.exit_global; fun parent_components thy (Ts, pname, renaming) = let @@ -430,7 +430,7 @@ let fun upd (n,v) = let - val nT = ProofContext.infer_type (LocalTheory.target_of lthy) n + val nT = ProofContext.infer_type (Local_Theory.target_of lthy) n in Context.proof_map (update_declinfo (Morphism.term phi (Free (n,nT)),v)) end; diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/SupInf.thy --- a/src/HOL/SupInf.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/SupInf.thy Sat Nov 14 09:40:27 2009 +0100 @@ -118,7 +118,7 @@ shows "(!!x. x \ X \ x \ a) \ (!!y. (!!x. x \ X \ x \ y) \ a \ y) \ Sup X = a" by (metis Sup_least Sup_upper add_le_cancel_left diff_add_cancel insert_absorb - insert_not_empty real_le_anti_sym) + insert_not_empty real_le_antisym) lemma Sup_le: fixes S :: "real set" @@ -317,7 +317,7 @@ fixes a :: real shows "(!!x. x \ X \ a \ x) \ (!!y. (!!x. x \ X \ y \ x) \ y \ a) \ Inf X = a" by (metis Inf_greatest Inf_lower add_le_cancel_left diff_add_cancel - insert_absorb insert_not_empty real_le_anti_sym) + insert_absorb insert_not_empty real_le_antisym) lemma Inf_ge: fixes S :: "real set" @@ -397,7 +397,7 @@ fixes S :: "real set" shows "finite S \ S \ {} \ a \ Inf S \ (\ x \ S. a \ x)" by (metis Inf_finite_Min Inf_finite_ge_iff Inf_finite_in Min_le - real_le_anti_sym real_le_linear) + real_le_antisym real_le_linear) lemma Inf_finite_gt_iff: fixes S :: "real set" diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Tools/Datatype/datatype_abs_proofs.ML --- a/src/HOL/Tools/Datatype/datatype_abs_proofs.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Tools/Datatype/datatype_abs_proofs.ML Sat Nov 14 09:40:27 2009 +0100 @@ -156,9 +156,8 @@ thy0 |> Sign.map_naming Name_Space.conceal |> Inductive.add_inductive_global (serial ()) - {quiet_mode = #quiet config, verbose = false, kind = "", - alt_name = Binding.name big_rec_name', coind = false, no_elim = false, no_ind = true, - skip_mono = true, fork_mono = false} + {quiet_mode = #quiet config, verbose = false, alt_name = Binding.name big_rec_name', + coind = false, no_elim = false, no_ind = true, skip_mono = true, fork_mono = false} (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts)) (map dest_Free rec_fns) (map (fn x => (Attrib.empty_binding, x)) rec_intr_ts) [] diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Tools/Datatype/datatype_rep_proofs.ML --- a/src/HOL/Tools/Datatype/datatype_rep_proofs.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Tools/Datatype/datatype_rep_proofs.ML Sat Nov 14 09:40:27 2009 +0100 @@ -175,9 +175,8 @@ thy1 |> Sign.map_naming Name_Space.conceal |> Inductive.add_inductive_global (serial ()) - {quiet_mode = #quiet config, verbose = false, kind = "", - alt_name = Binding.name big_rec_name, coind = false, no_elim = true, no_ind = false, - skip_mono = true, fork_mono = false} + {quiet_mode = #quiet config, verbose = false, alt_name = Binding.name big_rec_name, + coind = false, no_elim = true, no_ind = false, skip_mono = true, fork_mono = false} (map (fn s => ((Binding.name s, UnivT'), NoSyn)) rep_set_names') [] (map (fn x => (Attrib.empty_binding, x)) intr_ts) [] ||> Sign.restore_naming thy1 diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Tools/Function/fun.ML --- a/src/HOL/Tools/Function/fun.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Tools/Function/fun.ML Sat Nov 14 09:40:27 2009 +0100 @@ -153,11 +153,11 @@ fun gen_fun add config fixes statements int lthy = let val group = serial () in lthy - |> LocalTheory.set_group group + |> Local_Theory.set_group group |> add fixes statements config |> by_pat_completeness_auto int - |> LocalTheory.restore - |> LocalTheory.set_group group + |> Local_Theory.restore + |> Local_Theory.set_group group |> termination_by (Function_Common.get_termination_prover lthy) int end; diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Tools/Function/function.ML --- a/src/HOL/Tools/Function/function.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Tools/Function/function.ML Sat Nov 14 09:40:27 2009 +0100 @@ -43,9 +43,6 @@ [Simplifier.simp_add, Nitpick_Psimps.add] -fun note_theorem ((binding, atts), ths) = - LocalTheory.note Thm.generatedK ((binding, atts), ths) - fun mk_defname fixes = fixes |> map (fst o fst) |> space_implode "_" fun add_simps fnames post sort extra_qualify label mod_binding moreatts simps lthy = @@ -55,13 +52,14 @@ |> map (apfst (apfst extra_qualify)) val (saved_spec_simps, lthy) = - fold_map (LocalTheory.note Thm.generatedK) spec lthy + fold_map Local_Theory.note spec lthy val saved_simps = maps snd saved_spec_simps val simps_by_f = sort saved_simps fun add_for_f fname simps = - note_theorem ((mod_binding (Binding.qualify true fname (Binding.name label)), []), simps) + Local_Theory.note + ((mod_binding (Binding.qualify true fname (Binding.name label)), []), simps) #> snd in (saved_simps, @@ -100,14 +98,14 @@ |> addsmps (conceal_partial o Binding.qualify false "partial") "psimps" conceal_partial psimp_attribs psimps ||> fold_option (snd oo addsmps I "simps" I simp_attribs) trsimps - ||>> note_theorem ((conceal_partial (qualify "pinduct"), + ||>> Local_Theory.note ((conceal_partial (qualify "pinduct"), [Attrib.internal (K (Rule_Cases.case_names cnames)), Attrib.internal (K (Rule_Cases.consumes 1)), Attrib.internal (K (Induct.induct_pred ""))]), simple_pinducts) - ||>> note_theorem ((Binding.conceal (qualify "termination"), []), [termination]) - ||> (snd o note_theorem ((qualify "cases", + ||>> Local_Theory.note ((Binding.conceal (qualify "termination"), []), [termination]) + ||> (snd o Local_Theory.note ((qualify "cases", [Attrib.internal (K (Rule_Cases.case_names cnames))]), [cases])) - ||> fold_option (snd oo curry note_theorem (qualify "domintros", [])) domintros + ||> fold_option (snd oo curry Local_Theory.note (qualify "domintros", [])) domintros val cdata = FunctionCtxData { add_simps=addsmps, case_names=cnames, psimps=psimps', pinducts=snd pinducts', termination=termination', @@ -117,11 +115,11 @@ else Specification.print_consts lthy (K false) (map fst fixes) in lthy - |> LocalTheory.declaration false (add_function_data o morph_function_data cdata) + |> Local_Theory.declaration false (add_function_data o morph_function_data cdata) end in lthy - |> is_external ? LocalTheory.set_group (serial ()) + |> is_external ? Local_Theory.set_group (serial ()) |> Proof.theorem_i NONE afterqed [[(Logic.unprotect (concl_of goalstate), [])]] |> Proof.refine (Method.primitive_text (fn _ => goalstate)) |> Seq.hd end @@ -155,7 +153,7 @@ in lthy |> add_simps I "simps" I simp_attribs tsimps |> snd - |> note_theorem + |> Local_Theory.note ((qualify "induct", [Attrib.internal (K (Rule_Cases.case_names case_names))]), tinduct) |> snd @@ -177,12 +175,12 @@ fun termination term_opt lthy = lthy - |> LocalTheory.set_group (serial ()) + |> Local_Theory.set_group (serial ()) |> termination_proof term_opt; fun termination_cmd term_opt lthy = lthy - |> LocalTheory.set_group (serial ()) + |> Local_Theory.set_group (serial ()) |> termination_proof_cmd term_opt; diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Tools/Function/function_core.ML --- a/src/HOL/Tools/Function/function_core.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Tools/Function/function_core.ML Sat Nov 14 09:40:27 2009 +0100 @@ -456,11 +456,10 @@ let val ({intrs = intrs_gen, elims = [elim_gen], preds = [ Rdef ], induct, ...}, lthy) = lthy - |> LocalTheory.conceal + |> Local_Theory.conceal |> Inductive.add_inductive_i {quiet_mode = true, verbose = ! trace, - kind = "", alt_name = Binding.empty, coind = false, no_elim = false, @@ -471,7 +470,7 @@ [] (* no parameters *) (map (fn t => (Attrib.empty_binding, t)) intrs) (* intro rules *) [] (* no special monos *) - ||> LocalTheory.restore_naming lthy + ||> Local_Theory.restore_naming lthy val cert = cterm_of (ProofContext.theory_of lthy) fun requantify orig_intro thm = @@ -519,7 +518,7 @@ $ (default $ Bound 0) $ Abs ("y", ranT, G $ Bound 1 $ Bound 0)) |> Syntax.check_term lthy in - LocalTheory.define "" + Local_Theory.define "" ((Binding.name (function_name fname), mixfix), ((Binding.conceal (Binding.name fdefname), []), f_def)) lthy end @@ -929,7 +928,7 @@ PROFILE "def_rel" (define_recursion_relation (rel_name defname) domT ranT fvar f abstract_qglrs clauses RCss) lthy val (_, lthy) = - LocalTheory.abbrev Syntax.mode_default ((Binding.name (dom_name defname), NoSyn), mk_acc domT R) lthy + Local_Theory.abbrev Syntax.mode_default ((Binding.name (dom_name defname), NoSyn), mk_acc domT R) lthy val newthy = ProofContext.theory_of lthy val clauses = map (transfer_clause_ctx newthy) clauses diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Tools/Function/mutual.ML --- a/src/HOL/Tools/Function/mutual.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Tools/Function/mutual.ML Sat Nov 14 09:40:27 2009 +0100 @@ -146,7 +146,7 @@ fun def ((MutualPart {i=i, i'=i', fvar=(fname, fT), cargTs, f_def, ...}), (_, mixfix)) lthy = let val ((f, (_, f_defthm)), lthy') = - LocalTheory.define "" + Local_Theory.define "" ((Binding.name fname, mixfix), ((Binding.conceal (Binding.name (fname ^ "_def")), []), Term.subst_bound (fsum, f_def))) lthy diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Tools/Function/size.ML --- a/src/HOL/Tools/Function/size.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Tools/Function/size.ML Sat Nov 14 09:40:27 2009 +0100 @@ -130,7 +130,7 @@ fun define_overloaded (def_name, eq) lthy = let val (Free (c, _), rhs) = (Logic.dest_equals o Syntax.check_term lthy) eq; - val ((_, (_, thm)), lthy') = lthy |> LocalTheory.define Thm.definitionK + val ((_, (_, thm)), lthy') = lthy |> Local_Theory.define Thm.definitionK ((Binding.name c, NoSyn), ((Binding.name def_name, []), rhs)); val ctxt_thy = ProofContext.init (ProofContext.theory_of lthy'); val thm' = singleton (ProofContext.export lthy' ctxt_thy) thm; @@ -149,7 +149,7 @@ ||>> fold_map define_overloaded (def_names' ~~ map Logic.mk_equals (overloaded_size_fns ~~ map (app fs') rec_combs1)) ||> Class.prove_instantiation_instance (K (Class.intro_classes_tac [])) - ||> LocalTheory.exit_global; + ||> Local_Theory.exit_global; val ctxt = ProofContext.init thy'; diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Tools/Predicate_Compile/predicate_compile.ML --- a/src/HOL/Tools/Predicate_Compile/predicate_compile.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Tools/Predicate_Compile/predicate_compile.ML Sat Nov 14 09:40:27 2009 +0100 @@ -137,8 +137,8 @@ in if (is_inductify options) then let - val lthy' = LocalTheory.theory (preprocess options const) lthy - |> LocalTheory.checkpoint + val lthy' = Local_Theory.theory (preprocess options const) lthy + |> Local_Theory.checkpoint val const = case Predicate_Compile_Fun.pred_of_function (ProofContext.theory_of lthy') const of SOME c => c diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML --- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Sat Nov 14 09:40:27 2009 +0100 @@ -2411,9 +2411,9 @@ let val thy = ProofContext.theory_of lthy val const = prep_const thy raw_const - val lthy' = LocalTheory.theory (PredData.map + val lthy' = Local_Theory.theory (PredData.map (extend (fetch_pred_data thy) (depending_preds_of thy) const)) lthy - |> LocalTheory.checkpoint + |> Local_Theory.checkpoint val thy' = ProofContext.theory_of lthy' val preds = Graph.all_succs (PredData.get thy') [const] |> filter_out (has_elim thy') fun mk_cases const = @@ -2437,7 +2437,7 @@ val global_thms = ProofContext.export goal_ctxt (ProofContext.init (ProofContext.theory_of goal_ctxt)) (map the_single thms) in - goal_ctxt |> LocalTheory.theory (fold set_elim global_thms #> + goal_ctxt |> Local_Theory.theory (fold set_elim global_thms #> (if is_random options then (add_equations options [const] #> add_quickcheck_equations options [const]) diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML --- a/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_fun.ML Sat Nov 14 09:40:27 2009 +0100 @@ -355,9 +355,8 @@ thy |> Sign.map_naming Name_Space.conceal |> Inductive.add_inductive_global (serial ()) - {quiet_mode = false, verbose = false, kind = "", - alt_name = Binding.empty, coind = false, no_elim = false, - no_ind = false, skip_mono = false, fork_mono = false} + {quiet_mode = false, verbose = false, alt_name = Binding.empty, coind = false, + no_elim = false, no_ind = false, skip_mono = false, fork_mono = false} (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (distinct (op =) (map dest_Free preds))) pnames diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Tools/inductive.ML --- a/src/HOL/Tools/inductive.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Tools/inductive.ML Sat Nov 14 09:40:27 2009 +0100 @@ -39,8 +39,8 @@ val inductive_cases_i: (Attrib.binding * term list) list -> local_theory -> thm list list * local_theory type inductive_flags = - {quiet_mode: bool, verbose: bool, kind: string, alt_name: binding, - coind: bool, no_elim: bool, no_ind: bool, skip_mono: bool, fork_mono: bool} + {quiet_mode: bool, verbose: bool, alt_name: binding, coind: bool, + no_elim: bool, no_ind: bool, skip_mono: bool, fork_mono: bool} val add_inductive_i: inductive_flags -> ((binding * typ) * mixfix) list -> (string * typ) list -> (Attrib.binding * term) list -> thm list -> local_theory -> @@ -71,7 +71,7 @@ term list -> (Attrib.binding * term) list -> thm list -> term list -> (binding * mixfix) list -> local_theory -> inductive_result * local_theory - val declare_rules: string -> binding -> bool -> bool -> string list -> + val declare_rules: binding -> bool -> bool -> string list -> thm list -> binding list -> Attrib.src list list -> (thm * string list) list -> thm -> local_theory -> thm list * thm list * thm * local_theory val add_ind_def: add_ind_def @@ -469,7 +469,7 @@ val facts = args |> map (fn ((a, atts), props) => ((a, map (prep_att thy) atts), map (Thm.no_attributes o single o mk_cases lthy o prep_prop lthy) props)); - in lthy |> LocalTheory.notes Thm.generatedK facts |>> map snd end; + in lthy |> Local_Theory.notes facts |>> map snd end; val inductive_cases = gen_inductive_cases Attrib.intern_src Syntax.read_prop; val inductive_cases_i = gen_inductive_cases (K I) Syntax.check_prop; @@ -509,7 +509,8 @@ fun mk_ind_prem r = let - fun subst s = (case dest_predicate cs params s of + fun subst s = + (case dest_predicate cs params s of SOME (_, i, ys, (_, Ts)) => let val k = length Ts; @@ -520,10 +521,11 @@ HOLogic.mk_binop inductive_conj_name (list_comb (incr_boundvars k s, bs), P)) in (Q, case Ts of [] => SOME (s, P) | _ => NONE) end - | NONE => (case s of - (t $ u) => (fst (subst t) $ fst (subst u), NONE) - | (Abs (a, T, t)) => (Abs (a, T, fst (subst t)), NONE) - | _ => (s, NONE))); + | NONE => + (case s of + (t $ u) => (fst (subst t) $ fst (subst u), NONE) + | (Abs (a, T, t)) => (Abs (a, T, fst (subst t)), NONE) + | _ => (s, NONE))); fun mk_prem s prems = (case subst s of @@ -618,16 +620,17 @@ SOME (_, i, ts, (Ts, Us)) => let val l = length Us; - val zs = map Bound (l - 1 downto 0) + val zs = map Bound (l - 1 downto 0); in list_abs (map (pair "z") Us, list_comb (p, make_bool_args' bs i @ make_args argTs ((map (incr_boundvars l) ts ~~ Ts) @ (zs ~~ Us)))) end - | NONE => (case t of - t1 $ t2 => subst t1 $ subst t2 - | Abs (x, T, u) => Abs (x, T, subst u) - | _ => t)); + | NONE => + (case t of + t1 $ t2 => subst t1 $ subst t2 + | Abs (x, T, u) => Abs (x, T, subst u) + | _ => t)); (* transform an introduction rule into a conjunction *) (* [| p_i t; ... |] ==> p_j u *) @@ -662,13 +665,13 @@ else alt_name; val ((rec_const, (_, fp_def)), lthy') = lthy - |> LocalTheory.conceal - |> LocalTheory.define "" + |> Local_Theory.conceal + |> Local_Theory.define "" ((rec_name, case cnames_syn of [(_, syn)] => syn | _ => NoSyn), ((Binding.empty, [Attrib.internal (K Nitpick_Defs.add)]), fold_rev lambda params (Const (fp_name, (predT --> predT) --> predT) $ fp_fun))) - ||> LocalTheory.restore_naming lthy; + ||> Local_Theory.restore_naming lthy; val fp_def' = Simplifier.rewrite (HOL_basic_ss addsimps [fp_def]) (cterm_of (ProofContext.theory_of lthy') (list_comb (rec_const, params))); val specs = @@ -685,21 +688,21 @@ make_args argTs (xs ~~ Ts))))) end) (cnames_syn ~~ cs); val (consts_defs, lthy'') = lthy' - |> LocalTheory.conceal - |> fold_map (LocalTheory.define "") specs - ||> LocalTheory.restore_naming lthy'; + |> Local_Theory.conceal + |> fold_map (Local_Theory.define "") specs + ||> Local_Theory.restore_naming lthy'; val preds = (case cs of [_] => [rec_const] | _ => map #1 consts_defs); val mono = prove_mono quiet_mode skip_mono fork_mono predT fp_fun monos lthy''; val ((_, [mono']), lthy''') = - LocalTheory.note "" (apfst Binding.conceal Attrib.empty_binding, [mono]) lthy''; + Local_Theory.note (apfst Binding.conceal Attrib.empty_binding, [mono]) lthy''; in (lthy''', rec_name, mono', fp_def', map (#2 o #2) consts_defs, list_comb (rec_const, params), preds, argTs, bs, xs) end; -fun declare_rules kind rec_binding coind no_ind cnames - intrs intr_bindings intr_atts elims raw_induct lthy = +fun declare_rules rec_binding coind no_ind cnames + intrs intr_bindings intr_atts elims raw_induct lthy = let val rec_name = Binding.name_of rec_binding; fun rec_qualified qualified = Binding.qualify qualified rec_name; @@ -716,7 +719,7 @@ val (intrs', lthy1) = lthy |> - LocalTheory.notes kind + Local_Theory.notes (map (rec_qualified false) intr_bindings ~~ intr_atts ~~ map (fn th => [([th], [Attrib.internal (K (Context_Rules.intro_query NONE)), @@ -724,16 +727,16 @@ map (hd o snd); val (((_, elims'), (_, [induct'])), lthy2) = lthy1 |> - LocalTheory.note kind ((rec_qualified true (Binding.name "intros"), []), intrs') ||>> + Local_Theory.note ((rec_qualified true (Binding.name "intros"), []), intrs') ||>> fold_map (fn (name, (elim, cases)) => - LocalTheory.note kind + Local_Theory.note ((Binding.qualify true (Long_Name.base_name name) (Binding.name "cases"), [Attrib.internal (K (Rule_Cases.case_names cases)), Attrib.internal (K (Rule_Cases.consumes 1)), Attrib.internal (K (Induct.cases_pred name)), Attrib.internal (K (Context_Rules.elim_query NONE))]), [elim]) #> apfst (hd o snd)) (if null elims then [] else cnames ~~ elims) ||>> - LocalTheory.note kind + Local_Theory.note ((rec_qualified true (Binding.name (coind_prefix coind ^ "induct")), map (Attrib.internal o K) (#2 induct)), [rulify (#1 induct)]); @@ -742,7 +745,7 @@ else let val inducts = cnames ~~ Project_Rule.projects lthy2 (1 upto length cnames) induct' in lthy2 |> - LocalTheory.notes kind [((rec_qualified true (Binding.name "inducts"), []), + Local_Theory.notes [((rec_qualified true (Binding.name "inducts"), []), inducts |> map (fn (name, th) => ([th], [Attrib.internal (K ind_case_names), Attrib.internal (K (Rule_Cases.consumes 1)), @@ -751,8 +754,8 @@ in (intrs', elims', induct', lthy3) end; type inductive_flags = - {quiet_mode: bool, verbose: bool, kind: string, alt_name: binding, - coind: bool, no_elim: bool, no_ind: bool, skip_mono: bool, fork_mono: bool}; + {quiet_mode: bool, verbose: bool, alt_name: binding, coind: bool, + no_elim: bool, no_ind: bool, skip_mono: bool, fork_mono: bool}; type add_ind_def = inductive_flags -> @@ -760,8 +763,7 @@ term list -> (binding * mixfix) list -> local_theory -> inductive_result * local_theory; -fun add_ind_def - {quiet_mode, verbose, kind, alt_name, coind, no_elim, no_ind, skip_mono, fork_mono} +fun add_ind_def {quiet_mode, verbose, alt_name, coind, no_elim, no_ind, skip_mono, fork_mono} cs intros monos params cnames_syn lthy = let val _ = null cnames_syn andalso error "No inductive predicates given"; @@ -769,7 +771,7 @@ val _ = message (quiet_mode andalso not verbose) ("Proofs for " ^ coind_prefix coind ^ "inductive predicate(s) " ^ commas_quote names); - val cnames = map (LocalTheory.full_name lthy o #1) cnames_syn; (* FIXME *) + val cnames = map (Local_Theory.full_name lthy o #1) cnames_syn; (* FIXME *) val ((intr_names, intr_atts), intr_ts) = apfst split_list (split_list (map (check_rule lthy cs params) intros)); @@ -797,7 +799,7 @@ prove_indrule quiet_mode cs argTs bs xs rec_const params intr_ts mono fp_def rec_preds_defs lthy1); - val (intrs', elims', induct, lthy2) = declare_rules kind rec_name coind no_ind + val (intrs', elims', induct, lthy2) = declare_rules rec_name coind no_ind cnames intrs intr_names intr_atts elims raw_induct lthy1; val result = @@ -808,7 +810,7 @@ induct = induct}; val lthy3 = lthy2 - |> LocalTheory.declaration false (fn phi => + |> Local_Theory.declaration false (fn phi => let val result' = morph_result phi result; in put_inductives cnames (*global names!?*) ({names = cnames, coind = coind}, result') end); in (result, lthy3) end; @@ -817,7 +819,7 @@ (* external interfaces *) fun gen_add_inductive_i mk_def - (flags as {quiet_mode, verbose, kind, alt_name, coind, no_elim, no_ind, skip_mono, fork_mono}) + (flags as {quiet_mode, verbose, alt_name, coind, no_elim, no_ind, skip_mono, fork_mono}) cnames_syn pnames spec monos lthy = let val thy = ProofContext.theory_of lthy; @@ -870,7 +872,7 @@ in lthy |> mk_def flags cs intros monos ps preds - ||> fold (snd oo LocalTheory.abbrev Syntax.mode_default) abbrevs + ||> fold (snd oo Local_Theory.abbrev Syntax.mode_default) abbrevs end; fun gen_add_inductive mk_def verbose coind cnames_syn pnames_syn intro_srcs raw_monos int lthy = @@ -880,12 +882,11 @@ |> Specification.read_spec (cnames_syn @ pnames_syn) intro_srcs; val (cs, ps) = chop (length cnames_syn) vars; val monos = Attrib.eval_thms lthy raw_monos; - val flags = {quiet_mode = false, verbose = verbose, kind = Thm.generatedK, - alt_name = Binding.empty, coind = coind, no_elim = false, no_ind = false, - skip_mono = false, fork_mono = not int}; + val flags = {quiet_mode = false, verbose = verbose, alt_name = Binding.empty, + coind = coind, no_elim = false, no_ind = false, skip_mono = false, fork_mono = not int}; in lthy - |> LocalTheory.set_group (serial ()) + |> Local_Theory.set_group (serial ()) |> gen_add_inductive_i mk_def flags cs (map (apfst Binding.name_of o fst) ps) intrs monos end; @@ -897,9 +898,9 @@ val name = Sign.full_name thy (fst (fst (hd cnames_syn))); val ctxt' = thy |> Theory_Target.init NONE - |> LocalTheory.set_group group + |> Local_Theory.set_group group |> add_inductive_i flags cnames_syn pnames pre_intros monos |> snd - |> LocalTheory.exit; + |> Local_Theory.exit; val info = #2 (the_inductive ctxt' name); in (info, ProofContext.theory_of ctxt') end; diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Tools/inductive_realizer.ML --- a/src/HOL/Tools/inductive_realizer.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Tools/inductive_realizer.ML Sat Nov 14 09:40:27 2009 +0100 @@ -14,7 +14,7 @@ structure InductiveRealizer : INDUCTIVE_REALIZER = struct -(* FIXME: LocalTheory.note should return theorems with proper names! *) (* FIXME ?? *) +(* FIXME: Local_Theory.note should return theorems with proper names! *) (* FIXME ?? *) fun name_of_thm thm = (case Proofterm.fold_proof_atoms false (fn PThm (_, ((name, _, _), _)) => cons name | _ => I) [Thm.proof_of thm] [] of @@ -351,8 +351,8 @@ val (ind_info, thy3') = thy2 |> Inductive.add_inductive_global (serial ()) - {quiet_mode = false, verbose = false, kind = Thm.generatedK, alt_name = Binding.empty, - coind = false, no_elim = false, no_ind = false, skip_mono = false, fork_mono = false} + {quiet_mode = false, verbose = false, alt_name = Binding.empty, coind = false, + no_elim = false, no_ind = false, skip_mono = false, fork_mono = false} rlzpreds rlzparams (map (fn (rintr, intr) => ((Binding.name (Long_Name.base_name (name_of_thm intr)), []), subst_atomic rlzpreds' (Logic.unvarify rintr))) diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Tools/inductive_set.ML --- a/src/HOL/Tools/inductive_set.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Tools/inductive_set.ML Sat Nov 14 09:40:27 2009 +0100 @@ -224,7 +224,7 @@ map (fn (x, ps) => let val U = HOLogic.dest_setT (fastype_of x); - val x' = map_type (K (HOLogic.strip_ptupleT ps U ---> HOLogic.boolT)) x + val x' = map_type (K (HOLogic.strip_ptupleT ps U ---> HOLogic.boolT)) x; in (cterm_of thy x, cterm_of thy (HOLogic.Collect_const U $ @@ -405,7 +405,7 @@ (**** definition of inductive sets ****) fun add_ind_set_def - {quiet_mode, verbose, kind, alt_name, coind, no_elim, no_ind, skip_mono, fork_mono} + {quiet_mode, verbose, alt_name, coind, no_elim, no_ind, skip_mono, fork_mono} cs intros monos params cnames_syn lthy = let val thy = ProofContext.theory_of lthy; @@ -477,20 +477,20 @@ val monos' = map (to_pred [] (Context.Proof lthy)) monos; val ({preds, intrs, elims, raw_induct, ...}, lthy1) = Inductive.add_ind_def - {quiet_mode = quiet_mode, verbose = verbose, kind = kind, alt_name = Binding.empty, + {quiet_mode = quiet_mode, verbose = verbose, alt_name = Binding.empty, coind = coind, no_elim = no_elim, no_ind = no_ind, skip_mono = skip_mono, fork_mono = fork_mono} cs' intros' monos' params1 cnames_syn' lthy; (* define inductive sets using previously defined predicates *) val (defs, lthy2) = lthy1 - |> LocalTheory.conceal (* FIXME ?? *) - |> fold_map (LocalTheory.define "") + |> Local_Theory.conceal (* FIXME ?? *) + |> fold_map (Local_Theory.define "") (map (fn ((c_syn, (fs, U, _)), p) => (c_syn, (Attrib.empty_binding, fold_rev lambda params (HOLogic.Collect_const U $ HOLogic.mk_psplits fs U HOLogic.boolT (list_comb (p, params3)))))) (cnames_syn ~~ cs_info ~~ preds)) - ||> LocalTheory.restore_naming lthy1; + ||> Local_Theory.restore_naming lthy1; (* prove theorems for converting predicate to set notation *) val lthy3 = fold @@ -505,7 +505,7 @@ (K (REPEAT (rtac ext 1) THEN simp_tac (HOL_basic_ss addsimps [def, mem_Collect_eq, split_conv]) 1)) in - lthy |> LocalTheory.note kind ((Binding.name (s ^ "p_" ^ s ^ "_eq"), + lthy |> Local_Theory.note ((Binding.name (s ^ "p_" ^ s ^ "_eq"), [Attrib.internal (K pred_set_conv_att)]), [conv_thm]) |> snd end) (preds ~~ cs ~~ cs_info ~~ defs) lthy2; @@ -515,11 +515,11 @@ if Binding.is_empty alt_name then Binding.name (space_implode "_" (map (Binding.name_of o fst) cnames_syn)) else alt_name; - val cnames = map (LocalTheory.full_name lthy3 o #1) cnames_syn; (* FIXME *) + val cnames = map (Local_Theory.full_name lthy3 o #1) cnames_syn; (* FIXME *) val (intr_names, intr_atts) = split_list (map fst intros); val raw_induct' = to_set [] (Context.Proof lthy3) raw_induct; val (intrs', elims', induct, lthy4) = - Inductive.declare_rules kind rec_name coind no_ind cnames + Inductive.declare_rules rec_name coind no_ind cnames (map (to_set [] (Context.Proof lthy3)) intrs) intr_names intr_atts (map (fn th => (to_set [] (Context.Proof lthy3) th, map fst (fst (Rule_Cases.get th)))) elims) diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Tools/primrec.ML --- a/src/HOL/Tools/primrec.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Tools/primrec.ML Sat Nov 14 09:40:27 2009 +0100 @@ -259,7 +259,7 @@ val ((prefix, (fs, defs)), prove) = distill lthy fixes ts; in lthy - |> fold_map (LocalTheory.define Thm.definitionK) defs + |> fold_map (Local_Theory.define Thm.definitionK) defs |-> (fn defs => `(fn lthy => (prefix, prove lthy defs))) end; @@ -275,12 +275,10 @@ map (Attrib.internal o K) [Simplifier.simp_add, Nitpick_Simps.add]); in lthy - |> set_group ? LocalTheory.set_group (serial ()) + |> set_group ? Local_Theory.set_group (serial ()) |> add_primrec_simple fixes (map snd spec) - |-> (fn (prefix, simps) => fold_map (LocalTheory.note Thm.generatedK) - (attr_bindings prefix ~~ simps) - #-> (fn simps' => LocalTheory.note Thm.generatedK - (simp_attr_binding prefix, maps snd simps'))) + |-> (fn (prefix, simps) => fold_map Local_Theory.note (attr_bindings prefix ~~ simps) + #-> (fn simps' => Local_Theory.note (simp_attr_binding prefix, maps snd simps'))) |>> snd end; @@ -296,14 +294,14 @@ val lthy = Theory_Target.init NONE thy; val (simps, lthy') = add_primrec fixes specs lthy; val simps' = ProofContext.export lthy' lthy simps; - in (simps', LocalTheory.exit_global lthy') end; + in (simps', Local_Theory.exit_global lthy') end; fun add_primrec_overloaded ops fixes specs thy = let val lthy = Theory_Target.overloading ops thy; val (simps, lthy') = add_primrec fixes specs lthy; val simps' = ProofContext.export lthy' lthy simps; - in (simps', LocalTheory.exit_global lthy') end; + in (simps', Local_Theory.exit_global lthy') end; (* outer syntax *) diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Tools/quickcheck_generators.ML --- a/src/HOL/Tools/quickcheck_generators.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Tools/quickcheck_generators.ML Sat Nov 14 09:40:27 2009 +0100 @@ -190,7 +190,7 @@ in lthy |> random_aux_primrec aux_eq' - ||>> fold_map (LocalTheory.define Thm.definitionK) proj_defs + ||>> fold_map (Local_Theory.define Thm.definitionK) proj_defs |-> (fn (aux_simp, proj_defs) => prove_eqs aux_simp proj_defs) end; @@ -214,8 +214,8 @@ lthy |> random_aux_primrec_multi (name ^ prfx) proto_eqs |-> (fn proto_simps => prove_simps proto_simps) - |-> (fn simps => LocalTheory.note Thm.generatedK ((b, - Code.add_default_eqn_attrib :: map (Attrib.internal o K) + |-> (fn simps => Local_Theory.note + ((b, Code.add_default_eqn_attrib :: map (Attrib.internal o K) [Simplifier.simp_add, Nitpick_Simps.add]), simps)) |> snd end diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Transcendental.thy --- a/src/HOL/Transcendental.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Transcendental.thy Sat Nov 14 09:40:27 2009 +0100 @@ -1213,6 +1213,9 @@ apply (simp_all add: abs_if isCont_ln) done +lemma DERIV_ln_divide: "0 < x ==> DERIV ln x :> 1 / x" + by (rule DERIV_ln[THEN DERIV_cong], simp, simp add: divide_inverse) + lemma ln_series: assumes "0 < x" and "x < 2" shows "ln x = (\ n. (-1)^n * (1 / real (n + 1)) * (x - 1)^(Suc n))" (is "ln x = suminf (?f (x - 1))") proof - @@ -1702,9 +1705,8 @@ apply (drule_tac f = cos in Rolle) apply (drule_tac [5] f = cos in Rolle) apply (auto dest!: DERIV_cos [THEN DERIV_unique] simp add: differentiable_def) -apply (drule_tac y1 = xa in order_le_less_trans [THEN sin_gt_zero]) -apply (assumption, rule_tac y=y in order_less_le_trans, simp_all) -apply (drule_tac y1 = y in order_le_less_trans [THEN sin_gt_zero], assumption, simp_all) +apply (metis order_less_le_trans real_less_def sin_gt_zero) +apply (metis order_less_le_trans real_less_def sin_gt_zero) done lemma pi_half: "pi/2 = (THE x. 0 \ x & x \ 2 & cos x = 0)" @@ -2436,14 +2438,8 @@ apply (rule arctan_ubound [of x, THEN dense, THEN exE], clarify) apply (subgoal_tac "isCont arctan (tan (arctan x))", simp) apply (erule (1) isCont_inverse_function2 [where f=tan]) -apply (clarify, rule arctan_tan) -apply (erule (1) order_less_le_trans) -apply (erule (1) order_le_less_trans) -apply (clarify, rule isCont_tan) -apply (rule less_imp_neq [symmetric]) -apply (rule cos_gt_zero_pi) -apply (erule (1) order_less_le_trans) -apply (erule (1) order_le_less_trans) +apply (metis arctan_tan order_le_less_trans order_less_le_trans) +apply (metis cos_gt_zero_pi isCont_tan order_less_le_trans real_less_def) done lemma DERIV_arcsin: @@ -3119,8 +3115,7 @@ lemma polar_ex2: "y < 0 ==> \r a. x = r * cos a & y = r * sin a" apply (insert polar_ex1 [where x=x and y="-y"], simp, clarify) -apply (rule_tac x = r in exI) -apply (rule_tac x = "-a" in exI, simp) +apply (metis cos_minus minus_minus minus_mult_right sin_minus) done lemma polar_Ex: "\r a. x = r * cos a & y = r * sin a" diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/Transitive_Closure.thy --- a/src/HOL/Transitive_Closure.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/Transitive_Closure.thy Sat Nov 14 09:40:27 2009 +0100 @@ -575,6 +575,25 @@ "(x,y) \ R\<^sup>* = (x=y \ x\y \ (x,y) \ R\<^sup>+)" by (fast elim: trancl_into_rtrancl dest: rtranclD) +lemma trancl_unfold_right: "r^+ = r^* O r" +by (auto dest: tranclD2 intro: rtrancl_into_trancl1) + +lemma trancl_unfold_left: "r^+ = r O r^*" +by (auto dest: tranclD intro: rtrancl_into_trancl2) + + +text {* Simplifying nested closures *} + +lemma rtrancl_trancl_absorb[simp]: "(R^*)^+ = R^*" +by (simp add: trans_rtrancl) + +lemma trancl_rtrancl_absorb[simp]: "(R^+)^* = R^*" +by (subst reflcl_trancl[symmetric]) simp + +lemma rtrancl_reflcl_absorb[simp]: "(R^*)^= = R^*" +by auto + + text {* @{text Domain} and @{text Range} *} lemma Domain_rtrancl [simp]: "Domain (R^*) = UNIV" diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOL/ex/LocaleTest2.thy --- a/src/HOL/ex/LocaleTest2.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOL/ex/LocaleTest2.thy Sat Nov 14 09:40:27 2009 +0100 @@ -29,7 +29,7 @@ locale dpo = fixes le :: "['a, 'a] => bool" (infixl "\" 50) assumes refl [intro, simp]: "x \ x" - and anti_sym [intro]: "[| x \ y; y \ x |] ==> x = y" + and antisym [intro]: "[| x \ y; y \ x |] ==> x = y" and trans [trans]: "[| x \ y; y \ z |] ==> x \ z" begin @@ -87,7 +87,7 @@ assume inf: "is_inf x y i" assume inf': "is_inf x y i'" show ?thesis - proof (rule anti_sym) + proof (rule antisym) from inf' show "i \ i'" proof (rule is_inf_greatest) from inf show "i \ x" .. @@ -159,7 +159,7 @@ assume sup: "is_sup x y s" assume sup': "is_sup x y s'" show ?thesis - proof (rule anti_sym) + proof (rule antisym) from sup show "s \ s'" proof (rule is_sup_least) from sup' show "x \ s'" .. @@ -457,7 +457,7 @@ moreover { assume c: "x \ y | x \ z" from c have "?l = x" - by (metis (*anti_sym*) (*c*) (*circular*) (*join_assoc*)(* join_commute *) join_connection2 (*join_left*) join_related2 meet_connection(* meet_related2*) total trans) + by (metis (*antisym*) (*c*) (*circular*) (*join_assoc*)(* join_commute *) join_connection2 (*join_left*) join_related2 meet_connection(* meet_related2*) total trans) also from c have "... = ?r" by (metis join_commute join_related2 meet_connection meet_related2 total) finally have "?l = ?r" . } diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOLCF/Representable.thy --- a/src/HOLCF/Representable.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOLCF/Representable.thy Sat Nov 14 09:40:27 2009 +0100 @@ -6,6 +6,7 @@ theory Representable imports Algebraic Universal Ssum Sprod One ConvexPD +uses ("Tools/repdef.ML") begin subsection {* Class of representable types *} @@ -174,16 +175,21 @@ setup {* Sign.add_const_constraint (@{const_name prj}, SOME @{typ "udom \ 'a::pcpo"}) *} +definition + repdef_approx :: + "('a::pcpo \ udom) \ (udom \ 'a) \ udom alg_defl \ nat \ 'a \ 'a" +where + "repdef_approx Rep Abs t = (\i. \ x. Abs (cast\(approx i\t)\(Rep x)))" + lemma typedef_rep_class: fixes Rep :: "'a::pcpo \ udom" fixes Abs :: "udom \ 'a::pcpo" fixes t :: TypeRep assumes type: "type_definition Rep Abs {x. x ::: t}" assumes below: "op \ \ \x y. Rep x \ Rep y" - assumes emb: "emb = (\ x. Rep x)" - assumes prj: "prj = (\ x. Abs (cast\t\x))" - assumes approx: - "(approx :: nat \ 'a \ 'a) = (\i. prj oo cast\(approx i\t) oo emb)" + assumes emb: "emb \ (\ x. Rep x)" + assumes prj: "prj \ (\ x. Abs (cast\t\x))" + assumes approx: "(approx :: nat \ 'a \ 'a) \ repdef_approx Rep Abs t" shows "OFCLASS('a, rep_class)" proof have adm: "adm (\x. x \ {x. x ::: t})" @@ -199,6 +205,19 @@ apply (rule typedef_cont_Abs [OF type below adm]) apply simp_all done + have cast_cast_approx: + "\i x. cast\t\(cast\(approx i\t)\x) = cast\(approx i\t)\x" + apply (rule cast_fixed) + apply (rule subdeflationD) + apply (rule approx.below) + apply (rule cast_in_deflation) + done + have approx': + "approx = (\i. \(x::'a). prj\(cast\(approx i\t)\(emb\x)))" + unfolding approx repdef_approx_def + apply (subst cast_cast_approx [symmetric]) + apply (simp add: prj_beta [symmetric] emb_beta [symmetric]) + done have emb_in_deflation: "\x::'a. emb\x ::: t" using type_definition.Rep [OF type] by (simp add: emb_beta) @@ -216,22 +235,15 @@ apply (simp add: emb_prj cast.below) done show "chain (approx :: nat \ 'a \ 'a)" - unfolding approx by simp + unfolding approx' by simp show "\x::'a. (\i. approx i\x) = x" - unfolding approx + unfolding approx' apply (simp add: lub_distribs) apply (subst cast_fixed [OF emb_in_deflation]) apply (rule prj_emb) done - have cast_cast_approx: - "\i x. cast\t\(cast\(approx i\t)\x) = cast\(approx i\t)\x" - apply (rule cast_fixed) - apply (rule subdeflationD) - apply (rule approx.below) - apply (rule cast_in_deflation) - done show "\(i::nat) (x::'a). approx i\(approx i\x) = approx i\x" - unfolding approx + unfolding approx' apply simp apply (simp add: emb_prj) apply (simp add: cast_cast_approx) @@ -239,7 +251,7 @@ show "\i::nat. finite {x::'a. approx i\x = x}" apply (rule_tac B="(\x. prj\x) ` {x. cast\(approx i\t)\x = x}" in finite_subset) - apply (clarsimp simp add: approx) + apply (clarsimp simp add: approx') apply (drule_tac f="\x. emb\x" in arg_cong) apply (rule image_eqI) apply (rule prj_emb [symmetric]) @@ -269,8 +281,8 @@ fixes t :: TypeRep assumes type: "type_definition Rep Abs {x. x ::: t}" assumes below: "op \ \ \x y. Rep x \ Rep y" - assumes emb: "emb = (\ x. Rep x)" - assumes prj: "prj = (\ x. Abs (cast\t\x))" + assumes emb: "emb \ (\ x. Rep x)" + assumes prj: "prj \ (\ x. Abs (cast\t\x))" shows "REP('a) = t" proof - have adm: "adm (\x. x \ {x. x ::: t})" @@ -303,6 +315,11 @@ done qed +lemma adm_mem_Collect_in_deflation: "adm (\x. x \ {x. x ::: A})" +unfolding mem_Collect_eq by (rule adm_in_deflation) + +use "Tools/repdef.ML" + subsection {* Instances of class @{text rep} *} diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOLCF/Tools/fixrec.ML --- a/src/HOLCF/Tools/fixrec.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOLCF/Tools/fixrec.ML Sat Nov 14 09:40:27 2009 +0100 @@ -209,7 +209,7 @@ | defs (l::[]) r = [one_def l r] | defs (l::ls) r = one_def l (mk_fst r) :: defs ls (mk_snd r); val fixdefs = defs lhss fixpoint; - val define_all = fold_map (LocalTheory.define Thm.definitionK); + val define_all = fold_map (Local_Theory.define Thm.definitionK); val (fixdef_thms : (term * (string * thm)) list, lthy') = lthy |> define_all (map (apfst fst) fixes ~~ fixdefs); fun pair_equalI (thm1, thm2) = @{thm Pair_equalI} OF [thm1, thm2]; @@ -242,8 +242,7 @@ ((thm_name, [src]), [thm]) end; val (thmss, lthy'') = lthy' - |> fold_map (LocalTheory.note Thm.generatedK) - (induct_note :: map unfold_note unfold_thms); + |> fold_map Local_Theory.note (induct_note :: map unfold_note unfold_thms); in (lthy'', names, fixdef_thms, map snd unfold_thms) end; @@ -465,7 +464,7 @@ val simps2 : (Attrib.binding * thm list) list = map (apsnd (fn thm => [thm])) (flat simps); val (_, lthy'') = lthy' - |> fold_map (LocalTheory.note Thm.generatedK) (simps1 @ simps2); + |> fold_map Local_Theory.note (simps1 @ simps2); in lthy'' end diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOLCF/Tools/pcpodef.ML --- a/src/HOLCF/Tools/pcpodef.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOLCF/Tools/pcpodef.ML Sat Nov 14 09:40:27 2009 +0100 @@ -87,6 +87,7 @@ thy |> AxClass.prove_arity (full_tname, lhs_sorts, @{sort cpo}) (Tactic.rtac (@{thm typedef_cpo} OF cpo_thms) 1); + (* transfer thms so that they will know about the new cpo instance *) val cpo_thms' = map (Thm.transfer thy2) cpo_thms; fun make thm = Drule.standard (thm OF cpo_thms'); val ([adm, cont_Rep, cont_Abs, lub, thelub, compact], thy3) = @@ -152,7 +153,7 @@ fun declare_type_name a = Variable.declare_constraints (Logic.mk_type (TFree (a, dummyS))); -fun prepare prep_term _ name (t, vs, mx) raw_set opt_morphs thy = +fun prepare prep_term name (t, vs, mx) raw_set opt_morphs thy = let val _ = Theory.requires thy "Pcpodef" "pcpodefs"; val ctxt = ProofContext.init thy; @@ -167,7 +168,6 @@ (*lhs*) val defS = Sign.defaultS thy; val lhs_tfrees = map (fn v => (v, the_default defS (AList.lookup (op =) rhs_tfrees v))) vs; - val lhs_sorts = map snd lhs_tfrees; val tname = Binding.map_name (Syntax.type_name mx) t; val full_tname = Sign.full_name thy tname; val newT = Type (full_tname, map TFree lhs_tfrees); @@ -175,7 +175,7 @@ val morphs = opt_morphs |> the_default (Binding.prefix_name "Rep_" name, Binding.prefix_name "Abs_" name); in - (newT, oldT, set, morphs, lhs_sorts) + (newT, oldT, set, morphs) end fun add_podef def opt_name typ set opt_morphs tac thy = @@ -188,21 +188,20 @@ val lhs_tfrees = map dest_TFree (snd (dest_Type newT)); val RepC = Const (Rep_name, newT --> oldT); - val below_def = Logic.mk_equals (below_const newT, + val below_eqn = Logic.mk_equals (below_const newT, Abs ("x", newT, Abs ("y", newT, below_const oldT $ (RepC $ Bound 1) $ (RepC $ Bound 0)))); val lthy3 = thy2 |> Theory_Target.instantiation ([full_tname], lhs_tfrees, @{sort po}); - val below_def' = Syntax.check_term lthy3 below_def; - val ((_, (_, below_definition')), lthy4) = lthy3 + val ((_, (_, below_ldef)), lthy4) = lthy3 |> Specification.definition (NONE, - ((Binding.prefix_name "below_" (Binding.suffix_name "_def" name), []), below_def')); + ((Binding.prefix_name "below_" (Binding.suffix_name "_def" name), []), below_eqn)); val ctxt_thy = ProofContext.init (ProofContext.theory_of lthy4); - val below_definition = singleton (ProofContext.export lthy4 ctxt_thy) below_definition'; + val below_def = singleton (ProofContext.export lthy4 ctxt_thy) below_ldef; val thy5 = lthy4 |> Class.prove_instantiation_instance - (K (Tactic.rtac (@{thm typedef_po} OF [type_definition, below_definition]) 1)) - |> LocalTheory.exit_global; - in ((info, below_definition), thy5) end; + (K (Tactic.rtac (@{thm typedef_po} OF [type_definition, below_def]) 1)) + |> Local_Theory.exit_global; + in ((info, below_def), thy5) end; fun prepare_cpodef (prep_term: Proof.context -> 'a -> term) @@ -214,8 +213,8 @@ (thy: theory) : term * term * (thm -> thm -> theory -> (Typedef.info * cpo_info) * theory) = let - val (newT, oldT, set, morphs as (Rep_name, Abs_name), lhs_sorts) = - prepare prep_term def name typ raw_set opt_morphs thy; + val (newT, oldT, set, morphs as (Rep_name, Abs_name)) = + prepare prep_term name typ raw_set opt_morphs thy; val goal_nonempty = HOLogic.mk_Trueprop (HOLogic.mk_exists ("x", oldT, HOLogic.mk_mem (Free ("x", oldT), set))); @@ -247,8 +246,8 @@ (thy: theory) : term * term * (thm -> thm -> theory -> (Typedef.info * cpo_info * pcpo_info) * theory) = let - val (newT, oldT, set, morphs as (Rep_name, Abs_name), lhs_sorts) = - prepare prep_term def name typ raw_set opt_morphs thy; + val (newT, oldT, set, morphs as (Rep_name, Abs_name)) = + prepare prep_term name typ raw_set opt_morphs thy; val goal_UU_mem = HOLogic.mk_Trueprop (HOLogic.mk_mem (Const (@{const_name UU}, oldT), set)); diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOLCF/Tools/repdef.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/Tools/repdef.ML Sat Nov 14 09:40:27 2009 +0100 @@ -0,0 +1,181 @@ +(* Title: HOLCF/Tools/repdef.ML + Author: Brian Huffman + +Defining representable domains using algebraic deflations. +*) + +signature REPDEF = +sig + type rep_info = + { emb_def: thm, prj_def: thm, approx_def: thm, REP: thm } + + val add_repdef: bool -> binding option -> binding * string list * mixfix -> + term -> (binding * binding) option -> theory -> + (Typedef.info * Pcpodef.cpo_info * Pcpodef.pcpo_info * rep_info) * theory + + val repdef_cmd: (bool * binding) * (binding * string list * mixfix) * string + * (binding * binding) option -> theory -> theory +end; + +structure Repdef :> REPDEF = +struct + +(** type definitions **) + +type rep_info = + { emb_def: thm, prj_def: thm, approx_def: thm, REP: thm }; + +(* building terms *) + +fun adm_const T = Const (@{const_name adm}, (T --> HOLogic.boolT) --> HOLogic.boolT); +fun mk_adm (x, T, P) = adm_const T $ absfree (x, T, P); + +fun below_const T = Const (@{const_name below}, T --> T --> HOLogic.boolT); + +val natT = @{typ nat}; +val udomT = @{typ udom}; +fun alg_deflT T = Type (@{type_name alg_defl}, [T]); +fun cfunT (T, U) = Type (@{type_name "->"}, [T, U]); +fun emb_const T = Const (@{const_name emb}, cfunT (T, udomT)); +fun prj_const T = Const (@{const_name prj}, cfunT (udomT, T)); +fun approx_const T = Const (@{const_name approx}, natT --> cfunT (T, T)); + +fun LAM_const (T, U) = Const (@{const_name Abs_CFun}, (T --> U) --> cfunT (T, U)); +fun APP_const (T, U) = Const (@{const_name Rep_CFun}, cfunT (T, U) --> (T --> U)); +fun cast_const T = Const (@{const_name cast}, cfunT (alg_deflT T, cfunT (T, T))); +fun mk_cast (t, x) = + APP_const (udomT, udomT) + $ (APP_const (alg_deflT udomT, cfunT (udomT, udomT)) $ cast_const udomT $ t) + $ x; + +(* manipulating theorems *) + +(* proving class instances *) + +fun declare_type_name a = + Variable.declare_constraints (Logic.mk_type (TFree (a, dummyS))); + +fun gen_add_repdef + (prep_term: Proof.context -> 'a -> term) + (def: bool) + (name: binding) + (typ as (t, vs, mx) : binding * string list * mixfix) + (raw_defl: 'a) + (opt_morphs: (binding * binding) option) + (thy: theory) + : (Typedef.info * Pcpodef.cpo_info * Pcpodef.pcpo_info * rep_info) * theory = + let + val _ = Theory.requires thy "Representable" "repdefs"; + val ctxt = ProofContext.init thy; + + (*rhs*) + val defl = prep_term (ctxt |> fold declare_type_name vs) raw_defl; + val deflT = Term.fastype_of defl; + val _ = if deflT = @{typ "udom alg_defl"} then () + else error ("Not type udom alg_defl: " ^ quote (Syntax.string_of_typ ctxt deflT)); + val rhs_tfrees = Term.add_tfrees defl []; + + (*lhs*) + val defS = Sign.defaultS thy; + val lhs_tfrees = map (fn v => (v, the_default defS (AList.lookup (op =) rhs_tfrees v))) vs; + val lhs_sorts = map snd lhs_tfrees; + val tname = Binding.map_name (Syntax.type_name mx) t; + val full_tname = Sign.full_name thy tname; + val newT = Type (full_tname, map TFree lhs_tfrees); + + (*morphisms*) + val morphs = opt_morphs + |> the_default (Binding.prefix_name "Rep_" name, Binding.prefix_name "Abs_" name); + + (*set*) + val in_defl = @{term "in_deflation :: udom => udom alg_defl => bool"}; + val set = HOLogic.Collect_const udomT $ Abs ("x", udomT, in_defl $ Bound 0 $ defl); + + (*pcpodef*) + val tac1 = rtac @{thm CollectI} 1 THEN rtac @{thm bottom_in_deflation} 1; + val tac2 = rtac @{thm adm_mem_Collect_in_deflation} 1; + val ((info, cpo_info, pcpo_info), thy2) = thy + |> Pcpodef.add_pcpodef def (SOME name) typ set (SOME morphs) (tac1, tac2); + + (*definitions*) + val Rep_const = Const (#Rep_name info, newT --> udomT); + val Abs_const = Const (#Abs_name info, udomT --> newT); + val emb_eqn = Logic.mk_equals (emb_const newT, LAM_const (newT, udomT) $ Rep_const); + val prj_eqn = Logic.mk_equals (prj_const newT, LAM_const (udomT, newT) $ + Abs ("x", udomT, Abs_const $ mk_cast (defl, Bound 0))); + val repdef_approx_const = + Const (@{const_name repdef_approx}, (newT --> udomT) --> (udomT --> newT) + --> alg_deflT udomT --> natT --> cfunT (newT, newT)); + val approx_eqn = Logic.mk_equals (approx_const newT, + repdef_approx_const $ Rep_const $ Abs_const $ defl); + + (*instantiate class rep*) + val name_def = Binding.suffix_name "_def" name; + val ([emb_ldef, prj_ldef, approx_ldef], lthy3) = thy2 + |> Theory_Target.instantiation ([full_tname], lhs_tfrees, @{sort rep}) + |> fold_map Specification.definition + [ (NONE, ((Binding.prefix_name "emb_" name_def, []), emb_eqn)) + , (NONE, ((Binding.prefix_name "prj_" name_def, []), prj_eqn)) + , (NONE, ((Binding.prefix_name "approx_" name_def, []), approx_eqn)) ] + |>> map (snd o snd); + val ctxt_thy = ProofContext.init (ProofContext.theory_of lthy3); + val [emb_def, prj_def, approx_def] = + ProofContext.export lthy3 ctxt_thy [emb_ldef, prj_ldef, approx_ldef]; + val typedef_thms = + [#type_definition info, #below_def cpo_info, emb_def, prj_def, approx_def]; + val thy4 = lthy3 + |> Class.prove_instantiation_instance + (K (Tactic.rtac (@{thm typedef_rep_class} OF typedef_thms) 1)) + |> Local_Theory.exit_global; + + (*other theorems*) + val typedef_thms' = map (Thm.transfer thy4) + [#type_definition info, #below_def cpo_info, emb_def, prj_def]; + val ([REP_thm], thy5) = thy4 + |> Sign.add_path (Binding.name_of name) + |> PureThy.add_thms + [((Binding.prefix_name "REP_" name, + Drule.standard (@{thm typedef_REP} OF typedef_thms')), [])] + ||> Sign.parent_path; + + val rep_info = + { emb_def = emb_def, prj_def = prj_def, approx_def = approx_def, REP = REP_thm }; + in + ((info, cpo_info, pcpo_info, rep_info), thy5) + end + handle ERROR msg => + cat_error msg ("The error(s) above occurred in repdef " ^ quote (Binding.str_of name)); + +fun add_repdef def opt_name typ defl opt_morphs thy = + let + val name = the_default (#1 typ) opt_name; + in + gen_add_repdef Syntax.check_term def name typ defl opt_morphs thy + end; + +fun repdef_cmd ((def, name), typ, A, morphs) = + snd o gen_add_repdef Syntax.read_term def name typ A morphs; + +(** outer syntax **) + +local structure P = OuterParse and K = OuterKeyword in + +val repdef_decl = + Scan.optional (P.$$$ "(" |-- + ((P.$$$ "open" >> K false) -- Scan.option P.binding || P.binding >> (fn s => (true, SOME s))) + --| P.$$$ ")") (true, NONE) -- + (P.type_args -- P.binding) -- P.opt_infix -- (P.$$$ "=" |-- P.term) -- + Scan.option (P.$$$ "morphisms" |-- P.!!! (P.binding -- P.binding)); + +fun mk_repdef ((((((def, opt_name), (vs, t)), mx), A), morphs)) = + repdef_cmd + ((def, the_default (Binding.map_name (Syntax.type_name mx) t) opt_name), (t, vs, mx), A, morphs); + +val _ = + OuterSyntax.command "repdef" "HOLCF definition of representable domains" K.thy_goal + (repdef_decl >> + (Toplevel.print oo (Toplevel.theory o mk_repdef))); + +end; + +end; diff -r 3db22a5707f3 -r ac68c3ee4c2e src/HOLCF/ex/Domain_Proofs.thy --- a/src/HOLCF/ex/Domain_Proofs.thy Sat Nov 14 09:31:54 2009 +0100 +++ b/src/HOLCF/ex/Domain_Proofs.thy Sat Nov 14 09:40:27 2009 +0100 @@ -94,13 +94,13 @@ begin definition emb_foo :: "'a foo \ udom" -where "emb_foo = (\ x. Rep_foo x)" +where "emb_foo \ (\ x. Rep_foo x)" definition prj_foo :: "udom \ 'a foo" -where "prj_foo = (\ y. Abs_foo (cast\(foo_typ\REP('a))\y))" +where "prj_foo \ (\ y. Abs_foo (cast\(foo_typ\REP('a))\y))" definition approx_foo :: "nat \ 'a foo \ 'a foo" -where "approx_foo = (\i. prj oo cast\(approx i\(foo_typ\REP('a))) oo emb)" +where "approx_foo \ repdef_approx Rep_foo Abs_foo (foo_typ\REP('a))" instance apply (rule typedef_rep_class) @@ -117,13 +117,13 @@ begin definition emb_bar :: "'a bar \ udom" -where "emb_bar = (\ x. Rep_bar x)" +where "emb_bar \ (\ x. Rep_bar x)" definition prj_bar :: "udom \ 'a bar" -where "prj_bar = (\ y. Abs_bar (cast\(bar_typ\REP('a))\y))" +where "prj_bar \ (\ y. Abs_bar (cast\(bar_typ\REP('a))\y))" definition approx_bar :: "nat \ 'a bar \ 'a bar" -where "approx_bar = (\i. prj oo cast\(approx i\(bar_typ\REP('a))) oo emb)" +where "approx_bar \ repdef_approx Rep_bar Abs_bar (bar_typ\REP('a))" instance apply (rule typedef_rep_class) @@ -140,13 +140,13 @@ begin definition emb_baz :: "'a baz \ udom" -where "emb_baz = (\ x. Rep_baz x)" +where "emb_baz \ (\ x. Rep_baz x)" definition prj_baz :: "udom \ 'a baz" -where "prj_baz = (\ y. Abs_baz (cast\(baz_typ\REP('a))\y))" +where "prj_baz \ (\ y. Abs_baz (cast\(baz_typ\REP('a))\y))" definition approx_baz :: "nat \ 'a baz \ 'a baz" -where "approx_baz = (\i. prj oo cast\(approx i\(baz_typ\REP('a))) oo emb)" +where "approx_baz \ repdef_approx Rep_baz Abs_baz (baz_typ\REP('a))" instance apply (rule typedef_rep_class) diff -r 3db22a5707f3 -r ac68c3ee4c2e src/Pure/Isar/attrib.ML --- a/src/Pure/Isar/attrib.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/Pure/Isar/attrib.ML Sat Nov 14 09:40:27 2009 +0100 @@ -120,7 +120,7 @@ fun attribute thy = attribute_i thy o intern_src thy; -fun eval_thms ctxt args = ProofContext.note_thmss Thm.generatedK +fun eval_thms ctxt args = ProofContext.note_thmss "" [(Thm.empty_binding, args |> map (fn (a, atts) => (ProofContext.get_fact ctxt a, map (attribute (ProofContext.theory_of ctxt)) atts)))] ctxt |> fst |> maps snd; diff -r 3db22a5707f3 -r ac68c3ee4c2e src/Pure/Isar/class.ML --- a/src/Pure/Isar/class.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/Pure/Isar/class.ML Sat Nov 14 09:40:27 2009 +0100 @@ -281,7 +281,7 @@ in thy |> Expression.add_locale bname Binding.empty supexpr elems - |> snd |> LocalTheory.exit_global + |> snd |> Local_Theory.exit_global |> adjungate_axclass bname class base_sort sups supsort supparam_names global_syntax ||> Theory.checkpoint |-> (fn (param_map, params, assm_axiom) => diff -r 3db22a5707f3 -r ac68c3ee4c2e src/Pure/Isar/class_target.ML --- a/src/Pure/Isar/class_target.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/Pure/Isar/class_target.ML Sat Nov 14 09:40:27 2009 +0100 @@ -405,9 +405,9 @@ fun mk_instantiation (arities, params) = Instantiation { arities = arities, params = params }; -fun get_instantiation lthy = case Instantiation.get (LocalTheory.target_of lthy) +fun get_instantiation lthy = case Instantiation.get (Local_Theory.target_of lthy) of Instantiation data => data; -fun map_instantiation f = (LocalTheory.target o Instantiation.map) +fun map_instantiation f = (Local_Theory.target o Instantiation.map) (fn Instantiation { arities, params } => mk_instantiation (f (arities, params))); fun the_instantiation lthy = case get_instantiation lthy @@ -526,14 +526,14 @@ fun confirm_declaration b = (map_instantiation o apsnd) (filter_out (fn (_, (c', _)) => c' = Binding.name_of b)) - #> LocalTheory.target synchronize_inst_syntax + #> Local_Theory.target synchronize_inst_syntax fun gen_instantiation_instance do_proof after_qed lthy = let val (tycos, vs, sort) = (#arities o the_instantiation) lthy; val arities_proof = maps (fn tyco => Logic.mk_arities (tyco, map snd vs, sort)) tycos; fun after_qed' results = - LocalTheory.theory (fold (AxClass.add_arity o Thm.varifyT) results) + Local_Theory.theory (fold (AxClass.add_arity o Thm.varifyT) results) #> after_qed; in lthy @@ -548,7 +548,7 @@ (fn {context, ...} => tac context)) ts) lthy) I; fun prove_instantiation_exit tac = prove_instantiation_instance tac - #> LocalTheory.exit_global; + #> Local_Theory.exit_global; fun prove_instantiation_exit_result f tac x lthy = let diff -r 3db22a5707f3 -r ac68c3ee4c2e src/Pure/Isar/expression.ML --- a/src/Pure/Isar/expression.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/Pure/Isar/expression.ML Sat Nov 14 09:40:27 2009 +0100 @@ -775,7 +775,7 @@ |> Locale.register_locale binding (extraTs, params) (asm, rev defs) (a_intro, b_intro) axioms ([], []) (rev notes) (rev deps') |> Theory_Target.init (SOME name) - |> fold (fn (kind, facts) => LocalTheory.notes kind facts #> snd) notes'; + |> fold (fn (kind, facts) => Local_Theory.notes_kind kind facts #> snd) notes'; in (name, loc_ctxt) end; diff -r 3db22a5707f3 -r ac68c3ee4c2e src/Pure/Isar/isar_cmd.ML --- a/src/Pure/Isar/isar_cmd.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/Pure/Isar/isar_cmd.ML Sat Nov 14 09:40:27 2009 +0100 @@ -181,7 +181,7 @@ fun declaration pervasive (txt, pos) = txt |> ML_Context.expression pos "val declaration: Morphism.declaration" - ("Context.map_proof (LocalTheory.declaration " ^ Bool.toString pervasive ^ " declaration)") + ("Context.map_proof (Local_Theory.declaration " ^ Bool.toString pervasive ^ " declaration)") |> Context.proof_map; diff -r 3db22a5707f3 -r ac68c3ee4c2e src/Pure/Isar/isar_syn.ML --- a/src/Pure/Isar/isar_syn.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/Pure/Isar/isar_syn.ML Sat Nov 14 09:40:27 2009 +0100 @@ -288,7 +288,7 @@ (* use ML text *) fun propagate_env (context as Context.Proof lthy) = - Context.Proof (LocalTheory.map_contexts (ML_Env.inherit context) lthy) + Context.Proof (Local_Theory.map_contexts (ML_Env.inherit context) lthy) | propagate_env context = context; fun propagate_env_prf prf = Proof.map_contexts diff -r 3db22a5707f3 -r ac68c3ee4c2e src/Pure/Isar/local_theory.ML --- a/src/Pure/Isar/local_theory.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/Pure/Isar/local_theory.ML Sat Nov 14 09:40:27 2009 +0100 @@ -33,8 +33,10 @@ (term * term) * local_theory val define: string -> (binding * mixfix) * (Attrib.binding * term) -> local_theory -> (term * (string * thm)) * local_theory - val note: string -> Attrib.binding * thm list -> local_theory -> (string * thm list) * local_theory - val notes: string -> (Attrib.binding * (thm list * Attrib.src list) list) list -> + val note: Attrib.binding * thm list -> local_theory -> (string * thm list) * local_theory + val notes: (Attrib.binding * (thm list * Attrib.src list) list) list -> + local_theory -> (string * thm list) list * local_theory + val notes_kind: string -> (Attrib.binding * (thm list * Attrib.src list) list) list -> local_theory -> (string * thm list) list * local_theory val type_syntax: bool -> declaration -> local_theory -> local_theory val term_syntax: bool -> declaration -> local_theory -> local_theory @@ -49,7 +51,7 @@ val exit_result_global: (morphism -> 'a -> 'b) -> 'a * local_theory -> 'b * theory end; -structure LocalTheory: LOCAL_THEORY = +structure Local_Theory: LOCAL_THEORY = struct (** local theory data **) @@ -186,12 +188,13 @@ val pretty = operation #pretty; val abbrev = apsnd checkpoint ooo operation2 #abbrev; val define = apsnd checkpoint ooo operation2 #define; -val notes = apsnd checkpoint ooo operation2 #notes; +val notes_kind = apsnd checkpoint ooo operation2 #notes; val type_syntax = checkpoint ooo operation2 #type_syntax; val term_syntax = checkpoint ooo operation2 #term_syntax; val declaration = checkpoint ooo operation2 #declaration; -fun note kind (a, ths) = notes kind [(a, [(ths, [])])] #>> the_single; +val notes = notes_kind ""; +fun note (a, ths) = notes [(a, [(ths, [])])] #>> the_single; fun notation add mode raw_args lthy = let val args = map (apfst (Morphism.term (target_morphism lthy))) raw_args diff -r 3db22a5707f3 -r ac68c3ee4c2e src/Pure/Isar/overloading.ML --- a/src/Pure/Isar/overloading.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/Pure/Isar/overloading.ML Sat Nov 14 09:40:27 2009 +0100 @@ -126,8 +126,8 @@ fun init _ = []; ); -val get_overloading = OverloadingData.get o LocalTheory.target_of; -val map_overloading = LocalTheory.target o OverloadingData.map; +val get_overloading = OverloadingData.get o Local_Theory.target_of; +val map_overloading = Local_Theory.target o OverloadingData.map; fun operation lthy b = get_overloading lthy |> get_first (fn ((c, _), (v, checked)) => @@ -169,7 +169,7 @@ (b, Logic.mk_equals (Const (c, Term.fastype_of t), t)); fun confirm b = map_overloading (filter_out (fn (_, (c', _)) => c' = Binding.name_of b)) - #> LocalTheory.target synchronize_syntax + #> Local_Theory.target synchronize_syntax fun conclude lthy = let diff -r 3db22a5707f3 -r ac68c3ee4c2e src/Pure/Isar/spec_rules.ML --- a/src/Pure/Isar/spec_rules.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/Pure/Isar/spec_rules.ML Sat Nov 14 09:40:27 2009 +0100 @@ -41,7 +41,7 @@ val get = Item_Net.content o Rules.get o Context.Proof; val get_global = Item_Net.content o Rules.get o Context.Theory; -fun add class (ts, ths) = LocalTheory.declaration true +fun add class (ts, ths) = Local_Theory.declaration true (fn phi => let val ts' = map (Morphism.term phi) ts; diff -r 3db22a5707f3 -r ac68c3ee4c2e src/Pure/Isar/specification.ML --- a/src/Pure/Isar/specification.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/Pure/Isar/specification.ML Sat Nov 14 09:40:27 2009 +0100 @@ -202,18 +202,18 @@ in (b, mx) end); val name = Thm.def_binding_optional b raw_name; val ((lhs, (_, raw_th)), lthy2) = lthy - |> LocalTheory.define Thm.definitionK + |> Local_Theory.define Thm.definitionK (var, ((Binding.suffix_name "_raw" name, []), rhs)); val th = prove lthy2 raw_th; val lthy3 = lthy2 |> Spec_Rules.add Spec_Rules.Equational ([lhs], [th]); - val ((def_name, [th']), lthy4) = lthy3 - |> LocalTheory.note Thm.definitionK - ((name, Predicate_Compile_Preproc_Const_Defs.add_attrib :: - Code.add_default_eqn_attrib :: atts), [th]); + val ([(def_name, [th'])], lthy4) = lthy3 + |> Local_Theory.notes_kind Thm.definitionK + [((name, Predicate_Compile_Preproc_Const_Defs.add_attrib :: + Code.add_default_eqn_attrib :: atts), [([th], [])])]; - val lhs' = Morphism.term (LocalTheory.target_morphism lthy4) lhs; + val lhs' = Morphism.term (Local_Theory.target_morphism lthy4) lhs; val _ = if not do_print then () else print_consts lthy4 (member (op =) (Term.add_frees lhs' [])) [(x, T)]; @@ -243,7 +243,7 @@ in (b, mx) end); val lthy' = lthy |> ProofContext.set_syntax_mode mode (* FIXME ?!? *) - |> LocalTheory.abbrev mode (var, rhs) |> snd + |> Local_Theory.abbrev mode (var, rhs) |> snd |> ProofContext.restore_syntax_mode lthy; val _ = if not do_print then () else print_consts lthy' (K false) [(x, T)]; @@ -256,7 +256,7 @@ (* notation *) fun gen_notation prep_const add mode args lthy = - lthy |> LocalTheory.notation add mode (map (apfst (prep_const lthy)) args); + lthy |> Local_Theory.notation add mode (map (apfst (prep_const lthy)) args); val notation = gen_notation (K I); val notation_cmd = gen_notation ProofContext.read_const; @@ -270,7 +270,7 @@ val facts = raw_facts |> map (fn ((name, atts), bs) => ((name, map attrib atts), bs |> map (fn (b, more_atts) => (prep_fact lthy b, map attrib more_atts)))); - val (res, lthy') = lthy |> LocalTheory.notes kind facts; + val (res, lthy') = lthy |> Local_Theory.notes_kind kind facts; val _ = Proof_Display.print_results true lthy' ((kind, ""), res); in (res, lthy') end; @@ -345,7 +345,7 @@ fun gen_theorem prep_att prep_stmt kind before_qed after_qed (name, raw_atts) raw_elems raw_concl int lthy = let - val _ = LocalTheory.affirm lthy; + val _ = Local_Theory.affirm lthy; val thy = ProofContext.theory_of lthy; val attrib = prep_att thy; @@ -359,14 +359,15 @@ burrow (map Goal.norm_result o ProofContext.export goal_ctxt' lthy) results in lthy - |> LocalTheory.notes kind (map2 (fn (a, _) => fn ths => (a, [(ths, [])])) stmt results') + |> Local_Theory.notes_kind kind + (map2 (fn (a, _) => fn ths => (a, [(ths, [])])) stmt results') |> (fn (res, lthy') => if Binding.is_empty name andalso null atts then (Proof_Display.print_results true lthy' ((kind, ""), res); lthy') else let val ([(res_name, _)], lthy'') = lthy' - |> LocalTheory.notes kind [((name, atts), [(maps #2 res, [])])]; + |> Local_Theory.notes_kind kind [((name, atts), [(maps #2 res, [])])]; val _ = Proof_Display.print_results true lthy' ((kind, res_name), res); in lthy'' end) |> after_qed results' diff -r 3db22a5707f3 -r ac68c3ee4c2e src/Pure/Isar/theory_target.ML --- a/src/Pure/Isar/theory_target.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/Pure/Isar/theory_target.ML Sat Nov 14 09:40:27 2009 +0100 @@ -77,14 +77,14 @@ fun direct_decl decl = let val decl0 = Morphism.form decl in - LocalTheory.theory (Context.theory_map decl0) #> - LocalTheory.target (Context.proof_map decl0) + Local_Theory.theory (Context.theory_map decl0) #> + Local_Theory.target (Context.proof_map decl0) end; fun target_decl add (Target {target, ...}) pervasive decl lthy = let - val global_decl = Morphism.transform (LocalTheory.global_morphism lthy) decl; - val target_decl = Morphism.transform (LocalTheory.target_morphism lthy) decl; + val global_decl = Morphism.transform (Local_Theory.global_morphism lthy) decl; + val target_decl = Morphism.transform (Local_Theory.target_morphism lthy) decl; in if target = "" then lthy @@ -92,7 +92,7 @@ else lthy |> pervasive ? direct_decl global_decl - |> LocalTheory.target (add target target_decl) + |> Local_Theory.target (add target target_decl) end; in @@ -104,8 +104,8 @@ end; fun class_target (Target {target, ...}) f = - LocalTheory.raw_theory f #> - LocalTheory.target (Class_Target.refresh_syntax target); + Local_Theory.raw_theory f #> + Local_Theory.target (Class_Target.refresh_syntax target); (* notes *) @@ -161,19 +161,19 @@ val thy = ProofContext.theory_of lthy; val facts' = facts |> map (fn (a, bs) => (a, PureThy.burrow_fact (PureThy.name_multi - (LocalTheory.full_name lthy (fst a))) bs)) + (Local_Theory.full_name lthy (fst a))) bs)) |> PureThy.map_facts (import_export_proof lthy); val local_facts = PureThy.map_facts #1 facts' |> Attrib.map_facts (Attrib.attribute_i thy); val target_facts = PureThy.map_facts #1 facts' - |> is_locale ? Element.facts_map (Element.morph_ctxt (LocalTheory.target_morphism lthy)); + |> is_locale ? Element.facts_map (Element.morph_ctxt (Local_Theory.target_morphism lthy)); val global_facts = PureThy.map_facts #2 facts' |> Attrib.map_facts (if is_locale then K I else Attrib.attribute_i thy); in lthy - |> LocalTheory.theory (PureThy.note_thmss kind global_facts #> snd) - |> not is_locale ? LocalTheory.target (ProofContext.note_thmss kind global_facts #> snd) - |> is_locale ? LocalTheory.target (Locale.add_thmss target kind target_facts) + |> Local_Theory.theory (PureThy.note_thmss kind global_facts #> snd) + |> not is_locale ? Local_Theory.target (ProofContext.note_thmss kind global_facts #> snd) + |> is_locale ? Local_Theory.target (Locale.add_thmss target kind target_facts) |> ProofContext.note_thmss kind local_facts end; @@ -212,22 +212,22 @@ fun declare_const (ta as Target {target, is_locale, is_class, ...}) depends ((b, T), mx) lthy = let - val xs = filter depends (#1 (ProofContext.inferred_fixes (LocalTheory.target_of lthy))); + val xs = filter depends (#1 (ProofContext.inferred_fixes (Local_Theory.target_of lthy))); val U = map #2 xs ---> T; val (mx1, mx2, mx3) = fork_mixfix ta mx; val (const, lthy') = lthy |> (case Class_Target.instantiation_param lthy b of SOME c' => if mx3 <> NoSyn then syntax_error c' - else LocalTheory.theory_result (AxClass.declare_overloaded (c', U)) + else Local_Theory.theory_result (AxClass.declare_overloaded (c', U)) ##> Class_Target.confirm_declaration b | NONE => (case Overloading.operation lthy b of SOME (c', _) => if mx3 <> NoSyn then syntax_error c' - else LocalTheory.theory_result (Overloading.declare (c', U)) + else Local_Theory.theory_result (Overloading.declare (c', U)) ##> Overloading.confirm b - | NONE => LocalTheory.theory_result (Sign.declare_const ((b, U), mx3)))); + | NONE => Local_Theory.theory_result (Sign.declare_const ((b, U), mx3)))); val t = Term.list_comb (const, map Free xs); in lthy' @@ -242,7 +242,7 @@ fun abbrev (ta as Target {target, is_locale, is_class, ...}) prmode ((b, mx), t) lthy = let val thy_ctxt = ProofContext.init (ProofContext.theory_of lthy); - val target_ctxt = LocalTheory.target_of lthy; + val target_ctxt = Local_Theory.target_of lthy; val (mx1, mx2, mx3) = fork_mixfix ta mx; val t' = Assumption.export_term lthy target_ctxt t; @@ -253,14 +253,14 @@ in lthy |> (if is_locale then - LocalTheory.theory_result (Sign.add_abbrev PrintMode.internal (b, global_rhs)) + Local_Theory.theory_result (Sign.add_abbrev PrintMode.internal (b, global_rhs)) #-> (fn (lhs, _) => let val lhs' = Term.list_comb (Logic.unvarify lhs, xs) in term_syntax ta false (locale_const ta prmode ((b, mx2), lhs')) #> is_class ? class_target ta (Class_Target.abbrev target prmode ((b, mx1), t')) end) else - LocalTheory.theory + Local_Theory.theory (Sign.add_abbrev (#1 prmode) (b, global_rhs) #-> (fn (lhs, _) => Sign.notation true prmode [(lhs, mx3)]))) |> ProofContext.add_abbrev PrintMode.internal (b, t) |> snd @@ -278,7 +278,7 @@ val name' = Thm.def_binding_optional b name; val (rhs', rhs_conv) = LocalDefs.export_cterm lthy thy_ctxt (Thm.cterm_of thy rhs) |>> Thm.term_of; - val xs = Variable.add_fixed (LocalTheory.target_of lthy) rhs' []; + val xs = Variable.add_fixed (Local_Theory.target_of lthy) rhs' []; val T = Term.fastype_of rhs; (*const*) @@ -287,7 +287,7 @@ (*def*) val (global_def, lthy3) = lthy2 - |> LocalTheory.theory_result + |> Local_Theory.theory_result (case Overloading.operation lthy b of SOME (_, checked) => Overloading.define checked name' (fst (dest_Const lhs'), rhs') | NONE => @@ -310,7 +310,8 @@ local fun init_target _ NONE = global_target - | init_target thy (SOME target) = if Locale.defined thy (Locale.intern thy target) + | init_target thy (SOME target) = + if Locale.defined thy (Locale.intern thy target) then make_target target true (Class_Target.is_class thy target) ([], [], []) [] else error ("No such locale: " ^ quote target); @@ -323,7 +324,7 @@ fun init_lthy (ta as Target {target, instantiation, overloading, ...}) = Data.put ta #> - LocalTheory.init (Long_Name.base_name target) + Local_Theory.init (Long_Name.base_name target) {pretty = pretty ta, abbrev = abbrev ta, define = define ta, @@ -332,7 +333,7 @@ term_syntax = term_syntax ta, declaration = declaration ta, reinit = fn lthy => init_lthy_ctxt ta (ProofContext.theory_of lthy), - exit = LocalTheory.target_of o + exit = Local_Theory.target_of o (if not (null (#1 instantiation)) then Class_Target.conclude_instantiation else if not (null overloading) then Overloading.conclude else I)} diff -r 3db22a5707f3 -r ac68c3ee4c2e src/Pure/Isar/toplevel.ML --- a/src/Pure/Isar/toplevel.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/Pure/Isar/toplevel.ML Sat Nov 14 09:40:27 2009 +0100 @@ -105,16 +105,16 @@ type generic_theory = Context.generic; (*theory or local_theory*) val loc_init = Theory_Target.context; -val loc_exit = LocalTheory.exit_global; +val loc_exit = Local_Theory.exit_global; fun loc_begin loc (Context.Theory thy) = loc_init (the_default "-" loc) thy | loc_begin NONE (Context.Proof lthy) = lthy | loc_begin (SOME loc) (Context.Proof lthy) = loc_init loc (loc_exit lthy); fun loc_finish _ (Context.Theory _) = Context.Theory o loc_exit - | loc_finish NONE (Context.Proof _) = Context.Proof o LocalTheory.restore + | loc_finish NONE (Context.Proof _) = Context.Proof o Local_Theory.restore | loc_finish (SOME _) (Context.Proof lthy) = fn lthy' => - Context.Proof (LocalTheory.reinit (LocalTheory.raw_theory (K (loc_exit lthy')) lthy)); + Context.Proof (Local_Theory.reinit (Local_Theory.raw_theory (K (loc_exit lthy')) lthy)); (* datatype node *) @@ -193,7 +193,7 @@ (* print state *) -val pretty_context = LocalTheory.pretty o Context.cases (Theory_Target.init NONE) I; +val pretty_context = Local_Theory.pretty o Context.cases (Theory_Target.init NONE) I; fun print_state_context state = (case try node_of state of @@ -259,7 +259,7 @@ | stale_error some = some; fun map_theory f (Theory (gthy, ctxt)) = - Theory (Context.mapping f (LocalTheory.raw_theory f) gthy, ctxt) + Theory (Context.mapping f (Local_Theory.raw_theory f) gthy, ctxt) | map_theory _ node = node; in diff -r 3db22a5707f3 -r ac68c3ee4c2e src/Pure/more_thm.ML --- a/src/Pure/more_thm.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/Pure/more_thm.ML Sat Nov 14 09:40:27 2009 +0100 @@ -86,7 +86,6 @@ val put_name_hint: string -> thm -> thm val definitionK: string val theoremK: string - val generatedK : string val lemmaK: string val corollaryK: string val get_kind: thm -> string @@ -413,7 +412,6 @@ val definitionK = "definition"; val theoremK = "theorem"; -val generatedK = "generatedK" val lemmaK = "lemma"; val corollaryK = "corollary"; diff -r 3db22a5707f3 -r ac68c3ee4c2e src/Pure/simplifier.ML --- a/src/Pure/simplifier.ML Sat Nov 14 09:31:54 2009 +0100 +++ b/src/Pure/simplifier.ML Sat Nov 14 09:40:27 2009 +0100 @@ -177,7 +177,7 @@ fun gen_simproc prep {name, lhss, proc, identifier} lthy = let val b = Binding.name name; - val naming = LocalTheory.naming_of lthy; + val naming = Local_Theory.naming_of lthy; val simproc = make_simproc {name = Name_Space.full_name naming b, lhss = @@ -191,7 +191,7 @@ proc = proc, identifier = identifier}; in - lthy |> LocalTheory.declaration false (fn phi => + lthy |> Local_Theory.declaration false (fn phi => let val b' = Morphism.binding phi b; val simproc' = morph_simproc phi simproc; @@ -335,7 +335,8 @@ "declaration of Simplifier rewrite rule" #> Attrib.setup (Binding.name congN) (Attrib.add_del cong_add cong_del) "declaration of Simplifier congruence rule" #> - Attrib.setup (Binding.name "simproc") simproc_att "declaration of simplification procedures" #> + Attrib.setup (Binding.name "simproc") simproc_att + "declaration of simplification procedures" #> Attrib.setup (Binding.name "simplified") simplified "simplified rule"));