merged
authorwenzelm
Fri, 14 Mar 2014 17:32:11 +0100
changeset 56148 d94d6a9178b5
parent 56132 64eeda68e693 (diff)
parent 56147 9589605bcf41 (current diff)
child 56149 ede82d6e0abd
merged
NEWS
etc/isar-keywords.el
--- a/Admin/components/components.sha1	Fri Mar 14 16:54:01 2014 +0100
+++ b/Admin/components/components.sha1	Fri Mar 14 17:32:11 2014 +0100
@@ -80,3 +80,4 @@
 3a8f77822278fe9250890e357248bc678d8fac95  z3-3.2-1.tar.gz
 12ae71acde43bd7bed1e005c43034b208c0cba4c  z3-3.2.tar.gz
 d94a716502c8503d63952bcb4d4176fac8b28704  z3-4.0.tar.gz
+86e721296c400ada440e4a9ce11b9e845eec9e25  z3-4.3.0.tar.gz
--- a/Admin/components/main	Fri Mar 14 16:54:01 2014 +0100
+++ b/Admin/components/main	Fri Mar 14 17:32:11 2014 +0100
@@ -11,5 +11,6 @@
 scala-2.10.3
 spass-3.8ds
 z3-3.2-1
+z3-4.3.0
 xz-java-1.2-1
-ProofGeneral-4.2-1
\ No newline at end of file
+ProofGeneral-4.2-1
--- a/CONTRIBUTORS	Fri Mar 14 16:54:01 2014 +0100
+++ b/CONTRIBUTORS	Fri Mar 14 17:32:11 2014 +0100
@@ -9,6 +9,15 @@
 * March 2014: René Thiemann
   Improved code generation for multisets.
 
+* Fall 2013 and Winter 2014: Lorenz Panny, Dmitriy Traytel, and
+  Jasmin Blanchette, TUM
+  Various improvements to the BNF-based (co)datatype package, including
+  a more polished "primcorec" command, optimizations, and integration in
+  the "HOL" session.
+
+* Winter 2014: Sascha Boehme, QAware GmbH, and Jasmin Blanchette, TUM
+  "SMT2" module and "smt2" proof method, based on SMT-LIB 2 and Z3 4.3.
+
 * January 2014: Lars Hupel, TUM
   An improved, interactive simplifier trace with integration into the
   Isabelle/jEdit Prover IDE.
@@ -42,7 +51,7 @@
 
 * Spring and Summer 2013: Lorenz Panny, Dmitriy Traytel, and
   Jasmin Blanchette, TUM
-  Various improvements to BNF-based (co)datatype package, including
+  Various improvements to the BNF-based (co)datatype package, including
   "primrec_new" and "primcorec" commands and a compatibility layer.
 
 * Spring and Summer 2013: Ondrej Kuncar, TUM
--- a/NEWS	Fri Mar 14 16:54:01 2014 +0100
+++ b/NEWS	Fri Mar 14 17:32:11 2014 +0100
@@ -98,6 +98,14 @@
 
 *** HOL ***
 
+* Simplifier: Enhanced solver of preconditions of rewrite rules
+  can now deal with conjunctions.
+  For help with converting proofs, the old behaviour of the simplifier
+  can be restored like this:  declare/using [[simp_legacy_precond]]
+  This configuration option will disappear again in the future.
+
+* Dropped facts INF_comp, SUP_comp.  INCOMPATIBILITY.
+
 * HOL-Word:
   * Abandoned fact collection "word_arith_alts", which is a
   duplicate of "word_arith_wis".
@@ -154,7 +162,7 @@
     BNF/Equiv_Relations_More.thy
   INCOMPATIBILITY.
 
-* New datatype package:
+* New (co)datatype package:
   * "primcorec" is fully implemented.
   * Renamed commands:
       datatype_new_compat ~> datatype_compat
@@ -215,7 +223,14 @@
 * Theory reorganizations:
   * Big_Operators.thy ~> Groups_Big.thy and Lattices_Big.thy
 
+* SMT module:
+  * A new version of the SMT module, temporarily called "SMT2", uses SMT-LIB 2
+    and supports recent versions of Z3 (e.g., 4.3). The new proof method is
+    called "smt2", and the new Z3 is called "z3_new" in Sledgehammer and
+    elsewhere.
+
 * Sledgehammer:
+  - New prover "z3_new" with support for Isar proofs
   - New option:
       smt_proofs
   - Renamed options:
--- a/etc/isar-keywords.el	Fri Mar 14 16:54:01 2014 +0100
+++ b/etc/isar-keywords.el	Fri Mar 14 17:32:11 2014 +0100
@@ -244,6 +244,7 @@
     "sledgehammer"
     "sledgehammer_params"
     "smt_status"
+    "smt2_status"
     "solve_direct"
     "sorry"
     "spark_end"
@@ -447,6 +448,7 @@
     "refute"
     "sledgehammer"
     "smt_status"
+    "smt2_status"
     "solve_direct"
     "spark_status"
     "term"
--- a/src/Doc/Datatypes/Datatypes.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/Doc/Datatypes/Datatypes.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -481,7 +481,8 @@
 The syntactic entity \synt{target} can be used to specify a local
 context---e.g., @{text "(in linorder)"} \cite{isabelle-isar-ref}.
 
-The optional target is potentially followed by datatype-specific options:
+The optional target is optionally followed by one or both of the following
+options:
 
 \begin{itemize}
 \setlength{\itemsep}{0pt}
@@ -1355,7 +1356,7 @@
   \label{ssec:primrec-command-syntax} *}
 
 
-subsubsection {* \keyw{primrec\_new}
+subsubsection {* \keyw{primrec}
   \label{sssec:primrec-new} *}
 
 text {*
@@ -1364,7 +1365,10 @@
 \end{matharray}
 
 @{rail \<open>
-  @@{command primrec} target? fixes \<newline> @'where' (@{syntax pr_equation} + '|')
+  @@{command primrec} target? @{syntax pr_option}? fixes \<newline>
+  @'where' (@{syntax pr_equation} + '|')
+  ;
+  @{syntax_def pr_option}: '(' 'nonexhaustive' ')'
   ;
   @{syntax_def pr_equation}: thmdecl? prop
 \<close>}
@@ -1380,6 +1384,17 @@
 \synt{thmdecl} denotes an optional name for the formula that follows, and
 \synt{prop} denotes a HOL proposition \cite{isabelle-isar-ref}.
 
+The optional target is optionally followed by the following option:
+
+\begin{itemize}
+\setlength{\itemsep}{0pt}
+
+\item
+The @{text "nonexhaustive"} option indicates that the functions are not
+necessarily specified for all constructors. It can be used to suppress the
+warning that is normally emitted when some constructors are missing.
+\end{itemize}
+
 %%% TODO: elaborate on format of the equations
 %%% TODO: elaborate on mutual and nested-to-mutual
 *}
@@ -1815,8 +1830,8 @@
 
 text {*
 \noindent
-Constructs such as @{text "let"}---@{text "in"}, @{text
-"if"}---@{text "then"}---@{text "else"}, and @{text "case"}---@{text "of"} may
+Constructs such as @{text "let"}--@{text "in"}, @{text
+"if"}--@{text "then"}--@{text "else"}, and @{text "case"}--@{text "of"} may
 appear around constructors that guard corecursive calls:
 *}
 
@@ -2266,7 +2281,8 @@
 \synt{thmdecl} denotes an optional name for the formula that follows, and
 \synt{prop} denotes a HOL proposition \cite{isabelle-isar-ref}.
 
-The optional target is potentially followed by a corecursion-specific option:
+The optional target is optionally followed by one or both of the following
+options:
 
 \begin{itemize}
 \setlength{\itemsep}{0pt}
--- a/src/Doc/ProgProve/Logic.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/Doc/ProgProve/Logic.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -811,7 +811,7 @@
 step': "star' r x y \<Longrightarrow> r y z \<Longrightarrow> star' r x z"
 
 text{*
-The single @{text r} step is performer after rather than before the @{text star'}
+The single @{text r} step is performed after rather than before the @{text star'}
 steps. Prove @{prop "star' r x y \<Longrightarrow> star r x y"} and
 @{prop "star r x y \<Longrightarrow> star r' x y"}. You may need lemmas.
 Note that rule induction fails
@@ -829,7 +829,7 @@
 \begin{exercise}
 A context-free grammar can be seen as an inductive definition where each
 nonterminal $A$ is an inductively defined predicate on lists of terminal
-symbols: $A(w)$ mans that $w$ is in the language generated by $A$.
+symbols: $A(w)$ means that $w$ is in the language generated by $A$.
 For example, the production $S \to a S b$ can be viewed as the implication
 @{prop"S w \<Longrightarrow> S (a # w @ [b])"} where @{text a} and @{text b} are terminal symbols,
 i.e., elements of some alphabet. The alphabet can be defined like this:
--- a/src/Doc/Sledgehammer/document/root.tex	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/Doc/Sledgehammer/document/root.tex	Fri Mar 14 17:32:11 2014 +0100
@@ -154,12 +154,8 @@
 be run locally; in addition, AgsyHOL, E, E-SInE, E-ToFoF, iProver, iProver-Eq,
 LEO-II, Satallax, SNARK, Vampire, and Waldmeister are available remotely via
 System\-On\-TPTP \cite{sutcliffe-2000}. If you want better performance, you
-should at least install E and SPASS locally.
-
-The SMT solvers CVC3, Yices, and Z3 can be run locally, and CVC3 and
-Z3 can be run remotely on a TU M\"unchen server. If you want better performance
-and get the ability to replay proofs that rely on the \emph{smt} proof method
-without an Internet connection, you should at least have Z3 locally installed.
+should at least install E and SPASS locally. The SMT solvers CVC3, Yices, and Z3
+can be run locally.
 
 There are three main ways to install automatic provers on your machine:
 
@@ -171,9 +167,10 @@
 these otherwise remarkable tools.}
 For Z3, you must additionally set the variable
 \texttt{Z3\_NON\_COMMERCIAL} to ``yes'' to confirm that you are a
-noncommercial user, either in the environment in which Isabelle is
-launched or in your
-\texttt{\$ISABELLE\_HOME\_USER/etc/settings} file.
+noncommercial user---either in the environment in which Isabelle is
+launched, in your \texttt{\$ISABELLE\_HOME\_USER/etc/settings} file, or
+via the ``z3\_non\_commercial'' option under ``Plugins > Plugin Options
+> Isabelle > General'' in Isabelle/jEdit.
 
 \item[\labelitemi] Alternatively, you can download the Isabelle-aware CVC3, E,
 SPASS, and Z3 binary packages from \download. Extract the archives, then add a
@@ -231,7 +228,7 @@
 
 To check whether E, SPASS, Vampire, and/or Z3 are successfully installed, try
 out the example in \S\ref{first-steps}. If the remote versions of any of these
-provers is used (identified by the prefix ``\emph{remote\_\/}''), or if the
+provers is used (identified by the prefix ``\textit{remote\_\/}''), or if the
 local versions fail to solve the easy goal presented there, something must be
 wrong with the installation.
 
@@ -297,7 +294,7 @@
 conclusion is a (universally quantified) equation.
 
 For each successful prover, Sledgehammer gives a one-line \textit{metis} or
-\textit{smt} method call. Rough timings are shown in parentheses, indicating how
+\textit{smt2} method call. Rough timings are shown in parentheses, indicating how
 fast the call is. You can click the proof to insert it into the theory text.
 
 In addition, you can ask Sledgehammer for an Isar text proof by enabling the
@@ -308,8 +305,8 @@
 \postw
 
 When Isar proof construction is successful, it can yield proofs that are more
-readable and also faster than the \textit{metis} or \textit{smt} one-line proofs.
-This feature is experimental and is only available for ATPs.
+readable and also faster than the \textit{metis} or \textit{smt2} one-line
+proofs. This feature is experimental and is only available for ATPs.
 
 \section{Hints}
 \label{hints}
@@ -360,7 +357,7 @@
 
 \item[\labelitemi] \textbf{\textit{isar\_proofs}} (\S\ref{output-format}) specifies
 that Isar proofs should be generated, in addition to one-line \textit{metis} or
-\textit{smt} proofs. The length of the Isar proofs can be controlled by setting
+\textit{smt2} proofs. The length of the Isar proofs can be controlled by setting
 \textit{compress\_isar} (\S\ref{output-format}).
 
 \item[\labelitemi] \textbf{\textit{timeout}} (\S\ref{timeouts}) controls the
@@ -466,10 +463,10 @@
 obtain a step-by-step Isar proof. Since the steps are fairly small, \textit{metis}
 and the other Isabelle proof methods are more likely to be able to replay them.
 
-\item[\labelitemi] Try the \textit{smt} proof method instead of \textit{metis}.
+\item[\labelitemi] Try the \textit{smt2} proof method instead of \textit{metis}.
 It is usually stronger, but you need to either have Z3 available to replay the
 proofs, trust the SMT solver, or use certificates. See the documentation in the
-\emph{SMT} theory (\texttt{\$ISABELLE\_HOME/src/HOL/SMT.thy}) for details.
+\textit{SMT2} theory (\texttt{\$ISABELLE\_HOME/src/HOL/SMT2.thy}) for details.
 
 \item[\labelitemi] Try the \textit{blast} or \textit{auto} proof methods, passing
 the necessary facts via \textbf{unfolding}, \textbf{using}, \textit{intro}{:},
@@ -566,7 +563,7 @@
 Automatic provers frequently use many more facts than are necessary.
 Sledgehammer inclues a minimization tool that takes a set of facts returned by a
 given prover and repeatedly calls the same prover, \textit{metis}, or
-\textit{smt} with subsets of those axioms in order to find a minimal set.
+\textit{smt2} with subsets of those axioms in order to find a minimal set.
 Reducing the number of axioms typically improves Metis's speed and success rate,
 while also removing superfluous clutter from the proof scripts.
 
@@ -885,7 +882,7 @@
 environment variable \texttt{SATALLAX\_HOME} to the directory that contains the
 \texttt{satallax} executable. Sledgehammer requires version 2.2 or above.
 
-\item[\labelitemi] \textbf{\textit{smt}:} The \textit{smt} proof method with the
+\item[\labelitemi] \textbf{\textit{smt}:} The \textit{smt2} proof method with the
 current settings (usually:\ Z3 with proof reconstruction) can be used for proof
 search.
 
@@ -913,8 +910,18 @@
 Microsoft Research \cite{z3}. To use Z3, set the environment variable
 \texttt{Z3\_SOLVER} to the complete path of the executable, including the file
 name, and set \texttt{Z3\_NON\_COMMERCIAL} to ``yes'' to confirm that you are a
-noncommercial user. Sledgehammer has been tested with versions 3.0, 3.1, 3.2,
-and 4.0.
+noncommercial user---either in the environment in which Isabelle is
+launched, in your \texttt{\$ISABELLE\_HOME\_USER/etc/settings} file, or
+via the ``z3\_non\_commercial'' option under ``Plugins > Plugin Options
+> Isabelle > General'' in Isabelle/jEdit. Sledgehammer has been tested with
+versions 3.0, 3.1, 3.2, and 4.0.
+
+\item[\labelitemi] \textbf{\textit{z3\_new}:} Newer versions of Z3 (e.g., 4.3)
+are treated as a different prover by Isabelle. To use these, set the environment
+variable \texttt{Z3\_NEW\_SOLVER} to the complete path of the executable,
+including the file name. You also need to set \texttt{Z3\_NON\_COMMERCIAL} to
+``yes'', as described above. Sledgehammer has been tested with versions 4.3.0 and
+4.3.1.
 \end{enum}
 \end{sloppy}
 
@@ -1168,7 +1175,7 @@
 (unreconstructible using \textit{metis}). The type encodings are
 listed below, with an indication of their soundness in parentheses.
 An asterisk (*) indicates that the encoding is slightly incomplete for
-reconstruction with \textit{metis}, unless the \emph{strict} option (described
+reconstruction with \textit{metis}, unless the \textit{strict} option (described
 below) is enabled.
 
 \begin{enum}
@@ -1317,9 +1324,9 @@
 The collection of methods is roughly the same as for the \textbf{try0} command.
 
 \opsmart{smt\_proofs}{no\_smt\_proofs}
-Specifies whether the \textit{smt} proof method should be tried as an
+Specifies whether the \textit{smt2} proof method should be tried as an
 alternative to \textit{metis}.  If the option is set to \textit{smart} (the
-default), the \textit{smt} method is used for one-line proofs but not in Isar
+default), the \textit{smt2} method is used for one-line proofs but not in Isar
 proofs.
 \end{enum}
 
--- a/src/HOL/Auth/Guard/Guard_Yahalom.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Auth/Guard/Guard_Yahalom.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -218,12 +218,12 @@
 apply (rule Guard_extand, simp, blast)
 apply (case_tac "NAa=NB", clarify)
 apply (frule Says_imp_spies)
-apply (frule in_Guard_kparts_Crypt, simp+, blast, simp+)
+apply (frule in_Guard_kparts_Crypt, simp+)
 apply (frule_tac A=A and B=B and NA=NA and NB=NB and C=Ba in ya3_shrK, simp)
 apply (drule ya2'_imp_ya1'_parts, simp, blast, blast)
 apply (case_tac "NBa=NB", clarify)
 apply (frule Says_imp_spies)
-apply (frule in_Guard_kparts_Crypt, simp+, blast, simp+)
+apply (frule in_Guard_kparts_Crypt, simp+)
 apply (frule_tac A=A and B=B and NA=NA and NB=NB and C=Ba in ya3_shrK, simp)
 apply (drule NB_is_uniq_in_ya2', simp+, blast, simp+)
 apply (simp add: No_Nonce, blast)
@@ -239,7 +239,7 @@
 apply (frule_tac A=A and B=B and NA=NA and NB=NB and C=Aa in ya3_shrK, simp)
 apply (frule ya3'_imp_ya2', simp+, blast, clarify)
 apply (frule_tac A=B' in Says_imp_spies)
-apply (rotate_tac -1, frule in_Guard_kparts_Crypt, simp+, blast, simp+)
+apply (rotate_tac -1, frule in_Guard_kparts_Crypt, simp+)
 apply (frule_tac A=A and B=B and NA=NA and NB=NB and C=Ba in ya3_shrK, simp)
 apply (drule NB_is_uniq_in_ya2', simp+, blast, clarify)
 apply (drule ya3'_imp_ya3, simp+)
--- a/src/HOL/Auth/Public.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Auth/Public.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -414,7 +414,7 @@
 (*Tactic for possibility theorems*)
 fun possibility_tac ctxt =
     REPEAT (*omit used_Says so that Nonces start from different traces!*)
-    (ALLGOALS (simp_tac (ctxt delsimps [@{thm used_Says}]))
+    (ALLGOALS (simp_tac (ctxt setSolver safe_solver delsimps [@{thm used_Says}]))
      THEN
      REPEAT_FIRST (eq_assume_tac ORELSE' 
                    resolve_tac [refl, conjI, @{thm Nonce_supply}]))
--- a/src/HOL/Auth/ZhouGollmann.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Auth/ZhouGollmann.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -101,7 +101,7 @@
                      THEN zg.ZG2, THEN zg.Reception [of _ B A],
                      THEN zg.ZG3, THEN zg.Reception [of _ A TTP], 
                      THEN zg.ZG4])
-apply (possibility, auto)
+apply (basic_possibility, auto)
 done
 
 subsection {*Basic Lemmas*}
--- a/src/HOL/BNF_Cardinal_Order_Relation.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/BNF_Cardinal_Order_Relation.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -600,7 +600,7 @@
 shows "|B| \<le>o |B \<times> A|"
 proof(cases "B = {}", simp add: card_of_empty)
   assume *: "B \<noteq> {}"
-  have "fst `(B \<times> A) = B" unfolding image_def using assms by auto
+  have "fst `(B \<times> A) = B" using assms by auto
   thus ?thesis using inj_on_iff_surj[of B "B \<times> A"]
                      card_of_ordLeq[of B "B \<times> A"] * by blast
 qed
@@ -711,7 +711,7 @@
 
 corollary card_of_Sigma_Times:
 "\<forall>i \<in> I. |A i| \<le>o |B| \<Longrightarrow> |SIGMA i : I. A i| \<le>o |I \<times> B|"
-using card_of_Sigma_mono1[of I A "\<lambda>i. B"] .
+by (fact card_of_Sigma_mono1)
 
 lemma card_of_UNION_Sigma:
 "|\<Union>i \<in> I. A i| \<le>o |SIGMA i : I. A i|"
@@ -1652,7 +1652,7 @@
     hence "\<forall> a. \<exists> b. h a = b" unfolding Func_def by auto
     then obtain f where f: "\<forall> a. h a = f a" by blast
     hence "range f \<subseteq> B" using h unfolding Func_def by auto
-    thus "h \<in> (\<lambda>f a. f a) ` {f. range f \<subseteq> B}" using f unfolding image_def by auto
+    thus "h \<in> (\<lambda>f a. f a) ` {f. range f \<subseteq> B}" using f by auto
   qed(unfold Func_def fun_eq_iff, auto)
 qed
 
--- a/src/HOL/BNF_Constructions_on_Wellorders.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/BNF_Constructions_on_Wellorders.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -1653,9 +1653,9 @@
       qed(insert h, unfold Func_def Func_map_def, auto)
     qed
     moreover have "g \<in> Func A2 A1" unfolding g_def apply(rule Func_map[OF h])
-    using j2A2 B1 A2 unfolding j1_def image_def by (fast intro: inv_into_into)+
+    using j2A2 B1 A2 unfolding j1_def by (fast intro: inv_into_into)+
     ultimately show "h \<in> Func_map B2 f1 f2 ` Func A2 A1"
-    unfolding Func_map_def[abs_def] unfolding image_def by auto
+    unfolding Func_map_def[abs_def] by auto
   qed(insert B1 Func_map[OF _ _ A2(2)], auto)
 qed
 
--- a/src/HOL/Bali/Example.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Bali/Example.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -1280,8 +1280,6 @@
 apply   (rule conjI)
 apply    (rule eval_Is (* Init Object *))
 apply    (simp)
-apply    (rule conjI, rule HOL.refl)+
-apply    (rule HOL.refl)
 apply   (simp)
 apply   (rule conjI, rule_tac [2] HOL.refl)
 apply   (rule eval_Is (* Expr *))
--- a/src/HOL/Bali/Trans.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Bali/Trans.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -23,12 +23,7 @@
     | (FVar) accC statDeclC stat a fn where "v={accC,statDeclC,stat}(Lit a)..fn"
     | (AVar) a i where "v=(Lit a).[Lit i]"
   using ground LVar FVar AVar
-  apply (cases v)
-  apply (simp add: groundVar_def)
-  apply (simp add: groundVar_def,blast)
-  apply (simp add: groundVar_def,blast)
-  apply (simp add: groundVar_def)
-  done
+  by (cases v) (auto simp add: groundVar_def)
 
 definition
   groundExprs :: "expr list \<Rightarrow> bool"
--- a/src/HOL/Basic_BNFs.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Basic_BNFs.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -175,7 +175,7 @@
   thus "f \<circ> x = g \<circ> x" by auto
 next
   fix f show "range \<circ> op \<circ> f = op ` f \<circ> range"
-  unfolding image_def comp_def[abs_def] by auto
+    by (auto simp add: fun_eq_iff)
 next
   show "card_order (natLeq +c |UNIV| )" (is "_ (_ +c ?U)")
   apply (rule card_order_csum)
--- a/src/HOL/Complete_Lattices.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Complete_Lattices.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -20,10 +20,6 @@
 definition INFI :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" where
   INF_def: "INFI A f = \<Sqinter>(f ` A)"
 
-lemma INF_comp: -- {* FIXME drop *}
-  "INFI A (g \<circ> f) = INFI (f ` A) g"
-  by (simp add: INF_def image_comp)
-
 lemma INF_image [simp]: "INFI (f`A) g = INFI A (\<lambda>x. g (f x))"
   by (simp add: INF_def image_image)
 
@@ -39,10 +35,6 @@
 definition SUPR :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" where
   SUP_def: "SUPR A f = \<Squnion>(f ` A)"
 
-lemma SUP_comp: -- {* FIXME drop *}
-  "SUPR A (g \<circ> f) = SUPR (f ` A) g"
-  by (simp add: SUP_def image_comp)
-
 lemma SUP_image [simp]: "SUPR (f`A) g = SUPR A (%x. g (f x))"
   by (simp add: SUP_def image_image)
 
@@ -503,6 +495,21 @@
   "(\<Squnion>a\<in>A. f a) \<sqinter> (\<Squnion>b\<in>B. g b) = (\<Squnion>a\<in>A. \<Squnion>b\<in>B. f a \<sqinter> g b)"
   by (subst SUP_commute) (simp add: inf_SUP SUP_inf)
 
+context
+  fixes f :: "'a \<Rightarrow> 'b::complete_lattice"
+  assumes "mono f"
+begin
+
+lemma mono_Inf:
+  shows "f (\<Sqinter>A) \<le> (\<Sqinter>x\<in>A. f x)"
+  using `mono f` by (auto intro: complete_lattice_class.INF_greatest Inf_lower dest: monoD)
+
+lemma mono_Sup:
+  shows "(\<Squnion>x\<in>A. f x) \<le> f (\<Squnion>A)"
+  using `mono f` by (auto intro: complete_lattice_class.SUP_least Sup_upper dest: monoD)
+
+end
+
 end
 
 class complete_boolean_algebra = boolean_algebra + complete_distrib_lattice
--- a/src/HOL/Decision_Procs/Approximation.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Decision_Procs/Approximation.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -2613,7 +2613,7 @@
     and inequality: "u \<le> lx \<and> ux \<le> l'"
     by (cases "approx prec x vs", auto,
       cases "approx prec a vs", auto,
-      cases "approx prec b vs", auto, blast)
+      cases "approx prec b vs", auto)
   from inequality approx[OF AtLeastAtMost.prems(2) l_eq] approx[OF AtLeastAtMost.prems(2) u_eq] approx[OF AtLeastAtMost.prems(2) x_eq]
   show ?case by auto
 qed
@@ -2902,7 +2902,7 @@
                (Mult (Inverse (Num (Float (int k) 0)))
                      (Mult (Add (Var (Suc (Suc 0))) (Minus (Num c)))
                            (Var (Suc 0))))) [Some (l1, u1), Some (l2, u2), vs!x]"
-      by (auto elim!: lift_bin) blast
+      by (auto elim!: lift_bin)
 
     from bnd_c `x < length xs`
     have bnd: "bounded_by (xs[x:=c]) (vs[x:= Some (c,c)])"
--- a/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -473,7 +473,6 @@
       then have ?case using 4
         apply (cases "p +\<^sub>p p' = 0\<^sub>p")
         apply (auto simp add: Let_def)
-        apply blast
         done
     }
     ultimately have ?case by blast
--- a/src/HOL/Fun.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Fun.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -159,8 +159,8 @@
 unfolding inj_on_def by auto
 
 lemma inj_on_strict_subset:
-  "\<lbrakk> inj_on f B; A < B \<rbrakk> \<Longrightarrow> f`A < f`B"
-unfolding inj_on_def unfolding image_def by blast
+  "inj_on f B \<Longrightarrow> A \<subset> B \<Longrightarrow> f ` A \<subset> f ` B"
+  unfolding inj_on_def by blast
 
 lemma inj_comp:
   "inj f \<Longrightarrow> inj g \<Longrightarrow> inj (f \<circ> g)"
@@ -198,16 +198,14 @@
 by (unfold inj_on_def, blast)
 
 lemma inj_on_iff: "[| inj_on f A;  x:A;  y:A |] ==> (f(x)=f(y)) = (x=y)"
-by (blast dest!: inj_onD)
+  by (fact inj_on_eq_iff)
 
 lemma comp_inj_on:
      "[| inj_on f A;  inj_on g (f`A) |] ==> inj_on (g o f) A"
 by (simp add: comp_def inj_on_def)
 
 lemma inj_on_imageI: "inj_on (g o f) A \<Longrightarrow> inj_on g (f ` A)"
-apply(simp add:inj_on_def image_def)
-apply blast
-done
+  by (simp add: inj_on_def) blast
 
 lemma inj_on_image_iff: "\<lbrakk> ALL x:A. ALL y:A. (g(f x) = g(f y)) = (g x = g y);
   inj_on f A \<rbrakk> \<Longrightarrow> inj_on g (f ` A) = inj_on g A"
@@ -368,26 +366,26 @@
     using assms by(auto simp:bij_betw_def)
   let ?P = "%b a. a:A \<and> f a = b" let ?g = "%b. The (?P b)"
   { fix a b assume P: "?P b a"
-    hence ex1: "\<exists>a. ?P b a" using s unfolding image_def by blast
+    hence ex1: "\<exists>a. ?P b a" using s by blast
     hence uex1: "\<exists>!a. ?P b a" by(blast dest:inj_onD[OF i])
     hence " ?g b = a" using the1_equality[OF uex1, OF P] P by simp
   } note g = this
   have "inj_on ?g B"
   proof(rule inj_onI)
     fix x y assume "x:B" "y:B" "?g x = ?g y"
-    from s `x:B` obtain a1 where a1: "?P x a1" unfolding image_def by blast
-    from s `y:B` obtain a2 where a2: "?P y a2" unfolding image_def by blast
+    from s `x:B` obtain a1 where a1: "?P x a1" by blast
+    from s `y:B` obtain a2 where a2: "?P y a2" by blast
     from g[OF a1] a1 g[OF a2] a2 `?g x = ?g y` show "x=y" by simp
   qed
   moreover have "?g ` B = A"
-  proof(auto simp:image_def)
+  proof(auto simp: image_def)
     fix b assume "b:B"
-    with s obtain a where P: "?P b a" unfolding image_def by blast
+    with s obtain a where P: "?P b a" by blast
     thus "?g b \<in> A" using g[OF P] by auto
   next
     fix a assume "a:A"
-    then obtain b where P: "?P b a" using s unfolding image_def by blast
-    then have "b:B" using s unfolding image_def by blast
+    then obtain b where P: "?P b a" using s by blast
+    then have "b:B" using s by blast
     with g[OF P] show "\<exists>b\<in>B. a = ?g b" by blast
   qed
   ultimately show ?thesis by(auto simp:bij_betw_def)
@@ -614,8 +612,9 @@
 lemma fun_upd_twist: "a ~= c ==> (m(a:=b))(c:=d) = (m(c:=d))(a:=b)"
 by (rule ext, auto)
 
-lemma inj_on_fun_updI: "\<lbrakk> inj_on f A; y \<notin> f`A \<rbrakk> \<Longrightarrow> inj_on (f(x:=y)) A"
-by (fastforce simp:inj_on_def image_def)
+lemma inj_on_fun_updI:
+  "inj_on f A \<Longrightarrow> y \<notin> f ` A \<Longrightarrow> inj_on (f(x := y)) A"
+  by (fastforce simp: inj_on_def)
 
 lemma fun_upd_image:
      "f(x:=y) ` A = (if x \<in> A then insert y (f ` (A-{x})) else f ` A)"
@@ -750,7 +749,7 @@
 
 lemma inj_on_the_inv_into:
   "inj_on f A \<Longrightarrow> inj_on (the_inv_into A f) (f ` A)"
-by (auto intro: inj_onI simp: image_def the_inv_into_f_f)
+by (auto intro: inj_onI simp: the_inv_into_f_f)
 
 lemma bij_betw_the_inv_into:
   "bij_betw f A B \<Longrightarrow> bij_betw (the_inv_into A f) B A"
--- a/src/HOL/HOLCF/IOA/meta_theory/Seq.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/HOLCF/IOA/meta_theory/Seq.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -5,7 +5,7 @@
 header {* Partial, Finite and Infinite Sequences (lazy lists), modeled as domain *}
 
 theory Seq
-imports HOLCF
+imports "../../HOLCF"
 begin
 
 default_sort pcpo
--- a/src/HOL/HOLCF/IOA/meta_theory/SimCorrectness.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/HOLCF/IOA/meta_theory/SimCorrectness.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -99,7 +99,6 @@
 apply (erule conjE)
 apply (simp add: Let_def)
 apply (rule_tac x = "ex" in someI)
-apply (erule conjE)
 apply assumption
 done
 
--- a/src/HOL/HOLCF/IOA/meta_theory/TLS.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/HOLCF/IOA/meta_theory/TLS.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -196,10 +196,7 @@
 apply (tactic {* Seq_case_simp_tac @{context} "y" 1 *})
 apply (tactic {* pair_tac @{context} "a" 1 *})
 apply (tactic {* Seq_case_simp_tac @{context} "s" 1 *})
-apply blast
-apply fastforce
 apply (tactic {* pair_tac @{context} "a" 1 *})
- apply fastforce
 done
 
 end
--- a/src/HOL/HOLCF/Universal.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/HOLCF/Universal.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -111,12 +111,11 @@
   "ubasis_until P 0 = 0"
 | "finite S \<Longrightarrow> ubasis_until P (node i a S) =
     (if P (node i a S) then node i a S else ubasis_until P a)"
-    apply clarify
-    apply (rule_tac x=b in node_cases)
-     apply simp
+   apply clarify
+   apply (rule_tac x=b in node_cases)
     apply simp
-    apply fast
    apply simp
+   apply fast
   apply simp
  apply simp
 done
--- a/src/HOL/Hoare_Parallel/RG_Tran.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Hoare_Parallel/RG_Tran.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -196,7 +196,6 @@
 apply(erule_tac x="Some P2" in allE)
 apply(erule allE,erule impE, assumption)
 apply(drule div_seq,simp)
-apply force
 apply clarify
 apply(erule disjE)
  apply clarify
--- a/src/HOL/Imperative_HOL/ex/Linked_Lists.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Imperative_HOL/ex/Linked_Lists.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -918,7 +918,7 @@
   from 3(7)[OF no_inter2] obtain rs where rs_def: "refs_of' h1 r1 rs" by (rule list_of'_refs_of')
   from refs_of'_merge[OF refs_of'_pn qrs_def 3(6) no_inter this] p_in have p_rs: "p \<notin> set rs" by auto
   with 3(7)[OF no_inter2] 3(1-5) 3(8) p_rs rs_def p_stays
-  show ?case by auto
+  show ?case by (auto simp: list_of'_set_ref)
 next
   case (4 x xs' y ys' p q pn qn h1 r1 h')
   from 4(1) obtain prs where prs_def: "refs_of' h p prs" by (rule list_of'_refs_of')
@@ -935,7 +935,7 @@
   from 4(7)[OF no_inter2] obtain rs where rs_def: "refs_of' h1 r1 rs" by (rule list_of'_refs_of')
   from refs_of'_merge[OF prs_def refs_of'_qn 4(6) no_inter this] q_in have q_rs: "q \<notin> set rs" by auto
   with 4(7)[OF no_inter2] 4(1-5) 4(8) q_rs rs_def q_stays
-  show ?case by auto
+  show ?case by (auto simp: list_of'_set_ref)
 qed
 
 section {* Code generation *}
--- a/src/HOL/Isar_Examples/Hoare_Ex.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Isar_Examples/Hoare_Ex.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -59,7 +59,7 @@
   by hoare
 
 lemma "\<turnstile> \<lbrace>True\<rbrace> \<acute>M := a; \<acute>N := b \<lbrace>\<acute>M = a \<and> \<acute>N = b\<rbrace>"
-  by hoare simp
+  by hoare
 
 lemma
   "\<turnstile> \<lbrace>\<acute>M = a \<and> \<acute>N = b\<rbrace>
--- a/src/HOL/Library/Predicate_Compile_Alternative_Defs.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Library/Predicate_Compile_Alternative_Defs.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -204,7 +204,6 @@
     apply auto
     apply (case_tac xc)
     apply auto
-    apply fastforce
     done
 qed
 
--- a/src/HOL/Lifting_Option.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Lifting_Option.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -17,8 +17,8 @@
     | _ \<Rightarrow> False)"
 by (auto split: prod.split option.split)
 
-abbreviation (input) option_pred :: "('a \<Rightarrow> bool) \<Rightarrow> 'a option \<Rightarrow> bool" where
-  "option_pred \<equiv> case_option True"
+abbreviation (input) pred_option :: "('a \<Rightarrow> bool) \<Rightarrow> 'a option \<Rightarrow> bool" where
+  "pred_option \<equiv> case_option True"
 
 lemma rel_option_eq [relator_eq]:
   "rel_option (op =) = (op =)"
@@ -35,7 +35,7 @@
 
 lemma Domainp_option[relator_domain]:
   assumes "Domainp A = P"
-  shows "Domainp (rel_option A) = (option_pred P)"
+  shows "Domainp (rel_option A) = (pred_option P)"
 using assms unfolding Domainp_iff[abs_def] rel_option_iff[abs_def]
 by (auto iff: fun_eq_iff split: option.split)
 
@@ -64,7 +64,7 @@
   unfolding bi_unique_def split_option_all by simp
 
 lemma option_invariant_commute [invariant_commute]:
-  "rel_option (Lifting.invariant P) = Lifting.invariant (option_pred P)"
+  "rel_option (Lifting.invariant P) = Lifting.invariant (pred_option P)"
   by (auto simp add: fun_eq_iff Lifting.invariant_def split_option_all)
 
 subsection {* Quotient theorem for the Lifting package *}
--- a/src/HOL/Lifting_Product.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Lifting_Product.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -10,12 +10,12 @@
 
 subsection {* Relator and predicator properties *}
 
-definition prod_pred :: "('a \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> bool) \<Rightarrow> 'a \<times> 'b \<Rightarrow> bool"
-where "prod_pred R1 R2 = (\<lambda>(a, b). R1 a \<and> R2 b)"
+definition pred_prod :: "('a \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> bool) \<Rightarrow> 'a \<times> 'b \<Rightarrow> bool"
+where "pred_prod R1 R2 = (\<lambda>(a, b). R1 a \<and> R2 b)"
 
-lemma prod_pred_apply [simp]:
-  "prod_pred P1 P2 (a, b) \<longleftrightarrow> P1 a \<and> P2 b"
-  by (simp add: prod_pred_def)
+lemma pred_prod_apply [simp]:
+  "pred_prod P1 P2 (a, b) \<longleftrightarrow> P1 a \<and> P2 b"
+  by (simp add: pred_prod_def)
 
 lemmas rel_prod_eq[relator_eq] = prod.rel_eq
 lemmas rel_prod_mono[relator_mono] = prod.rel_mono
@@ -27,8 +27,8 @@
 lemma Domainp_prod[relator_domain]:
   assumes "Domainp T1 = P1"
   assumes "Domainp T2 = P2"
-  shows "Domainp (rel_prod T1 T2) = (prod_pred P1 P2)"
-using assms unfolding rel_prod_def prod_pred_def by blast
+  shows "Domainp (rel_prod T1 T2) = (pred_prod P1 P2)"
+using assms unfolding rel_prod_def pred_prod_def by blast
 
 lemma left_total_rel_prod [reflexivity_rule]:
   assumes "left_total R1"
@@ -62,8 +62,8 @@
   using assms unfolding bi_unique_def rel_prod_def by auto
 
 lemma prod_invariant_commute [invariant_commute]: 
-  "rel_prod (Lifting.invariant P1) (Lifting.invariant P2) = Lifting.invariant (prod_pred P1 P2)"
-  by (simp add: fun_eq_iff rel_prod_def prod_pred_def Lifting.invariant_def) blast
+  "rel_prod (Lifting.invariant P1) (Lifting.invariant P2) = Lifting.invariant (pred_prod P1 P2)"
+  by (simp add: fun_eq_iff rel_prod_def pred_prod_def Lifting.invariant_def) blast
 
 subsection {* Quotient theorem for the Lifting package *}
 
@@ -109,4 +109,3 @@
 end
 
 end
-
--- a/src/HOL/List.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/List.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -1299,7 +1299,7 @@
   case (snoc a xs)
   show ?case
   proof cases
-    assume "x = a" thus ?case using snoc by (metis set_simps(1) emptyE)
+    assume "x = a" thus ?case using snoc by (auto intro!: exI)
   next
     assume "x \<noteq> a" thus ?case using snoc by fastforce
   qed
@@ -1332,7 +1332,8 @@
   show ?case
   proof cases
     assume "P x"
-    thus ?thesis by simp (metis Un_upper1 contra_subsetD in_set_conv_decomp_first self_append_conv2 set_append)
+    hence "x # xs = [] @ x # xs \<and> P x \<and> (\<forall>y\<in>set []. \<not> P y)" by simp
+    thus ?thesis by fast
   next
     assume "\<not> P x"
     hence "\<exists>x\<in>set xs. P x" using Cons(2) by simp
@@ -1359,7 +1360,7 @@
   case (snoc x xs)
   show ?case
   proof cases
-    assume "P x" thus ?thesis by (metis emptyE set_empty)
+    assume "P x" thus ?thesis by (auto intro!: exI)
   next
     assume "\<not> P x"
     hence "\<exists>x\<in>set xs. P x" using snoc(2) by simp
@@ -1375,7 +1376,8 @@
 lemma split_list_last_prop_iff:
   "(\<exists>x \<in> set xs. P x) \<longleftrightarrow>
    (\<exists>ys x zs. xs = ys@x#zs \<and> P x \<and> (\<forall>z \<in> set zs. \<not> P z))"
-by (metis split_list_last_prop [where P=P] in_set_conv_decomp)
+  by rule (erule split_list_last_prop, auto)
+
 
 lemma finite_list: "finite A ==> EX xs. set xs = A"
   by (erule finite_induct) (auto simp add: set_simps(2) [symmetric] simp del: set_simps(2))
@@ -1773,7 +1775,7 @@
 done
 
 lemma list_update_nonempty[simp]: "xs[k:=x] = [] \<longleftrightarrow> xs=[]"
-by(metis length_0_conv length_list_update)
+by (simp only: length_0_conv[symmetric] length_list_update)
 
 lemma list_update_same_conv:
 "i < length xs ==> (xs[i := x] = xs) = (xs!i = x)"
@@ -1936,7 +1938,7 @@
 
 lemma snoc_eq_iff_butlast:
   "xs @ [x] = ys \<longleftrightarrow> (ys \<noteq> [] & butlast ys = xs & last ys = x)"
-by (metis append_butlast_last_id append_is_Nil_conv butlast_snoc last_snoc not_Cons_self)
+by fastforce
 
 
 subsubsection {* @{const take} and @{const drop} *}
@@ -2121,8 +2123,7 @@
   "m >= n \<Longrightarrow> set(drop m xs) <= set(drop n xs)"
 apply(induct xs arbitrary: m n)
 apply(auto simp:drop_Cons split:nat.split)
-apply (metis set_drop_subset subset_iff)
-done
+by (metis set_drop_subset subset_iff)
 
 lemma in_set_takeD: "x : set(take n xs) \<Longrightarrow> x : set xs"
 using set_take_subset by fast
@@ -3250,15 +3251,9 @@
  apply (case_tac j)
 apply (clarsimp simp add: set_conv_nth, simp)
 apply (rule conjI)
-(*TOO SLOW
-apply (metis Zero_neq_Suc gr0_conv_Suc in_set_conv_nth lessI less_trans_Suc nth_Cons' nth_Cons_Suc)
-*)
  apply (clarsimp simp add: set_conv_nth)
  apply (erule_tac x = 0 in allE, simp)
  apply (erule_tac x = "Suc i" in allE, simp, clarsimp)
-(*TOO SLOW
-apply (metis Suc_Suc_eq lessI less_trans_Suc nth_Cons_Suc)
-*)
 apply (erule_tac x = "Suc i" in allE, simp)
 apply (erule_tac x = "Suc j" in allE, simp)
 done
@@ -3403,8 +3398,7 @@
 
 lemma distinct_length_2_or_more:
 "distinct (a # b # xs) \<longleftrightarrow> (a \<noteq> b \<and> distinct (a # xs) \<and> distinct (b # xs))"
-by (metis distinct.simps(2) list.sel(1) hd_in_set list.simps(2) set_ConsD set_rev_mp
-      set_subset_Cons)
+by force
 
 lemma remdups_adj_Cons: "remdups_adj (x # xs) =
   (case remdups_adj xs of [] \<Rightarrow> [x] | y # xs \<Rightarrow> if x = y then y # xs else x # y # xs)"
@@ -3636,8 +3630,8 @@
   case Nil thus ?case by simp
 next
   case (Cons x xs) thus ?case
-    by(auto simp: nth_Cons' split: if_splits)
-      (metis One_nat_def diff_Suc_1 less_Suc_eq_0_disj)
+    apply(auto simp: nth_Cons' split: if_splits)
+    using diff_Suc_1[unfolded One_nat_def] less_Suc_eq_0_disj by fastforce
 qed
 
 lemma find_cong[fundef_cong]:
@@ -3683,8 +3677,8 @@
    (case List.extract P xs of
       None \<Rightarrow> None |
       Some (ys, y, zs) \<Rightarrow> Some (x#ys, y, zs)))"
-by(auto simp add: extract_def split: list.splits)
-  (metis comp_def dropWhile_eq_Nil_conv list.distinct(1))
+by(auto simp add: extract_def comp_def split: list.splits)
+  (metis dropWhile_eq_Nil_conv list.distinct(1))
 
 
 subsubsection {* @{const remove1} *}
@@ -3792,7 +3786,7 @@
 
 lemma map_removeAll_inj: "inj f \<Longrightarrow>
   map f (removeAll x xs) = removeAll (f x) (map f xs)"
-by(metis map_removeAll_inj_on subset_inj_on subset_UNIV)
+by (rule map_removeAll_inj_on, erule subset_inj_on, rule subset_UNIV)
 
 
 subsubsection {* @{const replicate} *}
@@ -3962,7 +3956,7 @@
     with * show ?thesis by blast
   qed
   then show ?case
-    using xs'_def ys'_def by metis
+    using xs'_def ys'_def by meson
 qed
 
 lemma comm_append_is_replicate:
@@ -3974,7 +3968,7 @@
 proof -
   obtain m n zs where "concat (replicate m zs) = xs"
     and "concat (replicate n zs) = ys"
-    using assms by (metis comm_append_are_replicate)
+    using comm_append_are_replicate[of xs ys, OF assms] by blast
   then have "m + n > 1" and "concat (replicate (m+n) zs) = xs @ ys"
     using `xs \<noteq> []` and `ys \<noteq> []`
     by (auto simp: replicate_add)
@@ -4511,10 +4505,11 @@
 qed
 
 lemma infinite_UNIV_listI: "~ finite(UNIV::'a list set)"
-apply(rule notI)
-apply(drule finite_maxlen)
-apply (metis UNIV_I length_replicate less_not_refl)
-done
+apply (rule notI)
+apply (drule finite_maxlen)
+apply clarsimp
+apply (erule_tac x = "replicate n undefined" in allE)
+by simp
 
 
 subsection {* Sorting *}
@@ -4726,7 +4721,7 @@
   proof(induct rule:list_induct2[OF 1])
     case 1 show ?case by simp
   next
-    case 2 thus ?case by (simp add:sorted_Cons)
+    case 2 thus ?case by (simp add: sorted_Cons)
        (metis Diff_insert_absorb antisym insertE insert_iff)
   qed
 qed
@@ -5660,10 +5655,10 @@
 by (simp add: listrel1_def Cons_eq_append_conv) (blast)
 
 lemma listrel1I1: "(x,y) \<in> r \<Longrightarrow> (x # xs, y # xs) \<in> listrel1 r"
-by (metis Cons_listrel1_Cons)
+by fast
 
 lemma listrel1I2: "(xs, ys) \<in> listrel1 r \<Longrightarrow> (x # xs, x # ys) \<in> listrel1 r"
-by (metis Cons_listrel1_Cons)
+by fast
 
 lemma append_listrel1I:
   "(xs, ys) \<in> listrel1 r \<and> us = vs \<or> xs = ys \<and> (us, vs) \<in> listrel1 r
@@ -5757,8 +5752,8 @@
 done
 
 lemma wf_listrel1_iff[simp]: "wf(listrel1 r) = wf r"
-by(metis wf_acc_iff in_lists_conv_set lists_accI lists_accD Cons_in_lists_iff)
-
+by (auto simp: wf_acc_iff
+      intro: lists_accD lists_accI[THEN Cons_in_lists_iff[THEN iffD1, THEN conjunct1]])
 
 subsubsection {* Lifting Relations to Lists: all elements *}
 
@@ -5901,7 +5896,7 @@
       case base show ?case by(auto simp add: listrel_iff_zip set_zip)
     next
       case (step ys zs)
-      thus ?case  by (metis listrel_reflcl_if_listrel1 listrel_rtrancl_trans)
+      thus ?case by (metis listrel_reflcl_if_listrel1 listrel_rtrancl_trans)
     qed
   qed
 qed
--- a/src/HOL/MicroJava/BV/BVNoTypeError.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/MicroJava/BV/BVNoTypeError.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -236,7 +236,7 @@
     then obtain stk loc C sig pc frs' where
       xcp [simp]: "xcp = None" and
       frs [simp]: "frs = (stk,loc,C,sig,pc)#frs'" 
-      by (clarsimp simp add: neq_Nil_conv) fast
+      by (clarsimp simp add: neq_Nil_conv)
 
     from conforms obtain  ST LT rT maxs maxl ins et where
       hconf:  "G \<turnstile>h hp \<surd>" and
@@ -245,7 +245,7 @@
       phi:    "Phi C sig ! pc = Some (ST,LT)" and
       frame:  "correct_frame G hp (ST,LT) maxl ins (stk,loc,C,sig,pc)" and
       frames: "correct_frames G hp Phi rT sig frs'"
-      by (auto simp add: correct_state_def) (rule that)
+      by (auto simp add: correct_state_def)
 
     from frame obtain
       stk: "approx_stk G hp stk ST" and
@@ -354,7 +354,7 @@
         obtain D fs where
           hp: "hp (the_Addr x) = Some (D,fs)" and
           D:  "G \<turnstile> D \<preceq>C C"
-          by - (drule non_npD, assumption, clarsimp, blast)
+          by - (drule non_npD, assumption, clarsimp)
         hence "hp (the_Addr x) \<noteq> None" (is ?P1) by simp
         moreover
         from wf mth hp D
--- a/src/HOL/MicroJava/BV/BVSpecTypeSafe.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/MicroJava/BV/BVSpecTypeSafe.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -889,7 +889,7 @@
     loc:    "hp ref = Some obj" and
     obj_ty: "obj_ty obj = Class D" and
     D:      "G \<turnstile> Class D \<preceq> X"
-    by (auto simp add: conf_def) blast
+    by (auto simp add: conf_def)
   
   with X_Ref obtain X' where X': "X = Class X'"
     by (blast dest: widen_Class)
@@ -1052,7 +1052,7 @@
       meth'': "method (G, D) sig = Some (D'', rT'', body)" and
       ST0':   "ST' = rev apTs @ Class D # ST0'" and
       len':   "length apTs = length pt" 
-      by clarsimp blast    
+      by clarsimp
 
     from f frame'
     obtain
@@ -1074,7 +1074,7 @@
       methD':  "method (G, D') sig = Some (mD, rT0, body')" and
       lessD':  "G \<turnstile> X \<preceq> Class D'" and
       suc_pc': "Suc pc' < length ins'"
-      by (clarsimp simp add: wt_instr_def eff_def norm_eff_def) blast
+      by (clarsimp simp add: wt_instr_def eff_def norm_eff_def)
 
     from len len' ST0 ST0'
     have "X = Class D" by simp
--- a/src/HOL/MicroJava/BV/EffectMono.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/MicroJava/BV/EffectMono.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -161,7 +161,7 @@
         where s2:  "s2 = (vT' # oT' # ST', LT')" and
               oT': "G\<turnstile> oT' \<preceq> oT" and
               vT': "G\<turnstile> vT' \<preceq> vT"
-        by - (cases s2, simp add: sup_state_Cons2, elim exE conjE, simp, rule that)
+        by - (cases s2, simp add: sup_state_Cons2, elim exE conjE, simp)
       moreover
       from vT' vT
       have "G \<turnstile> vT' \<preceq> b" by (rule widen_trans)
--- a/src/HOL/MicroJava/BV/JVMType.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/MicroJava/BV/JVMType.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -258,7 +258,7 @@
     assume l: "length (l#ls) = length b"
     
     then obtain b' bs where b: "b = b'#bs"
-      by (cases b) (simp, simp add: neq_Nil_conv, rule that)
+      by (cases b) (simp, simp add: neq_Nil_conv)
 
     with f
     have "\<forall>n. n < length ls \<longrightarrow> (G \<turnstile> (ls!n) <=o (bs!n))"
--- a/src/HOL/MicroJava/Comp/CorrComp.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/MicroJava/Comp/CorrComp.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -523,7 +523,6 @@
 apply (rule_tac ?instrs0.0 = "[load_default_val b]" in progression_transitive, simp)
 apply (rule progression_one_step)
 apply (simp (no_asm_simp) add: load_default_val_def)
-apply (rule conjI, simp)+ apply (rule HOL.refl)
 
 apply (rule progression_one_step)
 apply (simp (no_asm_simp))
@@ -816,7 +815,7 @@
 apply (rule_tac ?instrs0.0 = "[Dup]" in progression_transitive, simp)
 apply (rule progression_one_step)
 apply (simp add: gh_def)
-apply (rule conjI, simp)+ apply simp
+apply simp
 apply (rule progression_one_step)
 apply (simp add: gh_def)
 (* the following falls out of the general scheme *)
@@ -865,8 +864,6 @@
    (* Dup_x1 *)
    apply (rule progression_one_step)
    apply (simp add: gh_def)
-   apply (rule conjI, simp)+ apply simp
-
 
    (* Putfield \<longrightarrow> still looks nasty*)
    apply (rule progression_one_step)
@@ -948,7 +945,6 @@
 apply simp
 apply (rule_tac ?instrs0.0 = "[LitPush (Bool False)]" in progression_transitive, simp (no_asm_simp))
 apply (rule progression_one_step,  simp)
-apply (rule conjI, rule HOL.refl)+ apply (rule HOL.refl)
 
 apply (rule_tac ?instrs0.0 = "compExpr (gmb G CL S) e" in progression_transitive, rule HOL.refl)
 apply fast
@@ -958,18 +954,17 @@
 apply simp
 apply (rule_tac ?instrs0.0 = "[Ifcmpeq (2 + int (length (compStmt (gmb G CL S) c1)))]" in progression_transitive, simp)
 apply (rule progression_one_step) apply simp
-apply (rule conjI, rule HOL.refl)+ apply (rule HOL.refl)
 apply (rule_tac ?instrs0.0 = "(compStmt (gmb G CL S) c1)" in progression_transitive, simp)
 apply fast
 apply (rule_tac ?instrs1.0 = "[]" in jump_fwd_progression)
-apply (simp, rule conjI, (rule HOL.refl)+)
+apply (simp)
 apply simp apply (rule conjI, simp) apply (simp add: nat_add_distrib)
 apply (rule progression_refl)
 
  (* case b= False *)
 apply simp
 apply (rule_tac ?instrs1.0 = "compStmt (gmb G CL S) c2" in jump_fwd_progression)
-apply (simp, rule conjI, (rule HOL.refl)+)
+apply (simp)
 apply (simp, rule conjI, rule HOL.refl, simp add: nat_add_distrib)
 apply fast
 
@@ -992,12 +987,11 @@
 apply (rule_tac ?instrs0.0 = "[LitPush (Bool False)]" in progression_transitive, simp (no_asm_simp))
 apply (rule progression_one_step)
    apply simp 
-   apply (rule conjI, rule HOL.refl)+ apply (rule HOL.refl)
 
 apply (rule_tac ?instrs0.0 = "compExpr (gmb G CL S) e" in progression_transitive, rule HOL.refl)
 apply fast
 apply (rule_tac ?instrs1.0 = "[]" in jump_fwd_progression)
-apply (simp, rule conjI, rule HOL.refl, rule HOL.refl)
+apply (simp)
 apply (simp, rule conjI, rule HOL.refl, simp add: nat_add_distrib)
 apply (rule progression_refl)
 
@@ -1022,12 +1016,10 @@
 apply simp
 apply (rule jump_bwd_progression) 
 apply simp
-apply (rule conjI, (rule HOL.refl)+)
 
 apply (rule_tac ?instrs0.0 = "[LitPush (Bool False)]" in progression_transitive, simp (no_asm_simp))
 apply (rule progression_one_step)
    apply simp 
-   apply (rule conjI, simp)+ apply simp
 
 apply (rule_tac ?instrs0.0 = "compExpr (gmb G CL S) e" in progression_transitive, rule HOL.refl)
 apply fast
@@ -1038,7 +1030,6 @@
 
 apply (rule_tac ?instrs0.0 = "[Ifcmpeq (2 + int (length (compStmt (gmb G CL S) c)))]" in progression_transitive, simp)
 apply (rule progression_one_step) apply simp
-apply (rule conjI, rule HOL.refl)+ apply (rule HOL.refl)
 apply fast
 
  (* case b= False \<longrightarrow> contradiction*)
@@ -1126,7 +1117,6 @@
 apply (simp (no_asm_use) only: exec_instr.simps)
 apply (erule thin_rl, erule thin_rl, erule thin_rl)
 apply (simp add: compMethod_def raise_system_xcpt_def)
-apply (rule conjI, simp)+ apply (rule HOL.refl)
 
      (* get instructions of invoked method *)
 apply (simp (no_asm_simp) add: gis_def gmb_def compMethod_def)
@@ -1177,7 +1167,6 @@
 apply (frule non_npD) apply assumption
 apply (erule exE)+ apply simp
 apply (rule conf_obj_AddrI) apply simp 
-apply (rule conjI, (rule HOL.refl)+)
 apply (rule widen_Class_Class [THEN iffD1], rule widen.refl)
 
 
--- a/src/HOL/MicroJava/Comp/CorrCompTp.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/MicroJava/Comp/CorrCompTp.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -1268,7 +1268,6 @@
   apply (frule max_spec2mheads, (erule exE)+, (erule conjE)+)
   apply (rule exI)+
   apply (simp add: wf_prog_ws_prog [THEN comp_method])
-  apply auto
   done
 
 
@@ -2517,16 +2516,13 @@
                               (start_ST, start_LT C pTs (length lvars))))
                 = (start_ST, inited_LT C pTs lvars)") 
    prefer 2 apply (rule compTpInitLvars_LT_ST) apply (rule HOL.refl) apply assumption
-apply (simp only:)
 apply (subgoal_tac "(snd (compTpStmt (pns, lvars, blk, res) G blk
                        (start_ST, inited_LT C pTs lvars))) 
                 = (start_ST, inited_LT C pTs lvars)") 
    prefer 2 apply (erule conjE)+
    apply (rule compTpStmt_LT_ST) apply (rule HOL.refl) apply assumption+
    apply (simp only:)+ apply (simp (no_asm_simp) add: is_inited_LT_def)
-apply (simp only:)
-apply (rule compTpExpr_LT_ST) apply (rule HOL.refl) apply assumption+
-   apply (simp only:)+ apply (simp (no_asm_simp) add: is_inited_LT_def)
+   apply (simp (no_asm_simp) add: is_inited_LT_def)
 
 
    (* Return *)
--- a/src/HOL/MicroJava/J/Eval.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/MicroJava/J/Eval.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -226,7 +226,7 @@
   next
     case (Call a b c d e f g h i j k l m n o p q r s t u v s4)
     with Call_code show ?thesis
-      by(cases "s4")(simp, (erule meta_allE meta_impE|rule conjI refl| assumption)+)
+      by(cases "s4") auto
   qed(erule (3) that[OF refl]|assumption)+
 next
   case evals
--- a/src/HOL/MicroJava/J/WellForm.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/MicroJava/J/WellForm.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -565,7 +565,7 @@
 apply clarify
 apply (frule wf_prog_ws_prog)
 apply (frule fields_Object, assumption+)
-apply (simp only: is_class_Object) apply simp
+apply (simp only: is_class_Object)
 
 apply clarify
 apply (frule fields_rec)
--- a/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Multivariate_Analysis/Brouwer_Fixpoint.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -3752,43 +3752,70 @@
   qed
 qed
 
+definition unit_cube :: "'a::euclidean_space set"
+  where "unit_cube = {x. \<forall>i\<in>Basis. 0 \<le> x \<bullet> i \<and> x \<bullet> i \<le> 1}"
+
+lemma mem_unit_cube: "x \<in> unit_cube \<longleftrightarrow> (\<forall>i\<in>Basis. 0 \<le> x \<bullet> i \<and> x \<bullet> i \<le> 1)"
+  unfolding unit_cube_def by simp
+
+lemma bounded_unit_cube: "bounded unit_cube"
+  unfolding bounded_def
+proof (intro exI ballI)
+  fix y :: 'a assume y: "y \<in> unit_cube"
+  have "dist 0 y = norm y" by (rule dist_0_norm)
+  also have "\<dots> = norm (\<Sum>i\<in>Basis. (y \<bullet> i) *\<^sub>R i)" unfolding euclidean_representation ..
+  also have "\<dots> \<le> (\<Sum>i\<in>Basis. norm ((y \<bullet> i) *\<^sub>R i))" by (rule norm_setsum)
+  also have "\<dots> \<le> (\<Sum>i::'a\<in>Basis. 1)"
+    by (rule setsum_mono, simp add: y [unfolded mem_unit_cube])
+  finally show "dist 0 y \<le> (\<Sum>i::'a\<in>Basis. 1)" .
+qed
+
+lemma closed_unit_cube: "closed unit_cube"
+  unfolding unit_cube_def Collect_ball_eq Collect_conj_eq
+  by (rule closed_INT, auto intro!: closed_Collect_le)
+
+lemma compact_unit_cube: "compact unit_cube" (is "compact ?C")
+  unfolding compact_eq_seq_compact_metric
+  using bounded_unit_cube closed_unit_cube
+  by (rule bounded_closed_imp_seq_compact)
+
 lemma brouwer_cube:
-  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> 'a::ordered_euclidean_space"
-  assumes "continuous_on {0..(\<Sum>Basis)} f"
-    and "f ` {0..(\<Sum>Basis)} \<subseteq> {0..(\<Sum>Basis)}"
-  shows "\<exists>x\<in>{0..(\<Sum>Basis)}. f x = x"
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'a"
+  assumes "continuous_on unit_cube f"
+    and "f ` unit_cube \<subseteq> unit_cube"
+  shows "\<exists>x\<in>unit_cube. f x = x"
 proof (rule ccontr)
   def n \<equiv> "DIM('a)"
   have n: "1 \<le> n" "0 < n" "n \<noteq> 0"
     unfolding n_def by (auto simp add: Suc_le_eq DIM_positive)
   assume "\<not> ?thesis"
-  then have *: "\<not> (\<exists>x\<in>{0..\<Sum>Basis}. f x - x = 0)"
+  then have *: "\<not> (\<exists>x\<in>unit_cube. f x - x = 0)"
     by auto
   obtain d where
-      d: "d > 0" "\<And>x. x \<in> {0..\<Sum>Basis} \<Longrightarrow> d \<le> norm (f x - x)"
-    apply (rule brouwer_compactness_lemma[OF compact_interval _ *])
+      d: "d > 0" "\<And>x. x \<in> unit_cube \<Longrightarrow> d \<le> norm (f x - x)"
+    apply (rule brouwer_compactness_lemma[OF compact_unit_cube _ *])
     apply (rule continuous_on_intros assms)+
     apply blast
     done
-  have *: "\<forall>x. x \<in> {0..\<Sum>Basis} \<longrightarrow> f x \<in> {0..\<Sum>Basis}"
-    "\<forall>x. x \<in> {0..(\<Sum>Basis)::'a} \<longrightarrow> (\<forall>i\<in>Basis. True \<longrightarrow> 0 \<le> x \<bullet> i \<and> x \<bullet> i \<le> 1)"
+  have *: "\<forall>x. x \<in> unit_cube \<longrightarrow> f x \<in> unit_cube"
+    "\<forall>x. x \<in> (unit_cube::'a set) \<longrightarrow> (\<forall>i\<in>Basis. True \<longrightarrow> 0 \<le> x \<bullet> i \<and> x \<bullet> i \<le> 1)"
     using assms(2)[unfolded image_subset_iff Ball_def]
-    unfolding mem_interval
+    unfolding mem_unit_cube
     by auto
   obtain label :: "'a \<Rightarrow> 'a \<Rightarrow> nat" where
     "\<forall>x. \<forall>i\<in>Basis. label x i \<le> 1"
-    "\<forall>x. \<forall>i\<in>Basis. x \<in> {0..\<Sum>Basis} \<and> True \<and> x \<bullet> i = 0 \<longrightarrow> label x i = 0"
-    "\<forall>x. \<forall>i\<in>Basis. x \<in> {0..\<Sum>Basis} \<and> True \<and> x \<bullet> i = 1 \<longrightarrow> label x i = 1"
-    "\<forall>x. \<forall>i\<in>Basis. x \<in> {0..\<Sum>Basis} \<and> True \<and> label x i = 0 \<longrightarrow> x \<bullet> i \<le> f x \<bullet> i"
-    "\<forall>x. \<forall>i\<in>Basis. x \<in> {0..\<Sum>Basis} \<and> True \<and> label x i = 1 \<longrightarrow> f x \<bullet> i \<le> x \<bullet> i"
+    "\<forall>x. \<forall>i\<in>Basis. x \<in> unit_cube \<and> True \<and> x \<bullet> i = 0 \<longrightarrow> label x i = 0"
+    "\<forall>x. \<forall>i\<in>Basis. x \<in> unit_cube \<and> True \<and> x \<bullet> i = 1 \<longrightarrow> label x i = 1"
+    "\<forall>x. \<forall>i\<in>Basis. x \<in> unit_cube \<and> True \<and> label x i = 0 \<longrightarrow> x \<bullet> i \<le> f x \<bullet> i"
+    "\<forall>x. \<forall>i\<in>Basis. x \<in> unit_cube \<and> True \<and> label x i = 1 \<longrightarrow> f x \<bullet> i \<le> x \<bullet> i"
     using kuhn_labelling_lemma[OF *] by blast
   note label = this [rule_format]
-  have lem1: "\<forall>x\<in>{0..\<Sum>Basis}.\<forall>y\<in>{0..\<Sum>Basis}.\<forall>i\<in>Basis. label x i \<noteq> label y i \<longrightarrow>
+  have lem1: "\<forall>x\<in>unit_cube. \<forall>y\<in>unit_cube. \<forall>i\<in>Basis. label x i \<noteq> label y i \<longrightarrow>
     abs (f x \<bullet> i - x \<bullet> i) \<le> norm (f y - f x) + norm (y - x)"
   proof safe
     fix x y :: 'a
-    assume x: "x \<in> {0..\<Sum>Basis}"
-    assume y: "y \<in> {0..\<Sum>Basis}"
+    assume x: "x \<in> unit_cube"
+    assume y: "y \<in> unit_cube"
     fix i
     assume i: "label x i \<noteq> label y i" "i \<in> Basis"
     have *: "\<And>x y fx fy :: real. x \<le> fx \<and> fy \<le> y \<or> fx \<le> x \<and> y \<le> fy \<Longrightarrow>
@@ -3840,7 +3867,7 @@
     finally show "\<bar>f x \<bullet> i - x \<bullet> i\<bar> \<le> norm (f y - f x) + norm (y - x)"
       unfolding inner_simps .
   qed
-  have "\<exists>e>0. \<forall>x\<in>{0..\<Sum>Basis}. \<forall>y\<in>{0..\<Sum>Basis}. \<forall>z\<in>{0..\<Sum>Basis}. \<forall>i\<in>Basis.
+  have "\<exists>e>0. \<forall>x\<in>unit_cube. \<forall>y\<in>unit_cube. \<forall>z\<in>unit_cube. \<forall>i\<in>Basis.
     norm (x - z) < e \<and> norm (y - z) < e \<and> label x i \<noteq> label y i \<longrightarrow>
       abs ((f(z) - z)\<bullet>i) < d / (real n)"
   proof -
@@ -3850,12 +3877,12 @@
       unfolding n_def
       apply (auto simp:  DIM_positive)
       done
-    have *: "uniformly_continuous_on {0..\<Sum>Basis} f"
-      by (rule compact_uniformly_continuous[OF assms(1) compact_interval])
+    have *: "uniformly_continuous_on unit_cube f"
+      by (rule compact_uniformly_continuous[OF assms(1) compact_unit_cube])
     obtain e where e:
         "e > 0"
-        "\<And>x x'. x \<in> {0..\<Sum>Basis} \<Longrightarrow>
-          x' \<in> {0..\<Sum>Basis} \<Longrightarrow>
+        "\<And>x x'. x \<in> unit_cube \<Longrightarrow>
+          x' \<in> unit_cube \<Longrightarrow>
           norm (x' - x) < e \<Longrightarrow>
           norm (f x' - f x) < d / real n / 8"
       using *[unfolded uniformly_continuous_on_def,rule_format,OF d']
@@ -3869,7 +3896,7 @@
         using d' e by auto
       fix x y z i
       assume as:
-        "x \<in> {0..\<Sum>Basis}" "y \<in> {0..\<Sum>Basis}" "z \<in> {0..\<Sum>Basis}"
+        "x \<in> unit_cube" "y \<in> unit_cube" "z \<in> unit_cube"
         "norm (x - z) < min (e / 2) (d / real n / 8)"
         "norm (y - z) < min (e / 2) (d / real n / 8)"
         "label x i \<noteq> label y i"
@@ -3924,9 +3951,9 @@
   then
   obtain e where e:
     "e > 0"
-    "\<And>x y z i. x \<in> {0..\<Sum>Basis} \<Longrightarrow>
-      y \<in> {0..\<Sum>Basis} \<Longrightarrow>
-      z \<in> {0..\<Sum>Basis} \<Longrightarrow>
+    "\<And>x y z i. x \<in> unit_cube \<Longrightarrow>
+      y \<in> unit_cube \<Longrightarrow>
+      z \<in> unit_cube \<Longrightarrow>
       i \<in> Basis \<Longrightarrow>
       norm (x - z) < e \<and> norm (y - z) < e \<and> label x i \<noteq> label y i \<Longrightarrow>
       \<bar>(f z - z) \<bullet> i\<bar> < d / real n"
@@ -3987,8 +4014,8 @@
     assume as: "\<forall>i\<in>{1..n}. x i \<le> p" "i \<in> {1..n}"
     {
       assume "x i = p \<or> x i = 0"
-      have "(\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<in> {0::'a..\<Sum>Basis}"
-        unfolding mem_interval
+      have "(\<Sum>i\<in>Basis. (real (x (b' i)) / real p) *\<^sub>R i) \<in> (unit_cube::'a set)"
+        unfolding mem_unit_cube
         using as b'_Basis
         by (auto simp add: inner_simps bij_betw_def zero_le_divide_iff divide_le_eq_1)
     }
@@ -4021,8 +4048,8 @@
     have "\<forall>i\<in>Basis. q (b' i) \<in> {0..p}"
       using q(1) b'
       by (auto intro: less_imp_le simp: bij_betw_def)
-    then have "z \<in> {0..\<Sum>Basis}"
-      unfolding z_def mem_interval
+    then have "z \<in> unit_cube"
+      unfolding z_def mem_unit_cube
       using b'_Basis
       by (auto simp add: inner_simps bij_betw_def zero_le_divide_iff divide_le_eq_1)
     then have d_fz_z: "d \<le> norm (f z - z)"
@@ -4064,8 +4091,8 @@
     using q(1)[rule_format,OF b'_im]
     apply (auto simp add: Suc_le_eq)
     done
-  then have "r' \<in> {0..\<Sum>Basis}"
-    unfolding r'_def mem_interval
+  then have "r' \<in> unit_cube"
+    unfolding r'_def mem_unit_cube
     using b'_Basis
     by (auto simp add: inner_simps bij_betw_def zero_le_divide_iff divide_le_eq_1)
   def s' \<equiv> "(\<Sum>i\<in>Basis. (real (s (b' i)) / real p) *\<^sub>R i)::'a"
@@ -4075,12 +4102,12 @@
     using q(1)[rule_format,OF b'_im]
     apply (auto simp add: Suc_le_eq)
     done
-  then have "s' \<in> {0..\<Sum>Basis}"
-    unfolding s'_def mem_interval
+  then have "s' \<in> unit_cube"
+    unfolding s'_def mem_unit_cube
     using b'_Basis
     by (auto simp add: inner_simps bij_betw_def zero_le_divide_iff divide_le_eq_1)
-  have "z \<in> {0..\<Sum>Basis}"
-    unfolding z_def mem_interval
+  have "z \<in> unit_cube"
+    unfolding z_def mem_unit_cube
     using b'_Basis q(1)[rule_format,OF b'_im] `p > 0`
     by (auto simp add: inner_simps bij_betw_def zero_le_divide_iff divide_le_eq_1 less_imp_le)
   have *: "\<And>x. 1 + real x = real (Suc x)"
@@ -4118,7 +4145,7 @@
   then have "\<bar>(f z - z) \<bullet> i\<bar> < d / real n"
     using rs(3) i
     unfolding r'_def[symmetric] s'_def[symmetric] o_def bb'
-    by (intro e(2)[OF `r'\<in>{0..\<Sum>Basis}` `s'\<in>{0..\<Sum>Basis}` `z\<in>{0..\<Sum>Basis}`]) auto
+    by (intro e(2)[OF `r'\<in>unit_cube` `s'\<in>unit_cube` `z\<in>unit_cube`]) auto
   then show False
     using i by auto
 qed
@@ -4227,8 +4254,15 @@
 
 subsection {* The Brouwer theorem for any set with nonempty interior *}
 
+lemma convex_unit_cube: "convex unit_cube"
+  apply (rule is_interval_convex)
+  apply (clarsimp simp add: is_interval_def mem_unit_cube)
+  apply (drule (1) bspec)+
+  apply auto
+  done
+
 lemma brouwer_weak:
-  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> 'a::ordered_euclidean_space"
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'a"
   assumes "compact s"
     and "convex s"
     and "interior s \<noteq> {}"
@@ -4236,13 +4270,22 @@
     and "f ` s \<subseteq> s"
   obtains x where "x \<in> s" and "f x = x"
 proof -
-  have *: "interior {0::'a..\<Sum>Basis} \<noteq> {}"
-    unfolding interior_closed_interval interval_eq_empty
-    by auto
-  have *: "{0::'a..\<Sum>Basis} homeomorphic s"
-    using homeomorphic_convex_compact[OF convex_interval(1) compact_interval * assms(2,1,3)] .
-  have "\<forall>f. continuous_on {0::'a..\<Sum>Basis} f \<and> f ` {0::'a..\<Sum>Basis} \<subseteq> {0::'a..\<Sum>Basis} \<longrightarrow>
-    (\<exists>x\<in>{0::'a..\<Sum>Basis}. f x = x)"
+  let ?U = "unit_cube :: 'a set"
+  have "\<Sum>Basis /\<^sub>R 2 \<in> interior ?U"
+  proof (rule interiorI)
+    let ?I = "(\<Inter>i\<in>Basis. {x::'a. 0 < x \<bullet> i} \<inter> {x. x \<bullet> i < 1})"
+    show "open ?I"
+      by (intro open_INT finite_Basis ballI open_Int, auto intro: open_Collect_less)
+    show "\<Sum>Basis /\<^sub>R 2 \<in> ?I"
+      by simp
+    show "?I \<subseteq> unit_cube"
+      unfolding unit_cube_def by force
+  qed
+  then have *: "interior ?U \<noteq> {}" by fast
+  have *: "?U homeomorphic s"
+    using homeomorphic_convex_compact[OF convex_unit_cube compact_unit_cube * assms(2,1,3)] .
+  have "\<forall>f. continuous_on ?U f \<and> f ` ?U \<subseteq> ?U \<longrightarrow>
+    (\<exists>x\<in>?U. f x = x)"
     using brouwer_cube by auto
   then show ?thesis
     unfolding homeomorphic_fixpoint_property[OF *]
@@ -4260,7 +4303,7 @@
 text {* And in particular for a closed ball. *}
 
 lemma brouwer_ball:
-  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> 'a"
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'a"
   assumes "e > 0"
     and "continuous_on (cball a e) f"
     and "f ` cball a e \<subseteq> cball a e"
@@ -4274,7 +4317,7 @@
   a scaling and translation to put the set inside the unit cube. *}
 
 lemma brouwer:
-  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> 'a"
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'a"
   assumes "compact s"
     and "convex s"
     and "s \<noteq> {}"
@@ -4324,7 +4367,7 @@
 text {*So we get the no-retraction theorem. *}
 
 lemma no_retraction_cball:
-  fixes a :: "'a::ordered_euclidean_space"
+  fixes a :: "'a::euclidean_space"
   assumes "e > 0"
   shows "\<not> (frontier (cball a e) retract_of (cball a e))"
 proof
@@ -4363,7 +4406,7 @@
 
 subsection {*Bijections between intervals. *}
 
-definition interval_bij :: "'a \<times> 'a \<Rightarrow> 'a \<times> 'a \<Rightarrow> 'a \<Rightarrow> 'a::ordered_euclidean_space"
+definition interval_bij :: "'a \<times> 'a \<Rightarrow> 'a \<times> 'a \<Rightarrow> 'a \<Rightarrow> 'a::euclidean_space"
   where "interval_bij =
     (\<lambda>(a, b) (u, v) x. (\<Sum>i\<in>Basis. (u\<bullet>i + (x\<bullet>i - a\<bullet>i) / (b\<bullet>i - a\<bullet>i) * (v\<bullet>i - u\<bullet>i)) *\<^sub>R i))"
 
@@ -4374,7 +4417,7 @@
     field_simps inner_simps add_divide_distrib[symmetric] intro!: setsum_cong)
 
 lemma continuous_interval_bij:
-  fixes a b :: "'a::ordered_euclidean_space"
+  fixes a b :: "'a::euclidean_space"
   shows "continuous (at x) (interval_bij (a, b) (u, v))"
   by (auto simp add: divide_inverse interval_bij_def intro!: continuous_setsum continuous_intros)
 
@@ -4414,7 +4457,7 @@
 qed
 
 lemma interval_bij_bij:
-  "\<forall>(i::'a::ordered_euclidean_space)\<in>Basis. a\<bullet>i < b\<bullet>i \<and> u\<bullet>i < v\<bullet>i \<Longrightarrow>
+  "\<forall>(i::'a::euclidean_space)\<in>Basis. a\<bullet>i < b\<bullet>i \<and> u\<bullet>i < v\<bullet>i \<Longrightarrow>
     interval_bij (a, b) (u, v) (interval_bij (u, v) (a, b) x) = x"
   by (auto simp: interval_bij_def euclidean_eq_iff[where 'a='a])
 
--- a/src/HOL/Multivariate_Analysis/Derivative.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Multivariate_Analysis/Derivative.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -1245,7 +1245,7 @@
 subsection {* Proving surjectivity via Brouwer fixpoint theorem *}
 
 lemma brouwer_surjective:
-  fixes f :: "'n::ordered_euclidean_space \<Rightarrow> 'n"
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'n"
   assumes "compact t"
     and "convex t"
     and "t \<noteq> {}"
@@ -1266,7 +1266,7 @@
 qed
 
 lemma brouwer_surjective_cball:
-  fixes f :: "'n::ordered_euclidean_space \<Rightarrow> 'n"
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'n"
   assumes "e > 0"
     and "continuous_on (cball a e) f"
     and "\<forall>x\<in>s. \<forall>y\<in>cball a e. x + (y - f y) \<in> cball a e"
@@ -1282,7 +1282,7 @@
 text {* See Sussmann: "Multidifferential calculus", Theorem 2.1.1 *}
 
 lemma sussmann_open_mapping:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::ordered_euclidean_space"
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
   assumes "open s"
     and "continuous_on s f"
     and "x \<in> s"
@@ -1469,7 +1469,7 @@
 qed
 
 lemma has_derivative_inverse_strong:
-  fixes f :: "'n::ordered_euclidean_space \<Rightarrow> 'n"
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'n"
   assumes "open s"
     and "x \<in> s"
     and "continuous_on s f"
@@ -1548,7 +1548,7 @@
 text {* A rewrite based on the other domain. *}
 
 lemma has_derivative_inverse_strong_x:
-  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> 'a"
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'a"
   assumes "open s"
     and "g y \<in> s"
     and "continuous_on s f"
@@ -1564,7 +1564,7 @@
 text {* On a region. *}
 
 lemma has_derivative_inverse_on:
-  fixes f :: "'n::ordered_euclidean_space \<Rightarrow> 'n"
+  fixes f :: "'n::euclidean_space \<Rightarrow> 'n"
   assumes "open s"
     and "\<forall>x\<in>s. (f has_derivative f'(x)) (at x)"
     and "\<forall>x\<in>s. g (f x) = x"
--- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -2958,10 +2958,10 @@
   fix f
   assume *: "Ball f open" "s \<union> t \<subseteq> \<Union>f"
   from * `compact s` obtain s' where "s' \<subseteq> f \<and> finite s' \<and> s \<subseteq> \<Union>s'"
-    unfolding compact_eq_heine_borel by (auto elim!: allE[of _ f]) metis
+    unfolding compact_eq_heine_borel by (auto elim!: allE[of _ f])
   moreover
   from * `compact t` obtain t' where "t' \<subseteq> f \<and> finite t' \<and> t \<subseteq> \<Union>t'"
-    unfolding compact_eq_heine_borel by (auto elim!: allE[of _ f]) metis
+    unfolding compact_eq_heine_borel by (auto elim!: allE[of _ f])
   ultimately show "\<exists>f'\<subseteq>f. finite f' \<and> s \<union> t \<subseteq> \<Union>f'"
     by (auto intro!: exI[of _ "s' \<union> t'"])
 qed
--- a/src/HOL/Nominal/Examples/Class1.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Nominal/Examples/Class1.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -1702,32 +1702,13 @@
 next
   case (NotR y M d)
   then show ?case 
-    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod)
-    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(N,M{d:=(x).N},([(c,d)]\<bullet>M){c:=(x).N})")
-    apply(erule exE, simp add: fresh_prod)
-    apply(erule conjE)+
-    apply(simp add: fresh_fun_simp_NotR)
-    apply(simp add: trm.inject alpha)
-    apply(rule exists_fresh'(2)[OF fs_coname1])
-    done
+    by(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod)
 next
   case (AndR c1 M c2 M' c3)
   then show ?case
-    apply(simp)
-    apply(auto)
-    apply(simp add: fresh_prod calc_atm fresh_atm abs_fresh)
-    apply(simp add: fresh_prod calc_atm fresh_atm abs_fresh fresh_left)
-    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(N,M{c3:=(x).N},
-                  M'{c3:=(x).N},c1,c2,c3,([(c,c3)]\<bullet>M){c:=(x).N},([(c,c3)]\<bullet>M'){c:=(x).N})")
-    apply(erule exE, simp add: fresh_prod)
-    apply(erule conjE)+
-    apply(simp add: fresh_fun_simp_AndR)
-    apply (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh subst_fresh)
-    apply(rule exists_fresh'(2)[OF fs_coname1])
-    apply(simp add: fresh_prod calc_atm fresh_atm abs_fresh fresh_left)
-    apply(simp add: fresh_prod calc_atm fresh_atm abs_fresh fresh_left)
-    apply(auto simp add: trm.inject alpha)
-    done
+    apply(auto simp add: fresh_prod calc_atm fresh_atm abs_fresh fresh_left)
+    apply (metis (erased, hide_lams))
+    by metis
 next
   case AndL1
   then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod)
@@ -1737,25 +1718,11 @@
 next
   case (OrR1 d M e)
   then show ?case 
-    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod)
-    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(N,M{e:=(x).N},([(c,e)]\<bullet>M){c:=(x).N},d,e)")
-    apply(erule exE, simp add: fresh_prod)
-    apply(erule conjE)+
-    apply(simp add: fresh_fun_simp_OrR1)
-    apply(simp add: trm.inject alpha)
-    apply(rule exists_fresh'(2)[OF fs_coname1])
-    done
+    by(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod)
 next
   case (OrR2 d M e)
   then show ?case 
-    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod)
-    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(N,M{e:=(x).N},([(c,e)]\<bullet>M){c:=(x).N},d,e)")
-    apply(erule exE, simp add: fresh_prod)
-    apply(erule conjE)+
-    apply(simp add: fresh_fun_simp_OrR2)
-    apply(simp add: trm.inject alpha)
-    apply(rule exists_fresh'(2)[OF fs_coname1])
-    done
+    by(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod)
 next
   case (OrL x1 M x2 M' x3)
   then show ?case
@@ -1764,24 +1731,16 @@
   case ImpL
   then show ?case
     by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
+       metis
 next
   case (ImpR y d M e)
   then show ?case
-    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
-    apply(subgoal_tac "\<exists>a'::coname. a'\<sharp>(N,M{e:=(x).N},([(c,e)]\<bullet>M){c:=(x).N},d,e)")
-    apply(erule exE, simp add: fresh_prod)
-    apply(erule conjE)+
-    apply(simp add: fresh_fun_simp_ImpR)
-    apply(simp add: trm.inject alpha)
-    apply(rule exists_fresh'(2)[OF fs_coname1])
-    done
+    by(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
 next
   case (Cut d M y M')
   then show ?case
-    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
-    apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst])
-    apply(simp add: calc_atm)
-    done
+    by(simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
+      (metis crename.simps(1) crename_id crename_rename)
 qed
 
 lemma substc_rename2:
@@ -1885,14 +1844,7 @@
 next
   case (NotL d M z)
   then show ?case 
-    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod)
-    apply(subgoal_tac "\<exists>a'::name. a'\<sharp>(N,M{x:=<a>.N},([(y,x)]\<bullet>M){y:=<a>.N})")
-    apply(erule exE, simp add: fresh_prod)
-    apply(erule conjE)+
-    apply(simp add: fresh_fun_simp_NotL)
-    apply(simp add: trm.inject alpha)
-    apply(rule exists_fresh'(1)[OF fs_name1])
-    done
+    by(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod)
 next
   case (AndR c1 M c2 M' c3)
   then show ?case
@@ -1906,37 +1858,16 @@
 next
   case (AndL1 u M v)
   then show ?case 
-    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod)
-    apply(subgoal_tac "\<exists>a'::name. a'\<sharp>(N,M{x:=<a>.N},([(y,x)]\<bullet>M){y:=<a>.N},u,v)")
-    apply(erule exE, simp add: fresh_prod)
-    apply(erule conjE)+
-    apply(simp add: fresh_fun_simp_AndL1)
-    apply(simp add: trm.inject alpha)
-    apply(rule exists_fresh'(1)[OF fs_name1])
-    done
+    by(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod)
 next
   case (AndL2 u M v)
   then show ?case 
-    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod)
-    apply(subgoal_tac "\<exists>a'::name. a'\<sharp>(N,M{x:=<a>.N},([(y,x)]\<bullet>M){y:=<a>.N},u,v)")
-    apply(erule exE, simp add: fresh_prod)
-    apply(erule conjE)+
-    apply(simp add: fresh_fun_simp_AndL2)
-    apply(simp add: trm.inject alpha)
-    apply(rule exists_fresh'(1)[OF fs_name1])
-    done
+    by(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod)
 next
   case (OrL x1 M x2 M' x3)
   then show ?case
-    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
-    apply(subgoal_tac 
-    "\<exists>a'::name. a'\<sharp>(N,M{x:=<a>.N},M'{x:=<a>.N},([(y,x)]\<bullet>M){y:=<a>.N},([(y,x)]\<bullet>M'){y:=<a>.N},x1,x2)")
-    apply(erule exE, simp add: fresh_prod)
-    apply(erule conjE)+
-    apply(simp add: fresh_fun_simp_OrL)
-    apply(simp add: trm.inject alpha)
-    apply(rule exists_fresh'(1)[OF fs_name1])
-    done
+    by(simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
+      (metis (poly_guards_query))
 next 
   case ImpR
   then show ?case
@@ -1944,21 +1875,15 @@
 next
   case (ImpL d M v M' u)
   then show ?case
-    apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
-    apply(subgoal_tac 
-    "\<exists>a'::name. a'\<sharp>(N,M{u:=<a>.N},M'{u:=<a>.N},([(y,u)]\<bullet>M){y:=<a>.N},([(y,u)]\<bullet>M'){y:=<a>.N},d,v)")
-    apply(erule exE, simp add: fresh_prod)
-    apply(erule conjE)+
-    apply(simp add: fresh_fun_simp_ImpL)
-    apply(simp add: trm.inject alpha)
-    apply(rule exists_fresh'(1)[OF fs_name1])
-    done
+    by(simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
+       (metis (poly_guards_query))
 next
   case (Cut d M y M')
   then show ?case
     apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left)
     apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst])
     apply(simp add: calc_atm)
+    apply metis
     done
 qed
 
--- a/src/HOL/Nominal/Examples/Class2.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Nominal/Examples/Class2.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -2869,15 +2869,9 @@
 apply(rotate_tac 10)
 apply(drule_tac x="a" in meta_spec)
 apply(auto simp add: ty.inject)
-apply(blast)
-apply(blast)
-apply(blast)
 apply(rotate_tac 10)
 apply(drule_tac x="a" in meta_spec)
 apply(auto simp add: ty.inject)
-apply(blast)
-apply(blast)
-apply(blast)
 done
 
 termination
--- a/src/HOL/Old_Number_Theory/Pocklington.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Old_Number_Theory/Pocklington.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -615,7 +615,7 @@
             from x(3)[rule_format, of z] z(2,3) have "z=x"
               unfolding modeq_def mod_less[OF y(2)] by simp}
           with xm x(1,2) have "\<exists>!x. x \<in> ?S \<and> (?h \<circ> op * a) x = y"
-            unfolding modeq_def mod_less[OF y(2)] by auto }
+            unfolding modeq_def mod_less[OF y(2)] by safe auto }
         thus "\<forall>y\<in>{m. coprime m n \<and> m < n}.
        \<exists>!x. x \<in> {m. coprime m n \<and> m < n} \<and> ((\<lambda>m. m mod n) \<circ> op * a) x = y" by blast
       next
--- a/src/HOL/Product_Type.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Product_Type.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -1125,7 +1125,7 @@
 
 lemma swap_product:
   "(%(i, j). (j, i)) ` (A \<times> B) = B \<times> A"
-  by (simp add: split_def image_def) blast
+  by (simp add: split_def image_def set_eq_iff)
 
 lemma image_split_eq_Sigma:
   "(\<lambda>x. (f x, g x)) ` A = Sigma (f ` A) (\<lambda>x. g ` (f -` {x} \<inter> A))"
--- a/src/HOL/Proofs/Lambda/LambdaType.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Proofs/Lambda/LambdaType.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -130,10 +130,6 @@
   apply (case_tac ys)
   apply simp
   apply simp
-  apply atomize
-  apply (erule allE)+
-  apply (erule mp, rule conjI)
-  apply (rule refl)+
   done
 
 lemma types_snocE: "e \<tturnstile> ts @ [t] : Ts \<Longrightarrow>
@@ -142,7 +138,6 @@
   apply simp
   apply (case_tac "ts @ [t]")
   apply (simp add: types_snoc_eq)+
-  apply iprover
   done
 
 
--- a/src/HOL/Proofs/Lambda/WeakNorm.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Proofs/Lambda/WeakNorm.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -93,7 +93,7 @@
         next
           case (Cons a as)
           with argsT obtain T'' Ts where Us: "Us = T'' # Ts"
-            by (cases Us) (rule FalseE, simp+, erule that)
+            by (cases Us) (rule FalseE, simp)
           from varT and Us have varT: "e\<langle>i:T\<rangle> \<turnstile> Var x : T'' \<Rightarrow> Ts \<Rrightarrow> T'"
             by simp
           from varT eq have T: "T = T'' \<Rightarrow> Ts \<Rrightarrow> T'" by cases auto
--- a/src/HOL/Quotient.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Quotient.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -487,15 +487,7 @@
 
 lemma bex1_bexeq_reg:
   shows "(\<exists>!x\<in>Respects R. P x) \<longrightarrow> (Bex1_rel R (\<lambda>x. P x))"
-  apply (simp add: Ex1_def Bex1_rel_def in_respects)
-  apply clarify
-  apply auto
-  apply (rule bexI)
-  apply assumption
-  apply (simp add: in_respects)
-  apply (simp add: in_respects)
-  apply auto
-  done
+  by (auto simp add: Ex1_def Bex1_rel_def Bex_def Ball_def in_respects)
 
 lemma bex1_bexeq_reg_eqv:
   assumes a: "equivp R"
--- a/src/HOL/ROOT	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/ROOT	Fri Mar 14 17:32:11 2014 +0100
@@ -777,7 +777,8 @@
     "Boogie_Dijkstra.certs"
     "Boogie_Max.certs"
     "SMT_Examples.certs"
-    "SMT_Word_Examples.certs"
+    "SMT_Examples.certs2"
+    "SMT_Word_Examples.certs2"
     "VCC_Max.certs"
 
 session "HOL-SPARK" (main) in "SPARK" = "HOL-Word" +
--- a/src/HOL/Real.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Real.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -2197,7 +2197,20 @@
     times_real_inst.times_real uminus_real_inst.uminus_real
     zero_real_inst.zero_real
 
+
+subsection {* Setup for SMT *}
+
 ML_file "Tools/SMT/smt_real.ML"
 setup SMT_Real.setup
+ML_file "Tools/SMT2/smt2_real.ML"
+ML_file "Tools/SMT2/z3_new_real.ML"
+
+lemma [z3_new_rule]:
+  "0 + (x::real) = x"
+  "x + 0 = x"
+  "0 * x = 0"
+  "1 * x = x"
+  "x + y = y + x"
+  by auto
 
 end
--- a/src/HOL/Relation.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Relation.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -1135,7 +1135,4 @@
     (auto simp: comp_fun_commute.fold_insert comp_fun_commute_relcomp_fold insert_relcomp_fold
       cong: if_cong)
 
-
-
 end
-
--- a/src/HOL/SMT.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/SMT.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -31,14 +31,13 @@
 quantifier block.
 *}
 
-datatype pattern = Pattern
+typedecl pattern
 
-definition pat :: "'a \<Rightarrow> pattern" where "pat _ = Pattern"
-definition nopat :: "'a \<Rightarrow> pattern" where "nopat _ = Pattern"
+consts
+  pat :: "'a \<Rightarrow> pattern"
+  nopat :: "'a \<Rightarrow> pattern"
 
-definition trigger :: "pattern list list \<Rightarrow> bool \<Rightarrow> bool"
-where "trigger _ P = P"
-
+definition trigger :: "pattern list list \<Rightarrow> bool \<Rightarrow> bool" where "trigger _ P = P"
 
 
 subsection {* Quantifier weights *}
@@ -67,7 +66,6 @@
 *}
 
 
-
 subsection {* Higher-order encoding *}
 
 text {*
@@ -88,7 +86,6 @@
   fun_upd_upd fun_app_def
 
 
-
 subsection {* First-order logic *}
 
 text {*
@@ -107,7 +104,6 @@
 definition term_false where "term_false = False"
 
 
-
 subsection {* Integer division and modulo for Z3 *}
 
 definition z3div :: "int \<Rightarrow> int \<Rightarrow> int" where
@@ -117,7 +113,6 @@
   "z3mod k l = (if 0 \<le> l then k mod l else k mod (-l))"
 
 
-
 subsection {* Setup *}
 
 ML_file "Tools/SMT/smt_builtin.ML"
@@ -426,7 +421,7 @@
 
 
 hide_type (open) pattern
-hide_const Pattern fun_app term_true term_false z3div z3mod
+hide_const fun_app term_true term_false z3div z3mod
 hide_const (open) trigger pat nopat weight
 
 end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SMT2.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -0,0 +1,445 @@
+(*  Title:      HOL/SMT2.thy
+    Author:     Sascha Boehme, TU Muenchen
+*)
+
+header {* Bindings to Satisfiability Modulo Theories (SMT) solvers based on SMT-LIB 2 *}
+
+theory SMT2
+imports Record
+keywords "smt2_status" :: diag
+begin
+
+ML_file "Tools/SMT2/smt2_util.ML"
+ML_file "Tools/SMT2/smt2_failure.ML"
+ML_file "Tools/SMT2/smt2_config.ML"
+
+
+subsection {* Triggers for quantifier instantiation *}
+
+text {*
+Some SMT solvers support patterns as a quantifier instantiation
+heuristics.  Patterns may either be positive terms (tagged by "pat")
+triggering quantifier instantiations -- when the solver finds a
+term matching a positive pattern, it instantiates the corresponding
+quantifier accordingly -- or negative terms (tagged by "nopat")
+inhibiting quantifier instantiations.  A list of patterns
+of the same kind is called a multipattern, and all patterns in a
+multipattern are considered conjunctively for quantifier instantiation.
+A list of multipatterns is called a trigger, and their multipatterns
+act disjunctively during quantifier instantiation.  Each multipattern
+should mention at least all quantified variables of the preceding
+quantifier block.
+*}
+
+typedecl pattern
+
+consts
+  pat :: "'a \<Rightarrow> pattern"
+  nopat :: "'a \<Rightarrow> pattern"
+
+definition trigger :: "pattern list list \<Rightarrow> bool \<Rightarrow> bool" where "trigger _ P = P"
+
+
+subsection {* Quantifier weights *}
+
+text {*
+Weight annotations to quantifiers influence the priority of quantifier
+instantiations.  They should be handled with care for solvers, which support
+them, because incorrect choices of weights might render a problem unsolvable.
+*}
+
+definition weight :: "int \<Rightarrow> bool \<Rightarrow> bool" where "weight _ P = P"
+
+text {*
+Weights must be nonnegative.  The value @{text 0} is equivalent to providing
+no weight at all.
+
+Weights should only be used at quantifiers and only inside triggers (if the
+quantifier has triggers).  Valid usages of weights are as follows:
+
+\begin{itemize}
+\item
+@{term "\<forall>x. trigger [[pat (P x)]] (weight 2 (P x))"}
+\item
+@{term "\<forall>x. weight 3 (P x)"}
+\end{itemize}
+*}
+
+
+subsection {* Higher-order encoding *}
+
+text {*
+Application is made explicit for constants occurring with varying
+numbers of arguments.  This is achieved by the introduction of the
+following constant.
+*}
+
+definition fun_app :: "'a \<Rightarrow> 'a" where "fun_app f = f"
+
+text {*
+Some solvers support a theory of arrays which can be used to encode
+higher-order functions.  The following set of lemmas specifies the
+properties of such (extensional) arrays.
+*}
+
+lemmas array_rules = ext fun_upd_apply fun_upd_same fun_upd_other  fun_upd_upd fun_app_def
+
+
+subsection {* Normalization *}
+
+lemma case_bool_if[abs_def]: "case_bool x y P = (if P then x else y)"
+  by simp
+
+lemma nat_int': "\<forall>n. nat (int n) = n" by simp
+lemma int_nat_nneg: "\<forall>i. i \<ge> 0 \<longrightarrow> int (nat i) = i" by simp
+lemma int_nat_neg: "\<forall>i. i < 0 \<longrightarrow> int (nat i) = 0" by simp
+
+lemma nat_zero_as_int: "0 = nat 0" by (rule transfer_nat_int_numerals(1))
+lemma nat_one_as_int: "1 = nat 1" by (rule transfer_nat_int_numerals(2))
+lemma nat_numeral_as_int: "numeral = (\<lambda>i. nat (numeral i))" by simp
+lemma nat_less_as_int: "op < = (\<lambda>a b. int a < int b)" by simp
+lemma nat_leq_as_int: "op \<le> = (\<lambda>a b. int a <= int b)" by simp
+lemma Suc_as_int: "Suc = (\<lambda>a. nat (int a + 1))" by (rule ext) simp
+lemma nat_plus_as_int: "op + = (\<lambda>a b. nat (int a + int b))" by (rule ext)+ simp
+lemma nat_minus_as_int: "op - = (\<lambda>a b. nat (int a - int b))" by (rule ext)+ simp
+lemma nat_times_as_int: "op * = (\<lambda>a b. nat (int a * int b))" by (simp add: nat_mult_distrib)
+lemma nat_div_as_int: "op div = (\<lambda>a b. nat (int a div int b))" by (simp add: nat_div_distrib)
+lemma nat_mod_as_int: "op mod = (\<lambda>a b. nat (int a mod int b))" by (simp add: nat_mod_distrib)
+
+lemma int_Suc: "int (Suc n) = int n + 1" by simp
+lemma int_plus: "int (n + m) = int n + int m" by (rule of_nat_add)
+lemma int_minus: "int (n - m) = int (nat (int n - int m))" by auto
+
+lemmas Ex1_def_raw = Ex1_def[abs_def]
+lemmas Ball_def_raw = Ball_def[abs_def]
+lemmas Bex_def_raw = Bex_def[abs_def]
+lemmas abs_if_raw = abs_if[abs_def]
+lemmas min_def_raw = min_def[abs_def]
+lemmas max_def_raw = max_def[abs_def]
+
+
+subsection {* Integer division and modulo for Z3 *}
+
+text {*
+The following Z3-inspired definitions are overspecified for the case where @{text "l = 0"}. This
+Schönheitsfehler is corrected in the @{text div_as_z3div} and @{text mod_as_z3mod} theorems.
+*}
+
+definition z3div :: "int \<Rightarrow> int \<Rightarrow> int" where
+  "z3div k l = (if l \<ge> 0 then k div l else - (k div - l))"
+
+definition z3mod :: "int \<Rightarrow> int \<Rightarrow> int" where
+  "z3mod k l = k mod (if l \<ge> 0 then l else - l)"
+
+lemma div_as_z3div:
+  "\<forall>k l. k div l = (if l = 0 then 0 else if l > 0 then z3div k l else z3div (- k) (- l))"
+  by (simp add: z3div_def)
+
+lemma mod_as_z3mod:
+  "\<forall>k l. k mod l = (if l = 0 then k else if l > 0 then z3mod k l else - z3mod (- k) (- l))"
+  by (simp add: z3mod_def)
+
+
+subsection {* Setup *}
+
+ML_file "Tools/SMT2/smt2_builtin.ML"
+ML_file "Tools/SMT2/smt2_datatypes.ML"
+ML_file "Tools/SMT2/smt2_normalize.ML"
+ML_file "Tools/SMT2/smt2_translate.ML"
+ML_file "Tools/SMT2/smtlib2.ML"
+ML_file "Tools/SMT2/smtlib2_interface.ML"
+ML_file "Tools/SMT2/z3_new_model.ML"
+ML_file "Tools/SMT2/z3_new_proof.ML"
+ML_file "Tools/SMT2/smt2_solver.ML"
+ML_file "Tools/SMT2/z3_new_isar.ML"
+ML_file "Tools/SMT2/z3_new_interface.ML"
+ML_file "Tools/SMT2/z3_new_replay_util.ML"
+ML_file "Tools/SMT2/z3_new_replay_literals.ML"
+ML_file "Tools/SMT2/z3_new_replay_rules.ML"
+ML_file "Tools/SMT2/z3_new_replay_methods.ML"
+ML_file "Tools/SMT2/z3_new_replay.ML"
+ML_file "Tools/SMT2/smt2_systems.ML"
+
+method_setup smt2 = {*
+  Scan.optional Attrib.thms [] >>
+    (fn thms => fn ctxt =>
+      METHOD (fn facts => HEADGOAL (SMT2_Solver.smt2_tac ctxt (thms @ facts))))
+*} "apply an SMT solver to the current goal (based on SMT-LIB 2)"
+
+
+subsection {* Configuration *}
+
+text {*
+The current configuration can be printed by the command
+@{text smt2_status}, which shows the values of most options.
+*}
+
+
+
+subsection {* General configuration options *}
+
+text {*
+The option @{text smt2_solver} can be used to change the target SMT
+solver.  The possible values can be obtained from the @{text smt2_status}
+command.
+
+Due to licensing restrictions, Yices and Z3 are not installed/enabled
+by default.  Z3 is free for non-commercial applications and can be enabled
+by setting Isabelle system option @{text z3_non_commercial} to @{text yes}.
+*}
+
+declare [[ smt2_solver = z3_new ]]
+
+text {*
+Since SMT solvers are potentially non-terminating, there is a timeout
+(given in seconds) to restrict their runtime.  A value greater than
+120 (seconds) is in most cases not advisable.
+*}
+
+declare [[ smt2_timeout = 20 ]]
+
+text {*
+SMT solvers apply randomized heuristics.  In case a problem is not
+solvable by an SMT solver, changing the following option might help.
+*}
+
+declare [[ smt2_random_seed = 1 ]]
+
+text {*
+In general, the binding to SMT solvers runs as an oracle, i.e, the SMT
+solvers are fully trusted without additional checks.  The following
+option can cause the SMT solver to run in proof-producing mode, giving
+a checkable certificate.  This is currently only implemented for Z3.
+*}
+
+declare [[ smt2_oracle = false ]]
+
+text {*
+Each SMT solver provides several commandline options to tweak its
+behaviour.  They can be passed to the solver by setting the following
+options.
+*}
+
+(* declare [[ cvc3_options = "" ]] TODO *)
+(* declare [[ yices_options = "" ]] TODO *)
+(* declare [[ z3_options = "" ]] TODO *)
+
+text {*
+The SMT method provides an inference mechanism to detect simple triggers
+in quantified formulas, which might increase the number of problems
+solvable by SMT solvers (note: triggers guide quantifier instantiations
+in the SMT solver).  To turn it on, set the following option.
+*}
+
+declare [[ smt2_infer_triggers = false ]]
+
+text {*
+Enable the following option to use built-in support for div/mod, datatypes,
+and records in Z3.  Currently, this is implemented only in oracle mode.
+*}
+
+declare [[ z3_new_extensions = false ]]
+
+text {*
+The SMT method monomorphizes the given facts, that is, it tries to
+instantiate all schematic type variables with fixed types occurring
+in the problem.  This is a (possibly nonterminating) fixed-point
+construction whose cycles are limited by the following option.
+*}
+
+declare [[ monomorph_max_rounds = 5 ]]
+
+text {*
+In addition, the number of generated monomorphic instances is limited
+by the following option.
+*}
+
+declare [[ monomorph_max_new_instances = 500 ]]
+
+
+
+subsection {* Certificates *}
+
+text {*
+By setting the option @{text smt2_certificates} to the name of a file,
+all following applications of an SMT solver a cached in that file.
+Any further application of the same SMT solver (using the very same
+configuration) re-uses the cached certificate instead of invoking the
+solver.  An empty string disables caching certificates.
+
+The filename should be given as an explicit path.  It is good
+practice to use the name of the current theory (with ending
+@{text ".certs"} instead of @{text ".thy"}) as the certificates file.
+Certificate files should be used at most once in a certain theory context,
+to avoid race conditions with other concurrent accesses.
+*}
+
+declare [[ smt2_certificates = "" ]]
+
+text {*
+The option @{text smt2_read_only_certificates} controls whether only
+stored certificates are should be used or invocation of an SMT solver
+is allowed.  When set to @{text true}, no SMT solver will ever be
+invoked and only the existing certificates found in the configured
+cache are used;  when set to @{text false} and there is no cached
+certificate for some proposition, then the configured SMT solver is
+invoked.
+*}
+
+declare [[ smt2_read_only_certificates = false ]]
+
+
+
+subsection {* Tracing *}
+
+text {*
+The SMT method, when applied, traces important information.  To
+make it entirely silent, set the following option to @{text false}.
+*}
+
+declare [[ smt2_verbose = true ]]
+
+text {*
+For tracing the generated problem file given to the SMT solver as
+well as the returned result of the solver, the option
+@{text smt2_trace} should be set to @{text true}.
+*}
+
+declare [[ smt2_trace = false ]]
+
+text {*
+From the set of assumptions given to the SMT solver, those assumptions
+used in the proof are traced when the following option is set to
+@{term true}.  This only works for Z3 when it runs in non-oracle mode
+(see options @{text smt2_solver} and @{text smt2_oracle} above).
+*}
+
+declare [[ smt2_trace_used_facts = false ]]
+
+
+subsection {* Schematic rules for Z3 proof reconstruction *}
+
+text {*
+Several prof rules of Z3 are not very well documented.  There are two
+lemma groups which can turn failing Z3 proof reconstruction attempts
+into succeeding ones: the facts in @{text z3_rule} are tried prior to
+any implemented reconstruction procedure for all uncertain Z3 proof
+rules;  the facts in @{text z3_simp} are only fed to invocations of
+the simplifier when reconstructing theory-specific proof steps.
+*}
+
+lemmas [z3_new_rule] =
+  refl eq_commute conj_commute disj_commute simp_thms nnf_simps
+  ring_distribs field_simps times_divide_eq_right times_divide_eq_left
+  if_True if_False not_not
+
+lemma [z3_new_rule]:
+  "(P \<and> Q) = (\<not>(\<not>P \<or> \<not>Q))"
+  "(P \<and> Q) = (\<not>(\<not>Q \<or> \<not>P))"
+  "(\<not>P \<and> Q) = (\<not>(P \<or> \<not>Q))"
+  "(\<not>P \<and> Q) = (\<not>(\<not>Q \<or> P))"
+  "(P \<and> \<not>Q) = (\<not>(\<not>P \<or> Q))"
+  "(P \<and> \<not>Q) = (\<not>(Q \<or> \<not>P))"
+  "(\<not>P \<and> \<not>Q) = (\<not>(P \<or> Q))"
+  "(\<not>P \<and> \<not>Q) = (\<not>(Q \<or> P))"
+  by auto
+
+lemma [z3_new_rule]:
+  "(P \<longrightarrow> Q) = (Q \<or> \<not>P)"
+  "(\<not>P \<longrightarrow> Q) = (P \<or> Q)"
+  "(\<not>P \<longrightarrow> Q) = (Q \<or> P)"
+  "(True \<longrightarrow> P) = P"
+  "(P \<longrightarrow> True) = True"
+  "(False \<longrightarrow> P) = True"
+  "(P \<longrightarrow> P) = True"
+  by auto
+
+lemma [z3_new_rule]:
+  "((P = Q) \<longrightarrow> R) = (R | (Q = (\<not>P)))"
+  by auto
+
+lemma [z3_new_rule]:
+  "(\<not>True) = False"
+  "(\<not>False) = True"
+  "(x = x) = True"
+  "(P = True) = P"
+  "(True = P) = P"
+  "(P = False) = (\<not>P)"
+  "(False = P) = (\<not>P)"
+  "((\<not>P) = P) = False"
+  "(P = (\<not>P)) = False"
+  "((\<not>P) = (\<not>Q)) = (P = Q)"
+  "\<not>(P = (\<not>Q)) = (P = Q)"
+  "\<not>((\<not>P) = Q) = (P = Q)"
+  "(P \<noteq> Q) = (Q = (\<not>P))"
+  "(P = Q) = ((\<not>P \<or> Q) \<and> (P \<or> \<not>Q))"
+  "(P \<noteq> Q) = ((\<not>P \<or> \<not>Q) \<and> (P \<or> Q))"
+  by auto
+
+lemma [z3_new_rule]:
+  "(if P then P else \<not>P) = True"
+  "(if \<not>P then \<not>P else P) = True"
+  "(if P then True else False) = P"
+  "(if P then False else True) = (\<not>P)"
+  "(if P then Q else True) = ((\<not>P) \<or> Q)"
+  "(if P then Q else True) = (Q \<or> (\<not>P))"
+  "(if P then Q else \<not>Q) = (P = Q)"
+  "(if P then Q else \<not>Q) = (Q = P)"
+  "(if P then \<not>Q else Q) = (P = (\<not>Q))"
+  "(if P then \<not>Q else Q) = ((\<not>Q) = P)"
+  "(if \<not>P then x else y) = (if P then y else x)"
+  "(if P then (if Q then x else y) else x) = (if P \<and> (\<not>Q) then y else x)"
+  "(if P then (if Q then x else y) else x) = (if (\<not>Q) \<and> P then y else x)"
+  "(if P then (if Q then x else y) else y) = (if P \<and> Q then x else y)"
+  "(if P then (if Q then x else y) else y) = (if Q \<and> P then x else y)"
+  "(if P then x else if P then y else z) = (if P then x else z)"
+  "(if P then x else if Q then x else y) = (if P \<or> Q then x else y)"
+  "(if P then x else if Q then x else y) = (if Q \<or> P then x else y)"
+  "(if P then x = y else x = z) = (x = (if P then y else z))"
+  "(if P then x = y else y = z) = (y = (if P then x else z))"
+  "(if P then x = y else z = y) = (y = (if P then x else z))"
+  by auto
+
+lemma [z3_new_rule]:
+  "0 + (x::int) = x"
+  "x + 0 = x"
+  "x + x = 2 * x"
+  "0 * x = 0"
+  "1 * x = x"
+  "x + y = y + x"
+  by auto
+
+lemma [z3_new_rule]:  (* for def-axiom *)
+  "P = Q \<or> P \<or> Q"
+  "P = Q \<or> \<not>P \<or> \<not>Q"
+  "(\<not>P) = Q \<or> \<not>P \<or> Q"
+  "(\<not>P) = Q \<or> P \<or> \<not>Q"
+  "P = (\<not>Q) \<or> \<not>P \<or> Q"
+  "P = (\<not>Q) \<or> P \<or> \<not>Q"
+  "P \<noteq> Q \<or> P \<or> \<not>Q"
+  "P \<noteq> Q \<or> \<not>P \<or> Q"
+  "P \<noteq> (\<not>Q) \<or> P \<or> Q"
+  "(\<not>P) \<noteq> Q \<or> P \<or> Q"
+  "P \<or> Q \<or> P \<noteq> (\<not>Q)"
+  "P \<or> Q \<or> (\<not>P) \<noteq> Q"
+  "P \<or> \<not>Q \<or> P \<noteq> Q"
+  "\<not>P \<or> Q \<or> P \<noteq> Q"
+  "P \<or> y = (if P then x else y)"
+  "P \<or> (if P then x else y) = y"
+  "\<not>P \<or> x = (if P then x else y)"
+  "\<not>P \<or> (if P then x else y) = x"
+  "P \<or> R \<or> \<not>(if P then Q else R)"
+  "\<not>P \<or> Q \<or> \<not>(if P then Q else R)"
+  "\<not>(if P then Q else R) \<or> \<not>P \<or> Q"
+  "\<not>(if P then Q else R) \<or> P \<or> R"
+  "(if P then Q else R) \<or> \<not>P \<or> \<not>Q"
+  "(if P then Q else R) \<or> P \<or> \<not>R"
+  "(if P then \<not>Q else R) \<or> \<not>P \<or> Q"
+  "(if P then Q else \<not>R) \<or> P \<or> R"
+  by auto
+
+hide_type (open) pattern
+hide_const fun_app z3div z3mod
+hide_const (open) trigger pat nopat weight
+
+end
--- a/src/HOL/SMT_Examples/SMT_Examples.certs	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/SMT_Examples/SMT_Examples.certs	Fri Mar 14 17:32:11 2014 +0100
@@ -1,669 +1,3 @@
-23d01cdabb599769b54210e40617eea3d6c91e30 8 0
-#2 := false
-#1 := true
-#7 := (not true)
-#29 := (iff #7 false)
-#30 := [rewrite]: #29
-#28 := [asserted]: #7
-[mp #28 #30]: false
-unsat
-22e23526a38d50ce23abbe4dbfb697891cbcd840 22 0
-#2 := false
-decl f1 :: S1
-#3 := f1
-decl f3 :: S1
-#7 := f3
-#8 := (= f3 f1)
-#9 := (not #8)
-#10 := (or #8 #9)
-#11 := (not #10)
-#40 := (iff #11 false)
-#1 := true
-#35 := (not true)
-#38 := (iff #35 false)
-#39 := [rewrite]: #38
-#36 := (iff #11 #35)
-#33 := (iff #10 true)
-#34 := [rewrite]: #33
-#37 := [monotonicity #34]: #36
-#41 := [trans #37 #39]: #40
-#32 := [asserted]: #11
-[mp #32 #41]: false
-unsat
-121552dd328e0993a2c6099c592d9c3db7fff190 28 0
-#2 := false
-decl f1 :: S1
-#3 := f1
-decl f3 :: S1
-#7 := f3
-#8 := (= f3 f1)
-#1 := true
-#9 := (and #8 true)
-#10 := (iff #9 #8)
-#11 := (not #10)
-#46 := (iff #11 false)
-#41 := (not true)
-#44 := (iff #41 false)
-#45 := [rewrite]: #44
-#42 := (iff #11 #41)
-#39 := (iff #10 true)
-#34 := (iff #8 #8)
-#37 := (iff #34 true)
-#38 := [rewrite]: #37
-#35 := (iff #10 #34)
-#33 := [rewrite]: #10
-#36 := [monotonicity #33]: #35
-#40 := [trans #36 #38]: #39
-#43 := [monotonicity #40]: #42
-#47 := [trans #43 #45]: #46
-#32 := [asserted]: #11
-[mp #32 #47]: false
-unsat
-263480c8c5909524c36f6198f60c623fbcfc953d 41 0
-#2 := false
-decl f1 :: S1
-#3 := f1
-decl f4 :: S1
-#9 := f4
-#10 := (= f4 f1)
-decl f3 :: S1
-#7 := f3
-#8 := (= f3 f1)
-#11 := (or #8 #10)
-#64 := (iff #11 false)
-#59 := (or false false)
-#62 := (iff #59 false)
-#63 := [rewrite]: #62
-#60 := (iff #11 #59)
-#57 := (iff #10 false)
-#48 := (not #10)
-#12 := (not #8)
-#13 := (and #11 #12)
-#37 := (not #13)
-#38 := (or #37 #10)
-#41 := (not #38)
-#14 := (implies #13 #10)
-#15 := (not #14)
-#42 := (iff #15 #41)
-#39 := (iff #14 #38)
-#40 := [rewrite]: #39
-#43 := [monotonicity #40]: #42
-#36 := [asserted]: #15
-#46 := [mp #36 #43]: #41
-#49 := [not-or-elim #46]: #48
-#58 := [iff-false #49]: #57
-#55 := (iff #8 false)
-#44 := [not-or-elim #46]: #13
-#47 := [and-elim #44]: #12
-#56 := [iff-false #47]: #55
-#61 := [monotonicity #56 #58]: #60
-#65 := [trans #61 #63]: #64
-#45 := [and-elim #44]: #11
-[mp #45 #65]: false
-unsat
-79d9d246dd9d27e03e8f1ea895e790f3a4420bfd 55 0
-#2 := false
-decl f1 :: S1
-#3 := f1
-decl f3 :: S1
-#7 := f3
-#8 := (= f3 f1)
-decl f5 :: S1
-#12 := f5
-#13 := (= f5 f1)
-#16 := (and #8 #13)
-decl f4 :: S1
-#9 := f4
-#10 := (= f4 f1)
-#15 := (and #13 #10)
-#17 := (or #15 #16)
-#18 := (implies #8 #17)
-#19 := (or #18 #8)
-#11 := (and #8 #10)
-#14 := (or #11 #13)
-#20 := (implies #14 #19)
-#21 := (not #20)
-#71 := (iff #21 false)
-#43 := (not #8)
-#44 := (or #43 #17)
-#47 := (or #44 #8)
-#53 := (not #14)
-#54 := (or #53 #47)
-#59 := (not #54)
-#69 := (iff #59 false)
-#1 := true
-#64 := (not true)
-#67 := (iff #64 false)
-#68 := [rewrite]: #67
-#65 := (iff #59 #64)
-#62 := (iff #54 true)
-#63 := [rewrite]: #62
-#66 := [monotonicity #63]: #65
-#70 := [trans #66 #68]: #69
-#60 := (iff #21 #59)
-#57 := (iff #20 #54)
-#50 := (implies #14 #47)
-#55 := (iff #50 #54)
-#56 := [rewrite]: #55
-#51 := (iff #20 #50)
-#48 := (iff #19 #47)
-#45 := (iff #18 #44)
-#46 := [rewrite]: #45
-#49 := [monotonicity #46]: #48
-#52 := [monotonicity #49]: #51
-#58 := [trans #52 #56]: #57
-#61 := [monotonicity #58]: #60
-#72 := [trans #61 #70]: #71
-#42 := [asserted]: #21
-[mp #42 #72]: false
-unsat
-050883983ebe99dc3b7f24a011b1724b1b2c4dd9 33 0
-#2 := false
-decl f1 :: S1
-#3 := f1
-decl f6 :: S1
-#14 := f6
-#15 := (= f6 f1)
-decl f5 :: S1
-#12 := f5
-#13 := (= f5 f1)
-#16 := (and #13 #15)
-decl f4 :: S1
-#9 := f4
-#10 := (= f4 f1)
-decl f3 :: S1
-#7 := f3
-#8 := (= f3 f1)
-#11 := (and #8 #10)
-#17 := (or #11 #16)
-#18 := (implies #17 #17)
-#19 := (not #18)
-#48 := (iff #19 false)
-#1 := true
-#43 := (not true)
-#46 := (iff #43 false)
-#47 := [rewrite]: #46
-#44 := (iff #19 #43)
-#41 := (iff #18 true)
-#42 := [rewrite]: #41
-#45 := [monotonicity #42]: #44
-#49 := [trans #45 #47]: #48
-#40 := [asserted]: #19
-[mp #40 #49]: false
-unsat
-8575241c64c02491d277f6598ca57e576f5a6b45 60 0
-#2 := false
-decl f1 :: S1
-#3 := f1
-decl f3 :: S1
-#7 := f3
-#8 := (= f3 f1)
-#9 := (iff #8 #8)
-#10 := (iff #9 #8)
-#11 := (iff #10 #8)
-#12 := (iff #11 #8)
-#13 := (iff #12 #8)
-#14 := (iff #13 #8)
-#15 := (iff #14 #8)
-#16 := (iff #15 #8)
-#17 := (iff #16 #8)
-#18 := (not #17)
-#78 := (iff #18 false)
-#1 := true
-#73 := (not true)
-#76 := (iff #73 false)
-#77 := [rewrite]: #76
-#74 := (iff #18 #73)
-#71 := (iff #17 true)
-#40 := (iff #9 true)
-#41 := [rewrite]: #40
-#69 := (iff #17 #9)
-#42 := (iff true #8)
-#45 := (iff #42 #8)
-#46 := [rewrite]: #45
-#66 := (iff #16 #42)
-#64 := (iff #15 true)
-#62 := (iff #15 #9)
-#59 := (iff #14 #42)
-#57 := (iff #13 true)
-#55 := (iff #13 #9)
-#52 := (iff #12 #42)
-#50 := (iff #11 true)
-#48 := (iff #11 #9)
-#43 := (iff #10 #42)
-#44 := [monotonicity #41]: #43
-#47 := [trans #44 #46]: #11
-#49 := [monotonicity #47]: #48
-#51 := [trans #49 #41]: #50
-#53 := [monotonicity #51]: #52
-#54 := [trans #53 #46]: #13
-#56 := [monotonicity #54]: #55
-#58 := [trans #56 #41]: #57
-#60 := [monotonicity #58]: #59
-#61 := [trans #60 #46]: #15
-#63 := [monotonicity #61]: #62
-#65 := [trans #63 #41]: #64
-#67 := [monotonicity #65]: #66
-#68 := [trans #67 #46]: #17
-#70 := [monotonicity #68]: #69
-#72 := [trans #70 #41]: #71
-#75 := [monotonicity #72]: #74
-#79 := [trans #75 #77]: #78
-#39 := [asserted]: #18
-[mp #39 #79]: false
-unsat
-8434421285df70a7e1728b19173d86303151090b 165 0
-#2 := false
-decl f1 :: S1
-#3 := f1
-decl f6 :: S1
-#13 := f6
-#14 := (= f6 f1)
-decl f5 :: S1
-#11 := f5
-#12 := (= f5 f1)
-decl f4 :: S1
-#9 := f4
-#10 := (= f4 f1)
-decl f3 :: S1
-#7 := f3
-#8 := (= f3 f1)
-#75 := (or #8 #10 #12 #14)
-#215 := (iff #75 false)
-#210 := (or false false false false)
-#213 := (iff #210 false)
-#214 := [rewrite]: #213
-#211 := (iff #75 #210)
-#167 := (iff #14 false)
-#119 := (not #14)
-#122 := (or #119 #12)
-#175 := (iff #122 #119)
-#170 := (or #119 false)
-#173 := (iff #170 #119)
-#174 := [rewrite]: #173
-#171 := (iff #122 #170)
-#168 := (iff #12 false)
-#25 := (not #12)
-decl f11 :: S1
-#43 := f11
-#44 := (= f11 f1)
-#45 := (not #44)
-#46 := (and #44 #45)
-decl f10 :: S1
-#40 := f10
-#41 := (= f10 f1)
-#47 := (or #41 #46)
-#42 := (not #41)
-#48 := (and #42 #47)
-#49 := (or #12 #48)
-#50 := (not #49)
-#150 := (iff #50 #25)
-#148 := (iff #49 #12)
-#143 := (or #12 false)
-#146 := (iff #143 #12)
-#147 := [rewrite]: #146
-#144 := (iff #49 #143)
-#141 := (iff #48 false)
-#136 := (and #42 #41)
-#139 := (iff #136 false)
-#140 := [rewrite]: #139
-#137 := (iff #48 #136)
-#134 := (iff #47 #41)
-#129 := (or #41 false)
-#132 := (iff #129 #41)
-#133 := [rewrite]: #132
-#130 := (iff #47 #129)
-#126 := (iff #46 false)
-#128 := [rewrite]: #126
-#131 := [monotonicity #128]: #130
-#135 := [trans #131 #133]: #134
-#138 := [monotonicity #135]: #137
-#142 := [trans #138 #140]: #141
-#145 := [monotonicity #142]: #144
-#149 := [trans #145 #147]: #148
-#151 := [monotonicity #149]: #150
-#125 := [asserted]: #50
-#154 := [mp #125 #151]: #25
-#169 := [iff-false #154]: #168
-#172 := [monotonicity #169]: #171
-#176 := [trans #172 #174]: #175
-#37 := (or #14 false)
-#38 := (not #37)
-#39 := (or #38 #12)
-#123 := (iff #39 #122)
-#120 := (iff #38 #119)
-#116 := (iff #37 #14)
-#118 := [rewrite]: #116
-#121 := [monotonicity #118]: #120
-#124 := [monotonicity #121]: #123
-#115 := [asserted]: #39
-#127 := [mp #115 #124]: #122
-#166 := [mp #127 #176]: #119
-#177 := [iff-false #166]: #167
-#165 := (iff #10 false)
-#109 := (not #10)
-#112 := (or #109 #12)
-#183 := (iff #112 #109)
-#178 := (or #109 false)
-#181 := (iff #178 #109)
-#182 := [rewrite]: #181
-#179 := (iff #112 #178)
-#180 := [monotonicity #169]: #179
-#184 := [trans #180 #182]: #183
-decl f9 :: S1
-#30 := f9
-#31 := (= f9 f1)
-#32 := (not #31)
-#33 := (or #31 #32)
-#34 := (and #10 #33)
-#35 := (not #34)
-#36 := (or #35 #12)
-#113 := (iff #36 #112)
-#110 := (iff #35 #109)
-#107 := (iff #34 #10)
-#1 := true
-#102 := (and #10 true)
-#105 := (iff #102 #10)
-#106 := [rewrite]: #105
-#103 := (iff #34 #102)
-#99 := (iff #33 true)
-#101 := [rewrite]: #99
-#104 := [monotonicity #101]: #103
-#108 := [trans #104 #106]: #107
-#111 := [monotonicity #108]: #110
-#114 := [monotonicity #111]: #113
-#98 := [asserted]: #36
-#117 := [mp #98 #114]: #112
-#164 := [mp #117 #184]: #109
-#185 := [iff-false #164]: #165
-#163 := (iff #8 false)
-#92 := (not #8)
-#95 := (or #92 #10)
-#191 := (iff #95 #92)
-#186 := (or #92 false)
-#189 := (iff #186 #92)
-#190 := [rewrite]: #189
-#187 := (iff #95 #186)
-#188 := [monotonicity #185]: #187
-#192 := [trans #188 #190]: #191
-#26 := (and #12 #25)
-#27 := (or #8 #26)
-#28 := (not #27)
-#29 := (or #28 #10)
-#96 := (iff #29 #95)
-#93 := (iff #28 #92)
-#90 := (iff #27 #8)
-#85 := (or #8 false)
-#88 := (iff #85 #8)
-#89 := [rewrite]: #88
-#86 := (iff #27 #85)
-#79 := (iff #26 false)
-#84 := [rewrite]: #79
-#87 := [monotonicity #84]: #86
-#91 := [trans #87 #89]: #90
-#94 := [monotonicity #91]: #93
-#97 := [monotonicity #94]: #96
-#74 := [asserted]: #29
-#100 := [mp #74 #97]: #95
-#162 := [mp #100 #192]: #92
-#193 := [iff-false #162]: #163
-#212 := [monotonicity #193 #185 #169 #177]: #211
-#216 := [trans #212 #214]: #215
-#15 := (or #12 #14)
-#16 := (or #10 #15)
-#17 := (or #8 #16)
-#76 := (iff #17 #75)
-#77 := [rewrite]: #76
-#72 := [asserted]: #17
-#78 := [mp #72 #77]: #75
-[mp #78 #216]: false
-unsat
-2571c5d0e3c2bb55fd62ced2ec0c2fd2a4870074 59 0
-#2 := false
-decl f3 :: (-> S3 S2 S2)
-decl f6 :: S2
-#16 := f6
-decl f4 :: (-> S4 S2 S3)
-decl f7 :: S2
-#19 := f7
-decl f5 :: S4
-#7 := f5
-#21 := (f4 f5 f7)
-#22 := (f3 #21 f6)
-#18 := (f4 f5 f6)
-#20 := (f3 #18 f7)
-#23 := (= #20 #22)
-#57 := (not #23)
-#17 := (= f6 f6)
-#24 := (and #17 #23)
-#25 := (not #24)
-#58 := (iff #25 #57)
-#55 := (iff #24 #23)
-#1 := true
-#50 := (and true #23)
-#53 := (iff #50 #23)
-#54 := [rewrite]: #53
-#51 := (iff #24 #50)
-#48 := (iff #17 true)
-#49 := [rewrite]: #48
-#52 := [monotonicity #49]: #51
-#56 := [trans #52 #54]: #55
-#59 := [monotonicity #56]: #58
-#47 := [asserted]: #25
-#62 := [mp #47 #59]: #57
-#8 := (:var 1 S2)
-#10 := (:var 0 S2)
-#12 := (f4 f5 #10)
-#13 := (f3 #12 #8)
-#546 := (pattern #13)
-#9 := (f4 f5 #8)
-#11 := (f3 #9 #10)
-#545 := (pattern #11)
-#14 := (= #11 #13)
-#547 := (forall (vars (?v0 S2) (?v1 S2)) (:pat #545 #546) #14)
-#15 := (forall (vars (?v0 S2) (?v1 S2)) #14)
-#550 := (iff #15 #547)
-#548 := (iff #14 #14)
-#549 := [refl]: #548
-#551 := [quant-intro #549]: #550
-#70 := (~ #15 #15)
-#68 := (~ #14 #14)
-#69 := [refl]: #68
-#71 := [nnf-pos #69]: #70
-#46 := [asserted]: #15
-#61 := [mp~ #46 #71]: #15
-#552 := [mp #61 #551]: #547
-#130 := (not #547)
-#216 := (or #130 #23)
-#131 := [quant-inst #16 #19]: #216
-[unit-resolution #131 #552 #62]: false
-unsat
-53042978396971446eabf6039172bd47071e3fd3 67 0
-#2 := false
-decl f1 :: S1
-#3 := f1
-decl f3 :: (-> Int S1)
-decl ?v0!0 :: Int
-#55 := ?v0!0
-#56 := (f3 ?v0!0)
-#57 := (= #56 f1)
-#58 := (not #57)
-decl ?v1!1 :: Int
-#66 := ?v1!1
-#67 := (f3 ?v1!1)
-#68 := (= #67 f1)
-#69 := (or #57 #68)
-#70 := (not #69)
-#86 := (and #57 #70)
-#63 := (not #58)
-#76 := (and #63 #70)
-#87 := (iff #76 #86)
-#84 := (iff #63 #57)
-#85 := [rewrite]: #84
-#88 := [monotonicity #85]: #87
-#7 := (:var 0 Int)
-#8 := (f3 #7)
-#9 := (= #8 f1)
-#10 := (:var 1 Int)
-#11 := (f3 #10)
-#12 := (= #11 f1)
-#13 := (or #12 #9)
-#14 := (forall (vars (?v1 Int)) #13)
-#39 := (not #9)
-#40 := (or #39 #14)
-#43 := (forall (vars (?v0 Int)) #40)
-#46 := (not #43)
-#79 := (~ #46 #76)
-#50 := (or #57 #9)
-#52 := (forall (vars (?v1 Int)) #50)
-#59 := (or #58 #52)
-#60 := (not #59)
-#77 := (~ #60 #76)
-#71 := (not #52)
-#72 := (~ #71 #70)
-#73 := [sk]: #72
-#64 := (~ #63 #63)
-#65 := [refl]: #64
-#78 := [nnf-neg #65 #73]: #77
-#61 := (~ #46 #60)
-#62 := [sk]: #61
-#80 := [trans #62 #78]: #79
-#15 := (implies #9 #14)
-#16 := (forall (vars (?v0 Int)) #15)
-#17 := (not #16)
-#47 := (iff #17 #46)
-#44 := (iff #16 #43)
-#41 := (iff #15 #40)
-#42 := [rewrite]: #41
-#45 := [quant-intro #42]: #44
-#48 := [monotonicity #45]: #47
-#38 := [asserted]: #17
-#51 := [mp #38 #48]: #46
-#81 := [mp~ #51 #80]: #76
-#82 := [mp #81 #88]: #86
-#89 := [and-elim #82]: #70
-#90 := [not-or-elim #89]: #58
-#83 := [and-elim #82]: #57
-[unit-resolution #83 #90]: false
-unsat
-a69a9e8c5e31ec6b9da4cf96f47b52cf6b9404d9 117 0
-#2 := false
-decl f3 :: (-> S3 S2 S1)
-#10 := (:var 0 S2)
-decl f4 :: (-> S4 S1 S3)
-decl f6 :: S1
-#16 := f6
-decl f5 :: S4
-#7 := f5
-#17 := (f4 f5 f6)
-#18 := (f3 #17 #10)
-#573 := (pattern #18)
-decl f1 :: S1
-#3 := f1
-#19 := (= #18 f1)
-#76 := (not #19)
-#574 := (forall (vars (?v0 S2)) (:pat #573) #76)
-decl f7 :: S2
-#21 := f7
-#22 := (f3 #17 f7)
-#23 := (= #22 f1)
-#150 := (= f6 f1)
-#151 := (iff #23 #150)
-#8 := (:var 1 S1)
-#9 := (f4 f5 #8)
-#11 := (f3 #9 #10)
-#566 := (pattern #11)
-#13 := (= #8 f1)
-#12 := (= #11 f1)
-#14 := (iff #12 #13)
-#567 := (forall (vars (?v0 S1) (?v1 S2)) (:pat #566) #14)
-#15 := (forall (vars (?v0 S1) (?v1 S2)) #14)
-#570 := (iff #15 #567)
-#568 := (iff #14 #14)
-#569 := [refl]: #568
-#571 := [quant-intro #569]: #570
-#62 := (~ #15 #15)
-#60 := (~ #14 #14)
-#61 := [refl]: #60
-#63 := [nnf-pos #61]: #62
-#46 := [asserted]: #15
-#53 := [mp~ #46 #63]: #15
-#572 := [mp #53 #571]: #567
-#152 := (not #567)
-#228 := (or #152 #151)
-#561 := [quant-inst #16 #21]: #228
-#237 := [unit-resolution #561 #572]: #151
-decl ?v0!0 :: S2
-#66 := ?v0!0
-#67 := (f3 #17 ?v0!0)
-#68 := (= #67 f1)
-#236 := (iff #68 #150)
-#238 := (or #152 #236)
-#229 := [quant-inst #16 #66]: #238
-#227 := [unit-resolution #229 #572]: #236
-#240 := (not #236)
-#199 := (or #240 #150)
-#55 := (not #23)
-#215 := [hypothesis]: #55
-#83 := (or #68 #23)
-#79 := (forall (vars (?v0 S2)) #76)
-#82 := (or #79 #55)
-#84 := (and #83 #82)
-#20 := (exists (vars (?v0 S2)) #19)
-#48 := (not #20)
-#49 := (iff #48 #23)
-#85 := (~ #49 #84)
-#57 := (~ #23 #23)
-#65 := [refl]: #57
-#64 := (~ #55 #55)
-#56 := [refl]: #64
-#80 := (~ #48 #79)
-#77 := (~ #76 #76)
-#78 := [refl]: #77
-#81 := [nnf-neg #78]: #80
-#73 := (not #48)
-#74 := (~ #73 #68)
-#69 := (~ #20 #68)
-#70 := [sk]: #69
-#75 := [nnf-neg #70]: #74
-#86 := [nnf-pos #75 #81 #56 #65]: #85
-#24 := (iff #20 #23)
-#25 := (not #24)
-#50 := (iff #25 #49)
-#51 := [rewrite]: #50
-#47 := [asserted]: #25
-#54 := [mp #47 #51]: #49
-#87 := [mp~ #54 #86]: #84
-#90 := [and-elim #87]: #83
-#557 := [unit-resolution #90 #215]: #68
-#243 := (not #68)
-#222 := (or #240 #243 #150)
-#558 := [def-axiom]: #222
-#541 := [unit-resolution #558 #557]: #199
-#203 := [unit-resolution #541 #227]: #150
-#241 := (not #150)
-#562 := (not #151)
-#204 := (or #562 #241)
-#563 := (or #562 #23 #241)
-#564 := [def-axiom]: #563
-#205 := [unit-resolution #564 #215]: #204
-#206 := [unit-resolution #205 #203 #237]: false
-#543 := [lemma #206]: #23
-#579 := (or #574 #55)
-#580 := (iff #82 #579)
-#577 := (iff #79 #574)
-#575 := (iff #76 #76)
-#576 := [refl]: #575
-#578 := [quant-intro #576]: #577
-#581 := [monotonicity #578]: #580
-#91 := [and-elim #87]: #82
-#582 := [mp #91 #581]: #579
-#242 := [unit-resolution #582 #543]: #574
-#555 := (not #574)
-#214 := (or #555 #55)
-#219 := [quant-inst #21]: #214
-[unit-resolution #219 #543 #242]: false
-unsat
 d97439af6f5bc7794ab403d0f6cc318d103016a1 1288 0
 #2 := false
 decl f1 :: S1
@@ -1953,6 +1287,124 @@
 #1532 := [unit-resolution #769 #1531]: #20
 [unit-resolution #606 #1532 #1528]: false
 unsat
+a69a9e8c5e31ec6b9da4cf96f47b52cf6b9404d9 117 0
+#2 := false
+decl f3 :: (-> S3 S2 S1)
+#10 := (:var 0 S2)
+decl f4 :: (-> S4 S1 S3)
+decl f6 :: S1
+#16 := f6
+decl f5 :: S4
+#7 := f5
+#17 := (f4 f5 f6)
+#18 := (f3 #17 #10)
+#573 := (pattern #18)
+decl f1 :: S1
+#3 := f1
+#19 := (= #18 f1)
+#76 := (not #19)
+#574 := (forall (vars (?v0 S2)) (:pat #573) #76)
+decl f7 :: S2
+#21 := f7
+#22 := (f3 #17 f7)
+#23 := (= #22 f1)
+#150 := (= f6 f1)
+#151 := (iff #23 #150)
+#8 := (:var 1 S1)
+#9 := (f4 f5 #8)
+#11 := (f3 #9 #10)
+#566 := (pattern #11)
+#13 := (= #8 f1)
+#12 := (= #11 f1)
+#14 := (iff #12 #13)
+#567 := (forall (vars (?v0 S1) (?v1 S2)) (:pat #566) #14)
+#15 := (forall (vars (?v0 S1) (?v1 S2)) #14)
+#570 := (iff #15 #567)
+#568 := (iff #14 #14)
+#569 := [refl]: #568
+#571 := [quant-intro #569]: #570
+#62 := (~ #15 #15)
+#60 := (~ #14 #14)
+#61 := [refl]: #60
+#63 := [nnf-pos #61]: #62
+#46 := [asserted]: #15
+#53 := [mp~ #46 #63]: #15
+#572 := [mp #53 #571]: #567
+#152 := (not #567)
+#228 := (or #152 #151)
+#561 := [quant-inst #16 #21]: #228
+#237 := [unit-resolution #561 #572]: #151
+decl ?v0!0 :: S2
+#66 := ?v0!0
+#67 := (f3 #17 ?v0!0)
+#68 := (= #67 f1)
+#236 := (iff #68 #150)
+#238 := (or #152 #236)
+#229 := [quant-inst #16 #66]: #238
+#227 := [unit-resolution #229 #572]: #236
+#240 := (not #236)
+#199 := (or #240 #150)
+#55 := (not #23)
+#215 := [hypothesis]: #55
+#83 := (or #68 #23)
+#79 := (forall (vars (?v0 S2)) #76)
+#82 := (or #79 #55)
+#84 := (and #83 #82)
+#20 := (exists (vars (?v0 S2)) #19)
+#48 := (not #20)
+#49 := (iff #48 #23)
+#85 := (~ #49 #84)
+#57 := (~ #23 #23)
+#65 := [refl]: #57
+#64 := (~ #55 #55)
+#56 := [refl]: #64
+#80 := (~ #48 #79)
+#77 := (~ #76 #76)
+#78 := [refl]: #77
+#81 := [nnf-neg #78]: #80
+#73 := (not #48)
+#74 := (~ #73 #68)
+#69 := (~ #20 #68)
+#70 := [sk]: #69
+#75 := [nnf-neg #70]: #74
+#86 := [nnf-pos #75 #81 #56 #65]: #85
+#24 := (iff #20 #23)
+#25 := (not #24)
+#50 := (iff #25 #49)
+#51 := [rewrite]: #50
+#47 := [asserted]: #25
+#54 := [mp #47 #51]: #49
+#87 := [mp~ #54 #86]: #84
+#90 := [and-elim #87]: #83
+#557 := [unit-resolution #90 #215]: #68
+#243 := (not #68)
+#222 := (or #240 #243 #150)
+#558 := [def-axiom]: #222
+#541 := [unit-resolution #558 #557]: #199
+#203 := [unit-resolution #541 #227]: #150
+#241 := (not #150)
+#562 := (not #151)
+#204 := (or #562 #241)
+#563 := (or #562 #23 #241)
+#564 := [def-axiom]: #563
+#205 := [unit-resolution #564 #215]: #204
+#206 := [unit-resolution #205 #203 #237]: false
+#543 := [lemma #206]: #23
+#579 := (or #574 #55)
+#580 := (iff #82 #579)
+#577 := (iff #79 #574)
+#575 := (iff #76 #76)
+#576 := [refl]: #575
+#578 := [quant-intro #576]: #577
+#581 := [monotonicity #578]: #580
+#91 := [and-elim #87]: #82
+#582 := [mp #91 #581]: #579
+#242 := [unit-resolution #582 #543]: #574
+#555 := (not #574)
+#214 := (or #555 #55)
+#219 := [quant-inst #21]: #214
+[unit-resolution #219 #543 #242]: false
+unsat
 fdf61e060f49731790f4d6c8f9b26c21349c60b3 117 0
 #2 := false
 decl f1 :: S1
@@ -2071,6716 +1523,6 @@
 #603 := [unit-resolution #271 #618]: #602
 [unit-resolution #603 #601 #297]: false
 unsat
-5c792581e65682628e5c59ca9f3f8801e6aeba72 61 0
-#2 := false
-decl f1 :: S1
-#3 := f1
-decl f3 :: (-> S2 S1)
-decl f4 :: S2
-#7 := f4
-#8 := (f3 f4)
-#9 := (= #8 f1)
-decl f5 :: S2
-#18 := f5
-#19 := (f3 f5)
-#20 := (= #19 f1)
-#45 := (not #9)
-#46 := (or #45 #20)
-#49 := (not #46)
-#21 := (implies #9 #20)
-#22 := (not #21)
-#50 := (iff #22 #49)
-#47 := (iff #21 #46)
-#48 := [rewrite]: #47
-#51 := [monotonicity #48]: #50
-#44 := [asserted]: #22
-#54 := [mp #44 #51]: #49
-#52 := [not-or-elim #54]: #9
-#10 := (:var 0 S2)
-#11 := (f3 #10)
-#550 := (pattern #11)
-#12 := (= #11 f1)
-#15 := (not #12)
-#551 := (forall (vars (?v0 S2)) (:pat #550) #15)
-#16 := (forall (vars (?v0 S2)) #15)
-#554 := (iff #16 #551)
-#552 := (iff #15 #15)
-#553 := [refl]: #552
-#555 := [quant-intro #553]: #554
-#13 := (exists (vars (?v0 S2)) #12)
-#14 := (not #13)
-#60 := (~ #14 #16)
-#63 := (~ #15 #15)
-#64 := [refl]: #63
-#72 := [nnf-neg #64]: #60
-#17 := (if #9 #14 #16)
-#70 := (iff #17 #14)
-#1 := true
-#65 := (if true #14 #16)
-#68 := (iff #65 #14)
-#69 := [rewrite]: #68
-#66 := (iff #17 #65)
-#61 := (iff #9 true)
-#62 := [iff-true #52]: #61
-#67 := [monotonicity #62]: #66
-#71 := [trans #67 #69]: #70
-#43 := [asserted]: #17
-#59 := [mp #43 #71]: #14
-#57 := [mp~ #59 #72]: #16
-#556 := [mp #57 #555]: #551
-#135 := (not #551)
-#221 := (or #135 #45)
-#136 := [quant-inst #7]: #221
-[unit-resolution #136 #556 #52]: false
-unsat
-0ce3a745d60cdbf0fe26b07c5e76de09d459dd25 17 0
-#2 := false
-#7 := 3::Int
-#8 := (= 3::Int 3::Int)
-#9 := (not #8)
-#38 := (iff #9 false)
-#1 := true
-#33 := (not true)
-#36 := (iff #33 false)
-#37 := [rewrite]: #36
-#34 := (iff #9 #33)
-#31 := (iff #8 true)
-#32 := [rewrite]: #31
-#35 := [monotonicity #32]: #34
-#39 := [trans #35 #37]: #38
-#30 := [asserted]: #9
-[mp #30 #39]: false
-unsat
-1532b1dde71eb42ca0a012bb62d9bbadf37fa326 17 0
-#2 := false
-#7 := 3::Real
-#8 := (= 3::Real 3::Real)
-#9 := (not #8)
-#38 := (iff #9 false)
-#1 := true
-#33 := (not true)
-#36 := (iff #33 false)
-#37 := [rewrite]: #36
-#34 := (iff #9 #33)
-#31 := (iff #8 true)
-#32 := [rewrite]: #31
-#35 := [monotonicity #32]: #34
-#39 := [trans #35 #37]: #38
-#30 := [asserted]: #9
-[mp #30 #39]: false
-unsat
-94425abeeb45b838fcb1ab9c8323796e36a681e5 26 0
-#2 := false
-#10 := 4::Int
-#8 := 1::Int
-#7 := 3::Int
-#9 := (+ 3::Int 1::Int)
-#11 := (= #9 4::Int)
-#12 := (not #11)
-#47 := (iff #12 false)
-#1 := true
-#42 := (not true)
-#45 := (iff #42 false)
-#46 := [rewrite]: #45
-#43 := (iff #12 #42)
-#40 := (iff #11 true)
-#35 := (= 4::Int 4::Int)
-#38 := (iff #35 true)
-#39 := [rewrite]: #38
-#36 := (iff #11 #35)
-#34 := [rewrite]: #11
-#37 := [monotonicity #34]: #36
-#41 := [trans #37 #39]: #40
-#44 := [monotonicity #41]: #43
-#48 := [trans #44 #46]: #47
-#33 := [asserted]: #12
-[mp #33 #48]: false
-unsat
-673f00f23a414ea8ab1557752d859ea787c89c1b 41 0
-#2 := false
-decl f3 :: Int
-#7 := f3
-decl f5 :: Int
-#9 := f5
-#12 := (+ f5 f3)
-decl f4 :: Int
-#8 := f4
-#13 := (+ f4 #12)
-#10 := (+ f4 f5)
-#11 := (+ f3 #10)
-#14 := (= #11 #13)
-#15 := (not #14)
-#59 := (iff #15 false)
-#1 := true
-#54 := (not true)
-#57 := (iff #54 false)
-#58 := [rewrite]: #57
-#55 := (iff #15 #54)
-#52 := (iff #14 true)
-#47 := (= #11 #11)
-#50 := (iff #47 true)
-#51 := [rewrite]: #50
-#48 := (iff #14 #47)
-#45 := (= #13 #11)
-#37 := (+ f3 f5)
-#40 := (+ f4 #37)
-#43 := (= #40 #11)
-#44 := [rewrite]: #43
-#41 := (= #13 #40)
-#38 := (= #12 #37)
-#39 := [rewrite]: #38
-#42 := [monotonicity #39]: #41
-#46 := [trans #42 #44]: #45
-#49 := [monotonicity #46]: #48
-#53 := [trans #49 #51]: #52
-#56 := [monotonicity #53]: #55
-#60 := [trans #56 #58]: #59
-#36 := [asserted]: #15
-[mp #36 #60]: false
-unsat
-1f5e59fc26e6d68939e39d2fe658ebc1a264f509 35 0
-#2 := false
-#8 := 3::Int
-#9 := 8::Int
-#10 := (<= 3::Int 8::Int)
-#11 := (if #10 8::Int 3::Int)
-#7 := 5::Int
-#12 := (< 5::Int #11)
-#13 := (not #12)
-#58 := (iff #13 false)
-#1 := true
-#53 := (not true)
-#56 := (iff #53 false)
-#57 := [rewrite]: #56
-#54 := (iff #13 #53)
-#51 := (iff #12 true)
-#46 := (< 5::Int 8::Int)
-#49 := (iff #46 true)
-#50 := [rewrite]: #49
-#47 := (iff #12 #46)
-#44 := (= #11 8::Int)
-#39 := (if true 8::Int 3::Int)
-#42 := (= #39 8::Int)
-#43 := [rewrite]: #42
-#40 := (= #11 #39)
-#37 := (iff #10 true)
-#38 := [rewrite]: #37
-#41 := [monotonicity #38]: #40
-#45 := [trans #41 #43]: #44
-#48 := [monotonicity #45]: #47
-#52 := [trans #48 #50]: #51
-#55 := [monotonicity #52]: #54
-#59 := [trans #55 #57]: #58
-#34 := [asserted]: #13
-[mp #34 #59]: false
-unsat
-e7f019160a38d08774f8a2e816f96aa54c924fba 216 0
-#2 := false
-#10 := 0::Real
-decl f4 :: Real
-#8 := f4
-#43 := -1::Real
-#45 := (* -1::Real f4)
-decl f3 :: Real
-#7 := f3
-#44 := (* -1::Real f3)
-#46 := (+ #44 #45)
-#9 := (+ f3 f4)
-#71 := (>= #9 0::Real)
-#78 := (if #71 #9 #46)
-#153 := (* -1::Real #78)
-#181 := (+ #46 #153)
-#183 := (>= #181 0::Real)
-#134 := (= #46 #78)
-#72 := (not #71)
-#95 := (>= f4 0::Real)
-#96 := (not #95)
-#154 := (+ #9 #153)
-#156 := (>= #154 0::Real)
-#133 := (= #9 #78)
-#197 := (not #134)
-#192 := (not #183)
-#163 := [hypothesis]: #95
-#193 := (or #192 #96)
-#184 := [hypothesis]: #183
-#102 := (if #95 f4 #45)
-#114 := (* -1::Real #102)
-#83 := (>= f3 0::Real)
-#90 := (if #83 f3 #44)
-#113 := (* -1::Real #90)
-#115 := (+ #113 #114)
-#116 := (+ #78 #115)
-#117 := (<= #116 0::Real)
-#122 := (not #117)
-#18 := (- f4)
-#17 := (< f4 0::Real)
-#19 := (if #17 #18 f4)
-#15 := (- f3)
-#14 := (< f3 0::Real)
-#16 := (if #14 #15 f3)
-#20 := (+ #16 #19)
-#12 := (- #9)
-#11 := (< #9 0::Real)
-#13 := (if #11 #12 #9)
-#21 := (<= #13 #20)
-#22 := (not #21)
-#125 := (iff #22 #122)
-#59 := (if #17 #45 f4)
-#54 := (if #14 #44 f3)
-#62 := (+ #54 #59)
-#49 := (if #11 #46 #9)
-#65 := (<= #49 #62)
-#68 := (not #65)
-#123 := (iff #68 #122)
-#120 := (iff #65 #117)
-#107 := (+ #90 #102)
-#110 := (<= #78 #107)
-#118 := (iff #110 #117)
-#119 := [rewrite]: #118
-#111 := (iff #65 #110)
-#108 := (= #62 #107)
-#105 := (= #59 #102)
-#99 := (if #96 #45 f4)
-#103 := (= #99 #102)
-#104 := [rewrite]: #103
-#100 := (= #59 #99)
-#97 := (iff #17 #96)
-#98 := [rewrite]: #97
-#101 := [monotonicity #98]: #100
-#106 := [trans #101 #104]: #105
-#93 := (= #54 #90)
-#84 := (not #83)
-#87 := (if #84 #44 f3)
-#91 := (= #87 #90)
-#92 := [rewrite]: #91
-#88 := (= #54 #87)
-#85 := (iff #14 #84)
-#86 := [rewrite]: #85
-#89 := [monotonicity #86]: #88
-#94 := [trans #89 #92]: #93
-#109 := [monotonicity #94 #106]: #108
-#81 := (= #49 #78)
-#75 := (if #72 #46 #9)
-#79 := (= #75 #78)
-#80 := [rewrite]: #79
-#76 := (= #49 #75)
-#73 := (iff #11 #72)
-#74 := [rewrite]: #73
-#77 := [monotonicity #74]: #76
-#82 := [trans #77 #80]: #81
-#112 := [monotonicity #82 #109]: #111
-#121 := [trans #112 #119]: #120
-#124 := [monotonicity #121]: #123
-#69 := (iff #22 #68)
-#66 := (iff #21 #65)
-#63 := (= #20 #62)
-#60 := (= #19 #59)
-#57 := (= #18 #45)
-#58 := [rewrite]: #57
-#61 := [monotonicity #58]: #60
-#55 := (= #16 #54)
-#52 := (= #15 #44)
-#53 := [rewrite]: #52
-#56 := [monotonicity #53]: #55
-#64 := [monotonicity #56 #61]: #63
-#50 := (= #13 #49)
-#47 := (= #12 #46)
-#48 := [rewrite]: #47
-#51 := [monotonicity #48]: #50
-#67 := [monotonicity #51 #64]: #66
-#70 := [monotonicity #67]: #69
-#126 := [trans #70 #124]: #125
-#42 := [asserted]: #22
-#127 := [mp #42 #126]: #122
-#147 := (+ f4 #114)
-#148 := (<= #147 0::Real)
-#141 := (= f4 #102)
-#143 := (or #96 #141)
-#144 := [def-axiom]: #143
-#172 := [unit-resolution #144 #163]: #141
-#173 := (not #141)
-#174 := (or #173 #148)
-#175 := [th-lemma arith triangle-eq]: #174
-#176 := [unit-resolution #175 #172]: #148
-#152 := (+ #44 #113)
-#155 := (<= #152 0::Real)
-#130 := (= #44 #90)
-#178 := (or #84 #96)
-#150 := (+ f3 #113)
-#151 := (<= #150 0::Real)
-#129 := (= f3 #90)
-#157 := [hypothesis]: #83
-#137 := (or #84 #129)
-#138 := [def-axiom]: #137
-#158 := [unit-resolution #138 #157]: #129
-#159 := (not #129)
-#160 := (or #159 #151)
-#161 := [th-lemma arith triangle-eq]: #160
-#162 := [unit-resolution #161 #158]: #151
-#164 := (or #71 #84 #96)
-#165 := [th-lemma arith assign-bounds -1 -1]: #164
-#166 := [unit-resolution #165 #157 #163]: #71
-#135 := (or #72 #133)
-#136 := [def-axiom]: #135
-#167 := [unit-resolution #136 #166]: #133
-#168 := (not #133)
-#169 := (or #168 #156)
-#170 := [th-lemma arith triangle-eq]: #169
-#171 := [unit-resolution #170 #167]: #156
-#177 := [th-lemma arith farkas 1 -1 -1 1 #176 #171 #127 #162]: false
-#179 := [lemma #177]: #178
-#185 := [unit-resolution #179 #163]: #84
-#139 := (or #83 #130)
-#140 := [def-axiom]: #139
-#186 := [unit-resolution #140 #185]: #130
-#187 := (not #130)
-#188 := (or #187 #155)
-#189 := [th-lemma arith triangle-eq]: #188
-#190 := [unit-resolution #189 #186]: #155
-#191 := [th-lemma arith farkas 2 -1 -1 1 1 #163 #190 #176 #127 #184]: false
-#194 := [lemma #191]: #193
-#202 := [unit-resolution #194 #163]: #192
-#198 := (or #197 #183)
-#195 := [hypothesis]: #192
-#196 := [hypothesis]: #134
-#199 := [th-lemma arith triangle-eq]: #198
-#200 := [unit-resolution #199 #196 #195]: false
-#201 := [lemma #200]: #198
-#203 := [unit-resolution #201 #202]: #197
-#131 := (or #71 #134)
-#132 := [def-axiom]: #131
-#204 := [unit-resolution #132 #203]: #71
-#205 := [unit-resolution #136 #204]: #133
-#206 := [unit-resolution #170 #205]: #156
-#207 := [th-lemma arith farkas 2 1 1 1 1 #185 #190 #176 #127 #206]: false
-#208 := [lemma #207]: #96
-#149 := (+ #45 #114)
-#180 := (<= #149 0::Real)
-#142 := (= #45 #102)
-#145 := (or #95 #142)
-#146 := [def-axiom]: #145
-#213 := [unit-resolution #146 #208]: #142
-#214 := (not #142)
-#215 := (or #214 #180)
-#216 := [th-lemma arith triangle-eq]: #215
-#217 := [unit-resolution #216 #213]: #180
-#219 := (not #156)
-#220 := (not #151)
-#221 := (or #219 #220)
-#211 := [hypothesis]: #151
-#212 := [hypothesis]: #156
-#218 := [th-lemma arith farkas 2 1 1 1 1 #208 #217 #127 #212 #211]: false
-#222 := [lemma #218]: #221
-#227 := [unit-resolution #222 #162]: #219
-#223 := [hypothesis]: #219
-#224 := [hypothesis]: #133
-#225 := [unit-resolution #170 #224 #223]: false
-#226 := [lemma #225]: #169
-#228 := [unit-resolution #226 #227]: #168
-#229 := [unit-resolution #136 #228]: #72
-#230 := [unit-resolution #132 #229]: #134
-#231 := [unit-resolution #201 #230]: #183
-#232 := [th-lemma arith farkas 1/2 -1/2 -1/2 1/2 1 #231 #162 #217 #127 #157]: false
-#233 := [lemma #232]: #84
-#234 := (or #72 #83 #95)
-#235 := [th-lemma arith assign-bounds 1 1]: #234
-#236 := [unit-resolution #235 #233 #208]: #72
-#237 := [unit-resolution #132 #236]: #134
-#238 := [unit-resolution #201 #237]: #183
-#239 := [unit-resolution #140 #233]: #130
-#240 := [unit-resolution #189 #239]: #155
-[th-lemma arith farkas -1 -1 1 1 #240 #217 #127 #238]: false
-unsat
-9e5f324cc33eb4abf1be11d977dfdec45557ae46 42 0
-#2 := false
-decl f3 :: (-> S1 S2)
-decl f1 :: S1
-#3 := f1
-#12 := (f3 f1)
-decl f2 :: S1
-#4 := f2
-#8 := 3::Int
-#7 := 2::Int
-#9 := (< 2::Int 3::Int)
-#10 := (if #9 f1 f2)
-#11 := (f3 #10)
-#13 := (= #11 #12)
-#14 := (not #13)
-#60 := (iff #14 false)
-#1 := true
-#55 := (not true)
-#58 := (iff #55 false)
-#59 := [rewrite]: #58
-#56 := (iff #14 #55)
-#53 := (iff #13 true)
-#48 := (= #12 #12)
-#51 := (iff #48 true)
-#52 := [rewrite]: #51
-#49 := (iff #13 #48)
-#45 := (= #10 f1)
-#40 := (if true f1 f2)
-#43 := (= #40 f1)
-#44 := [rewrite]: #43
-#41 := (= #10 #40)
-#38 := (iff #9 true)
-#39 := [rewrite]: #38
-#42 := [monotonicity #39]: #41
-#46 := [trans #42 #44]: #45
-#47 := [monotonicity #46]: #13
-#50 := [monotonicity #47]: #49
-#54 := [trans #50 #52]: #53
-#57 := [monotonicity #54]: #56
-#61 := [trans #57 #59]: #60
-#35 := [asserted]: #14
-[mp #35 #61]: false
-unsat
-cc322c3513bba37f77e905b379b26c79239b69a4 49 0
-#2 := false
-#12 := 1::Int
-decl f3 :: Int
-#8 := f3
-#13 := (< f3 1::Int)
-#9 := 3::Int
-#10 := (+ f3 3::Int)
-#7 := 4::Int
-#11 := (<= 4::Int #10)
-#14 := (or #11 #13)
-#15 := (not #14)
-#69 := (iff #15 false)
-#37 := (+ 3::Int f3)
-#40 := (<= 4::Int #37)
-#43 := (or #40 #13)
-#46 := (not #43)
-#67 := (iff #46 false)
-#1 := true
-#62 := (not true)
-#65 := (iff #62 false)
-#66 := [rewrite]: #65
-#63 := (iff #46 #62)
-#60 := (iff #43 true)
-#51 := (>= f3 1::Int)
-#52 := (not #51)
-#55 := (or #51 #52)
-#58 := (iff #55 true)
-#59 := [rewrite]: #58
-#56 := (iff #43 #55)
-#53 := (iff #13 #52)
-#54 := [rewrite]: #53
-#49 := (iff #40 #51)
-#50 := [rewrite]: #49
-#57 := [monotonicity #50 #54]: #56
-#61 := [trans #57 #59]: #60
-#64 := [monotonicity #61]: #63
-#68 := [trans #64 #66]: #67
-#47 := (iff #15 #46)
-#44 := (iff #14 #43)
-#41 := (iff #11 #40)
-#38 := (= #10 #37)
-#39 := [rewrite]: #38
-#42 := [monotonicity #39]: #41
-#45 := [monotonicity #42]: #44
-#48 := [monotonicity #45]: #47
-#70 := [trans #48 #68]: #69
-#36 := [asserted]: #15
-[mp #36 #70]: false
-unsat
-75c4589e7d7ab0bf262babccc302883b71f9a923 63 0
-#2 := false
-#14 := 0::Int
-decl f4 :: Int
-#10 := f4
-#49 := -1::Int
-#52 := (* -1::Int f4)
-decl f3 :: Int
-#8 := f3
-#53 := (+ f3 #52)
-#70 := (>= #53 0::Int)
-#94 := (iff #70 false)
-#51 := -4::Int
-#87 := (>= -4::Int 0::Int)
-#86 := (iff #87 false)
-#93 := [rewrite]: #86
-#88 := (iff #70 #87)
-#54 := (= #53 -4::Int)
-#11 := 4::Int
-#12 := (+ f3 4::Int)
-#13 := (= f4 #12)
-#56 := (iff #13 #54)
-#39 := (+ 4::Int f3)
-#46 := (= f4 #39)
-#50 := (iff #46 #54)
-#55 := [rewrite]: #50
-#47 := (iff #13 #46)
-#44 := (= #12 #39)
-#45 := [rewrite]: #44
-#48 := [monotonicity #45]: #47
-#57 := [trans #48 #55]: #56
-#38 := [asserted]: #13
-#58 := [mp #38 #57]: #54
-#85 := [monotonicity #58]: #88
-#95 := [trans #85 #93]: #94
-#15 := (- f4 f3)
-#16 := (< 0::Int #15)
-#17 := (not #16)
-#81 := (iff #17 #70)
-#60 := (* -1::Int f3)
-#61 := (+ #60 f4)
-#64 := (< 0::Int #61)
-#67 := (not #64)
-#79 := (iff #67 #70)
-#71 := (not #70)
-#74 := (not #71)
-#77 := (iff #74 #70)
-#78 := [rewrite]: #77
-#75 := (iff #67 #74)
-#72 := (iff #64 #71)
-#73 := [rewrite]: #72
-#76 := [monotonicity #73]: #75
-#80 := [trans #76 #78]: #79
-#68 := (iff #17 #67)
-#65 := (iff #16 #64)
-#62 := (= #15 #61)
-#63 := [rewrite]: #62
-#66 := [monotonicity #63]: #65
-#69 := [monotonicity #66]: #68
-#82 := [trans #69 #80]: #81
-#59 := [asserted]: #17
-#83 := [mp #59 #82]: #70
-[mp #83 #95]: false
-unsat
-31769d5312feac1587c3f744c5c881fb2d86e85f 35 0
-#2 := false
-#9 := 5::Int
-#7 := 2::Int
-#8 := (+ 2::Int 2::Int)
-#10 := (= #8 5::Int)
-#11 := (not #10)
-#12 := (not #11)
-#56 := (iff #12 false)
-#1 := true
-#51 := (not true)
-#54 := (iff #51 false)
-#55 := [rewrite]: #54
-#52 := (iff #12 #51)
-#49 := (iff #11 true)
-#44 := (not false)
-#47 := (iff #44 true)
-#48 := [rewrite]: #47
-#45 := (iff #11 #44)
-#42 := (iff #10 false)
-#34 := 4::Int
-#37 := (= 4::Int 5::Int)
-#40 := (iff #37 false)
-#41 := [rewrite]: #40
-#38 := (iff #10 #37)
-#35 := (= #8 4::Int)
-#36 := [rewrite]: #35
-#39 := [monotonicity #36]: #38
-#43 := [trans #39 #41]: #42
-#46 := [monotonicity #43]: #45
-#50 := [trans #46 #48]: #49
-#53 := [monotonicity #50]: #52
-#57 := [trans #53 #55]: #56
-#33 := [asserted]: #12
-[mp #33 #57]: false
-unsat
-f8ba8c3ed7f7c7d5e49139b62e145fc6eee338f1 45 0
-#2 := false
-#14 := 4::Real
-decl f4 :: Real
-#11 := f4
-#10 := 7::Real
-#12 := (* 7::Real f4)
-decl f3 :: Real
-#8 := f3
-#7 := 3::Real
-#9 := (* 3::Real f3)
-#13 := (+ #9 #12)
-#48 := (>= #13 4::Real)
-#46 := (not #48)
-#15 := (< #13 4::Real)
-#47 := (iff #15 #46)
-#44 := [rewrite]: #47
-#41 := [asserted]: #15
-#45 := [mp #41 #44]: #46
-#16 := 2::Real
-#17 := (* 2::Real f3)
-#50 := (<= #17 3::Real)
-#51 := (not #50)
-#18 := (< 3::Real #17)
-#52 := (iff #18 #51)
-#53 := [rewrite]: #52
-#42 := [asserted]: #18
-#54 := [mp #42 #53]: #51
-#19 := 0::Real
-#58 := (>= f4 0::Real)
-#20 := (< f4 0::Real)
-#21 := (not #20)
-#65 := (iff #21 #58)
-#56 := (not #58)
-#60 := (not #56)
-#63 := (iff #60 #58)
-#64 := [rewrite]: #63
-#61 := (iff #21 #60)
-#57 := (iff #20 #56)
-#59 := [rewrite]: #57
-#62 := [monotonicity #59]: #61
-#66 := [trans #62 #64]: #65
-#43 := [asserted]: #21
-#67 := [mp #43 #66]: #58
-[th-lemma arith farkas 7 3/2 1 #67 #54 #45]: false
-unsat
-c61600e5a5dab4b2c2864caededa0b50f81df696 59 0
-#2 := false
-#19 := (not false)
-decl f4 :: Int
-#11 := f4
-#7 := 0::Int
-#15 := (<= 0::Int f4)
-#16 := (not #15)
-#17 := (or #16 #15)
-#9 := 1::Int
-#10 := (- 1::Int)
-#12 := (* #10 f4)
-decl f3 :: Int
-#8 := f3
-#13 := (+ f3 #12)
-#14 := (<= 0::Int #13)
-#18 := (or #14 #17)
-#20 := (iff #18 #19)
-#21 := (not #20)
-#77 := (iff #21 false)
-#1 := true
-#72 := (not true)
-#75 := (iff #72 false)
-#76 := [rewrite]: #75
-#73 := (iff #21 #72)
-#70 := (iff #20 true)
-#65 := (iff true true)
-#68 := (iff #65 true)
-#69 := [rewrite]: #68
-#66 := (iff #20 #65)
-#63 := (iff #19 true)
-#64 := [rewrite]: #63
-#61 := (iff #18 true)
-#42 := -1::Int
-#45 := (* -1::Int f4)
-#48 := (+ f3 #45)
-#51 := (<= 0::Int #48)
-#56 := (or #51 true)
-#59 := (iff #56 true)
-#60 := [rewrite]: #59
-#57 := (iff #18 #56)
-#54 := (iff #17 true)
-#55 := [rewrite]: #54
-#52 := (iff #14 #51)
-#49 := (= #13 #48)
-#46 := (= #12 #45)
-#43 := (= #10 -1::Int)
-#44 := [rewrite]: #43
-#47 := [monotonicity #44]: #46
-#50 := [monotonicity #47]: #49
-#53 := [monotonicity #50]: #52
-#58 := [monotonicity #53 #55]: #57
-#62 := [trans #58 #60]: #61
-#67 := [monotonicity #62 #64]: #66
-#71 := [trans #67 #69]: #70
-#74 := [monotonicity #71]: #73
-#78 := [trans #74 #76]: #77
-#41 := [asserted]: #21
-[mp #41 #78]: false
-unsat
-7f98d11cd70eeb0eb4aea9722e1648cd3cfdbe2c 439 0
-#2 := false
-decl f4 :: Int
-#8 := f4
-decl f3 :: Int
-#7 := f3
-#20 := (= f3 f4)
-#287 := (not #20)
-#24 := (= f4 f3)
-#312 := (not #24)
-#499 := (iff #312 #287)
-#458 := (iff #24 #20)
-#459 := [commutativity]: #458
-#500 := [monotonicity #459]: #499
-decl f5 :: Int
-#10 := f5
-#30 := (= f5 f4)
-#13 := (= f4 f5)
-#493 := (iff #13 #30)
-#491 := (iff #30 #13)
-#492 := [commutativity]: #491
-#494 := [symm #492]: #493
-#18 := (= f3 f5)
-#238 := (not #18)
-#28 := (= f5 f3)
-#337 := (not #28)
-#485 := (iff #337 #238)
-#483 := (iff #28 #18)
-#484 := [commutativity]: #483
-#486 := [monotonicity #484]: #485
-#55 := 0::Int
-#77 := -1::Int
-#102 := (* -1::Int f4)
-#103 := (+ f3 #102)
-#104 := (<= #103 0::Int)
-#105 := (not #104)
-#118 := (>= #103 0::Int)
-#78 := (* -1::Int f5)
-#96 := (+ f4 #78)
-#95 := (>= #96 0::Int)
-#94 := (not #95)
-#261 := (not #13)
-#435 := [hypothesis]: #261
-#127 := (<= #96 0::Int)
-#474 := (or #18 #13)
-#441 := [hypothesis]: #238
-#447 := (or #104 #18 #13)
-#436 := [hypothesis]: #105
-#300 := (or #127 #104)
-#128 := (not #127)
-#134 := (and #128 #105)
-#216 := (not #134)
-#309 := (iff #216 #300)
-#301 := (not #300)
-#304 := (not #301)
-#307 := (iff #304 #300)
-#308 := [rewrite]: #307
-#305 := (iff #216 #304)
-#302 := (iff #134 #301)
-#303 := [rewrite]: #302
-#306 := [monotonicity #303]: #305
-#310 := [trans #306 #308]: #309
-#37 := (and #30 #24)
-#79 := (+ f3 #78)
-#80 := (<= #79 0::Int)
-#81 := (not #80)
-#84 := (and #13 #81)
-#88 := (>= #79 0::Int)
-#87 := (not #88)
-#91 := (and #24 #87)
-#99 := (and #94 #81)
-#108 := (and #105 #28)
-#111 := (and #105 #87)
-#114 := (and #30 #105)
-#117 := (not #118)
-#121 := (and #28 #117)
-#124 := (and #81 #117)
-#131 := (and #128 #24)
-#137 := (and #20 #94)
-#140 := (and #18 #128)
-#143 := (and #87 #128)
-#146 := (and #117 #13)
-#149 := (and #117 #94)
-#197 := (or #149 #146 #143 #140 #137 #134 #131 #124 #121 #114 #111 #108 #99 #91 #84 #37)
-#202 := (not #197)
-#26 := (< f5 f3)
-#36 := (and #13 #26)
-#38 := (or #36 #37)
-#15 := (< f3 f5)
-#35 := (and #24 #15)
-#39 := (or #35 #38)
-#11 := (< f4 f5)
-#34 := (and #11 #26)
-#40 := (or #34 #39)
-#22 := (< f4 f3)
-#33 := (and #22 #28)
-#41 := (or #33 #40)
-#32 := (and #22 #15)
-#42 := (or #32 #41)
-#31 := (and #30 #22)
-#43 := (or #31 #42)
-#9 := (< f3 f4)
-#29 := (and #28 #9)
-#44 := (or #29 #43)
-#27 := (and #26 #9)
-#45 := (or #27 #44)
-#16 := (< f5 f4)
-#25 := (and #16 #24)
-#46 := (or #25 #45)
-#23 := (and #16 #22)
-#47 := (or #23 #46)
-#21 := (and #20 #11)
-#48 := (or #21 #47)
-#19 := (and #18 #16)
-#49 := (or #19 #48)
-#17 := (and #15 #16)
-#50 := (or #17 #49)
-#14 := (and #9 #13)
-#51 := (or #14 #50)
-#12 := (and #9 #11)
-#52 := (or #12 #51)
-#53 := (not #52)
-#203 := (iff #53 #202)
-#200 := (iff #52 #197)
-#152 := (or #84 #37)
-#155 := (or #91 #152)
-#158 := (or #99 #155)
-#161 := (or #108 #158)
-#164 := (or #111 #161)
-#167 := (or #114 #164)
-#170 := (or #121 #167)
-#173 := (or #124 #170)
-#176 := (or #131 #173)
-#179 := (or #134 #176)
-#182 := (or #137 #179)
-#185 := (or #140 #182)
-#188 := (or #143 #185)
-#191 := (or #146 #188)
-#194 := (or #149 #191)
-#198 := (iff #194 #197)
-#199 := [rewrite]: #198
-#195 := (iff #52 #194)
-#192 := (iff #51 #191)
-#189 := (iff #50 #188)
-#186 := (iff #49 #185)
-#183 := (iff #48 #182)
-#180 := (iff #47 #179)
-#177 := (iff #46 #176)
-#174 := (iff #45 #173)
-#171 := (iff #44 #170)
-#168 := (iff #43 #167)
-#165 := (iff #42 #164)
-#162 := (iff #41 #161)
-#159 := (iff #40 #158)
-#156 := (iff #39 #155)
-#153 := (iff #38 #152)
-#85 := (iff #36 #84)
-#82 := (iff #26 #81)
-#83 := [rewrite]: #82
-#86 := [monotonicity #83]: #85
-#154 := [monotonicity #86]: #153
-#92 := (iff #35 #91)
-#89 := (iff #15 #87)
-#90 := [rewrite]: #89
-#93 := [monotonicity #90]: #92
-#157 := [monotonicity #93 #154]: #156
-#100 := (iff #34 #99)
-#97 := (iff #11 #94)
-#98 := [rewrite]: #97
-#101 := [monotonicity #98 #83]: #100
-#160 := [monotonicity #101 #157]: #159
-#109 := (iff #33 #108)
-#106 := (iff #22 #105)
-#107 := [rewrite]: #106
-#110 := [monotonicity #107]: #109
-#163 := [monotonicity #110 #160]: #162
-#112 := (iff #32 #111)
-#113 := [monotonicity #107 #90]: #112
-#166 := [monotonicity #113 #163]: #165
-#115 := (iff #31 #114)
-#116 := [monotonicity #107]: #115
-#169 := [monotonicity #116 #166]: #168
-#122 := (iff #29 #121)
-#119 := (iff #9 #117)
-#120 := [rewrite]: #119
-#123 := [monotonicity #120]: #122
-#172 := [monotonicity #123 #169]: #171
-#125 := (iff #27 #124)
-#126 := [monotonicity #83 #120]: #125
-#175 := [monotonicity #126 #172]: #174
-#132 := (iff #25 #131)
-#129 := (iff #16 #128)
-#130 := [rewrite]: #129
-#133 := [monotonicity #130]: #132
-#178 := [monotonicity #133 #175]: #177
-#135 := (iff #23 #134)
-#136 := [monotonicity #130 #107]: #135
-#181 := [monotonicity #136 #178]: #180
-#138 := (iff #21 #137)
-#139 := [monotonicity #98]: #138
-#184 := [monotonicity #139 #181]: #183
-#141 := (iff #19 #140)
-#142 := [monotonicity #130]: #141
-#187 := [monotonicity #142 #184]: #186
-#144 := (iff #17 #143)
-#145 := [monotonicity #90 #130]: #144
-#190 := [monotonicity #145 #187]: #189
-#147 := (iff #14 #146)
-#148 := [monotonicity #120]: #147
-#193 := [monotonicity #148 #190]: #192
-#150 := (iff #12 #149)
-#151 := [monotonicity #120 #98]: #150
-#196 := [monotonicity #151 #193]: #195
-#201 := [trans #196 #199]: #200
-#204 := [monotonicity #201]: #203
-#74 := [asserted]: #53
-#205 := [mp #74 #204]: #202
-#217 := [not-or-elim #205]: #216
-#311 := [mp #217 #310]: #300
-#437 := [unit-resolution #311 #436]: #127
-#438 := (or #13 #128 #94)
-#439 := [th-lemma arith triangle-eq]: #438
-#440 := [unit-resolution #439 #437 #435]: #94
-#363 := (or #104 #88)
-#226 := (not #111)
-#372 := (iff #226 #363)
-#364 := (not #363)
-#367 := (not #364)
-#370 := (iff #367 #363)
-#371 := [rewrite]: #370
-#368 := (iff #226 #367)
-#365 := (iff #111 #364)
-#366 := [rewrite]: #365
-#369 := [monotonicity #366]: #368
-#373 := [trans #369 #371]: #372
-#227 := [not-or-elim #205]: #226
-#374 := [mp #227 #373]: #363
-#442 := [unit-resolution #374 #436]: #88
-#443 := (or #18 #81 #87)
-#444 := [th-lemma arith triangle-eq]: #443
-#445 := [unit-resolution #444 #442 #441]: #81
-#387 := (or #95 #80)
-#230 := (not #99)
-#396 := (iff #230 #387)
-#388 := (not #387)
-#391 := (not #388)
-#394 := (iff #391 #387)
-#395 := [rewrite]: #394
-#392 := (iff #230 #391)
-#389 := (iff #99 #388)
-#390 := [rewrite]: #389
-#393 := [monotonicity #390]: #392
-#397 := [trans #393 #395]: #396
-#231 := [not-or-elim #205]: #230
-#398 := [mp #231 #397]: #387
-#446 := [unit-resolution #398 #445 #440]: false
-#448 := [lemma #446]: #447
-#466 := [unit-resolution #448 #441 #435]: #104
-#464 := (or #80 #13 #105)
-#460 := (iff #20 #24)
-#461 := [symm #459]: #460
-#453 := [hypothesis]: #104
-#449 := [hypothesis]: #81
-#325 := (or #80 #118)
-#220 := (not #124)
-#334 := (iff #220 #325)
-#326 := (not #325)
-#329 := (not #326)
-#332 := (iff #329 #325)
-#333 := [rewrite]: #332
-#330 := (iff #220 #329)
-#327 := (iff #124 #326)
-#328 := [rewrite]: #327
-#331 := [monotonicity #328]: #330
-#335 := [trans #331 #333]: #334
-#221 := [not-or-elim #205]: #220
-#336 := [mp #221 #335]: #325
-#454 := [unit-resolution #336 #449]: #118
-#455 := (or #20 #105 #117)
-#456 := [th-lemma arith triangle-eq]: #455
-#457 := [unit-resolution #456 #454 #453]: #20
-#462 := [mp #457 #461]: #24
-#450 := [unit-resolution #398 #449]: #95
-#451 := [unit-resolution #439 #450 #435]: #128
-#313 := (or #127 #312)
-#218 := (not #131)
-#322 := (iff #218 #313)
-#314 := (not #313)
-#317 := (not #314)
-#320 := (iff #317 #313)
-#321 := [rewrite]: #320
-#318 := (iff #218 #317)
-#315 := (iff #131 #314)
-#316 := [rewrite]: #315
-#319 := [monotonicity #316]: #318
-#323 := [trans #319 #321]: #322
-#219 := [not-or-elim #205]: #218
-#324 := [mp #219 #323]: #313
-#452 := [unit-resolution #324 #451]: #312
-#463 := [unit-resolution #452 #462]: false
-#465 := [lemma #463]: #464
-#467 := [unit-resolution #465 #466 #435]: #80
-#468 := [unit-resolution #444 #467 #441]: #87
-#250 := (or #88 #127)
-#210 := (not #143)
-#239 := (iff #210 #250)
-#247 := (not #250)
-#246 := (not #247)
-#241 := (iff #246 #250)
-#242 := [rewrite]: #241
-#243 := (iff #210 #246)
-#248 := (iff #143 #247)
-#245 := [rewrite]: #248
-#244 := [monotonicity #245]: #243
-#240 := [trans #244 #242]: #239
-#211 := [not-or-elim #205]: #210
-#76 := [mp #211 #240]: #250
-#469 := [unit-resolution #76 #468]: #127
-#470 := [unit-resolution #439 #469 #435]: #94
-#271 := (or #118 #95)
-#206 := (not #149)
-#266 := (iff #206 #271)
-#272 := (not #271)
-#269 := (not #272)
-#268 := (iff #269 #271)
-#265 := [rewrite]: #268
-#270 := (iff #206 #269)
-#273 := (iff #149 #272)
-#274 := [rewrite]: #273
-#267 := [monotonicity #274]: #270
-#263 := [trans #267 #265]: #266
-#207 := [not-or-elim #205]: #206
-#264 := [mp #207 #263]: #271
-#471 := [unit-resolution #264 #470]: #118
-#288 := (or #287 #95)
-#214 := (not #137)
-#297 := (iff #214 #288)
-#289 := (not #288)
-#292 := (not #289)
-#295 := (iff #292 #288)
-#296 := [rewrite]: #295
-#293 := (iff #214 #292)
-#290 := (iff #137 #289)
-#291 := [rewrite]: #290
-#294 := [monotonicity #291]: #293
-#298 := [trans #294 #296]: #297
-#215 := [not-or-elim #205]: #214
-#299 := [mp #215 #298]: #288
-#472 := [unit-resolution #299 #470]: #287
-#473 := [unit-resolution #456 #472 #471 #466]: false
-#475 := [lemma #473]: #474
-#476 := [unit-resolution #475 #435]: #18
-#275 := (or #238 #127)
-#212 := (not #140)
-#284 := (iff #212 #275)
-#276 := (not #275)
-#279 := (not #276)
-#282 := (iff #279 #275)
-#283 := [rewrite]: #282
-#280 := (iff #212 #279)
-#277 := (iff #140 #276)
-#278 := [rewrite]: #277
-#281 := [monotonicity #278]: #280
-#285 := [trans #281 #283]: #284
-#213 := [not-or-elim #205]: #212
-#286 := [mp #213 #285]: #275
-#477 := [unit-resolution #286 #476]: #127
-#478 := [unit-resolution #439 #477 #435]: #94
-#479 := [unit-resolution #264 #478]: #118
-#480 := [unit-resolution #299 #478]: #287
-#481 := [unit-resolution #456 #480 #479]: #105
-#375 := (or #104 #337)
-#228 := (not #108)
-#384 := (iff #228 #375)
-#376 := (not #375)
-#379 := (not #376)
-#382 := (iff #379 #375)
-#383 := [rewrite]: #382
-#380 := (iff #228 #379)
-#377 := (iff #108 #376)
-#378 := [rewrite]: #377
-#381 := [monotonicity #378]: #380
-#385 := [trans #381 #383]: #384
-#229 := [not-or-elim #205]: #228
-#386 := [mp #229 #385]: #375
-#482 := [unit-resolution #386 #481]: #337
-#487 := [mp #482 #486]: #238
-#488 := [unit-resolution #476 #487]: false
-#489 := [lemma #488]: #13
-#495 := [mp #489 #494]: #30
-#350 := (not #30)
-#423 := (or #350 #312)
-#236 := (not #37)
-#432 := (iff #236 #423)
-#424 := (not #423)
-#427 := (not #424)
-#430 := (iff #427 #423)
-#431 := [rewrite]: #430
-#428 := (iff #236 #427)
-#425 := (iff #37 #424)
-#426 := [rewrite]: #425
-#429 := [monotonicity #426]: #428
-#433 := [trans #429 #431]: #432
-#237 := [not-or-elim #205]: #236
-#434 := [mp #237 #433]: #423
-#498 := [unit-resolution #434 #495]: #312
-#501 := [mp #498 #500]: #287
-#262 := (or #118 #261)
-#208 := (not #146)
-#251 := (iff #208 #262)
-#259 := (not #262)
-#258 := (not #259)
-#253 := (iff #258 #262)
-#254 := [rewrite]: #253
-#255 := (iff #208 #258)
-#260 := (iff #146 #259)
-#257 := [rewrite]: #260
-#256 := [monotonicity #257]: #255
-#252 := [trans #256 #254]: #251
-#209 := [not-or-elim #205]: #208
-#249 := [mp #209 #252]: #262
-#490 := [unit-resolution #249 #489]: #118
-#351 := (or #350 #104)
-#224 := (not #114)
-#360 := (iff #224 #351)
-#352 := (not #351)
-#355 := (not #352)
-#358 := (iff #355 #351)
-#359 := [rewrite]: #358
-#356 := (iff #224 #355)
-#353 := (iff #114 #352)
-#354 := [rewrite]: #353
-#357 := [monotonicity #354]: #356
-#361 := [trans #357 #359]: #360
-#225 := [not-or-elim #205]: #224
-#362 := [mp #225 #361]: #351
-#496 := [unit-resolution #362 #495]: #104
-#497 := [unit-resolution #456 #496 #490]: #20
-[unit-resolution #497 #501]: false
-unsat
-70bd6436662c1fd4b8c8a6f696914593051990e6 52 0
-#2 := false
-#11 := 1::Real
-decl f3 :: Real
-#7 := f3
-#9 := 2::Real
-#10 := (* 2::Real f3)
-#12 := (+ #10 1::Real)
-#8 := (+ f3 f3)
-#13 := (< #8 #12)
-#14 := (or false #13)
-#15 := (or #13 #14)
-#16 := (not #15)
-#72 := (iff #16 false)
-#40 := (+ 1::Real #10)
-#43 := (< #10 #40)
-#60 := (not #43)
-#70 := (iff #60 false)
-#1 := true
-#65 := (not true)
-#68 := (iff #65 false)
-#69 := [rewrite]: #68
-#66 := (iff #60 #65)
-#63 := (iff #43 true)
-#64 := [rewrite]: #63
-#67 := [monotonicity #64]: #66
-#71 := [trans #67 #69]: #70
-#61 := (iff #16 #60)
-#58 := (iff #15 #43)
-#53 := (or #43 #43)
-#56 := (iff #53 #43)
-#57 := [rewrite]: #56
-#54 := (iff #15 #53)
-#51 := (iff #14 #43)
-#46 := (or false #43)
-#49 := (iff #46 #43)
-#50 := [rewrite]: #49
-#47 := (iff #14 #46)
-#44 := (iff #13 #43)
-#41 := (= #12 #40)
-#42 := [rewrite]: #41
-#38 := (= #8 #10)
-#39 := [rewrite]: #38
-#45 := [monotonicity #39 #42]: #44
-#48 := [monotonicity #45]: #47
-#52 := [trans #48 #50]: #51
-#55 := [monotonicity #45 #52]: #54
-#59 := [trans #55 #57]: #58
-#62 := [monotonicity #59]: #61
-#73 := [trans #62 #71]: #72
-#37 := [asserted]: #16
-[mp #37 #73]: false
-unsat
-6e7ef563e385e00340c905e5fb44172a278ff733 2215 0
-#2 := false
-decl f12 :: Int
-#52 := f12
-decl f5 :: Int
-#13 := f5
-#64 := (= f5 f12)
-#9 := 0::Int
-#97 := -1::Int
-#235 := (* -1::Int f12)
-#733 := (+ f5 #235)
-#735 := (>= #733 0::Int)
-decl f10 :: Int
-#40 := f10
-#201 := (* -1::Int f10)
-#394 := (>= f10 0::Int)
-#401 := (if #394 f10 #201)
-#412 := (* -1::Int #401)
-#746 := (+ f10 #412)
-#748 := (>= #746 0::Int)
-#916 := (not #748)
-decl f11 :: Int
-#46 := f11
-#218 := (* -1::Int f11)
-#365 := (>= f11 0::Int)
-#372 := (if #365 f11 #218)
-#383 := (* -1::Int #372)
-#743 := (+ f11 #383)
-#745 := (>= #743 0::Int)
-#717 := (= f11 #372)
-#899 := (not #735)
-#900 := [hypothesis]: #899
-#1902 := (or #365 #735)
-decl f4 :: Int
-#8 := f4
-#98 := (* -1::Int f4)
-#568 := (>= f4 0::Int)
-#575 := (if #568 f4 #98)
-#586 := (* -1::Int #575)
-#985 := (+ f4 #586)
-#986 := (<= #985 0::Int)
-#1269 := (not #986)
-#888 := (<= #746 0::Int)
-#709 := (= f10 #401)
-#366 := (not #365)
-#1202 := [hypothesis]: #366
-#1880 := (or #394 #735 #365)
-#655 := (= f4 #575)
-decl f3 :: Int
-#7 := f3
-#116 := (* -1::Int f3)
-#539 := (>= f3 0::Int)
-#546 := (if #539 f3 #116)
-#557 := (* -1::Int #546)
-#761 := (+ f3 #557)
-#762 := (<= #761 0::Int)
-#669 := (= f3 #546)
-#1863 := (or #539 #365 #735)
-#395 := (not #394)
-decl f6 :: Int
-#16 := f6
-#510 := (>= f6 0::Int)
-#511 := (not #510)
-decl f9 :: Int
-#34 := f9
-#184 := (* -1::Int f9)
-#423 := (>= f9 0::Int)
-#430 := (if #423 f9 #184)
-#441 := (* -1::Int #430)
-#749 := (+ f9 #441)
-#751 := (>= #749 0::Int)
-#701 := (= f9 #430)
-#1430 := (>= #985 0::Int)
-#1498 := (not #1430)
-#587 := (+ f5 #586)
-#588 := (+ f3 #587)
-#649 := (<= #588 0::Int)
-#589 := (= #588 0::Int)
-decl f13 :: Int
-#58 := f13
-#65 := (= f4 f13)
-#66 := (and #64 #65)
-#336 := (>= f12 0::Int)
-#343 := (if #336 f12 #235)
-#354 := (* -1::Int #343)
-#355 := (+ f13 #354)
-#356 := (+ f11 #355)
-#357 := (= #356 0::Int)
-#362 := (not #357)
-#384 := (+ f12 #383)
-#385 := (+ f10 #384)
-#386 := (= #385 0::Int)
-#391 := (not #386)
-#413 := (+ f11 #412)
-#414 := (+ f9 #413)
-#415 := (= #414 0::Int)
-#420 := (not #415)
-#442 := (+ f10 #441)
-decl f8 :: Int
-#28 := f8
-#443 := (+ f8 #442)
-#444 := (= #443 0::Int)
-#449 := (not #444)
-#167 := (* -1::Int f8)
-#452 := (>= f8 0::Int)
-#459 := (if #452 f8 #167)
-#470 := (* -1::Int #459)
-#471 := (+ f9 #470)
-decl f7 :: Int
-#22 := f7
-#472 := (+ f7 #471)
-#473 := (= #472 0::Int)
-#478 := (not #473)
-#150 := (* -1::Int f7)
-#481 := (>= f7 0::Int)
-#488 := (if #481 f7 #150)
-#499 := (* -1::Int #488)
-#500 := (+ f8 #499)
-#501 := (+ f6 #500)
-#502 := (= #501 0::Int)
-#507 := (not #502)
-#133 := (* -1::Int f6)
-#517 := (if #510 f6 #133)
-#528 := (* -1::Int #517)
-#529 := (+ f7 #528)
-#530 := (+ f3 #529)
-#531 := (= #530 0::Int)
-#536 := (not #531)
-#558 := (+ f6 #557)
-#559 := (+ f4 #558)
-#560 := (= #559 0::Int)
-#565 := (not #560)
-#594 := (not #589)
-#624 := (or #594 #565 #536 #507 #478 #449 #420 #391 #362 #66)
-#629 := (not #624)
-#60 := (- f12)
-#59 := (< f12 0::Int)
-#61 := (if #59 #60 f12)
-#62 := (- #61 f11)
-#63 := (= f13 #62)
-#67 := (implies #63 #66)
-#54 := (- f11)
-#53 := (< f11 0::Int)
-#55 := (if #53 #54 f11)
-#56 := (- #55 f10)
-#57 := (= f12 #56)
-#68 := (implies #57 #67)
-#48 := (- f10)
-#47 := (< f10 0::Int)
-#49 := (if #47 #48 f10)
-#50 := (- #49 f9)
-#51 := (= f11 #50)
-#69 := (implies #51 #68)
-#42 := (- f9)
-#41 := (< f9 0::Int)
-#43 := (if #41 #42 f9)
-#44 := (- #43 f8)
-#45 := (= f10 #44)
-#70 := (implies #45 #69)
-#36 := (- f8)
-#35 := (< f8 0::Int)
-#37 := (if #35 #36 f8)
-#38 := (- #37 f7)
-#39 := (= f9 #38)
-#71 := (implies #39 #70)
-#30 := (- f7)
-#29 := (< f7 0::Int)
-#31 := (if #29 #30 f7)
-#32 := (- #31 f6)
-#33 := (= f8 #32)
-#72 := (implies #33 #71)
-#24 := (- f6)
-#23 := (< f6 0::Int)
-#25 := (if #23 #24 f6)
-#26 := (- #25 f3)
-#27 := (= f7 #26)
-#73 := (implies #27 #72)
-#18 := (- f3)
-#17 := (< f3 0::Int)
-#19 := (if #17 #18 f3)
-#20 := (- #19 f4)
-#21 := (= f6 #20)
-#74 := (implies #21 #73)
-#11 := (- f4)
-#10 := (< f4 0::Int)
-#12 := (if #10 #11 f4)
-#14 := (- #12 f5)
-#15 := (= f3 #14)
-#75 := (implies #15 #74)
-#76 := (not #75)
-#632 := (iff #76 #629)
-#238 := (if #59 #235 f12)
-#244 := (+ #218 #238)
-#249 := (= f13 #244)
-#255 := (not #249)
-#256 := (or #255 #66)
-#221 := (if #53 #218 f11)
-#227 := (+ #201 #221)
-#232 := (= f12 #227)
-#264 := (not #232)
-#265 := (or #264 #256)
-#204 := (if #47 #201 f10)
-#210 := (+ #184 #204)
-#215 := (= f11 #210)
-#273 := (not #215)
-#274 := (or #273 #265)
-#187 := (if #41 #184 f9)
-#193 := (+ #167 #187)
-#198 := (= f10 #193)
-#282 := (not #198)
-#283 := (or #282 #274)
-#170 := (if #35 #167 f8)
-#176 := (+ #150 #170)
-#181 := (= f9 #176)
-#291 := (not #181)
-#292 := (or #291 #283)
-#153 := (if #29 #150 f7)
-#159 := (+ #133 #153)
-#164 := (= f8 #159)
-#300 := (not #164)
-#301 := (or #300 #292)
-#136 := (if #23 #133 f6)
-#142 := (+ #116 #136)
-#147 := (= f7 #142)
-#309 := (not #147)
-#310 := (or #309 #301)
-#119 := (if #17 #116 f3)
-#125 := (+ #98 #119)
-#130 := (= f6 #125)
-#318 := (not #130)
-#319 := (or #318 #310)
-#101 := (if #10 #98 f4)
-#107 := (* -1::Int f5)
-#108 := (+ #107 #101)
-#113 := (= f3 #108)
-#327 := (not #113)
-#328 := (or #327 #319)
-#333 := (not #328)
-#630 := (iff #333 #629)
-#627 := (iff #328 #624)
-#597 := (or #362 #66)
-#600 := (or #391 #597)
-#603 := (or #420 #600)
-#606 := (or #449 #603)
-#609 := (or #478 #606)
-#612 := (or #507 #609)
-#615 := (or #536 #612)
-#618 := (or #565 #615)
-#621 := (or #594 #618)
-#625 := (iff #621 #624)
-#626 := [rewrite]: #625
-#622 := (iff #328 #621)
-#619 := (iff #319 #618)
-#616 := (iff #310 #615)
-#613 := (iff #301 #612)
-#610 := (iff #292 #609)
-#607 := (iff #283 #606)
-#604 := (iff #274 #603)
-#601 := (iff #265 #600)
-#598 := (iff #256 #597)
-#363 := (iff #255 #362)
-#360 := (iff #249 #357)
-#348 := (+ #218 #343)
-#351 := (= f13 #348)
-#358 := (iff #351 #357)
-#359 := [rewrite]: #358
-#352 := (iff #249 #351)
-#349 := (= #244 #348)
-#346 := (= #238 #343)
-#337 := (not #336)
-#340 := (if #337 #235 f12)
-#344 := (= #340 #343)
-#345 := [rewrite]: #344
-#341 := (= #238 #340)
-#338 := (iff #59 #337)
-#339 := [rewrite]: #338
-#342 := [monotonicity #339]: #341
-#347 := [trans #342 #345]: #346
-#350 := [monotonicity #347]: #349
-#353 := [monotonicity #350]: #352
-#361 := [trans #353 #359]: #360
-#364 := [monotonicity #361]: #363
-#599 := [monotonicity #364]: #598
-#392 := (iff #264 #391)
-#389 := (iff #232 #386)
-#377 := (+ #201 #372)
-#380 := (= f12 #377)
-#387 := (iff #380 #386)
-#388 := [rewrite]: #387
-#381 := (iff #232 #380)
-#378 := (= #227 #377)
-#375 := (= #221 #372)
-#369 := (if #366 #218 f11)
-#373 := (= #369 #372)
-#374 := [rewrite]: #373
-#370 := (= #221 #369)
-#367 := (iff #53 #366)
-#368 := [rewrite]: #367
-#371 := [monotonicity #368]: #370
-#376 := [trans #371 #374]: #375
-#379 := [monotonicity #376]: #378
-#382 := [monotonicity #379]: #381
-#390 := [trans #382 #388]: #389
-#393 := [monotonicity #390]: #392
-#602 := [monotonicity #393 #599]: #601
-#421 := (iff #273 #420)
-#418 := (iff #215 #415)
-#406 := (+ #184 #401)
-#409 := (= f11 #406)
-#416 := (iff #409 #415)
-#417 := [rewrite]: #416
-#410 := (iff #215 #409)
-#407 := (= #210 #406)
-#404 := (= #204 #401)
-#398 := (if #395 #201 f10)
-#402 := (= #398 #401)
-#403 := [rewrite]: #402
-#399 := (= #204 #398)
-#396 := (iff #47 #395)
-#397 := [rewrite]: #396
-#400 := [monotonicity #397]: #399
-#405 := [trans #400 #403]: #404
-#408 := [monotonicity #405]: #407
-#411 := [monotonicity #408]: #410
-#419 := [trans #411 #417]: #418
-#422 := [monotonicity #419]: #421
-#605 := [monotonicity #422 #602]: #604
-#450 := (iff #282 #449)
-#447 := (iff #198 #444)
-#435 := (+ #167 #430)
-#438 := (= f10 #435)
-#445 := (iff #438 #444)
-#446 := [rewrite]: #445
-#439 := (iff #198 #438)
-#436 := (= #193 #435)
-#433 := (= #187 #430)
-#424 := (not #423)
-#427 := (if #424 #184 f9)
-#431 := (= #427 #430)
-#432 := [rewrite]: #431
-#428 := (= #187 #427)
-#425 := (iff #41 #424)
-#426 := [rewrite]: #425
-#429 := [monotonicity #426]: #428
-#434 := [trans #429 #432]: #433
-#437 := [monotonicity #434]: #436
-#440 := [monotonicity #437]: #439
-#448 := [trans #440 #446]: #447
-#451 := [monotonicity #448]: #450
-#608 := [monotonicity #451 #605]: #607
-#479 := (iff #291 #478)
-#476 := (iff #181 #473)
-#464 := (+ #150 #459)
-#467 := (= f9 #464)
-#474 := (iff #467 #473)
-#475 := [rewrite]: #474
-#468 := (iff #181 #467)
-#465 := (= #176 #464)
-#462 := (= #170 #459)
-#453 := (not #452)
-#456 := (if #453 #167 f8)
-#460 := (= #456 #459)
-#461 := [rewrite]: #460
-#457 := (= #170 #456)
-#454 := (iff #35 #453)
-#455 := [rewrite]: #454
-#458 := [monotonicity #455]: #457
-#463 := [trans #458 #461]: #462
-#466 := [monotonicity #463]: #465
-#469 := [monotonicity #466]: #468
-#477 := [trans #469 #475]: #476
-#480 := [monotonicity #477]: #479
-#611 := [monotonicity #480 #608]: #610
-#508 := (iff #300 #507)
-#505 := (iff #164 #502)
-#493 := (+ #133 #488)
-#496 := (= f8 #493)
-#503 := (iff #496 #502)
-#504 := [rewrite]: #503
-#497 := (iff #164 #496)
-#494 := (= #159 #493)
-#491 := (= #153 #488)
-#482 := (not #481)
-#485 := (if #482 #150 f7)
-#489 := (= #485 #488)
-#490 := [rewrite]: #489
-#486 := (= #153 #485)
-#483 := (iff #29 #482)
-#484 := [rewrite]: #483
-#487 := [monotonicity #484]: #486
-#492 := [trans #487 #490]: #491
-#495 := [monotonicity #492]: #494
-#498 := [monotonicity #495]: #497
-#506 := [trans #498 #504]: #505
-#509 := [monotonicity #506]: #508
-#614 := [monotonicity #509 #611]: #613
-#537 := (iff #309 #536)
-#534 := (iff #147 #531)
-#522 := (+ #116 #517)
-#525 := (= f7 #522)
-#532 := (iff #525 #531)
-#533 := [rewrite]: #532
-#526 := (iff #147 #525)
-#523 := (= #142 #522)
-#520 := (= #136 #517)
-#514 := (if #511 #133 f6)
-#518 := (= #514 #517)
-#519 := [rewrite]: #518
-#515 := (= #136 #514)
-#512 := (iff #23 #511)
-#513 := [rewrite]: #512
-#516 := [monotonicity #513]: #515
-#521 := [trans #516 #519]: #520
-#524 := [monotonicity #521]: #523
-#527 := [monotonicity #524]: #526
-#535 := [trans #527 #533]: #534
-#538 := [monotonicity #535]: #537
-#617 := [monotonicity #538 #614]: #616
-#566 := (iff #318 #565)
-#563 := (iff #130 #560)
-#551 := (+ #98 #546)
-#554 := (= f6 #551)
-#561 := (iff #554 #560)
-#562 := [rewrite]: #561
-#555 := (iff #130 #554)
-#552 := (= #125 #551)
-#549 := (= #119 #546)
-#540 := (not #539)
-#543 := (if #540 #116 f3)
-#547 := (= #543 #546)
-#548 := [rewrite]: #547
-#544 := (= #119 #543)
-#541 := (iff #17 #540)
-#542 := [rewrite]: #541
-#545 := [monotonicity #542]: #544
-#550 := [trans #545 #548]: #549
-#553 := [monotonicity #550]: #552
-#556 := [monotonicity #553]: #555
-#564 := [trans #556 #562]: #563
-#567 := [monotonicity #564]: #566
-#620 := [monotonicity #567 #617]: #619
-#595 := (iff #327 #594)
-#592 := (iff #113 #589)
-#580 := (+ #107 #575)
-#583 := (= f3 #580)
-#590 := (iff #583 #589)
-#591 := [rewrite]: #590
-#584 := (iff #113 #583)
-#581 := (= #108 #580)
-#578 := (= #101 #575)
-#569 := (not #568)
-#572 := (if #569 #98 f4)
-#576 := (= #572 #575)
-#577 := [rewrite]: #576
-#573 := (= #101 #572)
-#570 := (iff #10 #569)
-#571 := [rewrite]: #570
-#574 := [monotonicity #571]: #573
-#579 := [trans #574 #577]: #578
-#582 := [monotonicity #579]: #581
-#585 := [monotonicity #582]: #584
-#593 := [trans #585 #591]: #592
-#596 := [monotonicity #593]: #595
-#623 := [monotonicity #596 #620]: #622
-#628 := [trans #623 #626]: #627
-#631 := [monotonicity #628]: #630
-#334 := (iff #76 #333)
-#331 := (iff #75 #328)
-#324 := (implies #113 #319)
-#329 := (iff #324 #328)
-#330 := [rewrite]: #329
-#325 := (iff #75 #324)
-#322 := (iff #74 #319)
-#315 := (implies #130 #310)
-#320 := (iff #315 #319)
-#321 := [rewrite]: #320
-#316 := (iff #74 #315)
-#313 := (iff #73 #310)
-#306 := (implies #147 #301)
-#311 := (iff #306 #310)
-#312 := [rewrite]: #311
-#307 := (iff #73 #306)
-#304 := (iff #72 #301)
-#297 := (implies #164 #292)
-#302 := (iff #297 #301)
-#303 := [rewrite]: #302
-#298 := (iff #72 #297)
-#295 := (iff #71 #292)
-#288 := (implies #181 #283)
-#293 := (iff #288 #292)
-#294 := [rewrite]: #293
-#289 := (iff #71 #288)
-#286 := (iff #70 #283)
-#279 := (implies #198 #274)
-#284 := (iff #279 #283)
-#285 := [rewrite]: #284
-#280 := (iff #70 #279)
-#277 := (iff #69 #274)
-#270 := (implies #215 #265)
-#275 := (iff #270 #274)
-#276 := [rewrite]: #275
-#271 := (iff #69 #270)
-#268 := (iff #68 #265)
-#261 := (implies #232 #256)
-#266 := (iff #261 #265)
-#267 := [rewrite]: #266
-#262 := (iff #68 #261)
-#259 := (iff #67 #256)
-#252 := (implies #249 #66)
-#257 := (iff #252 #256)
-#258 := [rewrite]: #257
-#253 := (iff #67 #252)
-#250 := (iff #63 #249)
-#247 := (= #62 #244)
-#241 := (- #238 f11)
-#245 := (= #241 #244)
-#246 := [rewrite]: #245
-#242 := (= #62 #241)
-#239 := (= #61 #238)
-#236 := (= #60 #235)
-#237 := [rewrite]: #236
-#240 := [monotonicity #237]: #239
-#243 := [monotonicity #240]: #242
-#248 := [trans #243 #246]: #247
-#251 := [monotonicity #248]: #250
-#254 := [monotonicity #251]: #253
-#260 := [trans #254 #258]: #259
-#233 := (iff #57 #232)
-#230 := (= #56 #227)
-#224 := (- #221 f10)
-#228 := (= #224 #227)
-#229 := [rewrite]: #228
-#225 := (= #56 #224)
-#222 := (= #55 #221)
-#219 := (= #54 #218)
-#220 := [rewrite]: #219
-#223 := [monotonicity #220]: #222
-#226 := [monotonicity #223]: #225
-#231 := [trans #226 #229]: #230
-#234 := [monotonicity #231]: #233
-#263 := [monotonicity #234 #260]: #262
-#269 := [trans #263 #267]: #268
-#216 := (iff #51 #215)
-#213 := (= #50 #210)
-#207 := (- #204 f9)
-#211 := (= #207 #210)
-#212 := [rewrite]: #211
-#208 := (= #50 #207)
-#205 := (= #49 #204)
-#202 := (= #48 #201)
-#203 := [rewrite]: #202
-#206 := [monotonicity #203]: #205
-#209 := [monotonicity #206]: #208
-#214 := [trans #209 #212]: #213
-#217 := [monotonicity #214]: #216
-#272 := [monotonicity #217 #269]: #271
-#278 := [trans #272 #276]: #277
-#199 := (iff #45 #198)
-#196 := (= #44 #193)
-#190 := (- #187 f8)
-#194 := (= #190 #193)
-#195 := [rewrite]: #194
-#191 := (= #44 #190)
-#188 := (= #43 #187)
-#185 := (= #42 #184)
-#186 := [rewrite]: #185
-#189 := [monotonicity #186]: #188
-#192 := [monotonicity #189]: #191
-#197 := [trans #192 #195]: #196
-#200 := [monotonicity #197]: #199
-#281 := [monotonicity #200 #278]: #280
-#287 := [trans #281 #285]: #286
-#182 := (iff #39 #181)
-#179 := (= #38 #176)
-#173 := (- #170 f7)
-#177 := (= #173 #176)
-#178 := [rewrite]: #177
-#174 := (= #38 #173)
-#171 := (= #37 #170)
-#168 := (= #36 #167)
-#169 := [rewrite]: #168
-#172 := [monotonicity #169]: #171
-#175 := [monotonicity #172]: #174
-#180 := [trans #175 #178]: #179
-#183 := [monotonicity #180]: #182
-#290 := [monotonicity #183 #287]: #289
-#296 := [trans #290 #294]: #295
-#165 := (iff #33 #164)
-#162 := (= #32 #159)
-#156 := (- #153 f6)
-#160 := (= #156 #159)
-#161 := [rewrite]: #160
-#157 := (= #32 #156)
-#154 := (= #31 #153)
-#151 := (= #30 #150)
-#152 := [rewrite]: #151
-#155 := [monotonicity #152]: #154
-#158 := [monotonicity #155]: #157
-#163 := [trans #158 #161]: #162
-#166 := [monotonicity #163]: #165
-#299 := [monotonicity #166 #296]: #298
-#305 := [trans #299 #303]: #304
-#148 := (iff #27 #147)
-#145 := (= #26 #142)
-#139 := (- #136 f3)
-#143 := (= #139 #142)
-#144 := [rewrite]: #143
-#140 := (= #26 #139)
-#137 := (= #25 #136)
-#134 := (= #24 #133)
-#135 := [rewrite]: #134
-#138 := [monotonicity #135]: #137
-#141 := [monotonicity #138]: #140
-#146 := [trans #141 #144]: #145
-#149 := [monotonicity #146]: #148
-#308 := [monotonicity #149 #305]: #307
-#314 := [trans #308 #312]: #313
-#131 := (iff #21 #130)
-#128 := (= #20 #125)
-#122 := (- #119 f4)
-#126 := (= #122 #125)
-#127 := [rewrite]: #126
-#123 := (= #20 #122)
-#120 := (= #19 #119)
-#117 := (= #18 #116)
-#118 := [rewrite]: #117
-#121 := [monotonicity #118]: #120
-#124 := [monotonicity #121]: #123
-#129 := [trans #124 #127]: #128
-#132 := [monotonicity #129]: #131
-#317 := [monotonicity #132 #314]: #316
-#323 := [trans #317 #321]: #322
-#114 := (iff #15 #113)
-#111 := (= #14 #108)
-#104 := (- #101 f5)
-#109 := (= #104 #108)
-#110 := [rewrite]: #109
-#105 := (= #14 #104)
-#102 := (= #12 #101)
-#99 := (= #11 #98)
-#100 := [rewrite]: #99
-#103 := [monotonicity #100]: #102
-#106 := [monotonicity #103]: #105
-#112 := [trans #106 #110]: #111
-#115 := [monotonicity #112]: #114
-#326 := [monotonicity #115 #323]: #325
-#332 := [trans #326 #330]: #331
-#335 := [monotonicity #332]: #334
-#633 := [trans #335 #631]: #632
-#96 := [asserted]: #76
-#634 := [mp #96 #633]: #629
-#635 := [not-or-elim #634]: #589
-#1489 := (or #594 #649)
-#1490 := [th-lemma arith triangle-eq]: #1489
-#1491 := [unit-resolution #1490 #635]: #649
-#675 := (<= #559 0::Int)
-#636 := [not-or-elim #634]: #560
-#1486 := (or #565 #675)
-#1487 := [th-lemma arith triangle-eq]: #1486
-#1488 := [unit-resolution #1487 #636]: #675
-#1251 := (+ #167 #470)
-#741 := (>= #1251 0::Int)
-#1066 := [hypothesis]: #424
-#1804 := (or #539 #423)
-#818 := [hypothesis]: #540
-#1760 := (or #394 #539 #423)
-#747 := (+ #201 #412)
-#1708 := (>= #747 0::Int)
-#710 := (= #201 #401)
-#1122 := [hypothesis]: #395
-#713 := (or #394 #710)
-#714 := [def-axiom]: #713
-#1709 := [unit-resolution #714 #1122]: #710
-#1230 := (not #710)
-#1710 := (or #1230 #1708)
-#1711 := [th-lemma arith triangle-eq]: #1710
-#1712 := [unit-resolution #1711 #1709]: #1708
-#683 := (<= #530 0::Int)
-#637 := [not-or-elim #634]: #531
-#895 := (or #536 #683)
-#896 := [th-lemma arith triangle-eq]: #895
-#897 := [unit-resolution #896 #637]: #683
-#760 := (+ f6 #528)
-#756 := (>= #760 0::Int)
-#677 := (= f6 #517)
-#1197 := (or #510 #423)
-#989 := [hypothesis]: #511
-#1188 := (or #481 #510 #423)
-#752 := (+ f8 #470)
-#988 := (<= #752 0::Int)
-#1014 := (not #988)
-#1062 := (+ #150 #499)
-#1161 := (<= #1062 0::Int)
-#686 := (= #150 #488)
-#891 := [hypothesis]: #482
-#689 := (or #481 #686)
-#690 := [def-axiom]: #689
-#1169 := [unit-resolution #690 #891]: #686
-#1094 := (not #686)
-#1170 := (or #1094 #1161)
-#1171 := [th-lemma arith triangle-eq]: #1170
-#1172 := [unit-resolution #1171 #1169]: #1161
-#927 := (+ #184 #441)
-#744 := (>= #927 0::Int)
-#702 := (= #184 #430)
-#705 := (or #423 #702)
-#706 := [def-axiom]: #705
-#1071 := [unit-resolution #706 #1066]: #702
-#954 := (not #702)
-#1173 := (or #954 #744)
-#1174 := [th-lemma arith triangle-eq]: #1173
-#1175 := [unit-resolution #1174 #1071]: #744
-#1166 := (or #394 #423 #481)
-#700 := (>= #472 0::Int)
-#639 := [not-or-elim #634]: #473
-#1011 := (or #478 #700)
-#1012 := [th-lemma arith triangle-eq]: #1011
-#1013 := [unit-resolution #1012 #639]: #700
-#928 := (<= #927 0::Int)
-#955 := (or #954 #928)
-#1027 := (not #928)
-#1028 := [hypothesis]: #1027
-#1029 := [hypothesis]: #702
-#956 := [th-lemma arith triangle-eq]: #955
-#1030 := [unit-resolution #956 #1029 #1028]: false
-#1031 := [lemma #1030]: #955
-#1072 := [unit-resolution #1031 #1071]: #928
-#708 := (>= #443 0::Int)
-#640 := [not-or-elim #634]: #444
-#905 := (or #449 #708)
-#906 := [th-lemma arith triangle-eq]: #905
-#907 := [unit-resolution #906 #640]: #708
-#1015 := (not #700)
-#1048 := (not #708)
-#1130 := (or #481 #394 #1048 #1014 #1015 #423 #1027)
-#1131 := [th-lemma arith assign-bounds 1 1 1 1 2 1]: #1130
-#1162 := [unit-resolution #1131 #1122 #1066 #907 #891 #1072 #1013]: #1014
-#693 := (= f8 #459)
-#1123 := (or #452 #423 #394 #1048 #1027)
-#1124 := [th-lemma arith assign-bounds 1 1 1 1]: #1123
-#1163 := [unit-resolution #1124 #1122 #907 #1072 #1066]: #452
-#695 := (or #453 #693)
-#696 := [def-axiom]: #695
-#1164 := [unit-resolution #696 #1163]: #693
-#1007 := (not #693)
-#1008 := (or #1007 #988)
-#1067 := [hypothesis]: #1014
-#1068 := [hypothesis]: #693
-#1009 := [th-lemma arith triangle-eq]: #1008
-#1069 := [unit-resolution #1009 #1068 #1067]: false
-#1070 := [lemma #1069]: #1008
-#1165 := [unit-resolution #1070 #1164 #1162]: false
-#1167 := [lemma #1165]: #1166
-#1176 := [unit-resolution #1167 #891 #1066]: #394
-#707 := (<= #443 0::Int)
-#834 := (or #449 #707)
-#835 := [th-lemma arith triangle-eq]: #834
-#836 := [unit-resolution #835 #640]: #707
-#692 := (>= #501 0::Int)
-#638 := [not-or-elim #634]: #502
-#867 := (or #507 #692)
-#868 := [th-lemma arith triangle-eq]: #867
-#869 := [unit-resolution #868 #638]: #692
-#1002 := (not #692)
-#1179 := (not #1161)
-#1178 := (not #707)
-#1177 := (not #744)
-#1180 := (or #1014 #1015 #1177 #1178 #481 #395 #1179 #1002 #510)
-#1181 := [th-lemma arith assign-bounds 1 1 1 3 1 2 2 2]: #1180
-#1182 := [unit-resolution #1181 #891 #869 #1013 #836 #1176 #989 #1175 #1172]: #1014
-#1183 := (or #452 #1179 #1002 #510 #481)
-#1184 := [th-lemma arith assign-bounds 1 1 1 1]: #1183
-#1185 := [unit-resolution #1184 #891 #869 #989 #1172]: #452
-#1186 := [unit-resolution #696 #1185]: #693
-#1187 := [unit-resolution #1070 #1186 #1182]: false
-#1189 := [lemma #1187]: #1188
-#1168 := [unit-resolution #1189 #989 #1066]: #481
-#1159 := (or #539 #423 #510)
-#755 := (+ f7 #499)
-#812 := (<= #755 0::Int)
-#685 := (= f7 #488)
-#982 := (+ #133 #528)
-#983 := (<= #982 0::Int)
-#678 := (= #133 #517)
-#681 := (or #510 #678)
-#682 := [def-axiom]: #681
-#990 := [unit-resolution #682 #989]: #678
-#991 := (not #678)
-#992 := (or #991 #983)
-#993 := [th-lemma arith triangle-eq]: #992
-#994 := [unit-resolution #993 #990]: #983
-#684 := (>= #530 0::Int)
-#814 := (or #536 #684)
-#815 := [th-lemma arith triangle-eq]: #814
-#816 := [unit-resolution #815 #637]: #684
-#871 := (not #684)
-#995 := (not #983)
-#996 := (or #481 #995 #510 #539 #871)
-#997 := [th-lemma arith assign-bounds 1 1 1 1]: #996
-#1152 := [unit-resolution #997 #818 #816 #994 #989]: #481
-#687 := (or #482 #685)
-#688 := [def-axiom]: #687
-#1153 := [unit-resolution #688 #1152]: #685
-#876 := (not #685)
-#877 := (or #876 #812)
-#878 := [th-lemma arith triangle-eq]: #877
-#1154 := [unit-resolution #878 #1153]: #812
-#1001 := (not #812)
-#1016 := (or #423 #510 #1014 #1015 #1001 #1002)
-#1017 := [th-lemma arith assign-bounds 1 1 1 1 1]: #1016
-#1155 := [unit-resolution #1017 #1154 #1013 #1066 #989 #869]: #1014
-#1003 := (or #452 #1001 #1002 #510 #995 #539 #871)
-#1004 := [th-lemma arith assign-bounds 1 1 2 1 1 1]: #1003
-#1156 := [unit-resolution #1004 #1154 #816 #869 #818 #994 #989]: #452
-#1157 := [unit-resolution #696 #1156]: #693
-#1158 := [unit-resolution #1070 #1157 #1155]: false
-#1160 := [lemma #1158]: #1159
-#1190 := [unit-resolution #1160 #989 #1066]: #539
-#984 := (>= #982 0::Int)
-#1021 := (or #991 #984)
-#1022 := [th-lemma arith triangle-eq]: #1021
-#1023 := [unit-resolution #1022 #990]: #984
-#1191 := [unit-resolution #688 #1168]: #685
-#1192 := [unit-resolution #878 #1191]: #812
-#1079 := (not #984)
-#1051 := (not #683)
-#1108 := (or #452 #1001 #1002 #482 #540 #1051 #1079)
-#1109 := [th-lemma arith assign-bounds -1/2 1/2 1 1/2 -1/2 1/2]: #1108
-#1193 := [unit-resolution #1109 #1192 #1023 #869 #1190 #1168 #897]: #452
-#1194 := [unit-resolution #1017 #1192 #1013 #1066 #989 #869]: #1014
-#1195 := [unit-resolution #1070 #1194]: #1007
-#1196 := [unit-resolution #696 #1195 #1193]: false
-#1198 := [lemma #1196]: #1197
-#1203 := [unit-resolution #1198 #1066]: #510
-#679 := (or #511 #677)
-#680 := [def-axiom]: #679
-#1209 := [unit-resolution #680 #1203]: #677
-#830 := (not #677)
-#958 := (or #830 #756)
-#959 := [th-lemma arith triangle-eq]: #958
-#1713 := [unit-resolution #959 #1209]: #756
-#750 := (<= #749 0::Int)
-#1268 := (not #750)
-#1550 := [unit-resolution #1031 #1028]: #954
-#1551 := [unit-resolution #706 #1550]: #423
-#1552 := (or #928 #1268 #424)
-#1553 := [th-lemma arith assign-bounds 1 -2]: #1552
-#1554 := [unit-resolution #1553 #1551 #1028]: #1268
-#703 := (or #424 #701)
-#704 := [def-axiom]: #703
-#1555 := [unit-resolution #704 #1551]: #701
-#909 := (not #701)
-#910 := (or #909 #750)
-#911 := [th-lemma arith triangle-eq]: #910
-#1556 := [unit-resolution #911 #1555 #1554]: false
-#1557 := [lemma #1556]: #928
-#758 := (+ #116 #557)
-#759 := (<= #758 0::Int)
-#670 := (= #116 #546)
-#673 := (or #539 #670)
-#674 := [def-axiom]: #673
-#819 := [unit-resolution #674 #818]: #670
-#804 := (not #670)
-#805 := (or #804 #759)
-#806 := [th-lemma arith triangle-eq]: #805
-#820 := [unit-resolution #806 #819]: #759
-#691 := (<= #501 0::Int)
-#785 := (or #507 #691)
-#786 := [th-lemma arith triangle-eq]: #785
-#787 := [unit-resolution #786 #638]: #691
-#757 := (>= #755 0::Int)
-#1705 := (or #481 #423)
-#1356 := (<= #1251 0::Int)
-#1439 := (not #1356)
-#754 := (>= #752 0::Int)
-#1434 := (or #988 #754)
-#1435 := [th-lemma arith farkas 1 1]: #1434
-#1436 := [unit-resolution #1435 #1067]: #754
-#1437 := [unit-resolution #1070 #1067]: #1007
-#1438 := [unit-resolution #696 #1437]: #453
-#797 := (not #754)
-#1440 := (or #797 #1439 #452)
-#1441 := [th-lemma arith assign-bounds 1 2]: #1440
-#1442 := [unit-resolution #1441 #1438 #1436]: #1439
-#694 := (= #167 #459)
-#697 := (or #452 #694)
-#698 := [def-axiom]: #697
-#1443 := [unit-resolution #698 #1438]: #694
-#1444 := (not #694)
-#1445 := (or #1444 #1356)
-#1446 := [th-lemma arith triangle-eq]: #1445
-#1447 := [unit-resolution #1446 #1443 #1442]: false
-#1448 := [lemma #1447]: #988
-#1362 := [hypothesis]: #453
-#1466 := [unit-resolution #698 #1362]: #694
-#1478 := (or #1444 #741)
-#1479 := [th-lemma arith triangle-eq]: #1478
-#1480 := [unit-resolution #1479 #1466]: #741
-#699 := (<= #472 0::Int)
-#789 := (or #478 #699)
-#790 := [th-lemma arith triangle-eq]: #789
-#791 := [unit-resolution #790 #639]: #699
-#1546 := (or #481 #452)
-#668 := (not #65)
-#734 := (<= #733 0::Int)
-#811 := (<= #760 0::Int)
-#1449 := (or #452 #1179 #510 #481)
-#1450 := [unit-resolution #1184 #869]: #1449
-#1451 := [unit-resolution #1450 #1172 #1362 #891]: #510
-#1452 := [unit-resolution #680 #1451]: #677
-#831 := (or #830 #811)
-#832 := [th-lemma arith triangle-eq]: #831
-#1453 := [unit-resolution #832 #1452]: #811
-#870 := (not #811)
-#1454 := (or #481 #511 #870 #539)
-#1035 := (or #481 #511 #870 #539 #871)
-#1036 := [th-lemma arith assign-bounds 1 1 1 1]: #1035
-#1455 := [unit-resolution #1036 #816]: #1454
-#1456 := [unit-resolution #1455 #1453 #891 #1451]: #539
-#671 := (or #540 #669)
-#672 := [def-axiom]: #671
-#1457 := [unit-resolution #672 #1456]: #669
-#776 := (not #669)
-#777 := (or #776 #762)
-#778 := [th-lemma arith triangle-eq]: #777
-#1458 := [unit-resolution #778 #1457]: #762
-#844 := (not #762)
-#1459 := (or #568 #844 #870 #481)
-#676 := (>= #559 0::Int)
-#771 := (or #565 #676)
-#772 := [th-lemma arith triangle-eq]: #771
-#773 := [unit-resolution #772 #636]: #676
-#823 := (not #676)
-#1387 := (or #568 #823 #844 #870 #871 #481)
-#1388 := [th-lemma arith assign-bounds 1 1 1 1 1]: #1387
-#1460 := [unit-resolution #1388 #816 #773]: #1459
-#1461 := [unit-resolution #1460 #1458 #891 #1453]: #568
-#653 := (or #569 #655)
-#654 := [def-axiom]: #653
-#1462 := [unit-resolution #654 #1461]: #655
-#1263 := (not #655)
-#1463 := (or #1263 #1430)
-#1464 := [th-lemma arith triangle-eq]: #1463
-#1465 := [unit-resolution #1464 #1462]: #1430
-#1200 := (<= #743 0::Int)
-#1467 := [unit-resolution #1446 #1466]: #1356
-#1468 := (or #423 #1439 #481 #1015 #452)
-#1469 := [th-lemma arith assign-bounds 1 1 1 1]: #1468
-#1470 := [unit-resolution #1469 #891 #1013 #1362 #1467]: #423
-#1471 := [unit-resolution #704 #1470]: #701
-#1472 := [unit-resolution #911 #1471]: #750
-#1376 := (or #452 #365 #1268)
-#854 := (not #709)
-#1267 := (not #888)
-#1252 := [hypothesis]: #750
-#716 := (>= #414 0::Int)
-#641 := [not-or-elim #634]: #415
-#1215 := (or #420 #716)
-#1216 := [th-lemma arith triangle-eq]: #1215
-#1217 := [unit-resolution #1216 #641]: #716
-#1240 := (not #716)
-#1363 := (or #1267 #365 #1240 #1268 #1048 #452)
-#1364 := [th-lemma arith assign-bounds 1 1 1 1 1]: #1363
-#1365 := [unit-resolution #1364 #1362 #1217 #1202 #1252 #907]: #1267
-#1219 := (or #854 #888)
-#1358 := [hypothesis]: #1267
-#1359 := [hypothesis]: #709
-#1220 := [th-lemma arith triangle-eq]: #1219
-#1360 := [unit-resolution #1220 #1359 #1358]: false
-#1361 := [lemma #1360]: #1219
-#1366 := [unit-resolution #1361 #1365]: #854
-#711 := (or #395 #709)
-#712 := [def-axiom]: #711
-#1367 := [unit-resolution #712 #1366]: #395
-#1368 := [unit-resolution #714 #1367]: #710
-#753 := (<= #747 0::Int)
-#1227 := (not #753)
-#1369 := (or #748 #365 #1240 #1268 #1048 #452)
-#1370 := [th-lemma arith assign-bounds 1 1 1 1 1]: #1369
-#1371 := [unit-resolution #1370 #1362 #1217 #1202 #907 #1252]: #748
-#1372 := (or #916 #1227 #394)
-#1373 := [th-lemma arith assign-bounds 1 2]: #1372
-#1374 := [unit-resolution #1373 #1367 #1371]: #1227
-#1231 := (or #1230 #753)
-#1228 := [hypothesis]: #1227
-#1229 := [hypothesis]: #710
-#1232 := [th-lemma arith triangle-eq]: #1231
-#1233 := [unit-resolution #1232 #1229 #1228]: false
-#1234 := [lemma #1233]: #1231
-#1375 := [unit-resolution #1234 #1374 #1368]: false
-#1377 := [lemma #1375]: #1376
-#1473 := [unit-resolution #1377 #1472 #1362]: #365
-#719 := (or #366 #717)
-#720 := [def-axiom]: #719
-#1474 := [unit-resolution #720 #1473]: #717
-#860 := (not #717)
-#1475 := (or #860 #1200)
-#1476 := [th-lemma arith triangle-eq]: #1475
-#1477 := [unit-resolution #1476 #1474]: #1200
-#1481 := (or #394 #481 #1268)
-#1273 := (or #394 #481 #1014 #1015 #1268 #1048)
-#1274 := [th-lemma arith assign-bounds 1 1 1 1 1]: #1273
-#1482 := [unit-resolution #1274 #907 #1448 #1013]: #1481
-#1483 := [unit-resolution #1482 #1472 #891]: #394
-#1484 := [unit-resolution #712 #1483]: #709
-#1485 := [unit-resolution #1361 #1484]: #888
-#724 := (>= #385 0::Int)
-#642 := [not-or-elim #634]: #386
-#1492 := (or #391 #724)
-#1493 := [th-lemma arith triangle-eq]: #1492
-#1494 := [unit-resolution #1493 #642]: #724
-#933 := (>= #761 0::Int)
-#1495 := (or #776 #933)
-#1496 := [th-lemma arith triangle-eq]: #1495
-#1497 := [unit-resolution #1496 #1457]: #933
-#1504 := (not #675)
-#1503 := (not #933)
-#1050 := (not #699)
-#1502 := (not #741)
-#1501 := (not #724)
-#1500 := (not #1200)
-#1499 := (not #649)
-#1505 := (or #734 #1498 #1499 #1179 #1002 #1500 #1501 #1502 #1050 #1503 #1504 #1267 #1240)
-#1506 := [th-lemma arith assign-bounds 1 -1 -1 1 -1 1 1 -1 1 -1 -1 1]: #1505
-#1507 := [unit-resolution #1506 #1497 #869 #791 #1217 #1494 #1491 #1488 #1172 #1485 #1480 #1477 #1465]: #734
-#1064 := (>= #1062 0::Int)
-#1095 := (or #1094 #1064)
-#1090 := (not #1064)
-#1065 := [hypothesis]: #1090
-#1093 := [hypothesis]: #686
-#1096 := [th-lemma arith triangle-eq]: #1095
-#1097 := [unit-resolution #1096 #1093 #1065]: false
-#1098 := [lemma #1097]: #1095
-#1208 := [unit-resolution #1098 #1169]: #1064
-#1264 := (or #1263 #986)
-#1265 := [th-lemma arith triangle-eq]: #1264
-#1508 := [unit-resolution #1265 #1462]: #986
-#855 := (or #854 #748)
-#856 := [th-lemma arith triangle-eq]: #855
-#1509 := [unit-resolution #856 #1484]: #748
-#650 := (>= #588 0::Int)
-#901 := (or #594 #650)
-#902 := [th-lemma arith triangle-eq]: #901
-#903 := [unit-resolution #902 #635]: #650
-#723 := (<= #385 0::Int)
-#780 := (or #391 #723)
-#781 := [th-lemma arith triangle-eq]: #780
-#782 := [unit-resolution #781 #642]: #723
-#715 := (<= #414 0::Int)
-#880 := (or #420 #715)
-#881 := [th-lemma arith triangle-eq]: #880
-#882 := [unit-resolution #881 #641]: #715
-#861 := (or #860 #745)
-#795 := (not #745)
-#1204 := [hypothesis]: #795
-#1205 := [hypothesis]: #717
-#862 := [th-lemma arith triangle-eq]: #861
-#1206 := [unit-resolution #862 #1205 #1204]: false
-#1207 := [lemma #1206]: #861
-#1510 := [unit-resolution #1207 #1474]: #745
-#947 := (not #715)
-#822 := (not #723)
-#1049 := (not #691)
-#948 := (not #650)
-#1511 := (or #735 #1269 #948 #1090 #1049 #795 #822 #1439 #1015 #844 #823 #916 #947)
-#1512 := [th-lemma arith assign-bounds 1 -1 -1 1 -1 1 1 -1 1 -1 -1 1]: #1511
-#1513 := [unit-resolution #1512 #1510 #787 #1013 #882 #782 #903 #773 #1458 #1509 #1508 #1208 #1467]: #735
-#949 := (not #734)
-#1514 := (or #64 #949 #899)
-#1515 := [th-lemma arith triangle-eq]: #1514
-#1516 := [unit-resolution #1515 #1513 #1507]: #64
-#667 := (not #64)
-#647 := (or #667 #668)
-#644 := (not #66)
-#660 := (iff #644 #647)
-#648 := (not #647)
-#663 := (not #648)
-#662 := (iff #663 #647)
-#659 := [rewrite]: #662
-#664 := (iff #644 #663)
-#665 := (iff #66 #648)
-#666 := [rewrite]: #665
-#661 := [monotonicity #666]: #664
-#657 := [trans #661 #659]: #660
-#645 := [not-or-elim #634]: #644
-#658 := [mp #645 #657]: #647
-#1517 := [unit-resolution #658 #1516]: #668
-#736 := (* -1::Int f13)
-#737 := (+ f4 #736)
-#739 := (>= #737 0::Int)
-#1431 := (+ #235 #354)
-#1433 := (>= #1431 0::Int)
-#726 := (= #235 #343)
-#1518 := (or #337 #795 #822 #452 #1439 #481 #1015 #916 #947)
-#1519 := [th-lemma arith assign-bounds 1 1 1 1 1 1 1 1]: #1518
-#1520 := [unit-resolution #1519 #891 #1013 #882 #782 #1362 #1510 #1509 #1467]: #337
-#729 := (or #336 #726)
-#730 := [def-axiom]: #729
-#1521 := [unit-resolution #730 #1520]: #726
-#1522 := (not #726)
-#1523 := (or #1522 #1433)
-#1524 := [th-lemma arith triangle-eq]: #1523
-#1525 := [unit-resolution #1524 #1521]: #1433
-#731 := (<= #356 0::Int)
-#643 := [not-or-elim #634]: #357
-#767 := (or #362 #731)
-#768 := [th-lemma arith triangle-eq]: #767
-#769 := [unit-resolution #768 #643]: #731
-#824 := (not #731)
-#1526 := (not #1433)
-#1527 := (or #739 #1526 #1500 #1501 #1502 #1050 #1267 #1240 #824 #844 #823 #870 #871 #1268 #1048)
-#1528 := [th-lemma arith assign-bounds -1 1 -1 -1 1 2 -2 1 1 -1 1 -1 1 -1]: #1527
-#1529 := [unit-resolution #1528 #1458 #791 #907 #1217 #1494 #769 #773 #816 #1453 #1472 #1485 #1480 #1477 #1525]: #739
-#738 := (<= #737 0::Int)
-#1432 := (<= #1431 0::Int)
-#1530 := (or #1522 #1432)
-#1531 := [th-lemma arith triangle-eq]: #1530
-#1532 := [unit-resolution #1531 #1521]: #1432
-#1533 := [unit-resolution #959 #1452]: #756
-#1407 := (or #909 #751)
-#1408 := [th-lemma arith triangle-eq]: #1407
-#1534 := [unit-resolution #1408 #1471]: #751
-#732 := (>= #356 0::Int)
-#1535 := (or #362 #732)
-#1536 := [th-lemma arith triangle-eq]: #1535
-#1537 := [unit-resolution #1536 #643]: #732
-#838 := (not #751)
-#917 := (not #756)
-#1539 := (not #732)
-#1538 := (not #1432)
-#1540 := (or #738 #1538 #795 #822 #1439 #1015 #916 #947 #1539 #1503 #1504 #917 #1051 #838 #1178)
-#1541 := [th-lemma arith assign-bounds -1 1 -1 -1 1 2 -2 1 1 -1 1 -1 1 -1]: #1540
-#1542 := [unit-resolution #1541 #1510 #1013 #836 #882 #782 #1537 #1488 #897 #1534 #1509 #1533 #1497 #1467 #1532]: #738
-#765 := (not #739)
-#825 := (not #738)
-#1543 := (or #65 #825 #765)
-#1544 := [th-lemma arith triangle-eq]: #1543
-#1545 := [unit-resolution #1544 #1542 #1529 #1517]: false
-#1547 := [lemma #1545]: #1546
-#1572 := [unit-resolution #1547 #1362]: #481
-#1594 := (or #1027 #1502 #482 #1050 #1048 #394)
-#1595 := [th-lemma arith assign-bounds -1 -1 1 -1 1]: #1594
-#1596 := [unit-resolution #1595 #1480 #907 #1572 #1557 #791]: #394
-#1597 := [unit-resolution #712 #1596]: #709
-#1598 := [unit-resolution #1361 #1597]: #888
-#1573 := [unit-resolution #688 #1572]: #685
-#1574 := [unit-resolution #878 #1573]: #812
-#1680 := (or #1161 #482 #1001)
-#1681 := [th-lemma arith assign-bounds 2 -1]: #1680
-#1682 := [unit-resolution #1681 #1574 #1572]: #1161
-#1549 := [hypothesis]: #870
-#1558 := [hypothesis]: #677
-#1559 := [unit-resolution #832 #1558 #1549]: false
-#1560 := [lemma #1559]: #831
-#1561 := [unit-resolution #1560 #1549]: #830
-#1562 := [unit-resolution #680 #1561]: #511
-#1304 := (or #811 #510 #995)
-#1305 := [th-lemma arith assign-bounds 2 1]: #1304
-#1563 := [unit-resolution #1305 #1562 #1549]: #995
-#1564 := [unit-resolution #682 #1562]: #678
-#1565 := [unit-resolution #993 #1564 #1563]: false
-#1566 := [lemma #1565]: #811
-#1575 := (or #452 #1001 #870 #539)
-#1040 := (or #452 #1001 #1002 #870 #539 #871)
-#1041 := [th-lemma arith assign-bounds 1 1 1 1 1]: #1040
-#1576 := [unit-resolution #1041 #869 #816]: #1575
-#1577 := [unit-resolution #1576 #1574 #1566 #1362]: #539
-#1578 := [unit-resolution #672 #1577]: #669
-#1579 := [unit-resolution #1496 #1578]: #933
-#1636 := (or #423 #452)
-#886 := (+ #98 #586)
-#1570 := (>= #886 0::Int)
-#656 := (= #98 #575)
-#1580 := (or #452 #1001 #482 #540 #1079)
-#1581 := [unit-resolution #1109 #869 #897]: #1580
-#1582 := [unit-resolution #1581 #1577 #1572 #1362 #1574]: #1079
-#1548 := [hypothesis]: #1079
-#1567 := [hypothesis]: #678
-#1568 := [unit-resolution #1022 #1567 #1548]: false
-#1569 := [lemma #1568]: #1021
-#1583 := [unit-resolution #1569 #1582]: #991
-#1584 := [unit-resolution #682 #1583]: #510
-#1585 := [unit-resolution #680 #1584]: #677
-#1586 := [unit-resolution #959 #1585]: #756
-#1587 := (or #569 #1504 #917 #1051 #1503 #1439 #1015 #423 #452)
-#1588 := [th-lemma arith assign-bounds 1 1 1 1 1 1 1 1]: #1587
-#1589 := [unit-resolution #1588 #1066 #897 #1362 #1013 #1488 #1586 #1579 #1467]: #569
-#651 := (or #568 #656)
-#652 := [def-axiom]: #651
-#1590 := [unit-resolution #652 #1589]: #656
-#922 := (not #656)
-#1591 := (or #922 #1570)
-#1592 := [th-lemma arith triangle-eq]: #1591
-#1593 := [unit-resolution #1592 #1590]: #1570
-#1599 := [unit-resolution #778 #1578]: #762
-#1602 := (or #365 #1267 #1027 #423 #452)
-#1600 := (or #365 #1267 #1240 #1027 #1048 #423 #452)
-#1601 := [th-lemma arith assign-bounds 1 1 1 1 2 1]: #1600
-#1603 := [unit-resolution #1601 #907 #1217]: #1602
-#1604 := [unit-resolution #1603 #1066 #1557 #1362 #1598]: #365
-#1605 := [unit-resolution #720 #1604]: #717
-#1606 := [unit-resolution #1476 #1605]: #1200
-#1607 := (not #1570)
-#1608 := (or #734 #1499 #1500 #1501 #1502 #1050 #823 #1267 #1240 #1001 #1002 #844 #1607 #870 #871)
-#1609 := [th-lemma arith assign-bounds -1 -1 1 1 -1 1 -1 1 -1 1 -1 1 -2 2]: #1608
-#1610 := [unit-resolution #1609 #1606 #816 #869 #791 #1217 #1494 #1491 #1599 #1566 #1574 #1598 #773 #1480 #1593]: #734
-#1611 := [unit-resolution #856 #1597]: #748
-#887 := (<= #886 0::Int)
-#923 := (or #922 #887)
-#915 := (not #887)
-#920 := [hypothesis]: #915
-#921 := [hypothesis]: #656
-#924 := [th-lemma arith triangle-eq]: #923
-#925 := [unit-resolution #924 #921 #920]: false
-#926 := [lemma #925]: #923
-#1612 := [unit-resolution #926 #1590]: #887
-#940 := (or #876 #757)
-#941 := [th-lemma arith triangle-eq]: #940
-#1613 := [unit-resolution #941 #1573]: #757
-#1614 := [unit-resolution #1207 #1605]: #745
-#794 := (not #757)
-#1615 := (or #735 #948 #795 #822 #1439 #1015 #1504 #916 #947 #794 #1049 #1503 #915 #917 #1051)
-#1616 := [th-lemma arith assign-bounds -1 -1 1 1 -1 1 -1 1 -1 1 -1 1 -2 2]: #1615
-#1617 := [unit-resolution #1616 #1614 #897 #787 #1013 #882 #782 #903 #1488 #1613 #1612 #1611 #1586 #1579 #1467]: #735
-#1618 := [unit-resolution #1515 #1617 #1610]: #64
-#1619 := [unit-resolution #658 #1618]: #668
-#740 := (+ f12 #354)
-#1571 := (<= #740 0::Int)
-#725 := (= f12 #343)
-#1620 := (or #336 #1500 #1501 #1267 #1240 #423)
-#1621 := [th-lemma arith assign-bounds 1 1 1 1 1]: #1620
-#1622 := [unit-resolution #1621 #1066 #1494 #1217 #1598 #1606]: #336
-#727 := (or #337 #725)
-#728 := [def-axiom]: #727
-#1623 := [unit-resolution #728 #1622]: #725
-#1394 := (not #725)
-#1624 := (or #1394 #1571)
-#1625 := [th-lemma arith triangle-eq]: #1624
-#1626 := [unit-resolution #1625 #1623]: #1571
-#1627 := (not #1571)
-#1628 := (or #738 #1627 #1500 #1501 #1539 #1504 #917 #1051 #1503 #1439 #1015 #1177 #1178)
-#1629 := [th-lemma arith assign-bounds 1 1 -1 -1 1 -1 1 -1 1 -1 -1 1]: #1628
-#1630 := [unit-resolution #1629 #1175 #1013 #836 #1494 #1537 #1488 #1586 #1579 #897 #1467 #1606 #1626]: #738
-#742 := (>= #740 0::Int)
-#1395 := (or #1394 #742)
-#1396 := [th-lemma arith triangle-eq]: #1395
-#1631 := [unit-resolution #1396 #1623]: #742
-#796 := (not #742)
-#1632 := (or #739 #796 #795 #822 #824 #823 #870 #871 #844 #1502 #1050 #1027 #1048)
-#1633 := [th-lemma arith assign-bounds 1 1 -1 -1 1 -1 1 -1 1 -1 -1 1]: #1632
-#1634 := [unit-resolution #1633 #1614 #791 #907 #782 #769 #773 #816 #1631 #1599 #1566 #1557 #1480]: #739
-#1635 := [unit-resolution #1544 #1634 #1630 #1619]: false
-#1637 := [lemma #1635]: #1636
-#1683 := [unit-resolution #1637 #1362]: #423
-#1684 := [unit-resolution #704 #1683]: #701
-#1685 := [unit-resolution #911 #1684]: #750
-#1686 := [unit-resolution #1377 #1685 #1362]: #365
-#1687 := [unit-resolution #720 #1686]: #717
-#1688 := [unit-resolution #1476 #1687]: #1200
-#1689 := [unit-resolution #1207 #1687]: #745
-#1663 := (or #735 #844 #916 #795 #1439 #794 #917 #1503)
-#1652 := [hypothesis]: #1356
-#784 := [hypothesis]: #745
-#913 := [hypothesis]: #748
-#889 := [hypothesis]: #762
-#1653 := [hypothesis]: #933
-#898 := [hypothesis]: #756
-#788 := [hypothesis]: #757
-#1654 := [unit-resolution #1616 #900 #897 #787 #1013 #882 #782 #903 #1488 #788 #784 #913 #898 #1653 #1652]: #915
-#1655 := [unit-resolution #926 #1654]: #922
-#1656 := [unit-resolution #652 #1655]: #568
-#1657 := [unit-resolution #654 #1656]: #655
-#1658 := [unit-resolution #1265 #1657]: #986
-#1659 := (or #1064 #794 #1504 #569 #917 #1051 #1503)
-#1660 := [th-lemma arith assign-bounds -1 2 -2 -2 2 -2]: #1659
-#1661 := [unit-resolution #1660 #1656 #897 #788 #898 #1488 #1653]: #1064
-#1662 := [unit-resolution #1512 #1661 #1658 #787 #1013 #882 #782 #903 #773 #889 #913 #784 #900 #1652]: false
-#1664 := [lemma #1662]: #1663
-#1690 := [unit-resolution #1664 #1599 #1611 #1689 #1467 #1613 #1586 #1579]: #735
-#1650 := (or #739 #795 #844 #1502 #1500 #1268 #1267)
-#1642 := [hypothesis]: #741
-#766 := [hypothesis]: #765
-#1643 := [unit-resolution #1633 #766 #791 #907 #782 #769 #773 #816 #784 #889 #1566 #1557 #1642]: #796
-#1385 := [hypothesis]: #888
-#1644 := [hypothesis]: #1200
-#1645 := [unit-resolution #1528 #766 #791 #907 #1217 #1494 #769 #1644 #889 #1566 #1252 #1385 #1642 #816 #773]: #1526
-#1638 := [hypothesis]: #1526
-#1639 := [hypothesis]: #726
-#1640 := [unit-resolution #1524 #1639 #1638]: false
-#1641 := [lemma #1640]: #1523
-#1646 := [unit-resolution #1641 #1645]: #1522
-#1647 := [unit-resolution #730 #1646]: #336
-#1648 := [unit-resolution #728 #1647]: #725
-#1649 := [unit-resolution #1396 #1648 #1643]: false
-#1651 := [lemma #1649]: #1650
-#1691 := [unit-resolution #1651 #1689 #1599 #1480 #1688 #1685 #1598]: #739
-#1692 := [unit-resolution #1408 #1684]: #751
-#1675 := (or #738 #795 #916 #917 #1503 #1439 #838)
-#813 := [hypothesis]: #751
-#1668 := [hypothesis]: #825
-#1669 := [unit-resolution #1541 #1668 #1013 #836 #882 #782 #1537 #1652 #784 #813 #913 #898 #1653 #897 #1488]: #1538
-#1665 := [hypothesis]: #1538
-#1666 := [unit-resolution #1531 #1639 #1665]: false
-#1667 := [lemma #1666]: #1530
-#1670 := [unit-resolution #1667 #1669]: #1522
-#1671 := [unit-resolution #730 #1670]: #336
-#1672 := [unit-resolution #728 #1671]: #725
-#1673 := [unit-resolution #1625 #1672]: #1571
-#1674 := [th-lemma arith farkas 1/2 -1/2 1 -1 -1/2 1/2 -1/2 1/2 -1/2 1/2 -1/2 1/2 1/2 1/2 -1/2 1 #784 #782 #913 #882 #1488 #898 #897 #1653 #1652 #1013 #1673 #1537 #1668 #813 #836 #1671]: false
-#1676 := [lemma #1674]: #1675
-#1693 := [unit-resolution #1676 #1689 #1611 #1586 #1579 #1467 #1692]: #738
-#1694 := [unit-resolution #1544 #1693 #1691]: #65
-#1695 := [unit-resolution #658 #1694]: #667
-#1696 := [unit-resolution #1515 #1695 #1690]: #949
-#1697 := [unit-resolution #1506 #1696 #869 #791 #1217 #1494 #1688 #1579 #1682 #1598 #1480 #1488 #1491]: #1498
-#1698 := [unit-resolution #1609 #1696 #816 #869 #791 #1217 #1494 #1688 #1599 #1566 #1574 #1598 #773 #1480 #1491]: #1607
-#1677 := [hypothesis]: #1607
-#1678 := [unit-resolution #1592 #921 #1677]: false
-#1679 := [lemma #1678]: #1591
-#1699 := [unit-resolution #1679 #1698]: #922
-#1700 := [unit-resolution #652 #1699]: #568
-#1701 := [unit-resolution #654 #1700]: #655
-#1702 := [unit-resolution #1464 #1701 #1697]: false
-#1703 := [lemma #1702]: #452
-#1704 := [th-lemma arith farkas 1 1 1 1 1 #1703 #891 #1013 #1066 #1448]: false
-#1706 := [lemma #1704]: #1705
-#1714 := [unit-resolution #1706 #1066]: #481
-#1715 := [unit-resolution #688 #1714]: #685
-#1716 := [unit-resolution #941 #1715]: #757
-#1717 := [unit-resolution #696 #1703]: #693
-#1044 := (or #1007 #754)
-#1045 := [th-lemma arith triangle-eq]: #1044
-#1718 := [unit-resolution #1045 #1717]: #754
-#1076 := (or #838 #423 #1027)
-#1077 := [th-lemma arith assign-bounds 2 1]: #1076
-#1719 := [unit-resolution #1077 #1066 #1557]: #838
-#1720 := (or #750 #751)
-#1721 := [th-lemma arith farkas 1 1]: #1720
-#1722 := [unit-resolution #1721 #1719]: #750
-#1723 := [unit-resolution #1234 #1709]: #753
-#1726 := (or #1177 #1268 #394 #365 #1227)
-#1724 := (or #1177 #1268 #394 #365 #1227 #1240)
-#1725 := [th-lemma arith assign-bounds 1 2 2 2 2]: #1724
-#1727 := [unit-resolution #1725 #1217]: #1726
-#1728 := [unit-resolution #1727 #1723 #1722 #1122 #1175]: #365
-#1729 := [unit-resolution #720 #1728]: #717
-#1730 := [unit-resolution #1207 #1729]: #745
-#821 := (not #759)
-#1731 := (or #568 #823 #797 #1050 #794 #1049 #821 #394 #1048 #1027 #917 #1051)
-#1732 := [th-lemma arith assign-bounds 1 1 1 2 2 1 1 1 1 1 1]: #1731
-#1733 := [unit-resolution #1732 #1122 #897 #787 #791 #907 #773 #1716 #1718 #820 #1713 #1557]: #568
-#1734 := [unit-resolution #654 #1733]: #655
-#1735 := [unit-resolution #1265 #1734]: #986
-#1736 := [th-lemma arith assign-bounds 1 -1 -1 -1 1 1 -1 1 -3 3 1 -2 2 -2 2 -1 #1735 #903 #773 #1730 #782 #882 #1718 #791 #1716 #787 #820 #907 #1557 #1713 #897 #1712]: #735
-#1707 := (>= #758 0::Int)
-#1737 := (or #804 #1707)
-#1738 := [th-lemma arith triangle-eq]: #1737
-#1739 := [unit-resolution #1738 #819]: #1707
-#1740 := [unit-resolution #878 #1715]: #812
-#1741 := [unit-resolution #1476 #1729]: #1200
-#1742 := [unit-resolution #1464 #1734]: #1430
-#1743 := [th-lemma arith assign-bounds 1 -1 -1 -1 1 1 -1 1 -3 3 1 -2 2 -2 2 -1 #1742 #1491 #1488 #1741 #1494 #1217 #1448 #1013 #1740 #869 #1739 #836 #1175 #1566 #816 #1723]: #734
-#1744 := [unit-resolution #1515 #1743 #1736]: #64
-#1745 := [unit-resolution #1373 #1723 #1122]: #916
-#1746 := (or #888 #748)
-#1747 := [th-lemma arith farkas 1 1]: #1746
-#1748 := [unit-resolution #1747 #1745]: #888
-#1749 := [unit-resolution #1621 #1741 #1494 #1217 #1066 #1748]: #336
-#1750 := [unit-resolution #728 #1749]: #725
-#1751 := [unit-resolution #1396 #1750]: #742
-#1060 := (or #539 #795 #796 #739)
-#770 := [hypothesis]: #742
-#1025 := (or #510 #795 #796 #739 #539)
-#998 := [unit-resolution #997 #989 #816 #818 #994]: #481
-#999 := [unit-resolution #688 #998]: #685
-#1000 := [unit-resolution #878 #999]: #812
-#1005 := [unit-resolution #1004 #989 #816 #869 #818 #994 #1000]: #452
-#1006 := [unit-resolution #696 #1005]: #693
-#1010 := [unit-resolution #1009 #1006]: #988
-#1018 := [unit-resolution #1017 #989 #1013 #869 #1000 #1010]: #423
-#1019 := [unit-resolution #704 #1018]: #701
-#1020 := [unit-resolution #911 #1019]: #750
-#1024 := [th-lemma arith farkas -1 -1 1 1 -1 -1 1 1 -1 1 1 -1 1 #907 #784 #782 #820 #773 #770 #769 #766 #1023 #897 #1010 #1013 #1020]: false
-#1026 := [lemma #1024]: #1025
-#987 := [unit-resolution #1026 #818 #770 #766 #784]: #510
-#1032 := [unit-resolution #680 #987]: #677
-#1033 := [unit-resolution #959 #1032]: #756
-#1034 := [unit-resolution #832 #1032]: #811
-#1037 := [unit-resolution #1036 #987 #816 #818 #1034]: #481
-#1038 := [unit-resolution #688 #1037]: #685
-#1039 := [unit-resolution #878 #1038]: #812
-#1042 := [unit-resolution #1041 #818 #869 #816 #1034 #1039]: #452
-#1043 := [unit-resolution #696 #1042]: #693
-#1046 := [unit-resolution #1045 #1043]: #754
-#1047 := [unit-resolution #941 #1038]: #757
-#1052 := (or #1027 #1048 #796 #824 #739 #794 #1049 #797 #1050 #795 #822 #821 #823 #917 #1051)
-#1053 := [th-lemma arith assign-bounds -1 -1 1 1 -2 2 -1 1 -1 1 1 -1 -1 1]: #1052
-#1054 := [unit-resolution #1053 #1047 #787 #791 #907 #782 #769 #766 #770 #784 #897 #1046 #820 #1033 #773]: #1027
-#1055 := [unit-resolution #1031 #1054]: #954
-#1056 := [unit-resolution #706 #1055]: #423
-#1057 := [unit-resolution #704 #1056]: #701
-#1058 := [unit-resolution #911 #1057]: #750
-#1059 := [th-lemma arith farkas 1/2 -1/2 1 -1 -1/2 1/2 1/2 -1/2 -1/2 1/2 1/2 -1/2 -1/2 1/2 -1/2 1 #1046 #791 #1047 #787 #1058 #907 #784 #782 #820 #773 #770 #769 #766 #1033 #897 #1056]: false
-#1061 := [lemma #1059]: #1060
-#1752 := [unit-resolution #1061 #1751 #818 #1730]: #739
-#1753 := [unit-resolution #1625 #1750]: #1571
-#1754 := (not #1707)
-#1755 := (or #738 #1504 #1627 #1500 #1501 #1539 #1178 #1177 #1001 #1002 #1014 #1015 #870 #871 #1754)
-#1756 := [th-lemma arith assign-bounds 1 1 1 -1 -1 1 -1 2 -2 1 -1 1 -1 -1]: #1755
-#1757 := [unit-resolution #1756 #1741 #869 #1013 #836 #1494 #1537 #1488 #1566 #1740 #1448 #1175 #816 #1753 #1739]: #738
-#1758 := [unit-resolution #1544 #1757 #1752]: #65
-#1759 := [unit-resolution #658 #1758 #1744]: false
-#1761 := [lemma #1759]: #1760
-#1774 := [unit-resolution #1761 #818 #1066]: #394
-#1775 := [unit-resolution #712 #1774]: #709
-#1776 := [unit-resolution #1361 #1775]: #888
-#1779 := (or #1177 #1268 #1267 #365 #395)
-#1777 := (or #1177 #1268 #1267 #1240 #365 #395)
-#1778 := [th-lemma arith assign-bounds 1 2 2 2 2]: #1777
-#1780 := [unit-resolution #1778 #1217]: #1779
-#1781 := [unit-resolution #1780 #1776 #1722 #1774 #1175]: #365
-#1782 := [unit-resolution #720 #1781]: #717
-#1783 := [unit-resolution #1476 #1782]: #1200
-#1784 := [unit-resolution #1207 #1782]: #745
-#1785 := [unit-resolution #1621 #1783 #1494 #1217 #1066 #1776]: #336
-#1786 := [unit-resolution #728 #1785]: #725
-#1787 := [unit-resolution #1396 #1786]: #742
-#1788 := [unit-resolution #1061 #1787 #818 #1784]: #739
-#1789 := [unit-resolution #1625 #1786]: #1571
-#1790 := [unit-resolution #1756 #1789 #869 #1013 #836 #1494 #1537 #1783 #1566 #1740 #1448 #1175 #816 #1488 #1739]: #738
-#1791 := [unit-resolution #1544 #1790 #1788]: #65
-#1792 := [unit-resolution #658 #1791]: #667
-#1793 := [unit-resolution #856 #1775]: #748
-#1772 := (or #735 #795 #1001 #1754 #916)
-#1284 := [hypothesis]: #812
-#1762 := [hypothesis]: #1707
-#1764 := (or #915 #1001 #1754 #735 #795 #916)
-#904 := [hypothesis]: #887
-#1763 := [th-lemma arith farkas 1 1 -1 1 -1 -1 -1 1 -1 1 1 -1 1 #1488 #1448 #1013 #1284 #869 #1762 #903 #900 #784 #782 #882 #913 #904]: false
-#1765 := [lemma #1763]: #1764
-#1766 := [unit-resolution #1765 #900 #1762 #1284 #784 #913]: #915
-#1767 := [unit-resolution #926 #1766]: #922
-#1768 := [unit-resolution #652 #1767]: #568
-#1769 := [unit-resolution #654 #1768]: #655
-#1770 := [unit-resolution #1265 #1769]: #986
-#1771 := [th-lemma arith farkas -1 1 1 -1 1 1 1 -1 1 -1 -1 -1 -2 1 #903 #900 #1488 #784 #782 #882 #1448 #1013 #1284 #869 #1762 #913 #1768 #1770]: false
-#1773 := [lemma #1771]: #1772
-#1794 := [unit-resolution #1773 #1784 #1740 #1739 #1793]: #735
-#1795 := [unit-resolution #1515 #1794 #1792]: #949
-#1796 := (or #1607 #823 #797 #1050 #794 #1049 #821 #1499 #734 #1500 #1501 #1240 #1267)
-#1797 := [th-lemma arith assign-bounds 1 1 -1 1 -1 -1 -1 1 -1 1 1 -1]: #1796
-#1798 := [unit-resolution #1797 #1795 #787 #791 #1217 #1494 #773 #1716 #1718 #820 #1776 #1783 #1491]: #1607
-#1799 := [unit-resolution #1679 #1798]: #922
-#1800 := [unit-resolution #652 #1799]: #568
-#1801 := [unit-resolution #654 #1800]: #655
-#1802 := [unit-resolution #1464 #1801]: #1430
-#1803 := [th-lemma arith farkas -1/2 -1/2 1/2 -3/2 3/2 1/2 -1 1 -1 1 1/2 -1/2 1/2 -1/2 1/2 1/2 -1/2 1 #1488 #1448 #1013 #1740 #869 #1739 #836 #1175 #1566 #816 #1802 #1491 #1795 #1783 #1494 #1217 #1776 #1774]: false
-#1805 := [lemma #1803]: #1804
-#1806 := [unit-resolution #1805 #1066]: #539
-#1807 := (or #741 #797 #794 #1049 #917 #1051 #540)
-#1808 := [th-lemma arith assign-bounds -1 -2 2 -2 2 -2]: #1807
-#1809 := [unit-resolution #1808 #1716 #787 #897 #1718 #1713 #1806]: #741
-#1810 := (or #394 #794 #1049 #1048 #1027 #917 #1051 #423 #540)
-#1811 := [th-lemma arith assign-bounds 1 1 1 1 1 1 1 1]: #1810
-#1812 := [unit-resolution #1811 #1066 #787 #897 #907 #1806 #1716 #1713 #1557]: #394
-#1813 := [unit-resolution #712 #1812]: #709
-#1814 := [unit-resolution #1361 #1813]: #888
-#1815 := (or #1161 #1049 #453 #482 #511)
-#1816 := [th-lemma arith assign-bounds -1 1 1 1]: #1815
-#1817 := [unit-resolution #1816 #1714 #787 #1703 #1203]: #1161
-#1818 := [unit-resolution #1780 #1814 #1722 #1812 #1175]: #365
-#1819 := [unit-resolution #720 #1818]: #717
-#1820 := [unit-resolution #1476 #1819]: #1200
-#1821 := [unit-resolution #672 #1806]: #669
-#1822 := [unit-resolution #1496 #1821]: #933
-#1823 := [unit-resolution #1207 #1819]: #745
-#1826 := (or #1356 #453)
-#1824 := (or #1356 #453 #1014)
-#1825 := [th-lemma arith assign-bounds 2 -1]: #1824
-#1827 := [unit-resolution #1825 #1448]: #1826
-#1828 := [unit-resolution #1827 #1703]: #1356
-#1829 := [unit-resolution #778 #1821]: #762
-#1830 := [unit-resolution #856 #1813]: #748
-#1831 := [unit-resolution #1664 #1830 #1829 #1822 #1828 #1716 #1713 #1823]: #735
-#1832 := [unit-resolution #1651 #1820 #1829 #1809 #1823 #1722 #1814]: #739
-#1833 := [unit-resolution #1621 #1820 #1494 #1217 #1066 #1814]: #336
-#1834 := [unit-resolution #728 #1833]: #725
-#1835 := [unit-resolution #1625 #1834]: #1571
-#1836 := [unit-resolution #1629 #1835 #1013 #836 #1494 #1537 #1822 #1713 #1820 #1175 #1828 #897 #1488]: #738
-#1837 := [unit-resolution #1544 #1836 #1832]: #65
-#1838 := [unit-resolution #658 #1837]: #667
-#1839 := [unit-resolution #1515 #1838 #1831]: #949
-#1840 := [unit-resolution #1506 #1839 #869 #791 #1217 #1494 #1822 #1820 #1817 #1814 #1809 #1488 #1491]: #1498
-#1073 := (or #759 #540 #844)
-#1074 := [th-lemma arith assign-bounds 2 -1]: #1073
-#1841 := [unit-resolution #1074 #1829 #1806]: #759
-#1842 := [unit-resolution #1797 #1839 #787 #791 #1217 #1494 #773 #1716 #1718 #1841 #1814 #1820 #1491]: #1607
-#1843 := [unit-resolution #1679 #1842]: #922
-#1844 := [unit-resolution #652 #1843]: #568
-#1845 := [unit-resolution #654 #1844]: #655
-#1846 := [unit-resolution #1464 #1845 #1840]: false
-#1847 := [lemma #1846]: #423
-#1849 := [unit-resolution #704 #1847]: #701
-#1850 := [unit-resolution #1408 #1849]: #751
-#1354 := (or #539 #511 #365 #838)
-#1335 := [hypothesis]: #510
-#1336 := [unit-resolution #680 #1335]: #677
-#1337 := [unit-resolution #832 #1336]: #811
-#1338 := [unit-resolution #1036 #818 #816 #1335 #1337]: #481
-#1339 := [unit-resolution #688 #1338]: #685
-#1340 := [unit-resolution #878 #1339]: #812
-#1341 := [unit-resolution #1041 #1340 #869 #818 #1337 #816]: #452
-#1342 := [unit-resolution #696 #1341]: #693
-#1343 := [unit-resolution #1045 #1342]: #754
-#1344 := (or #983 #511 #870)
-#1345 := [th-lemma arith assign-bounds 2 -1]: #1344
-#1346 := [unit-resolution #1345 #1337 #1335]: #983
-#1347 := [unit-resolution #941 #1339]: #757
-#1289 := (or #539 #794 #1227 #995 #838 #365 #1001 #870)
-#1282 := [hypothesis]: #983
-#1283 := [hypothesis]: #753
-#890 := [hypothesis]: #811
-#1285 := [unit-resolution #1041 #818 #869 #1284 #890 #816]: #452
-#1286 := [unit-resolution #696 #1285]: #693
-#1287 := [unit-resolution #1045 #1286]: #754
-#1288 := [th-lemma arith farkas 2 2 1 1 1 1 1 1 1 1 1 1 #1287 #791 #788 #1283 #1217 #787 #816 #818 #1282 #813 #836 #1202]: false
-#1290 := [lemma #1288]: #1289
-#1348 := [unit-resolution #1290 #1347 #818 #1346 #813 #1202 #1340 #1337]: #1227
-#1349 := [unit-resolution #1234 #1348]: #1230
-#1350 := [unit-resolution #714 #1349]: #394
-#1351 := [unit-resolution #712 #1350]: #709
-#1352 := [unit-resolution #1220 #1351]: #888
-#1353 := [th-lemma arith farkas 1 -1 -1 1 -1 -1 -1 1 1 #1352 #1347 #1217 #787 #1335 #1350 #1343 #791 #1202]: false
-#1355 := [lemma #1353]: #1354
-#1851 := [unit-resolution #1355 #818 #1850 #1202]: #511
-#1852 := [unit-resolution #911 #1849]: #750
-#1199 := (+ #218 #383)
-#1201 := (>= #1199 0::Int)
-#718 := (= #218 #372)
-#721 := (or #365 #718)
-#722 := [def-axiom]: #721
-#1226 := [unit-resolution #722 #1202]: #718
-#1235 := (not #718)
-#1236 := (or #1235 #1201)
-#1237 := [th-lemma arith triangle-eq]: #1236
-#1238 := [unit-resolution #1237 #1226]: #1201
-#1223 := (not #1201)
-#1278 := (or #481 #1268 #735 #1223 #510)
-#1214 := [hypothesis]: #1201
-#1253 := [unit-resolution #1184 #1172 #869 #989 #891]: #452
-#1254 := [unit-resolution #696 #1253]: #693
-#1255 := [unit-resolution #1070 #1254]: #988
-#1256 := [unit-resolution #997 #891 #816 #989 #994]: #539
-#1257 := [unit-resolution #672 #1256]: #669
-#1258 := [unit-resolution #778 #1257]: #762
-#1259 := (or #568 #540 #844 #823 #510)
-#1260 := [th-lemma arith assign-bounds 1 1 1 1]: #1259
-#1261 := [unit-resolution #1260 #1258 #773 #989 #1256]: #568
-#1262 := [unit-resolution #654 #1261]: #655
-#1266 := [unit-resolution #1265 #1262]: #986
-#1270 := (or #1267 #1240 #1268 #1048 #844 #1049 #823 #1090 #1014 #1015 #1223 #822 #1269 #948 #735)
-#1271 := [th-lemma arith assign-bounds -1 2 -2 1 1 -1 -1 1 -1 -1 1 1 -1 1]: #1270
-#1272 := [unit-resolution #1271 #1258 #787 #1013 #907 #1217 #782 #900 #773 #1266 #1255 #1252 #1208 #903 #1214]: #1267
-#1275 := [unit-resolution #1274 #891 #907 #1013 #1255 #1252]: #394
-#1276 := [unit-resolution #712 #1275]: #709
-#1277 := [unit-resolution #1220 #1276 #1272]: false
-#1279 := [lemma #1277]: #1278
-#1853 := [unit-resolution #1279 #1851 #900 #1238 #1852]: #481
-#1854 := [unit-resolution #688 #1853]: #685
-#1855 := [unit-resolution #878 #1854]: #812
-#1311 := (or #539 #510 #395 #838 #1001)
-#1306 := [unit-resolution #1305 #994 #989]: #811
-#1307 := [unit-resolution #1041 #818 #869 #1284 #1306 #816]: #452
-#1308 := [unit-resolution #696 #1307]: #693
-#1309 := [unit-resolution #1045 #1308]: #754
-#783 := [hypothesis]: #394
-#1310 := [th-lemma arith farkas 1 1 1 1 1 1 1 1 1 #989 #783 #791 #816 #818 #994 #813 #836 #1309]: false
-#1312 := [lemma #1310]: #1311
-#1856 := [unit-resolution #1312 #1855 #1850 #818 #1851]: #395
-#1857 := [unit-resolution #941 #1854]: #757
-#1858 := [unit-resolution #682 #1851]: #678
-#1859 := [unit-resolution #993 #1858]: #983
-#1860 := [unit-resolution #1290 #1859 #1566 #1850 #818 #1202 #1855 #1857]: #1227
-#1861 := [unit-resolution #1234 #1860]: #1230
-#1862 := [unit-resolution #714 #1861 #1856]: false
-#1864 := [lemma #1862]: #1863
-#1865 := [unit-resolution #1864 #1202 #900]: #539
-#1866 := [unit-resolution #672 #1865]: #669
-#1867 := [unit-resolution #778 #1866]: #762
-#1868 := [unit-resolution #1482 #1122 #1852]: #481
-#1869 := [unit-resolution #688 #1868]: #685
-#1870 := [unit-resolution #941 #1869]: #757
-#1871 := (or #511 #797 #1050 #794 #1049 #1227 #365 #1240 #394)
-#1872 := [th-lemma arith assign-bounds 1 1 1 1 1 1 1 1]: #1871
-#1873 := [unit-resolution #1872 #1122 #791 #787 #1217 #1202 #1870 #1718 #1723]: #511
-#1874 := (or #568 #540 #844 #510)
-#1875 := [unit-resolution #1260 #773]: #1874
-#1876 := [unit-resolution #1875 #1873 #1865 #1867]: #568
-#1877 := [unit-resolution #654 #1876]: #655
-#1878 := [unit-resolution #1265 #1877]: #986
-#1879 := [th-lemma arith farkas -1 1 1 -1 1 -1 -1 1 -1 1 1 -1 1 #903 #900 #1867 #773 #782 #1238 #1718 #791 #1870 #787 #1723 #1217 #1878]: false
-#1881 := [lemma #1879]: #1880
-#1882 := [unit-resolution #1881 #1202 #900]: #394
-#1883 := [unit-resolution #712 #1882]: #709
-#1884 := [unit-resolution #1361 #1883]: #888
-#1885 := (or #481 #735 #844 #1267 #1268 #1223 #870)
-#1392 := (or #481 #735 #844 #1267 #1014 #1268 #1223 #870)
-#1378 := [hypothesis]: #988
-#1386 := [unit-resolution #1271 #1208 #787 #1013 #907 #1217 #782 #900 #889 #1385 #1378 #1252 #773 #903 #1214]: #1269
-#1389 := [unit-resolution #1388 #891 #816 #890 #889 #773]: #568
-#1390 := [unit-resolution #654 #1389]: #655
-#1391 := [unit-resolution #1265 #1390 #1386]: false
-#1393 := [lemma #1391]: #1392
-#1886 := [unit-resolution #1393 #1448]: #1885
-#1887 := [unit-resolution #1886 #1884 #900 #1566 #1852 #1238 #1867]: #481
-#1888 := [unit-resolution #688 #1887]: #685
-#1889 := [unit-resolution #941 #1888]: #757
-#1890 := (or #1064 #797 #1050 #838 #395 #1178 #794)
-#1891 := [th-lemma arith assign-bounds -2 2 -2 -2 2 -1]: #1890
-#1892 := [unit-resolution #1891 #1882 #836 #1889 #1718 #1850 #791]: #1064
-#1893 := (or #1267 #1268 #844 #1090 #1223 #1269 #735)
-#1894 := [unit-resolution #1271 #787 #1013 #907 #1217 #782 #1448 #773 #903]: #1893
-#1895 := [unit-resolution #1894 #1892 #900 #1238 #1867 #1852 #1884]: #1269
-#1896 := [unit-resolution #878 #1888]: #812
-#1897 := (or #1090 #1001 #823 #568 #870 #871 #844)
-#1898 := [th-lemma arith assign-bounds 1 2 2 2 2 2]: #1897
-#1899 := [unit-resolution #1898 #1892 #816 #1867 #1566 #1896 #773]: #568
-#1900 := [unit-resolution #654 #1899]: #655
-#1901 := [unit-resolution #1265 #1900 #1895]: false
-#1903 := [lemma #1901]: #1902
-#1924 := [unit-resolution #1903 #900]: #365
-#1925 := [unit-resolution #720 #1924]: #717
-#2127 := [unit-resolution #1207 #1925]: #745
-#1967 := (or #394 #481)
-#1968 := [unit-resolution #1482 #1852]: #1967
-#2032 := [unit-resolution #1968 #891]: #394
-#2033 := [unit-resolution #712 #2032]: #709
-#2034 := [unit-resolution #856 #2033]: #748
-#1998 := (or #394 #539)
-#1969 := [unit-resolution #1968 #1122]: #481
-#1970 := [unit-resolution #688 #1969]: #685
-#1971 := [unit-resolution #941 #1970]: #757
-#1225 := (or #365 #539 #1227 #794)
-#1218 := (or #539 #794 #1227 #995 #365)
-#1931 := [hypothesis]: #1001
-#1935 := (or #812 #757)
-#1936 := [th-lemma arith farkas 1 1]: #1935
-#1937 := [unit-resolution #1936 #1931]: #757
-#1932 := [hypothesis]: #685
-#1933 := [unit-resolution #878 #1932 #1931]: false
-#1934 := [lemma #1933]: #877
-#1938 := [unit-resolution #1934 #1931]: #876
-#1939 := [unit-resolution #688 #1938]: #482
-#1940 := (or #794 #481 #1179)
-#1941 := [th-lemma arith assign-bounds 2 1]: #1940
-#1942 := [unit-resolution #1941 #1939 #1937]: #1179
-#1943 := [unit-resolution #690 #1939]: #686
-#1944 := [unit-resolution #1171 #1943 #1942]: false
-#1945 := [lemma #1944]: #812
-#1221 := [unit-resolution #1290 #1566 #1850 #1945]: #1218
-#1210 := [unit-resolution #1221 #1202 #818 #1283 #788]: #995
-#1211 := (or #539 #511 #365)
-#1212 := [unit-resolution #1355 #1850]: #1211
-#1213 := [unit-resolution #1212 #1202 #818]: #511
-#1222 := [unit-resolution #682 #1213]: #678
-#1224 := [unit-resolution #993 #1222 #1210]: false
-#1946 := [lemma #1224]: #1225
-#1972 := [unit-resolution #1946 #1723 #818 #1971]: #365
-#1973 := [unit-resolution #720 #1972]: #717
-#1974 := [unit-resolution #1476 #1973]: #1200
-#1913 := (or #568 #394 #539)
-#1904 := [hypothesis]: #569
-#1905 := [unit-resolution #1732 #1904 #897 #787 #791 #907 #773 #1122 #1718 #820 #1870 #1557]: #917
-#1908 := (or #568 #821 #539 #510)
-#1906 := (or #568 #821 #539 #823 #510)
-#1907 := [th-lemma arith assign-bounds 1 1 1 1]: #1906
-#1909 := [unit-resolution #1907 #773]: #1908
-#1910 := [unit-resolution #1909 #1904 #818 #820]: #510
-#1911 := [unit-resolution #680 #1910]: #677
-#1912 := [unit-resolution #959 #1911 #1905]: false
-#1914 := [lemma #1912]: #1913
-#1915 := [unit-resolution #1914 #1122 #818]: #568
-#1916 := [unit-resolution #654 #1915]: #655
-#1975 := [unit-resolution #1464 #1916]: #1430
-#1929 := (or #394 #735 #539)
-#1917 := [unit-resolution #1265 #1916]: #986
-#934 := (or #735 #734)
-#964 := [th-lemma arith farkas 1 1]: #934
-#965 := [unit-resolution #964 #900]: #734
-#1918 := (or #336 #1269 #948 #949 #539 #823 #821 #797 #1050 #794 #1049 #424)
-#1919 := [th-lemma arith assign-bounds 1 1 1 2 1 1 1 1 1 1 1]: #1918
-#1920 := [unit-resolution #1919 #1870 #773 #787 #791 #1847 #903 #965 #818 #1718 #820 #1917]: #336
-#1921 := [unit-resolution #728 #1920]: #725
-#1922 := [unit-resolution #1625 #1921]: #1571
-#1923 := [unit-resolution #878 #1869]: #812
-#1926 := [unit-resolution #1476 #1925]: #1200
-#1428 := (or #337 #735 #739)
-#1239 := [hypothesis]: #336
-#1357 := [unit-resolution #728 #1239]: #725
-#1397 := [unit-resolution #1396 #1357]: #742
-#1150 := (or #795 #796 #739 #735)
-#980 := (or #395 #795 #796 #739 #735)
-#853 := [unit-resolution #712 #783]: #709
-#857 := [unit-resolution #856 #853]: #748
-#763 := (or #739 #738)
-#800 := [th-lemma arith farkas 1 1]: #763
-#801 := [unit-resolution #800 #766]: #738
-#962 := (or #539 #795 #949 #796 #739 #395)
-#826 := (or #510 #821 #539 #795 #395 #822 #823 #796 #824 #825)
-#827 := [th-lemma arith assign-bounds 1 1 1 1 1 1 1 1 1]: #826
-#935 := [unit-resolution #827 #820 #818 #783 #782 #769 #801 #770 #784 #773]: #510
-#936 := [unit-resolution #680 #935]: #677
-#937 := [unit-resolution #832 #936]: #811
-#872 := (or #481 #870 #539 #871 #821 #795 #395 #822 #823 #796 #824 #825)
-#873 := [th-lemma arith assign-bounds 1 2 1 1 1 1 1 1 1 1 1]: #872
-#938 := [unit-resolution #873 #937 #816 #818 #783 #782 #769 #801 #770 #784 #820 #773]: #481
-#939 := [unit-resolution #688 #938]: #685
-#942 := [unit-resolution #941 #939]: #757
-#931 := (or #569 #795 #395 #796 #739)
-#929 := [hypothesis]: #568
-#930 := [th-lemma arith farkas 1 1 -1 1 -1 -1 1 #784 #783 #782 #770 #769 #766 #929]: false
-#932 := [lemma #930]: #931
-#943 := [unit-resolution #932 #783 #784 #770 #766]: #569
-#944 := [unit-resolution #652 #943]: #656
-#945 := [unit-resolution #926 #944]: #887
-#946 := [hypothesis]: #734
-#950 := (or #424 #395 #916 #947 #539 #795 #822 #948 #949 #915 #796 #824 #825)
-#951 := [th-lemma arith assign-bounds 1 1 1 1 2 2 1 1 1 1 1 1]: #950
-#952 := [unit-resolution #951 #818 #903 #783 #882 #782 #769 #946 #801 #770 #784 #857 #945]: #424
-#953 := [unit-resolution #706 #952]: #702
-#957 := [unit-resolution #956 #953]: #928
-#960 := [unit-resolution #959 #936]: #756
-#961 := [th-lemma arith farkas 1 1 1 1 1 1 2 2 1 1 -1 1 -1 -1 1 1 #787 #960 #897 #957 #857 #882 #784 #782 #903 #946 #945 #770 #769 #766 #907 #942]: false
-#963 := [lemma #961]: #962
-#966 := [unit-resolution #963 #783 #965 #770 #766 #784]: #539
-#967 := [unit-resolution #672 #966]: #669
-#968 := [unit-resolution #778 #967]: #762
-#845 := (or #510 #540 #844 #795 #395 #822 #823 #796 #824 #825)
-#846 := [th-lemma arith assign-bounds 1 1 1 1 1 1 1 1 1]: #845
-#969 := [unit-resolution #846 #968 #966 #783 #782 #769 #801 #770 #784 #773]: #510
-#970 := [unit-resolution #680 #969]: #677
-#971 := [unit-resolution #959 #970]: #756
-#972 := [unit-resolution #832 #970]: #811
-#893 := (or #481 #395 #870 #795 #796 #825 #844)
-#817 := [hypothesis]: #738
-#892 := [th-lemma arith farkas 1 1 1 1 1 1 1 1 1 -1 1 #891 #783 #890 #784 #782 #773 #770 #769 #817 #816 #889]: false
-#894 := [lemma #892]: #893
-#973 := [unit-resolution #894 #972 #968 #784 #770 #801 #783]: #481
-#974 := [unit-resolution #688 #973]: #685
-#975 := [unit-resolution #941 #974]: #757
-#918 := (or #915 #916 #794 #795 #796 #739 #735 #917 #424)
-#792 := [hypothesis]: #423
-#908 := [unit-resolution #704 #792]: #701
-#912 := [unit-resolution #911 #908]: #750
-#914 := [th-lemma arith farkas 1/2 -1/2 -1/2 1/2 1/2 -1/2 -1/2 1 -1 1/2 -1/2 -1/2 1/2 -1/2 1/2 -1/2 1 #913 #882 #912 #907 #788 #787 #904 #784 #782 #770 #769 #766 #903 #900 #898 #897 #792]: false
-#919 := [lemma #914]: #918
-#976 := [unit-resolution #919 #975 #945 #784 #770 #766 #900 #971 #857]: #424
-#977 := [unit-resolution #706 #976]: #702
-#978 := [unit-resolution #956 #977]: #928
-#979 := [th-lemma arith farkas 1 1 2 2 1 1 1 -1 1 1 -1 -1 1 -1 1 1 #857 #882 #784 #782 #903 #965 #945 #770 #769 #766 #907 #975 #787 #971 #897 #978]: false
-#981 := [lemma #979]: #980
-#1063 := [unit-resolution #981 #784 #770 #766 #900]: #395
-#1099 := [unit-resolution #1061 #784 #770 #766]: #539
-#1135 := (or #423 #394 #739 #796 #795)
-#1101 := [unit-resolution #672 #1099]: #669
-#1102 := [unit-resolution #778 #1101]: #762
-#1118 := [unit-resolution #1074 #1102 #1099]: #759
-#1116 := (or #510 #795 #796 #739)
-#1086 := (or #423 #510 #795 #796 #825 #540)
-#774 := [hypothesis]: #539
-#775 := [unit-resolution #672 #774]: #669
-#779 := [unit-resolution #778 #775]: #762
-#1075 := [unit-resolution #1074 #779 #774]: #759
-#1078 := [unit-resolution #1077 #1066 #1072]: #838
-#1080 := (or #751 #1048 #795 #822 #821 #823 #796 #824 #825 #1079 #1051 #1014 #1015)
-#1081 := [th-lemma arith assign-bounds 1 1 1 1 1 1 1 1 -1 1 1 -1]: #1080
-#1082 := [unit-resolution #1081 #1078 #1013 #907 #782 #769 #817 #770 #784 #1075 #1023 #897 #773]: #1014
-#1083 := [unit-resolution #1070 #1082]: #1007
-#1084 := [unit-resolution #696 #1083]: #453
-#1085 := [th-lemma arith farkas 1 1 1 1 1 1 1 1 1 1 1 1 1 #989 #1084 #1072 #907 #1066 #773 #784 #782 #770 #769 #817 #779 #774]: false
-#1087 := [lemma #1085]: #1086
-#1100 := [unit-resolution #1087 #989 #784 #770 #801 #1099]: #423
-#1091 := (or #1090 #795 #796 #825 #844 #510 #424)
-#1088 := [hypothesis]: #1064
-#1089 := [th-lemma arith farkas 1 1 1 1 1 1 1 1 1 -1 -1 -1 -1 -1 1 #1088 #907 #773 #784 #782 #770 #769 #817 #816 #994 #889 #989 #787 #912 #792]: false
-#1092 := [lemma #1089]: #1091
-#1103 := [unit-resolution #1092 #989 #770 #801 #1102 #784 #1100]: #1090
-#1104 := [unit-resolution #1098 #1103]: #1094
-#1105 := [unit-resolution #690 #1104]: #481
-#1106 := [unit-resolution #688 #1105]: #685
-#1107 := [unit-resolution #878 #1106]: #812
-#1110 := [unit-resolution #1109 #1105 #897 #869 #1099 #1107 #1023]: #452
-#1111 := [unit-resolution #696 #1110]: #693
-#1112 := [unit-resolution #1070 #1111]: #988
-#1113 := [unit-resolution #704 #1100]: #701
-#1114 := [unit-resolution #911 #1113]: #750
-#1115 := [th-lemma arith farkas -1 -1 -1 1 -1 1 1 -1 1 1 -2 1 -1 1 #907 #773 #784 #782 #770 #769 #897 #1023 #1102 #1114 #1099 #1112 #1013 #766]: false
-#1117 := [lemma #1115]: #1116
-#1119 := [unit-resolution #1117 #784 #770 #766]: #510
-#1120 := [unit-resolution #680 #1119]: #677
-#1121 := [unit-resolution #959 #1120]: #756
-#1125 := [unit-resolution #1124 #1066 #907 #1122 #1072]: #452
-#1126 := [unit-resolution #696 #1125]: #693
-#1127 := [unit-resolution #1045 #1126]: #754
-#1128 := [unit-resolution #1053 #1127 #787 #791 #907 #782 #769 #766 #770 #784 #1121 #1072 #1118 #897 #773]: #794
-#1129 := [unit-resolution #1070 #1126]: #988
-#1132 := [unit-resolution #1131 #1066 #1013 #907 #1122 #1072 #1129]: #481
-#1133 := [unit-resolution #688 #1132]: #685
-#1134 := [unit-resolution #941 #1133 #1128]: false
-#1136 := [lemma #1134]: #1135
-#1137 := [unit-resolution #1136 #1063 #766 #770 #784]: #423
-#1140 := (or #1090 #424 #795 #796 #739)
-#1138 := [unit-resolution #832 #1120]: #811
-#1139 := [th-lemma arith farkas -1 -1 1 -1 -1 -1 -1 1 -1 1 1 1 1 -1 1 #792 #1088 #787 #907 #1119 #773 #784 #782 #770 #769 #766 #1102 #1138 #816 #912]: false
-#1141 := [lemma #1139]: #1140
-#1142 := [unit-resolution #1141 #1137 #784 #770 #766]: #1090
-#1143 := [unit-resolution #1098 #1142]: #1094
-#1144 := [unit-resolution #690 #1143]: #481
-#1145 := [unit-resolution #688 #1144]: #685
-#1146 := [unit-resolution #941 #1145]: #757
-#1147 := [unit-resolution #704 #1137]: #701
-#1148 := [unit-resolution #911 #1147]: #750
-#1149 := [th-lemma arith farkas -1 1 -1 1 1 -1 -1 -1 1 #1121 #897 #1137 #1148 #787 #907 #1146 #1099 #1063]: false
-#1151 := [lemma #1149]: #1150
-#1398 := [unit-resolution #1151 #1397 #766 #900]: #795
-#1399 := [unit-resolution #1207 #1398]: #860
-#1400 := [unit-resolution #720 #1399]: #366
-#1249 := (or #423 #365 #337)
-#1241 := (or #1227 #1240 #337 #1223 #423 #822)
-#1242 := [th-lemma arith assign-bounds -1 -1 -1 1 1]: #1241
-#1243 := [unit-resolution #1242 #1066 #782 #1239 #1217 #1238]: #1227
-#1244 := [unit-resolution #1234 #1243]: #1230
-#1245 := [unit-resolution #714 #1244]: #394
-#1246 := [unit-resolution #712 #1245]: #709
-#1247 := [unit-resolution #1220 #1246]: #888
-#1248 := [th-lemma arith farkas 1 1 1 1 1 #1202 #1247 #1217 #1066 #1245]: false
-#1250 := [lemma #1248]: #1249
-#1401 := [unit-resolution #1250 #1400 #1239]: #423
-#1402 := [unit-resolution #704 #1401]: #701
-#1403 := [unit-resolution #911 #1402]: #750
-#1404 := [unit-resolution #1377 #1400 #1403]: #452
-#1405 := [unit-resolution #696 #1404]: #693
-#1406 := [unit-resolution #1070 #1405]: #988
-#1409 := [unit-resolution #1408 #1402]: #751
-#1333 := (or #510 #796 #838 #739 #735 #1268)
-#1280 := [unit-resolution #1151 #770 #766 #900]: #795
-#1313 := [unit-resolution #1207 #1280]: #860
-#1314 := [unit-resolution #720 #1313]: #366
-#1315 := [unit-resolution #722 #1314]: #718
-#1316 := [unit-resolution #1237 #1315]: #1201
-#1317 := [unit-resolution #1279 #989 #900 #1316 #1252]: #481
-#1318 := [unit-resolution #688 #1317]: #685
-#1319 := [unit-resolution #878 #1318]: #812
-#1302 := (or #1227 #796 #995 #838 #739 #1079 #482 #365 #870)
-#1281 := [hypothesis]: #481
-#1291 := [unit-resolution #688 #1281]: #685
-#1292 := [unit-resolution #878 #1291]: #812
-#1293 := [hypothesis]: #984
-#1294 := [unit-resolution #941 #1291]: #757
-#1295 := [unit-resolution #1290 #1283 #1294 #1282 #813 #1202 #1292 #890]: #539
-#1296 := [unit-resolution #1109 #1295 #1293 #869 #1281 #1292 #897]: #452
-#1297 := [unit-resolution #696 #1296]: #693
-#1298 := [unit-resolution #1045 #1297]: #754
-#1299 := [unit-resolution #672 #1295]: #669
-#1300 := [unit-resolution #778 #1299]: #762
-#1301 := [th-lemma arith farkas -1 1 -1 1 1 -1 -2 2 -2 2 -1 1 -1 1 -3 3 1 #770 #769 #1238 #782 #1300 #773 #1294 #1283 #1217 #787 #816 #1282 #813 #836 #1298 #791 #766]: false
-#1303 := [lemma #1301]: #1302
-#1320 := [unit-resolution #1303 #994 #770 #813 #766 #1023 #1317 #1314 #1306]: #1227
-#1321 := [unit-resolution #1234 #1320]: #1230
-#1322 := [unit-resolution #714 #1321]: #394
-#1323 := [unit-resolution #1312 #989 #1322 #813 #1319]: #539
-#1324 := [unit-resolution #672 #1323]: #669
-#1325 := [unit-resolution #778 #1324]: #762
-#1326 := [unit-resolution #1109 #1323 #1023 #869 #1317 #1319 #897]: #452
-#1327 := [unit-resolution #696 #1326]: #693
-#1328 := [unit-resolution #1045 #1327]: #754
-#1329 := [unit-resolution #941 #1318]: #757
-#1330 := [unit-resolution #712 #1322]: #709
-#1331 := [unit-resolution #1220 #1330]: #888
-#1332 := [th-lemma arith farkas -1 1 -1 1 -4 2 -2 -2 2 -3 3 1 -1 -1 1 -1 1 1 #770 #769 #1316 #782 #1322 #1331 #1329 #1217 #787 #1328 #791 #1325 #773 #816 #994 #813 #836 #766]: false
-#1334 := [lemma #1332]: #1333
-#1410 := [unit-resolution #1334 #1397 #1409 #766 #900 #1403]: #510
-#1411 := [unit-resolution #1355 #1410 #1400 #1409]: #539
-#1412 := [unit-resolution #680 #1410]: #677
-#1413 := [unit-resolution #959 #1412]: #756
-#1383 := (or #394 #917 #540 #424 #1014)
-#1379 := [unit-resolution #1274 #1122 #907 #1378 #1013 #912]: #481
-#1380 := [unit-resolution #688 #1379]: #685
-#1381 := [unit-resolution #941 #1380]: #757
-#1382 := [th-lemma arith farkas -1 1 -1 1 1 -1 -1 1 1 #787 #898 #897 #774 #792 #1122 #912 #907 #1381]: false
-#1384 := [lemma #1382]: #1383
-#1414 := [unit-resolution #1384 #1413 #1411 #1401 #1406]: #394
-#1415 := [unit-resolution #712 #1414]: #709
-#1416 := [unit-resolution #1361 #1415]: #888
-#1417 := (or #794 #1049 #917 #1051 #540 #1268 #1048 #1267 #1240 #365)
-#1418 := [th-lemma arith assign-bounds -1 1 -1 1 -1 1 -1 1 -1]: #1417
-#1419 := [unit-resolution #1418 #1400 #787 #907 #1217 #897 #1411 #1413 #1403 #1416]: #794
-#1420 := [unit-resolution #832 #1412]: #811
-#1421 := [unit-resolution #722 #1400]: #718
-#1422 := [unit-resolution #1237 #1421]: #1201
-#1423 := [unit-resolution #672 #1411]: #669
-#1424 := [unit-resolution #778 #1423]: #762
-#1425 := [unit-resolution #1393 #1424 #900 #1416 #1406 #1403 #1422 #1420]: #481
-#1426 := [unit-resolution #688 #1425]: #685
-#1427 := [unit-resolution #941 #1426 #1419]: false
-#1429 := [lemma #1427]: #1428
-#1927 := [unit-resolution #1429 #1920 #900]: #739
-#1928 := [th-lemma arith farkas -1 -1 1/2 -1/2 1/2 1/2 1/2 -1/2 -1/2 1/2 -1/2 1/2 -1/2 -1/2 1/2 1 #1537 #1927 #1917 #903 #900 #1926 #1488 #1494 #1739 #1448 #1013 #1923 #869 #1712 #882 #1922]: false
-#1930 := [lemma #1928]: #1929
-#1976 := [unit-resolution #1930 #1122 #818]: #735
-#1965 := (or #510 #539 #899 #794 #1227 #1498)
-#1947 := [unit-resolution #1946 #1283 #818 #788]: #365
-#1948 := [unit-resolution #720 #1947]: #717
-#1949 := [unit-resolution #1476 #1948]: #1200
-#1950 := (or #336 #1240 #1500 #1501 #1227 #510 #797 #1050 #794 #1049 #995 #871 #838 #1178 #539)
-#1951 := [th-lemma arith assign-bounds 1 1 1 1 1 3 3 1 1 2 2 2 2 2]: #1950
-#1952 := [unit-resolution #1951 #989 #816 #787 #791 #836 #1217 #1494 #818 #788 #1718 #1850 #1283 #994 #1949]: #336
-#1953 := [unit-resolution #728 #1952]: #725
-#1954 := [unit-resolution #1625 #1953]: #1571
-#1955 := [hypothesis]: #735
-#1956 := [hypothesis]: #1430
-#1957 := [th-lemma arith assign-bounds 1 -1 1 -1 -1 1 1 3 -3 1 -1 -1 -2 2 2 -2 #1217 #1949 #1956 #1491 #1488 #1494 #1739 #1718 #791 #788 #787 #1283 #994 #816 #1850 #836]: #734
-#1958 := [unit-resolution #1515 #1957 #1955]: #64
-#1959 := [unit-resolution #658 #1958]: #668
-#1960 := [unit-resolution #1207 #1948]: #745
-#1961 := [unit-resolution #1396 #1953]: #742
-#1962 := [unit-resolution #1061 #1961 #818 #1960]: #739
-#1963 := [unit-resolution #1544 #1962 #1959]: #825
-#1964 := [th-lemma arith farkas -1 -1 1 1 -1 -1 1 -1 -1 1 -1 1 1 #1537 #1963 #1949 #1488 #1494 #1739 #994 #816 #1718 #791 #1850 #836 #1954]: false
-#1966 := [lemma #1964]: #1965
-#1977 := [unit-resolution #1966 #1976 #818 #1971 #1723 #1975]: #510
-#1978 := (or #744 #838 #511 #797 #1050 #794 #1049)
-#1979 := [th-lemma arith assign-bounds -1 -2 -2 2 -2 2]: #1978
-#1980 := [unit-resolution #1979 #1971 #791 #787 #1718 #1850 #1977]: #744
-#1983 := (or #1177 #1500 #336 #1267)
-#1981 := (or #1177 #1268 #1500 #336 #1501 #1267 #1240)
-#1982 := [th-lemma arith assign-bounds 1 2 2 2 2 2]: #1981
-#1984 := [unit-resolution #1982 #1494 #1852 #1217]: #1983
-#1985 := [unit-resolution #1984 #1980 #1974 #1748]: #336
-#1986 := [unit-resolution #728 #1985]: #725
-#1987 := [unit-resolution #1396 #1986]: #742
-#1988 := [unit-resolution #1625 #1986]: #1571
-#1989 := (or #738 #1627 #1500 #1177 #1754)
-#1990 := [unit-resolution #1756 #869 #1013 #836 #1494 #1537 #1566 #1945 #1448 #816 #1488]: #1989
-#1991 := [unit-resolution #1990 #1988 #1739 #1980 #1974]: #738
-#1992 := [unit-resolution #1207 #1973]: #745
-#1993 := [unit-resolution #1061 #1987 #818 #1992]: #739
-#1994 := [unit-resolution #1544 #1993 #1991]: #65
-#1995 := [unit-resolution #658 #1994]: #667
-#1996 := [unit-resolution #1515 #1995 #1976]: #949
-#1997 := [th-lemma arith farkas -1 -1 1/2 1/2 -1/2 -1/2 1/2 -1/2 1/2 1/2 -1/2 1/2 1/2 -1/2 -1/2 1 #769 #1991 #1992 #773 #782 #820 #1718 #791 #1217 #1975 #1491 #1996 #1971 #787 #1723 #1987]: false
-#1999 := [lemma #1997]: #1998
-#2000 := [unit-resolution #1999 #818]: #394
-#2001 := (or #539 #510 #395)
-#2002 := [unit-resolution #1312 #1850 #1945]: #2001
-#2003 := [unit-resolution #2002 #2000 #818]: #510
-#2008 := (or #1090 #511 #539)
-#2006 := (or #1090 #1001 #870 #511 #539)
-#2004 := (or #1090 #1001 #870 #871 #511 #539)
-#2005 := [th-lemma arith assign-bounds 1 2 2 2 2]: #2004
-#2007 := [unit-resolution #2005 #816]: #2006
-#2009 := [unit-resolution #2007 #1566 #1945]: #2008
-#2010 := [unit-resolution #2009 #2003 #818]: #1090
-#2011 := (or #1064 #395 #794)
-#2012 := [unit-resolution #1891 #836 #1718 #1850 #791]: #2011
-#2013 := [unit-resolution #2012 #2010 #2000]: #794
-#2014 := (or #481 #511 #539)
-#2015 := [unit-resolution #1455 #1566]: #2014
-#2016 := [unit-resolution #2015 #2003 #818]: #481
-#2017 := [unit-resolution #688 #2016]: #685
-#2018 := [unit-resolution #941 #2017 #2013]: false
-#2019 := [lemma #2018]: #539
-#2023 := [unit-resolution #672 #2019]: #669
-#2024 := [unit-resolution #778 #2023]: #762
-#2035 := (or #568 #844 #481)
-#2036 := [unit-resolution #1460 #1566]: #2035
-#2037 := [unit-resolution #2036 #891 #2024]: #568
-#2038 := [unit-resolution #654 #2037]: #655
-#2039 := [unit-resolution #1265 #2038]: #986
-#2030 := (or #735 #1090 #1269 #916)
-#2025 := [hypothesis]: #986
-#2026 := (or #735 #1269 #1090 #795 #844 #916)
-#2027 := [unit-resolution #1512 #787 #1013 #882 #782 #903 #773 #1828]: #2026
-#2028 := [unit-resolution #2027 #900 #1088 #2025 #2024 #913]: #795
-#2029 := [unit-resolution #1207 #1925 #2028]: false
-#2031 := [lemma #2029]: #2030
-#2040 := [unit-resolution #2031 #1208 #2039 #2034]: #735
-#2041 := [unit-resolution #1464 #2038]: #1430
-#2068 := (or #510 #481)
-#2042 := [unit-resolution #1496 #2023]: #933
-#1848 := (<= #1199 0::Int)
-#2043 := (or #366 #947 #838 #1178 #916 #1179 #481 #510 #1002)
-#2044 := [th-lemma arith assign-bounds 1 1 1 1 1 1 1 1]: #2043
-#2045 := [unit-resolution #2044 #989 #869 #836 #882 #891 #1850 #2034 #1172]: #366
-#2046 := [unit-resolution #722 #2045]: #718
-#2047 := (or #1235 #1848)
-#2048 := [th-lemma arith triangle-eq]: #2047
-#2049 := [unit-resolution #2048 #2046]: #1848
-#2050 := (not #1848)
-#2051 := (or #734 #1503 #797 #1050 #947 #1498 #1499 #1504 #1501 #916 #1179 #1002 #2050 #838 #1178)
-#2052 := [th-lemma arith assign-bounds 1 1 -1 -1 1 -1 -1 1 1 -1 1 -1 2 -2]: #2051
-#2053 := [unit-resolution #2052 #2049 #869 #791 #836 #882 #1494 #1491 #1718 #1850 #2034 #2042 #1172 #2041 #1488]: #734
-#2054 := [unit-resolution #1515 #2053 #2040]: #64
-#2055 := [unit-resolution #658 #2054]: #668
-#2056 := [unit-resolution #1569 #990]: #984
-#2057 := (or #336 #797 #1050 #947 #1501 #916 #1179 #510 #1002 #2050 #838 #1178)
-#2058 := [th-lemma arith assign-bounds 1 1 1 1 1 1 1 1 1 2 2]: #2057
-#2059 := [unit-resolution #2058 #989 #791 #836 #882 #1494 #869 #1718 #1850 #2034 #1172 #2049]: #336
-#2060 := [unit-resolution #728 #2059]: #725
-#2061 := [unit-resolution #1625 #2060]: #1571
-#2062 := [th-lemma arith assign-bounds 1 -1 -1 -1 1 -3 3 -1 1 -1 1 1 2 -2 2 -2 #2061 #1537 #1494 #1718 #791 #1850 #836 #2042 #1488 #2056 #897 #2049 #882 #2034 #1172 #869]: #738
-#2063 := [unit-resolution #1361 #2033]: #888
-#2064 := [unit-resolution #1237 #2046]: #1201
-#2065 := [unit-resolution #1396 #2060]: #742
-#2066 := [th-lemma arith assign-bounds 1 -1 -1 -1 1 -3 3 -1 1 -1 1 1 2 -2 2 -2 #2065 #769 #782 #1448 #1013 #1852 #907 #2024 #773 #994 #816 #2064 #1217 #2063 #1208 #787]: #739
-#2067 := [unit-resolution #1544 #2066 #2062 #2055]: false
-#2069 := [lemma #2067]: #2068
-#2103 := [unit-resolution #2069 #891]: #510
-#2101 := (or #1235 #1090 #1267 #511 #899 #916 #1179 #1498)
-#2083 := [hypothesis]: #718
-#2084 := [unit-resolution #1237 #2083]: #1201
-#2085 := [unit-resolution #959 #1336]: #756
-#2086 := [hypothesis]: #1161
-#2087 := [unit-resolution #2048 #2083]: #1848
-#2088 := [unit-resolution #2052 #2087 #869 #791 #836 #882 #1494 #1491 #1718 #1850 #913 #2042 #2086 #1956 #1488]: #734
-#2089 := [unit-resolution #1515 #2088 #1955]: #64
-#2090 := [unit-resolution #658 #2089]: #668
-#2081 := (or #739 #1267 #1090 #1223 #511 #2050)
-#2071 := [hypothesis]: #1848
-#2073 := (or #1526 #739 #2050)
-#2070 := [hypothesis]: #1433
-#2072 := [th-lemma arith farkas -1 -1 -1 -1 1 1 1 -1 1 -1 1 -1 1 #769 #766 #1566 #2024 #773 #816 #1850 #836 #1718 #791 #1494 #2071 #2070]: false
-#2074 := [lemma #2072]: #2073
-#2075 := [unit-resolution #2074 #766 #2071]: #1526
-#2076 := [unit-resolution #1641 #2075]: #1522
-#2077 := [unit-resolution #730 #2076]: #336
-#2078 := [unit-resolution #728 #2077]: #725
-#2079 := [unit-resolution #1396 #2078]: #742
-#2080 := [th-lemma arith farkas -1/2 1/2 1 -1/2 -1 1 -1 1/2 -3/2 3/2 1/2 -1/2 -1/2 -1/2 -1/2 1/2 1/2 1 #1448 #1013 #1217 #782 #1385 #1088 #787 #1214 #1852 #907 #2079 #769 #766 #1566 #2024 #773 #816 #1335]: false
-#2082 := [lemma #2080]: #2081
-#2091 := [unit-resolution #2082 #2084 #1088 #1385 #1335 #2087]: #739
-#2092 := [unit-resolution #1544 #2091 #2090]: #825
-#2093 := (or #1538 #1539 #738 #917 #1503 #1504 #1051 #1268 #1048 #1014 #1015 #822 #1223)
-#2094 := [th-lemma arith assign-bounds -1 -1 -1 -1 1 1 1 -1 1 -1 1 -1]: #2093
-#2095 := [unit-resolution #2094 #2092 #1013 #907 #782 #1537 #897 #2085 #1448 #1852 #2042 #2084 #1488]: #1538
-#2096 := [unit-resolution #1667 #2095]: #1522
-#2097 := [unit-resolution #730 #2096]: #336
-#2098 := [unit-resolution #728 #2097]: #725
-#2099 := [unit-resolution #1625 #2098]: #1571
-#2100 := [th-lemma arith farkas -1 -1 -2 -1 -1 1 1 1 -1 1 -1 1 -1 1 #1537 #2092 #2097 #2085 #2042 #1488 #897 #1852 #907 #1448 #1013 #782 #2084 #2099]: false
-#2102 := [lemma #2100]: #2101
-#2104 := [unit-resolution #2102 #1208 #2063 #2103 #2040 #2034 #1172 #2041]: #1235
-#2105 := [unit-resolution #722 #2104]: #365
-#2106 := (or #741 #797 #947 #916 #838 #1178 #366)
-#2107 := [th-lemma arith assign-bounds -1 2 -2 -2 2 -2]: #2106
-#2108 := [unit-resolution #2107 #2105 #882 #1718 #1850 #2034 #836]: #741
-#2109 := [unit-resolution #720 #2105]: #717
-#2110 := [unit-resolution #1476 #2109]: #1200
-#2111 := (or #734 #1498 #1179 #1500 #1502 #1503 #1267)
-#2112 := [unit-resolution #1506 #869 #791 #1217 #1494 #1488 #1491]: #2111
-#2113 := [unit-resolution #2112 #2110 #2042 #2041 #1172 #2063 #2108]: #734
-#2114 := [unit-resolution #1515 #2113 #2040]: #64
-#2115 := [unit-resolution #680 #2103]: #677
-#2116 := [unit-resolution #959 #2115]: #756
-#2117 := [unit-resolution #1207 #2109]: #745
-#2118 := (or #738 #795 #916 #917 #1503)
-#2119 := [unit-resolution #1676 #1850 #1828]: #2118
-#2120 := [unit-resolution #2119 #2117 #2042 #2116 #2034]: #738
-#2121 := (or #739 #795 #844 #1502 #1500 #1267)
-#2122 := [unit-resolution #1651 #1852]: #2121
-#2123 := [unit-resolution #2122 #2108 #2117 #2024 #2110 #2063]: #739
-#2124 := [unit-resolution #1544 #2123 #2120]: #65
-#2125 := [unit-resolution #658 #2124 #2114]: false
-#2126 := [lemma #2125]: #481
-#2149 := [unit-resolution #688 #2126]: #685
-#2020 := [hypothesis]: #794
-#2021 := [unit-resolution #941 #1932 #2020]: false
-#2022 := [lemma #2021]: #940
-#2150 := [unit-resolution #2022 #2149]: #757
-#2147 := (or #510 #735)
-#2136 := (or #916 #1001 #482 #947 #510 #1002 #838 #1178 #366)
-#2137 := [th-lemma arith assign-bounds -1 1 -1 -1 1 1 -1 1]: #2136
-#2138 := [unit-resolution #2137 #989 #869 #836 #882 #1924 #1850 #2126 #1945]: #916
-#2130 := (not #1708)
-#2139 := [unit-resolution #1875 #989 #2019 #2024]: #568
-#2140 := [unit-resolution #654 #2139]: #655
-#2141 := [unit-resolution #1265 #2140]: #986
-#2131 := (or #2130 #1079 #1269 #735)
-#2128 := [hypothesis]: #1708
-#2129 := [th-lemma arith farkas 1 -1 -1 1 -3/2 3/2 1/2 -1/2 -1/2 1/2 1/2 -1/2 -1/2 1/2 -1/2 1/2 -1/2 1 #1293 #897 #1852 #907 #1448 #1013 #2128 #1945 #882 #869 #2127 #2024 #2025 #903 #900 #773 #782 #2019]: false
-#2132 := [lemma #2129]: #2131
-#2142 := [unit-resolution #2132 #2056 #2141 #900]: #2130
-#2133 := [hypothesis]: #2130
-#2134 := [unit-resolution #1711 #1229 #2133]: false
-#2135 := [lemma #2134]: #1710
-#2143 := [unit-resolution #2135 #2142]: #1230
-#2144 := [unit-resolution #714 #2143]: #394
-#2145 := [unit-resolution #712 #2144]: #709
-#2146 := [unit-resolution #856 #2145 #2138]: false
-#2148 := [lemma #2146]: #2147
-#2151 := [unit-resolution #2148 #900]: #510
-#2152 := [unit-resolution #680 #2151]: #677
-#2153 := [unit-resolution #959 #2152]: #756
-#2154 := (or #735 #844 #916 #795 #794 #917 #1503)
-#2155 := [unit-resolution #1664 #1828]: #2154
-#2156 := [unit-resolution #2155 #2153 #2042 #2024 #2150 #900 #2127]: #916
-#2159 := (or #394 #917 #540)
-#2157 := (or #394 #917 #540 #424)
-#2158 := [unit-resolution #1384 #1448]: #2157
-#2160 := [unit-resolution #2158 #1847]: #2159
-#2161 := [unit-resolution #2160 #2153 #2019]: #394
-#2162 := [unit-resolution #712 #2161]: #709
-#2163 := [unit-resolution #856 #2162 #2156]: false
-#2164 := [lemma #2163]: #735
-#2208 := (or #365 #510)
-#2187 := [unit-resolution #1464 #2140]: #1430
-#2188 := (or #1161 #482)
-#2189 := [unit-resolution #1681 #1945]: #2188
-#2190 := [unit-resolution #2189 #2126]: #1161
-#2165 := [unit-resolution #2048 #1226]: #1848
-#2185 := (or #394 #1079 #1269 #1498 #365 #995)
-#2168 := (or #336 #365 #2050 #394)
-#2166 := (or #336 #1501 #365 #2050 #394)
-#2167 := [th-lemma arith assign-bounds 1 1 1 1]: #2166
-#2169 := [unit-resolution #2167 #1494]: #2168
-#2170 := [unit-resolution #2169 #1122 #1202 #2165]: #336
-#2171 := [unit-resolution #728 #2170]: #725
-#2172 := [unit-resolution #1396 #2171]: #742
-#2173 := (or #1227 #796 #995 #739 #1079 #482 #365)
-#2174 := [unit-resolution #1303 #1566 #1850]: #2173
-#2175 := [unit-resolution #2174 #2172 #2126 #1293 #1202 #1282 #1723]: #739
-#2176 := [unit-resolution #2135 #1709]: #1708
-#2177 := (or #734 #2130 #1014 #1015 #1001 #947 #1002 #1503 #1498 #1499 #1504 #1501 #2050)
-#2178 := [th-lemma arith assign-bounds 1 -1 1 -1 -1 1 1 1 -1 -1 1 -1]: #2177
-#2179 := [unit-resolution #2178 #2176 #869 #1013 #882 #1494 #1491 #1945 #1448 #2042 #1956 #2165 #1488]: #734
-#2180 := [unit-resolution #1515 #2179 #2164]: #64
-#2181 := [unit-resolution #658 #2180]: #668
-#2182 := [unit-resolution #1544 #2181 #2175]: #825
-#2183 := [unit-resolution #1625 #2171]: #1571
-#2184 := [th-lemma arith farkas -1 1 1 -1 -2 2 -2 -1 1 -1 1 -1 1 -1 1 1 #2183 #1537 #1293 #897 #2025 #903 #2179 #1448 #1013 #1852 #907 #2024 #773 #782 #1238 #2182]: false
-#2186 := [lemma #2184]: #2185
-#2191 := [unit-resolution #2186 #1202 #2141 #2187 #2056 #994]: #394
-#2192 := [unit-resolution #712 #2191]: #709
-#2193 := [unit-resolution #856 #2192]: #748
-#2194 := [unit-resolution #2052 #2193 #869 #791 #836 #882 #1494 #1491 #1718 #1850 #2165 #2042 #2190 #2187 #1488]: #734
-#2195 := [unit-resolution #1515 #2194 #2164]: #64
-#2196 := [unit-resolution #658 #2195]: #668
-#2197 := [unit-resolution #1361 #2192]: #888
-#2198 := (or #753 #395 #1267)
-#2199 := [th-lemma arith assign-bounds 2 -1]: #2198
-#2200 := [unit-resolution #2199 #2197 #2191]: #753
-#2201 := [unit-resolution #2058 #2193 #791 #836 #882 #1494 #869 #1718 #1850 #989 #2190 #2165]: #336
-#2202 := [unit-resolution #728 #2201]: #725
-#2203 := [unit-resolution #1396 #2202]: #742
-#2204 := [unit-resolution #2174 #2203 #2126 #2056 #1202 #994 #2200]: #739
-#2205 := [unit-resolution #1544 #2204 #2196]: #825
-#2206 := [unit-resolution #1625 #2202]: #1571
-#2207 := [th-lemma arith farkas -1 1 1 -1 -2 2 -2 -1 1 -1 1 -1 1 -1 1 1 #2206 #1537 #2056 #897 #2141 #903 #2194 #1448 #1013 #1852 #907 #2024 #773 #782 #1238 #2205]: false
-#2209 := [lemma #2207]: #2208
-#2210 := [unit-resolution #2209 #989]: #365
-#2231 := [unit-resolution #2137 #2210 #869 #836 #882 #989 #1850 #2126 #1945]: #916
-#2229 := (or #2130 #510)
-#2211 := [unit-resolution #720 #2210]: #717
-#2212 := [unit-resolution #1476 #2211]: #1200
-#2213 := (or #1848 #1500 #366)
-#2214 := [th-lemma arith assign-bounds 1 -2]: #2213
-#2215 := [unit-resolution #2214 #2212 #2210]: #1848
-#2216 := [unit-resolution #2178 #2128 #869 #1013 #882 #1494 #1491 #1945 #1448 #2042 #2187 #2215 #1488]: #734
-#2217 := [unit-resolution #1515 #2216 #2164]: #64
-#2218 := [unit-resolution #658 #2217]: #668
-#2219 := [unit-resolution #1207 #2211]: #745
-#2220 := (or #336 #844 #1269 #948 #949 #823 #510)
-#2221 := [th-lemma arith assign-bounds 1 1 1 1 1 1]: #2220
-#2222 := [unit-resolution #2221 #2216 #773 #903 #989 #2024 #2141]: #336
-#2223 := [unit-resolution #728 #2222]: #725
-#2224 := [unit-resolution #1396 #2223]: #742
-#2225 := [unit-resolution #1117 #2224 #2219 #989]: #739
-#2226 := [unit-resolution #1544 #2225 #2218]: #825
-#2227 := [unit-resolution #1625 #2223]: #1571
-#2228 := [th-lemma arith farkas -2 2 -1 -1 1 -1 1 -1 -1 1 1 1 -1 -1 1 1 #1448 #1013 #1945 #882 #869 #2141 #903 #2216 #2227 #1537 #2226 #2056 #897 #1852 #907 #2128]: false
-#2230 := [lemma #2228]: #2229
-#2232 := [unit-resolution #2230 #989]: #2130
-#2233 := [unit-resolution #2135 #2232]: #1230
-#2234 := [unit-resolution #714 #2233]: #394
-#2235 := [unit-resolution #712 #2234]: #709
-#2236 := [unit-resolution #856 #2235 #2231]: false
-#2237 := [lemma #2236]: #510
-#2238 := [unit-resolution #680 #2237]: #677
-#2239 := [unit-resolution #959 #2238]: #756
-#2240 := [unit-resolution #2160 #2239 #2019]: #394
-#2241 := [unit-resolution #1979 #2237 #791 #787 #1718 #1850 #2150]: #744
-#2242 := [unit-resolution #712 #2240]: #709
-#2243 := [unit-resolution #1361 #2242]: #888
-#2244 := (or #1177 #1267 #365 #395)
-#2245 := [unit-resolution #1780 #1852]: #2244
-#2246 := [unit-resolution #2245 #2243 #2241 #2240]: #365
-#2247 := [unit-resolution #720 #2246]: #717
-#2248 := [unit-resolution #1476 #2247]: #1200
-#2249 := (or #741 #794 #917 #540)
-#2250 := [unit-resolution #1808 #787 #897 #1718]: #2249
-#2251 := [unit-resolution #2250 #2239 #2019 #2150]: #741
-#2252 := [unit-resolution #2012 #2240 #2150]: #1064
-#2253 := (or #1090 #568 #844)
-#2254 := [unit-resolution #1898 #816 #1945 #1566 #773]: #2253
-#2255 := [unit-resolution #2254 #2252 #2024]: #568
-#2256 := [unit-resolution #654 #2255]: #655
-#2257 := [unit-resolution #1464 #2256]: #1430
-#2258 := [unit-resolution #2112 #2257 #2042 #2251 #2190 #2243 #2248]: #734
-#2259 := [unit-resolution #1515 #2258 #2164]: #64
-#2260 := [unit-resolution #1207 #2247]: #745
-#2261 := [unit-resolution #856 #2242]: #748
-#2262 := [unit-resolution #2119 #2261 #2042 #2260 #2239]: #738
-#2263 := [unit-resolution #2122 #2248 #2251 #2024 #2260 #2243]: #739
-#2264 := [unit-resolution #1544 #2263 #2262]: #65
-[unit-resolution #658 #2264 #2259]: false
-unsat
-68356683e9cf34e34d65674fa3c8a62835e193a4 341 0
-#2 := false
-#24 := 0::Int
-decl f3 :: Int
-#7 := f3
-#433 := (<= f3 0::Int)
-#443 := (>= f3 0::Int)
-#754 := (not #443)
-#410 := (not #433)
-#755 := (or #410 #754)
-#716 := (not #755)
-#10 := 2::Int
-#763 := (mod f3 2::Int)
-#111 := -1::Int
-#420 := (* -1::Int #763)
-decl f4 :: (-> S2 Int Int)
-decl f5 :: (-> S3 Int S2)
-decl f6 :: S3
-#11 := f6
-#12 := (f5 f6 f3)
-#13 := (f4 #12 2::Int)
-#550 := (+ #13 #420)
-#757 := (= #550 0::Int)
-#706 := (not #757)
-#718 := (>= #550 0::Int)
-#663 := (not #718)
-#658 := [hypothesis]: #718
-#696 := (>= #763 0::Int)
-#1 := true
-#69 := [true-axiom]: true
-#659 := (or false #696)
-#660 := [th-lemma arith]: #659
-#661 := [unit-resolution #660 #69]: #696
-#99 := (>= #13 0::Int)
-#102 := (not #99)
-#8 := 1::Int
-#14 := (* 2::Int #13)
-#15 := (+ #14 1::Int)
-#16 := (+ f3 #15)
-#9 := (+ f3 1::Int)
-#17 := (<= #9 #16)
-#18 := (not #17)
-#107 := (iff #18 #102)
-#81 := (+ f3 #14)
-#82 := (+ 1::Int #81)
-#72 := (+ 1::Int f3)
-#87 := (<= #72 #82)
-#90 := (not #87)
-#105 := (iff #90 #102)
-#97 := (>= #14 0::Int)
-#93 := (not #97)
-#103 := (iff #93 #102)
-#100 := (iff #97 #99)
-#101 := [rewrite]: #100
-#104 := [monotonicity #101]: #103
-#94 := (iff #90 #93)
-#95 := (iff #87 #97)
-#96 := [rewrite]: #95
-#98 := [monotonicity #96]: #94
-#106 := [trans #98 #104]: #105
-#91 := (iff #18 #90)
-#88 := (iff #17 #87)
-#85 := (= #16 #82)
-#75 := (+ 1::Int #14)
-#78 := (+ f3 #75)
-#83 := (= #78 #82)
-#84 := [rewrite]: #83
-#79 := (= #16 #78)
-#76 := (= #15 #75)
-#77 := [rewrite]: #76
-#80 := [monotonicity #77]: #79
-#86 := [trans #80 #84]: #85
-#73 := (= #9 #72)
-#74 := [rewrite]: #73
-#89 := [monotonicity #74 #86]: #88
-#92 := [monotonicity #89]: #91
-#108 := [trans #92 #106]: #107
-#71 := [asserted]: #18
-#109 := [mp #71 #108]: #102
-#662 := [th-lemma arith farkas -1 1 1 #109 #661 #658]: false
-#664 := [lemma #662]: #663
-#673 := (or #706 #718)
-#653 := [th-lemma arith triangle-eq]: #673
-#654 := [unit-resolution #653 #664]: #706
-#645 := (or #716 #757)
-#742 := -2::Int
-#431 := (* -1::Int f3)
-#466 := (mod #431 -2::Int)
-#362 := (+ #13 #466)
-#461 := (= #362 0::Int)
-#740 := (if #755 #757 #461)
-#442 := (= #13 0::Int)
-#441 := (= f3 0::Int)
-#451 := (if #441 #442 #740)
-#22 := (:var 0 Int)
-#20 := (:var 1 Int)
-#42 := (f5 f6 #20)
-#43 := (f4 #42 #22)
-#776 := (pattern #43)
-#115 := (* -1::Int #22)
-#112 := (* -1::Int #20)
-#170 := (mod #112 #115)
-#285 := (+ #43 #170)
-#286 := (= #285 0::Int)
-#44 := (mod #20 #22)
-#282 := (* -1::Int #44)
-#283 := (+ #43 #282)
-#284 := (= #283 0::Int)
-#137 := (<= #22 0::Int)
-#144 := (>= #20 0::Int)
-#229 := (or #144 #137)
-#230 := (not #229)
-#133 := (<= #20 0::Int)
-#227 := (or #133 #137)
-#228 := (not #227)
-#233 := (or #228 #230)
-#287 := (if #233 #284 #286)
-#281 := (= #43 0::Int)
-#25 := (= #20 0::Int)
-#288 := (if #25 #281 #287)
-#280 := (= #43 #20)
-#26 := (= #22 0::Int)
-#289 := (if #26 #280 #288)
-#777 := (forall (vars (?v0 Int) (?v1 Int)) (:pat #776) #289)
-#292 := (forall (vars (?v0 Int) (?v1 Int)) #289)
-#780 := (iff #292 #777)
-#778 := (iff #289 #289)
-#779 := [refl]: #778
-#781 := [quant-intro #779]: #780
-#176 := (* -1::Int #170)
-#249 := (if #233 #44 #176)
-#252 := (if #25 0::Int #249)
-#255 := (if #26 #20 #252)
-#258 := (= #43 #255)
-#261 := (forall (vars (?v0 Int) (?v1 Int)) #258)
-#293 := (iff #261 #292)
-#290 := (iff #258 #289)
-#291 := [rewrite]: #290
-#294 := [quant-intro #291]: #293
-#138 := (not #137)
-#145 := (not #144)
-#148 := (and #145 #138)
-#134 := (not #133)
-#141 := (and #134 #138)
-#151 := (or #141 #148)
-#196 := (if #151 #44 #176)
-#199 := (if #25 0::Int #196)
-#202 := (if #26 #20 #199)
-#205 := (= #43 #202)
-#208 := (forall (vars (?v0 Int) (?v1 Int)) #205)
-#262 := (iff #208 #261)
-#259 := (iff #205 #258)
-#256 := (= #202 #255)
-#253 := (= #199 #252)
-#250 := (= #196 #249)
-#234 := (iff #151 #233)
-#231 := (iff #148 #230)
-#232 := [rewrite]: #231
-#221 := (iff #141 #228)
-#222 := [rewrite]: #221
-#235 := [monotonicity #222 #232]: #234
-#251 := [monotonicity #235]: #250
-#254 := [monotonicity #251]: #253
-#257 := [monotonicity #254]: #256
-#260 := [monotonicity #257]: #259
-#263 := [quant-intro #260]: #262
-#219 := (~ #208 #208)
-#218 := (~ #205 #205)
-#215 := [refl]: #218
-#220 := [nnf-pos #215]: #219
-#36 := (- #22)
-#35 := (- #20)
-#45 := (mod #35 #36)
-#46 := (- #45)
-#29 := (< 0::Int #22)
-#31 := (< #20 0::Int)
-#32 := (and #31 #29)
-#28 := (< 0::Int #20)
-#30 := (and #28 #29)
-#33 := (or #30 #32)
-#47 := (if #33 #44 #46)
-#48 := (if #25 0::Int #47)
-#49 := (if #26 #20 #48)
-#50 := (= #43 #49)
-#51 := (forall (vars (?v0 Int) (?v1 Int)) #50)
-#211 := (iff #51 #208)
-#181 := (if #33 #44 #176)
-#184 := (if #25 0::Int #181)
-#187 := (if #26 #20 #184)
-#190 := (= #43 #187)
-#193 := (forall (vars (?v0 Int) (?v1 Int)) #190)
-#209 := (iff #193 #208)
-#206 := (iff #190 #205)
-#203 := (= #187 #202)
-#200 := (= #184 #199)
-#197 := (= #181 #196)
-#152 := (iff #33 #151)
-#149 := (iff #32 #148)
-#139 := (iff #29 #138)
-#140 := [rewrite]: #139
-#146 := (iff #31 #145)
-#147 := [rewrite]: #146
-#150 := [monotonicity #147 #140]: #149
-#142 := (iff #30 #141)
-#135 := (iff #28 #134)
-#136 := [rewrite]: #135
-#143 := [monotonicity #136 #140]: #142
-#153 := [monotonicity #143 #150]: #152
-#198 := [monotonicity #153]: #197
-#201 := [monotonicity #198]: #200
-#204 := [monotonicity #201]: #203
-#207 := [monotonicity #204]: #206
-#210 := [quant-intro #207]: #209
-#194 := (iff #51 #193)
-#191 := (iff #50 #190)
-#188 := (= #49 #187)
-#185 := (= #48 #184)
-#182 := (= #47 #181)
-#179 := (= #46 #176)
-#173 := (- #170)
-#177 := (= #173 #176)
-#178 := [rewrite]: #177
-#174 := (= #46 #173)
-#171 := (= #45 #170)
-#116 := (= #36 #115)
-#117 := [rewrite]: #116
-#113 := (= #35 #112)
-#114 := [rewrite]: #113
-#172 := [monotonicity #114 #117]: #171
-#175 := [monotonicity #172]: #174
-#180 := [trans #175 #178]: #179
-#183 := [monotonicity #180]: #182
-#186 := [monotonicity #183]: #185
-#189 := [monotonicity #186]: #188
-#192 := [monotonicity #189]: #191
-#195 := [quant-intro #192]: #194
-#212 := [trans #195 #210]: #211
-#169 := [asserted]: #51
-#213 := [mp #169 #212]: #208
-#216 := [mp~ #213 #220]: #208
-#264 := [mp #216 #263]: #261
-#295 := [mp #264 #294]: #292
-#782 := [mp #295 #781]: #777
-#735 := (not #777)
-#724 := (or #735 #451)
-#432 := (* -1::Int 2::Int)
-#764 := (mod #431 #432)
-#765 := (+ #13 #764)
-#766 := (= #765 0::Int)
-#444 := (<= 2::Int 0::Int)
-#447 := (or #443 #444)
-#426 := (not #447)
-#445 := (or #433 #444)
-#446 := (not #445)
-#761 := (or #446 #426)
-#767 := (if #761 #757 #766)
-#762 := (if #441 #442 #767)
-#440 := (= #13 f3)
-#356 := (= 2::Int 0::Int)
-#768 := (if #356 #440 #762)
-#725 := (or #735 #768)
-#721 := (iff #725 #724)
-#727 := (iff #724 #724)
-#728 := [rewrite]: #727
-#734 := (iff #768 #451)
-#454 := (if false #440 #451)
-#448 := (iff #454 #451)
-#730 := [rewrite]: #448
-#732 := (iff #768 #454)
-#452 := (iff #762 #451)
-#737 := (iff #767 #740)
-#462 := (iff #766 #461)
-#738 := (= #765 #362)
-#467 := (= #764 #466)
-#743 := (= #432 -2::Int)
-#465 := [rewrite]: #743
-#468 := [monotonicity #465]: #467
-#739 := [monotonicity #468]: #738
-#736 := [monotonicity #739]: #462
-#753 := (iff #761 #755)
-#394 := (iff #426 #754)
-#389 := (iff #447 #443)
-#748 := (or #443 false)
-#745 := (iff #748 #443)
-#751 := [rewrite]: #745
-#749 := (iff #447 #748)
-#423 := (iff #444 false)
-#759 := [rewrite]: #423
-#750 := [monotonicity #759]: #749
-#752 := [trans #750 #751]: #389
-#395 := [monotonicity #752]: #394
-#746 := (iff #446 #410)
-#408 := (iff #445 #433)
-#419 := (or #433 false)
-#744 := (iff #419 #433)
-#407 := [rewrite]: #744
-#760 := (iff #445 #419)
-#403 := [monotonicity #759]: #760
-#409 := [trans #403 #407]: #408
-#747 := [monotonicity #409]: #746
-#756 := [monotonicity #747 #395]: #753
-#741 := [monotonicity #756 #736]: #737
-#453 := [monotonicity #741]: #452
-#758 := (iff #356 false)
-#418 := [rewrite]: #758
-#733 := [monotonicity #418 #453]: #732
-#731 := [trans #733 #730]: #734
-#722 := [monotonicity #731]: #721
-#723 := [trans #722 #728]: #721
-#726 := [quant-inst #7 #10]: #725
-#729 := [mp #726 #723]: #724
-#656 := [unit-resolution #729 #782]: #451
-#594 := (not #441)
-#593 := (not #451)
-#665 := (or #593 #594)
-#699 := (not #442)
-#657 := (or #699 #99)
-#694 := [th-lemma arith triangle-eq]: #657
-#695 := [unit-resolution #694 #109]: #699
-#553 := (or #593 #594 #442)
-#701 := [def-axiom]: #553
-#655 := [unit-resolution #701 #695]: #665
-#666 := [unit-resolution #655 #656]: #594
-#603 := (or #593 #441 #740)
-#698 := [def-axiom]: #603
-#644 := [unit-resolution #698 #666 #656]: #740
-#720 := (not #740)
-#549 := (or #720 #716 #757)
-#551 := [def-axiom]: #549
-#647 := [unit-resolution #551 #644]: #645
-#648 := [unit-resolution #647 #654]: #716
-#571 := (or #755 #433)
-#572 := [def-axiom]: #571
-#649 := [unit-resolution #572 #648]: #433
-#714 := (or #755 #443)
-#715 := [def-axiom]: #714
-#650 := [unit-resolution #715 #648]: #443
-#651 := (or #441 #410 #754)
-#646 := [th-lemma arith triangle-eq]: #651
-#652 := [unit-resolution #646 #666]: #755
-[unit-resolution #652 #650 #649]: false
-unsat
-1432b33c6328a1ffc0a07c49f1ba0f71ab4e0de0 343 0
-#2 := false
-#23 := 0::Int
-decl f3 :: Int
-#7 := f3
-#428 := (<= f3 0::Int)
-#438 := (>= f3 0::Int)
-#749 := (not #438)
-#405 := (not #428)
-#750 := (or #405 #749)
-#712 := (not #750)
-#10 := 2::Int
-#758 := (mod f3 2::Int)
-#106 := -1::Int
-#415 := (* -1::Int #758)
-decl f4 :: (-> S2 Int Int)
-decl f5 :: (-> S3 Int S2)
-decl f6 :: S3
-#8 := f6
-#9 := (f5 f6 f3)
-#11 := (f4 #9 2::Int)
-#545 := (+ #11 #415)
-#752 := (= #545 0::Int)
-#703 := (not #752)
-#713 := (<= #545 0::Int)
-#659 := (not #713)
-#663 := (>= #758 2::Int)
-#665 := (not #663)
-#1 := true
-#68 := [true-axiom]: true
-#654 := (or false #665)
-#655 := [th-lemma arith]: #654
-#656 := [unit-resolution #655 #68]: #665
-#657 := [hypothesis]: #713
-#97 := (>= #11 2::Int)
-#14 := 3::Int
-#15 := (+ f3 3::Int)
-#12 := (+ #11 #11)
-#13 := (+ f3 #12)
-#16 := (< #13 #15)
-#17 := (not #16)
-#102 := (iff #17 #97)
-#77 := (+ 3::Int f3)
-#71 := (* 2::Int #11)
-#74 := (+ f3 #71)
-#80 := (< #74 #77)
-#83 := (not #80)
-#100 := (iff #83 #97)
-#90 := (>= #71 3::Int)
-#98 := (iff #90 #97)
-#99 := [rewrite]: #98
-#95 := (iff #83 #90)
-#88 := (not #90)
-#87 := (not #88)
-#93 := (iff #87 #90)
-#94 := [rewrite]: #93
-#91 := (iff #83 #87)
-#89 := (iff #80 #88)
-#86 := [rewrite]: #89
-#92 := [monotonicity #86]: #91
-#96 := [trans #92 #94]: #95
-#101 := [trans #96 #99]: #100
-#84 := (iff #17 #83)
-#81 := (iff #16 #80)
-#78 := (= #15 #77)
-#79 := [rewrite]: #78
-#75 := (= #13 #74)
-#72 := (= #12 #71)
-#73 := [rewrite]: #72
-#76 := [monotonicity #73]: #75
-#82 := [monotonicity #76 #79]: #81
-#85 := [monotonicity #82]: #84
-#103 := [trans #85 #101]: #102
-#70 := [asserted]: #17
-#104 := [mp #70 #103]: #97
-#658 := [th-lemma arith farkas -1 1 1 #104 #657 #656]: false
-#660 := [lemma #658]: #659
-#648 := (or #703 #713)
-#649 := [th-lemma arith triangle-eq]: #648
-#651 := [unit-resolution #649 #660]: #703
-#641 := (or #712 #752)
-#737 := -2::Int
-#426 := (* -1::Int f3)
-#461 := (mod #426 -2::Int)
-#357 := (+ #11 #461)
-#456 := (= #357 0::Int)
-#735 := (if #750 #752 #456)
-#437 := (= #11 0::Int)
-#436 := (= f3 0::Int)
-#446 := (if #436 #437 #735)
-#21 := (:var 0 Int)
-#19 := (:var 1 Int)
-#41 := (f5 f6 #19)
-#42 := (f4 #41 #21)
-#771 := (pattern #42)
-#110 := (* -1::Int #21)
-#107 := (* -1::Int #19)
-#165 := (mod #107 #110)
-#280 := (+ #42 #165)
-#281 := (= #280 0::Int)
-#43 := (mod #19 #21)
-#277 := (* -1::Int #43)
-#278 := (+ #42 #277)
-#279 := (= #278 0::Int)
-#132 := (<= #21 0::Int)
-#139 := (>= #19 0::Int)
-#224 := (or #139 #132)
-#225 := (not #224)
-#128 := (<= #19 0::Int)
-#222 := (or #128 #132)
-#223 := (not #222)
-#228 := (or #223 #225)
-#282 := (if #228 #279 #281)
-#276 := (= #42 0::Int)
-#24 := (= #19 0::Int)
-#283 := (if #24 #276 #282)
-#275 := (= #42 #19)
-#25 := (= #21 0::Int)
-#284 := (if #25 #275 #283)
-#772 := (forall (vars (?v0 Int) (?v1 Int)) (:pat #771) #284)
-#287 := (forall (vars (?v0 Int) (?v1 Int)) #284)
-#775 := (iff #287 #772)
-#773 := (iff #284 #284)
-#774 := [refl]: #773
-#776 := [quant-intro #774]: #775
-#171 := (* -1::Int #165)
-#244 := (if #228 #43 #171)
-#247 := (if #24 0::Int #244)
-#250 := (if #25 #19 #247)
-#253 := (= #42 #250)
-#256 := (forall (vars (?v0 Int) (?v1 Int)) #253)
-#288 := (iff #256 #287)
-#285 := (iff #253 #284)
-#286 := [rewrite]: #285
-#289 := [quant-intro #286]: #288
-#133 := (not #132)
-#140 := (not #139)
-#143 := (and #140 #133)
-#129 := (not #128)
-#136 := (and #129 #133)
-#146 := (or #136 #143)
-#191 := (if #146 #43 #171)
-#194 := (if #24 0::Int #191)
-#197 := (if #25 #19 #194)
-#200 := (= #42 #197)
-#203 := (forall (vars (?v0 Int) (?v1 Int)) #200)
-#257 := (iff #203 #256)
-#254 := (iff #200 #253)
-#251 := (= #197 #250)
-#248 := (= #194 #247)
-#245 := (= #191 #244)
-#229 := (iff #146 #228)
-#226 := (iff #143 #225)
-#227 := [rewrite]: #226
-#216 := (iff #136 #223)
-#217 := [rewrite]: #216
-#230 := [monotonicity #217 #227]: #229
-#246 := [monotonicity #230]: #245
-#249 := [monotonicity #246]: #248
-#252 := [monotonicity #249]: #251
-#255 := [monotonicity #252]: #254
-#258 := [quant-intro #255]: #257
-#214 := (~ #203 #203)
-#213 := (~ #200 #200)
-#210 := [refl]: #213
-#215 := [nnf-pos #210]: #214
-#35 := (- #21)
-#34 := (- #19)
-#44 := (mod #34 #35)
-#45 := (- #44)
-#28 := (< 0::Int #21)
-#30 := (< #19 0::Int)
-#31 := (and #30 #28)
-#27 := (< 0::Int #19)
-#29 := (and #27 #28)
-#32 := (or #29 #31)
-#46 := (if #32 #43 #45)
-#47 := (if #24 0::Int #46)
-#48 := (if #25 #19 #47)
-#49 := (= #42 #48)
-#50 := (forall (vars (?v0 Int) (?v1 Int)) #49)
-#206 := (iff #50 #203)
-#176 := (if #32 #43 #171)
-#179 := (if #24 0::Int #176)
-#182 := (if #25 #19 #179)
-#185 := (= #42 #182)
-#188 := (forall (vars (?v0 Int) (?v1 Int)) #185)
-#204 := (iff #188 #203)
-#201 := (iff #185 #200)
-#198 := (= #182 #197)
-#195 := (= #179 #194)
-#192 := (= #176 #191)
-#147 := (iff #32 #146)
-#144 := (iff #31 #143)
-#134 := (iff #28 #133)
-#135 := [rewrite]: #134
-#141 := (iff #30 #140)
-#142 := [rewrite]: #141
-#145 := [monotonicity #142 #135]: #144
-#137 := (iff #29 #136)
-#130 := (iff #27 #129)
-#131 := [rewrite]: #130
-#138 := [monotonicity #131 #135]: #137
-#148 := [monotonicity #138 #145]: #147
-#193 := [monotonicity #148]: #192
-#196 := [monotonicity #193]: #195
-#199 := [monotonicity #196]: #198
-#202 := [monotonicity #199]: #201
-#205 := [quant-intro #202]: #204
-#189 := (iff #50 #188)
-#186 := (iff #49 #185)
-#183 := (= #48 #182)
-#180 := (= #47 #179)
-#177 := (= #46 #176)
-#174 := (= #45 #171)
-#168 := (- #165)
-#172 := (= #168 #171)
-#173 := [rewrite]: #172
-#169 := (= #45 #168)
-#166 := (= #44 #165)
-#111 := (= #35 #110)
-#112 := [rewrite]: #111
-#108 := (= #34 #107)
-#109 := [rewrite]: #108
-#167 := [monotonicity #109 #112]: #166
-#170 := [monotonicity #167]: #169
-#175 := [trans #170 #173]: #174
-#178 := [monotonicity #175]: #177
-#181 := [monotonicity #178]: #180
-#184 := [monotonicity #181]: #183
-#187 := [monotonicity #184]: #186
-#190 := [quant-intro #187]: #189
-#207 := [trans #190 #205]: #206
-#164 := [asserted]: #50
-#208 := [mp #164 #207]: #203
-#211 := [mp~ #208 #215]: #203
-#259 := [mp #211 #258]: #256
-#290 := [mp #259 #289]: #287
-#777 := [mp #290 #776]: #772
-#730 := (not #772)
-#719 := (or #730 #446)
-#427 := (* -1::Int 2::Int)
-#759 := (mod #426 #427)
-#760 := (+ #11 #759)
-#761 := (= #760 0::Int)
-#439 := (<= 2::Int 0::Int)
-#442 := (or #438 #439)
-#421 := (not #442)
-#440 := (or #428 #439)
-#441 := (not #440)
-#756 := (or #441 #421)
-#762 := (if #756 #752 #761)
-#757 := (if #436 #437 #762)
-#435 := (= #11 f3)
-#351 := (= 2::Int 0::Int)
-#763 := (if #351 #435 #757)
-#720 := (or #730 #763)
-#716 := (iff #720 #719)
-#722 := (iff #719 #719)
-#723 := [rewrite]: #722
-#729 := (iff #763 #446)
-#449 := (if false #435 #446)
-#443 := (iff #449 #446)
-#725 := [rewrite]: #443
-#727 := (iff #763 #449)
-#447 := (iff #757 #446)
-#732 := (iff #762 #735)
-#457 := (iff #761 #456)
-#733 := (= #760 #357)
-#462 := (= #759 #461)
-#738 := (= #427 -2::Int)
-#460 := [rewrite]: #738
-#463 := [monotonicity #460]: #462
-#734 := [monotonicity #463]: #733
-#731 := [monotonicity #734]: #457
-#748 := (iff #756 #750)
-#389 := (iff #421 #749)
-#384 := (iff #442 #438)
-#743 := (or #438 false)
-#740 := (iff #743 #438)
-#746 := [rewrite]: #740
-#744 := (iff #442 #743)
-#418 := (iff #439 false)
-#754 := [rewrite]: #418
-#745 := [monotonicity #754]: #744
-#747 := [trans #745 #746]: #384
-#390 := [monotonicity #747]: #389
-#741 := (iff #441 #405)
-#403 := (iff #440 #428)
-#414 := (or #428 false)
-#739 := (iff #414 #428)
-#402 := [rewrite]: #739
-#755 := (iff #440 #414)
-#398 := [monotonicity #754]: #755
-#404 := [trans #398 #402]: #403
-#742 := [monotonicity #404]: #741
-#751 := [monotonicity #742 #390]: #748
-#736 := [monotonicity #751 #731]: #732
-#448 := [monotonicity #736]: #447
-#753 := (iff #351 false)
-#413 := [rewrite]: #753
-#728 := [monotonicity #413 #448]: #727
-#726 := [trans #728 #725]: #729
-#717 := [monotonicity #726]: #716
-#718 := [trans #717 #723]: #716
-#721 := [quant-inst #7 #10]: #720
-#724 := [mp #721 #718]: #719
-#652 := [unit-resolution #724 #777]: #446
-#548 := (not #436)
-#589 := (not #446)
-#643 := (or #589 #548)
-#697 := (not #437)
-#565 := (<= #11 0::Int)
-#653 := (not #565)
-#690 := (not #97)
-#691 := (or #653 #690)
-#650 := [th-lemma arith farkas 1 1]: #691
-#661 := [unit-resolution #650 #104]: #653
-#639 := (or #697 #565)
-#640 := [th-lemma arith triangle-eq]: #639
-#642 := [unit-resolution #640 #661]: #697
-#696 := (or #589 #548 #437)
-#598 := [def-axiom]: #696
-#644 := [unit-resolution #598 #642]: #643
-#645 := [unit-resolution #644 #652]: #548
-#693 := (or #589 #436 #735)
-#694 := [def-axiom]: #693
-#646 := [unit-resolution #694 #645 #652]: #735
-#544 := (not #735)
-#546 := (or #544 #712 #752)
-#547 := [def-axiom]: #546
-#647 := [unit-resolution #547 #646]: #641
-#633 := [unit-resolution #647 #651]: #712
-#567 := (or #750 #428)
-#709 := [def-axiom]: #567
-#629 := [unit-resolution #709 #633]: #428
-#710 := (or #750 #438)
-#711 := [def-axiom]: #710
-#630 := [unit-resolution #711 #633]: #438
-#631 := (or #436 #405 #749)
-#634 := [th-lemma arith triangle-eq]: #631
-#635 := [unit-resolution #634 #645]: #750
-[unit-resolution #635 #630 #629]: false
-unsat
-6c2df05479a46eb0dc1434ea9ed59f4fae72c26e 101 0
-#2 := false
-#8 := 0::Real
-decl f3 :: Real
-#7 := f3
-#9 := (= f3 0::Real)
-#10 := (not #9)
-#45 := [asserted]: #10
-#100 := (<= f3 0::Real)
-#20 := 2::Real
-#47 := (* 2::Real f3)
-#102 := (<= #47 0::Real)
-#95 := (= #47 0::Real)
-#19 := 4::Real
-#14 := (- f3)
-#13 := (< f3 0::Real)
-#15 := (if #13 #14 f3)
-#12 := 1::Real
-#16 := (< 1::Real #15)
-#17 := (not #16)
-#18 := (or #16 #17)
-#21 := (if #18 4::Real 2::Real)
-#22 := (* #21 f3)
-#11 := (+ f3 f3)
-#23 := (= #11 #22)
-#24 := (not #23)
-#25 := (not #24)
-#96 := (iff #25 #95)
-#77 := (* 4::Real f3)
-#80 := (= #47 #77)
-#93 := (iff #80 #95)
-#94 := [rewrite]: #93
-#91 := (iff #25 #80)
-#83 := (not #80)
-#86 := (not #83)
-#89 := (iff #86 #80)
-#90 := [rewrite]: #89
-#87 := (iff #25 #86)
-#84 := (iff #24 #83)
-#81 := (iff #23 #80)
-#78 := (= #22 #77)
-#75 := (= #21 4::Real)
-#1 := true
-#70 := (if true 4::Real 2::Real)
-#73 := (= #70 4::Real)
-#74 := [rewrite]: #73
-#71 := (= #21 #70)
-#68 := (iff #18 true)
-#50 := -1::Real
-#51 := (* -1::Real f3)
-#54 := (if #13 #51 f3)
-#57 := (< 1::Real #54)
-#60 := (not #57)
-#63 := (or #57 #60)
-#66 := (iff #63 true)
-#67 := [rewrite]: #66
-#64 := (iff #18 #63)
-#61 := (iff #17 #60)
-#58 := (iff #16 #57)
-#55 := (= #15 #54)
-#52 := (= #14 #51)
-#53 := [rewrite]: #52
-#56 := [monotonicity #53]: #55
-#59 := [monotonicity #56]: #58
-#62 := [monotonicity #59]: #61
-#65 := [monotonicity #59 #62]: #64
-#69 := [trans #65 #67]: #68
-#72 := [monotonicity #69]: #71
-#76 := [trans #72 #74]: #75
-#79 := [monotonicity #76]: #78
-#48 := (= #11 #47)
-#49 := [rewrite]: #48
-#82 := [monotonicity #49 #79]: #81
-#85 := [monotonicity #82]: #84
-#88 := [monotonicity #85]: #87
-#92 := [trans #88 #90]: #91
-#97 := [trans #92 #94]: #96
-#46 := [asserted]: #25
-#98 := [mp #46 #97]: #95
-#104 := (not #95)
-#105 := (or #104 #102)
-#106 := [th-lemma arith triangle-eq]: #105
-#107 := [unit-resolution #106 #98]: #102
-#108 := (not #102)
-#109 := (or #100 #108)
-#110 := [th-lemma arith assign-bounds 1]: #109
-#111 := [unit-resolution #110 #107]: #100
-#101 := (>= f3 0::Real)
-#103 := (>= #47 0::Real)
-#112 := (or #104 #103)
-#113 := [th-lemma arith triangle-eq]: #112
-#114 := [unit-resolution #113 #98]: #103
-#115 := (not #103)
-#116 := (or #101 #115)
-#117 := [th-lemma arith assign-bounds 1]: #116
-#118 := [unit-resolution #117 #114]: #101
-#120 := (not #101)
-#119 := (not #100)
-#121 := (or #9 #119 #120)
-#122 := [th-lemma arith triangle-eq]: #121
-[unit-resolution #122 #118 #111 #45]: false
-unsat
-0eb09039097aac0255a0090f04ca5df53ea2d10a 24 0
-#2 := false
-#7 := (exists (vars (?v0 Int)) false)
-#8 := (not #7)
-#9 := (not #8)
-#45 := (iff #9 false)
-#1 := true
-#40 := (not true)
-#43 := (iff #40 false)
-#44 := [rewrite]: #43
-#41 := (iff #9 #40)
-#38 := (iff #8 true)
-#33 := (not false)
-#36 := (iff #33 true)
-#37 := [rewrite]: #36
-#34 := (iff #8 #33)
-#31 := (iff #7 false)
-#32 := [elim-unused]: #31
-#35 := [monotonicity #32]: #34
-#39 := [trans #35 #37]: #38
-#42 := [monotonicity #39]: #41
-#46 := [trans #42 #44]: #45
-#30 := [asserted]: #9
-[mp #30 #46]: false
-unsat
-9f8072a1ad3de2c920c120b81de67bceefc50c87 916 0
-#2 := false
-#22 := 1::Int
-decl f3 :: (-> S2 Int Int)
-#12 := 2::Int
-decl f4 :: (-> S3 Int S2)
-decl f7 :: Int
-#9 := f7
-decl f5 :: S3
-#7 := f5
-#24 := (f4 f5 f7)
-#25 := (f3 #24 2::Int)
-#1265 := (<= #25 1::Int)
-#14 := 0::Int
-#551 := (mod f7 2::Int)
-#84 := -1::Int
-#521 := (* -1::Int #551)
-#522 := (+ #25 #521)
-#920 := (<= #522 0::Int)
-#523 := (= #522 0::Int)
-decl f6 :: Int
-#8 := f6
-#10 := (+ f6 f7)
-#431 := (>= #10 0::Int)
-#426 := (= #10 0::Int)
-#746 := (mod #10 2::Int)
-#748 := (* -1::Int #746)
-#11 := (f4 f5 #10)
-#13 := (f3 #11 2::Int)
-#405 := (+ #13 #748)
-#535 := (= #405 0::Int)
-#686 := (not #535)
-#691 := (<= #405 0::Int)
-#1269 := [hypothesis]: #535
-#1270 := (or #686 #691)
-#1271 := [th-lemma arith triangle-eq]: #1270
-#1272 := [unit-resolution #1271 #1269]: #691
-#693 := (>= #405 0::Int)
-#1273 := (or #686 #693)
-#1626 := [th-lemma arith triangle-eq]: #1273
-#1627 := [unit-resolution #1626 #1269]: #693
-#1371 := (not #691)
-#1437 := (not #693)
-#1647 := (or #1437 #1371)
-#1274 := (div f7 2::Int)
-#447 := -2::Int
-#1287 := (* -2::Int #1274)
-#1288 := (+ #521 #1287)
-#1289 := (+ f7 #1288)
-#1286 := (= #1289 0::Int)
-#1349 := (not #1286)
-#1474 := [hypothesis]: #1349
-#1 := true
-#78 := [true-axiom]: true
-#1346 := (or false #1286)
-#1347 := [th-lemma arith]: #1346
-#1475 := [unit-resolution #1347 #78 #1474]: false
-#1476 := [lemma #1475]: #1286
-#1472 := (or #1349 #1437 #1371)
-#1296 := (>= #551 0::Int)
-#1398 := (or false #1296)
-#1399 := [th-lemma arith]: #1398
-#1400 := [unit-resolution #1399 #78]: #1296
-#1422 := (>= #1289 0::Int)
-#1444 := [hypothesis]: #1286
-#1445 := (or #1349 #1422)
-#1446 := [th-lemma arith triangle-eq]: #1445
-#1447 := [unit-resolution #1446 #1444]: #1422
-#19 := 3::Int
-#17 := 4::Int
-#16 := (f4 f5 f6)
-#18 := (f3 #16 4::Int)
-#539 := (>= #18 3::Int)
-#20 := (= #18 3::Int)
-#81 := [asserted]: #20
-#989 := (not #20)
-#1010 := (or #989 #539)
-#1011 := [th-lemma arith triangle-eq]: #1010
-#1012 := [unit-resolution #1011 #81]: #539
-#831 := (div f6 4::Int)
-#634 := -4::Int
-#847 := (* -4::Int #831)
-#672 := (mod f6 4::Int)
-#673 := (* -1::Int #672)
-#848 := (+ #673 #847)
-#849 := (+ f6 #848)
-#855 := (>= #849 0::Int)
-#846 := (= #849 0::Int)
-#993 := (or false #846)
-#994 := [th-lemma arith]: #993
-#995 := [unit-resolution #994 #78]: #846
-#996 := (not #846)
-#1013 := (or #996 #855)
-#1014 := [th-lemma arith triangle-eq]: #1013
-#1015 := [unit-resolution #1014 #995]: #855
-#531 := (>= #13 0::Int)
-#15 := (= #13 0::Int)
-#80 := [asserted]: #15
-#593 := (not #15)
-#1428 := (or #593 #531)
-#1429 := [th-lemma arith triangle-eq]: #1428
-#1430 := [unit-resolution #1429 #80]: #531
-#777 := (div #10 2::Int)
-#794 := (* -2::Int #777)
-#795 := (+ #748 #794)
-#796 := (+ f7 #795)
-#797 := (+ f6 #796)
-#1268 := (>= #797 0::Int)
-#792 := (= #797 0::Int)
-#1355 := (or false #792)
-#1356 := [th-lemma arith]: #1355
-#1357 := [unit-resolution #1356 #78]: #792
-#1358 := (not #792)
-#1431 := (or #1358 #1268)
-#1432 := [th-lemma arith triangle-eq]: #1431
-#1433 := [unit-resolution #1432 #1357]: #1268
-#1434 := [hypothesis]: #693
-#674 := (+ #18 #673)
-#571 := (>= #674 0::Int)
-#668 := (= #674 0::Int)
-#453 := (* -1::Int f6)
-#631 := (mod #453 -4::Int)
-#619 := (+ #18 #631)
-#624 := (= #619 0::Int)
-#681 := (>= f6 0::Int)
-#640 := (not #681)
-#667 := (<= f6 0::Int)
-#641 := (not #667)
-#630 := (or #641 #640)
-#627 := (if #630 #668 #624)
-#678 := (= f6 0::Int)
-#561 := (not #678)
-#670 := (= #18 0::Int)
-#566 := (not #670)
-#389 := (= 3::Int 0::Int)
-#396 := (iff #389 false)
-#397 := [rewrite]: #396
-#407 := [hypothesis]: #670
-#409 := (= 3::Int #18)
-#410 := [symm #81]: #409
-#391 := [trans #410 #407]: #389
-#398 := [mp #391 #397]: false
-#399 := [lemma #398]: #566
-#1204 := (or #561 #670)
-#601 := (if #678 #670 #627)
-#32 := (:var 0 Int)
-#30 := (:var 1 Int)
-#51 := (f4 f5 #30)
-#52 := (f3 #51 #32)
-#761 := (pattern #52)
-#88 := (* -1::Int #32)
-#85 := (* -1::Int #30)
-#143 := (mod #85 #88)
-#272 := (+ #52 #143)
-#273 := (= #272 0::Int)
-#53 := (mod #30 #32)
-#269 := (* -1::Int #53)
-#270 := (+ #52 #269)
-#271 := (= #270 0::Int)
-#110 := (<= #32 0::Int)
-#117 := (>= #30 0::Int)
-#216 := (or #117 #110)
-#217 := (not #216)
-#106 := (<= #30 0::Int)
-#212 := (or #106 #110)
-#213 := (not #212)
-#220 := (or #213 #217)
-#274 := (if #220 #271 #273)
-#268 := (= #52 0::Int)
-#34 := (= #30 0::Int)
-#275 := (if #34 #268 #274)
-#267 := (= #52 #30)
-#35 := (= #32 0::Int)
-#276 := (if #35 #267 #275)
-#762 := (forall (vars (?v0 Int) (?v1 Int)) (:pat #761) #276)
-#279 := (forall (vars (?v0 Int) (?v1 Int)) #276)
-#765 := (iff #279 #762)
-#763 := (iff #276 #276)
-#764 := [refl]: #763
-#766 := [quant-intro #764]: #765
-#149 := (* -1::Int #143)
-#236 := (if #220 #53 #149)
-#239 := (if #34 0::Int #236)
-#242 := (if #35 #30 #239)
-#245 := (= #52 #242)
-#248 := (forall (vars (?v0 Int) (?v1 Int)) #245)
-#280 := (iff #248 #279)
-#277 := (iff #245 #276)
-#278 := [rewrite]: #277
-#281 := [quant-intro #278]: #280
-#111 := (not #110)
-#118 := (not #117)
-#121 := (and #118 #111)
-#107 := (not #106)
-#114 := (and #107 #111)
-#124 := (or #114 #121)
-#169 := (if #124 #53 #149)
-#172 := (if #34 0::Int #169)
-#175 := (if #35 #30 #172)
-#178 := (= #52 #175)
-#181 := (forall (vars (?v0 Int) (?v1 Int)) #178)
-#249 := (iff #181 #248)
-#246 := (iff #178 #245)
-#243 := (= #175 #242)
-#240 := (= #172 #239)
-#237 := (= #169 #236)
-#221 := (iff #124 #220)
-#218 := (iff #121 #217)
-#219 := [rewrite]: #218
-#214 := (iff #114 #213)
-#215 := [rewrite]: #214
-#222 := [monotonicity #215 #219]: #221
-#238 := [monotonicity #222]: #237
-#241 := [monotonicity #238]: #240
-#244 := [monotonicity #241]: #243
-#247 := [monotonicity #244]: #246
-#250 := [quant-intro #247]: #249
-#191 := (~ #181 #181)
-#193 := (~ #178 #178)
-#190 := [refl]: #193
-#194 := [nnf-pos #190]: #191
-#45 := (- #32)
-#44 := (- #30)
-#54 := (mod #44 #45)
-#55 := (- #54)
-#38 := (< 0::Int #32)
-#40 := (< #30 0::Int)
-#41 := (and #40 #38)
-#37 := (< 0::Int #30)
-#39 := (and #37 #38)
-#42 := (or #39 #41)
-#56 := (if #42 #53 #55)
-#57 := (if #34 0::Int #56)
-#58 := (if #35 #30 #57)
-#59 := (= #52 #58)
-#60 := (forall (vars (?v0 Int) (?v1 Int)) #59)
-#184 := (iff #60 #181)
-#154 := (if #42 #53 #149)
-#157 := (if #34 0::Int #154)
-#160 := (if #35 #30 #157)
-#163 := (= #52 #160)
-#166 := (forall (vars (?v0 Int) (?v1 Int)) #163)
-#182 := (iff #166 #181)
-#179 := (iff #163 #178)
-#176 := (= #160 #175)
-#173 := (= #157 #172)
-#170 := (= #154 #169)
-#125 := (iff #42 #124)
-#122 := (iff #41 #121)
-#112 := (iff #38 #111)
-#113 := [rewrite]: #112
-#119 := (iff #40 #118)
-#120 := [rewrite]: #119
-#123 := [monotonicity #120 #113]: #122
-#115 := (iff #39 #114)
-#108 := (iff #37 #107)
-#109 := [rewrite]: #108
-#116 := [monotonicity #109 #113]: #115
-#126 := [monotonicity #116 #123]: #125
-#171 := [monotonicity #126]: #170
-#174 := [monotonicity #171]: #173
-#177 := [monotonicity #174]: #176
-#180 := [monotonicity #177]: #179
-#183 := [quant-intro #180]: #182
-#167 := (iff #60 #166)
-#164 := (iff #59 #163)
-#161 := (= #58 #160)
-#158 := (= #57 #157)
-#155 := (= #56 #154)
-#152 := (= #55 #149)
-#146 := (- #143)
-#150 := (= #146 #149)
-#151 := [rewrite]: #150
-#147 := (= #55 #146)
-#144 := (= #54 #143)
-#89 := (= #45 #88)
-#90 := [rewrite]: #89
-#86 := (= #44 #85)
-#87 := [rewrite]: #86
-#145 := [monotonicity #87 #90]: #144
-#148 := [monotonicity #145]: #147
-#153 := [trans #148 #151]: #152
-#156 := [monotonicity #153]: #155
-#159 := [monotonicity #156]: #158
-#162 := [monotonicity #159]: #161
-#165 := [monotonicity #162]: #164
-#168 := [quant-intro #165]: #167
-#185 := [trans #168 #183]: #184
-#142 := [asserted]: #60
-#186 := [mp #142 #185]: #181
-#195 := [mp~ #186 #194]: #181
-#251 := [mp #195 #250]: #248
-#282 := [mp #251 #281]: #279
-#767 := [mp #282 #766]: #762
-#555 := (not #762)
-#612 := (or #555 #601)
-#675 := (* -1::Int 4::Int)
-#659 := (mod #453 #675)
-#660 := (+ #18 #659)
-#662 := (= #660 0::Int)
-#669 := (<= 4::Int 0::Int)
-#677 := (or #681 #669)
-#682 := (not #677)
-#679 := (or #667 #669)
-#680 := (not #679)
-#671 := (or #680 #682)
-#663 := (if #671 #668 #662)
-#664 := (if #678 #670 #663)
-#676 := (= #18 f6)
-#689 := (= 4::Int 0::Int)
-#665 := (if #689 #676 #664)
-#615 := (or #555 #665)
-#617 := (iff #615 #612)
-#618 := (iff #612 #612)
-#598 := [rewrite]: #618
-#610 := (iff #665 #601)
-#496 := (if false #676 #601)
-#609 := (iff #496 #601)
-#614 := [rewrite]: #609
-#607 := (iff #665 #496)
-#602 := (iff #664 #601)
-#622 := (iff #663 #627)
-#625 := (iff #662 #624)
-#620 := (= #660 #619)
-#637 := (= #659 #631)
-#635 := (= #675 -4::Int)
-#636 := [rewrite]: #635
-#623 := [monotonicity #636]: #637
-#621 := [monotonicity #623]: #620
-#626 := [monotonicity #621]: #625
-#632 := (iff #671 #630)
-#651 := (iff #682 #640)
-#649 := (iff #677 #681)
-#644 := (or #681 false)
-#647 := (iff #644 #681)
-#648 := [rewrite]: #647
-#645 := (iff #677 #644)
-#652 := (iff #669 false)
-#653 := [rewrite]: #652
-#646 := [monotonicity #653]: #645
-#650 := [trans #646 #648]: #649
-#629 := [monotonicity #650]: #651
-#642 := (iff #680 #641)
-#638 := (iff #679 #667)
-#655 := (or #667 false)
-#654 := (iff #655 #667)
-#658 := [rewrite]: #654
-#656 := (iff #679 #655)
-#657 := [monotonicity #653]: #656
-#639 := [trans #657 #658]: #638
-#643 := [monotonicity #639]: #642
-#633 := [monotonicity #643 #629]: #632
-#628 := [monotonicity #633 #626]: #622
-#603 := [monotonicity #628]: #602
-#661 := (iff #689 false)
-#666 := [rewrite]: #661
-#608 := [monotonicity #666 #603]: #607
-#611 := [trans #608 #614]: #610
-#613 := [monotonicity #611]: #617
-#544 := [trans #613 #598]: #617
-#616 := [quant-inst #8 #17]: #615
-#599 := [mp #616 #544]: #612
-#1203 := [unit-resolution #599 #767]: #601
-#560 := (not #601)
-#562 := (or #560 #561 #670)
-#563 := [def-axiom]: #562
-#1205 := [unit-resolution #563 #1203]: #1204
-#1206 := [unit-resolution #1205 #399]: #561
-#1207 := (or #678 #627)
-#564 := (or #560 #678 #627)
-#565 := [def-axiom]: #564
-#1208 := [unit-resolution #565 #1203]: #1207
-#1209 := [unit-resolution #1208 #1206]: #627
-#606 := (not #630)
-#826 := [hypothesis]: #606
-#580 := (or #630 #667)
-#604 := [def-axiom]: #580
-#827 := [unit-resolution #604 #826]: #667
-#605 := (or #630 #681)
-#600 := [def-axiom]: #605
-#828 := [unit-resolution #600 #826]: #681
-#829 := (or #678 #641 #640)
-#830 := [th-lemma arith triangle-eq]: #829
-#879 := [unit-resolution #830 #828 #827 #1206]: false
-#880 := [lemma #879]: #630
-#582 := (not #627)
-#584 := (or #582 #606 #668)
-#585 := [def-axiom]: #584
-#1353 := [unit-resolution #585 #880 #1209]: #668
-#576 := (not #668)
-#1216 := (or #576 #571)
-#1217 := [th-lemma arith triangle-eq]: #1216
-#1435 := [unit-resolution #1217 #1353]: #571
-#1330 := (* -1::Int #1274)
-#1051 := (* -2::Int #831)
-#1331 := (+ #1051 #1330)
-#940 := (* -1::Int #777)
-#1332 := (+ #940 #1331)
-#1333 := (+ #748 #1332)
-#1334 := (+ #673 #1333)
-#1335 := (+ #18 #1334)
-#1336 := (+ #13 #1335)
-#1337 := (+ f7 #1336)
-#1338 := (+ f6 #1337)
-#1339 := (>= #1338 2::Int)
-#1369 := (not #1339)
-#921 := (>= #522 0::Int)
-#1362 := [hypothesis]: #691
-#1438 := (or #523 #1437 #1371)
-#532 := (<= #18 3::Int)
-#990 := (or #989 #532)
-#991 := [th-lemma arith triangle-eq]: #990
-#992 := [unit-resolution #991 #81]: #532
-#854 := (<= #849 0::Int)
-#997 := (or #996 #854)
-#998 := [th-lemma arith triangle-eq]: #997
-#999 := [unit-resolution #998 #995]: #854
-#545 := (<= f7 0::Int)
-#542 := (= f7 0::Int)
-#1190 := (not #523)
-#1308 := [hypothesis]: #1190
-#1420 := (or #542 #523)
-#347 := (* -1::Int f7)
-#507 := (mod #347 -2::Int)
-#504 := (+ #25 #507)
-#493 := (= #504 0::Int)
-#548 := (>= f7 0::Int)
-#497 := (not #548)
-#517 := (not #545)
-#502 := (or #517 #497)
-#476 := (if #502 #523 #493)
-#1255 := (not #542)
-#1412 := [hypothesis]: #1255
-#1406 := (or #542 #476)
-#543 := (= #25 0::Int)
-#480 := (if #542 #543 #476)
-#366 := (or #555 #480)
-#416 := (* -1::Int 2::Int)
-#524 := (mod #347 #416)
-#526 := (+ #25 #524)
-#527 := (= #526 0::Int)
-#418 := (<= 2::Int 0::Int)
-#549 := (or #548 #418)
-#550 := (not #549)
-#546 := (or #545 #418)
-#547 := (not #546)
-#533 := (or #547 #550)
-#528 := (if #533 #523 #527)
-#371 := (if #542 #543 #528)
-#541 := (= #25 f7)
-#341 := (= 2::Int 0::Int)
-#529 := (if #341 #541 #371)
-#351 := (or #555 #529)
-#352 := (iff #351 #366)
-#355 := (iff #366 #366)
-#342 := [rewrite]: #355
-#488 := (iff #529 #480)
-#483 := (if false #541 #480)
-#486 := (iff #483 #480)
-#487 := [rewrite]: #486
-#484 := (iff #529 #483)
-#481 := (iff #371 #480)
-#478 := (iff #528 #476)
-#491 := (iff #527 #493)
-#490 := (= #526 #504)
-#500 := (= #524 #507)
-#721 := (= #416 -2::Int)
-#725 := [rewrite]: #721
-#503 := [monotonicity #725]: #500
-#492 := [monotonicity #503]: #490
-#494 := [monotonicity #492]: #491
-#506 := (iff #533 #502)
-#498 := (iff #550 #497)
-#505 := (iff #549 #548)
-#511 := (or #548 false)
-#510 := (iff #511 #548)
-#515 := [rewrite]: #510
-#513 := (iff #549 #511)
-#404 := (iff #418 false)
-#392 := [rewrite]: #404
-#514 := [monotonicity #392]: #513
-#495 := [trans #514 #515]: #505
-#501 := [monotonicity #495]: #498
-#520 := (iff #547 #517)
-#518 := (iff #546 #545)
-#525 := (or #545 false)
-#512 := (iff #525 #545)
-#516 := [rewrite]: #512
-#530 := (iff #546 #525)
-#509 := [monotonicity #392]: #530
-#519 := [trans #509 #516]: #518
-#508 := [monotonicity #519]: #520
-#499 := [monotonicity #508 #501]: #506
-#479 := [monotonicity #499 #494]: #478
-#482 := [monotonicity #479]: #481
-#753 := (iff #341 false)
-#743 := [rewrite]: #753
-#485 := [monotonicity #743 #482]: #484
-#477 := [trans #485 #487]: #488
-#350 := [monotonicity #477]: #352
-#344 := [trans #350 #342]: #352
-#349 := [quant-inst #9 #12]: #351
-#345 := [mp #349 #344]: #366
-#1313 := [unit-resolution #345 #767]: #480
-#1254 := (not #480)
-#1258 := (or #1254 #542 #476)
-#1259 := [def-axiom]: #1258
-#1407 := [unit-resolution #1259 #1313]: #1406
-#1413 := [unit-resolution #1407 #1412]: #476
-#1410 := (or #548 #523)
-#1309 := [hypothesis]: #497
-#881 := (or #502 #548)
-#882 := [def-axiom]: #881
-#1310 := [unit-resolution #882 #1309]: #502
-#1311 := (or #1255 #548)
-#1312 := [th-lemma arith triangle-eq]: #1311
-#1295 := [unit-resolution #1312 #1309]: #1255
-#1408 := [unit-resolution #1407 #1295]: #476
-#883 := (not #502)
-#802 := (not #476)
-#1102 := (or #802 #883 #523)
-#1103 := [def-axiom]: #1102
-#1409 := [unit-resolution #1103 #1408 #1310 #1308]: false
-#1411 := [lemma #1409]: #1410
-#1414 := [unit-resolution #1411 #1308]: #548
-#1415 := (or #542 #517 #497)
-#1416 := [th-lemma arith triangle-eq]: #1415
-#1417 := [unit-resolution #1416 #1412 #1414]: #517
-#370 := (or #502 #545)
-#372 := [def-axiom]: #370
-#1418 := [unit-resolution #372 #1417]: #502
-#1419 := [unit-resolution #1103 #1418 #1413 #1308]: false
-#1421 := [lemma #1419]: #1420
-#1424 := [unit-resolution #1421 #1308]: #542
-#1425 := (or #1255 #545)
-#1426 := [th-lemma arith triangle-eq]: #1425
-#1427 := [unit-resolution #1426 #1424]: #545
-#570 := (<= #13 0::Int)
-#1364 := (or #593 #570)
-#1365 := [th-lemma arith triangle-eq]: #1364
-#1366 := [unit-resolution #1365 #80]: #570
-#1267 := (<= #797 0::Int)
-#1359 := (or #1358 #1267)
-#1360 := [th-lemma arith triangle-eq]: #1359
-#1361 := [unit-resolution #1360 #1357]: #1267
-#540 := (<= #674 0::Int)
-#1212 := (or #576 #540)
-#1213 := [th-lemma arith triangle-eq]: #1212
-#1354 := [unit-resolution #1213 #1353]: #540
-#1436 := [th-lemma arith gcd-test -1/2 -1/2 -1/2 -1/2 -1/2 -1/2 -1/2 -1/2 -1/2 -1/2 -1/2 -1/2 -1/2 -1/2 #1435 #1354 #1434 #1362 #1433 #1361 #1430 #1366 #1414 #1427 #1015 #999 #1012 #992]: false
-#1439 := [lemma #1436]: #1438
-#1448 := [unit-resolution #1439 #1434 #1362]: #523
-#1449 := (or #1190 #921)
-#1450 := [th-lemma arith triangle-eq]: #1449
-#1451 := [unit-resolution #1450 #1448]: #921
-#1266 := (>= #25 1::Int)
-#1344 := (not #1266)
-#1452 := (or #1190 #920)
-#1453 := [th-lemma arith triangle-eq]: #1452
-#1454 := [unit-resolution #1453 #1448]: #920
-#1302 := (>= #551 2::Int)
-#1303 := (not #1302)
-#1455 := (or false #1303)
-#1456 := [th-lemma arith]: #1455
-#1457 := [unit-resolution #1456 #78]: #1303
-#1458 := (not #920)
-#1459 := (or #1265 #1302 #1458)
-#1460 := [th-lemma arith assign-bounds 1 1]: #1459
-#1461 := [unit-resolution #1460 #1457 #1454]: #1265
-#1464 := (not #1265)
-#1467 := (or #1464 #1344)
-#26 := (= #25 1::Int)
-#189 := (not #26)
-#21 := (f3 #16 2::Int)
-#23 := (= #21 1::Int)
-#1248 := (or #606 #23)
-#884 := (div f6 2::Int)
-#1118 := (* -1::Int #884)
-#1119 := (+ #1051 #1118)
-#1120 := (+ #673 #1119)
-#448 := (mod f6 2::Int)
-#457 := (* -1::Int #448)
-#1121 := (+ #457 #1120)
-#1122 := (+ #18 #1121)
-#1123 := (+ f6 #1122)
-#1124 := (>= #1123 2::Int)
-#1134 := (not #1124)
-#1210 := [hypothesis]: #630
-#1211 := [unit-resolution #585 #1210 #1209]: #668
-#1214 := [unit-resolution #1213 #1211]: #540
-#1005 := (not #540)
-#1135 := (or #1134 #1005)
-#906 := (>= #448 0::Int)
-#1129 := (or false #906)
-#1130 := [th-lemma arith]: #1129
-#1131 := [unit-resolution #1130 #78]: #906
-#1000 := [hypothesis]: #540
-#897 := (* -2::Int #884)
-#898 := (+ #457 #897)
-#899 := (+ f6 #898)
-#904 := (<= #899 0::Int)
-#896 := (= #899 0::Int)
-#1076 := (or false #896)
-#1077 := [th-lemma arith]: #1076
-#1078 := [unit-resolution #1077 #78]: #896
-#1079 := (not #896)
-#1080 := (or #1079 #904)
-#1081 := [th-lemma arith triangle-eq]: #1080
-#1082 := [unit-resolution #1081 #1078]: #904
-#1132 := [hypothesis]: #1124
-#1133 := [th-lemma arith farkas -1 2 -1 -1 -1 1 #992 #1132 #999 #1082 #1000 #1131]: false
-#1136 := [lemma #1133]: #1135
-#1215 := [unit-resolution #1136 #1214]: #1134
-#1218 := [unit-resolution #1217 #1211]: #571
-#905 := (>= #899 0::Int)
-#1219 := (or #1079 #905)
-#1220 := [th-lemma arith triangle-eq]: #1219
-#1221 := [unit-resolution #1220 #1078]: #905
-#458 := (+ #21 #457)
-#369 := (>= #458 0::Int)
-#449 := (= #458 0::Int)
-#489 := (mod #453 -2::Int)
-#471 := (+ #21 #489)
-#474 := (= #471 0::Int)
-#455 := (if #630 #449 #474)
-#475 := (= #21 0::Int)
-#424 := (if #678 #475 #455)
-#375 := (or #555 #424)
-#459 := (mod #453 #416)
-#440 := (+ #21 #459)
-#441 := (= #440 0::Int)
-#462 := (or #681 #418)
-#464 := (not #462)
-#460 := (or #667 #418)
-#461 := (not #460)
-#463 := (or #461 #464)
-#442 := (if #463 #449 #441)
-#434 := (if #678 #475 #442)
-#467 := (= #21 f6)
-#443 := (if #341 #467 #434)
-#377 := (or #555 #443)
-#381 := (iff #377 #375)
-#382 := (iff #375 #375)
-#357 := [rewrite]: #382
-#384 := (iff #443 #424)
-#390 := (if false #467 #424)
-#385 := (iff #390 #424)
-#386 := [rewrite]: #385
-#402 := (iff #443 #390)
-#400 := (iff #434 #424)
-#456 := (iff #442 #455)
-#465 := (iff #441 #474)
-#472 := (= #440 #471)
-#469 := (= #459 #489)
-#470 := [monotonicity #725]: #469
-#473 := [monotonicity #470]: #472
-#454 := [monotonicity #473]: #465
-#466 := (iff #463 #630)
-#422 := (iff #464 #640)
-#420 := (iff #462 #681)
-#406 := (iff #462 #644)
-#419 := [monotonicity #392]: #406
-#421 := [trans #419 #648]: #420
-#423 := [monotonicity #421]: #422
-#414 := (iff #461 #641)
-#445 := (iff #460 #667)
-#444 := (iff #460 #655)
-#435 := [monotonicity #392]: #444
-#412 := [trans #435 #658]: #445
-#415 := [monotonicity #412]: #414
-#468 := [monotonicity #415 #423]: #466
-#413 := [monotonicity #468 #454]: #456
-#401 := [monotonicity #413]: #400
-#383 := [monotonicity #743 #401]: #402
-#387 := [trans #383 #386]: #384
-#376 := [monotonicity #387]: #381
-#361 := [trans #376 #357]: #381
-#378 := [quant-inst #8 #12]: #377
-#362 := [mp #378 #361]: #375
-#1222 := [unit-resolution #362 #767]: #424
-#348 := (not #424)
-#1223 := (or #348 #455)
-#353 := (or #348 #678 #455)
-#354 := [def-axiom]: #353
-#1224 := [unit-resolution #354 #1206]: #1223
-#1225 := [unit-resolution #1224 #1222]: #455
-#368 := (not #455)
-#373 := (or #368 #606 #449)
-#356 := [def-axiom]: #373
-#1226 := [unit-resolution #356 #1210 #1225]: #449
-#363 := (not #449)
-#1227 := (or #363 #369)
-#1228 := [th-lemma arith triangle-eq]: #1227
-#1229 := [unit-resolution #1228 #1226]: #369
-#346 := (>= #21 1::Int)
-#1084 := (not #346)
-#343 := (<= #21 1::Int)
-#912 := (>= #448 2::Int)
-#913 := (not #912)
-#1230 := (or false #913)
-#1231 := [th-lemma arith]: #1230
-#1232 := [unit-resolution #1231 #78]: #913
-#367 := (<= #458 0::Int)
-#1233 := (or #363 #367)
-#1234 := [th-lemma arith triangle-eq]: #1233
-#1235 := [unit-resolution #1234 #1226]: #367
-#1236 := (not #367)
-#1237 := (or #343 #912 #1236)
-#1238 := [th-lemma arith assign-bounds 1 1]: #1237
-#1239 := [unit-resolution #1238 #1235 #1232]: #343
-#1241 := (not #343)
-#1244 := (or #1241 #1084)
-#188 := (not #23)
-#1240 := [hypothesis]: #188
-#1242 := (or #23 #1241 #1084)
-#1243 := [th-lemma arith triangle-eq]: #1242
-#1245 := [unit-resolution #1243 #1240]: #1244
-#1246 := [unit-resolution #1245 #1239]: #1084
-#1247 := [th-lemma arith farkas -1/2 -1/2 1/2 -1/2 -1/2 -1/2 1 #1012 #1015 #1246 #1229 #1221 #1218 #1215]: false
-#1249 := [lemma #1247]: #1248
-#1462 := [unit-resolution #1249 #880]: #23
-#200 := (or #188 #189)
-#27 := (and #23 #26)
-#28 := (not #27)
-#209 := (iff #28 #200)
-#201 := (not #200)
-#204 := (not #201)
-#207 := (iff #204 #200)
-#208 := [rewrite]: #207
-#205 := (iff #28 #204)
-#202 := (iff #27 #201)
-#203 := [rewrite]: #202
-#206 := [monotonicity #203]: #205
-#210 := [trans #206 #208]: #209
-#82 := [asserted]: #28
-#211 := [mp #82 #210]: #200
-#1463 := [unit-resolution #211 #1462]: #189
-#1465 := (or #26 #1464 #1344)
-#1466 := [th-lemma arith triangle-eq]: #1465
-#1468 := [unit-resolution #1466 #1463]: #1467
-#1469 := [unit-resolution #1468 #1461]: #1344
-#1370 := (not #921)
-#1372 := (or #1369 #1370 #1371 #1266)
-#1345 := [hypothesis]: #1344
-#1294 := (<= #1289 0::Int)
-#1348 := [unit-resolution #1347 #78]: #1286
-#1350 := (or #1349 #1294)
-#1351 := [th-lemma arith triangle-eq]: #1350
-#1352 := [unit-resolution #1351 #1348]: #1294
-#1363 := [hypothesis]: #1339
-#1367 := [hypothesis]: #921
-#1368 := [th-lemma arith farkas -1 1 -2 1 1 1 1 1 1 1 #1367 #1366 #1363 #1362 #1361 #1354 #1352 #999 #992 #1345]: false
-#1373 := [lemma #1368]: #1372
-#1470 := [unit-resolution #1373 #1469 #1362 #1451]: #1369
-#1471 := [th-lemma arith farkas -2 1 1 1 1 1 1 1 1 #1470 #1435 #1434 #1433 #1430 #1015 #1012 #1447 #1400]: false
-#1473 := [lemma #1471]: #1472
-#1648 := [unit-resolution #1473 #1476]: #1647
-#1649 := [unit-resolution #1648 #1627 #1272]: false
-#1650 := [lemma #1649]: #686
-#1479 := (or #426 #535)
-#1423 := [hypothesis]: #686
-#723 := (+ #453 #347)
-#722 := (mod #723 -2::Int)
-#437 := (+ #13 #722)
-#717 := (= #437 0::Int)
-#741 := (not #431)
-#427 := (<= #10 0::Int)
-#735 := (not #427)
-#450 := (or #735 #741)
-#715 := (if #450 #535 #717)
-#589 := (not #426)
-#768 := [hypothesis]: #589
-#1441 := (or #426 #715)
-#720 := (if #426 #15 #715)
-#556 := (or #555 #720)
-#742 := (* -1::Int #10)
-#417 := (mod #742 #416)
-#749 := (+ #13 #417)
-#750 := (= #749 0::Int)
-#428 := (or #431 #418)
-#432 := (not #428)
-#429 := (or #427 #418)
-#430 := (not #429)
-#411 := (or #430 #432)
-#751 := (if #411 #535 #750)
-#752 := (if #426 #15 #751)
-#425 := (= #13 #10)
-#747 := (if #341 #425 #752)
-#557 := (or #555 #747)
-#700 := (iff #557 #556)
-#702 := (iff #556 #556)
-#696 := [rewrite]: #702
-#708 := (iff #747 #720)
-#745 := (* -1::Int #13)
-#388 := (+ f7 #745)
-#729 := (+ f6 #388)
-#744 := (= #729 0::Int)
-#711 := (if false #744 #720)
-#712 := (iff #711 #720)
-#713 := [rewrite]: #712
-#706 := (iff #747 #711)
-#709 := (iff #752 #720)
-#719 := (iff #751 #715)
-#718 := (iff #750 #717)
-#438 := (= #749 #437)
-#726 := (= #417 #722)
-#724 := (= #742 #723)
-#446 := [rewrite]: #724
-#436 := [monotonicity #446 #725]: #726
-#439 := [monotonicity #436]: #438
-#433 := [monotonicity #439]: #718
-#451 := (iff #411 #450)
-#727 := (iff #432 #741)
-#740 := (iff #428 #431)
-#374 := (or #431 false)
-#379 := (iff #374 #431)
-#380 := [rewrite]: #379
-#737 := (iff #428 #374)
-#739 := [monotonicity #392]: #737
-#738 := [trans #739 #380]: #740
-#728 := [monotonicity #738]: #727
-#730 := (iff #430 #735)
-#733 := (iff #429 #427)
-#393 := (or #427 false)
-#731 := (iff #393 #427)
-#732 := [rewrite]: #731
-#394 := (iff #429 #393)
-#395 := [monotonicity #392]: #394
-#734 := [trans #395 #732]: #733
-#736 := [monotonicity #734]: #730
-#452 := [monotonicity #736 #728]: #451
-#716 := [monotonicity #452 #433]: #719
-#710 := [monotonicity #716]: #709
-#408 := (iff #425 #744)
-#403 := [rewrite]: #408
-#707 := [monotonicity #743 #403 #710]: #706
-#714 := [trans #707 #713]: #708
-#701 := [monotonicity #714]: #700
-#697 := [trans #701 #696]: #700
-#699 := [quant-inst #10 #12]: #557
-#703 := [mp #699 #697]: #556
-#1440 := [unit-resolution #703 #767]: #720
-#587 := (not #720)
-#591 := (or #587 #426 #715)
-#592 := [def-axiom]: #591
-#1442 := [unit-resolution #592 #1440]: #1441
-#1443 := [unit-resolution #1442 #768]: #715
-#775 := (or #450 #426)
-#536 := (not #450)
-#769 := [hypothesis]: #536
-#704 := (or #450 #427)
-#698 := [def-axiom]: #704
-#770 := [unit-resolution #698 #769]: #427
-#705 := (or #450 #431)
-#534 := [def-axiom]: #705
-#771 := [unit-resolution #534 #769]: #431
-#772 := (or #426 #735 #741)
-#773 := [th-lemma arith triangle-eq]: #772
-#774 := [unit-resolution #773 #771 #770 #768]: false
-#776 := [lemma #774]: #775
-#1477 := [unit-resolution #776 #768]: #450
-#695 := (not #715)
-#577 := (or #695 #536 #535)
-#578 := [def-axiom]: #577
-#1478 := [unit-resolution #578 #1477 #1443 #1423]: false
-#1480 := [lemma #1478]: #1479
-#1651 := [unit-resolution #1480 #1650]: #426
-#1652 := (or #589 #431)
-#1653 := [th-lemma arith triangle-eq]: #1652
-#1654 := [unit-resolution #1653 #1651]: #431
-#1655 := (or #589 #427)
-#1656 := [th-lemma arith triangle-eq]: #1655
-#1657 := [unit-resolution #1656 #1651]: #427
-#1645 := (or #523 #741 #735)
-#1513 := [hypothesis]: #427
-#1580 := (or #497 #735 #667)
-#1022 := [hypothesis]: #641
-#1487 := [hypothesis]: #548
-#1579 := [th-lemma arith farkas -1 1 1 #1513 #1487 #1022]: false
-#1581 := [lemma #1579]: #1580
-#1641 := [unit-resolution #1581 #1414 #1513]: #667
-#1642 := [unit-resolution #830 #1206]: #630
-#1643 := [unit-resolution #1642 #1641]: #640
-#1573 := [hypothesis]: #431
-#1644 := [th-lemma arith farkas -1 1 1 #1573 #1643 #1427]: false
-#1646 := [lemma #1644]: #1645
-#1658 := [unit-resolution #1646 #1657 #1654]: #523
-#1659 := [unit-resolution #1453 #1658]: #920
-#1660 := (or #1265 #1458)
-#1623 := [hypothesis]: #1302
-#1624 := [unit-resolution #1456 #78 #1623]: false
-#1625 := [lemma #1624]: #1303
-#1661 := [unit-resolution #1460 #1625]: #1660
-#1662 := [unit-resolution #1661 #1659]: #1265
-#1503 := (+ #673 #1331)
-#1609 := (+ #521 #1503)
-#1610 := (+ #18 #1609)
-#1611 := (+ f7 #1610)
-#1612 := (+ f6 #1611)
-#1613 := (>= #1612 2::Int)
-#1620 := (not #1613)
-#1621 := (or #1620 #735)
-#1512 := [unit-resolution #1351 #1476]: #1294
-#1618 := [hypothesis]: #1613
-#1619 := [th-lemma arith farkas 2 -1 -1 -1 -1 -1 1 #1618 #1513 #1354 #999 #992 #1512 #1400]: false
-#1622 := [lemma #1619]: #1621
-#1663 := [unit-resolution #1622 #1657]: #1620
-#1664 := [unit-resolution #1450 #1658]: #921
-#1639 := (or #1370 #1613 #741 #1266)
-#1597 := [unit-resolution #1446 #1476]: #1422
-#1637 := [th-lemma arith #1573 #1345 #1367 #1435 #1015 #1012 #1597]: #1613
-#1636 := [hypothesis]: #1620
-#1638 := [unit-resolution #1636 #1637]: false
-#1640 := [lemma #1638]: #1639
-#1665 := [unit-resolution #1640 #1664 #1654 #1663]: #1266
-[unit-resolution #1468 #1665 #1662]: false
-unsat
-f966ee970dc5619d71e606afb53aade7fa8a1452 24 0
-#2 := false
-#7 := (exists (vars (?v0 Real)) false)
-#8 := (not #7)
-#9 := (not #8)
-#45 := (iff #9 false)
-#1 := true
-#40 := (not true)
-#43 := (iff #40 false)
-#44 := [rewrite]: #43
-#41 := (iff #9 #40)
-#38 := (iff #8 true)
-#33 := (not false)
-#36 := (iff #33 true)
-#37 := [rewrite]: #36
-#34 := (iff #8 #33)
-#31 := (iff #7 false)
-#32 := [elim-unused]: #31
-#35 := [monotonicity #32]: #34
-#39 := [trans #35 #37]: #38
-#42 := [monotonicity #39]: #41
-#46 := [trans #42 #44]: #45
-#30 := [asserted]: #9
-[mp #30 #46]: false
-unsat
-c4f4c8220660d1979009b33a643f0927bee816b1 1 0
-unsat
-e7ef76d73ccb9bc09d2b5368495a7a59d1bae3dc 1 0
-unsat
-db6426d59fdd57da8ca5d11de399761d1f1443de 1 0
-unsat
-a2da5fa16f268876e3dcbc1874e34212d0a36218 54 0
-#2 := false
-#11 := 1::Int
-#8 := 0::Int
-#135 := (= 0::Int 1::Int)
-#137 := (iff #135 false)
-#138 := [rewrite]: #137
-decl ?v1!0 :: Int
-#55 := ?v1!0
-#58 := (= ?v1!0 1::Int)
-decl ?v0!1 :: Int
-#56 := ?v0!1
-#57 := (= ?v0!1 0::Int)
-#50 := (and #57 #58)
-#59 := (= ?v0!1 ?v1!0)
-#60 := (not #59)
-#52 := (not #50)
-#61 := (or #52 #60)
-#62 := (not #61)
-#10 := (:var 0 Int)
-#7 := (:var 1 Int)
-#14 := (= #7 #10)
-#15 := (not #14)
-#12 := (= #10 1::Int)
-#9 := (= #7 0::Int)
-#13 := (and #9 #12)
-#39 := (not #13)
-#40 := (or #39 #15)
-#43 := (forall (vars (?v0 Int) (?v1 Int)) #40)
-#46 := (not #43)
-#63 := (~ #46 #62)
-#64 := [sk]: #63
-#16 := (implies #13 #15)
-#17 := (forall (vars (?v0 Int) (?v1 Int)) #16)
-#18 := (not #17)
-#47 := (iff #18 #46)
-#44 := (iff #17 #43)
-#41 := (iff #16 #40)
-#42 := [rewrite]: #41
-#45 := [quant-intro #42]: #44
-#48 := [monotonicity #45]: #47
-#38 := [asserted]: #18
-#51 := [mp #38 #48]: #46
-#67 := [mp~ #51 #64]: #62
-#70 := [not-or-elim #67]: #50
-#72 := [and-elim #70]: #58
-#133 := (= 0::Int ?v1!0)
-#73 := [not-or-elim #67]: #59
-#131 := (= 0::Int ?v0!1)
-#71 := [and-elim #70]: #57
-#132 := [symm #71]: #131
-#134 := [trans #132 #73]: #133
-#136 := [trans #134 #72]: #135
-[mp #136 #138]: false
-unsat
-46597b09986e0d4d045609318eeba242d6132e5c 82 0
-#2 := false
-#8 := (:var 0 Int)
-#10 := 0::Int
-#12 := (<= 0::Int #8)
-#11 := (< #8 0::Int)
-#13 := (or #11 #12)
-#7 := (:var 1 Int)
-#9 := (< #7 #8)
-#14 := (implies #9 #13)
-#15 := (forall (vars (?v1 Int)) #14)
-#16 := (exists (vars (?v0 Int)) #15)
-#17 := (not #16)
-#102 := (iff #17 false)
-#38 := (not #9)
-#39 := (or #38 #13)
-#42 := (forall (vars (?v1 Int)) #39)
-#45 := (exists (vars (?v0 Int)) #42)
-#48 := (not #45)
-#100 := (iff #48 false)
-#1 := true
-#95 := (not true)
-#98 := (iff #95 false)
-#99 := [rewrite]: #98
-#96 := (iff #48 #95)
-#93 := (iff #45 true)
-#88 := (exists (vars (?v0 Int)) true)
-#91 := (iff #88 true)
-#92 := [elim-unused]: #91
-#89 := (iff #45 #88)
-#86 := (iff #42 true)
-#81 := (forall (vars (?v1 Int)) true)
-#84 := (iff #81 true)
-#85 := [elim-unused]: #84
-#82 := (iff #42 #81)
-#79 := (iff #39 true)
-#53 := (>= #8 0::Int)
-#51 := (not #53)
-#71 := (or #51 #53)
-#57 := -1::Int
-#60 := (* -1::Int #8)
-#61 := (+ #7 #60)
-#59 := (>= #61 0::Int)
-#74 := (or #59 #71)
-#77 := (iff #74 true)
-#78 := [rewrite]: #77
-#75 := (iff #39 #74)
-#72 := (iff #13 #71)
-#55 := (iff #12 #53)
-#56 := [rewrite]: #55
-#52 := (iff #11 #51)
-#54 := [rewrite]: #52
-#73 := [monotonicity #54 #56]: #72
-#69 := (iff #38 #59)
-#58 := (not #59)
-#64 := (not #58)
-#67 := (iff #64 #59)
-#68 := [rewrite]: #67
-#65 := (iff #38 #64)
-#62 := (iff #9 #58)
-#63 := [rewrite]: #62
-#66 := [monotonicity #63]: #65
-#70 := [trans #66 #68]: #69
-#76 := [monotonicity #70 #73]: #75
-#80 := [trans #76 #78]: #79
-#83 := [quant-intro #80]: #82
-#87 := [trans #83 #85]: #86
-#90 := [quant-intro #87]: #89
-#94 := [trans #90 #92]: #93
-#97 := [monotonicity #94]: #96
-#101 := [trans #97 #99]: #100
-#49 := (iff #17 #48)
-#46 := (iff #16 #45)
-#43 := (iff #15 #42)
-#40 := (iff #14 #39)
-#41 := [rewrite]: #40
-#44 := [quant-intro #41]: #43
-#47 := [quant-intro #44]: #46
-#50 := [monotonicity #47]: #49
-#103 := [trans #50 #101]: #102
-#37 := [asserted]: #17
-[mp #37 #103]: false
-unsat
-aea13e787f95ed97feac7bd1dfc69160a5b8be70 78 0
-#2 := false
-#8 := (:var 0 Int)
-#10 := 2::Int
-#14 := (* 2::Int #8)
-#12 := 1::Int
-#7 := (:var 1 Int)
-#11 := (* 2::Int #7)
-#13 := (+ #11 1::Int)
-#15 := (< #13 #14)
-#9 := (< #7 #8)
-#16 := (implies #9 #15)
-#17 := (forall (vars (?v0 Int) (?v1 Int)) #16)
-#18 := (not #17)
-#98 := (iff #18 false)
-#40 := (+ 1::Int #11)
-#43 := (< #40 #14)
-#49 := (not #9)
-#50 := (or #49 #43)
-#55 := (forall (vars (?v0 Int) (?v1 Int)) #50)
-#58 := (not #55)
-#96 := (iff #58 false)
-#1 := true
-#91 := (not true)
-#94 := (iff #91 false)
-#95 := [rewrite]: #94
-#92 := (iff #58 #91)
-#89 := (iff #55 true)
-#84 := (forall (vars (?v0 Int) (?v1 Int)) true)
-#87 := (iff #84 true)
-#88 := [elim-unused]: #87
-#85 := (iff #55 #84)
-#82 := (iff #50 true)
-#20 := 0::Int
-#61 := -1::Int
-#64 := (* -1::Int #8)
-#65 := (+ #7 #64)
-#63 := (>= #65 0::Int)
-#62 := (not #63)
-#76 := (or #63 #62)
-#80 := (iff #76 true)
-#81 := [rewrite]: #80
-#78 := (iff #50 #76)
-#77 := (iff #43 #62)
-#75 := [rewrite]: #77
-#73 := (iff #49 #63)
-#68 := (not #62)
-#71 := (iff #68 #63)
-#72 := [rewrite]: #71
-#69 := (iff #49 #68)
-#66 := (iff #9 #62)
-#67 := [rewrite]: #66
-#70 := [monotonicity #67]: #69
-#74 := [trans #70 #72]: #73
-#79 := [monotonicity #74 #75]: #78
-#83 := [trans #79 #81]: #82
-#86 := [quant-intro #83]: #85
-#90 := [trans #86 #88]: #89
-#93 := [monotonicity #90]: #92
-#97 := [trans #93 #95]: #96
-#59 := (iff #18 #58)
-#56 := (iff #17 #55)
-#53 := (iff #16 #50)
-#46 := (implies #9 #43)
-#51 := (iff #46 #50)
-#52 := [rewrite]: #51
-#47 := (iff #16 #46)
-#44 := (iff #15 #43)
-#41 := (= #13 #40)
-#42 := [rewrite]: #41
-#45 := [monotonicity #42]: #44
-#48 := [monotonicity #45]: #47
-#54 := [trans #48 #52]: #53
-#57 := [quant-intro #54]: #56
-#60 := [monotonicity #57]: #59
-#99 := [trans #60 #97]: #98
-#39 := [asserted]: #18
-[mp #39 #99]: false
-unsat
-e6703a33319f0e5148dba82e8205956f98cd7b63 56 0
-#2 := false
-#12 := (:var 0 Int)
-#7 := 2::Int
-#13 := (* 2::Int #12)
-#10 := 1::Int
-#8 := (:var 1 Int)
-#9 := (* 2::Int #8)
-#11 := (+ #9 1::Int)
-#14 := (= #11 #13)
-#15 := (not #14)
-#16 := (forall (vars (?v0 Int) (?v1 Int)) #15)
-#17 := (not #16)
-#77 := (iff #17 false)
-#39 := (+ 1::Int #9)
-#42 := (= #39 #13)
-#45 := (not #42)
-#48 := (forall (vars (?v0 Int) (?v1 Int)) #45)
-#51 := (not #48)
-#75 := (iff #51 false)
-#1 := true
-#70 := (not true)
-#73 := (iff #70 false)
-#74 := [rewrite]: #73
-#71 := (iff #51 #70)
-#68 := (iff #48 true)
-#63 := (forall (vars (?v0 Int) (?v1 Int)) true)
-#66 := (iff #63 true)
-#67 := [elim-unused]: #66
-#64 := (iff #48 #63)
-#61 := (iff #45 true)
-#54 := (not false)
-#59 := (iff #54 true)
-#60 := [rewrite]: #59
-#55 := (iff #45 #54)
-#56 := (iff #42 false)
-#57 := [rewrite]: #56
-#58 := [monotonicity #57]: #55
-#62 := [trans #58 #60]: #61
-#65 := [quant-intro #62]: #64
-#69 := [trans #65 #67]: #68
-#72 := [monotonicity #69]: #71
-#76 := [trans #72 #74]: #75
-#52 := (iff #17 #51)
-#49 := (iff #16 #48)
-#46 := (iff #15 #45)
-#43 := (iff #14 #42)
-#40 := (= #11 #39)
-#41 := [rewrite]: #40
-#44 := [monotonicity #41]: #43
-#47 := [monotonicity #44]: #46
-#50 := [quant-intro #47]: #49
-#53 := [monotonicity #50]: #52
-#78 := [trans #53 #76]: #77
-#38 := [asserted]: #17
-[mp #38 #78]: false
-unsat
-8a770e2a15f5bbced47daef21d1d322e18a383fb 89 0
-#2 := false
-#7 := 2::Int
-decl ?v0!1 :: Int
-#71 := ?v0!1
-decl ?v1!0 :: Int
-#70 := ?v1!0
-#85 := (+ ?v1!0 ?v0!1)
-#94 := (= #85 2::Int)
-#109 := (not #94)
-#97 := (>= #85 2::Int)
-#100 := (not #97)
-#88 := (<= #85 2::Int)
-#91 := (not #88)
-#103 := (or #91 #94 #100)
-#106 := (not #103)
-#72 := (+ ?v0!1 ?v1!0)
-#74 := (>= #72 2::Int)
-#75 := (not #74)
-#67 := (= #72 2::Int)
-#73 := (<= #72 2::Int)
-#40 := (not #73)
-#76 := (or #40 #67 #75)
-#77 := (not #76)
-#107 := (iff #77 #106)
-#104 := (iff #76 #103)
-#101 := (iff #75 #100)
-#98 := (iff #74 #97)
-#86 := (= #72 #85)
-#87 := [rewrite]: #86
-#99 := [monotonicity #87]: #98
-#102 := [monotonicity #99]: #101
-#95 := (iff #67 #94)
-#96 := [monotonicity #87]: #95
-#92 := (iff #40 #91)
-#89 := (iff #73 #88)
-#90 := [monotonicity #87]: #89
-#93 := [monotonicity #90]: #92
-#105 := [monotonicity #93 #96 #102]: #104
-#108 := [monotonicity #105]: #107
-#9 := (:var 0 Int)
-#8 := (:var 1 Int)
-#10 := (+ #8 #9)
-#44 := (>= #10 2::Int)
-#41 := (not #44)
-#12 := (= #10 2::Int)
-#45 := (<= #10 2::Int)
-#46 := (not #45)
-#55 := (or #46 #12 #41)
-#60 := (forall (vars (?v0 Int) (?v1 Int)) #55)
-#63 := (not #60)
-#78 := (~ #63 #77)
-#79 := [sk]: #78
-#13 := (< #10 2::Int)
-#14 := (or #12 #13)
-#11 := (< 2::Int #10)
-#15 := (or #11 #14)
-#16 := (forall (vars (?v0 Int) (?v1 Int)) #15)
-#17 := (not #16)
-#64 := (iff #17 #63)
-#61 := (iff #16 #60)
-#58 := (iff #15 #55)
-#49 := (or #12 #41)
-#52 := (or #46 #49)
-#56 := (iff #52 #55)
-#57 := [rewrite]: #56
-#53 := (iff #15 #52)
-#50 := (iff #14 #49)
-#43 := (iff #13 #41)
-#42 := [rewrite]: #43
-#51 := [monotonicity #42]: #50
-#47 := (iff #11 #46)
-#48 := [rewrite]: #47
-#54 := [monotonicity #48 #51]: #53
-#59 := [trans #54 #57]: #58
-#62 := [quant-intro #59]: #61
-#65 := [monotonicity #62]: #64
-#38 := [asserted]: #17
-#66 := [mp #38 #65]: #63
-#82 := [mp~ #66 #79]: #77
-#83 := [mp #82 #108]: #106
-#110 := [not-or-elim #83]: #109
-#111 := [not-or-elim #83]: #97
-#173 := (or #94 #100)
-#84 := [not-or-elim #83]: #88
-#171 := (or #94 #91 #100)
-#172 := [th-lemma arith triangle-eq]: #171
-#174 := [unit-resolution #172 #84]: #173
-[unit-resolution #174 #111 #110]: false
-unsat
-c93368b1109e5b13c7d8bc3c33d69c60ba539127 89 0
-#2 := false
-#7 := 0::Int
-decl ?v0!0 :: Int
-#87 := ?v0!0
-#88 := (<= ?v0!0 0::Int)
-#157 := (not #88)
-#166 := [hypothesis]: #88
-#10 := 1::Int
-#89 := (>= ?v0!0 1::Int)
-#90 := (not #89)
-#167 := (or #90 #157)
-#168 := [th-lemma arith farkas 1 1]: #167
-#169 := [unit-resolution #168 #166]: #90
-#170 := (or #157 #89)
-#56 := -1::Int
-#83 := (<= ?v0!0 -1::Int)
-#84 := (not #83)
-#91 := (if #88 #90 #84)
-#92 := (not #91)
-#8 := (:var 0 Int)
-#57 := (<= #8 -1::Int)
-#58 := (not #57)
-#62 := (>= #8 1::Int)
-#61 := (not #62)
-#52 := (<= #8 0::Int)
-#68 := (if #52 #61 #58)
-#73 := (forall (vars (?v0 Int)) #68)
-#76 := (not #73)
-#93 := (~ #76 #92)
-#94 := [sk]: #93
-#13 := (< #8 1::Int)
-#11 := (+ #8 1::Int)
-#12 := (< 0::Int #11)
-#9 := (< 0::Int #8)
-#14 := (if #9 #12 #13)
-#15 := (forall (vars (?v0 Int)) #14)
-#16 := (not #15)
-#79 := (iff #16 #76)
-#37 := (+ 1::Int #8)
-#40 := (< 0::Int #37)
-#43 := (if #9 #40 #13)
-#46 := (forall (vars (?v0 Int)) #43)
-#49 := (not #46)
-#77 := (iff #49 #76)
-#74 := (iff #46 #73)
-#71 := (iff #43 #68)
-#53 := (not #52)
-#65 := (if #53 #58 #61)
-#69 := (iff #65 #68)
-#70 := [rewrite]: #69
-#66 := (iff #43 #65)
-#63 := (iff #13 #61)
-#64 := [rewrite]: #63
-#59 := (iff #40 #58)
-#60 := [rewrite]: #59
-#54 := (iff #9 #53)
-#55 := [rewrite]: #54
-#67 := [monotonicity #55 #60 #64]: #66
-#72 := [trans #67 #70]: #71
-#75 := [quant-intro #72]: #74
-#78 := [monotonicity #75]: #77
-#50 := (iff #16 #49)
-#47 := (iff #15 #46)
-#44 := (iff #14 #43)
-#41 := (iff #12 #40)
-#38 := (= #11 #37)
-#39 := [rewrite]: #38
-#42 := [monotonicity #39]: #41
-#45 := [monotonicity #42]: #44
-#48 := [quant-intro #45]: #47
-#51 := [monotonicity #48]: #50
-#80 := [trans #51 #78]: #79
-#36 := [asserted]: #16
-#81 := [mp #36 #80]: #76
-#97 := [mp~ #81 #94]: #92
-#162 := (or #91 #157 #89)
-#163 := [def-axiom]: #162
-#171 := [unit-resolution #163 #97]: #170
-#172 := [unit-resolution #171 #169 #166]: false
-#173 := [lemma #172]: #157
-#174 := (or #84 #88)
-#175 := [th-lemma arith farkas 1 1]: #174
-#176 := [unit-resolution #175 #173]: #84
-#177 := (or #88 #83)
-#164 := (or #91 #88 #83)
-#165 := [def-axiom]: #164
-#178 := [unit-resolution #165 #97]: #177
-[unit-resolution #178 #176 #173]: false
-unsat
-8578dab7bf88c7d119f9af2e5f7eaf948f1bdb87 84 0
-WARNING: failed to find a pattern for quantifier (quantifier id: k!10)
-#2 := false
-#8 := 0::Int
-#7 := (:var 0 Int)
-#49 := (<= #7 0::Int)
-#50 := (not #49)
-#47 := (>= #7 0::Int)
-#45 := (not #47)
-#53 := (or #45 #50)
-#56 := (forall (vars (?v0 Int)) #53)
-#525 := (not #56)
-#218 := (<= 0::Int 0::Int)
-#539 := (not #218)
-#207 := (>= 0::Int 0::Int)
-#201 := (not #207)
-#537 := (or #201 #539)
-#526 := (or #525 #537)
-#170 := (iff #526 #525)
-#527 := (or #525 false)
-#530 := (iff #527 #525)
-#169 := [rewrite]: #530
-#164 := (iff #526 #527)
-#523 := (iff #537 false)
-#182 := (or false false)
-#185 := (iff #182 false)
-#522 := [rewrite]: #185
-#183 := (iff #537 #182)
-#178 := (iff #539 false)
-#1 := true
-#543 := (not true)
-#222 := (iff #543 false)
-#544 := [rewrite]: #222
-#194 := (iff #539 #543)
-#198 := (iff #218 true)
-#535 := [rewrite]: #198
-#536 := [monotonicity #535]: #194
-#520 := [trans #536 #544]: #178
-#534 := (iff #201 false)
-#538 := (iff #201 #543)
-#541 := (iff #207 true)
-#542 := [rewrite]: #541
-#326 := [monotonicity #542]: #538
-#193 := [trans #326 #544]: #534
-#184 := [monotonicity #193 #520]: #183
-#524 := [trans #184 #522]: #523
-#528 := [monotonicity #524]: #164
-#531 := [trans #528 #169]: #170
-#521 := [quant-inst #8]: #526
-#529 := [mp #521 #531]: #525
-#69 := (~ #56 #56)
-#67 := (~ #53 #53)
-#68 := [refl]: #67
-#70 := [nnf-pos #68]: #69
-#10 := (< 0::Int #7)
-#9 := (< #7 0::Int)
-#11 := (or #9 #10)
-#12 := (forall (vars (?v0 Int)) #11)
-#13 := (if #12 false true)
-#14 := (not #13)
-#59 := (iff #14 #56)
-#57 := (iff #12 #56)
-#54 := (iff #11 #53)
-#51 := (iff #10 #50)
-#52 := [rewrite]: #51
-#46 := (iff #9 #45)
-#48 := [rewrite]: #46
-#55 := [monotonicity #48 #52]: #54
-#58 := [quant-intro #55]: #57
-#43 := (iff #14 #12)
-#35 := (not #12)
-#38 := (not #35)
-#41 := (iff #38 #12)
-#42 := [rewrite]: #41
-#39 := (iff #14 #38)
-#36 := (iff #13 #35)
-#37 := [rewrite]: #36
-#40 := [monotonicity #37]: #39
-#44 := [trans #40 #42]: #43
-#60 := [trans #44 #58]: #59
-#34 := [asserted]: #14
-#61 := [mp #34 #60]: #56
-#63 := [mp~ #61 #70]: #56
-[unit-resolution #63 #529]: false
-unsat
-f6f0c702e5caae5d1fc0a3e7862c44d261de6d47 63 0
-#2 := false
-#15 := 1::Int
-#12 := (:var 1 Int)
-#10 := 6::Int
-#11 := (- 6::Int)
-#13 := (* #11 #12)
-#8 := (:var 2 Int)
-#7 := 4::Int
-#9 := (* 4::Int #8)
-#14 := (+ #9 #13)
-#16 := (= #14 1::Int)
-#17 := (exists (vars (?v0 Int) (?v1 Int) (?v2 Int)) #16)
-#18 := (not #17)
-#19 := (not #18)
-#86 := (iff #19 false)
-#56 := (:var 0 Int)
-#41 := -6::Int
-#58 := (* -6::Int #56)
-#57 := (* 4::Int #12)
-#59 := (+ #57 #58)
-#60 := (= #59 1::Int)
-#61 := (exists (vars (?v0 Int) (?v1 Int)) #60)
-#84 := (iff #61 false)
-#77 := (exists (vars (?v0 Int) (?v1 Int)) false)
-#82 := (iff #77 false)
-#83 := [elim-unused]: #82
-#80 := (iff #61 #77)
-#78 := (iff #60 false)
-#79 := [rewrite]: #78
-#81 := [quant-intro #79]: #80
-#85 := [trans #81 #83]: #84
-#74 := (iff #19 #61)
-#66 := (not #61)
-#69 := (not #66)
-#72 := (iff #69 #61)
-#73 := [rewrite]: #72
-#70 := (iff #19 #69)
-#67 := (iff #18 #66)
-#64 := (iff #17 #61)
-#44 := (* -6::Int #12)
-#47 := (+ #9 #44)
-#50 := (= #47 1::Int)
-#53 := (exists (vars (?v0 Int) (?v1 Int) (?v2 Int)) #50)
-#62 := (iff #53 #61)
-#63 := [elim-unused]: #62
-#54 := (iff #17 #53)
-#51 := (iff #16 #50)
-#48 := (= #14 #47)
-#45 := (= #13 #44)
-#42 := (= #11 -6::Int)
-#43 := [rewrite]: #42
-#46 := [monotonicity #43]: #45
-#49 := [monotonicity #46]: #48
-#52 := [monotonicity #49]: #51
-#55 := [quant-intro #52]: #54
-#65 := [trans #55 #63]: #64
-#68 := [monotonicity #65]: #67
-#71 := [monotonicity #68]: #70
-#75 := [trans #71 #73]: #74
-#87 := [trans #75 #85]: #86
-#40 := [asserted]: #19
-[mp #40 #87]: false
-unsat
-252d255c564463d916bc68156eea8dbe7fb0be0a 165 0
-WARNING: failed to find a pattern for quantifier (quantifier id: k!10)
-#2 := false
-#7 := 0::Int
-#8 := (:var 0 Int)
-#55 := (<= #8 0::Int)
-#56 := (not #55)
-#52 := (>= #8 0::Int)
-#51 := (not #52)
-#59 := (or #51 #56)
-#62 := (forall (vars (?v0 Int)) #59)
-#95 := (not #62)
-#587 := (<= 0::Int 0::Int)
-#586 := (not #587)
-#585 := (>= 0::Int 0::Int)
-#248 := (not #585)
-#593 := (or #248 #586)
-#290 := (or #95 #593)
-#569 := (iff #290 #95)
-#292 := (or #95 false)
-#572 := (iff #292 #95)
-#287 := [rewrite]: #572
-#293 := (iff #290 #292)
-#576 := (iff #593 false)
-#578 := (or false false)
-#575 := (iff #578 false)
-#579 := [rewrite]: #575
-#300 := (iff #593 #578)
-#201 := (iff #586 false)
-#1 := true
-#594 := (not true)
-#592 := (iff #594 false)
-#595 := [rewrite]: #592
-#306 := (iff #586 #594)
-#304 := (iff #587 true)
-#305 := [rewrite]: #304
-#307 := [monotonicity #305]: #306
-#577 := [trans #307 #595]: #201
-#581 := (iff #248 false)
-#589 := (iff #248 #594)
-#233 := (iff #585 true)
-#234 := [rewrite]: #233
-#249 := [monotonicity #234]: #589
-#582 := [trans #249 #595]: #581
-#301 := [monotonicity #582 #577]: #300
-#580 := [trans #301 #579]: #576
-#571 := [monotonicity #580]: #293
-#573 := [trans #571 #287]: #569
-#291 := [quant-inst #7]: #290
-#570 := [mp #291 #573]: #95
-decl z3name!0 :: bool
-#92 := z3name!0
-#15 := 3::Int
-#39 := -1::Int
-#99 := (if z3name!0 -1::Int 3::Int)
-#284 := (= #99 3::Int)
-#604 := (not #284)
-#602 := (>= #99 3::Int)
-#259 := (not #602)
-#102 := (<= #99 0::Int)
-#65 := (if #62 -1::Int 3::Int)
-#71 := (<= #65 0::Int)
-#103 := (~ #71 #102)
-#100 := (= #65 #99)
-#97 := (~ #62 z3name!0)
-#88 := (or z3name!0 #95)
-#93 := (not z3name!0)
-#94 := (or #93 #62)
-#89 := (and #94 #88)
-#96 := [intro-def]: #89
-#98 := [apply-def #96]: #97
-#101 := [monotonicity #98]: #100
-#104 := [monotonicity #101]: #103
-#13 := 1::Int
-#14 := (- 1::Int)
-#10 := (< 0::Int #8)
-#9 := (< #8 0::Int)
-#11 := (or #9 #10)
-#12 := (forall (vars (?v0 Int)) #11)
-#16 := (if #12 #14 3::Int)
-#17 := (< 0::Int #16)
-#18 := (not #17)
-#84 := (iff #18 #71)
-#42 := (if #12 -1::Int 3::Int)
-#45 := (< 0::Int #42)
-#48 := (not #45)
-#82 := (iff #48 #71)
-#72 := (not #71)
-#77 := (not #72)
-#80 := (iff #77 #71)
-#81 := [rewrite]: #80
-#78 := (iff #48 #77)
-#75 := (iff #45 #72)
-#68 := (< 0::Int #65)
-#73 := (iff #68 #72)
-#74 := [rewrite]: #73
-#69 := (iff #45 #68)
-#66 := (= #42 #65)
-#63 := (iff #12 #62)
-#60 := (iff #11 #59)
-#57 := (iff #10 #56)
-#58 := [rewrite]: #57
-#53 := (iff #9 #51)
-#54 := [rewrite]: #53
-#61 := [monotonicity #54 #58]: #60
-#64 := [quant-intro #61]: #63
-#67 := [monotonicity #64]: #66
-#70 := [monotonicity #67]: #69
-#76 := [trans #70 #74]: #75
-#79 := [monotonicity #76]: #78
-#83 := [trans #79 #81]: #82
-#49 := (iff #18 #48)
-#46 := (iff #17 #45)
-#43 := (= #16 #42)
-#40 := (= #14 -1::Int)
-#41 := [rewrite]: #40
-#44 := [monotonicity #41]: #43
-#47 := [monotonicity #44]: #46
-#50 := [monotonicity #47]: #49
-#85 := [trans #50 #83]: #84
-#38 := [asserted]: #18
-#86 := [mp #38 #85]: #71
-#133 := [mp~ #86 #104]: #102
-#389 := (not #102)
-#596 := (or #259 #389)
-#270 := [th-lemma arith farkas 1 1]: #596
-#271 := [unit-resolution #270 #133]: #259
-#603 := [hypothesis]: #284
-#605 := (or #604 #602)
-#606 := [th-lemma arith triangle-eq]: #605
-#601 := [unit-resolution #606 #603 #271]: false
-#607 := [lemma #601]: #604
-#286 := (or z3name!0 #284)
-#265 := [def-axiom]: #286
-#574 := [unit-resolution #265 #607]: z3name!0
-decl ?v0!1 :: Int
-#115 := ?v0!1
-#118 := (<= ?v0!1 0::Int)
-#119 := (not #118)
-#116 := (>= ?v0!1 0::Int)
-#117 := (not #116)
-#120 := (or #117 #119)
-#121 := (not #120)
-#126 := (or z3name!0 #121)
-#129 := (and #94 #126)
-#130 := (~ #89 #129)
-#127 := (~ #88 #126)
-#122 := (~ #95 #121)
-#123 := [sk]: #122
-#113 := (~ z3name!0 z3name!0)
-#114 := [refl]: #113
-#128 := [monotonicity #114 #123]: #127
-#111 := (~ #94 #94)
-#109 := (~ #62 #62)
-#107 := (~ #59 #59)
-#108 := [refl]: #107
-#110 := [nnf-pos #108]: #109
-#105 := (~ #93 #93)
-#106 := [refl]: #105
-#112 := [monotonicity #106 #110]: #111
-#131 := [monotonicity #112 #128]: #130
-#132 := [mp~ #96 #131]: #129
-#136 := [and-elim #132]: #94
-#563 := [unit-resolution #136 #574]: #62
-[unit-resolution #563 #570]: false
-unsat
-302156fb98e1f9b5657a3c89c418d5e1813f274a 101 0
-#2 := false
-#7 := 0::Int
-decl ?v1!1 :: Int
-#92 := ?v1!1
-decl ?v2!0 :: Int
-#91 := ?v2!0
-#109 := (+ ?v2!0 ?v1!1)
-#112 := (<= #109 0::Int)
-#115 := (not #112)
-#87 := (<= ?v2!0 0::Int)
-#88 := (not #87)
-#93 := (<= ?v1!1 0::Int)
-#94 := (not #93)
-#95 := (and #94 #88)
-#96 := (not #95)
-#118 := (or #96 #115)
-#121 := (not #118)
-#97 := (+ ?v1!1 ?v2!0)
-#98 := (<= #97 0::Int)
-#99 := (not #98)
-#100 := (or #96 #99)
-#101 := (not #100)
-#122 := (iff #101 #121)
-#119 := (iff #100 #118)
-#116 := (iff #99 #115)
-#113 := (iff #98 #112)
-#110 := (= #97 #109)
-#111 := [rewrite]: #110
-#114 := [monotonicity #111]: #113
-#117 := [monotonicity #114]: #116
-#120 := [monotonicity #117]: #119
-#123 := [monotonicity #120]: #122
-#10 := (:var 0 Int)
-#8 := (:var 1 Int)
-#13 := (+ #8 #10)
-#70 := (<= #13 0::Int)
-#71 := (not #70)
-#60 := (<= #10 0::Int)
-#61 := (not #60)
-#56 := (<= #8 0::Int)
-#57 := (not #56)
-#64 := (and #57 #61)
-#67 := (not #64)
-#74 := (or #67 #71)
-#77 := (forall (vars (?v1 Int) (?v2 Int)) #74)
-#80 := (not #77)
-#102 := (~ #80 #101)
-#103 := [sk]: #102
-#14 := (< 0::Int #13)
-#11 := (< 0::Int #10)
-#9 := (< 0::Int #8)
-#12 := (and #9 #11)
-#15 := (implies #12 #14)
-#16 := (forall (vars (?v1 Int) (?v2 Int)) #15)
-#17 := (exists (vars (?v0 Int)) #16)
-#18 := (not #17)
-#83 := (iff #18 #80)
-#39 := (not #12)
-#40 := (or #39 #14)
-#43 := (forall (vars (?v1 Int) (?v2 Int)) #40)
-#53 := (not #43)
-#81 := (iff #53 #80)
-#78 := (iff #43 #77)
-#75 := (iff #40 #74)
-#72 := (iff #14 #71)
-#73 := [rewrite]: #72
-#68 := (iff #39 #67)
-#65 := (iff #12 #64)
-#62 := (iff #11 #61)
-#63 := [rewrite]: #62
-#58 := (iff #9 #57)
-#59 := [rewrite]: #58
-#66 := [monotonicity #59 #63]: #65
-#69 := [monotonicity #66]: #68
-#76 := [monotonicity #69 #73]: #75
-#79 := [quant-intro #76]: #78
-#82 := [monotonicity #79]: #81
-#54 := (iff #18 #53)
-#51 := (iff #17 #43)
-#46 := (exists (vars (?v0 Int)) #43)
-#49 := (iff #46 #43)
-#50 := [elim-unused]: #49
-#47 := (iff #17 #46)
-#44 := (iff #16 #43)
-#41 := (iff #15 #40)
-#42 := [rewrite]: #41
-#45 := [quant-intro #42]: #44
-#48 := [quant-intro #45]: #47
-#52 := [trans #48 #50]: #51
-#55 := [monotonicity #52]: #54
-#84 := [trans #55 #82]: #83
-#38 := [asserted]: #18
-#85 := [mp #38 #84]: #80
-#106 := [mp~ #85 #103]: #101
-#107 := [mp #106 #123]: #121
-#126 := [not-or-elim #107]: #112
-#108 := [not-or-elim #107]: #95
-#124 := [and-elim #108]: #94
-#125 := [and-elim #108]: #88
-[th-lemma arith farkas 1 1 1 #125 #124 #126]: false
-unsat
-bcc217c52aea6d752e93b67733058589bedd0079 99 0
-#2 := false
-#39 := -1::Int
-decl ?v1!1 :: Int
-#101 := ?v1!1
-#106 := (<= ?v1!1 -1::Int)
-#107 := (not #106)
-#10 := 0::Real
-decl ?v2!0 :: Real
-#100 := ?v2!0
-#102 := (<= ?v2!0 0::Real)
-#103 := (not #102)
-#7 := 0::Int
-#98 := (<= ?v1!1 0::Int)
-#99 := (not #98)
-#104 := (and #99 #103)
-#105 := (not #104)
-#108 := (or #105 #107)
-#109 := (not #108)
-#8 := (:var 1 Int)
-#81 := (<= #8 -1::Int)
-#82 := (not #81)
-#11 := (:var 0 Real)
-#71 := (<= #11 0::Real)
-#72 := (not #71)
-#67 := (<= #8 0::Int)
-#68 := (not #67)
-#75 := (and #68 #72)
-#78 := (not #75)
-#85 := (or #78 #82)
-#88 := (forall (vars (?v1 Int) (?v2 Real)) #85)
-#91 := (not #88)
-#110 := (~ #91 #109)
-#111 := [sk]: #110
-#14 := 1::Int
-#15 := (- 1::Int)
-#16 := (< #15 #8)
-#12 := (< 0::Real #11)
-#9 := (< 0::Int #8)
-#13 := (and #9 #12)
-#17 := (implies #13 #16)
-#18 := (forall (vars (?v1 Int) (?v2 Real)) #17)
-#19 := (exists (vars (?v0 Int)) #18)
-#20 := (not #19)
-#94 := (iff #20 #91)
-#42 := (< -1::Int #8)
-#48 := (not #13)
-#49 := (or #48 #42)
-#54 := (forall (vars (?v1 Int) (?v2 Real)) #49)
-#64 := (not #54)
-#92 := (iff #64 #91)
-#89 := (iff #54 #88)
-#86 := (iff #49 #85)
-#83 := (iff #42 #82)
-#84 := [rewrite]: #83
-#79 := (iff #48 #78)
-#76 := (iff #13 #75)
-#73 := (iff #12 #72)
-#74 := [rewrite]: #73
-#69 := (iff #9 #68)
-#70 := [rewrite]: #69
-#77 := [monotonicity #70 #74]: #76
-#80 := [monotonicity #77]: #79
-#87 := [monotonicity #80 #84]: #86
-#90 := [quant-intro #87]: #89
-#93 := [monotonicity #90]: #92
-#65 := (iff #20 #64)
-#62 := (iff #19 #54)
-#57 := (exists (vars (?v0 Int)) #54)
-#60 := (iff #57 #54)
-#61 := [elim-unused]: #60
-#58 := (iff #19 #57)
-#55 := (iff #18 #54)
-#52 := (iff #17 #49)
-#45 := (implies #13 #42)
-#50 := (iff #45 #49)
-#51 := [rewrite]: #50
-#46 := (iff #17 #45)
-#43 := (iff #16 #42)
-#40 := (= #15 -1::Int)
-#41 := [rewrite]: #40
-#44 := [monotonicity #41]: #43
-#47 := [monotonicity #44]: #46
-#53 := [trans #47 #51]: #52
-#56 := [quant-intro #53]: #55
-#59 := [quant-intro #56]: #58
-#63 := [trans #59 #61]: #62
-#66 := [monotonicity #63]: #65
-#95 := [trans #66 #93]: #94
-#38 := [asserted]: #20
-#96 := [mp #38 #95]: #91
-#114 := [mp~ #96 #111]: #109
-#120 := [not-or-elim #114]: #106
-#117 := [not-or-elim #114]: #104
-#118 := [and-elim #117]: #99
-#178 := (or #107 #98)
-#179 := [th-lemma arith farkas 1 1]: #178
-#180 := [unit-resolution #179 #118]: #107
-[unit-resolution #180 #120]: false
-unsat
-8a78832884e41117489fba88c88de0b5cacb832a 143 0
-#2 := false
-#10 := 0::Int
-#8 := (:var 0 Int)
-#68 := (<= #8 0::Int)
-#69 := (not #68)
-#146 := (not false)
-#149 := (or #146 #69)
-#152 := (not #149)
-#155 := (forall (vars (?v0 Int)) #152)
-#182 := (iff #155 false)
-#177 := (forall (vars (?v0 Int)) false)
-#180 := (iff #177 false)
-#181 := [elim-unused]: #180
-#178 := (iff #155 #177)
-#175 := (iff #152 false)
-#1 := true
-#170 := (not true)
-#173 := (iff #170 false)
-#174 := [rewrite]: #173
-#171 := (iff #152 #170)
-#168 := (iff #149 true)
-#163 := (or true #69)
-#166 := (iff #163 true)
-#167 := [rewrite]: #166
-#164 := (iff #149 #163)
-#161 := (iff #146 true)
-#162 := [rewrite]: #161
-#165 := [monotonicity #162]: #164
-#169 := [trans #165 #167]: #168
-#172 := [monotonicity #169]: #171
-#176 := [trans #172 #174]: #175
-#179 := [quant-intro #176]: #178
-#183 := [trans #179 #181]: #182
-#59 := -1::Int
-#60 := (* -1::Int #8)
-#7 := (:var 1 Int)
-#61 := (+ #7 #60)
-#62 := (<= #61 0::Int)
-#65 := (not #62)
-#72 := (or #65 #69)
-#75 := (forall (vars (?v1 Int)) #72)
-#78 := (not #75)
-#81 := (or #78 #69)
-#107 := (not #81)
-#125 := (forall (vars (?v0 Int)) #107)
-#158 := (iff #125 #155)
-#129 := (forall (vars (?v1 Int)) #69)
-#132 := (not #129)
-#135 := (or #132 #69)
-#138 := (not #135)
-#141 := (forall (vars (?v0 Int)) #138)
-#156 := (iff #141 #155)
-#157 := [rewrite]: #156
-#142 := (iff #125 #141)
-#143 := [rewrite]: #142
-#159 := [trans #143 #157]: #158
-#118 := (and #75 #68)
-#121 := (forall (vars (?v0 Int)) #118)
-#126 := (iff #121 #125)
-#115 := (iff #118 #107)
-#124 := [rewrite]: #115
-#127 := [quant-intro #124]: #126
-#103 := (not #69)
-#106 := (and #75 #103)
-#110 := (forall (vars (?v0 Int)) #106)
-#122 := (iff #110 #121)
-#119 := (iff #106 #118)
-#116 := (iff #103 #68)
-#117 := [rewrite]: #116
-#120 := [monotonicity #117]: #119
-#123 := [quant-intro #120]: #122
-#84 := (exists (vars (?v0 Int)) #81)
-#87 := (not #84)
-#111 := (~ #87 #110)
-#108 := (~ #107 #106)
-#104 := (~ #103 #103)
-#105 := [refl]: #104
-#94 := (not #78)
-#95 := (~ #94 #75)
-#100 := (~ #75 #75)
-#98 := (~ #72 #72)
-#99 := [refl]: #98
-#101 := [nnf-pos #99]: #100
-#102 := [nnf-neg #101]: #95
-#109 := [nnf-neg #102 #105]: #108
-#112 := [nnf-neg #109]: #111
-#11 := (< 0::Int #8)
-#9 := (<= #7 #8)
-#12 := (implies #9 #11)
-#13 := (forall (vars (?v1 Int)) #12)
-#14 := (implies #13 #11)
-#15 := (exists (vars (?v0 Int)) #14)
-#16 := (not #15)
-#90 := (iff #16 #87)
-#37 := (not #9)
-#38 := (or #37 #11)
-#41 := (forall (vars (?v1 Int)) #38)
-#47 := (not #41)
-#48 := (or #47 #11)
-#53 := (exists (vars (?v0 Int)) #48)
-#56 := (not #53)
-#88 := (iff #56 #87)
-#85 := (iff #53 #84)
-#82 := (iff #48 #81)
-#70 := (iff #11 #69)
-#71 := [rewrite]: #70
-#79 := (iff #47 #78)
-#76 := (iff #41 #75)
-#73 := (iff #38 #72)
-#66 := (iff #37 #65)
-#63 := (iff #9 #62)
-#64 := [rewrite]: #63
-#67 := [monotonicity #64]: #66
-#74 := [monotonicity #67 #71]: #73
-#77 := [quant-intro #74]: #76
-#80 := [monotonicity #77]: #79
-#83 := [monotonicity #80 #71]: #82
-#86 := [quant-intro #83]: #85
-#89 := [monotonicity #86]: #88
-#57 := (iff #16 #56)
-#54 := (iff #15 #53)
-#51 := (iff #14 #48)
-#44 := (implies #41 #11)
-#49 := (iff #44 #48)
-#50 := [rewrite]: #49
-#45 := (iff #14 #44)
-#42 := (iff #13 #41)
-#39 := (iff #12 #38)
-#40 := [rewrite]: #39
-#43 := [quant-intro #40]: #42
-#46 := [monotonicity #43]: #45
-#52 := [trans #46 #50]: #51
-#55 := [quant-intro #52]: #54
-#58 := [monotonicity #55]: #57
-#91 := [trans #58 #89]: #90
-#36 := [asserted]: #16
-#92 := [mp #36 #91]: #87
-#113 := [mp~ #92 #112]: #110
-#114 := [mp #113 #123]: #121
-#128 := [mp #114 #127]: #125
-#160 := [mp #128 #159]: #155
-[mp #160 #183]: false
-unsat
-ea961570b37add45bc63c8f0e3f6ddc653b28f42 67 0
-ERROR: line 11 column 83: invalid pattern.
-#2 := false
-decl f3 :: Int
-#8 := f3
-#10 := 2::Int
-#12 := (* 2::Int f3)
-#7 := (:var 0 Int)
-#11 := (* 2::Int #7)
-#13 := (< #11 #12)
-#9 := (< #7 f3)
-#14 := (implies #9 #13)
-#15 := (forall (vars (?v0 Int)) #14)
-#16 := (not #15)
-#85 := (iff #16 false)
-#38 := (not #9)
-#39 := (or #38 #13)
-#42 := (forall (vars (?v0 Int)) #39)
-#45 := (not #42)
-#83 := (iff #45 false)
-#1 := true
-#78 := (not true)
-#81 := (iff #78 false)
-#82 := [rewrite]: #81
-#79 := (iff #45 #78)
-#76 := (iff #42 true)
-#71 := (forall (vars (?v0 Int)) true)
-#74 := (iff #71 true)
-#75 := [elim-unused]: #74
-#72 := (iff #42 #71)
-#69 := (iff #39 true)
-#18 := 0::Int
-#48 := -1::Int
-#51 := (* -1::Int f3)
-#52 := (+ #7 #51)
-#50 := (>= #52 0::Int)
-#49 := (not #50)
-#63 := (or #50 #49)
-#67 := (iff #63 true)
-#68 := [rewrite]: #67
-#65 := (iff #39 #63)
-#64 := (iff #13 #49)
-#62 := [rewrite]: #64
-#60 := (iff #38 #50)
-#55 := (not #49)
-#58 := (iff #55 #50)
-#59 := [rewrite]: #58
-#56 := (iff #38 #55)
-#53 := (iff #9 #49)
-#54 := [rewrite]: #53
-#57 := [monotonicity #54]: #56
-#61 := [trans #57 #59]: #60
-#66 := [monotonicity #61 #62]: #65
-#70 := [trans #66 #68]: #69
-#73 := [quant-intro #70]: #72
-#77 := [trans #73 #75]: #76
-#80 := [monotonicity #77]: #79
-#84 := [trans #80 #82]: #83
-#46 := (iff #16 #45)
-#43 := (iff #15 #42)
-#40 := (iff #14 #39)
-#41 := [rewrite]: #40
-#44 := [quant-intro #41]: #43
-#47 := [monotonicity #44]: #46
-#86 := [trans #47 #84]: #85
-#37 := [asserted]: #16
-[mp #37 #86]: false
-unsat
-cc87973002902704adc7d85df3fb8affa4a44929 54 0
-#2 := false
-#10 := 1::Int
-decl ?v1!0 :: Int
-#66 := ?v1!0
-#69 := (>= ?v1!0 1::Int)
-#62 := (not #69)
-#7 := 0::Int
-#67 := (<= ?v1!0 0::Int)
-#68 := (not #67)
-#63 := (or #68 #62)
-#70 := (not #63)
-#8 := (:var 0 Int)
-#47 := (>= #8 1::Int)
-#45 := (not #47)
-#41 := (<= #8 0::Int)
-#42 := (not #41)
-#49 := (or #42 #45)
-#52 := (forall (vars (?v1 Int)) #49)
-#55 := (not #52)
-#71 := (~ #55 #70)
-#72 := [sk]: #71
-#11 := (< #8 1::Int)
-#9 := (< 0::Int #8)
-#12 := (or #9 #11)
-#13 := (forall (vars (?v0 Int) (?v1 Int)) #12)
-#14 := (not #13)
-#58 := (iff #14 #55)
-#35 := (forall (vars (?v1 Int)) #12)
-#38 := (not #35)
-#56 := (iff #38 #55)
-#53 := (iff #35 #52)
-#50 := (iff #12 #49)
-#46 := (iff #11 #45)
-#48 := [rewrite]: #46
-#43 := (iff #9 #42)
-#44 := [rewrite]: #43
-#51 := [monotonicity #44 #48]: #50
-#54 := [quant-intro #51]: #53
-#57 := [monotonicity #54]: #56
-#39 := (iff #14 #38)
-#36 := (iff #13 #35)
-#37 := [elim-unused]: #36
-#40 := [monotonicity #37]: #39
-#59 := [trans #40 #57]: #58
-#34 := [asserted]: #14
-#60 := [mp #34 #59]: #55
-#75 := [mp~ #60 #72]: #70
-#79 := [not-or-elim #75]: #69
-#78 := [not-or-elim #75]: #67
-#137 := (or #62 #68)
-#138 := [th-lemma arith farkas 1 1]: #137
-#139 := [unit-resolution #138 #78]: #62
-[unit-resolution #139 #79]: false
-unsat
-1d9e76ccce459de8771731a1c234c6d9e2aa3527 1 0
-unsat
-e46d82e75c1853418f786555dbc1a12ba5d54f6e 75 0
-#2 := false
-#9 := 1::Int
-decl f5 :: Int
-#11 := f5
-#15 := (+ f5 1::Int)
-decl f3 :: Int
-#7 := f3
-#16 := (* f3 #15)
-decl f4 :: Int
-#8 := f4
-#14 := (* f3 f4)
-#17 := (+ #14 #16)
-#10 := (+ f4 1::Int)
-#12 := (+ #10 f5)
-#13 := (* f3 #12)
-#18 := (= #13 #17)
-#19 := (not #18)
-#93 := (iff #19 false)
-#1 := true
-#88 := (not true)
-#91 := (iff #88 false)
-#92 := [rewrite]: #91
-#89 := (iff #19 #88)
-#86 := (iff #18 true)
-#56 := (* f3 f5)
-#57 := (+ #14 #56)
-#58 := (+ f3 #57)
-#81 := (= #58 #58)
-#84 := (iff #81 true)
-#85 := [rewrite]: #84
-#82 := (iff #18 #81)
-#79 := (= #17 #58)
-#69 := (+ f3 #56)
-#74 := (+ #14 #69)
-#77 := (= #74 #58)
-#78 := [rewrite]: #77
-#75 := (= #17 #74)
-#72 := (= #16 #69)
-#63 := (+ 1::Int f5)
-#66 := (* f3 #63)
-#70 := (= #66 #69)
-#71 := [rewrite]: #70
-#67 := (= #16 #66)
-#64 := (= #15 #63)
-#65 := [rewrite]: #64
-#68 := [monotonicity #65]: #67
-#73 := [trans #68 #71]: #72
-#76 := [monotonicity #73]: #75
-#80 := [trans #76 #78]: #79
-#61 := (= #13 #58)
-#47 := (+ f4 f5)
-#48 := (+ 1::Int #47)
-#53 := (* f3 #48)
-#59 := (= #53 #58)
-#60 := [rewrite]: #59
-#54 := (= #13 #53)
-#51 := (= #12 #48)
-#41 := (+ 1::Int f4)
-#44 := (+ #41 f5)
-#49 := (= #44 #48)
-#50 := [rewrite]: #49
-#45 := (= #12 #44)
-#42 := (= #10 #41)
-#43 := [rewrite]: #42
-#46 := [monotonicity #43]: #45
-#52 := [trans #46 #50]: #51
-#55 := [monotonicity #52]: #54
-#62 := [trans #55 #60]: #61
-#83 := [monotonicity #62 #80]: #82
-#87 := [trans #83 #85]: #86
-#90 := [monotonicity #87]: #89
-#94 := [trans #90 #92]: #93
-#40 := [asserted]: #19
-[mp #40 #94]: false
-unsat
-60242f59c15f3933ccbd1d4ed5e4e07293c9dd72 62 0
-#2 := false
-decl f4 :: Real
-#9 := f4
-decl f3 :: Real
-#7 := f3
-#15 := 2::Real
-#16 := (* 2::Real f3)
-#17 := (* #16 f4)
-#8 := 1::Real
-#12 := (- 1::Real f4)
-#13 := (* f3 #12)
-#10 := (+ 1::Real f4)
-#11 := (* f3 #10)
-#14 := (- #11 #13)
-#18 := (= #14 #17)
-#19 := (not #18)
-#81 := (iff #19 false)
-#1 := true
-#76 := (not true)
-#79 := (iff #76 false)
-#80 := [rewrite]: #79
-#77 := (iff #19 #76)
-#74 := (iff #18 true)
-#41 := (* f3 f4)
-#63 := (* 2::Real #41)
-#69 := (= #63 #63)
-#72 := (iff #69 true)
-#73 := [rewrite]: #72
-#70 := (iff #18 #69)
-#67 := (= #17 #63)
-#68 := [rewrite]: #67
-#65 := (= #14 #63)
-#45 := -1::Real
-#53 := (* -1::Real #41)
-#54 := (+ f3 #53)
-#42 := (+ f3 #41)
-#59 := (- #42 #54)
-#62 := (= #59 #63)
-#64 := [rewrite]: #62
-#60 := (= #14 #59)
-#57 := (= #13 #54)
-#46 := (* -1::Real f4)
-#47 := (+ 1::Real #46)
-#50 := (* f3 #47)
-#55 := (= #50 #54)
-#56 := [rewrite]: #55
-#51 := (= #13 #50)
-#48 := (= #12 #47)
-#49 := [rewrite]: #48
-#52 := [monotonicity #49]: #51
-#58 := [trans #52 #56]: #57
-#43 := (= #11 #42)
-#44 := [rewrite]: #43
-#61 := [monotonicity #44 #58]: #60
-#66 := [trans #61 #64]: #65
-#71 := [monotonicity #66 #68]: #70
-#75 := [trans #71 #73]: #74
-#78 := [monotonicity #75]: #77
-#82 := [trans #78 #80]: #81
-#40 := [asserted]: #19
-[mp #40 #82]: false
-unsat
-3ecab0bc7101d63e72b4fb9ac8a649c491da9533 141 0
-#2 := false
-decl f6 :: Int
-#12 := f6
-decl f7 :: Int
-#16 := f7
-decl f5 :: Int
-#11 := f5
-#27 := (+ f5 f7)
-#28 := (+ #27 f6)
-decl f4 :: Int
-#9 := f4
-#8 := 1::Int
-#10 := (+ 1::Int f4)
-#29 := (* #10 #28)
-#24 := (* f7 f4)
-#22 := (* #10 f7)
-#13 := (+ f5 f6)
-#19 := 2::Int
-#20 := (* 2::Int #10)
-#21 := (* #20 #13)
-#23 := (+ #21 #22)
-#25 := (+ #23 #24)
-decl f3 :: Int
-#7 := f3
-#26 := (+ f3 #25)
-#30 := (- #26 #29)
-#17 := (* f4 f7)
-#14 := (* #10 #13)
-#15 := (+ f3 #14)
-#18 := (+ #15 #17)
-#31 := (= #18 #30)
-#32 := (not #31)
-#157 := (iff #32 false)
-#1 := true
-#152 := (not true)
-#155 := (iff #152 false)
-#156 := [rewrite]: #155
-#153 := (iff #32 #152)
-#150 := (iff #31 true)
-#55 := (* f4 f6)
-#54 := (* f4 f5)
-#56 := (+ #54 #55)
-#67 := (+ #17 #56)
-#68 := (+ f6 #67)
-#69 := (+ f5 #68)
-#70 := (+ f3 #69)
-#144 := (= #70 #70)
-#148 := (iff #144 true)
-#149 := [rewrite]: #148
-#143 := (iff #31 #144)
-#146 := (= #30 #70)
-#131 := (+ f7 #67)
-#132 := (+ f6 #131)
-#133 := (+ f5 #132)
-#85 := (* 2::Int #55)
-#83 := (* 2::Int #54)
-#86 := (+ #83 #85)
-#112 := (* 2::Int #17)
-#113 := (+ #112 #86)
-#114 := (+ f7 #113)
-#84 := (* 2::Int f6)
-#115 := (+ #84 #114)
-#82 := (* 2::Int f5)
-#116 := (+ #82 #115)
-#121 := (+ f3 #116)
-#138 := (- #121 #133)
-#141 := (= #138 #70)
-#147 := [rewrite]: #141
-#139 := (= #30 #138)
-#136 := (= #29 #133)
-#124 := (+ f6 f7)
-#125 := (+ f5 #124)
-#128 := (* #10 #125)
-#134 := (= #128 #133)
-#135 := [rewrite]: #134
-#129 := (= #29 #128)
-#126 := (= #28 #125)
-#127 := [rewrite]: #126
-#130 := [monotonicity #127]: #129
-#137 := [trans #130 #135]: #136
-#122 := (= #26 #121)
-#119 := (= #25 #116)
-#99 := (+ #17 #86)
-#100 := (+ f7 #99)
-#101 := (+ #84 #100)
-#102 := (+ #82 #101)
-#109 := (+ #102 #17)
-#117 := (= #109 #116)
-#118 := [rewrite]: #117
-#110 := (= #25 #109)
-#107 := (= #24 #17)
-#108 := [rewrite]: #107
-#105 := (= #23 #102)
-#93 := (+ f7 #17)
-#87 := (+ #84 #86)
-#88 := (+ #82 #87)
-#96 := (+ #88 #93)
-#103 := (= #96 #102)
-#104 := [rewrite]: #103
-#97 := (= #23 #96)
-#94 := (= #22 #93)
-#95 := [rewrite]: #94
-#91 := (= #21 #88)
-#75 := (* 2::Int f4)
-#76 := (+ 2::Int #75)
-#79 := (* #76 #13)
-#89 := (= #79 #88)
-#90 := [rewrite]: #89
-#80 := (= #21 #79)
-#77 := (= #20 #76)
-#78 := [rewrite]: #77
-#81 := [monotonicity #78]: #80
-#92 := [trans #81 #90]: #91
-#98 := [monotonicity #92 #95]: #97
-#106 := [trans #98 #104]: #105
-#111 := [monotonicity #106 #108]: #110
-#120 := [trans #111 #118]: #119
-#123 := [monotonicity #120]: #122
-#140 := [monotonicity #123 #137]: #139
-#145 := [trans #140 #147]: #146
-#73 := (= #18 #70)
-#57 := (+ f6 #56)
-#58 := (+ f5 #57)
-#61 := (+ f3 #58)
-#64 := (+ #61 #17)
-#71 := (= #64 #70)
-#72 := [rewrite]: #71
-#65 := (= #18 #64)
-#62 := (= #15 #61)
-#59 := (= #14 #58)
-#60 := [rewrite]: #59
-#63 := [monotonicity #60]: #62
-#66 := [monotonicity #63]: #65
-#74 := [trans #66 #72]: #73
-#142 := [monotonicity #74 #145]: #143
-#151 := [trans #142 #149]: #150
-#154 := [monotonicity #151]: #153
-#158 := [trans #154 #156]: #157
-#53 := [asserted]: #32
-[mp #53 #158]: false
-unsat
 43550507f510d81bc4fb9ef8c1fd14424eaa9070 37 0
 #2 := false
 #10 := 0::Int
@@ -8819,1764 +1561,6 @@
 #53 := [not-or-elim #52]: #11
 [th-lemma arith farkas 1 1 1 #53 #57 #55]: false
 unsat
-f5067fa58c623377db978838e2294684a3fe7bb2 225 0
-#2 := false
-#24 := 0::Int
-decl f5 :: (-> S4 S3 Int)
-decl f3 :: (-> S2 Int S3)
-decl f7 :: S3
-#10 := f7
-decl f6 :: S4
-#9 := f6
-#11 := (f5 f6 f7)
-#8 := 2::Int
-#12 := (* 2::Int #11)
-decl f4 :: S2
-#7 := f4
-#13 := (f3 f4 #12)
-#276 := (f5 f6 #13)
-#185 := -1::Int
-#596 := (* -1::Int #276)
-#597 := (+ #12 #596)
-#577 := (<= #597 0::Int)
-#595 := (= #597 0::Int)
-#256 := (>= #11 0::Int)
-#579 := (= #276 0::Int)
-#436 := (not #579)
-#297 := (<= #276 0::Int)
-#533 := (not #297)
-#14 := 1::Int
-#544 := (>= #276 1::Int)
-#549 := (= #276 1::Int)
-#15 := (f3 f4 1::Int)
-#569 := (f5 f6 #15)
-#570 := (= #569 1::Int)
-#25 := (:var 0 Int)
-#27 := (f3 f4 #25)
-#607 := (pattern #27)
-#28 := (f5 f6 #27)
-#29 := (= #28 #25)
-#70 := (>= #25 0::Int)
-#71 := (not #70)
-#74 := (or #71 #29)
-#608 := (forall (vars (?v0 Int)) (:pat #607) #74)
-#77 := (forall (vars (?v0 Int)) #74)
-#611 := (iff #77 #608)
-#609 := (iff #74 #74)
-#610 := [refl]: #609
-#612 := [quant-intro #610]: #611
-#114 := (~ #77 #77)
-#113 := (~ #74 #74)
-#110 := [refl]: #113
-#115 := [nnf-pos #110]: #114
-#26 := (<= 0::Int #25)
-#30 := (implies #26 #29)
-#31 := (forall (vars (?v0 Int)) #30)
-#80 := (iff #31 #77)
-#61 := (not #26)
-#62 := (or #61 #29)
-#65 := (forall (vars (?v0 Int)) #62)
-#78 := (iff #65 #77)
-#75 := (iff #62 #74)
-#72 := (iff #61 #71)
-#68 := (iff #26 #70)
-#69 := [rewrite]: #68
-#73 := [monotonicity #69]: #72
-#76 := [monotonicity #73]: #75
-#79 := [quant-intro #76]: #78
-#66 := (iff #31 #65)
-#63 := (iff #30 #62)
-#64 := [rewrite]: #63
-#67 := [quant-intro #64]: #66
-#81 := [trans #67 #79]: #80
-#59 := [asserted]: #31
-#82 := [mp #59 #81]: #77
-#111 := [mp~ #82 #115]: #77
-#613 := [mp #111 #612]: #608
-#589 := (not #608)
-#555 := (or #589 #570)
-#299 := (>= 1::Int 0::Int)
-#192 := (not #299)
-#292 := (or #192 #570)
-#556 := (or #589 #292)
-#552 := (iff #556 #555)
-#558 := (iff #555 #555)
-#559 := [rewrite]: #558
-#562 := (iff #292 #570)
-#563 := (or false #570)
-#561 := (iff #563 #570)
-#565 := [rewrite]: #561
-#564 := (iff #292 #563)
-#284 := (iff #192 false)
-#1 := true
-#571 := (not true)
-#282 := (iff #571 false)
-#283 := [rewrite]: #282
-#568 := (iff #192 #571)
-#293 := (iff #299 true)
-#567 := [rewrite]: #293
-#572 := [monotonicity #567]: #568
-#285 := [trans #572 #283]: #284
-#278 := [monotonicity #285]: #564
-#566 := [trans #278 #565]: #562
-#553 := [monotonicity #566]: #552
-#554 := [trans #553 #559]: #552
-#557 := [quant-inst #14]: #556
-#560 := [mp #557 #554]: #555
-#383 := [unit-resolution #560 #613]: #570
-#536 := (= #276 #569)
-#16 := (= #13 #15)
-#17 := (not #16)
-#18 := (not #17)
-#56 := (iff #18 #16)
-#57 := [rewrite]: #56
-#55 := [asserted]: #18
-#60 := [mp #55 #57]: #16
-#424 := [monotonicity #60]: #536
-#425 := [trans #424 #383]: #549
-#384 := (not #549)
-#532 := (or #384 #544)
-#434 := [th-lemma arith triangle-eq]: #532
-#529 := [unit-resolution #434 #425]: #544
-#530 := (not #544)
-#418 := (or #530 #533)
-#433 := [th-lemma arith farkas 1 1]: #418
-#435 := [unit-resolution #433 #529]: #533
-#429 := (or #436 #297)
-#437 := [th-lemma arith triangle-eq]: #429
-#438 := [unit-resolution #437 #435]: #436
-#581 := (or #256 #579)
-#33 := (= #28 0::Int)
-#100 := (or #70 #33)
-#614 := (forall (vars (?v0 Int)) (:pat #607) #100)
-#103 := (forall (vars (?v0 Int)) #100)
-#617 := (iff #103 #614)
-#615 := (iff #100 #100)
-#616 := [refl]: #615
-#618 := [quant-intro #616]: #617
-#116 := (~ #103 #103)
-#124 := (~ #100 #100)
-#125 := [refl]: #124
-#117 := [nnf-pos #125]: #116
-#32 := (< #25 0::Int)
-#34 := (implies #32 #33)
-#35 := (forall (vars (?v0 Int)) #34)
-#106 := (iff #35 #103)
-#84 := (not #32)
-#85 := (or #84 #33)
-#88 := (forall (vars (?v0 Int)) #85)
-#104 := (iff #88 #103)
-#101 := (iff #85 #100)
-#98 := (iff #84 #70)
-#93 := (not #71)
-#96 := (iff #93 #70)
-#97 := [rewrite]: #96
-#94 := (iff #84 #93)
-#91 := (iff #32 #71)
-#92 := [rewrite]: #91
-#95 := [monotonicity #92]: #94
-#99 := [trans #95 #97]: #98
-#102 := [monotonicity #99]: #101
-#105 := [quant-intro #102]: #104
-#89 := (iff #35 #88)
-#86 := (iff #34 #85)
-#87 := [rewrite]: #86
-#90 := [quant-intro #87]: #89
-#107 := [trans #90 #105]: #106
-#83 := [asserted]: #35
-#108 := [mp #83 #107]: #103
-#126 := [mp~ #108 #117]: #103
-#619 := [mp #126 #618]: #614
-#219 := (not #614)
-#583 := (or #219 #256 #579)
-#271 := (>= #12 0::Int)
-#580 := (or #271 #579)
-#585 := (or #219 #580)
-#574 := (iff #585 #583)
-#225 := (or #219 #581)
-#587 := (iff #225 #583)
-#573 := [rewrite]: #587
-#586 := (iff #585 #225)
-#576 := (iff #580 #581)
-#592 := (iff #271 #256)
-#594 := [rewrite]: #592
-#582 := [monotonicity #594]: #576
-#584 := [monotonicity #582]: #586
-#281 := [trans #584 #573]: #574
-#224 := [quant-inst #12]: #585
-#296 := [mp #224 #281]: #583
-#439 := [unit-resolution #296 #619]: #581
-#440 := [unit-resolution #439 #438]: #256
-#250 := (not #256)
-#598 := (or #250 #595)
-#248 := (or #589 #250 #595)
-#273 := (= #276 #12)
-#272 := (not #271)
-#277 := (or #272 #273)
-#253 := (or #589 #277)
-#238 := (iff #253 #248)
-#249 := (or #589 #598)
-#575 := (iff #249 #248)
-#237 := [rewrite]: #575
-#591 := (iff #253 #249)
-#593 := (iff #277 #598)
-#261 := (iff #273 #595)
-#262 := [rewrite]: #261
-#381 := (iff #272 #250)
-#588 := [monotonicity #594]: #381
-#599 := [monotonicity #588 #262]: #593
-#233 := [monotonicity #599]: #591
-#239 := [trans #233 #237]: #238
-#590 := [quant-inst #12]: #253
-#240 := [mp #590 #239]: #248
-#441 := [unit-resolution #240 #613]: #598
-#534 := [unit-resolution #441 #440]: #595
-#531 := (not #595)
-#535 := (or #531 #577)
-#522 := [th-lemma arith triangle-eq]: #535
-#524 := [unit-resolution #522 #534]: #577
-#578 := (>= #597 0::Int)
-#516 := (or #531 #578)
-#513 := [th-lemma arith triangle-eq]: #516
-#515 := [unit-resolution #513 #534]: #578
-#550 := (<= #276 1::Int)
-#525 := (or #384 #550)
-#526 := [th-lemma arith triangle-eq]: #525
-#527 := [unit-resolution #526 #425]: #550
-[th-lemma arith gcd-test -1/2 -1/2 -1/2 -1/2 #529 #527 #515 #524]: false
-unsat
-4225ab6372dca8ebf6ba05ad5ea39526a6e2a129 55 0
-#2 := false
-#74 := 4::Int
-decl f3 :: (-> S2 S3 Int)
-decl f5 :: S3
-#8 := f5
-decl f4 :: S2
-#7 := f4
-#9 := (f3 f4 f5)
-#75 := (>= #9 4::Int)
-#76 := (not #75)
-#10 := 3::Int
-#65 := (>= #9 3::Int)
-#79 := (or #65 #76)
-#82 := (not #79)
-#14 := 7::Int
-#12 := 2::Int
-#13 := (* 2::Int #9)
-#15 := (< #13 7::Int)
-#11 := (< #9 3::Int)
-#16 := (implies #11 #15)
-#17 := (not #16)
-#85 := (iff #17 #82)
-#56 := (not #11)
-#57 := (or #56 #15)
-#60 := (not #57)
-#83 := (iff #60 #82)
-#80 := (iff #57 #79)
-#77 := (iff #15 #76)
-#78 := [rewrite]: #77
-#72 := (iff #56 #65)
-#63 := (not #65)
-#67 := (not #63)
-#70 := (iff #67 #65)
-#71 := [rewrite]: #70
-#68 := (iff #56 #67)
-#64 := (iff #11 #63)
-#66 := [rewrite]: #64
-#69 := [monotonicity #66]: #68
-#73 := [trans #69 #71]: #72
-#81 := [monotonicity #73 #78]: #80
-#84 := [monotonicity #81]: #83
-#61 := (iff #17 #60)
-#58 := (iff #16 #57)
-#59 := [rewrite]: #58
-#62 := [monotonicity #59]: #61
-#86 := [trans #62 #84]: #85
-#55 := [asserted]: #17
-#87 := [mp #55 #86]: #82
-#89 := [not-or-elim #87]: #75
-#88 := [not-or-elim #87]: #63
-#300 := (or #76 #65)
-#216 := [th-lemma arith farkas 1 1]: #300
-#301 := [unit-resolution #216 #88]: #76
-[unit-resolution #301 #89]: false
-unsat
-6b3381ed26844d4b649300d18bdcc49988752527 270 0
-#2 := false
-#7 := 0::Int
-decl f3 :: (-> S2 S3 Int)
-decl f5 :: (-> S4 Int S3)
-decl f7 :: S3
-#11 := f7
-decl f4 :: S2
-#8 := f4
-#12 := (f3 f4 f7)
-#10 := 1::Int
-#13 := (+ 1::Int #12)
-decl f6 :: S4
-#9 := f6
-#14 := (f5 f6 #13)
-#15 := (f3 f4 #14)
-#60 := -1::Int
-#61 := (* -1::Int #12)
-#62 := (+ #61 #15)
-#65 := (f5 f6 #62)
-#68 := (f3 f4 #65)
-#625 := (* -1::Int #15)
-#593 := (+ #625 #68)
-#597 := (+ #12 #593)
-#574 := (>= #597 0::Int)
-#594 := (= #597 0::Int)
-#631 := (+ #12 #625)
-#315 := (<= #631 0::Int)
-#614 := (<= #631 -1::Int)
-#621 := (= #631 -1::Int)
-#294 := (>= #12 -1::Int)
-#416 := (>= #12 0::Int)
-#545 := (= #12 0::Int)
-#218 := (f5 f6 #12)
-#564 := (f3 f4 #218)
-#466 := (= #564 0::Int)
-#550 := (not #416)
-#551 := [hypothesis]: #550
-#561 := (or #416 #466)
-#27 := (:var 0 Int)
-#29 := (f5 f6 #27)
-#639 := (pattern #29)
-#30 := (f3 f4 #29)
-#35 := (= #30 0::Int)
-#101 := (>= #27 0::Int)
-#132 := (or #101 #35)
-#646 := (forall (vars (?v0 Int)) (:pat #639) #132)
-#135 := (forall (vars (?v0 Int)) #132)
-#649 := (iff #135 #646)
-#647 := (iff #132 #132)
-#648 := [refl]: #647
-#650 := [quant-intro #648]: #649
-#148 := (~ #135 #135)
-#156 := (~ #132 #132)
-#157 := [refl]: #156
-#149 := [nnf-pos #157]: #148
-#34 := (< #27 0::Int)
-#36 := (implies #34 #35)
-#37 := (forall (vars (?v0 Int)) #36)
-#138 := (iff #37 #135)
-#116 := (not #34)
-#117 := (or #116 #35)
-#120 := (forall (vars (?v0 Int)) #117)
-#136 := (iff #120 #135)
-#133 := (iff #117 #132)
-#130 := (iff #116 #101)
-#103 := (not #101)
-#125 := (not #103)
-#128 := (iff #125 #101)
-#129 := [rewrite]: #128
-#126 := (iff #116 #125)
-#123 := (iff #34 #103)
-#124 := [rewrite]: #123
-#127 := [monotonicity #124]: #126
-#131 := [trans #127 #129]: #130
-#134 := [monotonicity #131]: #133
-#137 := [quant-intro #134]: #136
-#121 := (iff #37 #120)
-#118 := (iff #36 #117)
-#119 := [rewrite]: #118
-#122 := [quant-intro #119]: #121
-#139 := [trans #122 #137]: #138
-#115 := [asserted]: #37
-#140 := [mp #115 #139]: #135
-#158 := [mp~ #140 #149]: #135
-#651 := [mp #158 #650]: #646
-#616 := (not #646)
-#450 := (or #616 #416 #466)
-#465 := (or #616 #561)
-#468 := (iff #465 #450)
-#461 := [rewrite]: #468
-#467 := [quant-inst #12]: #465
-#469 := [mp #467 #461]: #450
-#552 := [unit-resolution #469 #651]: #561
-#546 := [unit-resolution #552 #551]: #466
-#540 := (= #12 #564)
-#537 := (= f7 #218)
-#303 := (= #218 f7)
-#22 := (:var 0 S3)
-#23 := (f3 f4 #22)
-#632 := (pattern #23)
-#24 := (f5 f6 #23)
-#25 := (= #24 #22)
-#633 := (forall (vars (?v0 S3)) (:pat #632) #25)
-#26 := (forall (vars (?v0 S3)) #25)
-#636 := (iff #26 #633)
-#634 := (iff #25 #25)
-#635 := [refl]: #634
-#637 := [quant-intro #635]: #636
-#154 := (~ #26 #26)
-#152 := (~ #25 #25)
-#153 := [refl]: #152
-#155 := [nnf-pos #153]: #154
-#91 := [asserted]: #26
-#144 := [mp~ #91 #155]: #26
-#638 := [mp #144 #637]: #633
-#305 := (not #633)
-#296 := (or #305 #303)
-#307 := [quant-inst #11]: #296
-#553 := [unit-resolution #307 #638]: #303
-#538 := [symm #553]: #537
-#541 := [monotonicity #538]: #540
-#542 := [trans #541 #546]: #545
-#543 := (not #545)
-#539 := (or #543 #416)
-#544 := [th-lemma arith triangle-eq]: #539
-#530 := [unit-resolution #544 #551 #542]: false
-#531 := [lemma #530]: #416
-#547 := (or #550 #294)
-#533 := [th-lemma arith farkas 1 1]: #547
-#534 := [unit-resolution #533 #531]: #294
-#628 := (not #294)
-#622 := (or #628 #621)
-#31 := (= #30 #27)
-#106 := (or #103 #31)
-#640 := (forall (vars (?v0 Int)) (:pat #639) #106)
-#109 := (forall (vars (?v0 Int)) #106)
-#643 := (iff #109 #640)
-#641 := (iff #106 #106)
-#642 := [refl]: #641
-#644 := [quant-intro #642]: #643
-#146 := (~ #109 #109)
-#145 := (~ #106 #106)
-#142 := [refl]: #145
-#147 := [nnf-pos #142]: #146
-#28 := (<= 0::Int #27)
-#32 := (implies #28 #31)
-#33 := (forall (vars (?v0 Int)) #32)
-#112 := (iff #33 #109)
-#93 := (not #28)
-#94 := (or #93 #31)
-#97 := (forall (vars (?v0 Int)) #94)
-#110 := (iff #97 #109)
-#107 := (iff #94 #106)
-#104 := (iff #93 #103)
-#100 := (iff #28 #101)
-#102 := [rewrite]: #100
-#105 := [monotonicity #102]: #104
-#108 := [monotonicity #105]: #107
-#111 := [quant-intro #108]: #110
-#98 := (iff #33 #97)
-#95 := (iff #32 #94)
-#96 := [rewrite]: #95
-#99 := [quant-intro #96]: #98
-#113 := [trans #99 #111]: #112
-#92 := [asserted]: #33
-#114 := [mp #92 #113]: #109
-#143 := [mp~ #114 #147]: #109
-#645 := [mp #143 #644]: #640
-#266 := (not #640)
-#607 := (or #266 #628 #621)
-#413 := (= #15 #13)
-#289 := (>= #13 0::Int)
-#624 := (not #289)
-#620 := (or #624 #413)
-#270 := (or #266 #620)
-#612 := (iff #270 #607)
-#272 := (or #266 #622)
-#610 := (iff #272 #607)
-#611 := [rewrite]: #610
-#273 := (iff #270 #272)
-#282 := (iff #620 #622)
-#281 := (iff #413 #621)
-#286 := [rewrite]: #281
-#629 := (iff #624 #628)
-#295 := (iff #289 #294)
-#627 := [rewrite]: #295
-#630 := [monotonicity #627]: #629
-#623 := [monotonicity #630 #286]: #282
-#609 := [monotonicity #623]: #273
-#613 := [trans #609 #611]: #612
-#271 := [quant-inst #13]: #270
-#608 := [mp #271 #613]: #607
-#535 := [unit-resolution #608 #645]: #622
-#532 := [unit-resolution #535 #534]: #621
-#536 := (not #621)
-#516 := (or #536 #614)
-#517 := [th-lemma arith triangle-eq]: #516
-#519 := [unit-resolution #517 #532]: #614
-#520 := (not #614)
-#521 := (or #520 #315)
-#522 := [th-lemma arith farkas 1 1]: #521
-#523 := [unit-resolution #522 #519]: #315
-#595 := (not #315)
-#588 := (or #595 #594)
-#585 := (or #266 #595 #594)
-#604 := (= #68 #62)
-#603 := (>= #62 0::Int)
-#600 := (not #603)
-#314 := (or #600 #604)
-#590 := (or #266 #314)
-#577 := (iff #590 #585)
-#586 := (or #266 #588)
-#434 := (iff #586 #585)
-#435 := [rewrite]: #434
-#592 := (iff #590 #586)
-#589 := (iff #314 #588)
-#598 := (iff #604 #594)
-#587 := [rewrite]: #598
-#596 := (iff #600 #595)
-#316 := (iff #603 #315)
-#317 := [rewrite]: #316
-#311 := [monotonicity #317]: #596
-#584 := [monotonicity #311 #587]: #589
-#433 := [monotonicity #584]: #592
-#578 := [trans #433 #435]: #577
-#591 := [quant-inst #62]: #590
-#579 := [mp #591 #578]: #585
-#524 := [unit-resolution #579 #645]: #588
-#525 := [unit-resolution #524 #523]: #594
-#526 := (not #594)
-#527 := (or #526 #574)
-#528 := [th-lemma arith triangle-eq]: #527
-#518 := [unit-resolution #528 #525]: #574
-#77 := (<= #68 0::Int)
-#17 := (- #15 #12)
-#18 := (f5 f6 #17)
-#19 := (f3 f4 #18)
-#16 := (* 0::Int #15)
-#20 := (< #16 #19)
-#21 := (not #20)
-#88 := (iff #21 #77)
-#71 := (< 0::Int #68)
-#74 := (not #71)
-#86 := (iff #74 #77)
-#78 := (not #77)
-#81 := (not #78)
-#84 := (iff #81 #77)
-#85 := [rewrite]: #84
-#82 := (iff #74 #81)
-#79 := (iff #71 #78)
-#80 := [rewrite]: #79
-#83 := [monotonicity #80]: #82
-#87 := [trans #83 #85]: #86
-#75 := (iff #21 #74)
-#72 := (iff #20 #71)
-#69 := (= #19 #68)
-#66 := (= #18 #65)
-#63 := (= #17 #62)
-#64 := [rewrite]: #63
-#67 := [monotonicity #64]: #66
-#70 := [monotonicity #67]: #69
-#58 := (= #16 0::Int)
-#59 := [rewrite]: #58
-#73 := [monotonicity #59 #70]: #72
-#76 := [monotonicity #73]: #75
-#89 := [trans #76 #87]: #88
-#57 := [asserted]: #21
-#90 := [mp #57 #89]: #77
-[th-lemma arith farkas -1 -1 1 #90 #519 #518]: false
-unsat
-b3acce989065928cb3ce15ce4113a910c6fff5aa 269 0
-#2 := false
-#7 := 0::Int
-decl f3 :: (-> S2 S3 Int)
-decl f5 :: (-> S4 Int S3)
-decl f7 :: S3
-#11 := f7
-decl f4 :: S2
-#8 := f4
-#12 := (f3 f4 f7)
-#10 := 1::Int
-#13 := (+ 1::Int #12)
-decl f6 :: S4
-#9 := f6
-#14 := (f5 f6 #13)
-#15 := (f3 f4 #14)
-#65 := -1::Int
-#66 := (+ -1::Int #15)
-#69 := (f5 f6 #66)
-#367 := (f3 f4 #69)
-#638 := (* -1::Int #367)
-#499 := (+ #12 #638)
-#459 := (>= #499 0::Int)
-#498 := (= #12 #367)
-#605 := (= f7 #69)
-#72 := (= #69 f7)
-#101 := (<= #15 0::Int)
-#173 := (iff #101 #72)
-#192 := (iff #173 #72)
-#1 := true
-#187 := (iff true #72)
-#190 := (iff #187 #72)
-#191 := [rewrite]: #190
-#188 := (iff #173 #187)
-#179 := (iff #101 true)
-#102 := (not #101)
-#105 := (iff #102 #72)
-#108 := (or #105 #102)
-#111 := (not #108)
-#16 := (< 0::Int #15)
-#17 := (if #16 true false)
-#22 := (not #17)
-#23 := (implies #22 false)
-#18 := (- #15 1::Int)
-#19 := (f5 f6 #18)
-#20 := (= #19 f7)
-#21 := (iff #17 #20)
-#24 := (or #21 #23)
-#25 := (or false #24)
-#26 := (not #25)
-#114 := (iff #26 #111)
-#75 := (iff #16 #72)
-#88 := (or #75 #16)
-#98 := (not #88)
-#112 := (iff #98 #111)
-#109 := (iff #88 #108)
-#103 := (iff #16 #102)
-#104 := [rewrite]: #103
-#106 := (iff #75 #105)
-#107 := [monotonicity #104]: #106
-#110 := [monotonicity #107 #104]: #109
-#113 := [monotonicity #110]: #112
-#99 := (iff #26 #98)
-#96 := (iff #25 #88)
-#91 := (or false #88)
-#94 := (iff #91 #88)
-#95 := [rewrite]: #94
-#92 := (iff #25 #91)
-#89 := (iff #24 #88)
-#86 := (iff #23 #16)
-#78 := (not #16)
-#81 := (implies #78 false)
-#84 := (iff #81 #16)
-#85 := [rewrite]: #84
-#82 := (iff #23 #81)
-#79 := (iff #22 #78)
-#63 := (iff #17 #16)
-#64 := [rewrite]: #63
-#80 := [monotonicity #64]: #79
-#83 := [monotonicity #80]: #82
-#87 := [trans #83 #85]: #86
-#76 := (iff #21 #75)
-#73 := (iff #20 #72)
-#70 := (= #19 #69)
-#67 := (= #18 #66)
-#68 := [rewrite]: #67
-#71 := [monotonicity #68]: #70
-#74 := [monotonicity #71]: #73
-#77 := [monotonicity #64 #74]: #76
-#90 := [monotonicity #77 #87]: #89
-#93 := [monotonicity #90]: #92
-#97 := [trans #93 #95]: #96
-#100 := [monotonicity #97]: #99
-#115 := [trans #100 #113]: #114
-#62 := [asserted]: #26
-#116 := [mp #62 #115]: #111
-#119 := [not-or-elim #116]: #101
-#180 := [iff-true #119]: #179
-#189 := [monotonicity #180]: #188
-#193 := [trans #189 #191]: #192
-#117 := (not #105)
-#174 := (iff #117 #173)
-#175 := [rewrite]: #174
-#118 := [not-or-elim #116]: #117
-#176 := [mp #118 #175]: #173
-#177 := [mp #176 #193]: #72
-#608 := [symm #177]: #605
-#513 := [monotonicity #608]: #498
-#514 := (not #498)
-#515 := (or #514 #459)
-#516 := [th-lemma arith triangle-eq]: #515
-#609 := [unit-resolution #516 #513]: #459
-#672 := (* -1::Int #15)
-#673 := (+ #12 #672)
-#654 := (<= #673 -1::Int)
-#671 := (= #673 -1::Int)
-#669 := (>= #12 -1::Int)
-#616 := (>= #367 0::Int)
-#621 := (= #367 0::Int)
-#646 := (>= #15 1::Int)
-#357 := (not #646)
-#606 := (or #357 #102)
-#610 := [th-lemma arith farkas 1 1]: #606
-#597 := [unit-resolution #610 #119]: #357
-#32 := (:var 0 Int)
-#34 := (f5 f6 #32)
-#682 := (pattern #34)
-#35 := (f3 f4 #34)
-#40 := (= #35 0::Int)
-#130 := (>= #32 0::Int)
-#161 := (or #130 #40)
-#689 := (forall (vars (?v0 Int)) (:pat #682) #161)
-#164 := (forall (vars (?v0 Int)) #161)
-#692 := (iff #164 #689)
-#690 := (iff #161 #161)
-#691 := [refl]: #690
-#693 := [quant-intro #691]: #692
-#197 := (~ #164 #164)
-#195 := (~ #161 #161)
-#196 := [refl]: #195
-#198 := [nnf-pos #196]: #197
-#39 := (< #32 0::Int)
-#41 := (implies #39 #40)
-#42 := (forall (vars (?v0 Int)) #41)
-#167 := (iff #42 #164)
-#145 := (not #39)
-#146 := (or #145 #40)
-#149 := (forall (vars (?v0 Int)) #146)
-#165 := (iff #149 #164)
-#162 := (iff #146 #161)
-#159 := (iff #145 #130)
-#132 := (not #130)
-#154 := (not #132)
-#157 := (iff #154 #130)
-#158 := [rewrite]: #157
-#155 := (iff #145 #154)
-#152 := (iff #39 #132)
-#153 := [rewrite]: #152
-#156 := [monotonicity #153]: #155
-#160 := [trans #156 #158]: #159
-#163 := [monotonicity #160]: #162
-#166 := [quant-intro #163]: #165
-#150 := (iff #42 #149)
-#147 := (iff #41 #146)
-#148 := [rewrite]: #147
-#151 := [quant-intro #148]: #150
-#168 := [trans #151 #166]: #167
-#144 := [asserted]: #42
-#169 := [mp #144 #168]: #164
-#199 := [mp~ #169 #198]: #164
-#694 := [mp #199 #693]: #689
-#660 := (not #689)
-#624 := (or #660 #646 #621)
-#644 := (>= #66 0::Int)
-#622 := (or #644 #621)
-#625 := (or #660 #622)
-#612 := (iff #625 #624)
-#623 := (or #646 #621)
-#626 := (or #660 #623)
-#458 := (iff #626 #624)
-#611 := [rewrite]: #458
-#455 := (iff #625 #626)
-#617 := (iff #622 #623)
-#643 := (iff #644 #646)
-#647 := [rewrite]: #643
-#618 := [monotonicity #647]: #617
-#457 := [monotonicity #618]: #455
-#614 := [trans #457 #611]: #612
-#619 := [quant-inst #66]: #625
-#615 := [mp #619 #614]: #624
-#599 := [unit-resolution #615 #694 #597]: #621
-#591 := (not #621)
-#588 := (or #591 #616)
-#590 := [th-lemma arith triangle-eq]: #588
-#600 := [unit-resolution #590 #599]: #616
-#602 := (not #459)
-#601 := (not #616)
-#598 := (or #669 #601 #602)
-#603 := [th-lemma arith assign-bounds 1 1]: #598
-#592 := [unit-resolution #603 #600 #609]: #669
-#663 := (not #669)
-#674 := (or #663 #671)
-#36 := (= #35 #32)
-#135 := (or #132 #36)
-#683 := (forall (vars (?v0 Int)) (:pat #682) #135)
-#138 := (forall (vars (?v0 Int)) #135)
-#686 := (iff #138 #683)
-#684 := (iff #135 #135)
-#685 := [refl]: #684
-#687 := [quant-intro #685]: #686
-#194 := (~ #138 #138)
-#182 := (~ #135 #135)
-#178 := [refl]: #182
-#171 := [nnf-pos #178]: #194
-#33 := (<= 0::Int #32)
-#37 := (implies #33 #36)
-#38 := (forall (vars (?v0 Int)) #37)
-#141 := (iff #38 #138)
-#122 := (not #33)
-#123 := (or #122 #36)
-#126 := (forall (vars (?v0 Int)) #123)
-#139 := (iff #126 #138)
-#136 := (iff #123 #135)
-#133 := (iff #122 #132)
-#129 := (iff #33 #130)
-#131 := [rewrite]: #129
-#134 := [monotonicity #131]: #133
-#137 := [monotonicity #134]: #136
-#140 := [quant-intro #137]: #139
-#127 := (iff #38 #126)
-#124 := (iff #37 #123)
-#125 := [rewrite]: #124
-#128 := [quant-intro #125]: #127
-#142 := [trans #128 #140]: #141
-#121 := [asserted]: #38
-#143 := [mp #121 #142]: #138
-#172 := [mp~ #143 #171]: #138
-#688 := [mp #172 #687]: #683
-#329 := (not #683)
-#665 := (or #329 #663 #671)
-#332 := (= #15 #13)
-#351 := (>= #13 0::Int)
-#352 := (not #351)
-#667 := (or #352 #332)
-#325 := (or #329 #667)
-#316 := (iff #325 #665)
-#309 := (or #329 #674)
-#314 := (iff #309 #665)
-#315 := [rewrite]: #314
-#650 := (iff #325 #309)
-#664 := (iff #667 #674)
-#670 := (iff #332 #671)
-#668 := [rewrite]: #670
-#337 := (iff #352 #663)
-#326 := (iff #351 #669)
-#456 := [rewrite]: #326
-#338 := [monotonicity #456]: #337
-#324 := [monotonicity #338 #668]: #664
-#313 := [monotonicity #324]: #650
-#652 := [trans #313 #315]: #316
-#666 := [quant-inst #13]: #325
-#653 := [mp #666 #652]: #665
-#593 := [unit-resolution #653 #688]: #674
-#594 := [unit-resolution #593 #592]: #671
-#595 := (not #671)
-#589 := (or #595 #654)
-#596 := [th-lemma arith triangle-eq]: #589
-#580 := [unit-resolution #596 #594]: #654
-[th-lemma arith farkas 1 -1 -1 1 #600 #119 #580 #609]: false
-unsat
-4f28f42d6f2b6fbb94a4ff1e55f0a807d8afe0f8 147 0
-#2 := false
-#10 := 0::Int
-decl f7 :: Int
-#9 := f7
-#54 := -1::Int
-#55 := (* -1::Int f7)
-#73 := (>= f7 0::Int)
-#80 := (if #73 f7 #55)
-#617 := (* -1::Int #80)
-#282 := (+ #55 #617)
-#625 := (<= #282 0::Int)
-#313 := (= #55 #80)
-#74 := (not #73)
-#280 := (+ f7 #617)
-#281 := (<= #280 0::Int)
-#228 := (= f7 #80)
-#283 := [hypothesis]: #73
-#229 := (or #74 #228)
-#314 := [def-axiom]: #229
-#619 := [unit-resolution #314 #283]: #228
-#620 := (not #228)
-#621 := (or #620 #281)
-#622 := [th-lemma arith triangle-eq]: #621
-#623 := [unit-resolution #622 #619]: #281
-#319 := (>= #80 0::Int)
-#316 := (not #319)
-decl f5 :: (-> S4 Int S3)
-#23 := (:var 0 Int)
-decl f6 :: S4
-#8 := f6
-#25 := (f5 f6 #23)
-#649 := (pattern #25)
-decl f3 :: (-> S2 S3 Int)
-decl f4 :: S2
-#7 := f4
-#26 := (f3 f4 #25)
-#27 := (= #26 #23)
-#110 := (>= #23 0::Int)
-#112 := (not #110)
-#115 := (or #112 #27)
-#650 := (forall (vars (?v0 Int)) (:pat #649) #115)
-#118 := (forall (vars (?v0 Int)) #115)
-#653 := (iff #118 #650)
-#651 := (iff #115 #115)
-#652 := [refl]: #651
-#654 := [quant-intro #652]: #653
-#155 := (~ #118 #118)
-#154 := (~ #115 #115)
-#151 := [refl]: #154
-#156 := [nnf-pos #151]: #155
-#24 := (<= 0::Int #23)
-#28 := (implies #24 #27)
-#29 := (forall (vars (?v0 Int)) #28)
-#121 := (iff #29 #118)
-#102 := (not #24)
-#103 := (or #102 #27)
-#106 := (forall (vars (?v0 Int)) #103)
-#119 := (iff #106 #118)
-#116 := (iff #103 #115)
-#113 := (iff #102 #112)
-#109 := (iff #24 #110)
-#111 := [rewrite]: #109
-#114 := [monotonicity #111]: #113
-#117 := [monotonicity #114]: #116
-#120 := [quant-intro #117]: #119
-#107 := (iff #29 #106)
-#104 := (iff #28 #103)
-#105 := [rewrite]: #104
-#108 := [quant-intro #105]: #107
-#122 := [trans #108 #120]: #121
-#101 := [asserted]: #29
-#123 := [mp #101 #122]: #118
-#152 := [mp~ #123 #156]: #118
-#655 := [mp #152 #654]: #650
-#85 := (f5 f6 #80)
-#88 := (f3 f4 #85)
-#91 := (= #88 #80)
-#94 := (not #91)
-#12 := (- f7)
-#11 := (< f7 0::Int)
-#13 := (if #11 #12 f7)
-#14 := (f5 f6 #13)
-#15 := (f3 f4 #14)
-#16 := (= #15 #13)
-#17 := (not #16)
-#97 := (iff #17 #94)
-#58 := (if #11 #55 f7)
-#61 := (f5 f6 #58)
-#64 := (f3 f4 #61)
-#67 := (= #64 #58)
-#70 := (not #67)
-#95 := (iff #70 #94)
-#92 := (iff #67 #91)
-#83 := (= #58 #80)
-#77 := (if #74 #55 f7)
-#81 := (= #77 #80)
-#82 := [rewrite]: #81
-#78 := (= #58 #77)
-#75 := (iff #11 #74)
-#76 := [rewrite]: #75
-#79 := [monotonicity #76]: #78
-#84 := [trans #79 #82]: #83
-#89 := (= #64 #88)
-#86 := (= #61 #85)
-#87 := [monotonicity #84]: #86
-#90 := [monotonicity #87]: #89
-#93 := [monotonicity #90 #84]: #92
-#96 := [monotonicity #93]: #95
-#71 := (iff #17 #70)
-#68 := (iff #16 #67)
-#59 := (= #13 #58)
-#56 := (= #12 #55)
-#57 := [rewrite]: #56
-#60 := [monotonicity #57]: #59
-#65 := (= #15 #64)
-#62 := (= #14 #61)
-#63 := [monotonicity #60]: #62
-#66 := [monotonicity #63]: #65
-#69 := [monotonicity #66 #60]: #68
-#72 := [monotonicity #69]: #71
-#98 := [trans #72 #96]: #97
-#53 := [asserted]: #17
-#99 := [mp #53 #98]: #94
-#630 := (not #650)
-#304 := (or #630 #316 #91)
-#636 := (or #316 #91)
-#305 := (or #630 #636)
-#638 := (iff #305 #304)
-#639 := [rewrite]: #638
-#637 := [quant-inst #80]: #305
-#640 := [mp #637 #639]: #304
-#618 := [unit-resolution #640 #99 #655]: #316
-#624 := [th-lemma arith farkas -1 1 1 #283 #618 #623]: false
-#262 := [lemma #624]: #74
-#315 := (or #73 #313)
-#306 := [def-axiom]: #315
-#267 := [unit-resolution #306 #262]: #313
-#268 := (not #313)
-#628 := (or #268 #625)
-#626 := [th-lemma arith triangle-eq]: #628
-#629 := [unit-resolution #626 #267]: #625
-#641 := (<= #80 0::Int)
-#615 := (or #641 #319)
-#616 := [th-lemma arith farkas 1 1]: #615
-#338 := [unit-resolution #616 #618]: #641
-[th-lemma arith farkas 1 1 1 #338 #262 #629]: false
-unsat
-7e6da58556dd56d85be0ea32c44b6f00c868dac5 431 0
-WARNING: For problems containing quantifiers, the model finding capabilities of Z3 work better when the formula does not contain nested quantifiers. You can use PULL_NESTED_QUANTIFIERS=true to eliminate nested quantifiers.
-#2 := false
-#446 := -1::Int
-decl f4 :: (-> S3 S2 Int)
-decl f7 :: (-> S4 Int S2)
-decl f9 :: S2
-#28 := f9
-decl f5 :: S3
-#11 := f5
-#29 := (f4 f5 f9)
-#27 := 4::Int
-#30 := (* 4::Int #29)
-#10 := 1::Int
-#112 := (+ 1::Int #30)
-decl f8 :: S4
-#17 := f8
-#115 := (f7 f8 #112)
-#362 := (f4 f5 #115)
-#662 := (* -1::Int #362)
-#673 := (+ #30 #662)
-#649 := (>= #673 -1::Int)
-#672 := (= #673 -1::Int)
-#41 := 0::Int
-#664 := (>= #29 0::Int)
-#644 := (= #362 0::Int)
-#593 := (not #644)
-#640 := (<= #362 0::Int)
-#628 := (not #640)
-#447 := (<= #362 1::Int)
-#752 := (not #447)
-decl f6 :: (-> S2 S2 S1)
-#7 := (:var 0 S2)
-#452 := (f6 #7 #115)
-#768 := (pattern #452)
-#451 := (= #7 #115)
-#18 := (f7 f8 1::Int)
-#19 := (= #7 #18)
-decl f1 :: S1
-#3 := f1
-#449 := (= #452 f1)
-#453 := (not #449)
-#432 := (or #453 #19 #451)
-#770 := (forall (vars (?v1 S2)) (:pat #768) #432)
-#426 := (not #770)
-#437 := (or #447 #426)
-#438 := (not #437)
-decl f3 :: (-> S2 S1)
-#118 := (f3 #115)
-#121 := (= #118 f1)
-#127 := (not #121)
-#771 := (or #127 #438)
-decl ?v1!0 :: (-> S2 S2)
-#772 := (?v1!0 #115)
-#767 := (= #772 #115)
-#425 := (= #772 #18)
-#773 := (f6 #772 #115)
-#774 := (= #773 f1)
-#769 := (not #774)
-#409 := (or #769 #425 #767)
-#766 := (not #409)
-#751 := (or #121 #447 #766)
-#413 := (not #751)
-#764 := (not #771)
-#414 := (or #764 #413)
-#415 := (not #414)
-#12 := (f4 f5 #7)
-#804 := (pattern #12)
-#8 := (f3 #7)
-#803 := (pattern #8)
-#219 := (?v1!0 #7)
-#222 := (= #219 #7)
-#221 := (= #219 #18)
-#202 := (f6 #219 #7)
-#203 := (= #202 f1)
-#220 := (not #203)
-#223 := (or #220 #221 #222)
-#224 := (not #223)
-#89 := (<= #12 1::Int)
-#9 := (= #8 f1)
-#266 := (or #9 #89 #224)
-#290 := (not #266)
-#14 := (:var 1 S2)
-#15 := (f6 #7 #14)
-#776 := (pattern #15)
-#20 := (= #7 #14)
-#16 := (= #15 f1)
-#73 := (not #16)
-#93 := (or #73 #19 #20)
-#777 := (forall (vars (?v1 S2)) (:pat #776) #93)
-#782 := (not #777)
-#785 := (or #89 #782)
-#788 := (not #785)
-#242 := (not #9)
-#791 := (or #242 #788)
-#794 := (not #791)
-#797 := (or #794 #290)
-#800 := (not #797)
-#805 := (forall (vars (?v0 S2)) (:pat #803 #804) #800)
-#96 := (forall (vars (?v1 S2)) #93)
-#225 := (not #96)
-#281 := (or #89 #225)
-#282 := (not #281)
-#283 := (or #242 #282)
-#289 := (not #283)
-#291 := (or #289 #290)
-#292 := (not #291)
-#297 := (forall (vars (?v0 S2)) #292)
-#806 := (iff #297 #805)
-#801 := (iff #292 #800)
-#798 := (iff #291 #797)
-#795 := (iff #289 #794)
-#792 := (iff #283 #791)
-#789 := (iff #282 #788)
-#786 := (iff #281 #785)
-#783 := (iff #225 #782)
-#780 := (iff #96 #777)
-#778 := (iff #93 #93)
-#779 := [refl]: #778
-#781 := [quant-intro #779]: #780
-#784 := [monotonicity #781]: #783
-#787 := [monotonicity #784]: #786
-#790 := [monotonicity #787]: #789
-#793 := [monotonicity #790]: #792
-#796 := [monotonicity #793]: #795
-#799 := [monotonicity #796]: #798
-#802 := [monotonicity #799]: #801
-#807 := [quant-intro #802]: #806
-#90 := (not #89)
-#99 := (and #90 #96)
-#248 := (or #242 #99)
-#271 := (and #248 #266)
-#274 := (forall (vars (?v0 S2)) #271)
-#298 := (iff #274 #297)
-#295 := (iff #271 #292)
-#286 := (and #283 #266)
-#293 := (iff #286 #292)
-#294 := [rewrite]: #293
-#287 := (iff #271 #286)
-#284 := (iff #248 #283)
-#214 := (iff #99 #282)
-#215 := [rewrite]: #214
-#285 := [monotonicity #215]: #284
-#288 := [monotonicity #285]: #287
-#296 := [trans #288 #294]: #295
-#299 := [quant-intro #296]: #298
-#216 := (not #90)
-#230 := (or #216 #224)
-#247 := (or #9 #230)
-#249 := (and #248 #247)
-#252 := (forall (vars (?v0 S2)) #249)
-#275 := (iff #252 #274)
-#272 := (iff #249 #271)
-#269 := (iff #247 #266)
-#260 := (or #89 #224)
-#263 := (or #9 #260)
-#267 := (iff #263 #266)
-#268 := [rewrite]: #267
-#264 := (iff #247 #263)
-#261 := (iff #230 #260)
-#258 := (iff #216 #89)
-#259 := [rewrite]: #258
-#262 := [monotonicity #259]: #261
-#265 := [monotonicity #262]: #264
-#270 := [trans #265 #268]: #269
-#273 := [monotonicity #270]: #272
-#276 := [quant-intro #273]: #275
-#102 := (iff #9 #99)
-#105 := (forall (vars (?v0 S2)) #102)
-#253 := (~ #105 #252)
-#250 := (~ #102 #249)
-#240 := (~ #99 #99)
-#238 := (~ #96 #96)
-#236 := (~ #93 #93)
-#237 := [refl]: #236
-#239 := [nnf-pos #237]: #238
-#234 := (~ #90 #90)
-#235 := [refl]: #234
-#241 := [monotonicity #235 #239]: #240
-#231 := (not #99)
-#232 := (~ #231 #230)
-#226 := (~ #225 #224)
-#227 := [sk]: #226
-#217 := (~ #216 #216)
-#218 := [refl]: #217
-#233 := [nnf-neg #218 #227]: #232
-#245 := (~ #9 #9)
-#246 := [refl]: #245
-#243 := (~ #242 #242)
-#244 := [refl]: #243
-#251 := [nnf-pos #244 #246 #233 #241]: #250
-#254 := [nnf-pos #251]: #253
-#21 := (or #19 #20)
-#22 := (implies #16 #21)
-#23 := (forall (vars (?v1 S2)) #22)
-#13 := (< 1::Int #12)
-#24 := (and #13 #23)
-#25 := (iff #9 #24)
-#26 := (forall (vars (?v0 S2)) #25)
-#108 := (iff #26 #105)
-#74 := (or #73 #21)
-#77 := (forall (vars (?v1 S2)) #74)
-#80 := (and #13 #77)
-#83 := (iff #9 #80)
-#86 := (forall (vars (?v0 S2)) #83)
-#106 := (iff #86 #105)
-#103 := (iff #83 #102)
-#100 := (iff #80 #99)
-#97 := (iff #77 #96)
-#94 := (iff #74 #93)
-#95 := [rewrite]: #94
-#98 := [quant-intro #95]: #97
-#91 := (iff #13 #90)
-#92 := [rewrite]: #91
-#101 := [monotonicity #92 #98]: #100
-#104 := [monotonicity #101]: #103
-#107 := [quant-intro #104]: #106
-#87 := (iff #26 #86)
-#84 := (iff #25 #83)
-#81 := (iff #24 #80)
-#78 := (iff #23 #77)
-#75 := (iff #22 #74)
-#76 := [rewrite]: #75
-#79 := [quant-intro #76]: #78
-#82 := [monotonicity #79]: #81
-#85 := [monotonicity #82]: #84
-#88 := [quant-intro #85]: #87
-#109 := [trans #88 #107]: #108
-#72 := [asserted]: #26
-#110 := [mp #72 #109]: #105
-#255 := [mp~ #110 #254]: #252
-#256 := [mp #255 #276]: #274
-#300 := [mp #256 #299]: #297
-#808 := [mp #300 #807]: #805
-#756 := (not #805)
-#753 := (or #756 #415)
-#757 := [quant-inst #115]: #753
-#566 := [unit-resolution #757 #808]: #415
-#730 := (or #414 #771)
-#736 := [def-axiom]: #730
-#621 := [unit-resolution #736 #566]: #771
-#602 := (or #764 #438)
-#138 := (>= #29 1::Int)
-#139 := (or #127 #138)
-#142 := (not #139)
-#35 := (<= 1::Int #29)
-#31 := (+ #30 1::Int)
-#32 := (f7 f8 #31)
-#33 := (f3 #32)
-#34 := (= #33 f1)
-#36 := (implies #34 #35)
-#37 := (not #36)
-#145 := (iff #37 #142)
-#128 := (or #127 #35)
-#133 := (not #128)
-#143 := (iff #133 #142)
-#140 := (iff #128 #139)
-#136 := (iff #35 #138)
-#137 := [rewrite]: #136
-#141 := [monotonicity #137]: #140
-#144 := [monotonicity #141]: #143
-#134 := (iff #37 #133)
-#131 := (iff #36 #128)
-#124 := (implies #121 #35)
-#129 := (iff #124 #128)
-#130 := [rewrite]: #129
-#125 := (iff #36 #124)
-#122 := (iff #34 #121)
-#119 := (= #33 #118)
-#116 := (= #32 #115)
-#113 := (= #31 #112)
-#114 := [rewrite]: #113
-#117 := [monotonicity #114]: #116
-#120 := [monotonicity #117]: #119
-#123 := [monotonicity #120]: #122
-#126 := [monotonicity #123]: #125
-#132 := [trans #126 #130]: #131
-#135 := [monotonicity #132]: #134
-#146 := [trans #135 #144]: #145
-#111 := [asserted]: #37
-#147 := [mp #111 #146]: #142
-#148 := [not-or-elim #147]: #121
-#744 := (or #764 #127 #438)
-#748 := [def-axiom]: #744
-#626 := [unit-resolution #748 #148]: #602
-#627 := [unit-resolution #626 #621]: #438
-#758 := (or #437 #752)
-#395 := [def-axiom]: #758
-#622 := [unit-resolution #395 #627]: #752
-#596 := (or #628 #447)
-#603 := [th-lemma arith farkas 1 1]: #596
-#562 := [unit-resolution #603 #622]: #628
-#595 := (or #593 #640)
-#597 := [th-lemma arith triangle-eq]: #595
-#604 := [unit-resolution #597 #562]: #593
-#623 := (or #664 #644)
-#42 := (:var 0 Int)
-#44 := (f7 f8 #42)
-#815 := (pattern #44)
-#45 := (f4 f5 #44)
-#50 := (= #45 0::Int)
-#162 := (>= #42 0::Int)
-#192 := (or #162 #50)
-#822 := (forall (vars (?v0 Int)) (:pat #815) #192)
-#195 := (forall (vars (?v0 Int)) #192)
-#825 := (iff #195 #822)
-#823 := (iff #192 #192)
-#824 := [refl]: #823
-#826 := [quant-intro #824]: #825
-#212 := (~ #195 #195)
-#278 := (~ #192 #192)
-#279 := [refl]: #278
-#213 := [nnf-pos #279]: #212
-#49 := (< #42 0::Int)
-#51 := (implies #49 #50)
-#52 := (forall (vars (?v0 Int)) #51)
-#198 := (iff #52 #195)
-#176 := (not #49)
-#177 := (or #176 #50)
-#180 := (forall (vars (?v0 Int)) #177)
-#196 := (iff #180 #195)
-#193 := (iff #177 #192)
-#190 := (iff #176 #162)
-#163 := (not #162)
-#185 := (not #163)
-#188 := (iff #185 #162)
-#189 := [rewrite]: #188
-#186 := (iff #176 #185)
-#183 := (iff #49 #163)
-#184 := [rewrite]: #183
-#187 := [monotonicity #184]: #186
-#191 := [trans #187 #189]: #190
-#194 := [monotonicity #191]: #193
-#197 := [quant-intro #194]: #196
-#181 := (iff #52 #180)
-#178 := (iff #51 #177)
-#179 := [rewrite]: #178
-#182 := [quant-intro #179]: #181
-#199 := [trans #182 #197]: #198
-#175 := [asserted]: #52
-#200 := [mp #175 #199]: #195
-#280 := [mp~ #200 #213]: #195
-#827 := [mp #280 #826]: #822
-#518 := (not #822)
-#629 := (or #518 #664 #644)
-#678 := (>= #112 0::Int)
-#650 := (or #678 #644)
-#630 := (or #518 #650)
-#638 := (iff #630 #629)
-#636 := (or #518 #623)
-#634 := (iff #636 #629)
-#637 := [rewrite]: #634
-#632 := (iff #630 #636)
-#624 := (iff #650 #623)
-#665 := (iff #678 #664)
-#666 := [rewrite]: #665
-#625 := [monotonicity #666]: #624
-#633 := [monotonicity #625]: #632
-#639 := [trans #633 #637]: #638
-#631 := [quant-inst #112]: #630
-#635 := [mp #631 #639]: #629
-#606 := [unit-resolution #635 #827]: #623
-#607 := [unit-resolution #606 #604]: #664
-#667 := (not #664)
-#651 := (or #667 #672)
-#46 := (= #45 #42)
-#166 := (or #163 #46)
-#816 := (forall (vars (?v0 Int)) (:pat #815) #166)
-#169 := (forall (vars (?v0 Int)) #166)
-#819 := (iff #169 #816)
-#817 := (iff #166 #166)
-#818 := [refl]: #817
-#820 := [quant-intro #818]: #819
-#210 := (~ #169 #169)
-#209 := (~ #166 #166)
-#206 := [refl]: #209
-#211 := [nnf-pos #206]: #210
-#43 := (<= 0::Int #42)
-#47 := (implies #43 #46)
-#48 := (forall (vars (?v0 Int)) #47)
-#172 := (iff #48 #169)
-#153 := (not #43)
-#154 := (or #153 #46)
-#157 := (forall (vars (?v0 Int)) #154)
-#170 := (iff #157 #169)
-#167 := (iff #154 #166)
-#164 := (iff #153 #163)
-#160 := (iff #43 #162)
-#161 := [rewrite]: #160
-#165 := [monotonicity #161]: #164
-#168 := [monotonicity #165]: #167
-#171 := [quant-intro #168]: #170
-#158 := (iff #48 #157)
-#155 := (iff #47 #154)
-#156 := [rewrite]: #155
-#159 := [quant-intro #156]: #158
-#173 := [trans #159 #171]: #172
-#152 := [asserted]: #48
-#174 := [mp #152 #173]: #169
-#207 := [mp~ #174 #211]: #169
-#821 := [mp #207 #820]: #816
-#655 := (not #816)
-#656 := (or #655 #667 #672)
-#661 := (= #362 #112)
-#679 := (not #678)
-#663 := (or #679 #661)
-#657 := (or #655 #663)
-#643 := (iff #657 #656)
-#653 := (or #655 #651)
-#641 := (iff #653 #656)
-#642 := [rewrite]: #641
-#659 := (iff #657 #653)
-#652 := (iff #663 #651)
-#670 := (iff #661 #672)
-#671 := [rewrite]: #670
-#668 := (iff #679 #667)
-#669 := [monotonicity #666]: #668
-#654 := [monotonicity #669 #671]: #652
-#645 := [monotonicity #654]: #659
-#646 := [trans #645 #642]: #643
-#658 := [quant-inst #112]: #657
-#647 := [mp #658 #646]: #656
-#608 := [unit-resolution #647 #821]: #651
-#618 := [unit-resolution #608 #607]: #672
-#598 := (not #672)
-#619 := (or #598 #649)
-#574 := [th-lemma arith triangle-eq]: #619
-#575 := [unit-resolution #574 #618]: #649
-#149 := (not #138)
-#150 := [not-or-elim #147]: #149
-[th-lemma arith farkas -4 1 1 #150 #622 #575]: false
-unsat
-f0add7d14def5da0b06e595882e28df041b2cf29 58 0
-#2 := false
-decl f8 :: S2
-#18 := f8
-decl f6 :: S2
-#14 := f6
-#20 := (= f6 f8)
-decl f3 :: (-> S4 S5 S2)
-decl f5 :: (-> S2 S3 S5)
-decl f7 :: S3
-#15 := f7
-#16 := (f5 f6 f7)
-decl f4 :: S4
-#7 := f4
-#17 := (f3 f4 #16)
-#19 := (= #17 f8)
-#45 := (not #19)
-#46 := (or #45 #20)
-#49 := (not #46)
-#21 := (implies #19 #20)
-#22 := (not #21)
-#50 := (iff #22 #49)
-#47 := (iff #21 #46)
-#48 := [rewrite]: #47
-#51 := [monotonicity #48]: #50
-#44 := [asserted]: #22
-#54 := [mp #44 #51]: #49
-#52 := [not-or-elim #54]: #19
-#125 := (= f6 #17)
-#124 := (= #17 f6)
-#9 := (:var 0 S3)
-#8 := (:var 1 S2)
-#10 := (f5 #8 #9)
-#540 := (pattern #10)
-#11 := (f3 f4 #10)
-#12 := (= #11 #8)
-#541 := (forall (vars (?v0 S2) (?v1 S3)) (:pat #540) #12)
-#13 := (forall (vars (?v0 S2) (?v1 S3)) #12)
-#544 := (iff #13 #541)
-#542 := (iff #12 #12)
-#543 := [refl]: #542
-#545 := [quant-intro #543]: #544
-#67 := (~ #13 #13)
-#65 := (~ #12 #12)
-#66 := [refl]: #65
-#68 := [nnf-pos #66]: #67
-#43 := [asserted]: #13
-#57 := [mp~ #43 #68]: #13
-#546 := [mp #57 #545]: #541
-#211 := (not #541)
-#126 := (or #211 #124)
-#212 := [quant-inst #14 #15]: #126
-#210 := [unit-resolution #212 #546]: #124
-#203 := [symm #210]: #125
-#214 := [trans #203 #52]: #20
-#53 := (not #20)
-#55 := [not-or-elim #54]: #53
-[unit-resolution #55 #214]: false
-unsat
-86345bce2206ce27e174d4b1d6d3e0182564f8a1 106 0
-#2 := false
-decl f11 :: (-> S9 S5 S3)
-decl f16 :: S5
-#34 := f16
-decl f12 :: S9
-#25 := f12
-#39 := (f11 f12 f16)
-decl f6 :: (-> S6 S7 S3)
-decl f13 :: S7
-#29 := f13
-decl f7 :: S6
-#14 := f7
-#38 := (f6 f7 f13)
-#40 := (= #38 #39)
-decl f5 :: (-> S2 S3 S5)
-decl f14 :: S3
-#30 := f14
-decl f15 :: S2
-#31 := f15
-#35 := (f5 f15 f14)
-#165 := (f11 f12 #35)
-#233 := (= #165 #39)
-#573 := (= #39 #165)
-#36 := (= f16 #35)
-decl f8 :: (-> S3 S2 S7)
-#32 := (f8 f14 f15)
-#33 := (= f13 #32)
-#37 := (and #33 #36)
-#68 := (not #37)
-#69 := (or #68 #40)
-#72 := (not #69)
-#41 := (implies #37 #40)
-#42 := (not #41)
-#73 := (iff #42 #72)
-#70 := (iff #41 #69)
-#71 := [rewrite]: #70
-#74 := [monotonicity #71]: #73
-#67 := [asserted]: #42
-#77 := [mp #67 #74]: #72
-#75 := [not-or-elim #77]: #37
-#78 := [and-elim #75]: #36
-#579 := [monotonicity #78]: #573
-#570 := [symm #579]: #233
-#213 := (= #38 #165)
-#569 := (= f14 #165)
-#251 := (= #165 f14)
-#9 := (:var 0 S3)
-#8 := (:var 1 S2)
-#10 := (f5 #8 #9)
-#580 := (pattern #10)
-#26 := (f11 f12 #10)
-#27 := (= #26 #9)
-#600 := (forall (vars (?v0 S2) (?v1 S3)) (:pat #580) #27)
-#28 := (forall (vars (?v0 S2) (?v1 S3)) #27)
-#603 := (iff #28 #600)
-#601 := (iff #27 #27)
-#602 := [refl]: #601
-#604 := [quant-intro #602]: #603
-#88 := (~ #28 #28)
-#107 := (~ #27 #27)
-#108 := [refl]: #107
-#89 := [nnf-pos #108]: #88
-#66 := [asserted]: #28
-#109 := [mp~ #66 #89]: #28
-#605 := [mp #109 #604]: #600
-#256 := (not #600)
-#253 := (or #256 #251)
-#257 := [quant-inst #31 #30]: #253
-#568 := [unit-resolution #257 #605]: #251
-#228 := [symm #568]: #569
-#229 := (= #38 f14)
-#254 := (f6 f7 #32)
-#255 := (= #254 f14)
-#16 := (:var 0 S2)
-#15 := (:var 1 S3)
-#17 := (f8 #15 #16)
-#587 := (pattern #17)
-#18 := (f6 f7 #17)
-#19 := (= #18 #15)
-#588 := (forall (vars (?v0 S3) (?v1 S2)) (:pat #587) #19)
-#20 := (forall (vars (?v0 S3) (?v1 S2)) #19)
-#591 := (iff #20 #588)
-#589 := (iff #19 #19)
-#590 := [refl]: #589
-#592 := [quant-intro #590]: #591
-#84 := (~ #20 #20)
-#83 := (~ #19 #19)
-#102 := [refl]: #83
-#85 := [nnf-pos #102]: #84
-#64 := [asserted]: #20
-#103 := [mp~ #64 #85]: #20
-#593 := [mp #103 #592]: #588
-#574 := (not #588)
-#230 := (or #574 #255)
-#361 := [quant-inst #30 #31]: #230
-#241 := [unit-resolution #361 #593]: #255
-#577 := (= #38 #254)
-#76 := [and-elim #75]: #33
-#578 := [monotonicity #76]: #577
-#571 := [trans #578 #241]: #229
-#555 := [trans #571 #228]: #213
-#217 := [trans #555 #570]: #40
-#79 := (not #40)
-#80 := [not-or-elim #77]: #79
-[unit-resolution #80 #217]: false
-unsat
-7180d528e452ef46d73483bf56a7d7018ee1b306 113 0
-#2 := false
-decl f3 :: (-> S2 S3 S4)
-decl f8 :: S3
-#30 := f8
-decl f11 :: S2
-#38 := f11
-#48 := (f3 f11 f8)
-decl f4 :: (-> S5 S4 S2)
-decl f13 :: S4
-#45 := f13
-decl f5 :: (-> S6 S3 S5)
-decl f10 :: S3
-#34 := f10
-decl f6 :: (-> S7 S2 S6)
-decl f12 :: S4
-#41 := f12
-decl f9 :: S3
-#31 := f9
-decl f7 :: S7
-#7 := f7
-#39 := (f6 f7 f11)
-#40 := (f5 #39 f9)
-#42 := (f4 #40 f12)
-#43 := (f6 f7 #42)
-#44 := (f5 #43 f10)
-#46 := (f4 #44 f13)
-#47 := (f3 #46 f8)
-#49 := (= #47 #48)
-#261 := (f3 #42 f8)
-#271 := (= #261 #48)
-#270 := (= #261 f12)
-#32 := (= f8 f9)
-#549 := (if #32 #270 #271)
-#23 := (:var 0 S3)
-#21 := (:var 1 S4)
-#19 := (:var 2 S3)
-#17 := (:var 3 S2)
-#18 := (f6 f7 #17)
-#20 := (f5 #18 #19)
-#22 := (f4 #20 #21)
-#24 := (f3 #22 #23)
-#593 := (pattern #24)
-#26 := (f3 #17 #23)
-#108 := (= #24 #26)
-#107 := (= #24 #21)
-#25 := (= #23 #19)
-#93 := (if #25 #107 #108)
-#594 := (forall (vars (?v0 S2) (?v1 S3) (?v2 S4) (?v3 S3)) (:pat #593) #93)
-#100 := (forall (vars (?v0 S2) (?v1 S3) (?v2 S4) (?v3 S3)) #93)
-#597 := (iff #100 #594)
-#595 := (iff #93 #93)
-#596 := [refl]: #595
-#598 := [quant-intro #596]: #597
-#27 := (if #25 #21 #26)
-#28 := (= #24 #27)
-#29 := (forall (vars (?v0 S2) (?v1 S3) (?v2 S4) (?v3 S3)) #28)
-#97 := (iff #29 #100)
-#94 := (iff #28 #93)
-#99 := [rewrite]: #94
-#98 := [quant-intro #99]: #97
-#91 := (~ #29 #29)
-#90 := (~ #28 #28)
-#105 := [refl]: #90
-#92 := [nnf-pos #105]: #91
-#73 := [asserted]: #29
-#106 := [mp~ #73 #92]: #29
-#95 := [mp #106 #98]: #100
-#599 := [mp #95 #598]: #594
-#236 := (not #594)
-#547 := (or #236 #549)
-#551 := [quant-inst #38 #31 #41 #30]: #547
-#550 := [unit-resolution #551 #599]: #549
-#548 := (not #549)
-#264 := (or #548 #271)
-#33 := (not #32)
-#35 := (= f8 f10)
-#36 := (not #35)
-#37 := (and #33 #36)
-#75 := (not #37)
-#76 := (or #75 #49)
-#79 := (not #76)
-#50 := (implies #37 #49)
-#51 := (not #50)
-#80 := (iff #51 #79)
-#77 := (iff #50 #76)
-#78 := [rewrite]: #77
-#81 := [monotonicity #78]: #80
-#74 := [asserted]: #51
-#84 := [mp #74 #81]: #79
-#82 := [not-or-elim #84]: #37
-#83 := [and-elim #82]: #33
-#542 := (or #548 #32 #271)
-#543 := [def-axiom]: #542
-#387 := [unit-resolution #543 #83]: #264
-#388 := [unit-resolution #387 #550]: #271
-#263 := (= #47 #261)
-#260 := (= #47 f13)
-#242 := (if #35 #260 #263)
-#367 := (or #236 #242)
-#574 := [quant-inst #42 #34 #45 #30]: #367
-#389 := [unit-resolution #574 #599]: #242
-#247 := (not #242)
-#531 := (or #247 #263)
-#85 := [and-elim #82]: #36
-#582 := (or #247 #35 #263)
-#583 := [def-axiom]: #582
-#532 := [unit-resolution #583 #85]: #531
-#533 := [unit-resolution #532 #389]: #263
-#529 := [trans #533 #388]: #49
-#86 := (not #49)
-#87 := [not-or-elim #84]: #86
-[unit-resolution #87 #529]: false
-unsat
-1c419ffe565f74df1755b00362bfce413a0bbb21 74 0
-#2 := false
-decl f1 :: S1
-#3 := f1
-decl f6 :: (-> S2 S3 S1)
-decl f5 :: S3
-#8 := f5
-decl f4 :: S2
-#7 := f4
-#11 := (f6 f4 f5)
-#12 := (= #11 f1)
-decl f3 :: (-> S2 S3 S1)
-#9 := (f3 f4 f5)
-#10 := (= #9 f1)
-#70 := (not #10)
-#77 := (iff #70 #12)
-#81 := (iff #77 false)
-#83 := (iff #10 false)
-#43 := (iff #10 #12)
-#59 := (or #43 #10 #12)
-#62 := (not #59)
-#1 := true
-#16 := (iff #12 true)
-#15 := (iff #10 true)
-#17 := (or #15 #16)
-#13 := (and #12 true)
-#14 := (iff #10 #13)
-#18 := (or #14 #17)
-#19 := (not #18)
-#65 := (iff #19 #62)
-#50 := (or #10 #12)
-#53 := (or #43 #50)
-#56 := (not #53)
-#63 := (iff #56 #62)
-#60 := (iff #53 #59)
-#61 := [rewrite]: #60
-#64 := [monotonicity #61]: #63
-#57 := (iff #19 #56)
-#54 := (iff #18 #53)
-#51 := (iff #17 #50)
-#48 := (iff #16 #12)
-#49 := [rewrite]: #48
-#46 := (iff #15 #10)
-#47 := [rewrite]: #46
-#52 := [monotonicity #47 #49]: #51
-#44 := (iff #14 #43)
-#41 := (iff #13 #12)
-#42 := [rewrite]: #41
-#45 := [monotonicity #42]: #44
-#55 := [monotonicity #45 #52]: #54
-#58 := [monotonicity #55]: #57
-#66 := [trans #58 #64]: #65
-#40 := [asserted]: #19
-#67 := [mp #40 #66]: #62
-#71 := [not-or-elim #67]: #70
-#84 := [iff-false #71]: #83
-#92 := (iff #77 #10)
-#87 := (iff #70 false)
-#90 := (iff #87 #10)
-#91 := [rewrite]: #90
-#88 := (iff #77 #87)
-#85 := (iff #12 false)
-#72 := (not #12)
-#73 := [not-or-elim #67]: #72
-#86 := [iff-false #73]: #85
-#89 := [monotonicity #86]: #88
-#93 := [trans #89 #91]: #92
-#82 := [trans #93 #84]: #81
-#68 := (not #43)
-#78 := (iff #68 #77)
-#79 := [rewrite]: #78
-#69 := [not-or-elim #67]: #68
-#80 := [mp #69 #79]: #77
-[mp #80 #82]: false
-unsat
 76d09b53549e91e8b6b69b6b905b5e8307464c6f 106 0
 #2 := false
 decl f7 :: S2
@@ -10684,1133 +1668,6 @@
 #215 := [quant-inst #19]: #210
 [unit-resolution #215 #568 #555]: false
 unsat
-1396ebdf2db554fa58d5de90d7aa27d442610f3c 29 0
-#2 := false
-#1 := true
-decl f1 :: S1
-#3 := f1
-decl f3 :: (-> S1 S1)
-decl f2 :: S1
-#4 := f2
-decl f4 :: (-> S2 S1)
-#7 := (:var 0 S2)
-#8 := (f4 #7)
-#9 := (= #8 f1)
-#10 := (exists (vars (?v0 S2)) #9)
-#11 := (if #10 f1 f2)
-#12 := (f3 #11)
-#13 := (= #12 f1)
-#14 := (implies #13 true)
-#15 := (not #14)
-#44 := (iff #15 false)
-#39 := (not true)
-#42 := (iff #39 false)
-#43 := [rewrite]: #42
-#40 := (iff #15 #39)
-#37 := (iff #14 true)
-#38 := [rewrite]: #37
-#41 := [monotonicity #38]: #40
-#45 := [trans #41 #43]: #44
-#36 := [asserted]: #15
-[mp #36 #45]: false
-unsat
-352ef3cbf5b05cf656dc82749237c3b497c01e97 113 0
-#2 := false
-decl f1 :: S1
-#3 := f1
-decl f3 :: (-> S2 Int S1)
-#21 := 42::Int
-decl f4 :: (-> S3 Int S2)
-#19 := 3::Int
-decl f6 :: S3
-#17 := f6
-#20 := (f4 f6 3::Int)
-#22 := (f3 #20 42::Int)
-#23 := (= #22 f1)
-decl f5 :: S3
-#7 := f5
-#139 := (f4 f5 3::Int)
-#223 := (f3 #139 42::Int)
-#224 := (= #223 f1)
-#10 := (:var 0 Int)
-#8 := (:var 1 Int)
-#9 := (f4 f5 #8)
-#11 := (f3 #9 #10)
-#12 := (pattern #11)
-#27 := 0::Int
-#49 := -1::Int
-#50 := (* -1::Int #10)
-#51 := (+ #8 #50)
-#52 := (<= #51 0::Int)
-#13 := (= #11 f1)
-#55 := (iff #13 #52)
-#58 := (forall (vars (?v0 Int) (?v1 Int)) (:pat #12) #55)
-#83 := (~ #58 #58)
-#81 := (~ #55 #55)
-#82 := [refl]: #81
-#84 := [nnf-pos #82]: #83
-#14 := (<= #8 #10)
-#15 := (iff #13 #14)
-#16 := (forall (vars (?v0 Int) (?v1 Int)) (:pat #12) #15)
-#59 := (iff #16 #58)
-#56 := (iff #15 #55)
-#53 := (iff #14 #52)
-#54 := [rewrite]: #53
-#57 := [monotonicity #54]: #56
-#60 := [quant-intro #57]: #59
-#46 := [asserted]: #16
-#61 := [mp #46 #60]: #58
-#73 := [mp~ #61 #84]: #58
-#190 := (not #58)
-#191 := (or #190 #224)
-#225 := (* -1::Int 42::Int)
-#216 := (+ 3::Int #225)
-#227 := (<= #216 0::Int)
-#228 := (iff #224 #227)
-#192 := (or #190 #228)
-#529 := (iff #192 #191)
-#531 := (iff #191 #191)
-#532 := [rewrite]: #531
-#186 := (iff #228 #224)
-#1 := true
-#201 := (iff #224 true)
-#202 := (iff #201 #224)
-#543 := [rewrite]: #202
-#206 := (iff #228 #201)
-#551 := (iff #227 true)
-#203 := -39::Int
-#547 := (<= -39::Int 0::Int)
-#550 := (iff #547 true)
-#545 := [rewrite]: #550
-#548 := (iff #227 #547)
-#214 := (= #216 -39::Int)
-#229 := -42::Int
-#209 := (+ 3::Int -42::Int)
-#333 := (= #209 -39::Int)
-#540 := [rewrite]: #333
-#544 := (= #216 #209)
-#226 := (= #225 -42::Int)
-#230 := [rewrite]: #226
-#546 := [monotonicity #230]: #544
-#215 := [trans #546 #540]: #214
-#549 := [monotonicity #215]: #548
-#541 := [trans #549 #545]: #551
-#542 := [monotonicity #541]: #206
-#527 := [trans #542 #543]: #186
-#530 := [monotonicity #527]: #529
-#533 := [trans #530 #532]: #529
-#193 := [quant-inst #19 #21]: #192
-#528 := [mp #193 #533]: #191
-#534 := [unit-resolution #528 #73]: #224
-#536 := (= #22 #223)
-#178 := (= #20 #139)
-#537 := (= #139 #20)
-#172 := (= f5 f6)
-#18 := (= f6 f5)
-#48 := (not #18)
-#62 := (or #48 #23)
-#65 := (not #62)
-#24 := (implies #18 #23)
-#25 := (not #24)
-#66 := (iff #25 #65)
-#63 := (iff #24 #62)
-#64 := [rewrite]: #63
-#67 := [monotonicity #64]: #66
-#47 := [asserted]: #25
-#70 := [mp #47 #67]: #65
-#68 := [not-or-elim #70]: #18
-#535 := [symm #68]: #172
-#177 := [monotonicity #535]: #537
-#538 := [symm #177]: #178
-#539 := [monotonicity #538]: #536
-#525 := [trans #539 #534]: #23
-#69 := (not #23)
-#71 := [not-or-elim #70]: #69
-[unit-resolution #71 #525]: false
-unsat
-8fa5494ea43f950aa9add5e070d1d34c34426a1b 29 0
-#2 := false
-#1 := true
-decl f1 :: S1
-#3 := f1
-decl f3 :: (-> S1 S1)
-decl f2 :: S1
-#4 := f2
-decl f4 :: (-> S2 S1)
-#7 := (:var 0 S2)
-#8 := (f4 #7)
-#9 := (= #8 f1)
-#10 := (forall (vars (?v0 S2)) #9)
-#11 := (if #10 f1 f2)
-#12 := (f3 #11)
-#13 := (= #12 f1)
-#14 := (implies #13 true)
-#15 := (not #14)
-#44 := (iff #15 false)
-#39 := (not true)
-#42 := (iff #39 false)
-#43 := [rewrite]: #42
-#40 := (iff #15 #39)
-#37 := (iff #14 true)
-#38 := [rewrite]: #37
-#41 := [monotonicity #38]: #40
-#45 := [trans #41 #43]: #44
-#36 := [asserted]: #15
-[mp #36 #45]: false
-unsat
-2fd48adc6f5c51aec7f5f7945dc6937d8ac8fd61 424 0
-#2 := false
-decl f9 :: (-> S6 S7 S7)
-decl f12 :: S7
-#22 := f12
-decl f13 :: (-> S9 S2 S6)
-decl f5 :: (-> S4 Int S2)
-#49 := 2::Int
-decl f6 :: S4
-#11 := f6
-#50 := (f5 f6 2::Int)
-decl f14 :: S9
-#28 := f14
-#51 := (f13 f14 #50)
-#52 := (f9 #51 f12)
-#14 := 1::Int
-#44 := (f5 f6 1::Int)
-#45 := (f13 f14 #44)
-#53 := (f9 #45 #52)
-#46 := (f9 #45 f12)
-#41 := 0::Int
-#42 := (f5 f6 0::Int)
-#43 := (f13 f14 #42)
-#47 := (f9 #43 #46)
-decl f10 :: (-> S8 S3 S6)
-decl f4 :: S3
-#7 := f4
-decl f11 :: S8
-#19 := f11
-#40 := (f10 f11 f4)
-#48 := (f9 #40 #47)
-#54 := (= #48 #53)
-#654 := (f9 #40 #46)
-decl f3 :: (-> S3 S2 S2)
-#337 := (f3 f4 #42)
-#338 := (f13 f14 #337)
-#656 := (f9 #338 #654)
-#321 := (= #656 #53)
-#353 := (= #53 #656)
-#391 := (= #52 #654)
-#248 := (f9 #40 f12)
-#596 := (f3 f4 #44)
-#593 := (f13 f14 #596)
-#597 := (f9 #593 #248)
-#389 := (= #597 #654)
-#584 := (= #654 #597)
-#31 := (:var 0 S7)
-#26 := (:var 2 S3)
-#27 := (f10 f11 #26)
-#36 := (f9 #27 #31)
-#29 := (:var 1 S2)
-#34 := (f3 #26 #29)
-#35 := (f13 f14 #34)
-#37 := (f9 #35 #36)
-#670 := (pattern #37)
-#30 := (f13 f14 #29)
-#32 := (f9 #30 #31)
-#33 := (f9 #27 #32)
-#669 := (pattern #33)
-#38 := (= #33 #37)
-#671 := (forall (vars (?v0 S3) (?v1 S2) (?v2 S7)) (:pat #669 #670) #38)
-#39 := (forall (vars (?v0 S3) (?v1 S2) (?v2 S7)) #38)
-#674 := (iff #39 #671)
-#672 := (iff #38 #38)
-#673 := [refl]: #672
-#675 := [quant-intro #673]: #674
-#161 := (~ #39 #39)
-#179 := (~ #38 #38)
-#180 := [refl]: #179
-#162 := [nnf-pos #180]: #161
-#103 := [asserted]: #39
-#181 := [mp~ #103 #162]: #39
-#676 := [mp #181 #675]: #671
-#323 := (not #671)
-#575 := (or #323 #584)
-#577 := [quant-inst #7 #44 #22]: #575
-#430 := [unit-resolution #577 #676]: #584
-#390 := [symm #430]: #389
-#387 := (= #52 #597)
-#435 := (= f12 #248)
-#332 := (= #248 f12)
-#20 := (:var 0 S3)
-#21 := (f10 f11 #20)
-#662 := (pattern #21)
-#23 := (f9 #21 f12)
-#24 := (= #23 f12)
-#663 := (forall (vars (?v0 S3)) (:pat #662) #24)
-#25 := (forall (vars (?v0 S3)) #24)
-#666 := (iff #25 #663)
-#664 := (iff #24 #24)
-#665 := [refl]: #664
-#667 := [quant-intro #665]: #666
-#159 := (~ #25 #25)
-#158 := (~ #24 #24)
-#177 := [refl]: #158
-#160 := [nnf-pos #177]: #159
-#102 := [asserted]: #25
-#178 := [mp~ #102 #160]: #25
-#668 := [mp #178 #667]: #663
-#335 := (not #663)
-#339 := (or #335 #332)
-#318 := [quant-inst #7]: #339
-#431 := [unit-resolution #318 #668]: #332
-#436 := [symm #431]: #435
-#384 := (= #51 #593)
-#399 := (= #50 #596)
-decl f7 :: (-> S5 S2 Int)
-decl f8 :: S5
-#12 := f8
-#254 := (f7 f8 #44)
-#580 := (+ 1::Int #254)
-#581 := (f5 f6 #580)
-#412 := (= #581 #596)
-#582 := (= #596 #581)
-#8 := (:var 0 S2)
-#9 := (f3 f4 #8)
-#10 := (pattern #9)
-#13 := (f7 f8 #8)
-#90 := (+ 1::Int #13)
-#93 := (f5 f6 #90)
-#96 := (= #9 #93)
-#99 := (forall (vars (?v0 S2)) (:pat #10) #96)
-#175 := (~ #99 #99)
-#173 := (~ #96 #96)
-#174 := [refl]: #173
-#176 := [nnf-pos #174]: #175
-#15 := (+ #13 1::Int)
-#16 := (f5 f6 #15)
-#17 := (= #9 #16)
-#18 := (forall (vars (?v0 S2)) (:pat #10) #17)
-#100 := (iff #18 #99)
-#97 := (iff #17 #96)
-#94 := (= #16 #93)
-#91 := (= #15 #90)
-#92 := [rewrite]: #91
-#95 := [monotonicity #92]: #94
-#98 := [monotonicity #95]: #97
-#101 := [quant-intro #98]: #100
-#89 := [asserted]: #18
-#104 := [mp #89 #101]: #99
-#157 := [mp~ #104 #176]: #99
-#585 := (not #99)
-#567 := (or #585 #582)
-#568 := [quant-inst #44]: #567
-#278 := [unit-resolution #568 #157]: #582
-#398 := [symm #278]: #412
-#400 := (= #50 #581)
-#522 := (f7 f8 #581)
-#450 := (f5 f6 #522)
-#451 := (= #450 #581)
-#677 := (pattern #13)
-#56 := (f5 f6 #13)
-#57 := (= #56 #8)
-#678 := (forall (vars (?v0 S2)) (:pat #677) #57)
-#58 := (forall (vars (?v0 S2)) #57)
-#681 := (iff #58 #678)
-#679 := (iff #57 #57)
-#680 := [refl]: #679
-#682 := [quant-intro #680]: #681
-#163 := (~ #58 #58)
-#182 := (~ #57 #57)
-#183 := [refl]: #182
-#164 := [nnf-pos #183]: #163
-#106 := [asserted]: #58
-#165 := [mp~ #106 #164]: #58
-#683 := [mp #165 #682]: #678
-#453 := (not #678)
-#458 := (or #453 #451)
-#441 := [quant-inst #581]: #458
-#437 := [unit-resolution #441 #683]: #451
-#408 := (= #50 #450)
-#407 := (= 2::Int #522)
-#410 := (= #522 2::Int)
-#247 := -1::Int
-#507 := (* -1::Int #522)
-#488 := (+ #254 #507)
-#484 := (<= #488 -1::Int)
-#452 := (= #488 -1::Int)
-#520 := (>= #254 -1::Int)
-#515 := (>= #254 1::Int)
-#631 := (= #254 1::Int)
-#59 := (:var 0 Int)
-#61 := (f5 f6 #59)
-#684 := (pattern #61)
-#62 := (f7 f8 #61)
-#63 := (= #62 #59)
-#117 := (>= #59 0::Int)
-#118 := (not #117)
-#121 := (or #118 #63)
-#685 := (forall (vars (?v0 Int)) (:pat #684) #121)
-#124 := (forall (vars (?v0 Int)) #121)
-#688 := (iff #124 #685)
-#686 := (iff #121 #121)
-#687 := [refl]: #686
-#689 := [quant-intro #687]: #688
-#167 := (~ #124 #124)
-#166 := (~ #121 #121)
-#184 := [refl]: #166
-#168 := [nnf-pos #184]: #167
-#60 := (<= 0::Int #59)
-#64 := (implies #60 #63)
-#65 := (forall (vars (?v0 Int)) #64)
-#127 := (iff #65 #124)
-#108 := (not #60)
-#109 := (or #108 #63)
-#112 := (forall (vars (?v0 Int)) #109)
-#125 := (iff #112 #124)
-#122 := (iff #109 #121)
-#119 := (iff #108 #118)
-#115 := (iff #60 #117)
-#116 := [rewrite]: #115
-#120 := [monotonicity #116]: #119
-#123 := [monotonicity #120]: #122
-#126 := [quant-intro #123]: #125
-#113 := (iff #65 #112)
-#110 := (iff #64 #109)
-#111 := [rewrite]: #110
-#114 := [quant-intro #111]: #113
-#128 := [trans #114 #126]: #127
-#107 := [asserted]: #65
-#129 := [mp #107 #128]: #124
-#185 := [mp~ #129 #168]: #124
-#690 := [mp #185 #689]: #685
-#641 := (not #685)
-#623 := (or #641 #631)
-#360 := (>= 1::Int 0::Int)
-#361 := (not #360)
-#632 := (or #361 #631)
-#627 := (or #641 #632)
-#628 := (iff #627 #623)
-#618 := (iff #623 #623)
-#619 := [rewrite]: #618
-#626 := (iff #632 #631)
-#344 := (or false #631)
-#347 := (iff #344 #631)
-#625 := [rewrite]: #347
-#345 := (iff #632 #344)
-#630 := (iff #361 false)
-#1 := true
-#651 := (not true)
-#652 := (iff #651 false)
-#311 := [rewrite]: #652
-#629 := (iff #361 #651)
-#354 := (iff #360 true)
-#355 := [rewrite]: #354
-#633 := [monotonicity #355]: #629
-#634 := [trans #633 #311]: #630
-#346 := [monotonicity #634]: #345
-#340 := [trans #346 #625]: #626
-#617 := [monotonicity #340]: #628
-#614 := [trans #617 #619]: #628
-#624 := [quant-inst #14]: #627
-#615 := [mp #624 #614]: #623
-#433 := [unit-resolution #615 #690]: #631
-#438 := (not #631)
-#417 := (or #438 #515)
-#420 := [th-lemma arith triangle-eq]: #417
-#424 := [unit-resolution #420 #433]: #515
-#426 := (not #515)
-#427 := (or #426 #520)
-#425 := [th-lemma arith farkas 1 1]: #427
-#428 := [unit-resolution #425 #424]: #520
-#525 := (not #520)
-#482 := (or #641 #525 #452)
-#518 := (= #522 #580)
-#516 := (>= #580 0::Int)
-#517 := (not #516)
-#519 := (or #517 #518)
-#489 := (or #641 #519)
-#493 := (iff #489 #482)
-#513 := (or #525 #452)
-#479 := (or #641 #513)
-#490 := (iff #479 #482)
-#492 := [rewrite]: #490
-#481 := (iff #489 #479)
-#508 := (iff #519 #513)
-#506 := (iff #518 #452)
-#512 := [rewrite]: #506
-#521 := (iff #517 #525)
-#523 := (iff #516 #520)
-#524 := [rewrite]: #523
-#526 := [monotonicity #524]: #521
-#514 := [monotonicity #526 #512]: #508
-#483 := [monotonicity #514]: #481
-#494 := [trans #483 #492]: #493
-#448 := [quant-inst #580]: #489
-#504 := [mp #448 #494]: #482
-#416 := [unit-resolution #504 #690 #428]: #452
-#419 := (not #452)
-#421 := (or #419 #484)
-#422 := [th-lemma arith triangle-eq]: #421
-#418 := [unit-resolution #422 #416]: #484
-#505 := (>= #488 -1::Int)
-#423 := (or #419 #505)
-#413 := [th-lemma arith triangle-eq]: #423
-#403 := [unit-resolution #413 #416]: #505
-#404 := (<= #254 1::Int)
-#405 := (or #438 #404)
-#406 := [th-lemma arith triangle-eq]: #405
-#409 := [unit-resolution #406 #433]: #404
-#414 := [th-lemma arith eq-propagate -1 -1 1 1 #424 #409 #403 #418]: #410
-#415 := [symm #414]: #407
-#411 := [monotonicity #415]: #408
-#401 := [trans #411 #437]: #400
-#402 := [trans #401 #398]: #399
-#386 := [monotonicity #402]: #384
-#388 := [monotonicity #386 #436]: #387
-#392 := [trans #388 #390]: #391
-#351 := (= #45 #338)
-#350 := (= #44 #337)
-#658 := (f7 f8 #42)
-#586 := (+ 1::Int #658)
-#578 := (f5 f6 #586)
-#357 := (= #578 #337)
-#587 := (= #337 #578)
-#590 := (or #585 #587)
-#579 := [quant-inst #42]: #590
-#393 := [unit-resolution #579 #157]: #587
-#367 := [symm #393]: #357
-#348 := (= #44 #578)
-#570 := (f7 f8 #578)
-#447 := (f5 f6 #570)
-#449 := (= #447 #578)
-#454 := (or #453 #449)
-#455 := [quant-inst #578]: #454
-#394 := [unit-resolution #455 #683]: #449
-#365 := (= #44 #447)
-#364 := (= 1::Int #570)
-#362 := (= #570 1::Int)
-#564 := (* -1::Int #658)
-#565 := (+ #570 #564)
-#538 := (<= #565 1::Int)
-#562 := (= #565 1::Int)
-#573 := (>= #658 -1::Int)
-#589 := (>= #658 0::Int)
-#659 := (= #658 0::Int)
-#642 := (or #641 #659)
-#443 := (>= 0::Int 0::Int)
-#650 := (not #443)
-#660 := (or #650 #659)
-#643 := (or #641 #660)
-#644 := (iff #643 #642)
-#645 := (iff #642 #642)
-#647 := [rewrite]: #645
-#639 := (iff #660 #659)
-#637 := (or false #659)
-#301 := (iff #637 #659)
-#302 := [rewrite]: #301
-#299 := (iff #660 #637)
-#653 := (iff #650 false)
-#310 := (iff #650 #651)
-#655 := (iff #443 true)
-#661 := [rewrite]: #655
-#315 := [monotonicity #661]: #310
-#295 := [trans #315 #311]: #653
-#300 := [monotonicity #295]: #299
-#640 := [trans #300 #302]: #639
-#281 := [monotonicity #640]: #644
-#286 := [trans #281 #647]: #644
-#638 := [quant-inst #41]: #643
-#287 := [mp #638 #286]: #642
-#395 := [unit-resolution #287 #690]: #659
-#396 := (not #659)
-#385 := (or #396 #589)
-#397 := [th-lemma arith triangle-eq]: #385
-#374 := [unit-resolution #397 #395]: #589
-#376 := (not #589)
-#377 := (or #376 #573)
-#378 := [th-lemma arith farkas 1 1]: #377
-#379 := [unit-resolution #378 #374]: #573
-#560 := (not #573)
-#551 := (or #641 #560 #562)
-#571 := (= #570 #586)
-#576 := (>= #586 0::Int)
-#583 := (not #576)
-#572 := (or #583 #571)
-#552 := (or #641 #572)
-#548 := (iff #552 #551)
-#547 := (or #560 #562)
-#554 := (or #641 #547)
-#557 := (iff #554 #551)
-#558 := [rewrite]: #557
-#555 := (iff #552 #554)
-#549 := (iff #572 #547)
-#566 := (iff #571 #562)
-#546 := [rewrite]: #566
-#561 := (iff #583 #560)
-#569 := (iff #576 #573)
-#574 := [rewrite]: #569
-#563 := [monotonicity #574]: #561
-#550 := [monotonicity #563 #546]: #549
-#556 := [monotonicity #550]: #555
-#559 := [trans #556 #558]: #548
-#553 := [quant-inst #586]: #552
-#537 := [mp #553 #559]: #551
-#380 := [unit-resolution #537 #690 #379]: #562
-#381 := (not #562)
-#382 := (or #381 #538)
-#375 := [th-lemma arith triangle-eq]: #382
-#383 := [unit-resolution #375 #380]: #538
-#540 := (>= #565 1::Int)
-#368 := (or #381 #540)
-#369 := [th-lemma arith triangle-eq]: #368
-#370 := [unit-resolution #369 #380]: #540
-#588 := (<= #658 0::Int)
-#372 := (or #396 #588)
-#371 := [th-lemma arith triangle-eq]: #372
-#373 := [unit-resolution #371 #395]: #588
-#363 := [th-lemma arith eq-propagate -1 -1 -1 -1 #374 #373 #370 #383]: #362
-#356 := [symm #363]: #364
-#366 := [monotonicity #356]: #365
-#349 := [trans #366 #394]: #348
-#341 := [trans #349 #367]: #350
-#352 := [monotonicity #341]: #351
-#319 := [monotonicity #352 #392]: #353
-#322 := [symm #319]: #321
-#312 := (= #48 #656)
-#324 := (or #323 #312)
-#657 := [quant-inst #7 #42 #46]: #324
-#342 := [unit-resolution #657 #676]: #312
-#313 := [trans #342 #322]: #54
-#55 := (not #54)
-#105 := [asserted]: #55
-[unit-resolution #105 #313]: false
-unsat
-7a4c9001ff099c38b0602b196e3bc37f301b1551 24 0
-#2 := false
-decl f1 :: S1
-#3 := f1
-decl f3 :: (-> S2 S1)
-#7 := (:var 0 S2)
-#8 := (f3 #7)
-#9 := (= #8 f1)
-#10 := (forall (vars (?v0 S2)) #9)
-#11 := (not #10)
-#12 := (or #10 #11)
-#13 := (not #12)
-#42 := (iff #13 false)
-#1 := true
-#37 := (not true)
-#40 := (iff #37 false)
-#41 := [rewrite]: #40
-#38 := (iff #13 #37)
-#35 := (iff #12 true)
-#36 := [rewrite]: #35
-#39 := [monotonicity #36]: #38
-#43 := [trans #39 #41]: #42
-#34 := [asserted]: #13
-[mp #34 #43]: false
-unsat
-5e86b4c9726ef5b2868f22c9ea608e9e3558803e 344 0
-#2 := false
-decl f7 :: (-> S5 Int S2)
-#28 := 6::Int
-decl f8 :: S5
-#14 := f8
-#29 := (f7 f8 6::Int)
-decl f3 :: (-> S3 S2 S2)
-decl f5 :: (-> S4 S2 Int)
-#21 := 4::Int
-#22 := (f7 f8 4::Int)
-decl f4 :: S3
-#7 := f4
-#23 := (f3 f4 #22)
-decl f6 :: S4
-#10 := f6
-#24 := (f5 f6 #23)
-#25 := (* 4::Int #24)
-#26 := (f7 f8 #25)
-#27 := (f3 f4 #26)
-#30 := (= #27 #29)
-#526 := (f3 f4 #29)
-#490 := (= #526 #29)
-#552 := (f5 f6 #29)
-#67 := -10::Int
-#528 := (+ -10::Int #552)
-#508 := (f7 f8 #528)
-#454 := (f3 f4 #508)
-#509 := (= #526 #454)
-#12 := 10::Int
-#525 := (>= #552 10::Int)
-#514 := (if #525 #509 #490)
-#8 := (:var 0 S2)
-#9 := (f3 f4 #8)
-#665 := (pattern #9)
-#11 := (f5 f6 #8)
-#664 := (pattern #11)
-#182 := (= #9 #8)
-#68 := (+ -10::Int #11)
-#71 := (f7 f8 #68)
-#74 := (f3 f4 #71)
-#181 := (= #9 #74)
-#88 := (>= #11 10::Int)
-#169 := (if #88 #181 #182)
-#666 := (forall (vars (?v0 S2)) (:pat #664 #665) #169)
-#184 := (forall (vars (?v0 S2)) #169)
-#669 := (iff #184 #666)
-#667 := (iff #169 #169)
-#668 := [refl]: #667
-#670 := [quant-intro #668]: #669
-#93 := (if #88 #74 #8)
-#98 := (= #9 #93)
-#101 := (forall (vars (?v0 S2)) #98)
-#185 := (iff #101 #184)
-#170 := (iff #98 #169)
-#183 := [rewrite]: #170
-#186 := [quant-intro #183]: #185
-#173 := (~ #101 #101)
-#171 := (~ #98 #98)
-#172 := [refl]: #171
-#174 := [nnf-pos #172]: #173
-#15 := (- #11 10::Int)
-#16 := (f7 f8 #15)
-#17 := (f3 f4 #16)
-#13 := (< #11 10::Int)
-#18 := (if #13 #8 #17)
-#19 := (= #9 #18)
-#20 := (forall (vars (?v0 S2)) #19)
-#104 := (iff #20 #101)
-#77 := (if #13 #8 #74)
-#80 := (= #9 #77)
-#83 := (forall (vars (?v0 S2)) #80)
-#102 := (iff #83 #101)
-#99 := (iff #80 #98)
-#96 := (= #77 #93)
-#86 := (not #88)
-#90 := (if #86 #8 #74)
-#94 := (= #90 #93)
-#95 := [rewrite]: #94
-#91 := (= #77 #90)
-#87 := (iff #13 #86)
-#89 := [rewrite]: #87
-#92 := [monotonicity #89]: #91
-#97 := [trans #92 #95]: #96
-#100 := [monotonicity #97]: #99
-#103 := [quant-intro #100]: #102
-#84 := (iff #20 #83)
-#81 := (iff #19 #80)
-#78 := (= #18 #77)
-#75 := (= #17 #74)
-#72 := (= #16 #71)
-#69 := (= #15 #68)
-#70 := [rewrite]: #69
-#73 := [monotonicity #70]: #72
-#76 := [monotonicity #73]: #75
-#79 := [monotonicity #76]: #78
-#82 := [monotonicity #79]: #81
-#85 := [quant-intro #82]: #84
-#105 := [trans #85 #103]: #104
-#66 := [asserted]: #20
-#106 := [mp #66 #105]: #101
-#159 := [mp~ #106 #174]: #101
-#187 := [mp #159 #186]: #184
-#671 := [mp #187 #670]: #666
-#320 := (not #666)
-#516 := (or #320 #514)
-#484 := [quant-inst #29]: #516
-#469 := [unit-resolution #484 #671]: #514
-#450 := (not #525)
-#515 := (<= #552 6::Int)
-#553 := (= #552 6::Int)
-#36 := (:var 0 Int)
-#38 := (f7 f8 #36)
-#678 := (pattern #38)
-#39 := (f5 f6 #38)
-#40 := (= #39 #36)
-#35 := 0::Int
-#119 := (>= #36 0::Int)
-#120 := (not #119)
-#123 := (or #120 #40)
-#679 := (forall (vars (?v0 Int)) (:pat #678) #123)
-#126 := (forall (vars (?v0 Int)) #123)
-#682 := (iff #126 #679)
-#680 := (iff #123 #123)
-#681 := [refl]: #680
-#683 := [quant-intro #681]: #682
-#165 := (~ #126 #126)
-#164 := (~ #123 #123)
-#176 := [refl]: #164
-#166 := [nnf-pos #176]: #165
-#37 := (<= 0::Int #36)
-#41 := (implies #37 #40)
-#42 := (forall (vars (?v0 Int)) #41)
-#129 := (iff #42 #126)
-#110 := (not #37)
-#111 := (or #110 #40)
-#114 := (forall (vars (?v0 Int)) #111)
-#127 := (iff #114 #126)
-#124 := (iff #111 #123)
-#121 := (iff #110 #120)
-#117 := (iff #37 #119)
-#118 := [rewrite]: #117
-#122 := [monotonicity #118]: #121
-#125 := [monotonicity #122]: #124
-#128 := [quant-intro #125]: #127
-#115 := (iff #42 #114)
-#112 := (iff #41 #111)
-#113 := [rewrite]: #112
-#116 := [quant-intro #113]: #115
-#130 := [trans #116 #128]: #129
-#109 := [asserted]: #42
-#131 := [mp #109 #130]: #126
-#177 := [mp~ #131 #166]: #126
-#684 := [mp #177 #683]: #679
-#611 := (not #679)
-#545 := (or #611 #553)
-#549 := (>= 6::Int 0::Int)
-#551 := (not #549)
-#554 := (or #551 #553)
-#546 := (or #611 #554)
-#547 := (iff #546 #545)
-#529 := (iff #545 #545)
-#530 := [rewrite]: #529
-#543 := (iff #554 #553)
-#550 := (or false #553)
-#540 := (iff #550 #553)
-#542 := [rewrite]: #540
-#561 := (iff #554 #550)
-#559 := (iff #551 false)
-#1 := true
-#619 := (not true)
-#616 := (iff #619 false)
-#617 := [rewrite]: #616
-#557 := (iff #551 #619)
-#555 := (iff #549 true)
-#556 := [rewrite]: #555
-#558 := [monotonicity #556]: #557
-#560 := [trans #558 #617]: #559
-#539 := [monotonicity #560]: #561
-#544 := [trans #539 #542]: #543
-#533 := [monotonicity #544]: #547
-#531 := [trans #533 #530]: #547
-#541 := [quant-inst #28]: #546
-#534 := [mp #541 #531]: #545
-#470 := [unit-resolution #534 #684]: #553
-#477 := (not #553)
-#478 := (or #477 #515)
-#479 := [th-lemma arith triangle-eq]: #478
-#464 := [unit-resolution #479 #470]: #515
-#480 := (not #515)
-#441 := (or #480 #450)
-#442 := [th-lemma arith farkas 1 1]: #441
-#449 := [unit-resolution #442 #464]: #450
-#491 := (not #514)
-#485 := (or #491 #525 #490)
-#492 := [def-axiom]: #485
-#451 := [unit-resolution #492 #449 #469]: #490
-#404 := (= #27 #526)
-#641 := (f5 f6 #26)
-#638 := (+ -10::Int #641)
-#345 := (f7 f8 #638)
-#360 := (f3 f4 #345)
-#403 := (= #360 #526)
-#416 := (= #345 #29)
-#411 := (= #638 6::Int)
-#312 := (f5 f6 #22)
-#249 := -1::Int
-#518 := (* -1::Int #312)
-#519 := (+ #24 #518)
-#524 := (<= #519 0::Int)
-#517 := (= #24 #312)
-#303 := (= #23 #22)
-#297 := (+ -10::Int #312)
-#639 := (f7 f8 #297)
-#301 := (f3 f4 #639)
-#302 := (= #23 #301)
-#317 := (>= #312 10::Int)
-#304 := (if #317 #302 #303)
-#643 := (or #320 #304)
-#644 := [quant-inst #22]: #643
-#452 := [unit-resolution #644 #671]: #304
-#640 := (not #317)
-#447 := (<= #312 4::Int)
-#625 := (= #312 4::Int)
-#612 := (or #611 #625)
-#256 := (>= 4::Int 0::Int)
-#633 := (not #256)
-#629 := (or #633 #625)
-#606 := (or #611 #629)
-#613 := (iff #606 #612)
-#608 := (iff #612 #612)
-#615 := [rewrite]: #608
-#609 := (iff #629 #625)
-#618 := (or false #625)
-#466 := (iff #618 #625)
-#467 := [rewrite]: #466
-#624 := (iff #629 #618)
-#622 := (iff #633 false)
-#620 := (iff #633 #619)
-#626 := (iff #256 true)
-#630 := [rewrite]: #626
-#621 := [monotonicity #630]: #620
-#623 := [trans #621 #617]: #622
-#465 := [monotonicity #623]: #624
-#610 := [trans #465 #467]: #609
-#614 := [monotonicity #610]: #613
-#444 := [trans #614 #615]: #613
-#607 := [quant-inst #21]: #606
-#446 := [mp #607 #444]: #612
-#453 := [unit-resolution #446 #684]: #625
-#455 := (not #625)
-#456 := (or #455 #447)
-#457 := [th-lemma arith triangle-eq]: #456
-#458 := [unit-resolution #457 #453]: #447
-#459 := (not #447)
-#460 := (or #459 #640)
-#443 := [th-lemma arith farkas 1 1]: #460
-#461 := [unit-resolution #443 #458]: #640
-#645 := (not #304)
-#647 := (or #645 #317 #303)
-#649 := [def-axiom]: #647
-#431 := [unit-resolution #649 #461 #452]: #303
-#434 := [monotonicity #431]: #517
-#436 := (not #517)
-#437 := (or #436 #524)
-#438 := [th-lemma arith triangle-eq]: #437
-#280 := [unit-resolution #438 #434]: #524
-#520 := (>= #519 0::Int)
-#439 := (or #436 #520)
-#435 := [th-lemma arith triangle-eq]: #439
-#440 := [unit-resolution #435 #434]: #520
-#600 := (>= #312 4::Int)
-#419 := (or #455 #600)
-#422 := [th-lemma arith triangle-eq]: #419
-#426 := [unit-resolution #422 #453]: #600
-#504 := (* -1::Int #641)
-#505 := (+ #25 #504)
-#582 := (<= #505 0::Int)
-#503 := (= #505 0::Int)
-#597 := (>= #24 0::Int)
-#429 := (not #520)
-#428 := (not #600)
-#427 := (or #597 #428 #429)
-#430 := [th-lemma arith assign-bounds 1 1]: #427
-#418 := [unit-resolution #430 #426 #440]: #597
-#499 := (not #597)
-#598 := (or #499 #503)
-#586 := (or #611 #499 #503)
-#593 := (= #641 #25)
-#596 := (>= #25 0::Int)
-#498 := (not #596)
-#594 := (or #498 #593)
-#588 := (or #611 #594)
-#587 := (iff #588 #586)
-#577 := (or #611 #598)
-#590 := (iff #577 #586)
-#591 := [rewrite]: #590
-#579 := (iff #588 #577)
-#595 := (iff #594 #598)
-#501 := (iff #593 #503)
-#502 := [rewrite]: #501
-#500 := (iff #498 #499)
-#482 := (iff #596 #597)
-#497 := [rewrite]: #482
-#493 := [monotonicity #497]: #500
-#599 := [monotonicity #493 #502]: #595
-#589 := [monotonicity #599]: #579
-#592 := [trans #589 #591]: #587
-#580 := [quant-inst #25]: #588
-#581 := [mp #580 #592]: #586
-#421 := [unit-resolution #581 #684]: #598
-#423 := [unit-resolution #421 #418]: #503
-#424 := (not #503)
-#420 := (or #424 #582)
-#425 := [th-lemma arith triangle-eq]: #420
-#415 := [unit-resolution #425 #423]: #582
-#583 := (>= #505 0::Int)
-#405 := (or #424 #583)
-#407 := [th-lemma arith triangle-eq]: #405
-#408 := [unit-resolution #407 #423]: #583
-#412 := [th-lemma arith eq-propagate 1 1 -4 -4 -4 -4 #408 #415 #426 #458 #440 #280]: #411
-#409 := [monotonicity #412]: #416
-#401 := [monotonicity #409]: #403
-#361 := (= #27 #360)
-#362 := (= #27 #26)
-#642 := (>= #641 10::Int)
-#363 := (if #642 #361 #362)
-#634 := (or #320 #363)
-#356 := [quant-inst #26]: #634
-#417 := [unit-resolution #356 #671]: #363
-#410 := (not #582)
-#413 := (or #642 #410 #428 #429)
-#414 := [th-lemma arith assign-bounds 1 4 4]: #413
-#400 := [unit-resolution #414 #426 #415 #440]: #642
-#631 := (not #642)
-#357 := (not #363)
-#635 := (or #357 #631 #361)
-#632 := [def-axiom]: #635
-#402 := [unit-resolution #632 #400 #417]: #361
-#386 := [trans #402 #401]: #404
-#388 := [trans #386 #451]: #30
-#31 := (not #30)
-#107 := [asserted]: #31
-[unit-resolution #107 #388]: false
-unsat
-013f2c4f5eccbaac1754336d2ce477a569c8d0cd 1 0
-unsat
-8954c874576a1a34e48535e83e9151ff299d36aa 95 0
-#2 := false
-decl f1 :: S1
-#3 := f1
-decl f3 :: (-> S3 S2 S1)
-decl f10 :: (-> S5 S6 S2)
-decl f12 :: (-> S3 S6)
-decl f6 :: S3
-#19 := f6
-#43 := (f12 f6)
-decl f11 :: S5
-#42 := f11
-#44 := (f10 f11 #43)
-decl f8 :: (-> S4 S2 S3)
-decl f9 :: S4
-#29 := f9
-#45 := (f8 f9 #44)
-#53 := (f3 #45 #44)
-#54 := (= #53 f1)
-#55 := (not #54)
-#140 := [asserted]: #55
-decl f4 :: S3
-#7 := f4
-#46 := (f12 f4)
-#47 := (f10 f11 #46)
-#50 := (f8 f9 #47)
-#51 := (f3 #50 #44)
-#52 := (= #51 f1)
-#139 := [asserted]: #52
-#48 := (f3 #45 #47)
-#49 := (= #48 f1)
-#138 := [asserted]: #49
-#8 := (:var 0 S2)
-#12 := (:var 1 S2)
-#34 := (f8 f9 #12)
-#35 := (f3 #34 #8)
-#30 := (:var 2 S2)
-#31 := (f8 f9 #30)
-#32 := (f3 #31 #12)
-#635 := (pattern #32 #35)
-#37 := (f3 #31 #8)
-#38 := (= #37 f1)
-#36 := (= #35 f1)
-#112 := (not #36)
-#33 := (= #32 f1)
-#120 := (not #33)
-#129 := (or #120 #112 #38)
-#636 := (forall (vars (?v0 S2) (?v1 S2) (?v2 S2)) (:pat #635) #129)
-#132 := (forall (vars (?v0 S2) (?v1 S2) (?v2 S2)) #129)
-#639 := (iff #132 #636)
-#637 := (iff #129 #129)
-#638 := [refl]: #637
-#640 := [quant-intro #638]: #639
-#146 := (~ #132 #132)
-#162 := (~ #129 #129)
-#163 := [refl]: #162
-#147 := [nnf-pos #163]: #146
-#39 := (implies #36 #38)
-#40 := (implies #33 #39)
-#41 := (forall (vars (?v0 S2) (?v1 S2) (?v2 S2)) #40)
-#135 := (iff #41 #132)
-#114 := (or #112 #38)
-#121 := (or #120 #114)
-#126 := (forall (vars (?v0 S2) (?v1 S2) (?v2 S2)) #121)
-#133 := (iff #126 #132)
-#130 := (iff #121 #129)
-#131 := [rewrite]: #130
-#134 := [quant-intro #131]: #133
-#127 := (iff #41 #126)
-#124 := (iff #40 #121)
-#117 := (implies #33 #114)
-#122 := (iff #117 #121)
-#123 := [rewrite]: #122
-#118 := (iff #40 #117)
-#115 := (iff #39 #114)
-#116 := [rewrite]: #115
-#119 := [monotonicity #116]: #118
-#125 := [trans #119 #123]: #124
-#128 := [quant-intro #125]: #127
-#136 := [trans #128 #134]: #135
-#111 := [asserted]: #41
-#137 := [mp #111 #136]: #132
-#164 := [mp~ #137 #147]: #132
-#641 := [mp #164 #640]: #636
-#305 := (not #52)
-#219 := (not #49)
-#307 := (not #636)
-#298 := (or #307 #219 #305 #54)
-#220 := (or #219 #305 #54)
-#309 := (or #307 #220)
-#311 := (iff #309 #298)
-#308 := [rewrite]: #311
-#310 := [quant-inst #44 #47 #44]: #309
-#312 := [mp #310 #308]: #298
-[unit-resolution #312 #641 #138 #139 #140]: false
-unsat
-b66bf263776a429b47555990b2282b5f0c94c465 59 0
-#2 := false
-decl f1 :: S1
-#3 := f1
-decl f22 :: (-> Int S1)
-#70 := 1::Int
-#71 := (f22 1::Int)
-#72 := (= #71 f1)
-#73 := (not #72)
-#163 := [asserted]: #73
-#57 := (:var 0 Int)
-#58 := (f22 #57)
-#695 := (pattern #58)
-#59 := (= #58 f1)
-#696 := (forall (vars (?v0 Int)) (:pat #695) #59)
-#160 := (forall (vars (?v0 Int)) #59)
-#699 := (iff #160 #696)
-#697 := (iff #59 #59)
-#698 := [refl]: #697
-#700 := [quant-intro #698]: #699
-#174 := (~ #160 #160)
-#192 := (~ #59 #59)
-#193 := [refl]: #192
-#175 := [nnf-pos #193]: #174
-decl f17 :: (-> S10 S1)
-decl f23 :: (-> S13 S10 S10)
-decl f26 :: S10
-#62 := f26
-decl f24 :: (-> S14 Int S13)
-decl f25 :: S14
-#60 := f25
-#61 := (f24 f25 #57)
-#63 := (f23 #61 f26)
-#64 := (f17 #63)
-#65 := (= #64 f1)
-#66 := (not #65)
-#67 := (or #65 #66)
-#68 := (and #59 #67)
-#69 := (forall (vars (?v0 Int)) #68)
-#161 := (iff #69 #160)
-#158 := (iff #68 #59)
-#1 := true
-#153 := (and #59 true)
-#156 := (iff #153 #59)
-#157 := [rewrite]: #156
-#154 := (iff #68 #153)
-#150 := (iff #67 true)
-#152 := [rewrite]: #150
-#155 := [monotonicity #152]: #154
-#159 := [trans #155 #157]: #158
-#162 := [quant-intro #159]: #161
-#149 := [asserted]: #69
-#165 := [mp #149 #162]: #160
-#194 := [mp~ #165 #175]: #160
-#701 := [mp #194 #700]: #696
-#253 := (not #696)
-#338 := (or #253 #72)
-#339 := [quant-inst #70]: #338
-[unit-resolution #339 #701 #163]: false
-unsat
 d9c8c0d6c38991be073d0ed9988535642e4f47a6 396 0
 #2 := false
 decl f12 :: (-> S9 S10 S4)
@@ -12208,3 +2065,404 @@
 #247 := [asserted]: #123
 [unit-resolution #247 #607]: false
 unsat
+c4f4c8220660d1979009b33a643f0927bee816b1 1 0
+unsat
+db6426d59fdd57da8ca5d11de399761d1f1443de 1 0
+unsat
+e7ef76d73ccb9bc09d2b5368495a7a59d1bae3dc 1 0
+unsat
+8578dab7bf88c7d119f9af2e5f7eaf948f1bdb87 84 0
+WARNING: failed to find a pattern for quantifier (quantifier id: k!10)
+#2 := false
+#8 := 0::Int
+#7 := (:var 0 Int)
+#49 := (<= #7 0::Int)
+#50 := (not #49)
+#47 := (>= #7 0::Int)
+#45 := (not #47)
+#53 := (or #45 #50)
+#56 := (forall (vars (?v0 Int)) #53)
+#525 := (not #56)
+#218 := (<= 0::Int 0::Int)
+#539 := (not #218)
+#207 := (>= 0::Int 0::Int)
+#201 := (not #207)
+#537 := (or #201 #539)
+#526 := (or #525 #537)
+#170 := (iff #526 #525)
+#527 := (or #525 false)
+#530 := (iff #527 #525)
+#169 := [rewrite]: #530
+#164 := (iff #526 #527)
+#523 := (iff #537 false)
+#182 := (or false false)
+#185 := (iff #182 false)
+#522 := [rewrite]: #185
+#183 := (iff #537 #182)
+#178 := (iff #539 false)
+#1 := true
+#543 := (not true)
+#222 := (iff #543 false)
+#544 := [rewrite]: #222
+#194 := (iff #539 #543)
+#198 := (iff #218 true)
+#535 := [rewrite]: #198
+#536 := [monotonicity #535]: #194
+#520 := [trans #536 #544]: #178
+#534 := (iff #201 false)
+#538 := (iff #201 #543)
+#541 := (iff #207 true)
+#542 := [rewrite]: #541
+#326 := [monotonicity #542]: #538
+#193 := [trans #326 #544]: #534
+#184 := [monotonicity #193 #520]: #183
+#524 := [trans #184 #522]: #523
+#528 := [monotonicity #524]: #164
+#531 := [trans #528 #169]: #170
+#521 := [quant-inst #8]: #526
+#529 := [mp #521 #531]: #525
+#69 := (~ #56 #56)
+#67 := (~ #53 #53)
+#68 := [refl]: #67
+#70 := [nnf-pos #68]: #69
+#10 := (< 0::Int #7)
+#9 := (< #7 0::Int)
+#11 := (or #9 #10)
+#12 := (forall (vars (?v0 Int)) #11)
+#13 := (if #12 false true)
+#14 := (not #13)
+#59 := (iff #14 #56)
+#57 := (iff #12 #56)
+#54 := (iff #11 #53)
+#51 := (iff #10 #50)
+#52 := [rewrite]: #51
+#46 := (iff #9 #45)
+#48 := [rewrite]: #46
+#55 := [monotonicity #48 #52]: #54
+#58 := [quant-intro #55]: #57
+#43 := (iff #14 #12)
+#35 := (not #12)
+#38 := (not #35)
+#41 := (iff #38 #12)
+#42 := [rewrite]: #41
+#39 := (iff #14 #38)
+#36 := (iff #13 #35)
+#37 := [rewrite]: #36
+#40 := [monotonicity #37]: #39
+#44 := [trans #40 #42]: #43
+#60 := [trans #44 #58]: #59
+#34 := [asserted]: #14
+#61 := [mp #34 #60]: #56
+#63 := [mp~ #61 #70]: #56
+[unit-resolution #63 #529]: false
+unsat
+252d255c564463d916bc68156eea8dbe7fb0be0a 165 0
+WARNING: failed to find a pattern for quantifier (quantifier id: k!10)
+#2 := false
+#7 := 0::Int
+#8 := (:var 0 Int)
+#55 := (<= #8 0::Int)
+#56 := (not #55)
+#52 := (>= #8 0::Int)
+#51 := (not #52)
+#59 := (or #51 #56)
+#62 := (forall (vars (?v0 Int)) #59)
+#95 := (not #62)
+#587 := (<= 0::Int 0::Int)
+#586 := (not #587)
+#585 := (>= 0::Int 0::Int)
+#248 := (not #585)
+#593 := (or #248 #586)
+#290 := (or #95 #593)
+#569 := (iff #290 #95)
+#292 := (or #95 false)
+#572 := (iff #292 #95)
+#287 := [rewrite]: #572
+#293 := (iff #290 #292)
+#576 := (iff #593 false)
+#578 := (or false false)
+#575 := (iff #578 false)
+#579 := [rewrite]: #575
+#300 := (iff #593 #578)
+#201 := (iff #586 false)
+#1 := true
+#594 := (not true)
+#592 := (iff #594 false)
+#595 := [rewrite]: #592
+#306 := (iff #586 #594)
+#304 := (iff #587 true)
+#305 := [rewrite]: #304
+#307 := [monotonicity #305]: #306
+#577 := [trans #307 #595]: #201
+#581 := (iff #248 false)
+#589 := (iff #248 #594)
+#233 := (iff #585 true)
+#234 := [rewrite]: #233
+#249 := [monotonicity #234]: #589
+#582 := [trans #249 #595]: #581
+#301 := [monotonicity #582 #577]: #300
+#580 := [trans #301 #579]: #576
+#571 := [monotonicity #580]: #293
+#573 := [trans #571 #287]: #569
+#291 := [quant-inst #7]: #290
+#570 := [mp #291 #573]: #95
+decl z3name!0 :: bool
+#92 := z3name!0
+#15 := 3::Int
+#39 := -1::Int
+#99 := (if z3name!0 -1::Int 3::Int)
+#284 := (= #99 3::Int)
+#604 := (not #284)
+#602 := (>= #99 3::Int)
+#259 := (not #602)
+#102 := (<= #99 0::Int)
+#65 := (if #62 -1::Int 3::Int)
+#71 := (<= #65 0::Int)
+#103 := (~ #71 #102)
+#100 := (= #65 #99)
+#97 := (~ #62 z3name!0)
+#88 := (or z3name!0 #95)
+#93 := (not z3name!0)
+#94 := (or #93 #62)
+#89 := (and #94 #88)
+#96 := [intro-def]: #89
+#98 := [apply-def #96]: #97
+#101 := [monotonicity #98]: #100
+#104 := [monotonicity #101]: #103
+#13 := 1::Int
+#14 := (- 1::Int)
+#10 := (< 0::Int #8)
+#9 := (< #8 0::Int)
+#11 := (or #9 #10)
+#12 := (forall (vars (?v0 Int)) #11)
+#16 := (if #12 #14 3::Int)
+#17 := (< 0::Int #16)
+#18 := (not #17)
+#84 := (iff #18 #71)
+#42 := (if #12 -1::Int 3::Int)
+#45 := (< 0::Int #42)
+#48 := (not #45)
+#82 := (iff #48 #71)
+#72 := (not #71)
+#77 := (not #72)
+#80 := (iff #77 #71)
+#81 := [rewrite]: #80
+#78 := (iff #48 #77)
+#75 := (iff #45 #72)
+#68 := (< 0::Int #65)
+#73 := (iff #68 #72)
+#74 := [rewrite]: #73
+#69 := (iff #45 #68)
+#66 := (= #42 #65)
+#63 := (iff #12 #62)
+#60 := (iff #11 #59)
+#57 := (iff #10 #56)
+#58 := [rewrite]: #57
+#53 := (iff #9 #51)
+#54 := [rewrite]: #53
+#61 := [monotonicity #54 #58]: #60
+#64 := [quant-intro #61]: #63
+#67 := [monotonicity #64]: #66
+#70 := [monotonicity #67]: #69
+#76 := [trans #70 #74]: #75
+#79 := [monotonicity #76]: #78
+#83 := [trans #79 #81]: #82
+#49 := (iff #18 #48)
+#46 := (iff #17 #45)
+#43 := (= #16 #42)
+#40 := (= #14 -1::Int)
+#41 := [rewrite]: #40
+#44 := [monotonicity #41]: #43
+#47 := [monotonicity #44]: #46
+#50 := [monotonicity #47]: #49
+#85 := [trans #50 #83]: #84
+#38 := [asserted]: #18
+#86 := [mp #38 #85]: #71
+#133 := [mp~ #86 #104]: #102
+#389 := (not #102)
+#596 := (or #259 #389)
+#270 := [th-lemma arith farkas 1 1]: #596
+#271 := [unit-resolution #270 #133]: #259
+#603 := [hypothesis]: #284
+#605 := (or #604 #602)
+#606 := [th-lemma arith triangle-eq]: #605
+#601 := [unit-resolution #606 #603 #271]: false
+#607 := [lemma #601]: #604
+#286 := (or z3name!0 #284)
+#265 := [def-axiom]: #286
+#574 := [unit-resolution #265 #607]: z3name!0
+decl ?v0!1 :: Int
+#115 := ?v0!1
+#118 := (<= ?v0!1 0::Int)
+#119 := (not #118)
+#116 := (>= ?v0!1 0::Int)
+#117 := (not #116)
+#120 := (or #117 #119)
+#121 := (not #120)
+#126 := (or z3name!0 #121)
+#129 := (and #94 #126)
+#130 := (~ #89 #129)
+#127 := (~ #88 #126)
+#122 := (~ #95 #121)
+#123 := [sk]: #122
+#113 := (~ z3name!0 z3name!0)
+#114 := [refl]: #113
+#128 := [monotonicity #114 #123]: #127
+#111 := (~ #94 #94)
+#109 := (~ #62 #62)
+#107 := (~ #59 #59)
+#108 := [refl]: #107
+#110 := [nnf-pos #108]: #109
+#105 := (~ #93 #93)
+#106 := [refl]: #105
+#112 := [monotonicity #106 #110]: #111
+#131 := [monotonicity #112 #128]: #130
+#132 := [mp~ #96 #131]: #129
+#136 := [and-elim #132]: #94
+#563 := [unit-resolution #136 #574]: #62
+[unit-resolution #563 #570]: false
+unsat
+8a78832884e41117489fba88c88de0b5cacb832a 143 0
+#2 := false
+#10 := 0::Int
+#8 := (:var 0 Int)
+#68 := (<= #8 0::Int)
+#69 := (not #68)
+#146 := (not false)
+#149 := (or #146 #69)
+#152 := (not #149)
+#155 := (forall (vars (?v0 Int)) #152)
+#182 := (iff #155 false)
+#177 := (forall (vars (?v0 Int)) false)
+#180 := (iff #177 false)
+#181 := [elim-unused]: #180
+#178 := (iff #155 #177)
+#175 := (iff #152 false)
+#1 := true
+#170 := (not true)
+#173 := (iff #170 false)
+#174 := [rewrite]: #173
+#171 := (iff #152 #170)
+#168 := (iff #149 true)
+#163 := (or true #69)
+#166 := (iff #163 true)
+#167 := [rewrite]: #166
+#164 := (iff #149 #163)
+#161 := (iff #146 true)
+#162 := [rewrite]: #161
+#165 := [monotonicity #162]: #164
+#169 := [trans #165 #167]: #168
+#172 := [monotonicity #169]: #171
+#176 := [trans #172 #174]: #175
+#179 := [quant-intro #176]: #178
+#183 := [trans #179 #181]: #182
+#59 := -1::Int
+#60 := (* -1::Int #8)
+#7 := (:var 1 Int)
+#61 := (+ #7 #60)
+#62 := (<= #61 0::Int)
+#65 := (not #62)
+#72 := (or #65 #69)
+#75 := (forall (vars (?v1 Int)) #72)
+#78 := (not #75)
+#81 := (or #78 #69)
+#107 := (not #81)
+#125 := (forall (vars (?v0 Int)) #107)
+#158 := (iff #125 #155)
+#129 := (forall (vars (?v1 Int)) #69)
+#132 := (not #129)
+#135 := (or #132 #69)
+#138 := (not #135)
+#141 := (forall (vars (?v0 Int)) #138)
+#156 := (iff #141 #155)
+#157 := [rewrite]: #156
+#142 := (iff #125 #141)
+#143 := [rewrite]: #142
+#159 := [trans #143 #157]: #158
+#118 := (and #75 #68)
+#121 := (forall (vars (?v0 Int)) #118)
+#126 := (iff #121 #125)
+#115 := (iff #118 #107)
+#124 := [rewrite]: #115
+#127 := [quant-intro #124]: #126
+#103 := (not #69)
+#106 := (and #75 #103)
+#110 := (forall (vars (?v0 Int)) #106)
+#122 := (iff #110 #121)
+#119 := (iff #106 #118)
+#116 := (iff #103 #68)
+#117 := [rewrite]: #116
+#120 := [monotonicity #117]: #119
+#123 := [quant-intro #120]: #122
+#84 := (exists (vars (?v0 Int)) #81)
+#87 := (not #84)
+#111 := (~ #87 #110)
+#108 := (~ #107 #106)
+#104 := (~ #103 #103)
+#105 := [refl]: #104
+#94 := (not #78)
+#95 := (~ #94 #75)
+#100 := (~ #75 #75)
+#98 := (~ #72 #72)
+#99 := [refl]: #98
+#101 := [nnf-pos #99]: #100
+#102 := [nnf-neg #101]: #95
+#109 := [nnf-neg #102 #105]: #108
+#112 := [nnf-neg #109]: #111
+#11 := (< 0::Int #8)
+#9 := (<= #7 #8)
+#12 := (implies #9 #11)
+#13 := (forall (vars (?v1 Int)) #12)
+#14 := (implies #13 #11)
+#15 := (exists (vars (?v0 Int)) #14)
+#16 := (not #15)
+#90 := (iff #16 #87)
+#37 := (not #9)
+#38 := (or #37 #11)
+#41 := (forall (vars (?v1 Int)) #38)
+#47 := (not #41)
+#48 := (or #47 #11)
+#53 := (exists (vars (?v0 Int)) #48)
+#56 := (not #53)
+#88 := (iff #56 #87)
+#85 := (iff #53 #84)
+#82 := (iff #48 #81)
+#70 := (iff #11 #69)
+#71 := [rewrite]: #70
+#79 := (iff #47 #78)
+#76 := (iff #41 #75)
+#73 := (iff #38 #72)
+#66 := (iff #37 #65)
+#63 := (iff #9 #62)
+#64 := [rewrite]: #63
+#67 := [monotonicity #64]: #66
+#74 := [monotonicity #67 #71]: #73
+#77 := [quant-intro #74]: #76
+#80 := [monotonicity #77]: #79
+#83 := [monotonicity #80 #71]: #82
+#86 := [quant-intro #83]: #85
+#89 := [monotonicity #86]: #88
+#57 := (iff #16 #56)
+#54 := (iff #15 #53)
+#51 := (iff #14 #48)
+#44 := (implies #41 #11)
+#49 := (iff #44 #48)
+#50 := [rewrite]: #49
+#45 := (iff #14 #44)
+#42 := (iff #13 #41)
+#39 := (iff #12 #38)
+#40 := [rewrite]: #39
+#43 := [quant-intro #40]: #42
+#46 := [monotonicity #43]: #45
+#52 := [trans #46 #50]: #51
+#55 := [quant-intro #52]: #54
+#58 := [monotonicity #55]: #57
+#91 := [trans #58 #89]: #90
+#36 := [asserted]: #16
+#92 := [mp #36 #91]: #87
+#113 := [mp~ #92 #112]: #110
+#114 := [mp #113 #123]: #121
+#128 := [mp #114 #127]: #125
+#160 := [mp #128 #159]: #155
+[mp #160 #183]: false
+unsat
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SMT_Examples/SMT_Examples.certs2	Fri Mar 14 17:32:11 2014 +0100
@@ -0,0 +1,4557 @@
+7a16ef230bca5702aa346494226903ec25809d32 6 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let ((@x28 (rewrite (= (not true) false))))
+(mp (asserted (not true)) @x28 false))))
+
+27731fc512042f0ea1785a47796a8bfd64c4a8cf 7 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let ((@x34 (monotonicity (rewrite (= (or |p$| (not |p$|)) true)) (= (not (or |p$| (not |p$|))) (not true)))))
+(let ((@x38 (trans @x34 (rewrite (= (not true) false)) (= (not (or |p$| (not |p$|))) false))))
+(mp (asserted (not (or |p$| (not |p$|)))) @x38 false)))))
+
+5330fb77bfecb903300c8a50f577df102088abaa 9 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let ((@x34 (monotonicity (rewrite (= (and |p$| true) |p$|)) (= (= (and |p$| true) |p$|) (= |p$| |p$|)))))
+(let ((@x38 (trans @x34 (rewrite (= (= |p$| |p$|) true)) (= (= (and |p$| true) |p$|) true))))
+(let ((@x41 (monotonicity @x38 (= (not (= (and |p$| true) |p$|)) (not true)))))
+(let ((@x45 (trans @x41 (rewrite (= (not true) false)) (= (not (= (and |p$| true) |p$|)) false))))
+(mp (asserted (not (= (and |p$| true) |p$|))) @x45 false)))))))
+
+c2e74b12f4c731d0ea3ac811d94ac5a723029e93 13 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let (($x8 (not |p$|)))
+(let (($x7 (or |p$| |q$|)))
+(let (($x9 (and $x7 $x8)))
+(let ((@x39 (monotonicity (rewrite (= (=> $x9 |q$|) (or (not $x9) |q$|))) (= (not (=> $x9 |q$|)) (not (or (not $x9) |q$|))))))
+(let ((@x40 (|not-or-elim| (mp (asserted (not (=> $x9 |q$|))) @x39 (not (or (not $x9) |q$|))) $x9)))
+(let ((@x43 (|and-elim| @x40 $x8)))
+(let ((@x45 (|not-or-elim| (mp (asserted (not (=> $x9 |q$|))) @x39 (not (or (not $x9) |q$|))) (not |q$|))))
+(let ((@x41 (|and-elim| @x40 $x7)))
+(|unit-resolution| @x41 @x45 @x43 false)))))))))))
+
+800409db22b453674c1b66520bda2d5bafbf81b4 11 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let (($x10 (and |c$| |d$|)))
+(let (($x7 (and |a$| |b$|)))
+(let (($x11 (or $x7 $x10)))
+(let (($x12 (=> $x11 $x11)))
+(let (($x13 (not $x12)))
+(let ((@x43 (trans (monotonicity (rewrite (= $x12 true)) (= $x13 (not true))) (rewrite (= (not true) false)) (= $x13 false))))
+(mp (asserted $x13) @x43 false)))))))))
+
+8ba22a36afac456bfdc7db71e8b371143686dc86 23 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let (($x11 (and |p1$| |p3$|)))
+(let (($x10 (and |p3$| |p2$|)))
+(let (($x12 (or $x10 $x11)))
+(let (($x13 (=> |p1$| $x12)))
+(let (($x14 (or $x13 |p1$|)))
+(let (($x7 (and |p1$| |p2$|)))
+(let (($x9 (or $x7 |p3$|)))
+(let (($x15 (=> $x9 $x14)))
+(let (($x16 (not $x15)))
+(let (($x38 (not |p1$|)))
+(let (($x39 (or $x38 $x12)))
+(let (($x42 (or $x39 |p1$|)))
+(let (($x48 (not $x9)))
+(let (($x49 (or $x48 $x42)))
+(let (($x54 (not $x49)))
+(let ((@x65 (trans (monotonicity (rewrite (= $x49 true)) (= $x54 (not true))) (rewrite (= (not true) false)) (= $x54 false))))
+(let ((@x47 (monotonicity (monotonicity (rewrite (= $x13 $x39)) (= $x14 $x42)) (= $x15 (=> $x9 $x42)))))
+(let ((@x56 (monotonicity (trans @x47 (rewrite (= (=> $x9 $x42) $x49)) (= $x15 $x49)) (= $x16 $x54))))
+(mp (asserted $x16) (trans @x56 @x65 (= $x16 false)) false)))))))))))))))))))))
+
+9d0d2643780c0052a3bf06c1fd96112084da5890 24 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let (($x6 (= |p$| |p$|)))
+(let (($x7 (= $x6 |p$|)))
+(let (($x8 (= $x7 |p$|)))
+(let (($x9 (= $x8 |p$|)))
+(let (($x10 (= $x9 |p$|)))
+(let (($x11 (= $x10 |p$|)))
+(let (($x12 (= $x11 |p$|)))
+(let (($x13 (= $x12 |p$|)))
+(let (($x14 (= $x13 |p$|)))
+(let (($x15 (not $x14)))
+(let ((@x38 (rewrite (= $x6 true))))
+(let ((@x43 (rewrite (= (= true |p$|) |p$|))))
+(let ((@x45 (trans (monotonicity @x38 (= $x7 (= true |p$|))) @x43 (= $x7 |p$|))))
+(let ((@x51 (monotonicity (trans (monotonicity @x45 (= $x8 $x6)) @x38 (= $x8 true)) (= $x9 (= true |p$|)))))
+(let ((@x57 (trans (monotonicity (trans @x51 @x43 (= $x9 |p$|)) (= $x10 $x6)) @x38 (= $x10 true))))
+(let ((@x61 (trans (monotonicity @x57 (= $x11 (= true |p$|))) @x43 (= $x11 |p$|))))
+(let ((@x67 (monotonicity (trans (monotonicity @x61 (= $x12 $x6)) @x38 (= $x12 true)) (= $x13 (= true |p$|)))))
+(let ((@x73 (trans (monotonicity (trans @x67 @x43 (= $x13 |p$|)) (= $x14 $x6)) @x38 (= $x14 true))))
+(let ((@x80 (trans (monotonicity @x73 (= $x15 (not true))) (rewrite (= (not true) false)) (= $x15 false))))
+(mp (asserted $x15) @x80 false))))))))))))))))))))))
+
+63439e1fd6656fc5a2376d7e5f00d0dd92c536a2 34 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let (($x98 (not |b$|)))
+(let (($x17 (not |c$|)))
+(let (($x36 (or |p$| (and |q$| (not |q$|)))))
+(let (($x37 (and (not |p$|) $x36)))
+(let (($x38 (or |c$| $x37)))
+(let (($x39 (not $x38)))
+(let ((@x120 (monotonicity (rewrite (= (and |q$| (not |q$|)) false)) (= $x36 (or |p$| false)))))
+(let ((@x127 (monotonicity (trans @x120 (rewrite (= (or |p$| false) |p$|)) (= $x36 |p$|)) (= $x37 (and (not |p$|) |p$|)))))
+(let ((@x131 (trans @x127 (rewrite (= (and (not |p$|) |p$|) false)) (= $x37 false))))
+(let ((@x138 (trans (monotonicity @x131 (= $x38 (or |c$| false))) (rewrite (= (or |c$| false) |c$|)) (= $x38 |c$|))))
+(let ((@x143 (mp (asserted $x39) (monotonicity @x138 (= $x39 $x17)) $x17)))
+(let (($x101 (or $x98 |c$|)))
+(let ((@x93 (monotonicity (rewrite (= (or |x$| (not |x$|)) true)) (= (and |b$| (or |x$| (not |x$|))) (and |b$| true)))))
+(let ((@x97 (trans @x93 (rewrite (= (and |b$| true) |b$|)) (= (and |b$| (or |x$| (not |x$|))) |b$|))))
+(let ((@x103 (monotonicity (monotonicity @x97 (= (not (and |b$| (or |x$| (not |x$|)))) $x98)) (= (or (not (and |b$| (or |x$| (not |x$|)))) |c$|) $x101))))
+(let ((@x106 (mp (asserted (or (not (and |b$| (or |x$| (not |x$|)))) |c$|)) @x103 $x101)))
+(let (($x108 (not |d$|)))
+(let (($x111 (or $x108 |c$|)))
+(let ((@x110 (monotonicity (rewrite (= (or |d$| false) |d$|)) (= (not (or |d$| false)) $x108))))
+(let ((@x116 (mp (asserted (or (not (or |d$| false)) |c$|)) (monotonicity @x110 (= (or (not (or |d$| false)) |c$|) $x111)) $x111)))
+(let (($x64 (or |a$| |b$| |c$| |d$|)))
+(let ((@x67 (mp (asserted (or |a$| (or |b$| (or |c$| |d$|)))) (rewrite (= (or |a$| (or |b$| (or |c$| |d$|))) $x64)) $x64)))
+(let ((@x160 (|unit-resolution| @x67 (|unit-resolution| @x106 @x143 $x98) @x143 (|unit-resolution| @x116 @x143 $x108) |a$|)))
+(let (($x81 (not |a$|)))
+(let (($x84 (or $x81 |b$|)))
+(let ((@x76 (monotonicity (rewrite (= (and |c$| $x17) false)) (= (or |a$| (and |c$| $x17)) (or |a$| false)))))
+(let ((@x80 (trans @x76 (rewrite (= (or |a$| false) |a$|)) (= (or |a$| (and |c$| $x17)) |a$|))))
+(let ((@x86 (monotonicity (monotonicity @x80 (= (not (or |a$| (and |c$| $x17))) $x81)) (= (or (not (or |a$| (and |c$| $x17))) |b$|) $x84))))
+(let ((@x89 (mp (asserted (or (not (or |a$| (and |c$| $x17))) |b$|)) @x86 $x84)))
+(|unit-resolution| @x89 @x160 (|unit-resolution| @x106 @x143 $x98) false))))))))))))))))))))))))))))))))
+
+c1a1d5a3f58100ecdaa72705a063eeccc5044c46 27 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let ((?x15 (|symm_f$| |b$| |a$|)))
+(let ((?x14 (|symm_f$| |a$| |b$|)))
+(let (($x16 (= ?x14 ?x15)))
+(let (($x50 (not $x16)))
+(let ((@x45 (monotonicity (rewrite (= (= |a$| |a$|) true)) (= (and (= |a$| |a$|) $x16) (and true $x16)))))
+(let ((@x49 (trans @x45 (rewrite (= (and true $x16) $x16)) (= (and (= |a$| |a$|) $x16) $x16))))
+(let ((@x55 (mp (asserted (not (and (= |a$| |a$|) $x16))) (monotonicity @x49 (= (not (and (= |a$| |a$|) $x16)) $x50)) $x50)))
+(let (($x59 (forall ((?v0 |A$|) (?v1 |A$|) )(!(let ((?x8 (|symm_f$| ?v1 ?v0)))
+(let ((?x7 (|symm_f$| ?v0 ?v1)))
+(= ?x7 ?x8))) :pattern ( (|symm_f$| ?v0 ?v1) ) :pattern ( (|symm_f$| ?v1 ?v0) )))
+))
+(let (($x10 (forall ((?v0 |A$|) (?v1 |A$|) )(let ((?x8 (|symm_f$| ?v1 ?v0)))
+(let ((?x7 (|symm_f$| ?v0 ?v1)))
+(= ?x7 ?x8))))
+))
+(let ((?x8 (|symm_f$| ?0 ?1)))
+(let ((?x7 (|symm_f$| ?1 ?0)))
+(let (($x9 (= ?x7 ?x8)))
+(let ((@x58 (|mp~| (asserted $x10) (|nnf-pos| (refl (|~| $x9 $x9)) (|~| $x10 $x10)) $x10)))
+(let ((@x66 (mp @x58 (|quant-intro| (refl (= $x9 $x9)) (= $x10 $x59)) $x59)))
+(let (($x70 (or (not $x59) $x16)))
+(let ((@x71 ((_ |quant-inst| |a$| |b$|) $x70)))
+(|unit-resolution| @x71 @x66 @x55 false)))))))))))))))))))
+
+d1ba851b4b433507a4e12ae0555630bd23204076 38 0
+unsat
+((set-logic AUFLIA)
+(declare-fun ?v0!0 () Int)
+(declare-fun ?v1!1 () Int)
+(proof
+(let (($x46 (|p$| ?v0!0)))
+(let (($x48 (not $x46)))
+(let (($x61 (not (or $x46 (|p$| ?v1!1)))))
+(let ((@x77 (monotonicity (rewrite (= (not $x48) $x46)) (= (and (not $x48) $x61) (and $x46 $x61)))))
+(let (($x55 (not $x48)))
+(let (($x65 (and $x55 $x61)))
+(let (($x39 (forall ((?v0 Int) )(let (($x10 (forall ((?v1 Int) )(let (($x6 (|p$| ?v1)))
+(or (|p$| ?v0) $x6)))
+))
+(or (not (|p$| ?v0)) $x10)))
+))
+(let (($x42 (not $x39)))
+(let (($x50 (forall ((?v1 Int) )(let (($x6 (|p$| ?v1)))
+(let (($x46 (|p$| ?v0!0)))
+(or $x46 $x6))))
+))
+(let ((@x67 (|nnf-neg| (refl (|~| $x55 $x55)) (sk (|~| (not $x50) $x61)) (|~| (not (or $x48 $x50)) $x65))))
+(let (($x12 (forall ((?v0 Int) )(let (($x10 (forall ((?v1 Int) )(let (($x6 (|p$| ?v1)))
+(or (|p$| ?v0) $x6)))
+))
+(let (($x6 (|p$| ?v0)))
+(=> $x6 $x10))))
+))
+(let (($x13 (not $x12)))
+(let (($x10 (forall ((?v1 Int) )(let (($x6 (|p$| ?v1)))
+(or (|p$| ?0) $x6)))
+))
+(let ((@x41 (|quant-intro| (rewrite (= (=> (|p$| ?0) $x10) (or (not (|p$| ?0)) $x10))) (= $x12 $x39))))
+(let ((@x70 (|mp~| (mp (asserted $x13) (monotonicity @x41 (= $x13 $x42)) $x42) (trans (sk (|~| $x42 (not (or $x48 $x50)))) @x67 (|~| $x42 $x65)) $x65)))
+(let ((@x79 (|not-or-elim| (|and-elim| (mp @x70 @x77 (and $x46 $x61)) $x61) $x48)))
+(let ((@x72 (|and-elim| (mp @x70 @x77 (and $x46 $x61)) $x46)))
+(|unit-resolution| @x72 @x79 false))))))))))))))))))))
+
+19f6b54cdb476573f91d167cec6fca10e0e66fc7 27 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let (($x72 (forall ((?v0 |A$|) )(!(let (($x8 (|p$| ?v0)))
+(not $x8)) :pattern ( (|p$| ?v0) )))
+))
+(let (($x6 (|p$| |x$|)))
+(let ((@x46 (monotonicity (rewrite (= (=> $x6 (|p$| |y$|)) (or (not $x6) (|p$| |y$|)))) (= (not (=> $x6 (|p$| |y$|))) (not (or (not $x6) (|p$| |y$|)))))))
+(let ((@x49 (mp (asserted (not (=> $x6 (|p$| |y$|)))) @x46 (not (or (not $x6) (|p$| |y$|))))))
+(let ((@x47 (|not-or-elim| @x49 $x6)))
+(let (($x40 (not $x6)))
+(let (($x75 (or $x40 $x72)))
+(let (($x12 (forall ((?v0 |A$|) )(let (($x8 (|p$| ?v0)))
+(not $x8)))
+))
+(let (($x62 (or $x40 $x12)))
+(let ((@x74 (|quant-intro| (refl (= (not (|p$| ?0)) (not (|p$| ?0)))) (= $x12 $x72))))
+(let (($x9 (exists ((?v0 |A$|) )(|p$| ?v0))
+))
+(let (($x13 (ite $x6 (not $x9) $x12)))
+(let ((@x58 (|nnf-neg| (refl (|~| (not (|p$| ?0)) (not (|p$| ?0)))) (|~| (not $x9) $x12))))
+(let ((@x65 (|nnf-pos| (refl (|~| $x6 $x6)) (refl (|~| $x40 $x40)) @x58 (|nnf-pos| (refl (|~| (not (|p$| ?0)) (not (|p$| ?0)))) (|~| $x12 $x12)) (|~| $x13 (and $x62 (or $x6 $x12))))))
+(let ((@x78 (mp (|and-elim| (|mp~| (asserted $x13) @x65 (and $x62 (or $x6 $x12))) $x62) (monotonicity @x74 (= $x62 $x75)) $x75)))
+(let (($x86 (or (not $x72) $x40)))
+(let ((@x87 ((_ |quant-inst| |x$|) $x86)))
+(|unit-resolution| @x87 @x47 (|unit-resolution| @x78 @x47 $x72) false))))))))))))))))))))
+
+e86ca8427589ec8e24e5a85d218331bfb59ff385 7 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let ((@x33 (monotonicity (rewrite (= (= 3 3) true)) (= (not (= 3 3)) (not true)))))
+(let ((@x37 (trans @x33 (rewrite (= (not true) false)) (= (not (= 3 3)) false))))
+(mp (asserted (not (= 3 3))) @x37 false)))))
+
+77108fa1aa6a8a356ebdd1a376316f26d90399cb 7 0
+unsat
+((set-logic AUFLIRA)
+(proof
+(let ((@x33 (monotonicity (rewrite (= (= 3.0 3.0) true)) (= (not (= 3.0 3.0)) (not true)))))
+(let ((@x37 (trans @x33 (rewrite (= (not true) false)) (= (not (= 3.0 3.0)) false))))
+(mp (asserted (not (= 3.0 3.0))) @x37 false)))))
+
+98abe835b7d13273c58720c5dadf713cd8637495 9 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let ((@x35 (monotonicity (rewrite (= (+ 3 1) 4)) (= (= (+ 3 1) 4) (= 4 4)))))
+(let ((@x39 (trans @x35 (rewrite (= (= 4 4) true)) (= (= (+ 3 1) 4) true))))
+(let ((@x42 (monotonicity @x39 (= (not (= (+ 3 1) 4)) (not true)))))
+(let ((@x46 (trans @x42 (rewrite (= (not true) false)) (= (not (= (+ 3 1) 4)) false))))
+(mp (asserted (not (= (+ 3 1) 4))) @x46 false)))))))
+
+0382c7d04a37d9ca60cac3282bc80f6b329ab12f 16 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let ((?x10 (+ |z$| |x$|)))
+(let ((?x11 (+ |y$| ?x10)))
+(let ((?x8 (+ |y$| |z$|)))
+(let ((?x9 (+ |x$| ?x8)))
+(let (($x12 (= ?x9 ?x11)))
+(let (($x13 (not $x12)))
+(let ((@x43 (monotonicity (rewrite (= ?x10 (+ |x$| |z$|))) (= ?x11 (+ |y$| (+ |x$| |z$|))))))
+(let ((@x47 (trans @x43 (rewrite (= (+ |y$| (+ |x$| |z$|)) (+ |x$| |y$| |z$|))) (= ?x11 (+ |x$| |y$| |z$|)))))
+(let ((@x50 (monotonicity (rewrite (= ?x9 (+ |x$| |y$| |z$|))) @x47 (= $x12 (= (+ |x$| |y$| |z$|) (+ |x$| |y$| |z$|))))))
+(let ((@x54 (trans @x50 (rewrite (= (= (+ |x$| |y$| |z$|) (+ |x$| |y$| |z$|)) true)) (= $x12 true))))
+(let ((@x61 (trans (monotonicity @x54 (= $x13 (not true))) (rewrite (= (not true) false)) (= $x13 false))))
+(mp (asserted $x13) @x61 false))))))))))))))
+
+c608fc7154ce1246a30c68f4d20c1d35cedba663 11 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let ((@x39 (monotonicity (rewrite (= (<= 3 8) true)) (= (ite (<= 3 8) 8 3) (ite true 8 3)))))
+(let ((@x43 (trans @x39 (rewrite (= (ite true 8 3) 8)) (= (ite (<= 3 8) 8 3) 8))))
+(let ((@x46 (monotonicity @x43 (= (< 5 (ite (<= 3 8) 8 3)) (< 5 8)))))
+(let ((@x50 (trans @x46 (rewrite (= (< 5 8) true)) (= (< 5 (ite (<= 3 8) 8 3)) true))))
+(let ((@x53 (monotonicity @x50 (= (not (< 5 (ite (<= 3 8) 8 3))) (not true)))))
+(let ((@x57 (trans @x53 (rewrite (= (not true) false)) (= (not (< 5 (ite (<= 3 8) 8 3))) false))))
+(mp (asserted (not (< 5 (ite (<= 3 8) 8 3)))) @x57 false)))))))))
+
+4bdd1f2f245666e5db75e9d320ea9e892060d851 88 0
+unsat
+((set-logic AUFLIRA)
+(proof
+(let ((?x42 (* (~ 1.0) |x$|)))
+(let (($x81 (>= |x$| 0.0)))
+(let ((?x88 (ite $x81 |x$| ?x42)))
+(let ((?x111 (* (~ 1.0) ?x88)))
+(let ((?x146 (+ |x$| ?x111)))
+(let (($x147 (<= ?x146 0.0)))
+(let (($x131 (= |x$| ?x88)))
+(let ((?x43 (* (~ 1.0) |y$|)))
+(let ((?x44 (+ ?x42 ?x43)))
+(let ((?x7 (+ |x$| |y$|)))
+(let (($x69 (>= ?x7 0.0)))
+(let ((?x76 (ite $x69 ?x7 ?x44)))
+(let ((?x149 (* (~ 1.0) ?x76)))
+(let ((?x177 (+ ?x44 ?x149)))
+(let (($x179 (>= ?x177 0.0)))
+(let (($x128 (= ?x44 ?x76)))
+(let (($x70 (not $x69)))
+(let (($x93 (>= |y$| 0.0)))
+(let (($x94 (not $x93)))
+(let (($x152 (>= (+ ?x7 ?x149) 0.0)))
+(let (($x127 (= ?x7 ?x76)))
+(let (($x188 (not $x179)))
+(let ((@x159 (hypothesis $x93)))
+(let ((?x100 (ite $x93 |y$| ?x43)))
+(let ((?x112 (* (~ 1.0) ?x100)))
+(let ((?x113 (+ ?x76 ?x111 ?x112)))
+(let (($x114 (<= ?x113 0.0)))
+(let (($x119 (not $x114)))
+(let ((?x18 (+ (ite (< |x$| 0.0) (- |x$|) |x$|) (ite (< |y$| 0.0) (- |y$|) |y$|))))
+(let (($x20 (not (<= (ite (< ?x7 0.0) (- ?x7) ?x7) ?x18))))
+(let (($x15 (< |y$| 0.0)))
+(let ((?x57 (ite $x15 ?x43 |y$|)))
+(let (($x12 (< |x$| 0.0)))
+(let ((?x52 (ite $x12 ?x42 |x$|)))
+(let ((?x60 (+ ?x52 ?x57)))
+(let (($x9 (< ?x7 0.0)))
+(let ((?x47 (ite $x9 ?x44 ?x7)))
+(let (($x63 (<= ?x47 ?x60)))
+(let ((@x104 (trans (monotonicity (rewrite (= $x15 $x94)) (= ?x57 (ite $x94 ?x43 |y$|))) (rewrite (= (ite $x94 ?x43 |y$|) ?x100)) (= ?x57 ?x100))))
+(let ((@x87 (monotonicity (rewrite (= $x12 (not $x81))) (= ?x52 (ite (not $x81) ?x42 |x$|)))))
+(let ((@x92 (trans @x87 (rewrite (= (ite (not $x81) ?x42 |x$|) ?x88)) (= ?x52 ?x88))))
+(let ((@x80 (trans (monotonicity (rewrite (= $x9 $x70)) (= ?x47 (ite $x70 ?x44 ?x7))) (rewrite (= (ite $x70 ?x44 ?x7) ?x76)) (= ?x47 ?x76))))
+(let ((@x110 (monotonicity @x80 (monotonicity @x92 @x104 (= ?x60 (+ ?x88 ?x100))) (= $x63 (<= ?x76 (+ ?x88 ?x100))))))
+(let ((@x118 (trans @x110 (rewrite (= (<= ?x76 (+ ?x88 ?x100)) $x114)) (= $x63 $x114))))
+(let ((@x59 (monotonicity (rewrite (= (- |y$|) ?x43)) (= (ite $x15 (- |y$|) |y$|) ?x57))))
+(let ((@x54 (monotonicity (rewrite (= (- |x$|) ?x42)) (= (ite $x12 (- |x$|) |x$|) ?x52))))
+(let ((@x49 (monotonicity (rewrite (= (- ?x7) ?x44)) (= (ite $x9 (- ?x7) ?x7) ?x47))))
+(let ((@x65 (monotonicity @x49 (monotonicity @x54 @x59 (= ?x18 ?x60)) (= (<= (ite $x9 (- ?x7) ?x7) ?x18) $x63))))
+(let ((@x123 (trans (monotonicity @x65 (= $x20 (not $x63))) (monotonicity @x118 (= (not $x63) $x119)) (= $x20 $x119))))
+(let ((@x124 (mp (asserted $x20) @x123 $x119)))
+(let (($x137 (= |y$| ?x100)))
+(let ((@x172 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x137) (<= (+ |y$| ?x112) 0.0))) (|unit-resolution| (|def-axiom| (or $x94 $x137)) @x159 $x137) (<= (+ |y$| ?x112) 0.0))))
+(let ((?x148 (+ ?x42 ?x111)))
+(let (($x151 (<= ?x148 0.0)))
+(let (($x132 (= ?x42 ?x88)))
+(let (($x82 (not $x81)))
+(let ((@x157 ((_ |th-lemma| arith triangle-eq) (or (not $x131) $x147))))
+(let ((@x158 (|unit-resolution| @x157 (|unit-resolution| (|def-axiom| (or $x82 $x131)) (hypothesis $x81) $x131) $x147)))
+(let ((@x162 (|unit-resolution| ((_ |th-lemma| arith assign-bounds -1 -1) (or $x69 $x82 $x94)) (hypothesis $x81) @x159 $x69)))
+(let ((@x126 (|def-axiom| (or $x70 $x127))))
+(let ((@x166 ((_ |th-lemma| arith triangle-eq) (or (not $x127) $x152))))
+(let ((@x173 ((_ |th-lemma| arith farkas 1 -1 -1 1) @x172 (|unit-resolution| @x166 (|unit-resolution| @x126 @x162 $x127) $x152) @x124 @x158 false)))
+(let ((@x136 (|def-axiom| (or $x81 $x132))))
+(let ((@x182 (|unit-resolution| @x136 (|unit-resolution| (lemma @x173 (or $x82 $x94)) @x159 $x82) $x132)))
+(let ((@x187 ((_ |th-lemma| arith farkas 2 -1 -1 1 1) @x159 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x132) $x151)) @x182 $x151) @x172 @x124 (hypothesis $x179) false)))
+(let ((@x196 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x128) $x179)) (hypothesis $x128) (hypothesis $x188) false)))
+(let ((@x197 (lemma @x196 (or (not $x128) $x179))))
+(let ((@x199 (|unit-resolution| @x197 (|unit-resolution| (lemma @x187 (or $x188 $x94)) @x159 $x188) (not $x128))))
+(let ((@x130 (|def-axiom| (or $x69 $x128))))
+(let ((@x202 (|unit-resolution| @x166 (|unit-resolution| @x126 (|unit-resolution| @x130 @x199 $x69) $x127) $x152)))
+(let ((@x203 ((_ |th-lemma| arith farkas 2 1 1 1 1) (|unit-resolution| (lemma @x173 (or $x82 $x94)) @x159 $x82) (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x132) $x151)) @x182 $x151) @x172 @x124 @x202 false)))
+(let ((@x204 (lemma @x203 $x94)))
+(let ((@x210 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 1) (or $x81 $x93 $x70)) (hypothesis $x69) @x204 $x81)))
+(let ((@x134 (|def-axiom| (or $x82 $x131))))
+(let ((@x214 (|unit-resolution| @x166 (|unit-resolution| @x126 (hypothesis $x69) $x127) $x152)))
+(let ((?x145 (+ ?x43 ?x112)))
+(let (($x176 (<= ?x145 0.0)))
+(let (($x138 (= ?x43 ?x100)))
+(let ((@x219 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x138) $x176)) (|unit-resolution| (|def-axiom| (or $x93 $x138)) @x204 $x138) $x176)))
+(let ((@x220 ((_ |th-lemma| arith farkas 2 1 1 1 1) @x204 @x219 @x124 @x214 (|unit-resolution| @x157 (|unit-resolution| @x134 @x210 $x131) $x147) false)))
+(let ((@x224 (|unit-resolution| @x197 (|unit-resolution| @x130 (lemma @x220 $x70) $x128) $x179)))
+(let ((@x229 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x132) $x151)) (hypothesis $x132) (lemma ((_ |th-lemma| arith farkas 1 -1 -1 1) @x219 @x124 @x224 (hypothesis $x151) false) (not $x151)) false)))
+(let ((@x232 (|unit-resolution| @x134 (|unit-resolution| @x136 (lemma @x229 (not $x132)) $x81) $x131)))
+((_ |th-lemma| arith farkas -2 1 -1 -1 1) (|unit-resolution| @x136 (lemma @x229 (not $x132)) $x81) @x219 @x124 @x224 (|unit-resolution| @x157 @x232 $x147) false))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))
+
+f8d266138153a7b5a745c746bbb489254a734ae0 16 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let ((?x10 (|p$| true)))
+(let (($x7 (< 2 3)))
+(let (($x8 (ite $x7 true false)))
+(let ((?x9 (|p$| $x8)))
+(let (($x11 (= ?x9 ?x10)))
+(let (($x12 (not $x11)))
+(let ((@x50 (monotonicity (monotonicity (rewrite (= $x7 true)) (= (|p$| $x7) ?x10)) (= (= (|p$| $x7) ?x10) (= ?x10 ?x10)))))
+(let ((@x54 (trans @x50 (rewrite (= (= ?x10 ?x10) true)) (= (= (|p$| $x7) ?x10) true))))
+(let ((@x61 (trans (monotonicity @x54 (= (not (= (|p$| $x7) ?x10)) (not true))) (rewrite (= (not true) false)) (= (not (= (|p$| $x7) ?x10)) false))))
+(let ((@x41 (monotonicity (monotonicity (rewrite (= $x8 $x7)) (= ?x9 (|p$| $x7))) (= $x11 (= (|p$| $x7) ?x10)))))
+(let ((@x44 (monotonicity @x41 (= $x12 (not (= (|p$| $x7) ?x10))))))
+(mp (asserted $x12) (trans @x44 @x61 (= $x12 false)) false))))))))))))))
+
+81a816463ea508b010daafde9e601b0b985afe71 16 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let (($x11 (< |x$| 1)))
+(let ((?x35 (+ 3 |x$|)))
+(let (($x38 (<= 4 ?x35)))
+(let (($x41 (or $x38 $x11)))
+(let (($x44 (not $x41)))
+(let ((@x55 (monotonicity (rewrite (= $x38 (>= |x$| 1))) (rewrite (= $x11 (not (>= |x$| 1)))) (= $x41 (or (>= |x$| 1) (not (>= |x$| 1)))))))
+(let ((@x59 (trans @x55 (rewrite (= (or (>= |x$| 1) (not (>= |x$| 1))) true)) (= $x41 true))))
+(let ((@x66 (trans (monotonicity @x59 (= $x44 (not true))) (rewrite (= (not true) false)) (= $x44 false))))
+(let ((@x40 (monotonicity (rewrite (= (+ |x$| 3) ?x35)) (= (<= 4 (+ |x$| 3)) $x38))))
+(let ((@x46 (monotonicity (monotonicity @x40 (= (or (<= 4 (+ |x$| 3)) $x11) $x41)) (= (not (or (<= 4 (+ |x$| 3)) $x11)) $x44))))
+(let ((@x68 (trans @x46 @x66 (= (not (or (<= 4 (+ |x$| 3)) $x11)) false))))
+(mp (asserted (not (or (<= 4 (+ |x$| 3)) $x11))) @x68 false))))))))))))))
+
+3da41aa632fdaf484d160ab8b5a2c83b931d3de7 18 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let (($x52 (= (+ |x$| (* (~ 1) |y$|)) (~ 4))))
+(let ((@x46 (monotonicity (rewrite (= (+ |x$| 4) (+ 4 |x$|))) (= (= |y$| (+ |x$| 4)) (= |y$| (+ 4 |x$|))))))
+(let ((@x55 (trans @x46 (rewrite (= (= |y$| (+ 4 |x$|)) $x52)) (= (= |y$| (+ |x$| 4)) $x52))))
+(let ((@x84 (monotonicity (mp (asserted (= |y$| (+ |x$| 4))) @x55 $x52) (= (>= (+ |x$| (* (~ 1) |y$|)) 0) (>= (~ 4) 0)))))
+(let ((@x88 (trans @x84 (rewrite (= (>= (~ 4) 0) false)) (= (>= (+ |x$| (* (~ 1) |y$|)) 0) false))))
+(let (($x68 (>= (+ |x$| (* (~ 1) |y$|)) 0)))
+(let ((@x74 (monotonicity (rewrite (= (< 0 (+ (* (~ 1) |x$|) |y$|)) (not $x68))) (= (not (< 0 (+ (* (~ 1) |x$|) |y$|))) (not (not $x68))))))
+(let ((@x78 (trans @x74 (rewrite (= (not (not $x68)) $x68)) (= (not (< 0 (+ (* (~ 1) |x$|) |y$|))) $x68))))
+(let (($x62 (< 0 (+ (* (~ 1) |x$|) |y$|))))
+(let (($x65 (not $x62)))
+(let (($x15 (not (< 0 (- |y$| |x$|)))))
+(let ((@x64 (monotonicity (rewrite (= (- |y$| |x$|) (+ (* (~ 1) |x$|) |y$|))) (= (< 0 (- |y$| |x$|)) $x62))))
+(let ((@x81 (mp (asserted $x15) (trans (monotonicity @x64 (= $x15 $x65)) @x78 (= $x15 $x68)) $x68)))
+(mp @x81 @x88 false))))))))))))))))
+
+e43b05132d28d45640c9d0131930806093dbb0e2 11 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let ((@x37 (monotonicity (rewrite (= (+ 2 2) 4)) (= (= (+ 2 2) 5) (= 4 5)))))
+(let ((@x41 (trans @x37 (rewrite (= (= 4 5) false)) (= (= (+ 2 2) 5) false))))
+(let ((@x44 (monotonicity @x41 (= (not (= (+ 2 2) 5)) (not false)))))
+(let ((@x48 (trans @x44 (rewrite (= (not false) true)) (= (not (= (+ 2 2) 5)) true))))
+(let ((@x51 (monotonicity @x48 (= (not (not (= (+ 2 2) 5))) (not true)))))
+(let ((@x55 (trans @x51 (rewrite (= (not true) false)) (= (not (not (= (+ 2 2) 5))) false))))
+(mp (asserted (not (not (= (+ 2 2) 5)))) @x55 false)))))))))
+
+135df42816691c806246099ddf6fc7f6b81a2f42 19 0
+unsat
+((set-logic AUFLIRA)
+(proof
+(let ((?x10 (* 7.0 |a$|)))
+(let ((?x7 (* 3.0 |x$|)))
+(let ((?x11 (+ ?x7 ?x10)))
+(let (($x46 (>= ?x11 4.0)))
+(let (($x44 (not $x46)))
+(let ((@x43 (mp (asserted (< ?x11 4.0)) (rewrite (= (< ?x11 4.0) $x44)) $x44)))
+(let ((?x15 (* 2.0 |x$|)))
+(let (($x48 (<= ?x15 3.0)))
+(let (($x49 (not $x48)))
+(let ((@x52 (mp (asserted (< 3.0 ?x15)) (rewrite (= (< 3.0 ?x15) $x49)) $x49)))
+(let (($x56 (>= |a$| 0.0)))
+(let ((@x60 (monotonicity (rewrite (= (< |a$| 0.0) (not $x56))) (= (not (< |a$| 0.0)) (not (not $x56))))))
+(let ((@x64 (trans @x60 (rewrite (= (not (not $x56)) $x56)) (= (not (< |a$| 0.0)) $x56))))
+(let ((@x65 (mp (asserted (not (< |a$| 0.0))) @x64 $x56)))
+((_ |th-lemma| arith farkas 7 3/2 1) @x65 @x52 @x43 false)))))))))))))))))
+
+de926642fcc1657dfaa079f1656df9cc74f3caaf 22 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let (($x17 (not false)))
+(let (($x13 (<= 0 |x$|)))
+(let (($x14 (not $x13)))
+(let (($x15 (or $x14 $x13)))
+(let ((?x8 (- 1)))
+(let ((?x10 (* ?x8 |x$|)))
+(let ((?x11 (+ |y$| ?x10)))
+(let (($x12 (<= 0 ?x11)))
+(let (($x16 (or $x12 $x15)))
+(let (($x18 (= $x16 $x17)))
+(let (($x19 (not $x18)))
+(let ((@x58 (rewrite (= (or (<= 0 (+ |y$| (* (~ 1) |x$|))) true) true))))
+(let ((@x48 (monotonicity (monotonicity (rewrite (= ?x8 (~ 1))) (= ?x10 (* (~ 1) |x$|))) (= ?x11 (+ |y$| (* (~ 1) |x$|))))))
+(let ((@x56 (monotonicity (monotonicity @x48 (= $x12 (<= 0 (+ |y$| (* (~ 1) |x$|))))) (rewrite (= $x15 true)) (= $x16 (or (<= 0 (+ |y$| (* (~ 1) |x$|))) true)))))
+(let ((@x65 (monotonicity (trans @x56 @x58 (= $x16 true)) (rewrite (= $x17 true)) (= $x18 (= true true)))))
+(let ((@x69 (trans @x65 (rewrite (= (= true true) true)) (= $x18 true))))
+(let ((@x76 (trans (monotonicity @x69 (= $x19 (not true))) (rewrite (= (not true) false)) (= $x19 false))))
+(mp (asserted $x19) @x76 false))))))))))))))))))))
+
+c19a59241b121ac2665c4fbd7ba1fa2d48fef984 159 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let (($x22 (= |m$| |n$|)))
+(let ((@x478 (symm (commutativity (= $x22 (= |n$| |m$|))) (= (= |n$| |m$|) $x22))))
+(let (($x18 (= |n$| |m$|)))
+(let ((?x100 (* (~ 1) |m$|)))
+(let ((?x101 (+ |n$| ?x100)))
+(let (($x116 (>= ?x101 0)))
+(let ((?x76 (* (~ 1) |n$a|)))
+(let ((?x94 (+ |m$| ?x76)))
+(let (($x125 (<= ?x94 0)))
+(let ((?x77 (+ |n$| ?x76)))
+(let (($x86 (>= ?x77 0)))
+(let (($x259 (or $x86 $x125)))
+(let ((@x265 (monotonicity (rewrite (= (and (not $x86) (not $x125)) (not $x259))) (= (not (and (not $x86) (not $x125))) (not (not $x259))))))
+(let ((@x269 (trans @x265 (rewrite (= (not (not $x259)) $x259)) (= (not (and (not $x86) (not $x125))) $x259))))
+(let (($x126 (not $x125)))
+(let (($x85 (not $x86)))
+(let (($x141 (and $x85 $x126)))
+(let (($x208 (not $x141)))
+(let (($x28 (= |n$a| |m$|)))
+(let (($x35 (and $x28 $x22)))
+(let (($x78 (<= ?x77 0)))
+(let (($x79 (not $x78)))
+(let (($x11 (= |m$| |n$a|)))
+(let (($x82 (and $x11 $x79)))
+(let (($x89 (and $x22 $x85)))
+(let (($x93 (>= ?x94 0)))
+(let (($x92 (not $x93)))
+(let (($x97 (and $x92 $x79)))
+(let (($x26 (= |n$a| |n$|)))
+(let (($x102 (<= ?x101 0)))
+(let (($x103 (not $x102)))
+(let (($x106 (and $x103 $x26)))
+(let (($x109 (and $x103 $x85)))
+(let (($x112 (and $x28 $x103)))
+(let (($x115 (not $x116)))
+(let (($x119 (and $x26 $x115)))
+(let (($x122 (and $x79 $x115)))
+(let (($x129 (and $x126 $x22)))
+(let (($x132 (and $x126 $x103)))
+(let (($x135 (and $x18 $x92)))
+(let (($x16 (= |n$| |n$a|)))
+(let (($x138 (and $x16 $x126)))
+(let (($x144 (and $x115 $x11)))
+(let (($x147 (and $x115 $x92)))
+(let (($x195 (or $x147 $x144 $x141 $x138 $x135 $x132 $x129 $x122 $x119 $x112 $x109 $x106 $x97 $x89 $x82 $x35)))
+(let (($x38 (or (and (< |m$| |n$a|) (< |n$a| |n$|)) (or (and $x22 (< |n$| |n$a|)) (or (and $x11 (< |n$a| |n$|)) $x35)))))
+(let (($x40 (or (and (< |m$| |n$|) (< |n$| |n$a|)) (or (and (< |m$| |n$|) $x26) $x38))))
+(let (($x43 (or (and (< |n$a| |n$|) (< |n$| |m$|)) (or (and $x26 (< |n$| |m$|)) (or (and $x28 (< |m$| |n$|)) $x40)))))
+(let (($x45 (or (and (< |n$a| |m$|) (< |m$| |n$|)) (or (and (< |n$a| |m$|) $x22) $x43))))
+(let (($x48 (or (and (< |n$| |n$a|) (< |n$a| |m$|)) (or (and $x16 (< |n$a| |m$|)) (or (and $x18 (< |m$| |n$a|)) $x45)))))
+(let (($x50 (or (and (< |n$| |m$|) (< |m$| |n$a|)) (or (and (< |n$| |m$|) $x11) $x48))))
+(let (($x51 (not $x50)))
+(let (($x168 (or $x119 (or $x112 (or $x109 (or $x106 (or $x97 (or $x89 (or $x82 $x35)))))))))
+(let (($x189 (or $x144 (or $x141 (or $x138 (or $x135 (or $x132 (or $x129 (or $x122 $x168)))))))))
+(let (($x187 (= $x48 (or $x141 (or $x138 (or $x135 (or $x132 (or $x129 (or $x122 $x168)))))))))
+(let (($x184 (= (or (and $x16 (< |n$a| |m$|)) (or (and $x18 (< |m$| |n$a|)) $x45)) (or $x138 (or $x135 (or $x132 (or $x129 (or $x122 $x168))))))))
+(let (($x181 (= (or (and $x18 (< |m$| |n$a|)) $x45) (or $x135 (or $x132 (or $x129 (or $x122 $x168)))))))
+(let (($x169 (= (or (and $x26 (< |n$| |m$|)) (or (and $x28 (< |m$| |n$|)) $x40)) $x168)))
+(let (($x166 (= (or (and $x28 (< |m$| |n$|)) $x40) (or $x112 (or $x109 (or $x106 (or $x97 (or $x89 (or $x82 $x35)))))))))
+(let (($x160 (= (or (and (< |m$| |n$|) $x26) $x38) (or $x106 (or $x97 (or $x89 (or $x82 $x35)))))))
+(let (($x154 (= (or (and $x22 (< |n$| |n$a|)) (or (and $x11 (< |n$a| |n$|)) $x35)) (or $x89 (or $x82 $x35)))))
+(let ((@x81 (rewrite (= (< |n$a| |n$|) $x79))))
+(let ((@x152 (monotonicity (monotonicity @x81 (= (and $x11 (< |n$a| |n$|)) $x82)) (= (or (and $x11 (< |n$a| |n$|)) $x35) (or $x82 $x35)))))
+(let ((@x88 (rewrite (= (< |n$| |n$a|) $x85))))
+(let ((@x155 (monotonicity (monotonicity @x88 (= (and $x22 (< |n$| |n$a|)) $x89)) @x152 $x154)))
+(let ((@x96 (rewrite (= (< |m$| |n$a|) $x92))))
+(let ((@x99 (monotonicity @x96 @x81 (= (and (< |m$| |n$a|) (< |n$a| |n$|)) $x97))))
+(let ((@x158 (monotonicity @x99 @x155 (= $x38 (or $x97 (or $x89 (or $x82 $x35)))))))
+(let ((@x105 (rewrite (= (< |m$| |n$|) $x103))))
+(let ((@x161 (monotonicity (monotonicity @x105 (= (and (< |m$| |n$|) $x26) $x106)) @x158 $x160)))
+(let ((@x111 (monotonicity @x105 @x88 (= (and (< |m$| |n$|) (< |n$| |n$a|)) $x109))))
+(let ((@x164 (monotonicity @x111 @x161 (= $x40 (or $x109 (or $x106 (or $x97 (or $x89 (or $x82 $x35)))))))))
+(let ((@x167 (monotonicity (monotonicity @x105 (= (and $x28 (< |m$| |n$|)) $x112)) @x164 $x166)))
+(let ((@x118 (rewrite (= (< |n$| |m$|) $x115))))
+(let ((@x170 (monotonicity (monotonicity @x118 (= (and $x26 (< |n$| |m$|)) $x119)) @x167 $x169)))
+(let ((@x124 (monotonicity @x81 @x118 (= (and (< |n$a| |n$|) (< |n$| |m$|)) $x122))))
+(let ((@x128 (rewrite (= (< |n$a| |m$|) $x126))))
+(let ((@x176 (monotonicity (monotonicity @x128 (= (and (< |n$a| |m$|) $x22) $x129)) (monotonicity @x124 @x170 (= $x43 (or $x122 $x168))) (= (or (and (< |n$a| |m$|) $x22) $x43) (or $x129 (or $x122 $x168))))))
+(let ((@x134 (monotonicity @x128 @x105 (= (and (< |n$a| |m$|) (< |m$| |n$|)) $x132))))
+(let ((@x179 (monotonicity @x134 @x176 (= $x45 (or $x132 (or $x129 (or $x122 $x168)))))))
+(let ((@x182 (monotonicity (monotonicity @x96 (= (and $x18 (< |m$| |n$a|)) $x135)) @x179 $x181)))
+(let ((@x185 (monotonicity (monotonicity @x128 (= (and $x16 (< |n$a| |m$|)) $x138)) @x182 $x184)))
+(let ((@x143 (monotonicity @x88 @x128 (= (and (< |n$| |n$a|) (< |n$a| |m$|)) $x141))))
+(let ((@x191 (monotonicity (monotonicity @x118 (= (and (< |n$| |m$|) $x11) $x144)) (monotonicity @x143 @x185 $x187) (= (or (and (< |n$| |m$|) $x11) $x48) $x189))))
+(let ((@x149 (monotonicity @x118 @x96 (= (and (< |n$| |m$|) (< |m$| |n$a|)) $x147))))
+(let ((@x199 (trans (monotonicity @x149 @x191 (= $x50 (or $x147 $x189))) (rewrite (= (or $x147 $x189) $x195)) (= $x50 $x195))))
+(let ((@x203 (mp (asserted $x51) (monotonicity @x199 (= $x51 (not $x195))) (not $x195))))
+(let ((@x270 (mp (|not-or-elim| @x203 $x208) @x269 $x259)))
+(let (($x271 (not $x16)))
+(let (($x272 (or $x271 $x125)))
+(let ((@x278 (monotonicity (rewrite (= $x138 (not $x272))) (= (not $x138) (not (not $x272))))))
+(let ((@x282 (trans @x278 (rewrite (= (not (not $x272)) $x272)) (= (not $x138) $x272))))
+(let ((@x283 (mp (|not-or-elim| @x203 (not $x138)) @x282 $x272)))
+(let (($x284 (not $x18)))
+(let (($x309 (not $x22)))
+(let ((@x432 (hypothesis $x79)))
+(let (($x384 (or $x93 $x78)))
+(let ((@x390 (monotonicity (rewrite (= $x97 (not $x384))) (= (not $x97) (not (not $x384))))))
+(let ((@x394 (trans @x390 (rewrite (= (not (not $x384)) $x384)) (= (not $x97) $x384))))
+(let ((@x395 (mp (|not-or-elim| @x203 (not $x97)) @x394 $x384)))
+(let (($x246 (not $x11)))
+(let (($x408 (or $x246 $x78)))
+(let ((@x414 (monotonicity (rewrite (= $x82 (not $x408))) (= (not $x82) (not (not $x408))))))
+(let ((@x418 (trans @x414 (rewrite (= (not (not $x408)) $x408)) (= (not $x82) $x408))))
+(let ((@x419 (mp (|not-or-elim| @x203 (not $x82)) @x418 $x408)))
+(let ((@x437 ((_ |th-lemma| arith triangle-eq) (or $x11 $x126 $x92))))
+(let ((@x438 (|unit-resolution| @x437 (|unit-resolution| @x419 @x432 $x246) (|unit-resolution| @x395 @x432 $x93) $x126)))
+(let (($x310 (or $x125 $x309)))
+(let ((@x316 (monotonicity (rewrite (= $x129 (not $x310))) (= (not $x129) (not (not $x310))))))
+(let ((@x320 (trans @x316 (rewrite (= (not (not $x310)) $x310)) (= (not $x129) $x310))))
+(let ((@x321 (mp (|not-or-elim| @x203 (not $x129)) @x320 $x310)))
+(let ((@x448 (mp (|unit-resolution| @x321 @x438 $x309) (monotonicity (commutativity (= $x22 $x18)) (= $x309 $x284)) $x284)))
+(let (($x322 (or $x78 $x116)))
+(let ((@x328 (monotonicity (rewrite (= $x122 (not $x322))) (= (not $x122) (not (not $x322))))))
+(let ((@x332 (trans @x328 (rewrite (= (not (not $x322)) $x322)) (= (not $x122) $x322))))
+(let ((@x333 (mp (|not-or-elim| @x203 (not $x122)) @x332 $x322)))
+(let (($x297 (or $x125 $x102)))
+(let ((@x303 (monotonicity (rewrite (= $x132 (not $x297))) (= (not $x132) (not (not $x297))))))
+(let ((@x307 (trans @x303 (rewrite (= (not (not $x297)) $x297)) (= (not $x132) $x297))))
+(let ((@x308 (mp (|not-or-elim| @x203 (not $x132)) @x307 $x297)))
+(let ((@x442 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or $x18 $x103 $x115)) (|unit-resolution| @x308 @x438 $x102) (|unit-resolution| @x333 @x432 $x116) $x18)))
+(let ((@x457 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or $x16 $x79 $x85)) (lemma (|unit-resolution| @x442 @x448 false) $x78) (or $x16 $x85))))
+(let ((@x458 (|unit-resolution| @x457 (|unit-resolution| @x283 (hypothesis $x126) $x271) (|unit-resolution| @x270 (hypothesis $x126) $x86) false)))
+(let ((@x459 (lemma @x458 $x125)))
+(let (($x73 (or $x116 $x93)))
+(let ((@x240 (monotonicity (rewrite (= $x147 (not $x73))) (= (not $x147) (not (not $x73))))))
+(let ((@x244 (trans @x240 (rewrite (= (not (not $x73)) $x73)) (= (not $x147) $x73))))
+(let ((@x245 (mp (|not-or-elim| @x203 (not $x147)) @x244 $x73)))
+(let (($x247 (or $x116 $x246)))
+(let ((@x253 (monotonicity (rewrite (= $x144 (not $x247))) (= (not $x144) (not (not $x247))))))
+(let ((@x257 (trans @x253 (rewrite (= (not (not $x247)) $x247)) (= (not $x144) $x247))))
+(let ((@x258 (mp (|not-or-elim| @x203 (not $x144)) @x257 $x247)))
+(let ((@x463 (|unit-resolution| @x437 (|unit-resolution| @x258 (hypothesis $x115) $x246) (|unit-resolution| @x245 (hypothesis $x115) $x93) @x459 false)))
+(let (($x334 (not $x26)))
+(let (($x372 (or $x102 $x334)))
+(let ((@x378 (monotonicity (rewrite (= $x106 (not $x372))) (= (not $x106) (not (not $x372))))))
+(let ((@x382 (trans @x378 (rewrite (= (not (not $x372)) $x372)) (= (not $x106) $x372))))
+(let ((@x383 (mp (|not-or-elim| @x203 (not $x106)) @x382 $x372)))
+(let ((@x473 (mp (|unit-resolution| @x383 (hypothesis $x103) $x334) (monotonicity (commutativity (= $x26 $x16)) (= $x334 $x271)) $x271)))
+(let (($x360 (or $x102 $x86)))
+(let ((@x366 (monotonicity (rewrite (= $x109 (not $x360))) (= (not $x109) (not (not $x360))))))
+(let ((@x370 (trans @x366 (rewrite (= (not (not $x360)) $x360)) (= (not $x109) $x360))))
+(let ((@x371 (mp (|not-or-elim| @x203 (not $x109)) @x370 $x360)))
+(let ((@x467 (|unit-resolution| @x457 (|unit-resolution| @x371 (hypothesis $x103) $x86) $x16)))
+(let ((@x476 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or $x18 $x103 $x115)) (lemma (|unit-resolution| @x467 @x473 false) $x102) (lemma @x463 $x116) $x18)))
+(let (($x285 (or $x284 $x93)))
+(let ((@x291 (monotonicity (rewrite (= $x135 (not $x285))) (= (not $x135) (not (not $x285))))))
+(let ((@x295 (trans @x291 (rewrite (= (not (not $x285)) $x285)) (= (not $x135) $x285))))
+(let ((@x296 (mp (|not-or-elim| @x203 (not $x135)) @x295 $x285)))
+(let ((@x486 (mp (|unit-resolution| @x437 (|unit-resolution| @x296 @x476 $x93) @x459 $x11) (symm (commutativity (= $x28 $x11)) (= $x11 $x28)) $x28)))
+(let (($x420 (or (not $x28) $x309)))
+(let ((@x426 (monotonicity (rewrite (= $x35 (not $x420))) (= (not $x35) (not (not $x420))))))
+(let ((@x430 (trans @x426 (rewrite (= (not (not $x420)) $x420)) (= (not $x35) $x420))))
+(let ((@x431 (mp (|not-or-elim| @x203 (not $x35)) @x430 $x420)))
+(|unit-resolution| @x431 @x486 (mp @x476 @x478 $x22) false)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))
+
+ad69b5703e25623b7fccdbcfa3db5949b2899f42 927 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let ((?x131 (* (~ 1) |x4$|)))
+(let (($x436 (>= |x4$| 0)))
+(let ((?x443 (ite $x436 |x4$| ?x131)))
+(let ((?x454 (* (~ 1) ?x443)))
+(let ((?x675 (+ |x4$| ?x454)))
+(let (($x676 (<= ?x675 0)))
+(let (($x782 (not $x676)))
+(let ((?x672 (+ ?x131 ?x454)))
+(let (($x673 (<= ?x672 0)))
+(let (($x743 (not $x673)))
+(let ((?x653 (* (~ 1) |x11$|)))
+(let ((?x654 (+ |x2$| ?x653)))
+(let (($x656 (>= ?x654 0)))
+(let (($x708 (not $x656)))
+(let (($x71 (= |x2$| |x11$|)))
+(let ((@x1263 (hypothesis $x656)))
+(let (($x655 (<= ?x654 0)))
+(let ((?x165 (* (~ 1) |x6$|)))
+(let (($x386 (>= |x6$| 0)))
+(let ((?x393 (ite $x386 |x6$| ?x165)))
+(let ((?x404 (* (~ 1) ?x393)))
+(let ((?x669 (+ |x6$| ?x404)))
+(let (($x934 (<= ?x669 0)))
+(let (($x610 (= |x6$| ?x393)))
+(let (($x411 (>= |x5$| 0)))
+(let (($x286 (>= |x9$| 0)))
+(let (($x671 (>= ?x669 0)))
+(let (($x287 (not $x286)))
+(let ((@x1426 (hypothesis $x287)))
+(let ((?x233 (* (~ 1) |x10$|)))
+(let (($x311 (>= |x10$| 0)))
+(let ((?x318 (ite $x311 |x10$| ?x233)))
+(let ((?x329 (* (~ 1) ?x318)))
+(let ((?x660 (+ |x10$| ?x329)))
+(let (($x1369 (<= ?x660 0)))
+(let (($x642 (= |x10$| ?x318)))
+(let (($x643 (= ?x233 ?x318)))
+(let (($x1119 (not $x643)))
+(let ((?x1101 (+ ?x233 ?x329)))
+(let (($x1247 (<= ?x1101 0)))
+(let (($x1259 (not $x1247)))
+(let ((?x216 (* (~ 1) |x9$|)))
+(let ((?x293 (ite $x286 |x9$| ?x216)))
+(let ((?x304 (* (~ 1) ?x293)))
+(let ((?x1498 (+ ?x216 ?x304)))
+(let (($x1543 (>= ?x1498 0)))
+(let (($x635 (= ?x216 ?x293)))
+(let ((@x639 (|def-axiom| (or $x286 $x635))))
+(let ((@x1553 (|unit-resolution| @x639 @x1426 $x635)))
+(let ((@x1572 ((_ |th-lemma| arith triangle-eq) (or (not $x635) $x1543))))
+(let ((@x1573 (|unit-resolution| @x1572 @x1553 $x1543)))
+(let ((?x182 (* (~ 1) |x7$|)))
+(let (($x361 (>= |x7$| 0)))
+(let ((?x368 (ite $x361 |x7$| ?x182)))
+(let ((?x379 (* (~ 1) ?x368)))
+(let ((?x666 (+ |x7$| ?x379)))
+(let (($x838 (<= ?x666 0)))
+(let (($x618 (= |x7$| ?x368)))
+(let (($x412 (not $x411)))
+(let ((@x842 (hypothesis $x412)))
+(let ((?x775 (+ ?x165 ?x404)))
+(let (($x778 (<= ?x775 0)))
+(let (($x611 (= ?x165 ?x393)))
+(let (($x387 (not $x386)))
+(let (($x362 (not $x361)))
+(let ((@x1025 (hypothesis $x362)))
+(let ((@x1024 (hypothesis $x386)))
+(let ((?x405 (+ |x5$| |x7$| ?x404)))
+(let (($x617 (>= ?x405 0)))
+(let (($x406 (= ?x405 0)))
+(let ((?x330 (+ |x9$| |x11$| ?x329)))
+(let (($x331 (= ?x330 0)))
+(let ((?x305 (+ |x8$| |x10$| ?x304)))
+(let (($x306 (= ?x305 0)))
+(let ((?x199 (* (~ 1) |x8$|)))
+(let (($x336 (>= |x8$| 0)))
+(let ((?x343 (ite $x336 |x8$| ?x199)))
+(let ((?x354 (* (~ 1) ?x343)))
+(let ((?x355 (+ |x7$| |x9$| ?x354)))
+(let (($x356 (= ?x355 0)))
+(let ((?x380 (+ |x6$| |x8$| ?x379)))
+(let (($x381 (= ?x380 0)))
+(let ((?x148 (* (~ 1) |x5$|)))
+(let ((?x418 (ite $x411 |x5$| ?x148)))
+(let ((?x429 (* (~ 1) ?x418)))
+(let ((?x430 (+ |x4$| |x6$| ?x429)))
+(let (($x431 (= ?x430 0)))
+(let ((?x455 (+ |x3$| |x5$| ?x454)))
+(let (($x456 (= ?x455 0)))
+(let ((?x114 (* (~ 1) |x3$|)))
+(let (($x461 (>= |x3$| 0)))
+(let ((?x468 (ite $x461 |x3$| ?x114)))
+(let ((?x479 (* (~ 1) ?x468)))
+(let ((?x480 (+ |x2$| |x4$| ?x479)))
+(let (($x481 (= ?x480 0)))
+(let ((?x96 (* (~ 1) |x2$|)))
+(let (($x486 (>= |x2$| 0)))
+(let ((?x493 (ite $x486 |x2$| ?x96)))
+(let ((?x504 (* (~ 1) ?x493)))
+(let ((?x505 (+ |x3$| |x1$| ?x504)))
+(let (($x506 (= ?x505 0)))
+(let (($x535 (and $x506 $x481 $x456 $x431 $x406 $x381 $x356 $x306 $x331)))
+(let (($x546 (not (or (not $x535) (and (= |x1$| |x10$|) $x71)))))
+(let (($x70 (= |x1$| |x10$|)))
+(let (($x72 (and $x70 $x71)))
+(let (($x62 (and (= |x10$| (- (ite (< |x9$| 0) (- |x9$|) |x9$|) |x8$|)) (= |x11$| (- (ite (< |x10$| 0) (- |x10$|) |x10$|) |x9$|)))))
+(let (($x64 (and (= |x8$| (- (ite (< |x7$| 0) (- |x7$|) |x7$|) |x6$|)) (and (= |x9$| (- (ite (< |x8$| 0) (- |x8$|) |x8$|) |x7$|)) $x62))))
+(let (($x66 (and (= |x6$| (- (ite (< |x5$| 0) (- |x5$|) |x5$|) |x4$|)) (and (= |x7$| (- (ite (< |x6$| 0) (- |x6$|) |x6$|) |x5$|)) $x64))))
+(let (($x68 (and (= |x4$| (- (ite (< |x3$| 0) (- |x3$|) |x3$|) |x2$|)) (and (= |x5$| (- (ite (< |x4$| 0) (- |x4$|) |x4$|) |x3$|)) $x66))))
+(let (($x73 (=> (and (= |x3$| (- (ite (< |x2$| 0) (- |x2$|) |x2$|) |x1$|)) $x68) $x72)))
+(let (($x74 (not $x73)))
+(let (($x57 (< |x10$| 0)))
+(let ((?x236 (ite $x57 ?x233 |x10$|)))
+(let ((?x242 (+ ?x216 ?x236)))
+(let (($x247 (= |x11$| ?x242)))
+(let (($x51 (< |x9$| 0)))
+(let ((?x219 (ite $x51 ?x216 |x9$|)))
+(let ((?x225 (+ ?x199 ?x219)))
+(let (($x230 (= |x10$| ?x225)))
+(let (($x250 (and $x230 $x247)))
+(let (($x45 (< |x8$| 0)))
+(let ((?x202 (ite $x45 ?x199 |x8$|)))
+(let ((?x208 (+ ?x182 ?x202)))
+(let (($x213 (= |x9$| ?x208)))
+(let (($x253 (and $x213 $x250)))
+(let (($x39 (< |x7$| 0)))
+(let ((?x185 (ite $x39 ?x182 |x7$|)))
+(let ((?x191 (+ ?x165 ?x185)))
+(let (($x196 (= |x8$| ?x191)))
+(let (($x256 (and $x196 $x253)))
+(let (($x33 (< |x6$| 0)))
+(let ((?x168 (ite $x33 ?x165 |x6$|)))
+(let ((?x174 (+ ?x148 ?x168)))
+(let (($x179 (= |x7$| ?x174)))
+(let (($x259 (and $x179 $x256)))
+(let (($x27 (< |x5$| 0)))
+(let ((?x151 (ite $x27 ?x148 |x5$|)))
+(let ((?x157 (+ ?x131 ?x151)))
+(let (($x162 (= |x6$| ?x157)))
+(let (($x262 (and $x162 $x259)))
+(let (($x21 (< |x4$| 0)))
+(let ((?x134 (ite $x21 ?x131 |x4$|)))
+(let ((?x140 (+ ?x114 ?x134)))
+(let (($x145 (= |x5$| ?x140)))
+(let (($x265 (and $x145 $x262)))
+(let (($x15 (< |x3$| 0)))
+(let ((?x117 (ite $x15 ?x114 |x3$|)))
+(let ((?x123 (+ ?x96 ?x117)))
+(let (($x128 (= |x4$| ?x123)))
+(let (($x268 (and $x128 $x265)))
+(let (($x8 (< |x2$| 0)))
+(let ((?x99 (ite $x8 ?x96 |x2$|)))
+(let ((?x106 (+ (* (~ 1) |x1$|) ?x99)))
+(let (($x111 (= |x3$| ?x106)))
+(let (($x271 (and $x111 $x268)))
+(let (($x278 (or (not $x271) $x72)))
+(let (($x526 (and $x456 (and $x431 (and $x406 (and $x381 (and $x356 (and $x306 $x331))))))))
+(let (($x524 (= $x262 (and $x431 (and $x406 (and $x381 (and $x356 (and $x306 $x331))))))))
+(let ((@x317 (monotonicity (rewrite (= $x57 (not $x311))) (= ?x236 (ite (not $x311) ?x233 |x10$|)))))
+(let ((@x322 (trans @x317 (rewrite (= (ite (not $x311) ?x233 |x10$|) ?x318)) (= ?x236 ?x318))))
+(let ((@x328 (monotonicity (monotonicity @x322 (= ?x242 (+ ?x216 ?x318))) (= $x247 (= |x11$| (+ ?x216 ?x318))))))
+(let ((@x335 (trans @x328 (rewrite (= (= |x11$| (+ ?x216 ?x318)) $x331)) (= $x247 $x331))))
+(let ((@x292 (monotonicity (rewrite (= $x51 $x287)) (= ?x219 (ite $x287 ?x216 |x9$|)))))
+(let ((@x300 (monotonicity (trans @x292 (rewrite (= (ite $x287 ?x216 |x9$|) ?x293)) (= ?x219 ?x293)) (= ?x225 (+ ?x199 ?x293)))))
+(let ((@x310 (trans (monotonicity @x300 (= $x230 (= |x10$| (+ ?x199 ?x293)))) (rewrite (= (= |x10$| (+ ?x199 ?x293)) $x306)) (= $x230 $x306))))
+(let ((@x342 (monotonicity (rewrite (= $x45 (not $x336))) (= ?x202 (ite (not $x336) ?x199 |x8$|)))))
+(let ((@x347 (trans @x342 (rewrite (= (ite (not $x336) ?x199 |x8$|) ?x343)) (= ?x202 ?x343))))
+(let ((@x353 (monotonicity (monotonicity @x347 (= ?x208 (+ ?x182 ?x343))) (= $x213 (= |x9$| (+ ?x182 ?x343))))))
+(let ((@x360 (trans @x353 (rewrite (= (= |x9$| (+ ?x182 ?x343)) $x356)) (= $x213 $x356))))
+(let ((@x516 (monotonicity @x360 (monotonicity @x310 @x335 (= $x250 (and $x306 $x331))) (= $x253 (and $x356 (and $x306 $x331))))))
+(let ((@x367 (monotonicity (rewrite (= $x39 $x362)) (= ?x185 (ite $x362 ?x182 |x7$|)))))
+(let ((@x375 (monotonicity (trans @x367 (rewrite (= (ite $x362 ?x182 |x7$|) ?x368)) (= ?x185 ?x368)) (= ?x191 (+ ?x165 ?x368)))))
+(let ((@x385 (trans (monotonicity @x375 (= $x196 (= |x8$| (+ ?x165 ?x368)))) (rewrite (= (= |x8$| (+ ?x165 ?x368)) $x381)) (= $x196 $x381))))
+(let ((@x519 (monotonicity @x385 @x516 (= $x256 (and $x381 (and $x356 (and $x306 $x331)))))))
+(let ((@x392 (monotonicity (rewrite (= $x33 $x387)) (= ?x168 (ite $x387 ?x165 |x6$|)))))
+(let ((@x400 (monotonicity (trans @x392 (rewrite (= (ite $x387 ?x165 |x6$|) ?x393)) (= ?x168 ?x393)) (= ?x174 (+ ?x148 ?x393)))))
+(let ((@x410 (trans (monotonicity @x400 (= $x179 (= |x7$| (+ ?x148 ?x393)))) (rewrite (= (= |x7$| (+ ?x148 ?x393)) $x406)) (= $x179 $x406))))
+(let ((@x522 (monotonicity @x410 @x519 (= $x259 (and $x406 (and $x381 (and $x356 (and $x306 $x331))))))))
+(let ((@x417 (monotonicity (rewrite (= $x27 $x412)) (= ?x151 (ite $x412 ?x148 |x5$|)))))
+(let ((@x425 (monotonicity (trans @x417 (rewrite (= (ite $x412 ?x148 |x5$|) ?x418)) (= ?x151 ?x418)) (= ?x157 (+ ?x131 ?x418)))))
+(let ((@x435 (trans (monotonicity @x425 (= $x162 (= |x6$| (+ ?x131 ?x418)))) (rewrite (= (= |x6$| (+ ?x131 ?x418)) $x431)) (= $x162 $x431))))
+(let ((@x442 (monotonicity (rewrite (= $x21 (not $x436))) (= ?x134 (ite (not $x436) ?x131 |x4$|)))))
+(let ((@x447 (trans @x442 (rewrite (= (ite (not $x436) ?x131 |x4$|) ?x443)) (= ?x134 ?x443))))
+(let ((@x453 (monotonicity (monotonicity @x447 (= ?x140 (+ ?x114 ?x443))) (= $x145 (= |x5$| (+ ?x114 ?x443))))))
+(let ((@x460 (trans @x453 (rewrite (= (= |x5$| (+ ?x114 ?x443)) $x456)) (= $x145 $x456))))
+(let ((@x467 (monotonicity (rewrite (= $x15 (not $x461))) (= ?x117 (ite (not $x461) ?x114 |x3$|)))))
+(let ((@x472 (trans @x467 (rewrite (= (ite (not $x461) ?x114 |x3$|) ?x468)) (= ?x117 ?x468))))
+(let ((@x478 (monotonicity (monotonicity @x472 (= ?x123 (+ ?x96 ?x468))) (= $x128 (= |x4$| (+ ?x96 ?x468))))))
+(let ((@x485 (trans @x478 (rewrite (= (= |x4$| (+ ?x96 ?x468)) $x481)) (= $x128 $x481))))
+(let ((@x531 (monotonicity @x485 (monotonicity @x460 (monotonicity @x435 @x522 $x524) (= $x265 $x526)) (= $x268 (and $x481 $x526)))))
+(let ((@x492 (monotonicity (rewrite (= $x8 (not $x486))) (= ?x99 (ite (not $x486) ?x96 |x2$|)))))
+(let ((@x497 (trans @x492 (rewrite (= (ite (not $x486) ?x96 |x2$|) ?x493)) (= ?x99 ?x493))))
+(let ((@x503 (monotonicity (monotonicity @x497 (= ?x106 (+ (* (~ 1) |x1$|) ?x493))) (= $x111 (= |x3$| (+ (* (~ 1) |x1$|) ?x493))))))
+(let ((@x510 (trans @x503 (rewrite (= (= |x3$| (+ (* (~ 1) |x1$|) ?x493)) $x506)) (= $x111 $x506))))
+(let ((@x539 (trans (monotonicity @x510 @x531 (= $x271 (and $x506 (and $x481 $x526)))) (rewrite (= (and $x506 (and $x481 $x526)) $x535)) (= $x271 $x535))))
+(let ((@x545 (monotonicity (monotonicity @x539 (= (not $x271) (not $x535))) (= $x278 (or (not $x535) $x72)))))
+(let ((@x238 (monotonicity (rewrite (= (- |x10$|) ?x233)) (= (ite $x57 (- |x10$|) |x10$|) ?x236))))
+(let ((@x241 (monotonicity @x238 (= (- (ite $x57 (- |x10$|) |x10$|) |x9$|) (- ?x236 |x9$|)))))
+(let ((@x246 (trans @x241 (rewrite (= (- ?x236 |x9$|) ?x242)) (= (- (ite $x57 (- |x10$|) |x10$|) |x9$|) ?x242))))
+(let ((@x249 (monotonicity @x246 (= (= |x11$| (- (ite $x57 (- |x10$|) |x10$|) |x9$|)) $x247))))
+(let ((@x221 (monotonicity (rewrite (= (- |x9$|) ?x216)) (= (ite $x51 (- |x9$|) |x9$|) ?x219))))
+(let ((@x224 (monotonicity @x221 (= (- (ite $x51 (- |x9$|) |x9$|) |x8$|) (- ?x219 |x8$|)))))
+(let ((@x229 (trans @x224 (rewrite (= (- ?x219 |x8$|) ?x225)) (= (- (ite $x51 (- |x9$|) |x9$|) |x8$|) ?x225))))
+(let ((@x232 (monotonicity @x229 (= (= |x10$| (- (ite $x51 (- |x9$|) |x9$|) |x8$|)) $x230))))
+(let ((@x204 (monotonicity (rewrite (= (- |x8$|) ?x199)) (= (ite $x45 (- |x8$|) |x8$|) ?x202))))
+(let ((@x207 (monotonicity @x204 (= (- (ite $x45 (- |x8$|) |x8$|) |x7$|) (- ?x202 |x7$|)))))
+(let ((@x212 (trans @x207 (rewrite (= (- ?x202 |x7$|) ?x208)) (= (- (ite $x45 (- |x8$|) |x8$|) |x7$|) ?x208))))
+(let ((@x215 (monotonicity @x212 (= (= |x9$| (- (ite $x45 (- |x8$|) |x8$|) |x7$|)) $x213))))
+(let ((@x255 (monotonicity @x215 (monotonicity @x232 @x249 (= $x62 $x250)) (= (and (= |x9$| (- (ite $x45 (- |x8$|) |x8$|) |x7$|)) $x62) $x253))))
+(let ((@x187 (monotonicity (rewrite (= (- |x7$|) ?x182)) (= (ite $x39 (- |x7$|) |x7$|) ?x185))))
+(let ((@x190 (monotonicity @x187 (= (- (ite $x39 (- |x7$|) |x7$|) |x6$|) (- ?x185 |x6$|)))))
+(let ((@x195 (trans @x190 (rewrite (= (- ?x185 |x6$|) ?x191)) (= (- (ite $x39 (- |x7$|) |x7$|) |x6$|) ?x191))))
+(let ((@x198 (monotonicity @x195 (= (= |x8$| (- (ite $x39 (- |x7$|) |x7$|) |x6$|)) $x196))))
+(let ((@x170 (monotonicity (rewrite (= (- |x6$|) ?x165)) (= (ite $x33 (- |x6$|) |x6$|) ?x168))))
+(let ((@x173 (monotonicity @x170 (= (- (ite $x33 (- |x6$|) |x6$|) |x5$|) (- ?x168 |x5$|)))))
+(let ((@x178 (trans @x173 (rewrite (= (- ?x168 |x5$|) ?x174)) (= (- (ite $x33 (- |x6$|) |x6$|) |x5$|) ?x174))))
+(let ((@x181 (monotonicity @x178 (= (= |x7$| (- (ite $x33 (- |x6$|) |x6$|) |x5$|)) $x179))))
+(let ((@x261 (monotonicity @x181 (monotonicity @x198 @x255 (= $x64 $x256)) (= (and (= |x7$| (- (ite $x33 (- |x6$|) |x6$|) |x5$|)) $x64) $x259))))
+(let ((@x153 (monotonicity (rewrite (= (- |x5$|) ?x148)) (= (ite $x27 (- |x5$|) |x5$|) ?x151))))
+(let ((@x156 (monotonicity @x153 (= (- (ite $x27 (- |x5$|) |x5$|) |x4$|) (- ?x151 |x4$|)))))
+(let ((@x161 (trans @x156 (rewrite (= (- ?x151 |x4$|) ?x157)) (= (- (ite $x27 (- |x5$|) |x5$|) |x4$|) ?x157))))
+(let ((@x164 (monotonicity @x161 (= (= |x6$| (- (ite $x27 (- |x5$|) |x5$|) |x4$|)) $x162))))
+(let ((@x136 (monotonicity (rewrite (= (- |x4$|) ?x131)) (= (ite $x21 (- |x4$|) |x4$|) ?x134))))
+(let ((@x139 (monotonicity @x136 (= (- (ite $x21 (- |x4$|) |x4$|) |x3$|) (- ?x134 |x3$|)))))
+(let ((@x144 (trans @x139 (rewrite (= (- ?x134 |x3$|) ?x140)) (= (- (ite $x21 (- |x4$|) |x4$|) |x3$|) ?x140))))
+(let ((@x147 (monotonicity @x144 (= (= |x5$| (- (ite $x21 (- |x4$|) |x4$|) |x3$|)) $x145))))
+(let ((@x267 (monotonicity @x147 (monotonicity @x164 @x261 (= $x66 $x262)) (= (and (= |x5$| (- (ite $x21 (- |x4$|) |x4$|) |x3$|)) $x66) $x265))))
+(let ((@x119 (monotonicity (rewrite (= (- |x3$|) ?x114)) (= (ite $x15 (- |x3$|) |x3$|) ?x117))))
+(let ((@x122 (monotonicity @x119 (= (- (ite $x15 (- |x3$|) |x3$|) |x2$|) (- ?x117 |x2$|)))))
+(let ((@x127 (trans @x122 (rewrite (= (- ?x117 |x2$|) ?x123)) (= (- (ite $x15 (- |x3$|) |x3$|) |x2$|) ?x123))))
+(let ((@x130 (monotonicity @x127 (= (= |x4$| (- (ite $x15 (- |x3$|) |x3$|) |x2$|)) $x128))))
+(let ((@x101 (monotonicity (rewrite (= (- |x2$|) ?x96)) (= (ite $x8 (- |x2$|) |x2$|) ?x99))))
+(let ((@x104 (monotonicity @x101 (= (- (ite $x8 (- |x2$|) |x2$|) |x1$|) (- ?x99 |x1$|)))))
+(let ((@x110 (trans @x104 (rewrite (= (- ?x99 |x1$|) ?x106)) (= (- (ite $x8 (- |x2$|) |x2$|) |x1$|) ?x106))))
+(let ((@x113 (monotonicity @x110 (= (= |x3$| (- (ite $x8 (- |x2$|) |x2$|) |x1$|)) $x111))))
+(let ((@x273 (monotonicity @x113 (monotonicity @x130 @x267 (= $x68 $x268)) (= (and (= |x3$| (- (ite $x8 (- |x2$|) |x2$|) |x1$|)) $x68) $x271))))
+(let ((@x282 (trans (monotonicity @x273 (= $x73 (=> $x271 $x72))) (rewrite (= (=> $x271 $x72) $x278)) (= $x73 $x278))))
+(let ((@x550 (trans (monotonicity @x282 (= $x74 (not $x278))) (monotonicity @x545 (= (not $x278) $x546)) (= $x74 $x546))))
+(let ((@x552 (|not-or-elim| (mp (asserted $x74) @x550 $x546) $x535)))
+(let ((@x557 (|and-elim| @x552 $x406)))
+(let ((@x851 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x406) $x617)) @x557 $x617)))
+(let ((@x948 ((_ |th-lemma| arith triangle-eq) (or (not $x610) $x934))))
+(let ((@x1027 (|unit-resolution| @x948 (|unit-resolution| (|def-axiom| (or $x387 $x610)) @x1024 $x610) $x934)))
+(let ((@x1030 (lemma ((_ |th-lemma| arith farkas 1 1 1 1 1) @x1027 @x851 @x1025 @x842 @x1024 false) (or $x361 $x411 $x387))))
+(let ((@x615 (|def-axiom| (or $x386 $x611))))
+(let ((@x1061 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x611) $x778)) (|unit-resolution| @x615 (|unit-resolution| @x1030 @x1025 @x842 $x387) $x611) $x778)))
+(let ((@x1062 ((_ |th-lemma| arith farkas 1 1 1 1 1) (|unit-resolution| @x1030 @x1025 @x842 $x387) @x1025 @x851 @x842 @x1061 false)))
+(let ((@x1064 (lemma @x1062 (or $x361 $x411))))
+(let ((@x621 (|def-axiom| (or $x362 $x618))))
+(let ((@x863 ((_ |th-lemma| arith triangle-eq) (or (not $x618) $x838))))
+(let ((@x1087 (|unit-resolution| @x863 (|unit-resolution| @x621 (|unit-resolution| @x1064 @x842 $x361) $x618) $x838)))
+(let ((?x663 (+ |x8$| ?x354)))
+(let (($x661 (<= ?x663 0)))
+(let (($x626 (= |x8$| ?x343)))
+(let (($x665 (>= ?x663 0)))
+(let ((@x1538 (|unit-resolution| ((_ |th-lemma| arith farkas 1 1) (or $x661 $x665)) (hypothesis (not $x661)) $x665)))
+(let (($x627 (= ?x199 ?x343)))
+(let (($x337 (not $x336)))
+(let ((@x1527 (hypothesis $x337)))
+(let ((@x631 (|def-axiom| (or $x336 $x627))))
+(let ((@x1528 (|unit-resolution| @x631 @x1527 $x627)))
+(let ((?x664 (+ ?x199 ?x354)))
+(let (($x873 (<= ?x664 0)))
+(let (($x1510 (not $x873)))
+(let ((@x856 (hypothesis $x665)))
+(let ((@x1516 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x627) $x873)) (hypothesis $x627) (hypothesis $x1510) false)))
+(let ((@x1517 (lemma @x1516 (or (not $x627) $x873))))
+(let ((@x1532 (|unit-resolution| @x1517 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 2 1) (or (not $x665) $x336 $x1510)) @x1527 @x856 $x1510) @x1528 false)))
+(let ((@x629 (|def-axiom| (or $x337 $x626))))
+(let ((@x1540 (|unit-resolution| @x629 (|unit-resolution| (lemma @x1532 (or $x336 (not $x665))) @x1538 $x336) $x626)))
+(let ((@x1127 ((_ |th-lemma| arith triangle-eq) (or (not $x626) $x661))))
+(let ((@x1542 (lemma (|unit-resolution| @x1127 @x1540 (hypothesis (not $x661)) false) $x661)))
+(let ((@x1206 (hypothesis $x838)))
+(let ((@x1211 (hypothesis $x661)))
+(let ((@x843 (hypothesis $x387)))
+(let (($x625 (>= ?x380 0)))
+(let ((@x558 (|and-elim| @x552 $x381)))
+(let ((@x833 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x381) $x625)) @x558 $x625)))
+(let (($x633 (>= ?x355 0)))
+(let ((@x559 (|and-elim| @x552 $x356)))
+(let ((@x1125 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x356) $x633)) @x559 $x633)))
+(let ((@x1429 (lemma ((_ |th-lemma| arith farkas 1 1 1 1 1 1) @x1125 @x1426 @x833 @x843 @x1211 @x1206 false) (or $x286 $x386 (not $x661) (not $x838)))))
+(let ((@x1915 (|unit-resolution| (|unit-resolution| @x1429 @x1542 (or $x286 $x386 (not $x838))) @x1087 @x1426 $x386)))
+(let ((@x613 (|def-axiom| (or $x387 $x610))))
+(let ((@x1917 (|unit-resolution| @x948 (|unit-resolution| @x613 @x1915 $x610) $x934)))
+(let ((?x678 (+ |x3$| ?x479)))
+(let (($x670 (>= ?x678 0)))
+(let (($x586 (= |x3$| ?x468)))
+(let ((?x929 (+ ?x148 ?x429)))
+(let (($x1022 (>= ?x929 0)))
+(let (($x603 (= ?x148 ?x418)))
+(let ((@x607 (|def-axiom| (or $x411 $x603))))
+(let ((@x994 (|unit-resolution| @x607 @x842 $x603)))
+(let ((@x1037 ((_ |th-lemma| arith triangle-eq) (or (not $x603) $x1022))))
+(let ((@x1038 (|unit-resolution| @x1037 @x994 $x1022)))
+(let (($x462 (not $x461)))
+(let ((@x686 (hypothesis $x462)))
+(let (($x601 (>= ?x455 0)))
+(let ((@x555 (|and-elim| @x552 $x456)))
+(let ((@x685 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x456) $x601)) @x555 $x601)))
+(let (($x608 (<= ?x430 0)))
+(let ((@x556 (|and-elim| @x552 $x431)))
+(let ((@x810 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x431) $x608)) @x556 $x608)))
+(let ((?x755 (+ |x5$| ?x429)))
+(let (($x773 (<= ?x755 0)))
+(let (($x931 (<= ?x929 0)))
+(let ((@x997 ((_ |th-lemma| arith triangle-eq) (or (not $x603) $x931))))
+(let ((@x998 (|unit-resolution| @x997 @x994 $x931)))
+(let ((@x1067 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 2) (or $x773 (not $x931) $x411)) @x998 @x842 $x773)))
+(let (($x609 (>= ?x430 0)))
+(let ((@x797 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x431) $x609)) @x556 $x609)))
+(let ((@x801 ((_ |th-lemma| arith assign-bounds 1 1 1 1 1) (or $x386 (not $x773) (not $x601) $x461 $x782 (not $x609)))))
+(let (($x594 (= |x4$| ?x443)))
+(let ((@x1070 ((_ |th-lemma| arith assign-bounds 1 1 1 1) (or $x436 (not $x931) $x411 (not $x609) $x386))))
+(let ((@x597 (|def-axiom| (or (not $x436) $x594))))
+(let ((@x1072 (|unit-resolution| @x597 (|unit-resolution| @x1070 @x843 @x797 @x842 @x998 $x436) $x594)))
+(let ((@x691 ((_ |th-lemma| arith triangle-eq) (or (not $x594) $x676))))
+(let ((@x1073 (|unit-resolution| @x691 @x1072 (|unit-resolution| @x801 @x843 @x797 @x1067 @x686 @x685 $x782) false)))
+(let ((@x1081 (|unit-resolution| ((_ |th-lemma| arith assign-bounds -1 1 -1 1 -1) (or $x743 (not $x601) $x461 (not $x1022) (not $x608) $x387)) (|unit-resolution| (lemma @x1073 (or $x386 $x461 $x411)) @x686 @x842 $x386) @x810 @x685 @x686 @x1038 $x743)))
+(let (($x595 (= ?x131 ?x443)))
+(let (($x437 (not $x436)))
+(let ((@x1082 (|unit-resolution| @x613 (|unit-resolution| (lemma @x1073 (or $x386 $x461 $x411)) @x686 @x842 $x386) $x610)))
+(let ((@x806 ((_ |th-lemma| arith triangle-eq) (or (not $x610) $x671))))
+(let (($x668 (>= ?x666 0)))
+(let ((@x924 ((_ |th-lemma| arith triangle-eq) (or (not $x618) $x668))))
+(let ((@x1086 (|unit-resolution| @x924 (|unit-resolution| @x621 (|unit-resolution| @x1064 @x842 $x361) $x618) $x668)))
+(let ((@x1092 ((_ |th-lemma| arith assign-bounds 1 1 1 1 1) (or $x336 (not $x625) (not $x838) (not $x934) (not $x617) $x411))))
+(let ((@x1093 (|unit-resolution| @x1092 (|unit-resolution| @x948 @x1082 $x934) @x833 @x842 @x851 @x1087 $x336)))
+(let ((@x681 (hypothesis $x668)))
+(let ((@x692 (|unit-resolution| @x691 (|unit-resolution| @x597 (hypothesis $x436) $x594) $x676)))
+(let ((@x687 (hypothesis $x436)))
+(let (($x616 (<= ?x405 0)))
+(let ((@x696 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x406) $x616)) @x557 $x616)))
+(let ((@x697 (hypothesis $x671)))
+(let (($x624 (<= ?x380 0)))
+(let ((@x701 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x381) $x624)) @x558 $x624)))
+(let ((@x702 (hypothesis $x336)))
+(let ((@x707 (lemma ((_ |th-lemma| arith farkas 1 -1 1 -1 1 -1 -1 1 1) @x702 @x701 @x697 @x696 @x687 @x692 @x686 @x685 @x681 false) (or $x461 $x337 (not $x671) $x437 (not $x668)))))
+(let ((@x1094 (|unit-resolution| @x707 @x1093 @x1086 @x686 (|unit-resolution| @x806 @x1082 $x671) $x437)))
+(let ((@x599 (|def-axiom| (or $x436 $x595))))
+(let ((@x738 ((_ |th-lemma| arith triangle-eq) (or (not $x595) $x673))))
+(let ((@x1098 (lemma (|unit-resolution| @x738 (|unit-resolution| @x599 @x1094 $x595) @x1081 false) (or $x461 $x411))))
+(let ((@x589 (|def-axiom| (or $x462 $x586))))
+(let ((@x1268 ((_ |th-lemma| arith triangle-eq) (or (not $x586) $x670))))
+(let ((@x1269 (|unit-resolution| @x1268 (|unit-resolution| @x589 (|unit-resolution| @x1098 @x842 $x461) $x586) $x670)))
+(let (($x667 (>= ?x675 0)))
+(let (($x1499 (<= ?x1498 0)))
+(let ((@x1556 ((_ |th-lemma| arith triangle-eq) (or (not $x635) $x1499))))
+(let ((@x1557 (|unit-resolution| @x1556 @x1553 $x1499)))
+(let (($x930 (>= ?x672 0)))
+(let ((@x964 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x595) $x930)) (|unit-resolution| @x599 (hypothesis $x437) $x595) $x930)))
+(let ((@x939 (|unit-resolution| @x738 (|unit-resolution| @x599 (hypothesis $x437) $x595) $x673)))
+(let ((@x1185 (hypothesis $x411)))
+(let (($x1090 (not $x838)))
+(let (($x837 (>= ?x775 0)))
+(let ((@x890 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x611) $x837)) (hypothesis $x611) (hypothesis (not $x837)) false)))
+(let ((@x891 (lemma @x890 (or (not $x611) $x837))))
+(let ((@x1133 (|unit-resolution| @x891 (|unit-resolution| @x615 @x843 $x611) $x837)))
+(let ((?x776 (+ ?x182 ?x379)))
+(let (($x777 (<= ?x776 0)))
+(let (($x900 (not $x777)))
+(let ((@x904 (hypothesis $x900)))
+(let (($x619 (= ?x182 ?x368)))
+(let (($x821 (not $x619)))
+(let ((@x823 ((_ |th-lemma| arith triangle-eq) (or $x821 $x777))))
+(let ((@x907 (lemma (|unit-resolution| @x823 (hypothesis $x619) @x904 false) (or $x821 $x777))))
+(let ((@x623 (|def-axiom| (or $x361 $x619))))
+(let ((@x1363 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 2 -1) (or $x777 $x362 $x1090)) (|unit-resolution| @x623 (|unit-resolution| @x907 @x904 $x821) $x361) @x904 $x1090)))
+(let ((@x1364 (|unit-resolution| @x621 (|unit-resolution| @x623 (|unit-resolution| @x907 @x904 $x821) $x361) $x618)))
+(let ((@x1366 (lemma (|unit-resolution| @x863 @x1364 @x1363 false) $x777)))
+(let ((@x1447 (|unit-resolution| ((_ |th-lemma| arith assign-bounds -1 1 -1 1 -1) (or $x900 (not $x625) $x336 (not $x837) (not $x616) $x412)) @x833 @x1366 @x696 (or $x336 (not $x837) $x412))))
+(let ((@x1476 (|unit-resolution| @x1127 (|unit-resolution| @x629 (|unit-resolution| @x1447 @x1133 @x1185 $x336) $x626) $x661)))
+(let ((?x1358 (+ ?x96 ?x504)))
+(let (($x1367 (<= ?x1358 0)))
+(let (($x579 (= ?x96 ?x493)))
+(let (($x487 (not $x486)))
+(let (($x602 (= |x5$| ?x418)))
+(let ((@x605 (|def-axiom| (or $x412 $x602))))
+(let ((@x792 ((_ |th-lemma| arith triangle-eq) (or (not $x602) $x773))))
+(let ((@x1187 (|unit-resolution| @x792 (|unit-resolution| @x605 @x1185 $x602) $x773)))
+(let ((@x761 (hypothesis $x437)))
+(let ((@x1357 (lemma ((_ |th-lemma| arith farkas 1 1 1 1 1) @x1185 @x797 @x761 @x843 @x1187 false) (or $x436 $x412 $x386))))
+(let ((@x826 ((_ |th-lemma| arith triangle-eq) (or (not $x594) $x667))))
+(let ((@x1468 (|unit-resolution| @x826 (|unit-resolution| @x597 (|unit-resolution| @x1357 @x843 @x1185 $x436) $x594) $x667)))
+(let ((@x1115 ((_ |th-lemma| arith triangle-eq) (or (not $x626) $x665))))
+(let ((@x1471 (|unit-resolution| @x1115 (|unit-resolution| @x629 (|unit-resolution| @x1447 @x1133 @x1185 $x336) $x626) $x665)))
+(let ((@x1472 (|unit-resolution| @x691 (|unit-resolution| @x597 (|unit-resolution| @x1357 @x843 @x1185 $x436) $x594) $x676)))
+(let ((@x1473 (|unit-resolution| (|unit-resolution| @x801 @x797 @x685 (or $x386 (not $x773) $x461 $x782)) @x1472 @x1187 @x843 $x461)))
+(let ((@x1475 (|unit-resolution| @x1268 (|unit-resolution| @x589 @x1473 $x586) $x670)))
+(let ((@x848 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x611) $x778)) (|unit-resolution| @x615 @x843 $x611) $x778)))
+(let ((?x657 (+ |x9$| ?x304)))
+(let (($x659 (>= ?x657 0)))
+(let (($x634 (= |x9$| ?x293)))
+(let (($x774 (>= ?x755 0)))
+(let ((@x789 ((_ |th-lemma| arith triangle-eq) (or (not $x602) $x774))))
+(let ((@x1477 (|unit-resolution| @x789 (|unit-resolution| @x605 @x1185 $x602) $x774)))
+(let (($x858 (not $x665)))
+(let (($x901 (not $x667)))
+(let (($x815 (not $x774)))
+(let (($x1196 (not $x661)))
+(let (($x798 (not $x773)))
+(let (($x564 (not $x70)))
+(let (($x658 (<= ?x657 0)))
+(let ((@x1379 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 2 1 1 1 1 1) (or $x286 $x361 (not $x633) $x900 (not $x625) $x386 $x1196)) @x1025 @x833 @x1125 @x843 @x1366 @x1211 $x286)))
+(let ((@x637 (|def-axiom| (or $x287 $x634))))
+(let ((@x1149 ((_ |th-lemma| arith triangle-eq) (or (not $x634) $x658))))
+(let (($x1354 (>= ?x776 0)))
+(let ((@x1385 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or $x821 $x1354)) (|unit-resolution| @x623 @x1025 $x619) $x1354)))
+(let ((@x1207 (hypothesis $x773)))
+(let ((@x866 (hypothesis $x676)))
+(let ((@x1388 (|unit-resolution| (|unit-resolution| @x801 @x797 @x685 (or $x386 $x798 $x461 $x782)) @x866 @x1207 @x843 $x461)))
+(let ((@x1390 (|unit-resolution| @x1268 (|unit-resolution| @x589 @x1388 $x586) $x670)))
+(let ((@x898 (hypothesis $x667)))
+(let (($x641 (>= ?x305 0)))
+(let ((@x560 (|and-elim| @x552 $x306)))
+(let ((@x1136 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x306) $x641)) @x560 $x641)))
+(let ((@x1199 ((_ |th-lemma| arith assign-bounds 1 1 1 1 1) (or $x361 $x311 $x1196 (not $x633) (not $x658) (not $x641)))))
+(let ((@x1393 (|unit-resolution| (|unit-resolution| @x1199 @x1136 @x1125 (or $x361 $x311 $x1196 (not $x658))) (|unit-resolution| @x1149 (|unit-resolution| @x637 @x1379 $x634) $x658) @x1211 @x1025 $x311)))
+(let ((@x645 (|def-axiom| (or (not $x311) $x642))))
+(let ((@x1396 ((_ |th-lemma| arith triangle-eq) (or (not $x642) $x1369))))
+(let (($x1139 (not $x658)))
+(let (($x1374 (not $x1354)))
+(let (($x1260 (not $x670)))
+(let (($x1104 (not $x778)))
+(let (($x1373 (not $x1369)))
+(let ((@x1137 (hypothesis $x658)))
+(let ((@x1370 (hypothesis $x1354)))
+(let (($x592 (<= ?x480 0)))
+(let ((@x554 (|and-elim| @x552 $x481)))
+(let ((@x1252 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x481) $x592)) @x554 $x592)))
+(let (($x600 (<= ?x455 0)))
+(let ((@x830 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x456) $x600)) @x555 $x600)))
+(let ((@x1249 (hypothesis $x670)))
+(let ((@x1248 (hypothesis $x778)))
+(let (($x764 (not $x655)))
+(let ((@x1253 (hypothesis $x764)))
+(let (($x649 (>= ?x330 0)))
+(let ((@x561 (|and-elim| @x552 $x331)))
+(let ((@x1256 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x331) $x649)) @x561 $x649)))
+(let ((@x1371 (hypothesis $x1369)))
+(let ((@x1372 ((_ |th-lemma| arith farkas -1 1 -1 -1 1 -1 -1 -1 1 1 -1 1 1) @x1136 @x1371 @x1256 @x1253 @x1248 @x851 @x898 @x1249 @x830 @x1252 @x1370 @x701 @x1137 false)))
+(let ((@x1376 (lemma @x1372 (or $x655 $x1373 $x1104 $x901 $x1260 $x1374 $x1139))))
+(let ((@x1398 (|unit-resolution| @x1376 (|unit-resolution| @x1396 (|unit-resolution| @x645 @x1393 $x642) $x1369) @x848 @x898 @x1390 @x1385 (|unit-resolution| @x1149 (|unit-resolution| @x637 @x1379 $x634) $x658) $x655)))
+(let ((@x1277 ((_ |th-lemma| arith triangle-eq) (or $x71 $x764 $x708))))
+(let (($x565 (not $x71)))
+(let (($x566 (or $x564 $x565)))
+(let ((@x572 (monotonicity (rewrite (= $x72 (not $x566))) (= (not $x72) (not (not $x566))))))
+(let ((@x576 (trans @x572 (rewrite (= (not (not $x566)) $x566)) (= (not $x72) $x566))))
+(let ((@x577 (mp (|not-or-elim| (mp (asserted $x74) @x550 $x546) (not $x72)) @x576 $x566)))
+(let ((?x650 (+ |x1$| ?x233)))
+(let (($x652 (>= ?x650 0)))
+(let (($x632 (<= ?x355 0)))
+(let ((@x855 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x356) $x632)) @x559 $x632)))
+(let ((@x897 (hypothesis $x774)))
+(let (($x585 (>= ?x505 0)))
+(let ((@x553 (|and-elim| @x552 $x506)))
+(let ((@x1284 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x506) $x585)) @x553 $x585)))
+(let ((@x1404 ((_ |th-lemma| arith assign-bounds 1 1 1 1 1 1 1 1) (or $x487 $x1260 (not $x592) (not $x600) $x901 $x361 (not $x617) $x386 $x1104))))
+(let ((@x1406 (|unit-resolution| @x1404 @x830 @x851 @x1252 (or $x487 $x1260 $x901 $x361 $x386 $x1104))))
+(let ((@x583 (|def-axiom| (or $x486 $x579))))
+(let ((@x1408 (|unit-resolution| @x583 (|unit-resolution| @x1406 @x1025 @x843 @x848 @x898 @x1390 $x487) $x579)))
+(let ((@x1411 ((_ |th-lemma| arith triangle-eq) (or (not $x579) $x1367))))
+(let ((@x1413 ((_ |th-lemma| arith assign-bounds 1 -1 -1 1 -3 3 2 -2 -2 2 1 -1 -1 1 -1 1 -1) (|unit-resolution| @x1411 @x1408 $x1367) @x1284 @x897 @x810 @x898 @x830 @x848 @x851 @x1390 @x1252 (|unit-resolution| @x1396 (|unit-resolution| @x645 @x1393 $x642) $x1369) @x1256 @x1263 @x855 @x1385 @x701 @x856 $x652)))
+(let (($x651 (<= ?x650 0)))
+(let (($x648 (<= ?x330 0)))
+(let ((@x713 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x331) $x648)) @x561 $x648)))
+(let (($x662 (>= ?x660 0)))
+(let ((@x1165 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x642) $x662)) (hypothesis $x642) (hypothesis (not $x662)) false)))
+(let ((@x1166 (lemma @x1165 (or (not $x642) $x662))))
+(let (($x593 (>= ?x480 0)))
+(let ((@x718 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x481) $x593)) @x554 $x593)))
+(let (($x679 (<= ?x678 0)))
+(let ((@x723 ((_ |th-lemma| arith triangle-eq) (or (not $x586) $x679))))
+(let (($x584 (<= ?x505 0)))
+(let ((@x1296 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x506) $x584)) @x553 $x584)))
+(let (($x1368 (>= ?x1358 0)))
+(let ((@x1419 ((_ |th-lemma| arith assign-bounds 1 -1 -1 1 -3 3 2 -2 -2 2 1 -1 -1 1 -1 1 -1) (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x579) $x1368)) @x1408 $x1368) @x1296 @x1207 @x797 @x866 @x685 @x1133 @x696 (|unit-resolution| @x723 (|unit-resolution| @x589 @x1388 $x586) $x679) @x718 (|unit-resolution| @x1166 (|unit-resolution| @x645 @x1393 $x642) $x662) @x713 @x1398 @x1125 @x1366 @x833 @x1211 $x651)))
+(let ((@x1304 ((_ |th-lemma| arith triangle-eq) (or $x70 (not $x651) (not $x652)))))
+(let ((@x1420 (|unit-resolution| @x1304 @x1419 @x1413 (|unit-resolution| @x577 (|unit-resolution| @x1277 @x1398 @x1263 $x71) $x564) false)))
+(let ((@x1478 (|unit-resolution| (lemma @x1420 (or $x361 $x798 $x782 $x1196 $x815 $x901 $x708 $x858 $x386)) @x1263 @x1472 @x1476 @x1477 @x1468 @x1187 @x1471 @x843 $x361)))
+(let ((@x1481 (|unit-resolution| @x1429 (|unit-resolution| @x863 (|unit-resolution| @x621 @x1478 $x618) $x838) @x1476 @x843 $x286)))
+(let ((@x1144 ((_ |th-lemma| arith triangle-eq) (or (not $x634) $x659))))
+(let ((@x1483 (|unit-resolution| @x1144 (|unit-resolution| @x637 @x1481 $x634) $x659)))
+(let (($x1302 (not $x652)))
+(let ((@x729 (hypothesis $x659)))
+(let (($x640 (<= ?x305 0)))
+(let ((@x728 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x306) $x640)) @x560 $x640)))
+(let ((@x1258 ((_ |th-lemma| arith farkas 1/2 -1 -1/2 -1/2 1/2 1/2 -1/2 1/2 -1/2 1/2 -1/2 1/2 -1/2 1/2 1) @x681 @x855 @x701 (hypothesis $x1247) @x1256 @x1253 @x1252 @x1249 @x830 @x729 @x728 @x898 @x1248 @x851 @x856 false)))
+(let ((@x1262 (lemma @x1258 (or $x655 (not $x668) $x1259 $x1260 (not $x659) $x901 $x1104 $x858))))
+(let ((@x1309 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or $x1119 $x1247)) (hypothesis $x643) (hypothesis $x1259) false)))
+(let ((@x1310 (lemma @x1309 (or $x1119 $x1247))))
+(let ((@x1424 (|unit-resolution| @x1310 (|unit-resolution| @x1262 @x1253 @x856 @x1249 @x681 @x898 @x1248 @x729 $x1259) $x1119)))
+(let ((@x647 (|def-axiom| (or $x311 $x643))))
+(let ((@x1431 (|unit-resolution| @x1396 (|unit-resolution| @x645 (|unit-resolution| @x647 @x1424 $x311) $x642) $x1369)))
+(let ((@x1432 ((_ |th-lemma| arith farkas -2 -1 2 1 -1 2 -1 1 1 -1 1 1 1 -1 -1 1) @x855 @x701 @x856 @x729 @x728 (|unit-resolution| @x647 @x1424 $x311) @x1431 @x1256 @x1253 @x1248 @x851 @x898 @x1249 @x830 @x1252 @x681 false)))
+(let ((@x1485 (|unit-resolution| (lemma @x1432 (or $x655 $x858 (not $x659) $x1104 $x901 $x1260 (not $x668))) @x1483 @x1471 @x848 @x1468 @x1475 (|unit-resolution| @x924 (|unit-resolution| @x621 @x1478 $x618) $x668) $x655)))
+(let ((@x1449 (|unit-resolution| @x629 (|unit-resolution| @x1447 (hypothesis $x837) @x1185 $x336) $x626)))
+(let ((@x865 (hypothesis $x837)))
+(let (($x1301 (not $x651)))
+(let ((@x1318 (hypothesis $x1301)))
+(let ((?x1142 (+ |x2$| ?x504)))
+(let (($x1237 (>= ?x1142 0)))
+(let (($x578 (= |x2$| ?x493)))
+(let (($x1409 (not $x579)))
+(let (($x1437 (not $x1368)))
+(let ((@x867 (hypothesis $x679)))
+(let ((@x1436 ((_ |th-lemma| arith farkas -1 1 1 -1 -2 -1 2 1 1 -1 -1 1 -1 1 1) @x1137 @x1136 @x865 @x696 @x866 @x867 @x685 @x718 @x1125 @x1211 @x1296 @x1318 @x1207 @x797 (hypothesis $x1368) false)))
+(let ((@x1439 (lemma @x1436 (or $x1437 $x1139 (not $x837) $x782 (not $x679) $x1196 $x651 $x798))))
+(let ((@x1451 (|unit-resolution| @x1439 @x1318 @x865 @x866 @x867 (|unit-resolution| @x1127 @x1449 $x661) @x1137 @x1187 $x1437)))
+(let ((@x1441 (hypothesis $x579)))
+(let ((@x1442 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or $x1409 $x1368)) @x1441 (hypothesis $x1437) false)))
+(let ((@x1443 (lemma @x1442 (or $x1409 $x1368))))
+(let ((@x581 (|def-axiom| (or $x487 $x578))))
+(let ((@x1454 (|unit-resolution| @x581 (|unit-resolution| @x583 (|unit-resolution| @x1443 @x1451 $x1409) $x486) $x578)))
+(let ((@x1298 ((_ |th-lemma| arith triangle-eq) (or (not $x578) $x1237))))
+(let ((@x1456 ((_ |th-lemma| arith farkas 1/2 -1/2 1/2 -1/2 1/2 -1/2 1/2 -1/2 1/2 1/2 -1/2 1/2 -1/2 1) @x1249 @x1252 (|unit-resolution| @x1298 @x1454 $x1237) @x1296 @x1318 @x1187 @x797 @x1137 @x1136 @x865 @x696 @x1125 (|unit-resolution| @x1127 @x1449 $x661) @x1185 false)))
+(let ((@x1490 (|unit-resolution| (lemma @x1456 (or $x651 $x1260 $x1139 (not $x837) $x412 $x782 (not $x679))) (|unit-resolution| @x1149 (|unit-resolution| @x637 @x1481 $x634) $x658) @x1475 @x1133 @x1185 @x1472 (|unit-resolution| @x723 (|unit-resolution| @x589 @x1473 $x586) $x679) $x651)))
+(let ((@x1491 (|unit-resolution| @x1304 @x1490 (|unit-resolution| @x577 (|unit-resolution| @x1277 @x1485 @x1263 $x71) $x564) $x1302)))
+(let (($x1236 (<= ?x1142 0)))
+(let ((@x1291 ((_ |th-lemma| arith triangle-eq) (or (not $x578) $x1236))))
+(let ((@x1461 (|unit-resolution| @x1291 (|unit-resolution| @x581 (hypothesis $x486) $x578) $x1236)))
+(let ((@x1463 ((_ |th-lemma| arith farkas -1 1 -1 1 -1 1 1 -1 -1 1 1 -1 -2 -2 2 1) @x1284 (hypothesis $x1302) @x897 @x810 @x729 @x728 @x1248 @x851 @x1249 @x1252 @x855 @x856 (hypothesis $x486) @x898 @x830 @x1461 false)))
+(let ((@x1465 (lemma @x1463 (or $x487 $x652 $x815 (not $x659) $x1104 $x1260 $x858 $x901))))
+(let ((@x1493 (|unit-resolution| @x583 (|unit-resolution| @x1465 @x1491 @x1477 @x1483 @x848 @x1475 @x1471 @x1468 $x487) $x579)))
+(let ((@x1495 ((_ |th-lemma| arith farkas -1 1 -1 1 -1 1 1 -1 -1 1 1 -1 -2 2 1) @x1284 @x1491 @x1477 @x810 @x1483 @x728 @x848 @x851 @x1475 @x1252 @x855 @x1471 @x1468 @x830 (|unit-resolution| @x1411 @x1493 $x1367) false)))
+(let (($x704 (not $x671)))
+(let ((@x1150 (|unit-resolution| @x1149 (|unit-resolution| @x637 (hypothesis $x286) $x634) $x658)))
+(let ((@x1076 (hypothesis $x286)))
+(let (($x312 (not $x311)))
+(let (($x1162 (not $x642)))
+(let (($x732 (not $x662)))
+(let ((@x1145 (|unit-resolution| @x1144 (|unit-resolution| @x637 @x1076 $x634) $x659)))
+(let ((@x709 (hypothesis $x708)))
+(let ((@x714 (hypothesis $x662)))
+(let (($x845 (not $x611)))
+(let (($x870 (not $x837)))
+(let ((?x674 (+ ?x114 ?x479)))
+(let (($x677 (<= ?x674 0)))
+(let (($x587 (= ?x114 ?x468)))
+(let ((@x591 (|def-axiom| (or $x461 $x587))))
+(let ((@x760 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x587) $x677)) (|unit-resolution| @x591 @x686 $x587) $x677)))
+(let ((@x942 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 2 1) (or $x676 $x436 $x743)) @x939 @x761 $x676)))
+(let ((@x864 (|unit-resolution| @x863 (|unit-resolution| @x621 (hypothesis $x361) $x618) $x838)))
+(let ((@x839 (hypothesis $x361)))
+(let ((@x868 ((_ |th-lemma| arith farkas -1 1 -1 1 -1 -1 1 1 -1 1 1 -1 -2 1) @x833 @x867 @x729 @x728 @x718 @x714 @x713 @x709 @x685 @x866 @x696 @x865 @x839 @x864 false)))
+(let ((@x877 (|unit-resolution| (lemma @x868 (or $x362 (not $x679) (not $x659) $x732 $x656 $x782 $x870)) @x865 @x729 @x714 @x709 @x866 @x867 $x362)))
+(let ((@x880 ((_ |th-lemma| arith farkas -1 1 -1 1 -1 -1 1 1 -1 1 1 -1 1) @x833 @x867 @x729 @x728 @x718 @x714 @x713 @x709 @x685 @x866 @x696 @x865 (|unit-resolution| @x823 (|unit-resolution| @x623 @x877 $x619) $x777) false)))
+(let ((@x882 (lemma @x880 (or $x870 (not $x679) (not $x659) $x732 $x656 $x782))))
+(let ((@x943 (|unit-resolution| @x882 @x942 @x729 @x714 @x709 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 2) (or $x679 (not $x677) $x461)) @x760 @x686 $x679) $x870)))
+(let ((@x946 (|unit-resolution| @x613 (|unit-resolution| @x615 (|unit-resolution| @x891 @x943 $x845) $x386) $x610)))
+(let ((@x952 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 1 1 1) (or $x411 $x743 (not $x601) $x461 $x436)) @x761 @x685 @x686 @x939 $x411)))
+(let ((@x958 ((_ |th-lemma| arith assign-bounds 1 1 1 1 1) (or $x361 (not $x934) (not $x617) $x436 $x798 (not $x609)))))
+(let ((@x959 (|unit-resolution| @x958 @x761 @x851 @x797 (|unit-resolution| @x792 (|unit-resolution| @x605 @x952 $x602) $x773) (|unit-resolution| @x948 @x946 $x934) $x361)))
+(let ((@x965 ((_ |th-lemma| arith farkas -1 -1 1 1 -1 -1 1 1 1 -1 -1 1 1) @x833 @x729 @x728 @x760 @x718 @x714 @x713 @x709 (|unit-resolution| @x948 @x946 $x934) @x851 @x964 @x830 (|unit-resolution| @x863 (|unit-resolution| @x621 @x959 $x618) $x838) false)))
+(let ((@x972 (|unit-resolution| (lemma @x965 (or $x436 (not $x659) $x732 $x656 $x461)) @x686 @x714 @x709 @x729 $x436)))
+(let ((@x976 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 1 1 1) (or $x411 (not $x601) $x461 $x437 $x782)) (|unit-resolution| @x691 (|unit-resolution| @x597 @x972 $x594) $x676) @x685 @x686 @x972 $x411)))
+(let ((@x979 (|unit-resolution| @x882 (|unit-resolution| @x691 (|unit-resolution| @x597 @x972 $x594) $x676) @x729 @x714 @x709 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 2) (or $x679 (not $x677) $x461)) @x760 @x686 $x679) $x870)))
+(let ((@x982 (|unit-resolution| @x613 (|unit-resolution| @x615 (|unit-resolution| @x891 @x979 $x845) $x386) $x610)))
+(let ((@x933 ((_ |th-lemma| arith farkas -1 -1 1 1 -1 -1 1 1 -1 1 -2 2 -1 1 1) @x833 @x729 @x728 (hypothesis $x677) @x718 @x714 @x713 @x709 @x697 @x696 @x897 @x810 @x898 @x830 (hypothesis $x777) false)))
+(let ((@x969 (lemma @x933 (or $x900 (not $x659) (not $x677) $x732 $x656 $x704 $x815 $x901))))
+(let ((@x984 (|unit-resolution| @x969 @x760 @x729 @x714 @x709 (|unit-resolution| @x806 @x982 $x671) (|unit-resolution| @x789 (|unit-resolution| @x605 @x976 $x602) $x774) (|unit-resolution| @x826 (|unit-resolution| @x597 @x972 $x594) $x667) $x900)))
+(let ((@x987 (|unit-resolution| @x621 (|unit-resolution| @x623 (|unit-resolution| @x907 @x984 $x821) $x361) $x618)))
+(let ((@x989 ((_ |th-lemma| arith farkas -1 -1 1 1 -1 -1 1 1 -1 1 -2 2 -2 -1 1 1) @x833 @x729 @x728 @x760 @x718 @x714 @x713 @x709 (|unit-resolution| @x806 @x982 $x671) @x696 (|unit-resolution| @x789 (|unit-resolution| @x605 @x976 $x602) $x774) @x810 (|unit-resolution| @x623 (|unit-resolution| @x907 @x984 $x821) $x361) (|unit-resolution| @x826 (|unit-resolution| @x597 @x972 $x594) $x667) @x830 (|unit-resolution| @x863 @x987 $x838) false)))
+(let ((@x970 (|unit-resolution| (lemma @x989 (or $x461 (not $x659) $x732 $x656)) @x714 @x729 @x709 $x461)))
+(let ((@x992 (|unit-resolution| @x723 (|unit-resolution| @x589 @x970 $x586) $x679)))
+(let ((@x1009 (|unit-resolution| @x891 (|unit-resolution| @x882 @x942 @x729 @x714 @x709 @x992 $x870) $x845)))
+(let ((@x1012 (|unit-resolution| @x948 (|unit-resolution| @x613 (|unit-resolution| @x615 @x1009 $x386) $x610) $x934)))
+(let ((@x751 (|unit-resolution| ((_ |th-lemma| arith farkas 1 1) (or $x656 $x655)) @x709 $x655)))
+(let ((@x999 (hypothesis $x934)))
+(let ((@x1002 ((_ |th-lemma| arith assign-bounds 1 1 1 1 1 2) (or $x361 (not $x934) (not $x617) $x436 (not $x609) (not $x931) $x411))))
+(let ((@x1004 (|unit-resolution| @x621 (|unit-resolution| @x1002 @x842 @x797 @x851 @x761 @x999 @x998 $x361) $x618)))
+(let ((@x762 (hypothesis $x655)))
+(let ((@x1006 ((_ |th-lemma| arith farkas 1 1 1 2 1 1 1 1 1 1 1 1 1 2 1) @x833 @x999 @x851 @x842 @x729 @x728 @x718 @x714 @x713 @x762 @x685 @x939 @x867 @x761 (|unit-resolution| @x863 @x1004 $x838) false)))
+(let ((@x1008 (lemma @x1006 (or $x411 (not $x934) (not $x659) $x732 $x764 (not $x679) $x436))))
+(let ((@x1014 (|unit-resolution| @x605 (|unit-resolution| @x1008 @x1012 @x729 @x714 @x751 @x992 @x761 $x411) $x602)))
+(let ((@x1016 (|unit-resolution| @x958 (|unit-resolution| @x792 @x1014 $x773) @x851 @x761 @x1012 @x797 $x361)))
+(let ((@x1019 ((_ |th-lemma| arith farkas -1/2 1/2 -1/2 1/2 -1/2 1/2 1/2 -1/2 1/2 1/2 -1/2 -1/2 -1/2 1) @x830 @x964 (|unit-resolution| @x863 (|unit-resolution| @x621 @x1016 $x618) $x838) @x833 @x1012 @x851 @x729 @x728 @x718 @x714 @x713 @x709 @x992 @x970 false)))
+(let ((@x1023 (|unit-resolution| (lemma @x1019 (or $x436 (not $x659) $x732 $x656)) @x714 @x729 @x709 $x436)))
+(let ((@x1033 (|unit-resolution| @x882 (|unit-resolution| @x691 (|unit-resolution| @x597 @x1023 $x594) $x676) @x729 @x714 @x709 @x992 $x870)))
+(let ((@x1035 (|unit-resolution| @x615 (|unit-resolution| @x891 @x1033 $x845) $x386)))
+(let ((@x1041 (|unit-resolution| @x863 (|unit-resolution| @x621 (|unit-resolution| @x1030 @x842 @x1035 $x361) $x618) $x838)))
+(let ((@x1044 ((_ |th-lemma| arith farkas -1 1 -1 1 1 -1 1 1 -1 -1 -1 1 -1 1 1) (|unit-resolution| @x948 (|unit-resolution| @x613 @x1035 $x610) $x934) @x851 @x1041 @x833 @x729 @x728 @x718 @x714 @x713 @x709 @x992 @x1038 @x810 @x970 @x1035 false)))
+(let ((@x1049 (|unit-resolution| (lemma @x1044 (or $x411 (not $x659) $x732 $x656)) @x714 @x729 @x709 $x411)))
+(let ((@x895 (|unit-resolution| @x723 (|unit-resolution| @x589 (hypothesis $x461) $x586) $x679)))
+(let ((@x899 ((_ |th-lemma| arith farkas -1/2 1/2 1 -1 -1/2 1/2 -1/2 1/2 -1/2 1/2 1/2 -1/2 -1/2 -1/2 1/2 1) @x830 @x898 @x897 @x810 (hypothesis $x777) @x833 @x895 @x729 @x728 @x718 @x714 @x713 @x709 @x696 @x697 (hypothesis $x461) false)))
+(let ((@x903 (lemma @x899 (or $x900 $x901 $x815 (not $x659) $x732 $x656 $x704 $x462))))
+(let ((@x1052 (|unit-resolution| @x903 (|unit-resolution| @x789 (|unit-resolution| @x605 @x1049 $x602) $x774) @x970 @x729 @x714 @x709 (|unit-resolution| @x826 (|unit-resolution| @x597 @x1023 $x594) $x667) (|unit-resolution| @x806 (|unit-resolution| @x613 @x1035 $x610) $x671) $x900)))
+(let ((@x1055 (|unit-resolution| @x621 (|unit-resolution| @x623 (|unit-resolution| @x907 @x1052 $x821) $x361) $x618)))
+(let ((@x1057 ((_ |th-lemma| arith farkas 1 -1 1/2 -1/2 1 1/2 -1/2 -1/2 1/2 1/2 -1/2 1/2 1/2 -1/2 -1/2 -1/2 1) (|unit-resolution| @x789 (|unit-resolution| @x605 @x1049 $x602) $x774) @x810 (|unit-resolution| @x826 (|unit-resolution| @x597 @x1023 $x594) $x667) @x830 (|unit-resolution| @x623 (|unit-resolution| @x907 @x1052 $x821) $x361) (|unit-resolution| @x806 (|unit-resolution| @x613 @x1035 $x610) $x671) @x696 (|unit-resolution| @x863 @x1055 $x838) @x833 @x729 @x728 @x718 @x714 @x713 @x709 @x992 @x970 false)))
+(let ((@x1167 (|unit-resolution| (lemma @x1057 (or $x732 (not $x659) $x656)) @x709 @x1145 $x732)))
+(let ((@x1169 (|unit-resolution| @x645 (|unit-resolution| @x1166 @x1167 $x1162) $x312)))
+(let ((@x1191 ((_ |th-lemma| arith assign-bounds 1 1 1 1) (or $x336 $x311 $x1139 (not $x641) $x287))))
+(let ((@x1216 (|unit-resolution| @x629 (|unit-resolution| @x1191 @x1169 @x1136 @x1076 @x1150 $x336) $x626)))
+(let ((@x1217 (|unit-resolution| @x1127 @x1216 $x661)))
+(let ((@x1131 (|unit-resolution| @x723 (|unit-resolution| @x589 (|unit-resolution| @x1098 @x842 $x461) $x586) $x679)))
+(let (($x1103 (>= ?x1101 0)))
+(let ((@x1158 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or $x1119 $x1103)) (hypothesis $x643) (hypothesis (not $x1103)) false)))
+(let ((@x1159 (lemma @x1158 (or $x1119 $x1103))))
+(let ((@x1110 (|unit-resolution| ((_ |th-lemma| arith farkas 1 1) (or $x934 $x671)) (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 2) (or $x704 $x1104 $x386)) @x848 @x843 $x704) $x934)))
+(let ((@x1112 (|unit-resolution| @x629 (|unit-resolution| @x1092 @x1110 @x833 @x851 @x842 @x1087 $x336) $x626)))
+(let ((@x841 (hypothesis $x311)))
+(let ((@x860 (lemma ((_ |th-lemma| arith farkas 1 1 1 1 1 1 1 1 1) @x856 @x855 @x851 @x843 @x729 @x728 @x848 @x842 @x841 false) (or $x411 $x858 $x386 (not $x659) $x312))))
+(let ((@x1117 (|unit-resolution| @x860 (|unit-resolution| @x1115 @x1112 $x665) @x842 @x729 @x843 $x312)))
+(let ((@x1122 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or $x1119 $x1103)) (|unit-resolution| @x647 @x1117 $x643) $x1103)))
+(let ((@x1138 ((_ |th-lemma| arith farkas 1 -1 1 -1 -1 1 -1 -1 1 -1 1 -1 -2 2 1) @x833 @x1137 @x1136 @x1087 @x696 @x1133 @x713 @x709 @x718 (|unit-resolution| @x691 @x1072 $x676) @x685 @x1131 (|unit-resolution| @x1127 @x1112 $x661) @x1125 @x1122 false)))
+(let ((@x1172 (|unit-resolution| (lemma @x1138 (or $x386 $x1139 $x656 $x411 (not $x659))) @x842 @x709 @x1150 @x1145 $x386)))
+(let ((@x1152 ((_ |th-lemma| arith farkas -1/2 1/2 1/2 -1/2 1/2 -1/2 1/2 -1/2 -1/2 1/2 -1/2 1/2 -1/2 1) @x701 @x681 @x697 @x696 (hypothesis $x1103) @x1150 @x1136 @x713 @x709 @x718 @x866 @x685 @x867 @x1076 false)))
+(let ((@x1155 (lemma @x1152 (or (not $x1103) (not $x668) $x704 $x656 $x782 (not $x679) $x287))))
+(let ((@x1175 (|unit-resolution| @x1155 (|unit-resolution| @x806 (|unit-resolution| @x613 @x1172 $x610) $x671) (|unit-resolution| @x1159 (|unit-resolution| @x647 @x1169 $x643) $x1103) @x709 @x1131 @x1086 @x1076 $x782)))
+(let ((@x1177 (|unit-resolution| @x1092 @x1087 @x833 @x842 (|unit-resolution| @x948 (|unit-resolution| @x613 @x1172 $x610) $x934) @x851 $x336)))
+(let ((@x1102 (lemma ((_ |th-lemma| arith farkas 1 1 1 1 1 1 1 1 1) @x856 @x701 @x1086 @x855 @x761 @x998 @x842 @x797 @x1076 false) (or $x436 $x858 $x411 $x287))))
+(let ((@x1180 (|unit-resolution| @x1102 (|unit-resolution| @x1115 (|unit-resolution| @x629 @x1177 $x626) $x665) @x842 @x1076 $x436)))
+(let ((@x1184 (lemma (|unit-resolution| @x691 (|unit-resolution| @x597 @x1180 $x594) @x1175 false) (or $x411 $x287 $x656))))
+(let ((@x1220 (|unit-resolution| @x789 (|unit-resolution| @x605 (|unit-resolution| @x1184 @x709 @x1076 $x411) $x602) $x774)))
+(let ((@x1193 (|unit-resolution| @x629 (|unit-resolution| @x1191 (hypothesis $x312) @x1136 @x1076 @x1150 $x336) $x626)))
+(let ((@x1188 (hypothesis $x312)))
+(let ((@x1200 (|unit-resolution| @x1199 (|unit-resolution| @x1127 @x1193 $x661) @x1136 @x1188 @x1150 @x1125 $x361)))
+(let ((@x1203 ((_ |th-lemma| arith farkas -1 1 -1 -1 -1 1 1 -1 1) @x1185 @x701 (|unit-resolution| @x924 (|unit-resolution| @x621 @x1200 $x618) $x668) @x1076 (|unit-resolution| @x1115 @x1193 $x665) @x855 @x761 @x797 @x1187 false)))
+(let ((@x1205 (lemma @x1203 (or $x436 $x412 $x287 $x311))))
+(let ((@x1221 (|unit-resolution| @x1205 (|unit-resolution| @x1184 @x709 @x1076 $x411) @x1076 @x1169 $x436)))
+(let (($x816 (not $x608)))
+(let (($x1197 (not $x633)))
+(let (($x1189 (not $x641)))
+(let (($x741 (not $x616)))
+(let ((@x1224 (|unit-resolution| ((_ |th-lemma| arith assign-bounds -1 -1 -1 1 1 -1 1 1 -1) (or $x704 $x741 $x311 $x1139 $x1189 $x815 $x1196 $x1197 $x437 $x816)) @x1169 @x696 @x1125 @x1136 @x810 @x1150 @x1221 @x1220 @x1217 $x704)))
+(let ((@x1225 (|unit-resolution| @x792 (|unit-resolution| @x605 (|unit-resolution| @x1184 @x709 @x1076 $x411) $x602) $x773)))
+(let ((@x1229 (|unit-resolution| @x621 (|unit-resolution| @x1199 @x1217 @x1136 @x1169 @x1150 @x1125 $x361) $x618)))
+(let ((@x1209 (|unit-resolution| @x589 (|unit-resolution| @x801 @x843 @x797 @x1207 @x866 @x685 $x461) $x586)))
+(let ((@x1212 ((_ |th-lemma| arith farkas -1 -2 2 -1 1 1 -1 -1 1 -1 1 -1 -1 1 1) @x696 @x1211 @x1125 @x1137 @x1136 (hypothesis $x1103) @x713 @x709 @x718 (|unit-resolution| @x723 @x1209 $x679) @x833 @x1206 @x866 @x685 @x1133 false)))
+(let ((@x1231 (|unit-resolution| (lemma @x1212 (or $x386 $x1196 $x1139 (not $x1103) $x656 $x1090 $x782 $x798)) @x1217 @x1150 (|unit-resolution| @x1159 (|unit-resolution| @x647 @x1169 $x643) $x1103) @x709 (|unit-resolution| @x863 @x1229 $x838) (|unit-resolution| @x691 (|unit-resolution| @x597 @x1221 $x594) $x676) @x1225 $x386)))
+(let ((@x1235 (lemma (|unit-resolution| @x806 (|unit-resolution| @x613 @x1231 $x610) @x1224 false) (or $x656 $x287))))
+(let ((@x1502 (|unit-resolution| @x1235 (|unit-resolution| (lemma @x1495 (or $x708 $x412 $x386)) @x843 @x1185 $x708) $x287)))
+(let ((@x1504 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 2 1 1 1 1 1) (or $x286 $x361 $x1197 $x900 (not $x625) $x386 $x1196)) @x1502 @x833 @x1125 @x843 @x1366 @x1476 $x361)))
+(let ((@x1506 (|unit-resolution| @x863 (|unit-resolution| @x621 @x1504 $x618) (|unit-resolution| @x1429 @x1502 @x1476 @x843 $x1090) false)))
+(let ((@x1508 (lemma @x1506 (or $x386 $x412))))
+(let ((@x1815 (|unit-resolution| @x1508 @x1185 $x386)))
+(let (($x1513 (not $x627)))
+(let ((@x1519 (|unit-resolution| ((_ |th-lemma| arith assign-bounds -1 -1 -1 1 1) (or $x1510 $x1197 $x387 $x1374 (not $x624) $x286)) @x1385 @x1125 @x1426 @x1024 @x701 $x1510)))
+(let ((@x1522 (|unit-resolution| @x629 (|unit-resolution| @x631 (|unit-resolution| @x1517 @x1519 $x1513) $x336) $x626)))
+(let ((@x1524 ((_ |th-lemma| arith farkas 1 1 1 1 1) @x1426 @x1125 (|unit-resolution| @x1127 @x1522 $x661) @x1025 (|unit-resolution| @x631 (|unit-resolution| @x1517 @x1519 $x1513) $x336) false)))
+(let ((@x1526 (lemma @x1524 (or $x361 $x286 $x387))))
+(let ((@x1826 (|unit-resolution| @x924 (|unit-resolution| @x621 (|unit-resolution| @x1526 @x1815 @x1426 $x361) $x618) $x668)))
+(let (($x705 (not $x668)))
+(let ((@x1734 (|unit-resolution| @x806 (|unit-resolution| @x613 @x1024 $x610) $x671)))
+(let ((@x1670 (|unit-resolution| @x924 (|unit-resolution| @x621 @x839 $x618) $x668)))
+(let (($x1500 (>= ?x664 0)))
+(let ((@x1546 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or $x1513 $x1500)) @x1528 $x1500)))
+(let ((@x1547 (|unit-resolution| @x1517 @x1528 $x873)))
+(let ((@x1550 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 1 1 1 1 1 1 1) (or $x437 $x815 $x816 $x704 $x741 $x1510 $x1197 $x286 $x336)) @x1426 @x696 @x1527 @x1125 @x810 @x697 @x1477 @x1547 $x437)))
+(let ((@x1552 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x595) $x930)) (|unit-resolution| @x599 @x1550 $x595) $x930)))
+(let ((@x1558 (|unit-resolution| @x738 (|unit-resolution| @x599 @x1550 $x595) $x673)))
+(let (($x740 (not $x624)))
+(let (($x742 (not $x601)))
+(let ((@x1560 ((_ |th-lemma| arith assign-bounds 1 1 1 1 1 1 1 1 1 2 2) (or $x461 $x815 $x816 $x742 $x705 $x740 $x1510 $x1197 $x286 $x743 $x704 $x741))))
+(let ((@x1561 (|unit-resolution| @x1560 @x1426 @x810 @x696 @x701 @x1125 @x685 @x697 @x681 @x1558 @x1477 @x1547 $x461)))
+(let ((@x1566 ((_ |th-lemma| arith assign-bounds 1 1 1 1 1 1 1 1) (or $x311 (not $x1499) $x1189 $x286 $x705 $x412 $x704 $x741 $x740))))
+(let ((@x1568 (|unit-resolution| @x645 (|unit-resolution| @x1566 @x1557 @x701 @x1185 @x1136 @x1426 @x697 @x681 @x696 $x311) $x642)))
+(let ((@x1570 ((_ |th-lemma| arith assign-bounds -1 1 1 -1 -1 -1 -3 3 1 -1 1 1 -2 2 2 -2) (|unit-resolution| @x1396 @x1568 $x1369) @x1256 (|unit-resolution| @x1268 (|unit-resolution| @x589 @x1561 $x586) $x670) @x1252 @x830 @x1206 @x999 @x851 @x833 @x1557 @x1136 @x1552 @x1187 @x797 @x1546 @x855 $x655)))
+(let ((@x1574 (|unit-resolution| @x723 (|unit-resolution| @x589 @x1561 $x586) $x679)))
+(let ((@x1576 ((_ |th-lemma| arith assign-bounds -1 1 1 -1 -1 -1 -3 3 1 -1 1 1 -2 2 2 -2) (|unit-resolution| @x1166 @x1568 $x662) @x713 @x1574 @x718 @x685 @x681 @x697 @x696 @x701 @x1573 @x728 @x1558 @x1477 @x810 @x1547 @x1125 $x656)))
+(let (($x813 (not $x593)))
+(let (($x869 (not $x679)))
+(let (($x1579 (or $x486 $x286 $x336 $x869 $x813 $x742 $x705 $x704 $x741 $x740 $x743 $x815 $x816 $x1510 $x1197)))
+(let ((@x1581 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 2 1 1 1 1 1 3 3 1 1 2 2 2 2) $x1579) @x1426 @x685 @x810 @x696 @x701 @x1527 @x1125 @x718 @x697 @x681 @x1558 @x1477 @x1574 @x1547 $x486)))
+(let (($x812 (not $x640)))
+(let (($x1586 (not $x1543)))
+(let (($x1585 (not $x585)))
+(let (($x1584 (not $x1236)))
+(let (($x1587 (or $x652 $x1584 $x1585 $x815 $x816 $x1510 $x1197 $x704 $x741 $x869 $x813 $x1586 $x812)))
+(let ((@x1589 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 -1 -1 1 1 -1 -1 1 1 -1 -1 1) $x1587) @x1574 @x810 @x696 @x1125 @x728 @x1284 @x697 @x1477 @x718 @x1547 @x1573 (|unit-resolution| @x1291 (|unit-resolution| @x581 @x1581 $x578) $x1236) $x652)))
+(let (($x1564 (not $x1499)))
+(let (($x1401 (not $x592)))
+(let (($x956 (not $x617)))
+(let (($x955 (not $x934)))
+(let (($x1593 (not $x632)))
+(let (($x1592 (not $x1500)))
+(let (($x799 (not $x609)))
+(let (($x1591 (not $x584)))
+(let (($x1321 (not $x1237)))
+(let (($x1594 (or $x651 $x1321 $x1591 $x798 $x799 $x1592 $x1593 $x955 $x956 $x1260 $x1401 $x1564 $x1189)))
+(let ((@x1596 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 -1 -1 1 1 -1 -1 1 1 -1 -1 1) $x1594) (|unit-resolution| @x1268 (|unit-resolution| @x589 @x1561 $x586) $x670) @x797 @x851 @x855 @x1136 @x1296 @x1187 @x1252 @x999 (|unit-resolution| @x1298 (|unit-resolution| @x581 @x1581 $x578) $x1237) @x1557 @x1546 $x651)))
+(let ((@x1597 (|unit-resolution| @x1304 @x1596 @x1589 (|unit-resolution| @x577 (|unit-resolution| @x1277 @x1576 @x1570 $x71) $x564) false)))
+(let ((@x1671 (|unit-resolution| (lemma @x1597 (or $x286 $x955 $x704 $x336 $x705 $x1090 $x412)) @x1670 @x697 @x1527 @x999 @x864 @x1185 $x286)))
+(let ((@x1673 (|unit-resolution| @x1149 (|unit-resolution| @x637 @x1671 $x634) $x658)))
+(let ((@x1676 (|unit-resolution| (|unit-resolution| @x1191 @x1136 (or $x336 $x311 $x1139 $x287)) @x1673 @x1671 @x1527 $x311)))
+(let ((@x1677 (|unit-resolution| @x1235 @x1671 $x656)))
+(let (($x1654 (or $x655 $x705 $x704 $x1139 $x1104 $x815 $x1564 $x798 $x955 $x1592 $x1090 $x708 $x312)))
+(let ((@x1602 (|unit-resolution| @x1396 (|unit-resolution| @x645 @x841 $x642) $x1369)))
+(let ((@x1600 (hypothesis $x1500)))
+(let ((@x1623 (hypothesis $x1499)))
+(let ((@x1604 ((_ |th-lemma| arith farkas 2 2 2 2 1 1 1 1 1 1 1 1 1 1) (hypothesis $x487) @x1602 @x1256 @x1263 @x1136 @x761 @x1207 @x797 @x999 @x851 @x1600 @x855 @x841 @x1137 false)))
+(let ((@x1620 (|unit-resolution| (lemma @x1604 (or $x486 $x708 $x436 $x798 $x955 $x1592 $x312 $x1139)) @x761 @x1263 @x1207 @x999 @x1600 @x841 @x1137 $x486)))
+(let (($x1626 (not $x930)))
+(let (($x1089 (not $x625)))
+(let (($x1402 (not $x600)))
+(let (($x1625 (not $x649)))
+(let (($x1627 (or $x1301 $x1584 $x1585 $x798 $x799 $x1592 $x1593 $x955 $x956 $x1373 $x1625 $x655 $x1402 $x1090 $x1089 $x1626)))
+(let ((@x1629 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 -1 1 -1 -1 1 2 -2 1 -1 -1 1 1 -1 -1) $x1627) @x964 @x797 @x851 @x833 @x855 @x1256 @x1284 @x1253 @x1207 @x999 @x830 @x1206 @x1602 @x1600 (|unit-resolution| @x1291 (|unit-resolution| @x581 @x1620 $x578) $x1236) $x1301)))
+(let ((@x1630 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 -1 -1 1 1 -1 -1 1 1 -1 -1 1) $x1594) @x1629 @x797 @x851 @x855 @x1136 @x1623 @x1207 (|unit-resolution| @x1298 (|unit-resolution| @x581 @x1620 $x578) $x1237) @x999 @x1252 @x1296 @x1600 $x1260)))
+(let (($x757 (not $x587)))
+(let (($x1607 (>= ?x674 0)))
+(let (($x1611 (not $x1607)))
+(let ((@x1609 (hypothesis $x673)))
+(let ((@x1610 ((_ |th-lemma| arith farkas 1 1 -1 1 -1 -1 -1 -1 1 1 -1 1 1) @x685 @x697 @x696 @x681 @x701 @x1609 @x1252 @x1371 @x1256 @x1253 @x1137 @x1136 (hypothesis $x1607) false)))
+(let ((@x1613 (lemma @x1610 (or $x1611 $x704 $x705 $x743 $x1373 $x655 $x1139))))
+(let ((@x1618 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or $x757 $x1607)) (hypothesis $x587) (hypothesis $x1611) false)))
+(let ((@x1619 (lemma @x1618 (or $x757 $x1607))))
+(let ((@x1632 (|unit-resolution| @x1619 (|unit-resolution| @x1613 @x939 @x681 @x697 @x1602 @x1253 @x1137 $x1611) $x757)))
+(let ((@x1635 (|unit-resolution| @x1268 (|unit-resolution| @x589 (|unit-resolution| @x591 @x1632 $x461) $x586) @x1630 false)))
+(let ((@x1637 (lemma @x1635 (or $x436 $x705 $x704 $x655 $x1139 $x1564 $x798 $x955 $x1592 $x1090 $x708 $x312))))
+(let ((@x1638 (|unit-resolution| @x1637 @x1253 @x697 @x681 @x1137 @x1623 @x1207 @x999 @x1600 @x1206 @x1263 @x841 $x436)))
+(let ((@x1641 (|unit-resolution| ((_ |th-lemma| arith assign-bounds -1 -2 -2 2 -2 2) (or $x1354 $x705 $x437 $x815 $x816 $x704 $x741)) @x1638 @x696 @x697 @x681 @x897 @x810 $x1354)))
+(let ((@x1644 (|unit-resolution| @x1376 (|unit-resolution| @x826 (|unit-resolution| @x597 @x1638 $x594) $x667) @x1248 @x1137 @x1253 @x1641 @x1602 $x1260)))
+(let ((@x1648 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 2 -1) (or $x673 $x437 $x782)) (|unit-resolution| @x691 (|unit-resolution| @x597 @x1638 $x594) $x676) @x1638 $x673)))
+(let ((@x1650 (|unit-resolution| @x1619 (|unit-resolution| @x1613 @x1648 @x681 @x697 @x1602 @x1253 @x1137 $x1611) $x757)))
+(let ((@x1653 (|unit-resolution| @x1268 (|unit-resolution| @x589 (|unit-resolution| @x591 @x1650 $x461) $x586) @x1644 false)))
+(let ((@x1681 (|unit-resolution| (lemma @x1653 $x1654) @x1670 @x697 @x1673 @x1248 @x1477 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 -2) (or $x1499 $x1139 $x287)) @x1673 @x1671 $x1499) @x1187 @x999 @x1546 @x864 @x1677 @x1676 $x655)))
+(let (($x1665 (or $x436 $x815 $x1510 $x704 $x764 $x705 $x708 $x798 $x955 $x1090 $x1592 $x312 $x1139)))
+(let (($x1658 (or $x652 $x1584 $x1585 $x798 $x799 $x1592 $x1593 $x955 $x956 $x1373 $x1625 $x708 $x1402 $x1090 $x1089 $x1626)))
+(let ((@x1660 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 1 1 1 1 1 2 2 1 1 1 1 1 -1 -1) $x1658) (|unit-resolution| @x1291 (|unit-resolution| @x581 @x1620 $x578) $x1236) @x797 @x851 @x833 @x855 @x1256 @x964 @x1263 @x1207 @x999 @x830 @x1206 @x1602 @x1600 @x1284 $x652)))
+(let ((@x1661 (|unit-resolution| @x1304 @x1660 (|unit-resolution| @x577 (|unit-resolution| @x1277 @x762 @x1263 $x71) $x564) $x1301)))
+(let ((@x1664 ((_ |th-lemma| arith farkas 1 -1 1 -1 -1 1 2 -2 1 -1 -1 1 1 -1 -1 1) (|unit-resolution| @x1298 (|unit-resolution| @x581 @x1620 $x578) $x1237) @x1296 @x897 @x810 (hypothesis $x873) @x1125 @x697 @x696 (|unit-resolution| @x1166 (|unit-resolution| @x645 @x841 $x642) $x662) @x713 @x762 @x685 @x681 @x701 @x939 @x1661 false)))
+(let ((@x1682 (|unit-resolution| (lemma @x1664 $x1665) @x1681 @x1547 @x697 @x1477 @x1670 @x1677 @x1187 @x999 @x864 @x1546 @x1676 @x1673 $x436)))
+(let ((@x1694 (|unit-resolution| ((_ |th-lemma| arith assign-bounds -2 2 -2 -2 2 -1) (or $x930 $x815 $x816 $x704 $x362 $x741 $x901)) @x696 @x810 (or $x930 $x815 $x704 $x362 $x901))))
+(let ((@x1695 (|unit-resolution| @x1694 (|unit-resolution| @x826 (|unit-resolution| @x597 @x1682 $x594) $x667) @x697 @x839 @x1477 $x930)))
+(let ((@x1667 ((_ |th-lemma| arith farkas 1 -1 1 -1 -1 -1 1 1 -1 1 1) @x681 @x701 @x697 @x696 (hypothesis $x487) @x1371 @x1256 @x1263 @x1137 @x1136 @x1185 false)))
+(let ((@x1669 (lemma @x1667 (or $x486 $x705 $x704 $x1373 $x708 $x1139 $x412))))
+(let ((@x1696 (|unit-resolution| @x1669 @x1670 @x697 (|unit-resolution| @x1396 (|unit-resolution| @x645 @x1676 $x642) $x1369) @x1677 @x1673 @x1185 $x486)))
+(let ((@x1699 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 1 1 1 1 1 2 2 1 1 1 1 1 -1 -1) $x1658) (|unit-resolution| @x1291 (|unit-resolution| @x581 @x1696 $x578) $x1236) @x797 @x851 @x833 @x855 @x1256 @x1695 @x1677 @x1187 @x999 @x830 @x864 (|unit-resolution| @x1396 (|unit-resolution| @x645 @x1676 $x642) $x1369) @x1546 @x1284 $x652)))
+(let ((@x1700 (|unit-resolution| @x1304 @x1699 (|unit-resolution| @x577 (|unit-resolution| @x1277 @x1681 @x1677 $x71) $x564) $x1301)))
+(let ((@x1702 ((_ |th-lemma| arith farkas -2 -1 1 -1 -1 1 1 -1 -2 2 -1 1 1 -1 -1 1 1) @x1682 (|unit-resolution| @x1298 (|unit-resolution| @x581 @x1696 $x578) $x1237) @x1296 @x1700 @x1477 @x810 @x1547 @x1125 @x697 @x696 (|unit-resolution| @x1166 (|unit-resolution| @x645 @x1676 $x642) $x662) @x713 @x1681 @x685 @x1670 @x701 (|unit-resolution| @x691 (|unit-resolution| @x597 @x1682 $x594) $x676) false)))
+(let ((@x1736 (|unit-resolution| (lemma @x1702 (or $x362 $x704 $x955 $x412 $x1104 $x336)) @x1527 @x1027 @x1185 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 2 -1) (or $x778 $x387 $x955)) @x1027 @x1024 $x778) @x1734 $x362)))
+(let ((@x1737 (|unit-resolution| (|unit-resolution| ((_ |th-lemma| arith assign-bounds 2 1) (or $x705 $x361 $x900)) @x1366 (or $x705 $x361)) @x1736 $x705)))
+(let ((@x1741 (|unit-resolution| @x1149 (|unit-resolution| @x637 (|unit-resolution| @x1526 @x1736 @x1024 $x286) $x634) $x658)))
+(let ((@x1743 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or $x821 $x1354)) (|unit-resolution| @x623 @x1736 $x619) $x1354)))
+(let ((@x1744 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 2 2 1 2 2 2) (or $x1374 $x1139 $x1189 $x1090 $x1197 $x1196 $x311)) @x1743 @x1542 @x1741 (|unit-resolution| ((_ |th-lemma| arith farkas 1 1) (or $x838 $x668)) @x1737 $x838) @x1136 @x1125 $x311)))
+(let ((@x1747 (|unit-resolution| @x1235 (|unit-resolution| @x1526 @x1736 @x1024 $x286) $x656)))
+(let ((@x1750 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 1 1 1 1 1) (or $x486 $x336 $x1373 $x1625 $x708 $x1139 $x1189)) @x1527 @x1136 @x1256 @x1747 @x1741 (|unit-resolution| @x1396 (|unit-resolution| @x645 @x1744 $x642) $x1369) $x486)))
+(let ((@x1719 (|unit-resolution| @x958 @x851 @x797 (or $x361 $x955 $x436 $x798))))
+(let ((@x1755 (|unit-resolution| @x826 (|unit-resolution| @x597 (|unit-resolution| @x1719 @x1736 @x1027 @x1187 $x436) $x594) $x667)))
+(let (($x1756 (or $x652 $x901 $x1584 $x1585 $x815 $x816 $x1592 $x1593 $x1373 $x1625 $x708 $x1402 $x1089 $x900)))
+(let ((@x1758 (|unit-resolution| ((_ |th-lemma| arith assign-bounds -1 1 -1 -1 1 -1 1 1 -1 -1 1 -1 1) $x1756) @x1747 @x810 @x833 @x855 @x1256 @x1284 @x830 @x1477 @x1755 @x1366 (|unit-resolution| @x1396 (|unit-resolution| @x645 @x1744 $x642) $x1369) @x1546 (|unit-resolution| @x1291 (|unit-resolution| @x581 @x1750 $x578) $x1236) $x652)))
+(let ((@x1709 (|unit-resolution| (|unit-resolution| ((_ |th-lemma| arith assign-bounds 2 1) (or $x705 $x361 $x900)) @x1366 (or $x705 $x361)) @x1025 $x705)))
+(let ((@x1715 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 2 2 1 2 2 2) (or $x1374 $x1139 $x1189 $x1090 $x1197 $x1196 $x311)) @x1385 @x1542 @x1137 (|unit-resolution| ((_ |th-lemma| arith farkas 1 1) (or $x838 $x668)) @x1709 $x838) @x1136 @x1125 $x311)))
+(let ((@x1722 (|unit-resolution| @x691 (|unit-resolution| @x597 (|unit-resolution| @x1719 @x1025 @x999 @x1207 $x436) $x594) $x676)))
+(let (($x1723 (or $x1611 $x955 $x956 $x1401 $x1373 $x1625 $x655 $x1139 $x1189 $x798 $x799 $x782 $x742 $x740 $x1374)))
+(let ((@x1725 (|unit-resolution| ((_ |th-lemma| arith assign-bounds -1 1 -1 -1 1 1 -1 1 -2 2 -1 1 -1 1) $x1723) @x1253 @x797 @x851 @x701 @x1136 @x1256 @x685 @x1137 @x1722 @x1207 @x999 (|unit-resolution| @x1396 (|unit-resolution| @x645 @x1715 $x642) $x1369) @x1385 @x1252 $x1611)))
+(let ((@x1726 (|unit-resolution| @x826 (|unit-resolution| @x597 (|unit-resolution| @x1719 @x1025 @x999 @x1207 $x436) $x594) $x667)))
+(let ((@x1729 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 1 1 1 1 1 1 1) (or $x462 $x361 $x901 $x815 $x816 $x1402 $x1089 $x900 $x336)) @x1025 @x810 @x830 @x833 @x1527 @x897 @x1726 @x1366 $x462)))
+(let ((@x1733 (lemma (|unit-resolution| @x1619 (|unit-resolution| @x591 @x1729 $x587) @x1725 false) (or $x655 $x1139 $x798 $x955 $x361 $x336 $x815))))
+(let ((@x1760 (|unit-resolution| @x1277 (|unit-resolution| @x1733 @x1741 @x1187 @x1027 @x1736 @x1527 @x1477 $x655) @x1747 $x71)))
+(let ((@x1765 (|unit-resolution| @x691 (|unit-resolution| @x597 (|unit-resolution| @x1719 @x1736 @x1027 @x1187 $x436) $x594) $x676)))
+(let ((@x1766 ((_ |th-lemma| arith farkas -1 1 -1 -1 1 -1 1 1 -1 -1 1 -1 1 1) @x1765 (|unit-resolution| @x1298 (|unit-resolution| @x581 @x1750 $x578) $x1237) @x1296 @x1187 @x797 @x1547 @x1125 (|unit-resolution| @x1166 (|unit-resolution| @x645 @x1744 $x642) $x662) @x713 (|unit-resolution| @x1733 @x1741 @x1187 @x1027 @x1736 @x1527 @x1477 $x655) @x685 @x701 @x1743 (|unit-resolution| @x1304 (|unit-resolution| @x577 @x1760 $x564) @x1758 $x1301) false)))
+(let ((@x1768 (lemma @x1766 (or $x336 $x387 $x412))))
+(let ((@x1829 (|unit-resolution| (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 -2) (or $x873 $x1196 $x337)) @x1542 (or $x873 $x337)) (|unit-resolution| @x1768 @x1185 @x1815 $x336) $x873)))
+(let ((@x1820 (|unit-resolution| @x806 (|unit-resolution| @x613 @x1815 $x610) $x671)))
+(let ((@x1805 (hypothesis $x1139)))
+(let ((@x1807 (|unit-resolution| @x1556 @x1553 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 2 1) (or $x658 $x286 $x1564)) @x1426 @x1805 $x1564) false)))
+(let ((@x1811 (|unit-resolution| @x637 (|unit-resolution| (lemma @x1807 (or $x286 $x658)) @x1805 $x286) $x634)))
+(let ((@x1813 (lemma (|unit-resolution| @x1149 @x1811 @x1805 false) $x658)))
+(let (($x1791 (or $x1586 $x815 $x816 $x704 $x741 $x1510 $x1197 $x1139 $x461 $x742 $x705 $x740 $x743)))
+(let ((@x1831 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 2 2 4 4 2 2 1 2 2 2 2 2) $x1791) @x810 @x696 @x701 @x1125 @x1813 @x685 (or $x1586 $x815 $x704 $x1510 $x461 $x705 $x743))))
+(let ((@x1833 (|unit-resolution| @x589 (|unit-resolution| @x1831 @x1820 @x1573 @x1826 @x1609 @x1477 @x1829 $x461) $x586)))
+(let ((@x1836 (|unit-resolution| @x1566 @x701 @x1136 @x696 (or $x311 $x1564 $x286 $x705 $x412 $x704))))
+(let ((@x1838 (|unit-resolution| @x645 (|unit-resolution| @x1836 @x1820 @x1557 @x1426 @x1185 @x1826 $x311) $x642)))
+(let ((@x1842 (|unit-resolution| (|unit-resolution| @x1669 @x1813 (or $x486 $x705 $x704 $x1373 $x708 $x412)) (|unit-resolution| @x1396 @x1838 $x1369) @x1263 @x1826 @x1820 @x1185 $x486)))
+(let ((@x1846 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 -1 -1 1 1 -1 -1 1 1 -1 -1 1) $x1587) @x810 @x696 @x1125 @x728 @x718 @x1284 (or $x652 $x1584 $x815 $x1510 $x704 $x869 $x1586))))
+(let ((@x1847 (|unit-resolution| @x1846 (|unit-resolution| @x1291 (|unit-resolution| @x581 @x1842 $x578) $x1236) @x1573 @x1820 @x1477 (|unit-resolution| @x723 @x1833 $x679) @x1829 $x652)))
+(let ((@x1818 (|unit-resolution| @x1115 (|unit-resolution| @x629 (|unit-resolution| @x1768 @x1185 @x1815 $x336) $x626) $x665)))
+(let ((@x1821 ((_ |th-lemma| arith farkas -1 1/3 -1/3 4/3 1/3 -1/3 -1/3 1/3 1/3 -1/3 1/3 -2/3 2/3 2/3 -2/3 1/3 -1/3 1) @x701 @x1820 @x696 @x1185 @x1249 @x1252 @x1371 @x1256 @x1253 @x1623 @x1136 @x1187 @x797 @x1818 @x855 (hypothesis $x930) @x830 @x681 false)))
+(let ((@x1849 (|unit-resolution| (lemma @x1821 (or $x655 $x412 $x1260 $x1373 $x1564 $x1626 $x705)) @x1185 (|unit-resolution| @x1268 @x1833 $x670) (|unit-resolution| @x1396 @x1838 $x1369) @x1557 (hypothesis $x930) @x1826 $x655)))
+(let ((@x1852 (|unit-resolution| @x1304 (|unit-resolution| @x577 (|unit-resolution| @x1277 @x1849 @x1263 $x71) $x564) @x1847 $x1301)))
+(let ((@x1855 ((_ |th-lemma| arith farkas 1 1 1 1 1 1 1 1 1 1 1 1 2 2 2 1) @x701 @x1820 @x696 (|unit-resolution| @x1268 @x1833 $x670) @x1252 (|unit-resolution| @x1166 @x1838 $x662) @x713 @x1849 @x1557 @x1136 @x1609 @x685 (|unit-resolution| @x1298 (|unit-resolution| @x581 @x1842 $x578) $x1237) @x1296 @x1852 @x1826 false)))
+(let ((@x1858 (|unit-resolution| (lemma @x1855 (or $x412 $x743 $x708 $x1626 $x286)) @x939 @x1263 @x964 @x1426 $x412)))
+(let ((@x1860 (|unit-resolution| @x997 (|unit-resolution| @x607 @x1858 $x603) $x931)))
+(let ((@x1861 (|unit-resolution| @x1037 (|unit-resolution| @x607 @x1858 $x603) $x1022)))
+(let ((@x1865 (|unit-resolution| @x863 (|unit-resolution| @x621 (|unit-resolution| @x1064 @x1858 $x361) $x618) $x838)))
+(let ((@x1868 (|unit-resolution| (|unit-resolution| @x1070 @x797 (or $x436 (not $x931) $x411 $x386)) @x1860 @x761 @x1858 $x386)))
+(let ((@x1874 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 2 2 2 2 2) (or (not $x1022) $x798 $x336 $x1090 $x955 $x956 $x1089)) @x833 @x851 (or (not $x1022) $x798 $x336 $x1090 $x955))))
+(let ((@x1875 (|unit-resolution| @x1874 (|unit-resolution| @x948 (|unit-resolution| @x613 @x1868 $x610) $x934) @x1865 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 2) (or $x773 (not $x931) $x411)) @x1860 @x1858 $x773) @x1861 $x336)))
+(let ((@x1877 (|unit-resolution| @x1115 (|unit-resolution| @x629 @x1875 $x626) $x665)))
+(let ((@x1878 (|unit-resolution| @x924 (|unit-resolution| @x621 (|unit-resolution| @x1064 @x1858 $x361) $x618) $x668)))
+(let ((@x1879 (|unit-resolution| @x806 (|unit-resolution| @x613 @x1868 $x610) $x671)))
+(let (($x1000 (not $x931)))
+(let ((@x1881 ((_ |th-lemma| arith assign-bounds 2 2 1 1 1 1 1 1 1 1 1) (or $x311 $x705 $x740 $x704 $x741 $x1564 $x1189 $x436 $x799 $x858 $x1593 $x1000))))
+(let ((@x1882 (|unit-resolution| @x1881 @x761 @x696 @x701 @x855 @x1136 @x797 @x1879 @x1878 @x1877 @x1860 @x1557 $x311)))
+(let ((@x1887 (|unit-resolution| @x1268 (|unit-resolution| @x589 (|unit-resolution| @x1098 @x1858 $x461) $x586) $x670)))
+(let ((@x1888 (|unit-resolution| @x723 (|unit-resolution| @x589 (|unit-resolution| @x1098 @x1858 $x461) $x586) $x679)))
+(let ((@x1892 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 2 2 2 2 2) (or (not $x1022) $x798 $x486 $x782 $x869 $x742 $x813)) @x685 @x718 (or (not $x1022) $x798 $x486 $x782 $x869))))
+(let ((@x1893 (|unit-resolution| @x1892 @x1861 @x942 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 2) (or $x773 $x1000 $x411)) @x1860 @x1858 $x773) @x1888 $x486)))
+(let (($x1078 (not $x1022)))
+(let (($x1896 (or $x652 $x1090 $x1089 $x955 $x956 $x869 $x813 $x1586 $x812 $x1584 $x1585 $x816 $x1196 $x1197 $x1078)))
+(let ((@x1898 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 2 -2 1 -1 1 -1 -1 1 1 -1 1 1 -1 -1) $x1896) @x1888 @x810 @x851 @x833 @x1125 @x728 @x1284 @x718 (|unit-resolution| @x948 (|unit-resolution| @x613 @x1868 $x610) $x934) @x1865 @x1861 @x1542 @x1573 (|unit-resolution| @x1291 (|unit-resolution| @x581 @x1893 $x578) $x1236) $x652)))
+(let (($x1900 (or $x651 $x705 $x740 $x704 $x741 $x1260 $x1401 $x1564 $x1189 $x1321 $x1591 $x799 $x858 $x1593 $x1000)))
+(let ((@x1902 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 2 -2 1 -1 1 -1 -1 1 1 -1 1 1 -1 -1) $x1900) @x1879 @x797 @x696 @x701 @x855 @x1136 @x1296 @x1252 @x1878 @x1877 @x1887 @x1860 (|unit-resolution| @x1298 (|unit-resolution| @x581 @x1893 $x578) $x1237) @x1557 $x651)))
+(let ((@x1905 (|unit-resolution| @x1277 (|unit-resolution| @x577 (|unit-resolution| @x1304 @x1902 @x1898 $x70) $x565) @x1263 $x764)))
+(let ((@x1906 ((_ |th-lemma| arith farkas -1 -1 -1 1 -3 3 -1 1 -1 1 1 -1 -2 -2 2 2 1) @x1256 @x1905 @x964 @x830 @x1878 @x701 @x1879 @x696 @x1887 @x1252 @x1557 @x1136 @x797 @x1877 @x855 @x1860 (|unit-resolution| @x1396 (|unit-resolution| @x645 @x1882 $x642) $x1369) false)))
+(let ((@x1919 (|unit-resolution| @x597 (|unit-resolution| (lemma @x1906 (or $x436 $x708 $x286)) @x1426 @x1263 $x436) $x594)))
+(let ((@x1922 (|unit-resolution| @x1892 @x1038 (|unit-resolution| @x691 @x1919 $x676) @x1067 @x1131 $x486)))
+(let ((@x1925 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 2 -2 1 -1 1 -1 -1 1 1 -1 1 1 -1 -1) $x1896) @x1917 @x810 @x851 @x833 @x1125 @x728 @x1284 @x718 @x1131 @x1087 @x1038 @x1542 @x1573 (|unit-resolution| @x1291 (|unit-resolution| @x581 @x1922 $x578) $x1236) $x652)))
+(let ((@x1929 (|unit-resolution| @x629 (|unit-resolution| @x1874 @x1917 @x1087 @x1067 @x1038 $x336) $x626)))
+(let ((@x1931 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 2 -2 1 -1 1 -1 -1 1 1 -1 1 1 -1 -1) $x1900) (|unit-resolution| @x1115 @x1929 $x665) @x797 @x696 @x701 @x855 @x1136 @x1296 @x1252 @x1086 (|unit-resolution| @x806 (|unit-resolution| @x613 @x1915 $x610) $x671) @x1269 @x998 (|unit-resolution| @x1298 (|unit-resolution| @x581 @x1922 $x578) $x1237) @x1557 $x651)))
+(let ((@x1934 (|unit-resolution| @x1277 (|unit-resolution| @x577 (|unit-resolution| @x1304 @x1931 @x1925 $x70) $x565) @x1263 $x764)))
+(let ((@x1910 ((_ |th-lemma| arith farkas -1 -1 -1 1 -1 1 1 -1 1 -1 -1 1 1) @x1256 @x1253 @x898 @x830 @x1249 @x1252 @x1206 @x833 @x999 @x851 (hypothesis $x1543) @x728 (hypothesis $x1247) false)))
+(let ((@x1935 (|unit-resolution| (lemma @x1910 (or $x1259 $x655 $x901 $x1260 $x1090 $x955 $x1586)) @x1934 (|unit-resolution| @x826 @x1919 $x667) @x1269 @x1087 @x1917 @x1573 $x1259)))
+(let ((@x1938 (|unit-resolution| @x645 (|unit-resolution| @x647 (|unit-resolution| @x1310 @x1935 $x1119) $x311) $x642)))
+(let ((@x1940 ((_ |th-lemma| arith farkas -1 -1 -2 -1 1 -1 1 1 -1 1 -1 -1 1 1) @x1256 @x1934 (|unit-resolution| @x647 (|unit-resolution| @x1310 @x1935 $x1119) $x311) (|unit-resolution| @x826 @x1919 $x667) @x830 @x1269 @x1252 @x1087 @x833 @x1917 @x851 @x1573 @x728 (|unit-resolution| @x1396 @x1938 $x1369) false)))
+(let ((@x1943 (|unit-resolution| (lemma @x1940 (or $x411 $x708 $x286)) @x1426 @x1263 $x411)))
+(let ((@x1944 (|unit-resolution| @x1508 @x1943 $x386)))
+(let ((@x1948 (|unit-resolution| (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 -2) (or $x873 $x1196 $x337)) @x1542 (or $x873 $x337)) (|unit-resolution| @x1768 @x1943 @x1944 $x336) $x873)))
+(let ((@x1950 (|unit-resolution| ((_ |th-lemma| arith assign-bounds -1 -1 -1 1 1) (or $x1510 $x1197 $x387 $x1374 $x740 $x286)) @x1125 @x701 (or $x1510 $x387 $x1374 $x286))))
+(let ((@x1956 (|unit-resolution| @x924 (|unit-resolution| @x621 (|unit-resolution| @x1526 @x1944 @x1426 $x361) $x618) $x668)))
+(let ((@x1958 (|unit-resolution| ((_ |th-lemma| arith assign-bounds -1 -2 -2 2 -2 2) (or $x1354 $x705 $x437 $x815 $x816 $x704 $x741)) @x696 @x810 (or $x1354 $x705 $x437 $x815 $x704))))
+(let ((@x1959 (|unit-resolution| @x1958 @x1956 (|unit-resolution| @x789 (|unit-resolution| @x605 @x1943 $x602) $x774) (|unit-resolution| (lemma @x1906 (or $x436 $x708 $x286)) @x1426 @x1263 $x436) (|unit-resolution| @x1950 @x1948 @x1426 @x1944 $x1374) (|unit-resolution| @x806 (|unit-resolution| @x613 @x1944 $x610) $x671) false)))
+(let ((@x1992 (|unit-resolution| (lemma @x1959 (or $x286 $x708)) @x1263 $x286)))
+(let ((@x1240 (|unit-resolution| @x613 (|unit-resolution| @x1070 @x761 @x797 @x998 @x842 $x386) $x610)))
+(let ((@x1242 (|unit-resolution| @x1092 (|unit-resolution| @x948 @x1240 $x934) @x833 @x842 @x1087 @x851 $x336)))
+(let ((@x1244 (|unit-resolution| @x1115 (|unit-resolution| @x629 @x1242 $x626) (|unit-resolution| @x1102 @x761 @x842 @x1076 $x858) false)))
+(let ((@x1325 (|unit-resolution| @x597 (|unit-resolution| (lemma @x1244 (or $x436 $x411 $x287)) @x842 @x1076 $x436) $x594)))
+(let ((@x1265 (|unit-resolution| @x629 (|unit-resolution| @x1092 @x1110 @x833 @x842 @x1087 @x851 $x336) $x626)))
+(let ((@x1270 (|unit-resolution| @x860 (|unit-resolution| @x1115 @x1265 $x665) @x842 @x729 @x843 $x312)))
+(let ((@x1274 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or $x1119 $x1247)) (|unit-resolution| @x647 @x1270 $x643) $x1247)))
+(let ((@x1275 (|unit-resolution| @x1262 @x1274 @x1086 @x1269 @x729 @x898 @x848 (|unit-resolution| @x1115 @x1265 $x665) $x655)))
+(let ((@x1287 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 1 1 1 1) (or $x486 $x813 $x411 $x782 $x742 $x869)) @x866 @x685 @x842 @x1131 @x718 $x486)))
+(let ((@x1293 ((_ |th-lemma| arith assign-bounds 1 -3/2 3/2 -1 1/2 -1/2 1/2 -1/2 -1 1 1/2 -1/2 -1/2 1/2 1/2 -1/2 1/2) (|unit-resolution| @x1291 (|unit-resolution| @x581 @x1287 $x578) $x1236) @x718 @x1131 @x1284 @x1087 @x729 @x728 @x833 @x1038 @x810 @x848 @x851 (|unit-resolution| @x1159 (|unit-resolution| @x647 @x1270 $x643) $x1103) @x713 @x1275 @x685 @x866 $x652)))
+(let ((@x1300 ((_ |th-lemma| arith assign-bounds 1 -3/2 3/2 -1 1/2 -1/2 1/2 -1/2 -1 1 1/2 -1/2 -1/2 1/2 1/2 -1/2 1/2) (|unit-resolution| @x1298 (|unit-resolution| @x581 @x1287 $x578) $x1237) @x1252 @x1269 @x1296 @x1086 @x1137 @x1136 @x701 @x998 @x797 @x1133 @x696 @x1274 @x1256 @x1263 @x830 @x898 $x651)))
+(let ((@x1305 (|unit-resolution| @x1304 @x1300 @x1293 (|unit-resolution| @x577 (|unit-resolution| @x1277 @x1275 @x1263 $x71) $x564) false)))
+(let ((@x1329 (|unit-resolution| (lemma @x1305 (or $x386 $x1139 $x708 $x901 (not $x659) $x782 $x411)) (|unit-resolution| @x826 @x1325 $x667) (|unit-resolution| @x1235 @x1076 $x656) @x1150 @x1145 (|unit-resolution| @x691 @x1325 $x676) @x842 $x386)))
+(let ((@x1331 (|unit-resolution| @x948 (|unit-resolution| @x613 @x1329 $x610) $x934)))
+(let ((@x1333 ((_ |th-lemma| arith assign-bounds 2 -1) (or $x778 $x387 $x955))))
+(let ((@x1336 (|unit-resolution| @x629 (|unit-resolution| @x1092 @x1331 @x833 @x842 @x1087 @x851 $x336) $x626)))
+(let ((@x1337 (|unit-resolution| @x1115 @x1336 $x665)))
+(let ((@x1313 (|unit-resolution| @x629 (|unit-resolution| @x1092 @x1027 @x833 @x842 @x1087 @x851 $x336) $x626)))
+(let ((@x1315 ((_ |th-lemma| arith farkas -1 -1 -1 1 -1 1 -1 1 1) @x1024 @x841 @x729 @x728 @x851 @x842 (|unit-resolution| @x1115 @x1313 $x665) @x855 @x1027 false)))
+(let ((@x1338 (|unit-resolution| (lemma @x1315 (or $x312 $x387 (not $x659) $x411)) @x1329 @x1145 @x842 $x312)))
+(let ((@x1341 (|unit-resolution| @x1262 (|unit-resolution| @x1310 (|unit-resolution| @x647 @x1338 $x643) $x1247) @x1337 @x1269 @x1145 (|unit-resolution| @x826 @x1325 $x667) (|unit-resolution| @x1333 @x1331 @x1329 $x778) @x1086 $x655)))
+(let ((@x1343 (|unit-resolution| @x577 (|unit-resolution| @x1277 @x1341 (|unit-resolution| @x1235 @x1076 $x656) $x71) $x564)))
+(let ((@x1344 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 1 1 1 1) (or $x486 $x813 $x411 $x782 $x742 $x869)) (|unit-resolution| @x691 @x1325 $x676) @x685 @x842 @x1131 @x718 $x486)))
+(let ((@x1320 ((_ |th-lemma| arith farkas 1 -1 1/2 -1/2 1/2 -1/2 1/2 -1/2 1/2 -1/2 1/2 -1/2 1/2 -1/2 1/2 1) @x681 @x701 @x856 @x855 @x697 @x1150 @x1136 @x696 (hypothesis $x1237) @x1252 @x1249 @x1296 @x1318 (hypothesis $x931) @x797 @x1076 false)))
+(let ((@x1323 (lemma @x1320 (or $x651 $x705 $x858 $x704 $x1321 $x1260 $x1000 $x287))))
+(let ((@x1348 (|unit-resolution| @x1323 @x1086 @x1337 (|unit-resolution| @x806 (|unit-resolution| @x613 @x1329 $x610) $x671) (|unit-resolution| @x1298 (|unit-resolution| @x581 @x1344 $x578) $x1237) @x1269 @x998 @x1076 $x651)))
+(let ((@x1351 ((_ |th-lemma| arith farkas -1/2 1/2 -1/2 1/2 1/2 -1/2 -1/2 1/2 -1/2 1/2 -1/2 1/2 -1/2 1) @x1331 @x1145 @x728 @x851 @x1337 @x855 (|unit-resolution| @x1291 (|unit-resolution| @x581 @x1344 $x578) $x1236) @x718 @x1131 @x1284 (|unit-resolution| @x1304 @x1348 @x1343 $x1302) @x1038 @x810 @x1329 false)))
+(let ((@x1353 (lemma @x1351 (or $x411 $x287))))
+(let ((@x1993 (|unit-resolution| @x1353 @x1992 $x411)))
+(let ((@x1994 (|unit-resolution| @x1508 @x1993 $x386)))
+(let ((@x1996 (|unit-resolution| @x948 (|unit-resolution| @x613 @x1994 $x610) $x934)))
+(let ((@x1998 (|unit-resolution| @x792 (|unit-resolution| @x605 @x1993 $x602) $x773)))
+(let ((@x1964 (|unit-resolution| @x613 (|unit-resolution| @x1508 (|unit-resolution| @x1353 @x1076 $x411) $x386) $x610)))
+(let ((@x1967 (|unit-resolution| @x789 (|unit-resolution| @x605 (|unit-resolution| @x1353 @x1076 $x411) $x602) $x774)))
+(let ((@x1970 (|unit-resolution| ((_ |th-lemma| arith assign-bounds -1 -1 -1 1 1 -1 1 1 -1) (or $x704 $x741 $x311 $x1139 $x1189 $x815 $x1196 $x1197 $x437 $x816)) @x696 @x1125 @x1136 @x810 (or $x704 $x311 $x1139 $x815 $x1196 $x437))))
+(let ((@x1973 (|unit-resolution| (|unit-resolution| @x1970 @x1542 @x1813 (or $x704 $x311 $x815 $x437)) (|unit-resolution| @x1205 @x1188 @x1076 (|unit-resolution| @x1353 @x1076 $x411) $x436) @x1188 @x1967 (|unit-resolution| @x806 @x1964 $x671) false)))
+(let ((@x2008 (|unit-resolution| @x1115 (|unit-resolution| @x629 (|unit-resolution| @x1768 @x1993 @x1994 $x336) $x626) $x665)))
+(let ((@x2012 (|unit-resolution| @x1144 (|unit-resolution| @x637 @x1992 $x634) $x659)))
+(let ((@x2049 (lemma ((_ |th-lemma| arith farkas 1 -1 1 -1 -1 -1 1 -1 1 1) @x729 @x728 @x856 @x855 @x1207 @x761 @x797 @x999 @x851 @x841 false) (or $x436 (not $x659) $x858 $x798 $x955 $x312))))
+(let ((@x2050 (|unit-resolution| @x2049 @x2012 @x2008 @x1998 @x1996 (|unit-resolution| (lemma @x1973 (or $x311 $x287)) @x1992 $x311) $x436)))
+(let ((@x2000 (|unit-resolution| @x645 (|unit-resolution| (lemma @x1973 (or $x311 $x287)) @x1992 $x311) $x642)))
+(let ((@x2001 (|unit-resolution| @x1396 @x2000 $x1369)))
+(let ((@x2002 (|unit-resolution| @x1333 @x1996 @x1994 $x778)))
+(let ((@x2053 (|unit-resolution| @x806 (|unit-resolution| @x613 @x1994 $x610) $x671)))
+(let ((@x2006 (|unit-resolution| @x1768 @x1993 @x1994 $x336)))
+(let ((@x2027 (|unit-resolution| @x691 (|unit-resolution| @x597 (|unit-resolution| @x1719 @x1025 @x1996 @x1998 $x436) $x594) $x676)))
+(let ((@x2028 (|unit-resolution| @x826 (|unit-resolution| @x597 (|unit-resolution| @x1719 @x1025 @x1996 @x1998 $x436) $x594) $x667)))
+(let ((@x1982 (|unit-resolution| (|unit-resolution| @x1376 @x1813 (or $x655 $x1373 $x1104 $x901 $x1260 $x1374)) @x1253 @x1370 @x898 @x1248 @x1371 $x1260)))
+(let ((@x1984 (|unit-resolution| ((_ |th-lemma| arith assign-bounds -1 1 -1 -1 1 1 -1 1 -2 2 -1 1 -1 1) $x1723) @x797 @x851 @x701 @x1136 @x1256 @x1813 @x685 @x1252 (or $x1611 $x955 $x1373 $x655 $x798 $x782 $x1374))))
+(let ((@x1986 (|unit-resolution| @x1619 (|unit-resolution| @x1984 @x1253 @x1370 @x866 @x1207 @x999 @x1371 $x1611) $x757)))
+(let ((@x1989 (|unit-resolution| @x1268 (|unit-resolution| @x589 (|unit-resolution| @x591 @x1986 $x461) $x586) @x1982 false)))
+(let ((@x1991 (lemma @x1989 (or $x655 $x1374 $x901 $x1104 $x1373 $x782 $x798 $x955))))
+(let ((@x2029 (|unit-resolution| @x1991 @x1385 @x2028 @x2002 @x2001 @x2027 @x1998 @x1996 $x655)))
+(let ((@x2009 (|unit-resolution| @x789 (|unit-resolution| @x605 @x1993 $x602) $x774)))
+(let ((@x2004 (|unit-resolution| @x1277 (|unit-resolution| @x1991 @x1370 @x898 @x2002 @x2001 @x866 @x1998 @x1996 $x655) @x1263 $x71)))
+(let ((@x2010 (|unit-resolution| @x1166 @x2000 $x662)))
+(let (($x731 (not $x659)))
+(let (($x814 (not $x648)))
+(let (($x2015 (or $x652 $x1585 $x732 $x814 $x764 $x901 $x1402 $x858 $x1593 $x815 $x816 $x900 $x1089 $x731 $x812 (not $x1367))))
+(let ((@x2017 (|unit-resolution| ((_ |th-lemma| arith assign-bounds -1 -1 1 1 -1 1 -1 1 -1 1 1 -1 -2 2 1) $x2015) (|unit-resolution| @x1411 @x1441 $x1367) @x810 @x833 @x855 @x728 @x713 @x1284 (|unit-resolution| @x1991 @x1370 @x898 @x2002 @x2001 @x866 @x1998 @x1996 $x655) @x2012 @x2010 @x2009 @x2008 @x898 @x1366 @x830 $x652)))
+(let (($x2019 (or $x651 $x1591 $x1373 $x1625 $x708 $x782 $x742 $x1196 $x1197 $x798 $x799 $x1374 $x740 $x1139 $x1189 $x1437)))
+(let ((@x2021 (|unit-resolution| ((_ |th-lemma| arith assign-bounds -1 -1 1 1 -1 1 -1 1 -1 1 1 -1 -2 2 1) $x2019) (|unit-resolution| @x1443 @x1441 $x1368) @x797 @x701 @x1125 @x1136 @x1256 @x1296 @x1263 @x1813 @x866 @x1998 @x1542 @x2001 @x1370 @x685 $x651)))
+(let ((@x2022 (|unit-resolution| @x1304 @x2021 @x2017 (|unit-resolution| @x577 @x2004 $x564) false)))
+(let ((@x2032 (|unit-resolution| (lemma @x2022 (or $x1409 $x708 $x782 $x1374 $x901)) @x2027 @x1263 @x1385 @x2028 $x1409)))
+(let ((@x2035 (|unit-resolution| @x1291 (|unit-resolution| @x581 (|unit-resolution| @x583 @x2032 $x486) $x578) $x1236)))
+(let ((@x2038 (|unit-resolution| ((_ |th-lemma| arith assign-bounds -1 -2 -2 2 2 -2 2) (or $x1500 $x858 $x487 $x732 $x814 $x764 $x731 $x812)) @x2029 @x713 @x728 @x2012 @x2010 @x2008 (|unit-resolution| @x583 @x2032 $x486) $x1500)))
+(let ((@x2040 (|unit-resolution| ((_ |th-lemma| arith assign-bounds -1 1 -1 -1 1 -1 1 1 -1 -1 1 -1 1) $x1756) @x810 @x833 @x855 @x1256 @x1366 @x830 @x1284 (or $x652 $x901 $x1584 $x815 $x1592 $x1373 $x708))))
+(let ((@x2042 (|unit-resolution| @x1304 (|unit-resolution| @x2040 @x2038 @x2035 @x1263 @x2009 @x2028 @x2001 $x652) (|unit-resolution| @x577 (|unit-resolution| @x1277 @x2029 @x1263 $x71) $x564) $x1301)))
+(let ((@x2043 (|unit-resolution| @x1298 (|unit-resolution| @x581 (|unit-resolution| @x583 @x2032 $x486) $x578) $x1237)))
+(let ((@x2044 ((_ |th-lemma| arith farkas 1/2 -1/2 -1/2 1/2 -1/2 1/2 -1/2 1/2 -1/2 1/2 -1/2 1/2 1/2 -1/2 1) @x2010 @x713 @x2029 @x2043 @x1296 @x2042 @x2027 @x685 @x1542 @x1125 @x1998 @x797 @x1385 @x701 @x2006 false)))
+(let ((@x2055 (|unit-resolution| @x621 (|unit-resolution| (lemma @x2044 (or $x361 $x708)) @x1263 $x361) $x618)))
+(let ((@x1979 (lemma (|unit-resolution| @x924 (hypothesis $x618) (hypothesis $x705) false) (or (not $x618) $x668))))
+(let ((@x2056 (|unit-resolution| @x1979 @x2055 $x668)))
+(let ((@x2059 (|unit-resolution| @x1991 (|unit-resolution| @x826 (|unit-resolution| @x597 @x2050 $x594) $x667) (|unit-resolution| @x1958 @x2050 @x2009 @x2056 @x2053 $x1354) @x2002 @x2001 (|unit-resolution| @x691 (|unit-resolution| @x597 @x2050 $x594) $x676) @x1998 @x1996 $x655)))
+(let ((@x2061 (|unit-resolution| (|unit-resolution| @x1669 @x1813 (or $x486 $x705 $x704 $x1373 $x708 $x412)) @x2056 @x1263 @x2001 @x2053 @x1993 $x486)))
+(let ((@x2063 (|unit-resolution| @x589 (|unit-resolution| @x707 @x2050 @x2053 @x2006 @x2056 $x461) $x586)))
+(let ((@x2065 (|unit-resolution| @x1465 (|unit-resolution| @x1268 @x2063 $x670) @x2009 @x2012 @x2002 @x2061 @x2008 (|unit-resolution| @x826 (|unit-resolution| @x597 @x2050 $x594) $x667) $x652)))
+(let ((@x2071 (|unit-resolution| @x1323 (|unit-resolution| @x1268 @x2063 $x670) @x1992 @x2008 @x2053 (|unit-resolution| @x1298 (|unit-resolution| @x581 @x2061 $x578) $x1237) (|unit-resolution| ((_ |th-lemma| arith assign-bounds 2 -1) (or $x931 $x412 $x798)) @x1998 @x1993 $x931) @x2056 $x651)))
+(let ((@x2073 (|unit-resolution| @x577 (|unit-resolution| @x1304 @x2071 @x2065 $x70) (|unit-resolution| @x1277 @x2059 @x1263 $x71) false)))
+(let ((@x2074 (lemma @x2073 $x708)))
+(let ((@x1771 (|unit-resolution| @x621 (|unit-resolution| @x1526 (|unit-resolution| @x1235 @x709 $x287) @x1024 $x361) $x618)))
+(let ((@x1772 (|unit-resolution| @x924 @x1771 $x668)))
+(let ((@x1773 (|unit-resolution| @x1768 @x1185 @x1024 $x336)))
+(let ((@x1769 (|unit-resolution| @x1235 @x709 $x287)))
+(let ((@x1776 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 -1 1 -1 -1 1 -1 1) (or $x437 $x815 $x816 $x704 $x741 $x1196 $x337 $x286 $x1197)) @x1769 @x696 @x1773 @x1125 @x810 @x1734 @x1477 @x1542 $x437)))
+(let ((@x1782 (|unit-resolution| @x1566 (|unit-resolution| @x1556 (|unit-resolution| @x639 @x1769 $x635) $x1499) @x701 @x1185 @x1136 @x1769 @x1734 @x1772 @x696 $x311)))
+(let ((@x1790 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 2 1) (or $x658 $x286 $x1564)) (|unit-resolution| @x1556 (|unit-resolution| @x639 @x1769 $x635) $x1499) @x1769 $x658)))
+(let ((@x1793 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 2 2 4 4 2 2 1 2 2 2 2 2) $x1791) (|unit-resolution| @x738 (|unit-resolution| @x599 @x1776 $x595) $x673) @x810 @x696 @x701 @x1125 @x1790 @x1734 @x1772 (|unit-resolution| @x1572 (|unit-resolution| @x639 @x1769 $x635) $x1543) @x1477 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 -2) (or $x873 $x1196 $x337)) @x1773 @x1542 $x873) @x685 $x461)))
+(let ((@x1796 ((_ |th-lemma| arith farkas 4 -1 3 -3 -1 1 1 -1 -1 1 -1 2 -2 -2 2 -1 1 1) @x1773 @x701 @x1734 @x696 (|unit-resolution| @x723 (|unit-resolution| @x589 @x1793 $x586) $x679) @x718 (|unit-resolution| @x1166 (|unit-resolution| @x645 @x1782 $x642) $x662) @x713 @x709 (|unit-resolution| @x1572 (|unit-resolution| @x639 @x1769 $x635) $x1543) @x728 @x1477 @x810 @x1542 @x1125 (|unit-resolution| @x738 (|unit-resolution| @x599 @x1776 $x595) $x673) @x685 @x1772 false)))
+(let ((@x2081 (|unit-resolution| (lemma @x1796 (or $x656 $x412 $x387)) @x1815 @x1185 @x2074 false)))
+(let ((@x2082 (lemma @x2081 $x412)))
+(let ((@x2100 (|unit-resolution| @x863 (|unit-resolution| @x621 (|unit-resolution| @x1064 @x2082 $x361) $x618) $x838)))
+(let ((@x2117 (|unit-resolution| @x1572 (|unit-resolution| @x639 (|unit-resolution| @x1235 @x2074 $x287) $x635) $x1543)))
+(let ((@x2101 (|unit-resolution| (|unit-resolution| @x1429 @x1542 (or $x286 $x386 $x1090)) @x2100 (|unit-resolution| @x1235 @x2074 $x287) $x386)))
+(let ((@x2090 (|unit-resolution| @x1556 (|unit-resolution| @x639 (|unit-resolution| @x1235 @x2074 $x287) $x635) $x1499)))
+(let ((@x2078 (|unit-resolution| @x997 @x994 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 2) (or $x815 $x1000 $x411)) @x842 @x897 $x1000) false)))
+(let ((@x2097 (|unit-resolution| ((_ |th-lemma| arith farkas 1 1) (or $x774 $x773)) (|unit-resolution| (lemma @x2078 (or $x411 $x815)) @x2082 $x815) $x773)))
+(let ((@x2104 (|unit-resolution| @x1874 (|unit-resolution| @x948 (|unit-resolution| @x613 @x2101 $x610) $x934) @x2100 @x2097 (|unit-resolution| @x1037 (|unit-resolution| @x607 @x2082 $x603) $x1022) $x336)))
+(let ((@x2107 (|unit-resolution| @x1979 (|unit-resolution| @x621 (|unit-resolution| @x1064 @x2082 $x361) $x618) $x668)))
+(let ((@x2109 (|unit-resolution| @x1881 @x1188 @x696 @x701 @x855 @x1136 @x797 (|unit-resolution| @x806 (|unit-resolution| @x613 @x2101 $x610) $x671) @x2107 (|unit-resolution| @x1115 (|unit-resolution| @x629 @x2104 $x626) $x665) (|unit-resolution| @x997 (|unit-resolution| @x607 @x2082 $x603) $x931) @x2090 $x436)))
+(let ((@x2114 (|unit-resolution| @x723 (|unit-resolution| @x589 (|unit-resolution| @x1098 @x2082 $x461) $x586) $x679)))
+(let ((@x2115 ((_ |th-lemma| arith farkas 1 1 -1 -1 1 -1 -1 1 -1 -1 1 -1 1) @x1136 (|unit-resolution| @x806 (|unit-resolution| @x613 @x2101 $x610) $x671) @x696 @x2090 @x2107 @x701 @x2114 @x718 @x713 @x2074 @x685 (|unit-resolution| @x691 (|unit-resolution| @x597 @x2109 $x594) $x676) (|unit-resolution| @x1159 (|unit-resolution| @x647 @x1188 $x643) $x1103) false)))
+(let ((@x2119 (|unit-resolution| @x1166 (|unit-resolution| @x645 (lemma @x2115 $x311) $x642) $x662)))
+(let ((@x2120 ((_ |th-lemma| arith farkas 1 -1 1 -1 1 1 -1 -1 3 -3 2 -2 2 -2 1 -1 1) @x2114 @x718 @x728 @x2119 @x713 (|unit-resolution| @x948 (|unit-resolution| @x613 @x2101 $x610) $x934) @x851 @x2117 @x2100 @x833 @x1542 @x1125 @x810 (|unit-resolution| @x1037 (|unit-resolution| @x607 @x2082 $x603) $x1022) @x1609 @x685 @x2074 false)))
+(let ((@x2121 (lemma @x2120 $x743)))
+(let (($x736 (not $x595)))
+(let ((@x2125 (|unit-resolution| @x599 (lemma (|unit-resolution| @x738 (hypothesis $x595) @x2121 false) $x736) $x436)))
+(|unit-resolution| @x691 (|unit-resolution| @x597 @x2125 $x594) (|unit-resolution| ((_ |th-lemma| arith assign-bounds 2 -1) (or $x673 $x437 $x782)) @x2125 @x2121 $x782) false)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))
+
+8778062e40723924421e3a1f0c912b62e43b9b81 20 0
+unsat
+((set-logic AUFLIRA)
+(proof
+(let ((?x8 (* 2.0 |x$|)))
+(let ((?x10 (+ ?x8 1.0)))
+(let ((?x6 (+ |x$| |x$|)))
+(let (($x11 (< ?x6 ?x10)))
+(let (($x12 (or false $x11)))
+(let (($x13 (or $x11 $x12)))
+(let (($x14 (not $x13)))
+(let ((@x65 (monotonicity (rewrite (= (< ?x8 (+ 1.0 ?x8)) true)) (= (not (< ?x8 (+ 1.0 ?x8))) (not true)))))
+(let ((@x69 (trans @x65 (rewrite (= (not true) false)) (= (not (< ?x8 (+ 1.0 ?x8))) false))))
+(let ((?x38 (+ 1.0 ?x8)))
+(let (($x41 (< ?x8 ?x38)))
+(let ((@x43 (monotonicity (rewrite (= ?x6 ?x8)) (rewrite (= ?x10 ?x38)) (= $x11 $x41))))
+(let ((@x50 (trans (monotonicity @x43 (= $x12 (or false $x41))) (rewrite (= (or false $x41) $x41)) (= $x12 $x41))))
+(let ((@x57 (trans (monotonicity @x43 @x50 (= $x13 (or $x41 $x41))) (rewrite (= (or $x41 $x41) $x41)) (= $x13 $x41))))
+(let ((@x60 (monotonicity @x57 (= $x14 (not $x41)))))
+(mp (asserted $x14) (trans @x60 @x69 (= $x14 false)) false))))))))))))))))))
+
+bbf5431bd7e9448dc98de52e9b465f05ca123636 113 0
+unsat
+((set-logic <null>)
+(proof
+(let ((?x215 (mod |x$| 2)))
+(let ((?x249 (* (~ 1) ?x215)))
+(let ((?x9 (|mod$| |x$| 2)))
+(let ((?x250 (+ ?x9 ?x249)))
+(let (($x267 (>= ?x250 0)))
+(let (($x251 (= ?x250 0)))
+(let (($x203 (forall ((?v0 Int) (?v1 Int) )(!(let ((?x30 (mod ?v0 ?v1)))
+(let ((?x99 (* (~ 1) ?v1)))
+(let ((?x96 (* (~ 1) ?v0)))
+(let ((?x142 (mod ?x96 ?x99)))
+(let ((?x148 (* (~ 1) ?x142)))
+(let (($x117 (<= ?v1 0)))
+(let ((?x168 (ite $x117 ?x148 ?x30)))
+(let (($x19 (= ?v1 0)))
+(let ((?x173 (ite $x19 ?v0 ?x168)))
+(let ((?x29 (|mod$| ?v0 ?v1)))
+(= ?x29 ?x173))))))))))) :pattern ( (|mod$| ?v0 ?v1) )))
+))
+(let (($x179 (forall ((?v0 Int) (?v1 Int) )(let ((?x30 (mod ?v0 ?v1)))
+(let ((?x99 (* (~ 1) ?v1)))
+(let ((?x96 (* (~ 1) ?v0)))
+(let ((?x142 (mod ?x96 ?x99)))
+(let ((?x148 (* (~ 1) ?x142)))
+(let (($x117 (<= ?v1 0)))
+(let ((?x168 (ite $x117 ?x148 ?x30)))
+(let (($x19 (= ?v1 0)))
+(let ((?x173 (ite $x19 ?v0 ?x168)))
+(let ((?x29 (|mod$| ?v0 ?v1)))
+(= ?x29 ?x173))))))))))))
+))
+(let ((?x30 (mod ?1 ?0)))
+(let ((?x99 (* (~ 1) ?0)))
+(let ((?x96 (* (~ 1) ?1)))
+(let ((?x142 (mod ?x96 ?x99)))
+(let ((?x148 (* (~ 1) ?x142)))
+(let (($x117 (<= ?0 0)))
+(let ((?x168 (ite $x117 ?x148 ?x30)))
+(let (($x19 (= ?0 0)))
+(let ((?x173 (ite $x19 ?1 ?x168)))
+(let ((?x29 (|mod$| ?1 ?0)))
+(let (($x176 (= ?x29 ?x173)))
+(let (($x36 (forall ((?v0 Int) (?v1 Int) )(let (($x19 (= ?v1 0)))
+(let ((?x34 (ite $x19 ?v0 (ite (< 0 ?v1) (mod ?v0 ?v1) (- (mod (- ?v0) (- ?v1)))))))
+(let ((?x29 (|mod$| ?v0 ?v1)))
+(= ?x29 ?x34)))))
+))
+(let (($x162 (forall ((?v0 Int) (?v1 Int) )(let ((?x99 (* (~ 1) ?v1)))
+(let ((?x96 (* (~ 1) ?v0)))
+(let ((?x142 (mod ?x96 ?x99)))
+(let ((?x148 (* (~ 1) ?x142)))
+(let ((?x30 (mod ?v0 ?v1)))
+(let (($x20 (< 0 ?v1)))
+(let ((?x153 (ite $x20 ?x30 ?x148)))
+(let (($x19 (= ?v1 0)))
+(let ((?x156 (ite $x19 ?v0 ?x153)))
+(let ((?x29 (|mod$| ?v0 ?v1)))
+(= ?x29 ?x156))))))))))))
+))
+(let ((@x167 (monotonicity (rewrite (= (< 0 ?0) (not $x117))) (= (ite (< 0 ?0) ?x30 ?x148) (ite (not $x117) ?x30 ?x148)))))
+(let ((@x172 (trans @x167 (rewrite (= (ite (not $x117) ?x30 ?x148) ?x168)) (= (ite (< 0 ?0) ?x30 ?x148) ?x168))))
+(let ((@x175 (monotonicity @x172 (= (ite $x19 ?1 (ite (< 0 ?0) ?x30 ?x148)) ?x173))))
+(let ((@x178 (monotonicity @x175 (= (= ?x29 (ite $x19 ?1 (ite (< 0 ?0) ?x30 ?x148))) $x176))))
+(let (($x20 (< 0 ?0)))
+(let ((?x153 (ite $x20 ?x30 ?x148)))
+(let ((?x156 (ite $x19 ?1 ?x153)))
+(let (($x159 (= ?x29 ?x156)))
+(let (($x160 (= (= ?x29 (ite $x19 ?1 (ite $x20 ?x30 (- (mod (- ?1) (- ?0)))))) $x159)))
+(let ((@x144 (monotonicity (rewrite (= (- ?1) ?x96)) (rewrite (= (- ?0) ?x99)) (= (mod (- ?1) (- ?0)) ?x142))))
+(let ((@x152 (trans (monotonicity @x144 (= (- (mod (- ?1) (- ?0))) (- ?x142))) (rewrite (= (- ?x142) ?x148)) (= (- (mod (- ?1) (- ?0))) ?x148))))
+(let ((@x155 (monotonicity @x152 (= (ite $x20 ?x30 (- (mod (- ?1) (- ?0)))) ?x153))))
+(let ((@x158 (monotonicity @x155 (= (ite $x19 ?1 (ite $x20 ?x30 (- (mod (- ?1) (- ?0))))) ?x156))))
+(let ((@x183 (trans (|quant-intro| (monotonicity @x158 $x160) (= $x36 $x162)) (|quant-intro| @x178 (= $x162 $x179)) (= $x36 $x179))))
+(let ((@x194 (|mp~| (mp (asserted $x36) @x183 $x179) (|nnf-pos| (refl (|~| $x176 $x176)) (|~| $x179 $x179)) $x179)))
+(let ((@x208 (mp @x194 (|quant-intro| (refl (= $x176 $x176)) (= $x179 $x203)) $x203)))
+(let (($x257 (or (not $x203) $x251)))
+(let ((?x212 (* (~ 1) 2)))
+(let ((?x211 (* (~ 1) |x$|)))
+(let ((?x213 (mod ?x211 ?x212)))
+(let ((?x214 (* (~ 1) ?x213)))
+(let (($x210 (<= 2 0)))
+(let ((?x216 (ite $x210 ?x214 ?x215)))
+(let (($x209 (= 2 0)))
+(let ((?x217 (ite $x209 |x$| ?x216)))
+(let (($x218 (= ?x9 ?x217)))
+(let ((@x231 (monotonicity (monotonicity (rewrite (= ?x212 (~ 2))) (= ?x213 (mod ?x211 (~ 2)))) (= ?x214 (* (~ 1) (mod ?x211 (~ 2)))))))
+(let ((@x234 (monotonicity (rewrite (= $x210 false)) @x231 (= ?x216 (ite false (* (~ 1) (mod ?x211 (~ 2))) ?x215)))))
+(let ((@x238 (trans @x234 (rewrite (= (ite false (* (~ 1) (mod ?x211 (~ 2))) ?x215) ?x215)) (= ?x216 ?x215))))
+(let ((@x241 (monotonicity (rewrite (= $x209 false)) @x238 (= ?x217 (ite false |x$| ?x215)))))
+(let ((@x248 (monotonicity (trans @x241 (rewrite (= (ite false |x$| ?x215) ?x215)) (= ?x217 ?x215)) (= $x218 (= ?x9 ?x215)))))
+(let ((@x261 (monotonicity (trans @x248 (rewrite (= (= ?x9 ?x215) $x251)) (= $x218 $x251)) (= (or (not $x203) $x218) $x257))))
+(let ((@x264 (trans @x261 (rewrite (= $x257 $x257)) (= (or (not $x203) $x218) $x257))))
+(let ((@x265 (mp ((_ |quant-inst| |x$| 2) (or (not $x203) $x218)) @x264 $x257)))
+(let ((@x324 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x251) $x267)) (|unit-resolution| @x265 @x208 $x251) $x267)))
+(let (($x292 (>= ?x215 0)))
+(let (($x83 (>= ?x9 0)))
+(let (($x86 (not $x83)))
+(let (($x14 (not (<= (+ |x$| 1) (+ |x$| (+ (* 2 ?x9) 1))))))
+(let ((@x88 (monotonicity (rewrite (= (>= (* 2 ?x9) 0) $x83)) (= (not (>= (* 2 ?x9) 0)) $x86))))
+(let ((?x10 (* 2 ?x9)))
+(let ((?x66 (+ 1 |x$| ?x10)))
+(let (($x71 (<= (+ 1 |x$|) ?x66)))
+(let (($x74 (not $x71)))
+(let ((@x82 (monotonicity (rewrite (= $x71 (>= ?x10 0))) (= $x74 (not (>= ?x10 0))))))
+(let ((@x65 (monotonicity (rewrite (= (+ ?x10 1) (+ 1 ?x10))) (= (+ |x$| (+ ?x10 1)) (+ |x$| (+ 1 ?x10))))))
+(let ((@x70 (trans @x65 (rewrite (= (+ |x$| (+ 1 ?x10)) ?x66)) (= (+ |x$| (+ ?x10 1)) ?x66))))
+(let ((@x73 (monotonicity (rewrite (= (+ |x$| 1) (+ 1 |x$|))) @x70 (= (<= (+ |x$| 1) (+ |x$| (+ ?x10 1))) $x71))))
+(let ((@x92 (trans (monotonicity @x73 (= $x14 $x74)) (trans @x82 @x88 (= $x74 $x86)) (= $x14 $x86))))
+(let ((@x93 (mp (asserted $x14) @x92 $x86)))
+((_ |th-lemma| arith farkas -1 1 1) @x93 (|unit-resolution| ((_ |th-lemma| arith) (or false $x292)) (|true-axiom| true) $x292) @x324 false))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))
+
+b5183bee77d63a5b887fd6f1c6035b47d90e65cb 112 0
+unsat
+((set-logic <null>)
+(proof
+(let ((?x211 (mod |x$| 2)))
+(let (($x305 (>= ?x211 2)))
+(let (($x306 (not $x305)))
+(let ((?x245 (* (~ 1) ?x211)))
+(let ((?x7 (|mod$| |x$| 2)))
+(let ((?x246 (+ ?x7 ?x245)))
+(let (($x262 (<= ?x246 0)))
+(let (($x247 (= ?x246 0)))
+(let (($x199 (forall ((?v0 Int) (?v1 Int) )(!(let ((?x29 (mod ?v0 ?v1)))
+(let ((?x95 (* (~ 1) ?v1)))
+(let ((?x92 (* (~ 1) ?v0)))
+(let ((?x138 (mod ?x92 ?x95)))
+(let ((?x144 (* (~ 1) ?x138)))
+(let (($x113 (<= ?v1 0)))
+(let ((?x164 (ite $x113 ?x144 ?x29)))
+(let (($x18 (= ?v1 0)))
+(let ((?x169 (ite $x18 ?v0 ?x164)))
+(let ((?x28 (|mod$| ?v0 ?v1)))
+(= ?x28 ?x169))))))))))) :pattern ( (|mod$| ?v0 ?v1) )))
+))
+(let (($x175 (forall ((?v0 Int) (?v1 Int) )(let ((?x29 (mod ?v0 ?v1)))
+(let ((?x95 (* (~ 1) ?v1)))
+(let ((?x92 (* (~ 1) ?v0)))
+(let ((?x138 (mod ?x92 ?x95)))
+(let ((?x144 (* (~ 1) ?x138)))
+(let (($x113 (<= ?v1 0)))
+(let ((?x164 (ite $x113 ?x144 ?x29)))
+(let (($x18 (= ?v1 0)))
+(let ((?x169 (ite $x18 ?v0 ?x164)))
+(let ((?x28 (|mod$| ?v0 ?v1)))
+(= ?x28 ?x169))))))))))))
+))
+(let ((?x29 (mod ?1 ?0)))
+(let ((?x95 (* (~ 1) ?0)))
+(let ((?x92 (* (~ 1) ?1)))
+(let ((?x138 (mod ?x92 ?x95)))
+(let ((?x144 (* (~ 1) ?x138)))
+(let (($x113 (<= ?0 0)))
+(let ((?x164 (ite $x113 ?x144 ?x29)))
+(let (($x18 (= ?0 0)))
+(let ((?x169 (ite $x18 ?1 ?x164)))
+(let ((?x28 (|mod$| ?1 ?0)))
+(let (($x172 (= ?x28 ?x169)))
+(let (($x35 (forall ((?v0 Int) (?v1 Int) )(let (($x18 (= ?v1 0)))
+(let ((?x33 (ite $x18 ?v0 (ite (< 0 ?v1) (mod ?v0 ?v1) (- (mod (- ?v0) (- ?v1)))))))
+(let ((?x28 (|mod$| ?v0 ?v1)))
+(= ?x28 ?x33)))))
+))
+(let (($x158 (forall ((?v0 Int) (?v1 Int) )(let ((?x95 (* (~ 1) ?v1)))
+(let ((?x92 (* (~ 1) ?v0)))
+(let ((?x138 (mod ?x92 ?x95)))
+(let ((?x144 (* (~ 1) ?x138)))
+(let ((?x29 (mod ?v0 ?v1)))
+(let (($x19 (< 0 ?v1)))
+(let ((?x149 (ite $x19 ?x29 ?x144)))
+(let (($x18 (= ?v1 0)))
+(let ((?x152 (ite $x18 ?v0 ?x149)))
+(let ((?x28 (|mod$| ?v0 ?v1)))
+(= ?x28 ?x152))))))))))))
+))
+(let ((@x163 (monotonicity (rewrite (= (< 0 ?0) (not $x113))) (= (ite (< 0 ?0) ?x29 ?x144) (ite (not $x113) ?x29 ?x144)))))
+(let ((@x168 (trans @x163 (rewrite (= (ite (not $x113) ?x29 ?x144) ?x164)) (= (ite (< 0 ?0) ?x29 ?x144) ?x164))))
+(let ((@x171 (monotonicity @x168 (= (ite $x18 ?1 (ite (< 0 ?0) ?x29 ?x144)) ?x169))))
+(let ((@x174 (monotonicity @x171 (= (= ?x28 (ite $x18 ?1 (ite (< 0 ?0) ?x29 ?x144))) $x172))))
+(let (($x19 (< 0 ?0)))
+(let ((?x149 (ite $x19 ?x29 ?x144)))
+(let ((?x152 (ite $x18 ?1 ?x149)))
+(let (($x155 (= ?x28 ?x152)))
+(let (($x156 (= (= ?x28 (ite $x18 ?1 (ite $x19 ?x29 (- (mod (- ?1) (- ?0)))))) $x155)))
+(let ((@x140 (monotonicity (rewrite (= (- ?1) ?x92)) (rewrite (= (- ?0) ?x95)) (= (mod (- ?1) (- ?0)) ?x138))))
+(let ((@x148 (trans (monotonicity @x140 (= (- (mod (- ?1) (- ?0))) (- ?x138))) (rewrite (= (- ?x138) ?x144)) (= (- (mod (- ?1) (- ?0))) ?x144))))
+(let ((@x151 (monotonicity @x148 (= (ite $x19 ?x29 (- (mod (- ?1) (- ?0)))) ?x149))))
+(let ((@x154 (monotonicity @x151 (= (ite $x18 ?1 (ite $x19 ?x29 (- (mod (- ?1) (- ?0))))) ?x152))))
+(let ((@x179 (trans (|quant-intro| (monotonicity @x154 $x156) (= $x35 $x158)) (|quant-intro| @x174 (= $x158 $x175)) (= $x35 $x175))))
+(let ((@x190 (|mp~| (mp (asserted $x35) @x179 $x175) (|nnf-pos| (refl (|~| $x172 $x172)) (|~| $x175 $x175)) $x175)))
+(let ((@x204 (mp @x190 (|quant-intro| (refl (= $x172 $x172)) (= $x175 $x199)) $x199)))
+(let (($x253 (or (not $x199) $x247)))
+(let ((?x208 (* (~ 1) 2)))
+(let ((?x207 (* (~ 1) |x$|)))
+(let ((?x209 (mod ?x207 ?x208)))
+(let ((?x210 (* (~ 1) ?x209)))
+(let (($x206 (<= 2 0)))
+(let ((?x212 (ite $x206 ?x210 ?x211)))
+(let (($x205 (= 2 0)))
+(let ((?x213 (ite $x205 |x$| ?x212)))
+(let (($x214 (= ?x7 ?x213)))
+(let ((@x227 (monotonicity (monotonicity (rewrite (= ?x208 (~ 2))) (= ?x209 (mod ?x207 (~ 2)))) (= ?x210 (* (~ 1) (mod ?x207 (~ 2)))))))
+(let ((@x230 (monotonicity (rewrite (= $x206 false)) @x227 (= ?x212 (ite false (* (~ 1) (mod ?x207 (~ 2))) ?x211)))))
+(let ((@x234 (trans @x230 (rewrite (= (ite false (* (~ 1) (mod ?x207 (~ 2))) ?x211) ?x211)) (= ?x212 ?x211))))
+(let ((@x237 (monotonicity (rewrite (= $x205 false)) @x234 (= ?x213 (ite false |x$| ?x211)))))
+(let ((@x244 (monotonicity (trans @x237 (rewrite (= (ite false |x$| ?x211) ?x211)) (= ?x213 ?x211)) (= $x214 (= ?x7 ?x211)))))
+(let ((@x257 (monotonicity (trans @x244 (rewrite (= (= ?x7 ?x211) $x247)) (= $x214 $x247)) (= (or (not $x199) $x214) $x253))))
+(let ((@x260 (trans @x257 (rewrite (= $x253 $x253)) (= (or (not $x199) $x214) $x253))))
+(let ((@x261 (mp ((_ |quant-inst| |x$| 2) (or (not $x199) $x214)) @x260 $x253)))
+(let ((@x323 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x247) $x262)) (|unit-resolution| @x261 @x204 $x247) $x262)))
+(let (($x82 (>= ?x7 2)))
+(let ((?x56 (* 2 ?x7)))
+(let (($x75 (>= ?x56 3)))
+(let (($x65 (< (+ |x$| ?x56) (+ 3 |x$|))))
+(let (($x68 (not $x65)))
+(let ((@x77 (monotonicity (rewrite (= $x65 (not $x75))) (= $x68 (not (not $x75))))))
+(let ((@x86 (trans (trans @x77 (rewrite (= (not (not $x75)) $x75)) (= $x68 $x75)) (rewrite (= $x75 $x82)) (= $x68 $x82))))
+(let ((@x61 (monotonicity (rewrite (= (+ ?x7 ?x7) ?x56)) (= (+ |x$| (+ ?x7 ?x7)) (+ |x$| ?x56)))))
+(let ((@x67 (monotonicity @x61 (rewrite (= (+ |x$| 3) (+ 3 |x$|))) (= (< (+ |x$| (+ ?x7 ?x7)) (+ |x$| 3)) $x65))))
+(let ((@x70 (monotonicity @x67 (= (not (< (+ |x$| (+ ?x7 ?x7)) (+ |x$| 3))) $x68))))
+(let ((@x88 (trans @x70 @x86 (= (not (< (+ |x$| (+ ?x7 ?x7)) (+ |x$| 3))) $x82))))
+(let ((@x89 (mp (asserted (not (< (+ |x$| (+ ?x7 ?x7)) (+ |x$| 3)))) @x88 $x82)))
+((_ |th-lemma| arith farkas -1 1 1) @x89 @x323 (|unit-resolution| ((_ |th-lemma| arith) (or false $x306)) (|true-axiom| true) $x306) false)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))
+
+fb75370f1b646783db5a9a683587f9b2b11bd686 32 0
+unsat
+((set-logic <null>)
+(proof
+(let (($x7 (= |x$| 0.0)))
+(let (($x8 (not $x7)))
+(let ((@x43 (asserted $x8)))
+(let (($x99 (<= |x$| 0.0)))
+(let ((?x45 (* 2.0 |x$|)))
+(let (($x97 (<= ?x45 0.0)))
+(let (($x93 (= ?x45 0.0)))
+(let (($x14 (< 1.0 (ite (< |x$| 0.0) (- |x$|) |x$|))))
+(let (($x16 (or $x14 (not $x14))))
+(let ((?x19 (ite $x16 4.0 2.0)))
+(let (($x23 (not (not (= (+ |x$| |x$|) (* ?x19 |x$|))))))
+(let ((@x88 (rewrite (= (not (not (= ?x45 (* 4.0 |x$|)))) (= ?x45 (* 4.0 |x$|))))))
+(let (($x82 (= (not (= (+ |x$| |x$|) (* ?x19 |x$|))) (not (= ?x45 (* 4.0 |x$|))))))
+(let (($x55 (< 1.0 (ite (< |x$| 0.0) (* (~ 1.0) |x$|) |x$|))))
+(let (($x53 (= (ite (< |x$| 0.0) (- |x$|) |x$|) (ite (< |x$| 0.0) (* (~ 1.0) |x$|) |x$|))))
+(let ((@x57 (monotonicity (monotonicity (rewrite (= (- |x$|) (* (~ 1.0) |x$|))) $x53) (= $x14 $x55))))
+(let ((@x63 (monotonicity @x57 (monotonicity @x57 (= (not $x14) (not $x55))) (= $x16 (or $x55 (not $x55))))))
+(let ((@x67 (trans @x63 (rewrite (= (or $x55 (not $x55)) true)) (= $x16 true))))
+(let ((@x74 (trans (monotonicity @x67 (= ?x19 (ite true 4.0 2.0))) (rewrite (= (ite true 4.0 2.0) 4.0)) (= ?x19 4.0))))
+(let ((@x80 (monotonicity (rewrite (= (+ |x$| |x$|) ?x45)) (monotonicity @x74 (= (* ?x19 |x$|) (* 4.0 |x$|))) (= (= (+ |x$| |x$|) (* ?x19 |x$|)) (= ?x45 (* 4.0 |x$|))))))
+(let ((@x86 (monotonicity (monotonicity @x80 $x82) (= $x23 (not (not (= ?x45 (* 4.0 |x$|))))))))
+(let ((@x95 (trans (trans @x86 @x88 (= $x23 (= ?x45 (* 4.0 |x$|)))) (rewrite (= (= ?x45 (* 4.0 |x$|)) $x93)) (= $x23 $x93))))
+(let ((@x96 (mp (asserted $x23) @x95 $x93)))
+(let ((@x108 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1) (or $x99 (not $x97))) (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x93) $x97)) @x96 $x97) $x99)))
+(let (($x100 (>= |x$| 0.0)))
+(let (($x98 (>= ?x45 0.0)))
+(let ((@x115 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1) (or $x100 (not $x98))) (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x93) $x98)) @x96 $x98) $x100)))
+(|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or $x7 (not $x99) (not $x100))) @x115 @x108 @x43 false))))))))))))))))))))))))))))))
+
+bba1efa8562001b979c24cfd840c5185f0dad8b2 242 0
+unsat
+((set-logic <null>)
+(proof
+(let ((?x471 (div |m$| 2)))
+(let ((?x531 (* (~ 1) ?x471)))
+(let ((?x426 (mod |m$| 2)))
+(let ((?x453 (* (~ 1) ?x426)))
+(let ((?x368 (div |n$| 4)))
+(let ((?x541 (* (~ 2) ?x368)))
+(let ((?x316 (mod |n$| 4)))
+(let ((?x350 (* (~ 1) ?x316)))
+(let ((?x7 (+ |n$| |m$|)))
+(let ((?x259 (div ?x7 2)))
+(let ((?x540 (* (~ 1) ?x259)))
+(let ((?x201 (mod ?x7 2)))
+(let ((?x240 (* (~ 1) ?x201)))
+(let ((?x13 (|mod$| |n$| 4)))
+(let ((?x9 (|mod$| ?x7 2)))
+(let ((?x534 (+ |n$| |m$| ?x9 ?x13 ?x240 ?x540 ?x350 ?x541 ?x453 ?x531)))
+(let (($x535 (>= ?x534 2)))
+(let (($x492 (>= ?x426 0)))
+(let ((@x62 (|true-axiom| true)))
+(let ((?x484 (* (~ 2) ?x471)))
+(let ((?x485 (+ |m$| ?x453 ?x484)))
+(let (($x490 (<= ?x485 0)))
+(let (($x483 (= ?x485 0)))
+(let ((@x379 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x483) $x490)) (|unit-resolution| ((_ |th-lemma| arith) (or false $x483)) @x62 $x483) $x490)))
+(let ((?x351 (+ ?x13 ?x350)))
+(let (($x366 (<= ?x351 0)))
+(let (($x352 (= ?x351 0)))
+(let (($x189 (forall ((?v0 Int) (?v1 Int) )(!(let ((?x37 (mod ?v0 ?v1)))
+(let ((?x71 (* (~ 1) ?v1)))
+(let ((?x68 (* (~ 1) ?v0)))
+(let ((?x114 (mod ?x68 ?x71)))
+(let ((?x120 (* (~ 1) ?x114)))
+(let (($x89 (<= ?v1 0)))
+(let ((?x140 (ite $x89 ?x120 ?x37)))
+(let (($x26 (= ?v1 0)))
+(let ((?x145 (ite $x26 ?v0 ?x140)))
+(let ((?x36 (|mod$| ?v0 ?v1)))
+(= ?x36 ?x145))))))))))) :pattern ( (|mod$| ?v0 ?v1) )))
+))
+(let (($x151 (forall ((?v0 Int) (?v1 Int) )(let ((?x37 (mod ?v0 ?v1)))
+(let ((?x71 (* (~ 1) ?v1)))
+(let ((?x68 (* (~ 1) ?v0)))
+(let ((?x114 (mod ?x68 ?x71)))
+(let ((?x120 (* (~ 1) ?x114)))
+(let (($x89 (<= ?v1 0)))
+(let ((?x140 (ite $x89 ?x120 ?x37)))
+(let (($x26 (= ?v1 0)))
+(let ((?x145 (ite $x26 ?v0 ?x140)))
+(let ((?x36 (|mod$| ?v0 ?v1)))
+(= ?x36 ?x145))))))))))))
+))
+(let ((?x37 (mod ?1 ?0)))
+(let ((?x71 (* (~ 1) ?0)))
+(let ((?x68 (* (~ 1) ?1)))
+(let ((?x114 (mod ?x68 ?x71)))
+(let ((?x120 (* (~ 1) ?x114)))
+(let (($x89 (<= ?0 0)))
+(let ((?x140 (ite $x89 ?x120 ?x37)))
+(let (($x26 (= ?0 0)))
+(let ((?x145 (ite $x26 ?1 ?x140)))
+(let ((?x36 (|mod$| ?1 ?0)))
+(let (($x148 (= ?x36 ?x145)))
+(let (($x43 (forall ((?v0 Int) (?v1 Int) )(let (($x26 (= ?v1 0)))
+(let ((?x41 (ite $x26 ?v0 (ite (< 0 ?v1) (mod ?v0 ?v1) (- (mod (- ?v0) (- ?v1)))))))
+(let ((?x36 (|mod$| ?v0 ?v1)))
+(= ?x36 ?x41)))))
+))
+(let (($x134 (forall ((?v0 Int) (?v1 Int) )(let ((?x71 (* (~ 1) ?v1)))
+(let ((?x68 (* (~ 1) ?v0)))
+(let ((?x114 (mod ?x68 ?x71)))
+(let ((?x120 (* (~ 1) ?x114)))
+(let ((?x37 (mod ?v0 ?v1)))
+(let (($x27 (< 0 ?v1)))
+(let ((?x125 (ite $x27 ?x37 ?x120)))
+(let (($x26 (= ?v1 0)))
+(let ((?x128 (ite $x26 ?v0 ?x125)))
+(let ((?x36 (|mod$| ?v0 ?v1)))
+(= ?x36 ?x128))))))))))))
+))
+(let ((@x139 (monotonicity (rewrite (= (< 0 ?0) (not $x89))) (= (ite (< 0 ?0) ?x37 ?x120) (ite (not $x89) ?x37 ?x120)))))
+(let ((@x144 (trans @x139 (rewrite (= (ite (not $x89) ?x37 ?x120) ?x140)) (= (ite (< 0 ?0) ?x37 ?x120) ?x140))))
+(let ((@x147 (monotonicity @x144 (= (ite $x26 ?1 (ite (< 0 ?0) ?x37 ?x120)) ?x145))))
+(let ((@x150 (monotonicity @x147 (= (= ?x36 (ite $x26 ?1 (ite (< 0 ?0) ?x37 ?x120))) $x148))))
+(let (($x27 (< 0 ?0)))
+(let ((?x125 (ite $x27 ?x37 ?x120)))
+(let ((?x128 (ite $x26 ?1 ?x125)))
+(let (($x131 (= ?x36 ?x128)))
+(let (($x132 (= (= ?x36 (ite $x26 ?1 (ite $x27 ?x37 (- (mod (- ?1) (- ?0)))))) $x131)))
+(let ((@x116 (monotonicity (rewrite (= (- ?1) ?x68)) (rewrite (= (- ?0) ?x71)) (= (mod (- ?1) (- ?0)) ?x114))))
+(let ((@x124 (trans (monotonicity @x116 (= (- (mod (- ?1) (- ?0))) (- ?x114))) (rewrite (= (- ?x114) ?x120)) (= (- (mod (- ?1) (- ?0))) ?x120))))
+(let ((@x127 (monotonicity @x124 (= (ite $x27 ?x37 (- (mod (- ?1) (- ?0)))) ?x125))))
+(let ((@x130 (monotonicity @x127 (= (ite $x26 ?1 (ite $x27 ?x37 (- (mod (- ?1) (- ?0))))) ?x128))))
+(let ((@x155 (trans (|quant-intro| (monotonicity @x130 $x132) (= $x43 $x134)) (|quant-intro| @x150 (= $x134 $x151)) (= $x43 $x151))))
+(let ((@x166 (|mp~| (mp (asserted $x43) @x155 $x151) (|nnf-pos| (refl (|~| $x148 $x148)) (|~| $x151 $x151)) $x151)))
+(let ((@x194 (mp @x166 (|quant-intro| (refl (= $x148 $x148)) (= $x151 $x189)) $x189)))
+(let (($x247 (not $x189)))
+(let (($x357 (or $x247 $x352)))
+(let ((?x317 (ite (<= 4 0) (* (~ 1) (mod (* (~ 1) |n$|) (* (~ 1) 4))) ?x316)))
+(let ((?x318 (ite (= 4 0) |n$| ?x317)))
+(let (($x319 (= ?x13 ?x318)))
+(let ((@x337 (rewrite (= (ite false (* (~ 1) (mod (* (~ 1) |n$|) (~ 4))) ?x316) ?x316))))
+(let (($x331 (= (* (~ 1) (mod (* (~ 1) |n$|) (* (~ 1) 4))) (* (~ 1) (mod (* (~ 1) |n$|) (~ 4))))))
+(let ((@x329 (monotonicity (rewrite (= (* (~ 1) 4) (~ 4))) (= (mod (* (~ 1) |n$|) (* (~ 1) 4)) (mod (* (~ 1) |n$|) (~ 4))))))
+(let ((@x335 (monotonicity (rewrite (= (<= 4 0) false)) (monotonicity @x329 $x331) (= ?x317 (ite false (* (~ 1) (mod (* (~ 1) |n$|) (~ 4))) ?x316)))))
+(let ((@x342 (monotonicity (rewrite (= (= 4 0) false)) (trans @x335 @x337 (= ?x317 ?x316)) (= ?x318 (ite false |n$| ?x316)))))
+(let ((@x349 (monotonicity (trans @x342 (rewrite (= (ite false |n$| ?x316) ?x316)) (= ?x318 ?x316)) (= $x319 (= ?x13 ?x316)))))
+(let ((@x361 (monotonicity (trans @x349 (rewrite (= (= ?x13 ?x316) $x352)) (= $x319 $x352)) (= (or $x247 $x319) $x357))))
+(let ((@x365 (mp ((_ |quant-inst| |n$| 4) (or $x247 $x319)) (trans @x361 (rewrite (= $x357 $x357)) (= (or $x247 $x319) $x357)) $x357)))
+(let ((@x570 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x352) $x366)) (|unit-resolution| @x365 @x194 $x352) $x366)))
+(let ((?x275 (* (~ 2) ?x259)))
+(let ((?x276 (+ |n$| |m$| ?x240 ?x275)))
+(let (($x281 (<= ?x276 0)))
+(let (($x274 (= ?x276 0)))
+(let ((@x560 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x274) $x281)) (|unit-resolution| ((_ |th-lemma| arith) (or false $x274)) @x62 $x274) $x281)))
+(let ((?x241 (+ ?x9 ?x240)))
+(let (($x257 (<= ?x241 0)))
+(let (($x242 (= ?x241 0)))
+(let (($x248 (or $x247 $x242)))
+(let (($x196 (<= 2 0)))
+(let ((?x202 (ite $x196 (* (~ 1) (mod (* (~ 1) ?x7) (* (~ 1) 2))) ?x201)))
+(let (($x195 (= 2 0)))
+(let ((?x203 (ite $x195 ?x7 ?x202)))
+(let (($x204 (= ?x9 ?x203)))
+(let ((?x223 (ite false (* (~ 1) (mod (+ (* (~ 1) |n$|) (* (~ 1) |m$|)) (~ 2))) ?x201)))
+(let (($x221 (= (* (~ 1) (mod (* (~ 1) ?x7) (* (~ 1) 2))) (* (~ 1) (mod (+ (* (~ 1) |n$|) (* (~ 1) |m$|)) (~ 2))))))
+(let (($x218 (= (mod (* (~ 1) ?x7) (* (~ 1) 2)) (mod (+ (* (~ 1) |n$|) (* (~ 1) |m$|)) (~ 2)))))
+(let ((@x216 (rewrite (= (* (~ 1) 2) (~ 2)))))
+(let ((@x219 (monotonicity (rewrite (= (* (~ 1) ?x7) (+ (* (~ 1) |n$|) (* (~ 1) |m$|)))) @x216 $x218)))
+(let ((@x208 (rewrite (= $x196 false))))
+(let ((@x229 (trans (monotonicity @x208 (monotonicity @x219 $x221) (= ?x202 ?x223)) (rewrite (= ?x223 ?x201)) (= ?x202 ?x201))))
+(let ((@x206 (rewrite (= $x195 false))))
+(let ((@x236 (trans (monotonicity @x206 @x229 (= ?x203 (ite false ?x7 ?x201))) (rewrite (= (ite false ?x7 ?x201) ?x201)) (= ?x203 ?x201))))
+(let ((@x246 (trans (monotonicity @x236 (= $x204 (= ?x9 ?x201))) (rewrite (= (= ?x9 ?x201) $x242)) (= $x204 $x242))))
+(let ((@x255 (trans (monotonicity @x246 (= (or $x247 $x204) $x248)) (rewrite (= $x248 $x248)) (= (or $x247 $x204) $x248))))
+(let ((@x256 (mp ((_ |quant-inst| (+ |n$| |m$|) 2) (or $x247 $x204)) @x255 $x248)))
+(let ((@x565 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x242) $x257)) (|unit-resolution| @x256 @x194 $x242) $x257)))
+(let ((?x384 (* (~ 4) ?x368)))
+(let ((?x385 (+ |n$| ?x350 ?x384)))
+(let (($x390 (<= ?x385 0)))
+(let (($x383 (= ?x385 0)))
+(let ((@x577 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x383) $x390)) (|unit-resolution| ((_ |th-lemma| arith) (or false $x383)) @x62 $x383) $x390)))
+(let (($x422 (<= ?x13 3)))
+(let (($x15 (= ?x13 3)))
+(let ((@x64 (asserted $x15)))
+(let ((@x581 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x15) $x422)) @x64 $x422)))
+(let (($x420 (<= ?x9 0)))
+(let (($x11 (= ?x9 0)))
+(let ((@x63 (asserted $x11)))
+(let ((@x553 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x11) $x420)) @x63 $x420)))
+(let ((@x494 ((_ |th-lemma| arith farkas -1 -1 2 -1 -1 -1 -1 -1 1) @x553 @x581 (hypothesis $x535) @x577 @x565 @x560 @x570 @x379 (|unit-resolution| ((_ |th-lemma| arith) (or false $x492)) @x62 $x492) false)))
+(let (($x304 (>= ?x485 0)))
+(let ((@x648 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x483) $x304)) (|unit-resolution| ((_ |th-lemma| arith) (or false $x483)) @x62 $x483) $x304)))
+(let (($x367 (>= ?x351 0)))
+(let ((@x473 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x352) $x367)) (|unit-resolution| @x365 @x194 $x352) $x367)))
+(let (($x421 (>= ?x9 0)))
+(let (($x282 (>= ?x276 0)))
+(let ((@x371 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x274) $x282)) (|unit-resolution| ((_ |th-lemma| arith) (or false $x274)) @x62 $x274) $x282)))
+(let (($x258 (>= ?x241 0)))
+(let ((@x377 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x242) $x258)) (|unit-resolution| @x256 @x194 $x242) $x258)))
+(let (($x391 (>= ?x385 0)))
+(let ((@x474 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x383) $x391)) (|unit-resolution| ((_ |th-lemma| arith) (or false $x383)) @x62 $x383) $x391)))
+(let (($x423 (>= ?x13 3)))
+(let ((@x261 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x15) $x423)) @x64 $x423)))
+(let ((?x19 (|mod$| |m$| 2)))
+(let ((?x454 (+ ?x19 ?x453)))
+(let (($x263 (>= ?x454 0)))
+(let (($x386 (= ?x454 0)))
+(let (($x486 (or $x247 $x386)))
+(let ((?x198 (* (~ 1) 2)))
+(let ((?x210 (* (~ 1) |m$|)))
+(let ((?x424 (mod ?x210 ?x198)))
+(let ((?x425 (* (~ 1) ?x424)))
+(let ((?x427 (ite $x196 ?x425 ?x426)))
+(let ((?x428 (ite $x195 |m$| ?x427)))
+(let (($x429 (= ?x19 ?x428)))
+(let ((@x594 (monotonicity (monotonicity @x216 (= ?x424 (mod ?x210 (~ 2)))) (= ?x425 (* (~ 1) (mod ?x210 (~ 2)))))))
+(let ((@x596 (monotonicity @x208 @x594 (= ?x427 (ite false (* (~ 1) (mod ?x210 (~ 2))) ?x426)))))
+(let ((@x603 (trans @x596 (rewrite (= (ite false (* (~ 1) (mod ?x210 (~ 2))) ?x426) ?x426)) (= ?x427 ?x426))))
+(let ((@x414 (trans (monotonicity @x206 @x603 (= ?x428 (ite false |m$| ?x426))) (rewrite (= (ite false |m$| ?x426) ?x426)) (= ?x428 ?x426))))
+(let ((@x482 (trans (monotonicity @x414 (= $x429 (= ?x19 ?x426))) (rewrite (= (= ?x19 ?x426) $x386)) (= $x429 $x386))))
+(let ((@x511 (trans (monotonicity @x482 (= (or $x247 $x429) $x486)) (rewrite (= $x486 $x486)) (= (or $x247 $x429) $x486))))
+(let ((@x512 (mp ((_ |quant-inst| |m$| 2) (or $x247 $x429)) @x511 $x486)))
+(let ((@x653 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x386) $x263)) (|unit-resolution| @x512 @x194 $x386) $x263)))
+(let (($x271 (>= ?x19 1)))
+(let (($x666 (not $x271)))
+(let (($x509 (<= ?x19 1)))
+(let (($x498 (>= ?x426 2)))
+(let (($x635 (not $x498)))
+(let (($x469 (<= ?x454 0)))
+(let ((@x659 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x386) $x469)) (|unit-resolution| @x512 @x194 $x386) $x469)))
+(let ((@x663 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 1) (or $x509 $x498 (not $x469))) @x659 (|unit-resolution| ((_ |th-lemma| arith) (or false $x635)) @x62 $x635) $x509)))
+(let (($x20 (= ?x19 1)))
+(let (($x168 (not $x20)))
+(let ((?x16 (|mod$| |n$| 2)))
+(let (($x18 (= ?x16 1)))
+(let (($x280 (>= ?x16 1)))
+(let (($x606 (not $x280)))
+(let (($x279 (<= ?x16 1)))
+(let ((?x430 (mod |n$| 2)))
+(let ((?x437 (* (~ 1) ?x430)))
+(let ((?x438 (+ ?x16 ?x437)))
+(let (($x455 (<= ?x438 0)))
+(let (($x439 (= ?x438 0)))
+(let (($x444 (or $x247 $x439)))
+(let ((?x209 (* (~ 1) |n$|)))
+(let ((?x461 (mod ?x209 ?x198)))
+(let ((?x462 (* (~ 1) ?x461)))
+(let ((?x431 (ite $x196 ?x462 ?x430)))
+(let ((?x432 (ite $x195 |n$| ?x431)))
+(let (($x433 (= ?x16 ?x432)))
+(let ((@x522 (monotonicity (monotonicity @x216 (= ?x461 (mod ?x209 (~ 2)))) (= ?x462 (* (~ 1) (mod ?x209 (~ 2)))))))
+(let ((@x521 (monotonicity @x208 @x522 (= ?x431 (ite false (* (~ 1) (mod ?x209 (~ 2))) ?x430)))))
+(let ((@x288 (trans @x521 (rewrite (= (ite false (* (~ 1) (mod ?x209 (~ 2))) ?x430) ?x430)) (= ?x431 ?x430))))
+(let ((@x538 (trans (monotonicity @x206 @x288 (= ?x432 (ite false |n$| ?x430))) (rewrite (= (ite false |n$| ?x430) ?x430)) (= ?x432 ?x430))))
+(let ((@x443 (trans (monotonicity @x538 (= $x433 (= ?x16 ?x430))) (rewrite (= (= ?x16 ?x430) $x439)) (= $x433 $x439))))
+(let ((@x451 (trans (monotonicity @x443 (= (or $x247 $x433) $x444)) (rewrite (= $x444 $x444)) (= (or $x247 $x433) $x444))))
+(let ((@x460 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x439) $x455)) (|unit-resolution| (mp ((_ |quant-inst| |n$| 2) (or $x247 $x433)) @x451 $x444) @x194 $x439) $x455)))
+(let ((@x463 (|unit-resolution| ((_ |th-lemma| arith) (or false (not (>= ?x430 2)))) @x62 (not (>= ?x430 2)))))
+(let ((@x295 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 1) (or $x279 (>= ?x430 2) (not $x455))) @x463 @x460 $x279)))
+(let ((@x292 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or $x18 (not $x279) $x606)) (hypothesis (not $x18)) (or (not $x279) $x606))))
+(let (($x623 (or (not (>= (+ |n$| ?x13 ?x350 ?x541 (* (~ 1) (div |n$| 2))) 2)) $x280)))
+(let ((?x491 (+ |n$| ?x437 (* (~ 2) (div |n$| 2)))))
+(let (($x397 (<= ?x491 0)))
+(let (($x508 (= ?x491 0)))
+(let ((@x614 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x508) $x397)) (|unit-resolution| ((_ |th-lemma| arith) (or false $x508)) @x62 $x508) $x397)))
+(let (($x601 (>= (+ |n$| ?x13 ?x350 ?x541 (* (~ 1) (div |n$| 2))) 2)))
+(let ((@x620 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x439) (>= ?x438 0))) (|unit-resolution| (mp ((_ |quant-inst| |n$| 2) (or $x247 $x433)) @x451 $x444) @x194 $x439) (>= ?x438 0))))
+(let ((@x621 ((_ |th-lemma| arith farkas -1 -2 1 1 1 1 1) @x620 (hypothesis $x601) @x614 @x570 @x577 @x581 (hypothesis $x606) false)))
+(let ((@x403 (|unit-resolution| (lemma @x621 $x623) (|unit-resolution| @x292 @x295 $x606) (not $x601))))
+(let ((@x406 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x508) (>= ?x491 0))) (|unit-resolution| ((_ |th-lemma| arith) (or false $x508)) @x62 $x508) (>= ?x491 0))))
+(let ((@x411 (|unit-resolution| ((_ |th-lemma| arith) (or false (>= ?x430 0))) @x62 (>= ?x430 0))))
+(let (($x169 (or (not $x18) $x168)))
+(let ((@x175 (monotonicity (rewrite (= (and $x18 $x20) (not $x169))) (= (not (and $x18 $x20)) (not (not $x169))))))
+(let ((@x179 (trans @x175 (rewrite (= (not (not $x169)) $x169)) (= (not (and $x18 $x20)) $x169))))
+(let ((@x180 (mp (asserted (not (and $x18 $x20))) @x179 $x169)))
+(let ((@x664 (|unit-resolution| @x180 (lemma ((_ |th-lemma| arith farkas -1/2 -1/2 -1/2 -1/2 -1/2 1) @x411 @x406 @x473 @x474 @x261 @x403 false) $x18) $x168)))
+(let ((@x670 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or $x20 (not $x509) $x666)) @x664 (or (not $x509) $x666))))
+((_ |th-lemma| arith farkas 1/2 -1/2 -1/2 -1/2 -1/2 -1/2 -1/2 -1/2 -1/2 1) (|unit-resolution| @x670 @x663 $x666) @x653 @x261 @x474 @x377 @x371 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x11) $x421)) @x63 $x421) @x473 @x648 (lemma @x494 (not $x535)) false)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))
+
+5e6ffeb79676694a9ab7732936a1e448ef9134cd 12 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let (($x6 (exists ((?v0 Int) )false)
+))
+(let (($x5 (not $x6)))
+(let (($x7 (not $x5)))
+(let ((@x33 (monotonicity (|elim-unused| (= $x6 false)) (= $x5 (not false)))))
+(let ((@x40 (monotonicity (trans @x33 (rewrite (= (not false) true)) (= $x5 true)) (= $x7 (not true)))))
+(let ((@x44 (trans @x40 (rewrite (= (not true) false)) (= $x7 false))))
+(mp (asserted $x7) @x44 false)))))))))
+
+d02a9dcd83dfb0d3e50a887c4f5274a79c10c85e 12 0
+unsat
+((set-logic AUFLIRA)
+(proof
+(let (($x6 (exists ((?v0 Real) )false)
+))
+(let (($x5 (not $x6)))
+(let (($x7 (not $x5)))
+(let ((@x33 (monotonicity (|elim-unused| (= $x6 false)) (= $x5 (not false)))))
+(let ((@x40 (monotonicity (trans @x33 (rewrite (= (not false) true)) (= $x5 true)) (= $x7 (not true)))))
+(let ((@x44 (trans @x40 (rewrite (= (not true) false)) (= $x7 false))))
+(mp (asserted $x7) @x44 false)))))))))
+
+1a820d07cf476448545d144873b309b9cfc3a238 2 0
+unknown
+(error "line 6 column 10: proof is not available")
+23ed6364c527ef515dd659de8b496cfd59df4ec7 2 0
+unknown
+(error "line 6 column 10: proof is not available")
+9722ff2e938783a69202c957c858fb219ec0cdb0 2 0
+unknown
+(error "line 6 column 10: proof is not available")
+cb87115705dc568881932b35aa82751f3f97049c 22 0
+unsat
+((set-logic AUFLIA)
+(declare-fun ?v1!0 () Int)
+(declare-fun ?v0!1 () Int)
+(proof
+(let (($x51 (= ?v1!0 1)))
+(let (($x57 (not (or (not (and (= ?v0!1 0) $x51)) (not (= ?v0!1 ?v1!0))))))
+(let (($x41 (forall ((?v0 Int) (?v1 Int) )(or (not (and (= ?v0 0) (= ?v1 1))) (not (= ?v0 ?v1))))
+))
+(let (($x44 (not $x41)))
+(let (($x15 (forall ((?v0 Int) (?v1 Int) )(=> (and (= ?v0 0) (= ?v1 1)) (not (= ?v0 ?v1))))
+))
+(let (($x16 (not $x15)))
+(let (($x39 (= (=> (and (= ?1 0) (= ?0 1)) (not (= ?1 ?0))) (or (not (and (= ?1 0) (= ?0 1))) (not (= ?1 ?0))))))
+(let ((@x46 (monotonicity (|quant-intro| (rewrite $x39) (= $x15 $x41)) (= $x16 $x44))))
+(let ((@x63 (|not-or-elim| (|mp~| (mp (asserted $x16) @x46 $x44) (sk (|~| $x44 $x57)) $x57) (and (= ?v0!1 0) $x51))))
+(let ((@x65 (|and-elim| @x63 $x51)))
+(let (($x54 (= ?v0!1 ?v1!0)))
+(let ((@x66 (|not-or-elim| (|mp~| (mp (asserted $x16) @x46 $x44) (sk (|~| $x44 $x57)) $x57) $x54)))
+(let ((@x68 (trans (symm (|and-elim| @x63 (= ?v0!1 0)) (= 0 ?v0!1)) @x66 (= 0 ?v1!0))))
+(mp (trans @x68 @x65 (= 0 1)) (rewrite (= (= 0 1) false)) false))))))))))))))))
+
+d2d0a7794c4de3708d5541374fd9e0075ad5fa36 55 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let (($x14 (exists ((?v0 Int) )(forall ((?v1 Int) )(let (($x10 (<= 0 ?v1)))
+(let (($x9 (< ?v1 0)))
+(let (($x11 (or $x9 $x10)))
+(let (($x7 (< ?v0 ?v1)))
+(=> $x7 $x11))))))
+)
+))
+(let (($x15 (not $x14)))
+(let (($x43 (exists ((?v0 Int) )(forall ((?v1 Int) )(let (($x10 (<= 0 ?v1)))
+(let (($x9 (< ?v1 0)))
+(let (($x11 (or $x9 $x10)))
+(let (($x7 (< ?v0 ?v1)))
+(let (($x36 (not $x7)))
+(or $x36 $x11)))))))
+)
+))
+(let (($x46 (not $x43)))
+(let (($x86 (exists ((?v0 Int) )true)
+))
+(let (($x40 (forall ((?v1 Int) )(let (($x10 (<= 0 ?v1)))
+(let (($x9 (< ?v1 0)))
+(let (($x11 (or $x9 $x10)))
+(let (($x7 (< ?0 ?v1)))
+(let (($x36 (not $x7)))
+(or $x36 $x11)))))))
+))
+(let (($x79 (forall ((?v1 Int) )true)
+))
+(let (($x10 (<= 0 ?0)))
+(let (($x9 (< ?0 0)))
+(let (($x11 (or $x9 $x10)))
+(let (($x7 (< ?1 ?0)))
+(let (($x36 (not $x7)))
+(let (($x37 (or $x36 $x11)))
+(let (($x58 (<= (+ ?0 (* (~ 1) ?1)) 0)))
+(let ((@x76 (rewrite (= (or $x58 (or (not (>= ?0 0)) (>= ?0 0))) true))))
+(let ((@x71 (monotonicity (rewrite (= $x9 (not (>= ?0 0)))) (rewrite (= $x10 (>= ?0 0))) (= $x11 (or (not (>= ?0 0)) (>= ?0 0))))))
+(let ((@x64 (monotonicity (rewrite (= $x7 (not $x58))) (= $x36 (not (not $x58))))))
+(let ((@x74 (monotonicity (trans @x64 (rewrite (= (not (not $x58)) $x58)) (= $x36 $x58)) @x71 (= $x37 (or $x58 (or (not (>= ?0 0)) (>= ?0 0)))))))
+(let ((@x85 (trans (|quant-intro| (trans @x74 @x76 (= $x37 true)) (= $x40 $x79)) (|elim-unused| (= $x79 true)) (= $x40 true))))
+(let ((@x92 (trans (|quant-intro| @x85 (= $x43 $x86)) (|elim-unused| (= $x86 true)) (= $x43 true))))
+(let ((@x99 (trans (monotonicity @x92 (= $x46 (not true))) (rewrite (= (not true) false)) (= $x46 false))))
+(let (($x13 (forall ((?v1 Int) )(let (($x10 (<= 0 ?v1)))
+(let (($x9 (< ?v1 0)))
+(let (($x11 (or $x9 $x10)))
+(let (($x7 (< ?0 ?v1)))
+(=> $x7 $x11))))))
+))
+(let ((@x45 (|quant-intro| (|quant-intro| (rewrite (= (=> $x7 $x11) $x37)) (= $x13 $x40)) (= $x14 $x43))))
+(let ((@x48 (monotonicity @x45 (= $x15 $x46))))
+(mp (asserted $x15) (trans @x48 @x99 (= $x15 false)) false)))))))))))))))))))))))))))
+
+75e4df9c205f7c15d7e530b5e1e97635aed16d82 42 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let (($x15 (forall ((?v0 Int) (?v1 Int) )(let ((?x12 (* 2 ?v1)))
+(let ((?x9 (* 2 ?v0)))
+(let ((?x11 (+ ?x9 1)))
+(let (($x13 (< ?x11 ?x12)))
+(let (($x7 (< ?v0 ?v1)))
+(=> $x7 $x13)))))))
+))
+(let (($x16 (not $x15)))
+(let (($x53 (forall ((?v0 Int) (?v1 Int) )(let ((?x12 (* 2 ?v1)))
+(let ((?x9 (* 2 ?v0)))
+(let ((?x38 (+ 1 ?x9)))
+(let (($x41 (< ?x38 ?x12)))
+(let (($x7 (< ?v0 ?v1)))
+(let (($x47 (not $x7)))
+(or $x47 $x41))))))))
+))
+(let (($x56 (not $x53)))
+(let (($x82 (forall ((?v0 Int) (?v1 Int) )true)
+))
+(let ((?x12 (* 2 ?0)))
+(let ((?x9 (* 2 ?1)))
+(let ((?x38 (+ 1 ?x9)))
+(let (($x41 (< ?x38 ?x12)))
+(let (($x7 (< ?1 ?0)))
+(let (($x47 (not $x7)))
+(let (($x48 (or $x47 $x41)))
+(let (($x61 (>= (+ ?1 (* (~ 1) ?0)) 0)))
+(let (($x60 (not $x61)))
+(let ((@x72 (trans (monotonicity (rewrite (= $x7 $x60)) (= $x47 (not $x60))) (rewrite (= (not $x60) $x61)) (= $x47 $x61))))
+(let ((@x77 (monotonicity @x72 (rewrite (= $x41 $x60)) (= $x48 (or $x61 $x60)))))
+(let ((@x84 (|quant-intro| (trans @x77 (rewrite (= (or $x61 $x60) true)) (= $x48 true)) (= $x53 $x82))))
+(let ((@x91 (monotonicity (trans @x84 (|elim-unused| (= $x82 true)) (= $x53 true)) (= $x56 (not true)))))
+(let ((@x95 (trans @x91 (rewrite (= (not true) false)) (= $x56 false))))
+(let ((@x43 (monotonicity (rewrite (= (+ ?x9 1) ?x38)) (= (< (+ ?x9 1) ?x12) $x41))))
+(let ((@x46 (monotonicity @x43 (= (=> $x7 (< (+ ?x9 1) ?x12)) (=> $x7 $x41)))))
+(let ((@x52 (trans @x46 (rewrite (= (=> $x7 $x41) $x48)) (= (=> $x7 (< (+ ?x9 1) ?x12)) $x48))))
+(let ((@x58 (monotonicity (|quant-intro| @x52 (= $x15 $x53)) (= $x16 $x56))))
+(mp (asserted $x16) (trans @x58 @x95 (= $x16 false)) false))))))))))))))))))))))))))
+
+591a3beb5d84c4e2d6724bf947c2fd4fa44c6bbc 32 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let (($x14 (forall ((?v0 Int) (?v1 Int) )(let ((?x11 (* 2 ?v1)))
+(let ((?x8 (* 2 ?v0)))
+(let ((?x10 (+ ?x8 1)))
+(let (($x12 (= ?x10 ?x11)))
+(not $x12))))))
+))
+(let (($x15 (not $x14)))
+(let (($x46 (forall ((?v0 Int) (?v1 Int) )(let ((?x11 (* 2 ?v1)))
+(let ((?x8 (* 2 ?v0)))
+(let ((?x37 (+ 1 ?x8)))
+(let (($x40 (= ?x37 ?x11)))
+(not $x40))))))
+))
+(let (($x49 (not $x46)))
+(let (($x61 (forall ((?v0 Int) (?v1 Int) )true)
+))
+(let ((?x11 (* 2 ?0)))
+(let ((?x8 (* 2 ?1)))
+(let ((?x37 (+ 1 ?x8)))
+(let (($x40 (= ?x37 ?x11)))
+(let (($x43 (not $x40)))
+(let ((@x60 (trans (monotonicity (rewrite (= $x40 false)) (= $x43 (not false))) (rewrite (= (not false) true)) (= $x43 true))))
+(let ((@x67 (trans (|quant-intro| @x60 (= $x46 $x61)) (|elim-unused| (= $x61 true)) (= $x46 true))))
+(let ((@x74 (trans (monotonicity @x67 (= $x49 (not true))) (rewrite (= (not true) false)) (= $x49 false))))
+(let ((@x42 (monotonicity (rewrite (= (+ ?x8 1) ?x37)) (= (= (+ ?x8 1) ?x11) $x40))))
+(let ((@x48 (|quant-intro| (monotonicity @x42 (= (not (= (+ ?x8 1) ?x11)) $x43)) (= $x14 $x46))))
+(let ((@x51 (monotonicity @x48 (= $x15 $x49))))
+(mp (asserted $x15) (trans @x51 @x74 (= $x15 false)) false)))))))))))))))))))
+
+ed6b5d78e5a12fb3a6471e02bc6e89ebc78c1a34 43 0
+unsat
+((set-logic AUFLIA)
+(declare-fun ?v0!1 () Int)
+(declare-fun ?v1!0 () Int)
+(proof
+(let ((?x78 (+ ?v1!0 ?v0!1)))
+(let (($x90 (>= ?x78 2)))
+(let (($x93 (not $x90)))
+(let (($x87 (= ?x78 2)))
+(let (($x81 (<= ?x78 2)))
+(let (($x84 (not $x81)))
+(let (($x71 (or (not (<= (+ ?v0!1 ?v1!0) 2)) (= (+ ?v0!1 ?v1!0) 2) (not (>= (+ ?v0!1 ?v1!0) 2)))))
+(let (($x72 (not $x71)))
+(let ((@x80 (rewrite (= (+ ?v0!1 ?v1!0) ?x78))))
+(let ((@x95 (monotonicity (monotonicity @x80 (= (>= (+ ?v0!1 ?v1!0) 2) $x90)) (= (not (>= (+ ?v0!1 ?v1!0) 2)) $x93))))
+(let ((@x86 (monotonicity (monotonicity @x80 (= (<= (+ ?v0!1 ?v1!0) 2) $x81)) (= (not (<= (+ ?v0!1 ?v1!0) 2)) $x84))))
+(let ((@x98 (monotonicity @x86 (monotonicity @x80 (= (= (+ ?v0!1 ?v1!0) 2) $x87)) @x95 (= $x71 (or $x84 $x87 $x93)))))
+(let (($x58 (forall ((?v0 Int) (?v1 Int) )(let (($x39 (not (>= (+ ?v0 ?v1) 2))))
+(let ((?x8 (+ ?v0 ?v1)))
+(let (($x10 (= ?x8 2)))
+(let (($x44 (not (<= ?x8 2))))
+(or $x44 $x10 $x39))))))
+))
+(let (($x61 (not $x58)))
+(let (($x14 (forall ((?v0 Int) (?v1 Int) )(or (< 2 (+ ?v0 ?v1)) (or (= (+ ?v0 ?v1) 2) (< (+ ?v0 ?v1) 2))))
+))
+(let (($x15 (not $x14)))
+(let (($x39 (not (>= (+ ?1 ?0) 2))))
+(let ((?x8 (+ ?1 ?0)))
+(let (($x10 (= ?x8 2)))
+(let (($x44 (not (<= ?x8 2))))
+(let (($x53 (or $x44 $x10 $x39)))
+(let (($x13 (or (< 2 ?x8) (or $x10 (< ?x8 2)))))
+(let ((@x49 (monotonicity (rewrite (= (< ?x8 2) $x39)) (= (or $x10 (< ?x8 2)) (or $x10 $x39)))))
+(let ((@x52 (monotonicity (rewrite (= (< 2 ?x8) $x44)) @x49 (= $x13 (or $x44 (or $x10 $x39))))))
+(let ((@x57 (trans @x52 (rewrite (= (or $x44 (or $x10 $x39)) $x53)) (= $x13 $x53))))
+(let ((@x64 (mp (asserted $x15) (monotonicity (|quant-intro| @x57 (= $x14 $x58)) (= $x15 $x61)) $x61)))
+(let ((@x76 (mp (|mp~| @x64 (sk (|~| $x61 $x72)) $x72) (monotonicity @x98 (= $x72 (not (or $x84 $x87 $x93)))) (not (or $x84 $x87 $x93)))))
+(let ((@x103 (|not-or-elim| @x76 (not $x87))))
+(let ((@x104 (|not-or-elim| @x76 $x90)))
+(let ((@x77 (|not-or-elim| @x76 $x81)))
+(|unit-resolution| (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or $x87 $x84 $x93)) @x77 (or $x87 $x93)) @x104 @x103 false)))))))))))))))))))))))))))))))))
+
+092a8c0984dc61be1fab786d699cb79093b7c5f2 46 0
+unsat
+((set-logic AUFLIA)
+(declare-fun ?v0!0 () Int)
+(proof
+(let (($x81 (<= ?v0!0 0)))
+(let (($x84 (<= ?v0!0 (~ 1))))
+(let (($x85 (not $x84)))
+(let ((@x103 (|unit-resolution| ((_ |th-lemma| arith farkas 1 1) (or $x85 $x81)) (hypothesis (not $x81)) $x85)))
+(let (($x82 (>= ?v0!0 1)))
+(let (($x83 (not $x82)))
+(let (($x86 (ite $x81 $x83 $x85)))
+(let (($x87 (not $x86)))
+(let (($x71 (forall ((?v0 Int) )(let (($x56 (not (<= ?v0 (~ 1)))))
+(let (($x59 (not (>= ?v0 1))))
+(ite (<= ?v0 0) $x59 $x56))))
+))
+(let (($x74 (not $x71)))
+(let (($x13 (forall ((?v0 Int) )(let (($x11 (< ?v0 1)))
+(let (($x7 (< 0 ?v0)))
+(ite $x7 (< 0 (+ ?v0 1)) $x11))))
+))
+(let (($x14 (not $x13)))
+(let (($x44 (forall ((?v0 Int) )(let (($x11 (< ?v0 1)))
+(let (($x38 (< 0 (+ 1 ?v0))))
+(let (($x7 (< 0 ?v0)))
+(ite $x7 $x38 $x11)))))
+))
+(let (($x56 (not (<= ?0 (~ 1)))))
+(let (($x59 (not (>= ?0 1))))
+(let (($x66 (ite (<= ?0 0) $x59 $x56)))
+(let (($x11 (< ?0 1)))
+(let (($x38 (< 0 (+ 1 ?0))))
+(let (($x7 (< 0 ?0)))
+(let (($x41 (ite $x7 $x38 $x11)))
+(let ((@x65 (monotonicity (rewrite (= $x7 (not (<= ?0 0)))) (rewrite (= $x38 $x56)) (rewrite (= $x11 $x59)) (= $x41 (ite (not (<= ?0 0)) $x56 $x59)))))
+(let ((@x70 (trans @x65 (rewrite (= (ite (not (<= ?0 0)) $x56 $x59) $x66)) (= $x41 $x66))))
+(let ((@x76 (monotonicity (|quant-intro| @x70 (= $x44 $x71)) (= (not $x44) $x74))))
+(let ((@x40 (monotonicity (rewrite (= (+ ?0 1) (+ 1 ?0))) (= (< 0 (+ ?0 1)) $x38))))
+(let ((@x43 (monotonicity @x40 (= (ite $x7 (< 0 (+ ?0 1)) $x11) $x41))))
+(let ((@x49 (monotonicity (|quant-intro| @x43 (= $x13 $x44)) (= $x14 (not $x44)))))
+(let ((@x90 (|mp~| (mp (asserted $x14) (trans @x49 @x76 (= $x14 $x74)) $x74) (sk (|~| $x74 $x87)) $x87)))
+(let ((@x106 (|unit-resolution| (|unit-resolution| (|def-axiom| (or $x86 $x81 $x84)) @x90 (or $x81 $x84)) @x103 (hypothesis (not $x81)) false)))
+(let ((@x107 (lemma @x106 $x81)))
+(let ((@x112 (|unit-resolution| (|def-axiom| (or $x86 (not $x81) $x82)) @x90 (or (not $x81) $x82))))
+(|unit-resolution| @x112 (|unit-resolution| ((_ |th-lemma| arith farkas 1 1) (or $x83 (not $x81))) @x107 $x83) @x107 false)))))))))))))))))))))))))))))))))
+
+a5ac1b17c244dd3aef6c7c1289500bca02b482d0 2 0
+unknown
+(error "line 6 column 10: proof is not available")
+f697308f77e23a8625c050b2aa7a7f131faf390d 2 0
+unknown
+(error "line 6 column 10: proof is not available")
+ad02a5379962c0c41aaf5c95191947c03228f5e6 39 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let (($x16 (exists ((?v0 Int) (?v1 Int) (?v2 Int) )(let ((?x11 (- 6)))
+(let ((?x12 (* ?x11 ?v1)))
+(let ((?x9 (* 4 ?v0)))
+(let ((?x13 (+ ?x9 ?x12)))
+(= ?x13 1))))))
+))
+(let (($x7 (not $x16)))
+(let (($x17 (not $x7)))
+(let (($x59 (exists ((?v0 Int) (?v1 Int) )(let ((?x56 (* (~ 6) ?v1)))
+(let ((?x55 (* 4 ?v0)))
+(let ((?x57 (+ ?x55 ?x56)))
+(= ?x57 1)))))
+))
+(let (($x75 (exists ((?v0 Int) (?v1 Int) )false)
+))
+(let ((@x79 (|quant-intro| (rewrite (= (= (+ (* 4 ?1) (* (~ 6) ?0)) 1) false)) (= $x59 $x75))))
+(let ((@x83 (trans @x79 (|elim-unused| (= $x75 false)) (= $x59 false))))
+(let (($x51 (exists ((?v0 Int) (?v1 Int) (?v2 Int) )(let ((?x42 (* (~ 6) ?v1)))
+(let ((?x9 (* 4 ?v0)))
+(let ((?x45 (+ ?x9 ?x42)))
+(= ?x45 1)))))
+))
+(let ((?x42 (* (~ 6) ?1)))
+(let ((?x9 (* 4 ?2)))
+(let ((?x45 (+ ?x9 ?x42)))
+(let (($x48 (= ?x45 1)))
+(let ((?x11 (- 6)))
+(let ((?x12 (* ?x11 ?1)))
+(let ((?x13 (+ ?x9 ?x12)))
+(let (($x15 (= ?x13 1)))
+(let ((@x47 (monotonicity (monotonicity (rewrite (= ?x11 (~ 6))) (= ?x12 ?x42)) (= ?x13 ?x45))))
+(let ((@x63 (trans (|quant-intro| (monotonicity @x47 (= $x15 $x48)) (= $x16 $x51)) (|elim-unused| (= $x51 $x59)) (= $x16 $x59))))
+(let ((@x69 (monotonicity (monotonicity @x63 (= $x7 (not $x59))) (= $x17 (not (not $x59))))))
+(let ((@x73 (trans @x69 (rewrite (= (not (not $x59)) $x59)) (= $x17 $x59))))
+(mp (asserted $x17) (trans @x73 @x83 (= $x17 false)) false)))))))))))))))))))))))
+
+8a423de6b668d07fe5d90dcad74d7fbb1fcb9c11 52 0
+unsat
+((set-logic AUFLIA)
+(declare-fun ?v1!1 () Int)
+(declare-fun ?v2!0 () Int)
+(proof
+(let ((?x103 (+ ?v2!0 ?v1!1)))
+(let (($x104 (<= ?x103 0)))
+(let (($x106 (or (not (and (not (<= ?v1!1 0)) (not (<= ?v2!0 0)))) (not $x104))))
+(let (($x86 (forall ((?v1 Int) (?v2 Int) )(or (not (and (not (<= ?v1 0)) (not (<= ?v2 0)))) (not (<= (+ ?v2 ?v1) 0))))
+))
+(let (($x89 (not $x86)))
+(let (($x15 (exists ((?v0 Int) )(forall ((?v1 Int) (?v2 Int) )(let (($x10 (and (< 0 ?v1) (< 0 ?v2))))
+(=> $x10 (< 0 (+ ?v1 ?v2)))))
+)
+))
+(let (($x16 (not $x15)))
+(let (($x52 (forall ((?v1 Int) (?v2 Int) )(let ((?x37 (+ ?v2 ?v1)))
+(let (($x40 (< 0 ?x37)))
+(or (not (and (< 0 ?v1) (< 0 ?v2))) $x40))))
+))
+(let (($x83 (or (not (and (not (<= ?1 0)) (not (<= ?0 0)))) (not (<= (+ ?0 ?1) 0)))))
+(let ((?x37 (+ ?0 ?1)))
+(let (($x40 (< 0 ?x37)))
+(let (($x47 (or (not (and (< 0 ?1) (< 0 ?0))) $x40)))
+(let (($x77 (= (not (and (< 0 ?1) (< 0 ?0))) (not (and (not (<= ?1 0)) (not (<= ?0 0)))))))
+(let (($x10 (and (< 0 ?1) (< 0 ?0))))
+(let ((@x75 (monotonicity (rewrite (= (< 0 ?1) (not (<= ?1 0)))) (rewrite (= (< 0 ?0) (not (<= ?0 0)))) (= $x10 (and (not (<= ?1 0)) (not (<= ?0 0)))))))
+(let ((@x85 (monotonicity (monotonicity @x75 $x77) (rewrite (= $x40 (not (<= ?x37 0)))) (= $x47 $x83))))
+(let ((@x91 (monotonicity (|quant-intro| @x85 (= $x52 $x86)) (= (not $x52) $x89))))
+(let (($x55 (exists ((?v0 Int) )(forall ((?v1 Int) (?v2 Int) )(let ((?x37 (+ ?v2 ?v1)))
+(let (($x40 (< 0 ?x37)))
+(or (not (and (< 0 ?v1) (< 0 ?v2))) $x40))))
+)
+))
+(let (($x14 (forall ((?v1 Int) (?v2 Int) )(let (($x10 (and (< 0 ?v1) (< 0 ?v2))))
+(=> $x10 (< 0 (+ ?v1 ?v2)))))
+))
+(let ((@x42 (monotonicity (rewrite (= (+ ?1 ?0) ?x37)) (= (< 0 (+ ?1 ?0)) $x40))))
+(let ((@x45 (monotonicity @x42 (= (=> $x10 (< 0 (+ ?1 ?0))) (=> $x10 $x40)))))
+(let ((@x51 (trans @x45 (rewrite (= (=> $x10 $x40) $x47)) (= (=> $x10 (< 0 (+ ?1 ?0))) $x47))))
+(let ((@x61 (trans (|quant-intro| (|quant-intro| @x51 (= $x14 $x52)) (= $x15 $x55)) (|elim-unused| (= $x55 $x52)) (= $x15 $x52))))
+(let ((@x93 (trans (monotonicity @x61 (= $x16 (not $x52))) @x91 (= $x16 $x89))))
+(let ((@x110 (|mp~| (mp (asserted $x16) @x93 $x89) (sk (|~| $x89 (not $x106))) (not $x106))))
+(let ((@x116 (|not-or-elim| @x110 $x104)))
+(let (($x97 (<= ?v1!1 0)))
+(let (($x98 (not $x97)))
+(let ((@x114 (|and-elim| (|not-or-elim| @x110 (and $x98 (not (<= ?v2!0 0)))) $x98)))
+(let (($x99 (<= ?v2!0 0)))
+(let (($x100 (not $x99)))
+(let ((@x115 (|and-elim| (|not-or-elim| @x110 (and $x98 $x100)) $x100)))
+((_ |th-lemma| arith farkas 1 1 1) @x115 @x114 @x116 false)))))))))))))))))))))))))))))))))))
+
+7c93c190dc21779c8214786ce8c1fd4de433814f 46 0
+unsat
+((set-logic AUFLIRA)
+(declare-fun ?v1!1 () Int)
+(declare-fun ?v2!0 () Real)
+(proof
+(let (($x103 (<= ?v1!1 (~ 1))))
+(let (($x104 (not $x103)))
+(let (($x105 (or (not (and (not (<= ?v1!1 0)) (not (<= ?v2!0 0.0)))) $x104)))
+(let (($x86 (forall ((?v1 Int) (?v2 Real) )(or (not (and (not (<= ?v1 0)) (not (<= ?v2 0.0)))) (not (<= ?v1 (~ 1)))))
+))
+(let (($x89 (not $x86)))
+(let (($x18 (exists ((?v0 Int) )(forall ((?v1 Int) (?v2 Real) )(let (($x12 (and (< 0 ?v1) (< 0.0 ?v2))))
+(=> $x12 (< (- 1) ?v1))))
+)
+))
+(let (($x5 (not $x18)))
+(let (($x52 (forall ((?v1 Int) (?v2 Real) )(let (($x40 (< (~ 1) ?v1)))
+(or (not (and (< 0 ?v1) (< 0.0 ?v2))) $x40)))
+))
+(let (($x83 (or (not (and (not (<= ?1 0)) (not (<= ?0 0.0)))) (not (<= ?1 (~ 1))))))
+(let (($x40 (< (~ 1) ?1)))
+(let (($x47 (or (not (and (< 0 ?1) (< 0.0 ?0))) $x40)))
+(let (($x77 (= (not (and (< 0 ?1) (< 0.0 ?0))) (not (and (not (<= ?1 0)) (not (<= ?0 0.0)))))))
+(let (($x12 (and (< 0 ?1) (< 0.0 ?0))))
+(let ((@x75 (monotonicity (rewrite (= (< 0 ?1) (not (<= ?1 0)))) (rewrite (= (< 0.0 ?0) (not (<= ?0 0.0)))) (= $x12 (and (not (<= ?1 0)) (not (<= ?0 0.0)))))))
+(let ((@x85 (monotonicity (monotonicity @x75 $x77) (rewrite (= $x40 (not (<= ?1 (~ 1))))) (= $x47 $x83))))
+(let ((@x91 (monotonicity (|quant-intro| @x85 (= $x52 $x86)) (= (not $x52) $x89))))
+(let (($x55 (exists ((?v0 Int) )(forall ((?v1 Int) (?v2 Real) )(let (($x40 (< (~ 1) ?v1)))
+(or (not (and (< 0 ?v1) (< 0.0 ?v2))) $x40)))
+)
+))
+(let (($x17 (forall ((?v1 Int) (?v2 Real) )(let (($x12 (and (< 0 ?v1) (< 0.0 ?v2))))
+(=> $x12 (< (- 1) ?v1))))
+))
+(let ((@x42 (monotonicity (rewrite (= (- 1) (~ 1))) (= (< (- 1) ?1) $x40))))
+(let ((@x45 (monotonicity @x42 (= (=> $x12 (< (- 1) ?1)) (=> $x12 $x40)))))
+(let ((@x51 (trans @x45 (rewrite (= (=> $x12 $x40) $x47)) (= (=> $x12 (< (- 1) ?1)) $x47))))
+(let ((@x61 (trans (|quant-intro| (|quant-intro| @x51 (= $x17 $x52)) (= $x18 $x55)) (|elim-unused| (= $x55 $x52)) (= $x18 $x52))))
+(let ((@x93 (trans (monotonicity @x61 (= $x5 (not $x52))) @x91 (= $x5 $x89))))
+(let ((@x109 (|mp~| (mp (asserted $x5) @x93 $x89) (sk (|~| $x89 (not $x105))) (not $x105))))
+(let ((@x115 (|not-or-elim| @x109 $x103)))
+(let (($x97 (<= ?v1!1 0)))
+(let (($x98 (not $x97)))
+(let ((@x113 (|and-elim| (|not-or-elim| @x109 (and $x98 (not (<= ?v2!0 0.0)))) $x98)))
+(|unit-resolution| (|unit-resolution| ((_ |th-lemma| arith farkas 1 1) (or $x104 $x97)) @x113 $x104) @x115 false)))))))))))))))))))))))))))))))
+
+39a3c02f6687608102bd092d376b8901575e5356 2 0
+unknown
+(error "line 6 column 10: proof is not available")
+3bdf1da3a49c7c0ce726c85bf6e844aabdd6afa0 36 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let (($x13 (forall ((?v0 Int) )(let ((?x10 (* 2 |a$|)))
+(let ((?x9 (* 2 ?v0)))
+(let (($x11 (< ?x9 ?x10)))
+(let (($x7 (< ?v0 |a$|)))
+(=> $x7 $x11))))))
+))
+(let (($x14 (not $x13)))
+(let (($x40 (forall ((?v0 Int) )(let ((?x10 (* 2 |a$|)))
+(let ((?x9 (* 2 ?v0)))
+(let (($x11 (< ?x9 ?x10)))
+(let (($x7 (< ?v0 |a$|)))
+(let (($x36 (not $x7)))
+(or $x36 $x11)))))))
+))
+(let (($x43 (not $x40)))
+(let (($x69 (forall ((?v0 Int) )true)
+))
+(let ((?x10 (* 2 |a$|)))
+(let ((?x9 (* 2 ?0)))
+(let (($x11 (< ?x9 ?x10)))
+(let (($x7 (< ?0 |a$|)))
+(let (($x36 (not $x7)))
+(let (($x37 (or $x36 $x11)))
+(let (($x48 (>= (+ ?0 (* (~ 1) |a$|)) 0)))
+(let (($x47 (not $x48)))
+(let ((@x59 (trans (monotonicity (rewrite (= $x7 $x47)) (= $x36 (not $x47))) (rewrite (= (not $x47) $x48)) (= $x36 $x48))))
+(let ((@x64 (monotonicity @x59 (rewrite (= $x11 $x47)) (= $x37 (or $x48 $x47)))))
+(let ((@x71 (|quant-intro| (trans @x64 (rewrite (= (or $x48 $x47) true)) (= $x37 true)) (= $x40 $x69))))
+(let ((@x78 (monotonicity (trans @x71 (|elim-unused| (= $x69 true)) (= $x40 true)) (= $x43 (not true)))))
+(let ((@x82 (trans @x78 (rewrite (= (not true) false)) (= $x43 false))))
+(let ((@x45 (monotonicity (|quant-intro| (rewrite (= (=> $x7 $x11) $x37)) (= $x13 $x40)) (= $x14 $x43))))
+(mp (asserted $x14) (trans @x45 @x82 (= $x14 false)) false))))))))))))))))))))))
+
+0737ab0e9619ba68da155fd5dcce2691243e7d8d 24 0
+unsat
+((set-logic AUFLIA)
+(declare-fun ?v1!0 () Int)
+(proof
+(let (($x62 (>= ?v1!0 1)))
+(let (($x50 (forall ((?v1 Int) )(or (not (<= ?v1 0)) (not (>= ?v1 1))))
+))
+(let (($x53 (not $x50)))
+(let (($x12 (forall ((?v0 Int) (?v1 Int) )(or (< 0 ?v1) (< ?v1 1)))
+))
+(let (($x5 (not $x12)))
+(let (($x33 (forall ((?v1 Int) )(or (< 0 ?v1) (< ?v1 1)))
+))
+(let (($x11 (or (< 0 ?0) (< ?0 1))))
+(let ((@x49 (monotonicity (rewrite (= (< 0 ?0) (not (<= ?0 0)))) (rewrite (= (< ?0 1) (not (>= ?0 1)))) (= $x11 (or (not (<= ?0 0)) (not (>= ?0 1)))))))
+(let ((@x55 (monotonicity (|quant-intro| @x49 (= $x33 $x50)) (= (not $x33) $x53))))
+(let ((@x57 (trans (monotonicity (|elim-unused| (= $x12 $x33)) (= $x5 (not $x33))) @x55 (= $x5 $x53))))
+(let ((@x68 (|mp~| (mp (asserted $x5) @x57 $x53) (sk (|~| $x53 (not (or (not (<= ?v1!0 0)) (not $x62))))) (not (or (not (<= ?v1!0 0)) (not $x62))))))
+(let ((@x72 (|not-or-elim| @x68 $x62)))
+(let (($x63 (not $x62)))
+(let (($x60 (<= ?v1!0 0)))
+(let ((@x71 (|not-or-elim| @x68 $x60)))
+(|unit-resolution| (|unit-resolution| ((_ |th-lemma| arith farkas 1 1) (or $x63 (not $x60))) @x71 $x63) @x72 false))))))))))))))))))
+
+40bda8c0d6f113a97f5dc1db7c4465fe4cb9ac06 26 0
+unsat
+((set-logic <null>)
+(proof
+(let (($x56 (<= |b$| 0)))
+(let (($x60 (or (not (and (not (<= |a$| 0)) (not (<= (* |a$| |b$|) 0)))) (not $x56))))
+(let (($x63 (not $x60)))
+(let (($x14 (not (=> (and (< 0 |a$|) (< 0 (* |a$| |b$|))) (< 0 |b$|)))))
+(let (($x12 (< 0 |b$|)))
+(let (($x36 (or (not (and (< 0 |a$|) (< 0 (* |a$| |b$|)))) $x12)))
+(let (($x54 (= (not (and (< 0 |a$|) (< 0 (* |a$| |b$|)))) (not (and (not (<= |a$| 0)) (not (<= (* |a$| |b$|) 0)))))))
+(let ((?x9 (* |a$| |b$|)))
+(let (($x46 (<= ?x9 0)))
+(let (($x47 (not $x46)))
+(let (($x42 (<= |a$| 0)))
+(let (($x43 (not $x42)))
+(let (($x50 (and $x43 $x47)))
+(let (($x11 (and (< 0 |a$|) (< 0 ?x9))))
+(let ((@x52 (monotonicity (rewrite (= (< 0 |a$|) $x43)) (rewrite (= (< 0 ?x9) $x47)) (= $x11 $x50))))
+(let ((@x62 (monotonicity (monotonicity @x52 $x54) (rewrite (= $x12 (not $x56))) (= $x36 $x60))))
+(let ((@x41 (monotonicity (rewrite (= (=> $x11 $x12) $x36)) (= $x14 (not $x36)))))
+(let ((@x67 (trans @x41 (monotonicity @x62 (= (not $x36) $x63)) (= $x14 $x63))))
+(let ((@x72 (|not-or-elim| (mp (asserted $x14) @x67 $x63) $x56)))
+(let ((@x70 (|and-elim| (|not-or-elim| (mp (asserted $x14) @x67 $x63) $x50) $x43)))
+(let ((@x71 (|and-elim| (|not-or-elim| (mp (asserted $x14) @x67 $x63) $x50) $x47)))
+((_ |th-lemma| arith farkas 1 1 1) @x71 @x70 @x72 false))))))))))))))))))))))))
+
+3f9914c501829bba4f4b416ce311c4f49855326d 26 0
+unsat
+((set-logic <null>)
+(proof
+(let ((?x13 (+ |y$| 1)))
+(let ((?x14 (* |a$| ?x13)))
+(let ((?x12 (* |a$| |x$|)))
+(let ((?x15 (+ ?x12 ?x14)))
+(let ((?x8 (+ |x$| 1)))
+(let ((?x10 (+ ?x8 |y$|)))
+(let ((?x11 (* |a$| ?x10)))
+(let (($x16 (= ?x11 ?x15)))
+(let (($x17 (not $x16)))
+(let (($x80 (= (= (+ |a$| ?x12 (* |a$| |y$|)) (+ |a$| ?x12 (* |a$| |y$|))) true)))
+(let (($x78 (= $x16 (= (+ |a$| ?x12 (* |a$| |y$|)) (+ |a$| ?x12 (* |a$| |y$|))))))
+(let ((@x74 (rewrite (= (+ ?x12 (+ |a$| (* |a$| |y$|))) (+ |a$| ?x12 (* |a$| |y$|))))))
+(let ((@x64 (monotonicity (rewrite (= ?x13 (+ 1 |y$|))) (= ?x14 (* |a$| (+ 1 |y$|))))))
+(let ((@x69 (trans @x64 (rewrite (= (* |a$| (+ 1 |y$|)) (+ |a$| (* |a$| |y$|)))) (= ?x14 (+ |a$| (* |a$| |y$|))))))
+(let ((@x76 (trans (monotonicity @x69 (= ?x15 (+ ?x12 (+ |a$| (* |a$| |y$|))))) @x74 (= ?x15 (+ |a$| ?x12 (* |a$| |y$|))))))
+(let ((@x56 (rewrite (= (* |a$| (+ 1 |x$| |y$|)) (+ |a$| ?x12 (* |a$| |y$|))))))
+(let ((@x44 (monotonicity (rewrite (= ?x8 (+ 1 |x$|))) (= ?x10 (+ (+ 1 |x$|) |y$|)))))
+(let ((@x49 (trans @x44 (rewrite (= (+ (+ 1 |x$|) |y$|) (+ 1 |x$| |y$|))) (= ?x10 (+ 1 |x$| |y$|)))))
+(let ((@x58 (trans (monotonicity @x49 (= ?x11 (* |a$| (+ 1 |x$| |y$|)))) @x56 (= ?x11 (+ |a$| ?x12 (* |a$| |y$|))))))
+(let ((@x86 (monotonicity (trans (monotonicity @x58 @x76 $x78) (rewrite $x80) (= $x16 true)) (= $x17 (not true)))))
+(let ((@x90 (trans @x86 (rewrite (= (not true) false)) (= $x17 false))))
+(mp (asserted $x17) @x90 false))))))))))))))))))))))))
+
+f65cca85cf5c1c666974448574788ae3b34595b7 23 0
+unsat
+((set-logic <null>)
+(proof
+(let ((?x14 (* 2.0 |x$|)))
+(let ((?x15 (* ?x14 |y$|)))
+(let ((?x10 (- 1.0 |y$|)))
+(let ((?x11 (* |x$| ?x10)))
+(let ((?x8 (+ 1.0 |y$|)))
+(let ((?x9 (* |x$| ?x8)))
+(let ((?x12 (- ?x9 ?x11)))
+(let (($x16 (= ?x12 ?x15)))
+(let (($x17 (not $x16)))
+(let ((@x71 (rewrite (= (= (* 2.0 (* |x$| |y$|)) (* 2.0 (* |x$| |y$|))) true))))
+(let ((?x39 (* |x$| |y$|)))
+(let ((?x61 (* 2.0 ?x39)))
+(let ((@x54 (rewrite (= (* |x$| (+ 1.0 (* (~ 1.0) |y$|))) (+ |x$| (* (~ 1.0) ?x39))))))
+(let ((@x50 (monotonicity (rewrite (= ?x10 (+ 1.0 (* (~ 1.0) |y$|)))) (= ?x11 (* |x$| (+ 1.0 (* (~ 1.0) |y$|)))))))
+(let ((@x59 (monotonicity (rewrite (= ?x9 (+ |x$| ?x39))) (trans @x50 @x54 (= ?x11 (+ |x$| (* (~ 1.0) ?x39)))) (= ?x12 (- (+ |x$| ?x39) (+ |x$| (* (~ 1.0) ?x39)))))))
+(let ((@x64 (trans @x59 (rewrite (= (- (+ |x$| ?x39) (+ |x$| (* (~ 1.0) ?x39))) ?x61)) (= ?x12 ?x61))))
+(let ((@x73 (trans (monotonicity @x64 (rewrite (= ?x15 ?x61)) (= $x16 (= ?x61 ?x61))) @x71 (= $x16 true))))
+(let ((@x80 (trans (monotonicity @x73 (= $x17 (not true))) (rewrite (= (not true) false)) (= $x17 false))))
+(mp (asserted $x17) @x80 false)))))))))))))))))))))
+
+2643ba95811453f95121cf28c15c748e73c8c127 51 0
+unsat
+((set-logic <null>)
+(proof
+(let ((?x25 (+ |b$| |d$|)))
+(let ((?x26 (+ ?x25 |e$|)))
+(let ((?x8 (+ 1 |p$|)))
+(let ((?x27 (* ?x8 ?x26)))
+(let ((?x22 (* |d$| |p$|)))
+(let ((?x20 (* ?x8 |d$|)))
+(let ((?x11 (+ |b$| |e$|)))
+(let ((?x18 (* 2 ?x8)))
+(let ((?x19 (* ?x18 ?x11)))
+(let ((?x21 (+ ?x19 ?x20)))
+(let ((?x23 (+ ?x21 ?x22)))
+(let ((?x24 (+ |u$| ?x23)))
+(let ((?x28 (- ?x24 ?x27)))
+(let ((?x15 (* |p$| |d$|)))
+(let ((?x12 (* ?x8 ?x11)))
+(let ((?x13 (+ |u$| ?x12)))
+(let ((?x16 (+ ?x13 ?x15)))
+(let (($x29 (= ?x16 ?x28)))
+(let (($x30 (not $x29)))
+(let ((?x53 (* |p$| |e$|)))
+(let ((?x52 (* |p$| |b$|)))
+(let ((?x68 (+ |u$| |b$| |e$| ?x15 ?x52 ?x53)))
+(let ((?x125 (+ |b$| |e$| |d$| ?x15 ?x52 ?x53)))
+(let ((?x83 (* 2 ?x53)))
+(let ((?x81 (* 2 ?x52)))
+(let ((?x82 (* 2 |e$|)))
+(let ((?x80 (* 2 |b$|)))
+(let ((?x114 (+ |u$| ?x80 ?x82 |d$| (* 2 ?x15) ?x81 ?x83)))
+(let ((@x124 (monotonicity (rewrite (= ?x26 (+ |b$| |e$| |d$|))) (= ?x27 (* ?x8 (+ |b$| |e$| |d$|))))))
+(let ((@x129 (trans @x124 (rewrite (= (* ?x8 (+ |b$| |e$| |d$|)) ?x125)) (= ?x27 ?x125))))
+(let ((@x116 (rewrite (= (+ |u$| (+ ?x80 ?x82 |d$| (* 2 ?x15) ?x81 ?x83)) ?x114))))
+(let ((?x106 (+ ?x80 ?x82 |d$| (* 2 ?x15) ?x81 ?x83)))
+(let ((?x95 (+ ?x80 ?x82 |d$| ?x15 ?x81 ?x83)))
+(let ((@x86 (rewrite (= (* (+ 2 (* 2 |p$|)) ?x11) (+ ?x80 ?x82 ?x81 ?x83)))))
+(let ((@x79 (monotonicity (rewrite (= ?x18 (+ 2 (* 2 |p$|)))) (= ?x19 (* (+ 2 (* 2 |p$|)) ?x11)))))
+(let ((@x94 (monotonicity (trans @x79 @x86 (= ?x19 (+ ?x80 ?x82 ?x81 ?x83))) (rewrite (= ?x20 (+ |d$| ?x15))) (= ?x21 (+ (+ ?x80 ?x82 ?x81 ?x83) (+ |d$| ?x15))))))
+(let ((@x99 (trans @x94 (rewrite (= (+ (+ ?x80 ?x82 ?x81 ?x83) (+ |d$| ?x15)) ?x95)) (= ?x21 ?x95))))
+(let ((@x110 (trans (monotonicity @x99 (rewrite (= ?x22 ?x15)) (= ?x23 (+ ?x95 ?x15))) (rewrite (= (+ ?x95 ?x15) ?x106)) (= ?x23 ?x106))))
+(let ((@x118 (trans (monotonicity @x110 (= ?x24 (+ |u$| ?x106))) @x116 (= ?x24 ?x114))))
+(let ((@x137 (trans (monotonicity @x118 @x129 (= ?x28 (- ?x114 ?x125))) (rewrite (= (- ?x114 ?x125) ?x68)) (= ?x28 ?x68))))
+(let ((@x62 (rewrite (= (+ |u$| (+ |b$| |e$| ?x52 ?x53)) (+ |u$| |b$| |e$| ?x52 ?x53)))))
+(let ((@x59 (monotonicity (rewrite (= ?x12 (+ |b$| |e$| ?x52 ?x53))) (= ?x13 (+ |u$| (+ |b$| |e$| ?x52 ?x53))))))
+(let ((@x67 (monotonicity (trans @x59 @x62 (= ?x13 (+ |u$| |b$| |e$| ?x52 ?x53))) (= ?x16 (+ (+ |u$| |b$| |e$| ?x52 ?x53) ?x15)))))
+(let ((@x72 (trans @x67 (rewrite (= (+ (+ |u$| |b$| |e$| ?x52 ?x53) ?x15) ?x68)) (= ?x16 ?x68))))
+(let ((@x143 (trans (monotonicity @x72 @x137 (= $x29 (= ?x68 ?x68))) (rewrite (= (= ?x68 ?x68) true)) (= $x29 true))))
+(let ((@x150 (trans (monotonicity @x143 (= $x30 (not true))) (rewrite (= (not true) false)) (= $x30 false))))
+(mp (asserted $x30) @x150 false)))))))))))))))))))))))))))))))))))))))))))))))))
+
+9377273e8e637d8916ed13b81bd56a602ea76d29 126 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let ((?x7 (|of_nat$| |x$|)))
+(let ((?x8 (* 2 ?x7)))
+(let ((?x9 (|nat$| ?x8)))
+(let ((?x147 (|of_nat$| ?x9)))
+(let ((?x160 (* (~ 1) ?x147)))
+(let ((?x161 (+ ?x8 ?x160)))
+(let (($x177 (<= ?x161 0)))
+(let (($x158 (= ?x161 0)))
+(let (($x150 (>= ?x7 0)))
+(let (($x239 (>= ?x147 1)))
+(let (($x237 (= ?x147 1)))
+(let ((?x11 (|nat$| 1)))
+(let ((?x200 (|of_nat$| ?x11)))
+(let (($x201 (= ?x200 1)))
+(let (($x128 (forall ((?v0 Int) )(!(let ((?x23 (|nat$| ?v0)))
+(let ((?x24 (|of_nat$| ?x23)))
+(let (($x25 (= ?x24 ?v0)))
+(let (($x66 (>= ?v0 0)))
+(let (($x67 (not $x66)))
+(or $x67 $x25)))))) :pattern ( (|nat$| ?v0) )))
+))
+(let (($x73 (forall ((?v0 Int) )(let ((?x23 (|nat$| ?v0)))
+(let ((?x24 (|of_nat$| ?x23)))
+(let (($x25 (= ?x24 ?v0)))
+(let (($x66 (>= ?v0 0)))
+(let (($x67 (not $x66)))
+(or $x67 $x25)))))))
+))
+(let ((?x23 (|nat$| ?0)))
+(let ((?x24 (|of_nat$| ?x23)))
+(let (($x25 (= ?x24 ?0)))
+(let (($x66 (>= ?0 0)))
+(let (($x67 (not $x66)))
+(let (($x70 (or $x67 $x25)))
+(let (($x27 (forall ((?v0 Int) )(let ((?x23 (|nat$| ?v0)))
+(let ((?x24 (|of_nat$| ?x23)))
+(let (($x25 (= ?x24 ?v0)))
+(let (($x22 (<= 0 ?v0)))
+(=> $x22 $x25))))))
+))
+(let (($x61 (forall ((?v0 Int) )(let ((?x23 (|nat$| ?v0)))
+(let ((?x24 (|of_nat$| ?x23)))
+(let (($x25 (= ?x24 ?v0)))
+(or (not (<= 0 ?v0)) $x25)))))
+))
+(let ((@x69 (monotonicity (rewrite (= (<= 0 ?0) $x66)) (= (not (<= 0 ?0)) $x67))))
+(let ((@x75 (|quant-intro| (monotonicity @x69 (= (or (not (<= 0 ?0)) $x25) $x70)) (= $x61 $x73))))
+(let ((@x60 (rewrite (= (=> (<= 0 ?0) $x25) (or (not (<= 0 ?0)) $x25)))))
+(let ((@x78 (mp (asserted $x27) (trans (|quant-intro| @x60 (= $x27 $x61)) @x75 (= $x27 $x73)) $x73)))
+(let ((@x133 (mp (|mp~| @x78 (|nnf-pos| (refl (|~| $x70 $x70)) (|~| $x73 $x73)) $x73) (|quant-intro| (refl (= $x70 $x70)) (= $x73 $x128)) $x128)))
+(let (($x165 (not $x128)))
+(let (($x219 (or $x165 $x201)))
+(let ((@x204 (rewrite (= (>= 1 0) true))))
+(let ((@x211 (trans (monotonicity @x204 (= (not (>= 1 0)) (not true))) (rewrite (= (not true) false)) (= (not (>= 1 0)) false))))
+(let ((@x214 (monotonicity @x211 (= (or (not (>= 1 0)) $x201) (or false $x201)))))
+(let ((@x218 (trans @x214 (rewrite (= (or false $x201) $x201)) (= (or (not (>= 1 0)) $x201) $x201))))
+(let ((@x223 (monotonicity @x218 (= (or $x165 (or (not (>= 1 0)) $x201)) $x219))))
+(let ((@x226 (trans @x223 (rewrite (= $x219 $x219)) (= (or $x165 (or (not (>= 1 0)) $x201)) $x219))))
+(let ((@x227 (mp ((_ |quant-inst| 1) (or $x165 (or (not (>= 1 0)) $x201))) @x226 $x219)))
+(let (($x12 (= ?x9 ?x11)))
+(let ((@x56 (mp (asserted (not (not $x12))) (rewrite (= (not (not $x12)) $x12)) $x12)))
+(let ((@x252 (trans (monotonicity @x56 (= ?x147 ?x200)) (|unit-resolution| @x227 @x133 $x201) $x237)))
+(let ((@x261 (|unit-resolution| ((_ |th-lemma| arith farkas 1 1) (or (not $x239) (not (<= ?x147 0)))) (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x237) $x239)) @x252 $x239) (not (<= ?x147 0)))))
+(let ((@x265 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not (= ?x147 0)) (<= ?x147 0))) @x261 (not (= ?x147 0)))))
+(let (($x179 (= ?x147 0)))
+(let (($x181 (or $x150 $x179)))
+(let (($x134 (forall ((?v0 Int) )(!(let ((?x23 (|nat$| ?v0)))
+(let ((?x24 (|of_nat$| ?x23)))
+(let (($x29 (= ?x24 0)))
+(let (($x66 (>= ?v0 0)))
+(or $x66 $x29))))) :pattern ( (|nat$| ?v0) )))
+))
+(let (($x99 (forall ((?v0 Int) )(let ((?x23 (|nat$| ?v0)))
+(let ((?x24 (|of_nat$| ?x23)))
+(let (($x29 (= ?x24 0)))
+(let (($x66 (>= ?v0 0)))
+(or $x66 $x29))))))
+))
+(let ((@x138 (|quant-intro| (refl (= (or $x66 (= ?x24 0)) (or $x66 (= ?x24 0)))) (= $x99 $x134))))
+(let ((@x118 (|nnf-pos| (refl (|~| (or $x66 (= ?x24 0)) (or $x66 (= ?x24 0)))) (|~| $x99 $x99))))
+(let (($x31 (forall ((?v0 Int) )(let ((?x23 (|nat$| ?v0)))
+(let ((?x24 (|of_nat$| ?x23)))
+(let (($x29 (= ?x24 0)))
+(let (($x28 (< ?v0 0)))
+(=> $x28 $x29))))))
+))
+(let (($x84 (forall ((?v0 Int) )(let ((?x23 (|nat$| ?v0)))
+(let ((?x24 (|of_nat$| ?x23)))
+(let (($x29 (= ?x24 0)))
+(let (($x28 (< ?v0 0)))
+(let (($x80 (not $x28)))
+(or $x80 $x29)))))))
+))
+(let (($x29 (= ?x24 0)))
+(let (($x96 (or $x66 $x29)))
+(let (($x28 (< ?0 0)))
+(let (($x80 (not $x28)))
+(let (($x81 (or $x80 $x29)))
+(let ((@x95 (trans (monotonicity (rewrite (= $x28 $x67)) (= $x80 (not $x67))) (rewrite (= (not $x67) $x66)) (= $x80 $x66))))
+(let ((@x103 (trans (|quant-intro| (rewrite (= (=> $x28 $x29) $x81)) (= $x31 $x84)) (|quant-intro| (monotonicity @x95 (= $x81 $x96)) (= $x84 $x99)) (= $x31 $x99))))
+(let ((@x139 (mp (|mp~| (mp (asserted $x31) @x103 $x99) @x118 $x99) @x138 $x134)))
+(let (($x184 (not $x134)))
+(let (($x185 (or $x184 $x150 $x179)))
+(let ((@x152 (rewrite (= (>= ?x8 0) $x150))))
+(let ((@x190 (monotonicity (monotonicity @x152 (= (or (>= ?x8 0) $x179) $x181)) (= (or $x184 (or (>= ?x8 0) $x179)) (or $x184 $x181)))))
+(let ((@x194 (trans @x190 (rewrite (= (or $x184 $x181) $x185)) (= (or $x184 (or (>= ?x8 0) $x179)) $x185))))
+(let ((@x195 (mp ((_ |quant-inst| (* 2 ?x7)) (or $x184 (or (>= ?x8 0) $x179))) @x194 $x185)))
+(let (($x153 (not $x150)))
+(let (($x162 (or $x153 $x158)))
+(let (($x166 (or $x165 $x153 $x158)))
+(let (($x148 (= ?x147 ?x8)))
+(let (($x142 (>= ?x8 0)))
+(let (($x143 (not $x142)))
+(let (($x149 (or $x143 $x148)))
+(let (($x167 (or $x165 $x149)))
+(let ((@x164 (monotonicity (monotonicity @x152 (= $x143 $x153)) (rewrite (= $x148 $x158)) (= $x149 $x162))))
+(let ((@x175 (trans (monotonicity @x164 (= $x167 (or $x165 $x162))) (rewrite (= (or $x165 $x162) $x166)) (= $x167 $x166))))
+(let ((@x176 (mp ((_ |quant-inst| (* 2 ?x7)) $x167) @x175 $x166)))
+(let ((@x269 (|unit-resolution| (|unit-resolution| @x176 @x133 $x162) (|unit-resolution| (|unit-resolution| @x195 @x139 $x181) @x265 $x150) $x158)))
+(let (($x178 (>= ?x161 0)))
+(let (($x238 (<= ?x147 1)))
+((_ |th-lemma| arith gcd-test -1/2 -1/2 -1/2 -1/2) (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x237) $x239)) @x252 $x239) (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x237) $x238)) @x252 $x238) (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x158) $x178)) @x269 $x178) (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x158) $x177)) @x269 $x177) false))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))
+
+7540f10f61e5a987b0848b309bd25f2a2ae1cd0a 22 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let ((?x6 (|of_nat$| |a$|)))
+(let (($x71 (>= ?x6 4)))
+(let (($x78 (not (or (>= ?x6 3) (not $x71)))))
+(let (($x12 (< (* 2 ?x6) 7)))
+(let (($x8 (< ?x6 3)))
+(let (($x52 (not $x8)))
+(let (($x53 (or $x52 $x12)))
+(let ((@x65 (monotonicity (rewrite (= $x8 (not (>= ?x6 3)))) (= $x52 (not (not (>= ?x6 3)))))))
+(let ((@x69 (trans @x65 (rewrite (= (not (not (>= ?x6 3))) (>= ?x6 3))) (= $x52 (>= ?x6 3)))))
+(let ((@x77 (monotonicity @x69 (rewrite (= $x12 (not $x71))) (= $x53 (or (>= ?x6 3) (not $x71))))))
+(let ((@x58 (monotonicity (rewrite (= (=> $x8 $x12) $x53)) (= (not (=> $x8 $x12)) (not $x53)))))
+(let ((@x82 (trans @x58 (monotonicity @x77 (= (not $x53) $x78)) (= (not (=> $x8 $x12)) $x78))))
+(let ((@x85 (|not-or-elim| (mp (asserted (not (=> $x8 $x12))) @x82 $x78) $x71)))
+(let (($x72 (not $x71)))
+(let (($x61 (>= ?x6 3)))
+(let (($x59 (not $x61)))
+(let ((@x84 (|not-or-elim| (mp (asserted (not (=> $x8 $x12))) @x82 $x78) $x59)))
+(|unit-resolution| (|unit-resolution| ((_ |th-lemma| arith farkas 1 1) (or $x72 $x61)) @x84 $x72) @x85 false))))))))))))))))))))
+
+3a9a1f0f87885c249813dfb78d14e1062fc20ce3 147 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let ((?x8 (|of_nat$| |y$|)))
+(let ((?x9 (+ 1 ?x8)))
+(let ((?x10 (|nat$| ?x9)))
+(let ((?x11 (|of_nat$| ?x10)))
+(let ((?x57 (* (~ 1) ?x8)))
+(let ((?x58 (+ ?x57 ?x11)))
+(let ((?x61 (|nat$| ?x58)))
+(let ((?x64 (|of_nat$| ?x61)))
+(let ((?x195 (* (~ 1) ?x11)))
+(let ((?x246 (+ ?x8 ?x195 ?x64)))
+(let (($x265 (>= ?x246 0)))
+(let (($x247 (= ?x246 0)))
+(let ((?x196 (+ ?x8 ?x195)))
+(let (($x240 (<= ?x196 0)))
+(let (($x215 (<= ?x196 (~ 1))))
+(let (($x197 (= ?x196 (~ 1))))
+(let (($x189 (>= ?x8 (~ 1))))
+(let (($x283 (>= ?x8 0)))
+(let ((?x172 (|nat$| ?x8)))
+(let ((?x284 (|of_nat$| ?x172)))
+(let (($x285 (= ?x284 0)))
+(let (($x286 (or $x283 $x285)))
+(let (($x166 (forall ((?v0 Int) )(!(let ((?x25 (|nat$| ?v0)))
+(let ((?x26 (|of_nat$| ?x25)))
+(let (($x31 (= ?x26 0)))
+(let (($x97 (>= ?v0 0)))
+(or $x97 $x31))))) :pattern ( (|nat$| ?v0) )))
+))
+(let (($x131 (forall ((?v0 Int) )(let ((?x25 (|nat$| ?v0)))
+(let ((?x26 (|of_nat$| ?x25)))
+(let (($x31 (= ?x26 0)))
+(let (($x97 (>= ?v0 0)))
+(or $x97 $x31))))))
+))
+(let ((?x25 (|nat$| ?0)))
+(let ((?x26 (|of_nat$| ?x25)))
+(let (($x31 (= ?x26 0)))
+(let (($x97 (>= ?0 0)))
+(let (($x128 (or $x97 $x31)))
+(let (($x33 (forall ((?v0 Int) )(let ((?x25 (|nat$| ?v0)))
+(let ((?x26 (|of_nat$| ?x25)))
+(let (($x31 (= ?x26 0)))
+(let (($x30 (< ?v0 0)))
+(=> $x30 $x31))))))
+))
+(let (($x116 (forall ((?v0 Int) )(let ((?x25 (|nat$| ?v0)))
+(let ((?x26 (|of_nat$| ?x25)))
+(let (($x31 (= ?x26 0)))
+(let (($x30 (< ?v0 0)))
+(let (($x112 (not $x30)))
+(or $x112 $x31)))))))
+))
+(let ((@x123 (monotonicity (rewrite (= (< ?0 0) (not $x97))) (= (not (< ?0 0)) (not (not $x97))))))
+(let ((@x127 (trans @x123 (rewrite (= (not (not $x97)) $x97)) (= (not (< ?0 0)) $x97))))
+(let ((@x133 (|quant-intro| (monotonicity @x127 (= (or (not (< ?0 0)) $x31) $x128)) (= $x116 $x131))))
+(let ((@x115 (rewrite (= (=> (< ?0 0) $x31) (or (not (< ?0 0)) $x31)))))
+(let ((@x136 (mp (asserted $x33) (trans (|quant-intro| @x115 (= $x33 $x116)) @x133 (= $x33 $x131)) $x131)))
+(let ((@x171 (mp (|mp~| @x136 (|nnf-pos| (refl (|~| $x128 $x128)) (|~| $x131 $x131)) $x131) (|quant-intro| (refl (= $x128 $x128)) (= $x131 $x166)) $x166)))
+(let (($x222 (not $x166)))
+(let (($x289 (or $x222 $x283 $x285)))
+(let ((@x294 (mp ((_ |quant-inst| (|of_nat$| |y$|)) (or $x222 $x286)) (rewrite (= (or $x222 $x286) $x289)) $x289)))
+(let ((@x316 (|unit-resolution| (|unit-resolution| @x294 @x171 $x286) (hypothesis (not $x283)) $x285)))
+(let (($x173 (= ?x172 |y$|)))
+(let (($x153 (forall ((?v0 |Nat$|) )(!(= (|nat$| (|of_nat$| ?v0)) ?v0) :pattern ( (|of_nat$| ?v0) )))
+))
+(let (($x22 (forall ((?v0 |Nat$|) )(= (|nat$| (|of_nat$| ?v0)) ?v0))
+))
+(let ((@x155 (refl (= (= (|nat$| (|of_nat$| ?0)) ?0) (= (|nat$| (|of_nat$| ?0)) ?0)))))
+(let ((@x140 (refl (|~| (= (|nat$| (|of_nat$| ?0)) ?0) (= (|nat$| (|of_nat$| ?0)) ?0)))))
+(let ((@x158 (mp (|mp~| (asserted $x22) (|nnf-pos| @x140 (|~| $x22 $x22)) $x22) (|quant-intro| @x155 (= $x22 $x153)) $x153)))
+(let (($x176 (not $x153)))
+(let (($x177 (or $x176 $x173)))
+(let ((@x178 ((_ |quant-inst| |y$|) $x177)))
+(let ((@x321 (monotonicity (symm (|unit-resolution| @x178 @x158 $x173) (= |y$| ?x172)) (= ?x8 ?x284))))
+(let ((@x326 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not (= ?x8 0)) $x283)) (hypothesis (not $x283)) (trans @x321 @x316 (= ?x8 0)) false)))
+(let ((@x329 (|unit-resolution| ((_ |th-lemma| arith farkas 1 1) (or (not $x283) $x189)) (lemma @x326 $x283) $x189)))
+(let (($x192 (not $x189)))
+(let (($x200 (or $x192 $x197)))
+(let (($x160 (forall ((?v0 Int) )(!(let ((?x25 (|nat$| ?v0)))
+(let ((?x26 (|of_nat$| ?x25)))
+(let (($x27 (= ?x26 ?v0)))
+(let (($x97 (>= ?v0 0)))
+(let (($x99 (not $x97)))
+(or $x99 $x27)))))) :pattern ( (|nat$| ?v0) )))
+))
+(let (($x105 (forall ((?v0 Int) )(let ((?x25 (|nat$| ?v0)))
+(let ((?x26 (|of_nat$| ?x25)))
+(let (($x27 (= ?x26 ?v0)))
+(let (($x97 (>= ?v0 0)))
+(let (($x99 (not $x97)))
+(or $x99 $x27)))))))
+))
+(let ((@x162 (refl (= (or (not $x97) (= ?x26 ?0)) (or (not $x97) (= ?x26 ?0))))))
+(let ((@x143 (refl (|~| (or (not $x97) (= ?x26 ?0)) (or (not $x97) (= ?x26 ?0))))))
+(let (($x29 (forall ((?v0 Int) )(let ((?x25 (|nat$| ?v0)))
+(let ((?x26 (|of_nat$| ?x25)))
+(let (($x27 (= ?x26 ?v0)))
+(let (($x24 (<= 0 ?v0)))
+(=> $x24 $x27))))))
+))
+(let (($x93 (forall ((?v0 Int) )(let ((?x25 (|nat$| ?v0)))
+(let ((?x26 (|of_nat$| ?x25)))
+(let (($x27 (= ?x26 ?v0)))
+(or (not (<= 0 ?v0)) $x27)))))
+))
+(let (($x27 (= ?x26 ?0)))
+(let (($x99 (not $x97)))
+(let (($x102 (or $x99 $x27)))
+(let (($x90 (or (not (<= 0 ?0)) $x27)))
+(let ((@x101 (monotonicity (rewrite (= (<= 0 ?0) $x97)) (= (not (<= 0 ?0)) $x99))))
+(let ((@x95 (|quant-intro| (rewrite (= (=> (<= 0 ?0) $x27) $x90)) (= $x29 $x93))))
+(let ((@x109 (trans @x95 (|quant-intro| (monotonicity @x101 (= $x90 $x102)) (= $x93 $x105)) (= $x29 $x105))))
+(let ((@x146 (|mp~| (mp (asserted $x29) @x109 $x105) (|nnf-pos| @x143 (|~| $x105 $x105)) $x105)))
+(let ((@x165 (mp @x146 (|quant-intro| @x162 (= $x105 $x160)) $x160)))
+(let (($x203 (not $x160)))
+(let (($x204 (or $x203 $x192 $x197)))
+(let (($x188 (or (not (>= ?x9 0)) (= ?x11 ?x9))))
+(let (($x205 (or $x203 $x188)))
+(let ((@x194 (monotonicity (rewrite (= (>= ?x9 0) $x189)) (= (not (>= ?x9 0)) $x192))))
+(let ((@x202 (monotonicity @x194 (rewrite (= (= ?x11 ?x9) $x197)) (= $x188 $x200))))
+(let ((@x213 (trans (monotonicity @x202 (= $x205 (or $x203 $x200))) (rewrite (= (or $x203 $x200) $x204)) (= $x205 $x204))))
+(let ((@x214 (mp ((_ |quant-inst| (+ 1 ?x8)) $x205) @x213 $x204)))
+(let ((@x335 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x197) $x215)) (|unit-resolution| (|unit-resolution| @x214 @x165 $x200) @x329 $x197) $x215)))
+(let (($x243 (not $x240)))
+(let (($x250 (or $x243 $x247)))
+(let (($x253 (or $x203 $x243 $x247)))
+(let (($x239 (or (not (>= ?x58 0)) (= ?x64 ?x58))))
+(let (($x254 (or $x203 $x239)))
+(let ((@x245 (monotonicity (rewrite (= (>= ?x58 0) $x240)) (= (not (>= ?x58 0)) $x243))))
+(let ((@x252 (monotonicity @x245 (rewrite (= (= ?x64 ?x58) $x247)) (= $x239 $x250))))
+(let ((@x262 (trans (monotonicity @x252 (= $x254 (or $x203 $x250))) (rewrite (= (or $x203 $x250) $x253)) (= $x254 $x253))))
+(let ((@x263 (mp ((_ |quant-inst| (+ ?x57 ?x11)) $x254) @x262 $x253)))
+(let ((@x341 (|unit-resolution| (|unit-resolution| @x263 @x165 $x250) (|unit-resolution| ((_ |th-lemma| arith farkas 1 1) (or (not $x215) $x240)) @x335 $x240) $x247)))
+(let (($x73 (<= ?x64 0)))
+(let ((@x79 (monotonicity (rewrite (= (< 0 ?x64) (not $x73))) (= (not (< 0 ?x64)) (not (not $x73))))))
+(let ((@x83 (trans @x79 (rewrite (= (not (not $x73)) $x73)) (= (not (< 0 ?x64)) $x73))))
+(let (($x67 (< 0 ?x64)))
+(let (($x70 (not $x67)))
+(let (($x17 (not (< (* 0 ?x11) (|of_nat$| (|nat$| (- ?x11 ?x8)))))))
+(let ((@x63 (monotonicity (rewrite (= (- ?x11 ?x8) ?x58)) (= (|nat$| (- ?x11 ?x8)) ?x61))))
+(let ((@x69 (monotonicity (rewrite (= (* 0 ?x11) 0)) (monotonicity @x63 (= (|of_nat$| (|nat$| (- ?x11 ?x8))) ?x64)) (= (< (* 0 ?x11) (|of_nat$| (|nat$| (- ?x11 ?x8)))) $x67))))
+(let ((@x86 (mp (asserted $x17) (trans (monotonicity @x69 (= $x17 $x70)) @x83 (= $x17 $x73)) $x73)))
+((_ |th-lemma| arith farkas -1 -1 1) @x86 @x335 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x247) $x265)) @x341 $x265) false)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))
+
+72e646f619a773762ccf2e62425eb512a9cd35f3 144 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let ((?x8 (|of_nat$| |y$|)))
+(let ((?x9 (+ 1 ?x8)))
+(let ((?x10 (|nat$| ?x9)))
+(let ((?x11 (|of_nat$| ?x10)))
+(let ((?x62 (+ (~ 1) ?x11)))
+(let ((?x65 (|nat$| ?x62)))
+(let ((?x281 (|of_nat$| ?x65)))
+(let ((?x291 (* (~ 1) ?x281)))
+(let ((?x330 (+ ?x8 ?x291)))
+(let (($x332 (>= ?x330 0)))
+(let (($x329 (= ?x8 ?x281)))
+(let (($x68 (= ?x65 |y$|)))
+(let (($x102 (<= ?x11 0)))
+(let (($x112 (not (or (= (not $x102) $x68) (not $x102)))))
+(let (($x19 (=> (not (ite (< 0 ?x11) true false)) false)))
+(let (($x12 (< 0 ?x11)))
+(let (($x13 (ite $x12 true false)))
+(let (($x17 (= $x13 (= (|nat$| (- ?x11 1)) |y$|))))
+(let (($x21 (or false (or $x17 $x19))))
+(let (($x22 (not $x21)))
+(let (($x74 (= $x12 $x68)))
+(let (($x89 (or $x74 $x12)))
+(let ((@x108 (monotonicity (rewrite (= $x12 (not $x102))) (= $x74 (= (not $x102) $x68)))))
+(let ((@x111 (monotonicity @x108 (rewrite (= $x12 (not $x102))) (= $x89 (or (= (not $x102) $x68) (not $x102))))))
+(let ((@x84 (monotonicity (monotonicity (rewrite (= $x13 $x12)) (= (not $x13) (not $x12))) (= $x19 (=> (not $x12) false)))))
+(let ((@x88 (trans @x84 (rewrite (= (=> (not $x12) false) $x12)) (= $x19 $x12))))
+(let ((@x67 (monotonicity (rewrite (= (- ?x11 1) ?x62)) (= (|nat$| (- ?x11 1)) ?x65))))
+(let ((@x73 (monotonicity (rewrite (= $x13 $x12)) (monotonicity @x67 (= (= (|nat$| (- ?x11 1)) |y$|) $x68)) (= $x17 (= $x12 $x68)))))
+(let ((@x91 (monotonicity (trans @x73 (rewrite (= (= $x12 $x68) $x74)) (= $x17 $x74)) @x88 (= (or $x17 $x19) $x89))))
+(let ((@x98 (trans (monotonicity @x91 (= $x21 (or false $x89))) (rewrite (= (or false $x89) $x89)) (= $x21 $x89))))
+(let ((@x116 (trans (monotonicity @x98 (= $x22 (not $x89))) (monotonicity @x111 (= (not $x89) $x112)) (= $x22 $x112))))
+(let ((@x120 (|not-or-elim| (mp (asserted $x22) @x116 $x112) $x102)))
+(let (($x171 (= $x102 $x68)))
+(let ((@x119 (|not-or-elim| (mp (asserted $x22) @x116 $x112) (not (= (not $x102) $x68)))))
+(let ((@x173 (mp @x119 (rewrite (= (not (= (not $x102) $x68)) $x171)) $x171)))
+(let ((@x219 (|unit-resolution| (|def-axiom| (or (not $x102) $x68 (not $x171))) @x173 (or (not $x102) $x68))))
+(let ((@x345 (monotonicity (symm (|unit-resolution| @x219 @x120 $x68) (= |y$| ?x65)) $x329)))
+(let ((?x241 (* (~ 1) ?x11)))
+(let ((?x242 (+ ?x8 ?x241)))
+(let (($x259 (<= ?x242 (~ 1))))
+(let (($x240 (= ?x242 (~ 1))))
+(let (($x233 (>= ?x8 (~ 1))))
+(let (($x328 (>= ?x281 0)))
+(let (($x311 (= ?x281 0)))
+(let (($x284 (>= ?x11 1)))
+(let (($x287 (not $x284)))
+(let (($x204 (forall ((?v0 Int) )(!(let ((?x30 (|nat$| ?v0)))
+(let ((?x31 (|of_nat$| ?x30)))
+(let (($x36 (= ?x31 0)))
+(let (($x131 (>= ?v0 0)))
+(or $x131 $x36))))) :pattern ( (|nat$| ?v0) )))
+))
+(let (($x165 (forall ((?v0 Int) )(let ((?x30 (|nat$| ?v0)))
+(let ((?x31 (|of_nat$| ?x30)))
+(let (($x36 (= ?x31 0)))
+(let (($x131 (>= ?v0 0)))
+(or $x131 $x36))))))
+))
+(let ((?x30 (|nat$| ?0)))
+(let ((?x31 (|of_nat$| ?x30)))
+(let (($x36 (= ?x31 0)))
+(let (($x131 (>= ?0 0)))
+(let (($x162 (or $x131 $x36)))
+(let (($x38 (forall ((?v0 Int) )(let ((?x30 (|nat$| ?v0)))
+(let ((?x31 (|of_nat$| ?x30)))
+(let (($x36 (= ?x31 0)))
+(let (($x35 (< ?v0 0)))
+(=> $x35 $x36))))))
+))
+(let (($x150 (forall ((?v0 Int) )(let ((?x30 (|nat$| ?v0)))
+(let ((?x31 (|of_nat$| ?x30)))
+(let (($x36 (= ?x31 0)))
+(let (($x35 (< ?v0 0)))
+(let (($x146 (not $x35)))
+(or $x146 $x36)))))))
+))
+(let ((@x157 (monotonicity (rewrite (= (< ?0 0) (not $x131))) (= (not (< ?0 0)) (not (not $x131))))))
+(let ((@x161 (trans @x157 (rewrite (= (not (not $x131)) $x131)) (= (not (< ?0 0)) $x131))))
+(let ((@x167 (|quant-intro| (monotonicity @x161 (= (or (not (< ?0 0)) $x36) $x162)) (= $x150 $x165))))
+(let ((@x149 (rewrite (= (=> (< ?0 0) $x36) (or (not (< ?0 0)) $x36)))))
+(let ((@x170 (mp (asserted $x38) (trans (|quant-intro| @x149 (= $x38 $x150)) @x167 (= $x38 $x165)) $x165)))
+(let ((@x209 (mp (|mp~| @x170 (|nnf-pos| (refl (|~| $x162 $x162)) (|~| $x165 $x165)) $x165) (|quant-intro| (refl (= $x162 $x162)) (= $x165 $x204)) $x204)))
+(let (($x266 (not $x204)))
+(let (($x316 (or $x266 $x284 $x311)))
+(let ((@x286 (rewrite (= (>= ?x62 0) $x284))))
+(let ((@x321 (monotonicity (monotonicity @x286 (= (or (>= ?x62 0) $x311) (or $x284 $x311))) (= (or $x266 (or (>= ?x62 0) $x311)) (or $x266 (or $x284 $x311))))))
+(let ((@x325 (trans @x321 (rewrite (= (or $x266 (or $x284 $x311)) $x316)) (= (or $x266 (or (>= ?x62 0) $x311)) $x316))))
+(let ((@x326 (mp ((_ |quant-inst| (+ (~ 1) ?x11)) (or $x266 (or (>= ?x62 0) $x311))) @x325 $x316)))
+(let ((@x353 (|unit-resolution| @x326 @x209 (|unit-resolution| ((_ |th-lemma| arith farkas 1 1) (or $x287 (not $x102))) @x120 $x287) $x311)))
+(let ((@x362 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 1) (or $x233 (not $x328) (not $x332))) (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x311) $x328)) @x353 $x328) (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x329) $x332)) @x345 $x332) $x233)))
+(let (($x236 (not $x233)))
+(let (($x244 (or $x236 $x240)))
+(let (($x198 (forall ((?v0 Int) )(!(let ((?x30 (|nat$| ?v0)))
+(let ((?x31 (|of_nat$| ?x30)))
+(let (($x32 (= ?x31 ?v0)))
+(let (($x131 (>= ?v0 0)))
+(let (($x133 (not $x131)))
+(or $x133 $x32)))))) :pattern ( (|nat$| ?v0) )))
+))
+(let (($x139 (forall ((?v0 Int) )(let ((?x30 (|nat$| ?v0)))
+(let ((?x31 (|of_nat$| ?x30)))
+(let (($x32 (= ?x31 ?v0)))
+(let (($x131 (>= ?v0 0)))
+(let (($x133 (not $x131)))
+(or $x133 $x32)))))))
+))
+(let ((@x200 (refl (= (or (not $x131) (= ?x31 ?0)) (or (not $x131) (= ?x31 ?0))))))
+(let ((@x181 (refl (|~| (or (not $x131) (= ?x31 ?0)) (or (not $x131) (= ?x31 ?0))))))
+(let (($x34 (forall ((?v0 Int) )(let ((?x30 (|nat$| ?v0)))
+(let ((?x31 (|of_nat$| ?x30)))
+(let (($x32 (= ?x31 ?v0)))
+(let (($x29 (<= 0 ?v0)))
+(=> $x29 $x32))))))
+))
+(let (($x127 (forall ((?v0 Int) )(let ((?x30 (|nat$| ?v0)))
+(let ((?x31 (|of_nat$| ?x30)))
+(let (($x32 (= ?x31 ?v0)))
+(or (not (<= 0 ?v0)) $x32)))))
+))
+(let (($x32 (= ?x31 ?0)))
+(let (($x133 (not $x131)))
+(let (($x136 (or $x133 $x32)))
+(let (($x124 (or (not (<= 0 ?0)) $x32)))
+(let ((@x135 (monotonicity (rewrite (= (<= 0 ?0) $x131)) (= (not (<= 0 ?0)) $x133))))
+(let ((@x129 (|quant-intro| (rewrite (= (=> (<= 0 ?0) $x32) $x124)) (= $x34 $x127))))
+(let ((@x143 (trans @x129 (|quant-intro| (monotonicity @x135 (= $x124 $x136)) (= $x127 $x139)) (= $x34 $x139))))
+(let ((@x184 (|mp~| (mp (asserted $x34) @x143 $x139) (|nnf-pos| @x181 (|~| $x139 $x139)) $x139)))
+(let ((@x203 (mp @x184 (|quant-intro| @x200 (= $x139 $x198)) $x198)))
+(let (($x247 (not $x198)))
+(let (($x248 (or $x247 $x236 $x240)))
+(let (($x231 (= ?x11 ?x9)))
+(let (($x227 (>= ?x9 0)))
+(let (($x228 (not $x227)))
+(let (($x232 (or $x228 $x231)))
+(let (($x249 (or $x247 $x232)))
+(let ((@x246 (monotonicity (monotonicity (rewrite (= $x227 $x233)) (= $x228 $x236)) (rewrite (= $x231 $x240)) (= $x232 $x244))))
+(let ((@x257 (trans (monotonicity @x246 (= $x249 (or $x247 $x244))) (rewrite (= (or $x247 $x244) $x248)) (= $x249 $x248))))
+(let ((@x258 (mp ((_ |quant-inst| (+ 1 ?x8)) $x249) @x257 $x248)))
+(let ((@x368 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x240) $x259)) (|unit-resolution| (|unit-resolution| @x258 @x203 $x244) @x362 $x240) $x259)))
+((_ |th-lemma| arith farkas 1 -1 -1 1) (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x311) $x328)) @x353 $x328) @x120 @x368 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x329) $x332)) @x345 $x332) false))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))
+
+bda6be21dc699a816acc75c786757ad36dd913c8 78 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let ((?x51 (* (~ 1) |x$|)))
+(let (($x69 (>= |x$| 0)))
+(let ((?x76 (ite $x69 |x$| ?x51)))
+(let ((?x213 (* (~ 1) ?x76)))
+(let ((?x216 (+ ?x51 ?x213)))
+(let (($x226 (<= ?x216 0)))
+(let (($x182 (= ?x51 ?x76)))
+(let (($x70 (not $x69)))
+(let (($x181 (= |x$| ?x76)))
+(let ((@x222 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x181) (<= (+ |x$| ?x213) 0))) (|unit-resolution| (|def-axiom| (or $x70 $x181)) (hypothesis $x69) $x181) (<= (+ |x$| ?x213) 0))))
+(let (($x189 (>= ?x76 0)))
+(let (($x190 (not $x189)))
+(let (($x169 (forall ((?v0 Int) )(!(let ((?x21 (|nat$| ?v0)))
+(let ((?x22 (|of_nat$| ?x21)))
+(let (($x23 (= ?x22 ?v0)))
+(let (($x106 (>= ?v0 0)))
+(let (($x108 (not $x106)))
+(or $x108 $x23)))))) :pattern ( (|nat$| ?v0) )))
+))
+(let (($x114 (forall ((?v0 Int) )(let ((?x21 (|nat$| ?v0)))
+(let ((?x22 (|of_nat$| ?x21)))
+(let (($x23 (= ?x22 ?v0)))
+(let (($x106 (>= ?v0 0)))
+(let (($x108 (not $x106)))
+(or $x108 $x23)))))))
+))
+(let ((?x21 (|nat$| ?0)))
+(let ((?x22 (|of_nat$| ?x21)))
+(let (($x23 (= ?x22 ?0)))
+(let (($x106 (>= ?0 0)))
+(let (($x108 (not $x106)))
+(let (($x111 (or $x108 $x23)))
+(let (($x25 (forall ((?v0 Int) )(let ((?x21 (|nat$| ?v0)))
+(let ((?x22 (|of_nat$| ?x21)))
+(let (($x23 (= ?x22 ?v0)))
+(let (($x20 (<= 0 ?v0)))
+(=> $x20 $x23))))))
+))
+(let (($x102 (forall ((?v0 Int) )(let ((?x21 (|nat$| ?v0)))
+(let ((?x22 (|of_nat$| ?x21)))
+(let (($x23 (= ?x22 ?v0)))
+(or (not (<= 0 ?v0)) $x23)))))
+))
+(let ((@x110 (monotonicity (rewrite (= (<= 0 ?0) $x106)) (= (not (<= 0 ?0)) $x108))))
+(let ((@x116 (|quant-intro| (monotonicity @x110 (= (or (not (<= 0 ?0)) $x23) $x111)) (= $x102 $x114))))
+(let ((@x101 (rewrite (= (=> (<= 0 ?0) $x23) (or (not (<= 0 ?0)) $x23)))))
+(let ((@x119 (mp (asserted $x25) (trans (|quant-intro| @x101 (= $x25 $x102)) @x116 (= $x25 $x114)) $x114)))
+(let ((@x174 (mp (|mp~| @x119 (|nnf-pos| (refl (|~| $x111 $x111)) (|~| $x114 $x114)) $x114) (|quant-intro| (refl (= $x111 $x111)) (= $x114 $x169)) $x169)))
+(let ((?x81 (|nat$| ?x76)))
+(let ((?x84 (|of_nat$| ?x81)))
+(let (($x87 (= ?x84 ?x76)))
+(let (($x90 (not $x87)))
+(let (($x7 (< |x$| 0)))
+(let ((?x9 (ite $x7 (- |x$|) |x$|)))
+(let (($x13 (not (= (|of_nat$| (|nat$| ?x9)) ?x9))))
+(let (($x91 (= (not (= (|of_nat$| (|nat$| (ite $x7 ?x51 |x$|))) (ite $x7 ?x51 |x$|))) $x90)))
+(let ((?x54 (ite $x7 ?x51 |x$|)))
+(let ((?x57 (|nat$| ?x54)))
+(let ((?x60 (|of_nat$| ?x57)))
+(let (($x63 (= ?x60 ?x54)))
+(let ((@x80 (trans (monotonicity (rewrite (= $x7 $x70)) (= ?x54 (ite $x70 ?x51 |x$|))) (rewrite (= (ite $x70 ?x51 |x$|) ?x76)) (= ?x54 ?x76))))
+(let ((@x89 (monotonicity (monotonicity (monotonicity @x80 (= ?x57 ?x81)) (= ?x60 ?x84)) @x80 (= $x63 $x87))))
+(let ((@x59 (monotonicity (monotonicity (rewrite (= (- |x$|) ?x51)) (= ?x9 ?x54)) (= (|nat$| ?x9) ?x57))))
+(let ((@x65 (monotonicity (monotonicity @x59 (= (|of_nat$| (|nat$| ?x9)) ?x60)) (monotonicity (rewrite (= (- |x$|) ?x51)) (= ?x9 ?x54)) (= (= (|of_nat$| (|nat$| ?x9)) ?x9) $x63))))
+(let ((@x94 (trans (monotonicity @x65 (= $x13 (not $x63))) (monotonicity @x89 $x91) (= $x13 $x90))))
+(let ((@x95 (mp (asserted $x13) @x94 $x90)))
+(let (($x198 (or (not $x169) $x190 $x87)))
+(let ((@x203 (mp ((_ |quant-inst| (ite $x69 |x$| ?x51)) (or (not $x169) (or $x190 $x87))) (rewrite (= (or (not $x169) (or $x190 $x87)) $x198)) $x198)))
+(let ((@x224 ((_ |th-lemma| arith farkas -1 1 1) (hypothesis $x69) (|unit-resolution| @x203 @x95 @x174 $x190) @x222 false)))
+(let ((@x225 (lemma @x224 $x70)))
+(let ((@x232 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x182) $x226)) (|unit-resolution| (|def-axiom| (or $x69 $x182)) @x225 $x182) $x226)))
+(let (($x205 (<= ?x76 0)))
+(let ((@x235 (|unit-resolution| ((_ |th-lemma| arith farkas 1 1) (or $x205 $x189)) (|unit-resolution| @x203 @x95 @x174 $x190) $x205)))
+((_ |th-lemma| arith farkas 1 1 1) @x235 @x225 @x232 false)))))))))))))))))))))))))))))))))))))))))))))))))))))))
+
+777f9032b35b7e1c0dce62f13109424c73d5d094 312 0
+unsat
+((set-logic AUFLIA)
+(declare-fun ?v1!0 (|Nat$|) |Nat$|)
+(proof
+(let ((?x23 (|of_nat$| |m$|)))
+(let ((?x24 (* 4 ?x23)))
+(let ((?x110 (+ 1 ?x24)))
+(let ((?x113 (|nat$| ?x110)))
+(let ((?x344 (|of_nat$| ?x113)))
+(let ((?x491 (* (~ 1) ?x344)))
+(let ((?x492 (+ ?x24 ?x491)))
+(let (($x509 (>= ?x492 (~ 1))))
+(let (($x489 (= ?x492 (~ 1))))
+(let (($x481 (>= ?x23 0)))
+(let (($x345 (<= ?x344 1)))
+(let (($x373 (not $x345)))
+(let (($x351 (forall ((?v1 |Nat$|) )(!(let ((?x23 (|of_nat$| |m$|)))
+(let ((?x24 (* 4 ?x23)))
+(let ((?x110 (+ 1 ?x24)))
+(let ((?x113 (|nat$| ?x110)))
+(let (($x348 (= ?v1 ?x113)))
+(let ((?x12 (|nat$| 1)))
+(let (($x13 (= ?v1 ?x12)))
+(let (($x346 (|dvd$| ?v1 ?x113)))
+(let (($x347 (not $x346)))
+(or $x347 $x13 $x348)))))))))) :pattern ( (|dvd$| ?v1 (|nat$| (+ 1 (* 4 (|of_nat$| |m$|))))) )))
+))
+(let (($x352 (not $x351)))
+(let (($x353 (or $x345 $x352)))
+(let (($x354 (not $x353)))
+(let (($x116 (|prime_nat$| ?x113)))
+(let (($x122 (not $x116)))
+(let (($x355 (or $x122 $x354)))
+(let ((?x357 (?v1!0 ?x113)))
+(let (($x361 (= ?x357 ?x113)))
+(let ((?x12 (|nat$| 1)))
+(let (($x360 (= ?x357 ?x12)))
+(let (($x358 (|dvd$| ?x357 ?x113)))
+(let (($x359 (not $x358)))
+(let (($x362 (or $x359 $x360 $x361)))
+(let (($x363 (not $x362)))
+(let (($x364 (or $x116 $x345 $x363)))
+(let (($x365 (not $x364)))
+(let (($x356 (not $x355)))
+(let (($x366 (or $x356 $x365)))
+(let (($x367 (not $x366)))
+(let (($x321 (forall ((?v0 |Nat$|) )(!(let (($x217 (or (not (|dvd$| (?v1!0 ?v0) ?v0)) (= (?v1!0 ?v0) (|nat$| 1)) (= (?v1!0 ?v0) ?v0))))
+(let (($x218 (not $x217)))
+(let ((?x8 (|of_nat$| ?v0)))
+(let (($x87 (<= ?x8 1)))
+(let (($x6 (|prime_nat$| ?v0)))
+(let (($x245 (or $x6 $x87 $x218)))
+(let (($x293 (forall ((?v1 |Nat$|) )(!(let ((?x12 (|nat$| 1)))
+(let (($x13 (= ?v1 ?x12)))
+(or (not (|dvd$| ?v1 ?v0)) $x13 (= ?v1 ?v0)))) :pattern ( (|dvd$| ?v1 ?v0) )))
+))
+(let (($x198 (not $x6)))
+(not (or (not (or $x198 (not (or $x87 (not $x293))))) (not $x245))))))))))) :pattern ( (|prime_nat$| ?v0) ) :pattern ( (|of_nat$| ?v0) )))
+))
+(let (($x288 (forall ((?v0 |Nat$|) )(let (($x217 (or (not (|dvd$| (?v1!0 ?v0) ?v0)) (= (?v1!0 ?v0) (|nat$| 1)) (= (?v1!0 ?v0) ?v0))))
+(let (($x218 (not $x217)))
+(let ((?x8 (|of_nat$| ?v0)))
+(let (($x87 (<= ?x8 1)))
+(let (($x6 (|prime_nat$| ?v0)))
+(let (($x245 (or $x6 $x87 $x218)))
+(let (($x94 (forall ((?v1 |Nat$|) )(let ((?x12 (|nat$| 1)))
+(let (($x13 (= ?v1 ?x12)))
+(or (not (|dvd$| ?v1 ?v0)) $x13 (= ?v1 ?v0)))))
+))
+(let (($x219 (not $x94)))
+(let (($x271 (not (or $x87 $x219))))
+(let (($x198 (not $x6)))
+(let (($x274 (or $x198 $x271)))
+(not (or (not $x274) (not $x245)))))))))))))))
+))
+(let (($x217 (or (not (|dvd$| (?v1!0 ?0) ?0)) (= (?v1!0 ?0) ?x12) (= (?v1!0 ?0) ?0))))
+(let (($x218 (not $x217)))
+(let ((?x8 (|of_nat$| ?0)))
+(let (($x87 (<= ?x8 1)))
+(let (($x6 (|prime_nat$| ?0)))
+(let (($x245 (or $x6 $x87 $x218)))
+(let (($x293 (forall ((?v1 |Nat$|) )(!(let ((?x12 (|nat$| 1)))
+(let (($x13 (= ?v1 ?x12)))
+(or (not (|dvd$| ?v1 ?0)) $x13 (= ?v1 ?0)))) :pattern ( (|dvd$| ?v1 ?0) )))
+))
+(let (($x198 (not $x6)))
+(let (($x94 (forall ((?v1 |Nat$|) )(let ((?x12 (|nat$| 1)))
+(let (($x13 (= ?v1 ?x12)))
+(or (not (|dvd$| ?v1 ?0)) $x13 (= ?v1 ?0)))))
+))
+(let (($x219 (not $x94)))
+(let (($x271 (not (or $x87 $x219))))
+(let (($x274 (or $x198 $x271)))
+(let (($x283 (not (or (not $x274) (not $x245)))))
+(let (($x317 (= $x283 (not (or (not (or $x198 (not (or $x87 (not $x293))))) (not $x245))))))
+(let (($x314 (= (or (not $x274) (not $x245)) (or (not (or $x198 (not (or $x87 (not $x293))))) (not $x245)))))
+(let (($x13 (= ?0 ?x12)))
+(let (($x91 (or (not (|dvd$| ?0 ?1)) $x13 (= ?0 ?1))))
+(let ((@x300 (monotonicity (|quant-intro| (refl (= $x91 $x91)) (= $x94 $x293)) (= $x219 (not $x293)))))
+(let ((@x306 (monotonicity (monotonicity @x300 (= (or $x87 $x219) (or $x87 (not $x293)))) (= $x271 (not (or $x87 (not $x293)))))))
+(let ((@x312 (monotonicity (monotonicity @x306 (= $x274 (or $x198 (not (or $x87 (not $x293)))))) (= (not $x274) (not (or $x198 (not (or $x87 (not $x293)))))))))
+(let ((@x323 (|quant-intro| (monotonicity (monotonicity @x312 $x314) $x317) (= $x288 $x321))))
+(let (($x253 (forall ((?v0 |Nat$|) )(let (($x217 (or (not (|dvd$| (?v1!0 ?v0) ?v0)) (= (?v1!0 ?v0) (|nat$| 1)) (= (?v1!0 ?v0) ?v0))))
+(let (($x218 (not $x217)))
+(let ((?x8 (|of_nat$| ?v0)))
+(let (($x87 (<= ?x8 1)))
+(let (($x6 (|prime_nat$| ?v0)))
+(let (($x245 (or $x6 $x87 $x218)))
+(let (($x94 (forall ((?v1 |Nat$|) )(let ((?x12 (|nat$| 1)))
+(let (($x13 (= ?v1 ?x12)))
+(or (not (|dvd$| ?v1 ?v0)) $x13 (= ?v1 ?v0)))))
+))
+(let (($x88 (not $x87)))
+(let (($x97 (and $x88 $x94)))
+(let (($x198 (not $x6)))
+(let (($x227 (or $x198 $x97)))
+(and $x227 $x245)))))))))))))
+))
+(let ((@x276 (monotonicity (rewrite (= (and (not $x87) $x94) $x271)) (= (or $x198 (and (not $x87) $x94)) $x274))))
+(let ((@x279 (monotonicity @x276 (= (and (or $x198 (and (not $x87) $x94)) $x245) (and $x274 $x245)))))
+(let ((@x287 (trans @x279 (rewrite (= (and $x274 $x245) $x283)) (= (and (or $x198 (and (not $x87) $x94)) $x245) $x283))))
+(let (($x231 (forall ((?v0 |Nat$|) )(let (($x217 (or (not (|dvd$| (?v1!0 ?v0) ?v0)) (= (?v1!0 ?v0) (|nat$| 1)) (= (?v1!0 ?v0) ?v0))))
+(let (($x218 (not $x217)))
+(let ((?x8 (|of_nat$| ?v0)))
+(let (($x87 (<= ?x8 1)))
+(let (($x88 (not $x87)))
+(let (($x209 (not $x88)))
+(let (($x222 (or $x209 $x218)))
+(let (($x6 (|prime_nat$| ?v0)))
+(let (($x226 (or $x6 $x222)))
+(let (($x94 (forall ((?v1 |Nat$|) )(let ((?x12 (|nat$| 1)))
+(let (($x13 (= ?v1 ?x12)))
+(or (not (|dvd$| ?v1 ?v0)) $x13 (= ?v1 ?v0)))))
+))
+(let (($x97 (and $x88 $x94)))
+(let (($x198 (not $x6)))
+(let (($x227 (or $x198 $x97)))
+(and $x227 $x226)))))))))))))))
+))
+(let (($x88 (not $x87)))
+(let (($x97 (and $x88 $x94)))
+(let (($x227 (or $x198 $x97)))
+(let (($x250 (and $x227 $x245)))
+(let (($x209 (not $x88)))
+(let (($x222 (or $x209 $x218)))
+(let (($x226 (or $x6 $x222)))
+(let (($x228 (and $x227 $x226)))
+(let ((@x244 (monotonicity (monotonicity (rewrite (= $x209 $x87)) (= $x222 (or $x87 $x218))) (= $x226 (or $x6 (or $x87 $x218))))))
+(let ((@x249 (trans @x244 (rewrite (= (or $x6 (or $x87 $x218)) $x245)) (= $x226 $x245))))
+(let (($x103 (forall ((?v0 |Nat$|) )(let (($x94 (forall ((?v1 |Nat$|) )(let ((?x12 (|nat$| 1)))
+(let (($x13 (= ?v1 ?x12)))
+(or (not (|dvd$| ?v1 ?v0)) $x13 (= ?v1 ?v0)))))
+))
+(let ((?x8 (|of_nat$| ?v0)))
+(let (($x87 (<= ?x8 1)))
+(let (($x88 (not $x87)))
+(let (($x97 (and $x88 $x94)))
+(let (($x6 (|prime_nat$| ?v0)))
+(= $x6 $x97))))))))
+))
+(let ((@x225 (|nnf-neg| (refl (|~| $x209 $x209)) (sk (|~| $x219 $x218)) (|~| (not $x97) $x222))))
+(let ((@x208 (monotonicity (refl (|~| $x88 $x88)) (|nnf-pos| (refl (|~| $x91 $x91)) (|~| $x94 $x94)) (|~| $x97 $x97))))
+(let ((@x230 (|nnf-pos| (refl (|~| $x6 $x6)) (refl (|~| $x198 $x198)) @x208 @x225 (|~| (= $x6 $x97) $x228))))
+(let (($x20 (forall ((?v0 |Nat$|) )(let (($x17 (forall ((?v1 |Nat$|) )(let (($x11 (|dvd$| ?v1 ?v0)))
+(=> $x11 (or (= ?v1 (|nat$| 1)) (= ?v1 ?v0)))))
+))
+(let ((?x8 (|of_nat$| ?v0)))
+(let (($x9 (< 1 ?x8)))
+(let (($x6 (|prime_nat$| ?v0)))
+(= $x6 (and $x9 $x17)))))))
+))
+(let (($x84 (forall ((?v0 |Nat$|) )(let (($x70 (forall ((?v1 |Nat$|) )(or (not (|dvd$| ?v1 ?v0)) (or (= ?v1 (|nat$| 1)) (= ?v1 ?v0))))
+))
+(let ((?x8 (|of_nat$| ?v0)))
+(let (($x9 (< 1 ?x8)))
+(let (($x73 (and $x9 $x70)))
+(let (($x6 (|prime_nat$| ?v0)))
+(= $x6 $x73)))))))
+))
+(let (($x100 (= $x6 $x97)))
+(let (($x70 (forall ((?v1 |Nat$|) )(or (not (|dvd$| ?v1 ?0)) (or (= ?v1 (|nat$| 1)) (= ?v1 ?0))))
+))
+(let (($x9 (< 1 ?x8)))
+(let (($x73 (and $x9 $x70)))
+(let (($x79 (= $x6 $x73)))
+(let ((@x93 (rewrite (= (or (not (|dvd$| ?0 ?1)) (or $x13 (= ?0 ?1))) $x91))))
+(let ((@x99 (monotonicity (rewrite (= $x9 $x88)) (|quant-intro| @x93 (= $x70 $x94)) (= $x73 $x97))))
+(let (($x17 (forall ((?v1 |Nat$|) )(let (($x11 (|dvd$| ?v1 ?0)))
+(=> $x11 (or (= ?v1 (|nat$| 1)) (= ?v1 ?0)))))
+))
+(let (($x19 (= $x6 (and $x9 $x17))))
+(let (($x67 (or (not (|dvd$| ?0 ?1)) (or $x13 (= ?0 ?1)))))
+(let ((@x72 (|quant-intro| (rewrite (= (=> (|dvd$| ?0 ?1) (or $x13 (= ?0 ?1))) $x67)) (= $x17 $x70))))
+(let ((@x78 (monotonicity (monotonicity @x72 (= (and $x9 $x17) $x73)) (= $x19 (= $x6 $x73)))))
+(let ((@x86 (|quant-intro| (trans @x78 (rewrite (= (= $x6 $x73) $x79)) (= $x19 $x79)) (= $x20 $x84))))
+(let ((@x107 (trans @x86 (|quant-intro| (monotonicity @x99 (= $x79 $x100)) (= $x84 $x103)) (= $x20 $x103))))
+(let ((@x234 (|mp~| (mp (asserted $x20) @x107 $x103) (|nnf-pos| @x230 (|~| $x103 $x231)) $x231)))
+(let ((@x235 (mp @x234 (|quant-intro| (monotonicity @x249 (= $x228 $x250)) (= $x231 $x253)) $x253)))
+(let ((@x324 (mp (mp @x235 (|quant-intro| @x287 (= $x253 $x288)) $x288) @x323 $x321)))
+(let (($x371 (or (not $x321) $x367)))
+(let ((@x372 ((_ |quant-inst| (|nat$| ?x110)) $x371)))
+(let ((@x530 (|unit-resolution| (|def-axiom| (or $x366 $x355)) (|unit-resolution| @x372 @x324 $x367) $x355)))
+(let (($x137 (not (or $x122 (>= ?x23 1)))))
+(let (($x28 (<= 1 ?x23)))
+(let (($x29 (=> (|prime_nat$| (|nat$| (+ ?x24 1))) $x28)))
+(let (($x30 (not $x29)))
+(let ((@x136 (monotonicity (rewrite (= $x28 (>= ?x23 1))) (= (or $x122 $x28) (or $x122 (>= ?x23 1))))))
+(let ((@x115 (monotonicity (rewrite (= (+ ?x24 1) ?x110)) (= (|nat$| (+ ?x24 1)) ?x113))))
+(let ((@x121 (monotonicity (monotonicity @x115 (= (|prime_nat$| (|nat$| (+ ?x24 1))) $x116)) (= $x29 (=> $x116 $x28)))))
+(let ((@x127 (trans @x121 (rewrite (= (=> $x116 $x28) (or $x122 $x28))) (= $x29 (or $x122 $x28)))))
+(let ((@x141 (trans (monotonicity @x127 (= $x30 (not (or $x122 $x28)))) (monotonicity @x136 (= (not (or $x122 $x28)) $x137)) (= $x30 $x137))))
+(let ((@x143 (|not-or-elim| (mp (asserted $x30) @x141 $x137) $x116)))
+(let ((@x533 (|unit-resolution| (|unit-resolution| (|def-axiom| (or $x356 $x122 $x354)) @x143 (or $x356 $x354)) @x530 $x354)))
+(let ((@x538 (|unit-resolution| ((_ |th-lemma| arith farkas 1 1) (or (not (<= ?x344 0)) $x345)) (|unit-resolution| (|def-axiom| (or $x353 $x373)) @x533 $x373) (not (<= ?x344 0)))))
+(let ((@x542 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not (= ?x344 0)) (<= ?x344 0))) @x538 (not (= ?x344 0)))))
+(let (($x510 (= ?x344 0)))
+(let (($x512 (or $x481 $x510)))
+(let (($x338 (forall ((?v0 Int) )(!(let ((?x37 (|nat$| ?v0)))
+(let ((?x38 (|of_nat$| ?x37)))
+(let (($x43 (= ?x38 0)))
+(let (($x157 (>= ?v0 0)))
+(or $x157 $x43))))) :pattern ( (|nat$| ?v0) )))
+))
+(let (($x190 (forall ((?v0 Int) )(let ((?x37 (|nat$| ?v0)))
+(let ((?x38 (|of_nat$| ?x37)))
+(let (($x43 (= ?x38 0)))
+(let (($x157 (>= ?v0 0)))
+(or $x157 $x43))))))
+))
+(let ((?x37 (|nat$| ?0)))
+(let ((?x38 (|of_nat$| ?x37)))
+(let (($x43 (= ?x38 0)))
+(let (($x157 (>= ?0 0)))
+(let (($x187 (or $x157 $x43)))
+(let (($x45 (forall ((?v0 Int) )(let ((?x37 (|nat$| ?v0)))
+(let ((?x38 (|of_nat$| ?x37)))
+(let (($x43 (= ?x38 0)))
+(let (($x42 (< ?v0 0)))
+(=> $x42 $x43))))))
+))
+(let (($x175 (forall ((?v0 Int) )(let ((?x37 (|nat$| ?v0)))
+(let ((?x38 (|of_nat$| ?x37)))
+(let (($x43 (= ?x38 0)))
+(let (($x42 (< ?v0 0)))
+(let (($x171 (not $x42)))
+(or $x171 $x43)))))))
+))
+(let ((@x182 (monotonicity (rewrite (= (< ?0 0) (not $x157))) (= (not (< ?0 0)) (not (not $x157))))))
+(let ((@x186 (trans @x182 (rewrite (= (not (not $x157)) $x157)) (= (not (< ?0 0)) $x157))))
+(let ((@x192 (|quant-intro| (monotonicity @x186 (= (or (not (< ?0 0)) $x43) $x187)) (= $x175 $x190))))
+(let ((@x174 (rewrite (= (=> (< ?0 0) $x43) (or (not (< ?0 0)) $x43)))))
+(let ((@x195 (mp (asserted $x45) (trans (|quant-intro| @x174 (= $x45 $x175)) @x192 (= $x45 $x190)) $x190)))
+(let ((@x343 (mp (|mp~| @x195 (|nnf-pos| (refl (|~| $x187 $x187)) (|~| $x190 $x190)) $x190) (|quant-intro| (refl (= $x187 $x187)) (= $x190 $x338)) $x338)))
+(let (($x515 (not $x338)))
+(let (($x516 (or $x515 $x481 $x510)))
+(let ((@x483 (rewrite (= (>= ?x110 0) $x481))))
+(let ((@x514 (monotonicity @x483 (= (or (>= ?x110 0) $x510) $x512))))
+(let ((@x521 (monotonicity @x514 (= (or $x515 (or (>= ?x110 0) $x510)) (or $x515 $x512)))))
+(let ((@x525 (trans @x521 (rewrite (= (or $x515 $x512) $x516)) (= (or $x515 (or (>= ?x110 0) $x510)) $x516))))
+(let ((@x526 (mp ((_ |quant-inst| (+ 1 ?x24)) (or $x515 (or (>= ?x110 0) $x510))) @x525 $x516)))
+(let (($x484 (not $x481)))
+(let (($x493 (or $x484 $x489)))
+(let (($x332 (forall ((?v0 Int) )(!(let ((?x37 (|nat$| ?v0)))
+(let ((?x38 (|of_nat$| ?x37)))
+(let (($x39 (= ?x38 ?v0)))
+(let (($x157 (>= ?v0 0)))
+(let (($x158 (not $x157)))
+(or $x158 $x39)))))) :pattern ( (|nat$| ?v0) )))
+))
+(let (($x164 (forall ((?v0 Int) )(let ((?x37 (|nat$| ?v0)))
+(let ((?x38 (|of_nat$| ?x37)))
+(let (($x39 (= ?x38 ?v0)))
+(let (($x157 (>= ?v0 0)))
+(let (($x158 (not $x157)))
+(or $x158 $x39)))))))
+))
+(let ((@x334 (refl (= (or (not $x157) (= ?x38 ?0)) (or (not $x157) (= ?x38 ?0))))))
+(let ((@x261 (refl (|~| (or (not $x157) (= ?x38 ?0)) (or (not $x157) (= ?x38 ?0))))))
+(let (($x41 (forall ((?v0 Int) )(let ((?x37 (|nat$| ?v0)))
+(let ((?x38 (|of_nat$| ?x37)))
+(let (($x39 (= ?x38 ?v0)))
+(let (($x36 (<= 0 ?v0)))
+(=> $x36 $x39))))))
+))
+(let (($x152 (forall ((?v0 Int) )(let ((?x37 (|nat$| ?v0)))
+(let ((?x38 (|of_nat$| ?x37)))
+(let (($x39 (= ?x38 ?v0)))
+(or (not (<= 0 ?v0)) $x39)))))
+))
+(let (($x39 (= ?x38 ?0)))
+(let (($x158 (not $x157)))
+(let (($x161 (or $x158 $x39)))
+(let (($x149 (or (not (<= 0 ?0)) $x39)))
+(let ((@x160 (monotonicity (rewrite (= (<= 0 ?0) $x157)) (= (not (<= 0 ?0)) $x158))))
+(let ((@x154 (|quant-intro| (rewrite (= (=> (<= 0 ?0) $x39) $x149)) (= $x41 $x152))))
+(let ((@x168 (trans @x154 (|quant-intro| (monotonicity @x160 (= $x149 $x161)) (= $x152 $x164)) (= $x41 $x164))))
+(let ((@x264 (|mp~| (mp (asserted $x41) @x168 $x164) (|nnf-pos| @x261 (|~| $x164 $x164)) $x164)))
+(let ((@x337 (mp @x264 (|quant-intro| @x334 (= $x164 $x332)) $x332)))
+(let (($x496 (not $x332)))
+(let (($x497 (or $x496 $x484 $x489)))
+(let (($x479 (= ?x344 ?x110)))
+(let (($x474 (>= ?x110 0)))
+(let (($x475 (not $x474)))
+(let (($x480 (or $x475 $x479)))
+(let (($x498 (or $x496 $x480)))
+(let ((@x495 (monotonicity (monotonicity @x483 (= $x475 $x484)) (rewrite (= $x479 $x489)) (= $x480 $x493))))
+(let ((@x506 (trans (monotonicity @x495 (= $x498 (or $x496 $x493))) (rewrite (= (or $x496 $x493) $x497)) (= $x498 $x497))))
+(let ((@x507 (mp ((_ |quant-inst| (+ 1 ?x24)) $x498) @x506 $x497)))
+(let ((@x546 (|unit-resolution| (|unit-resolution| @x507 @x337 $x493) (|unit-resolution| (|unit-resolution| @x526 @x343 $x512) @x542 $x481) $x489)))
+(let ((@x145 (|not-or-elim| (mp (asserted $x30) @x141 $x137) (not (>= ?x23 1)))))
+((_ |th-lemma| arith farkas -4 1 1) @x145 (|unit-resolution| (|def-axiom| (or $x353 $x373)) @x533 $x373) (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x489) $x509)) @x546 $x509) false)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))
+
+c7efedca31e5e8360d3c81014f43c447bc784df3 23 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let (($x17 (= |x$| |a$|)))
+(let ((?x13 (|pair$| |x$| |y$|)))
+(let ((?x14 (|fst$| ?x13)))
+(let (($x16 (= ?x14 |a$|)))
+(let ((@x48 (monotonicity (rewrite (= (=> $x16 $x17) (or (not $x16) $x17))) (= (not (=> $x16 $x17)) (not (or (not $x16) $x17))))))
+(let ((@x49 (|not-or-elim| (mp (asserted (not (=> $x16 $x17))) @x48 (not (or (not $x16) $x17))) $x16)))
+(let (($x65 (= ?x14 |x$|)))
+(let (($x59 (forall ((?v0 |A$|) (?v1 |B$|) )(!(= (|fst$| (|pair$| ?v0 ?v1)) ?v0) :pattern ( (|pair$| ?v0 ?v1) )))
+))
+(let (($x10 (forall ((?v0 |A$|) (?v1 |B$|) )(= (|fst$| (|pair$| ?v0 ?v1)) ?v0))
+))
+(let (($x9 (= (|fst$| (|pair$| ?1 ?0)) ?1)))
+(let ((@x57 (|mp~| (asserted $x10) (|nnf-pos| (refl (|~| $x9 $x9)) (|~| $x10 $x10)) $x10)))
+(let ((@x64 (mp @x57 (|quant-intro| (refl (= $x9 $x9)) (= $x10 $x59)) $x59)))
+(let (($x69 (or (not $x59) $x65)))
+(let ((@x70 ((_ |quant-inst| |x$| |y$|) $x69)))
+(let ((@x72 (trans (symm (|unit-resolution| @x70 @x64 $x65) (= |x$| ?x14)) @x49 $x17)))
+(let ((@x52 (|not-or-elim| (mp (asserted (not (=> $x16 $x17))) @x48 (not (or (not $x16) $x17))) (not $x17))))
+(|unit-resolution| @x52 @x72 false)))))))))))))))))))
+
+aea156a69bb23683148bfccfaa255f874937bce0 42 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let ((?x33 (|snd$a| |p2$|)))
+(let ((?x32 (|fst$a| |p1$|)))
+(let (($x34 (= ?x32 ?x33)))
+(let ((?x29 (|pair$| |y$| |x$|)))
+(let (($x30 (= |p2$| ?x29)))
+(let ((?x26 (|pair$a| |x$| |y$|)))
+(let (($x27 (= |p1$| ?x26)))
+(let (($x31 (and $x27 $x30)))
+(let ((@x68 (monotonicity (rewrite (= (=> $x31 $x34) (or (not $x31) $x34))) (= (not (=> $x31 $x34)) (not (or (not $x31) $x34))))))
+(let ((@x69 (|not-or-elim| (mp (asserted (not (=> $x31 $x34))) @x68 (not (or (not $x31) $x34))) $x31)))
+(let ((@x72 (|and-elim| @x69 $x30)))
+(let ((@x150 (symm (monotonicity @x72 (= ?x33 (|snd$a| ?x29))) (= (|snd$a| ?x29) ?x33))))
+(let ((?x123 (|snd$a| ?x29)))
+(let (($x124 (= ?x123 |x$|)))
+(let (($x115 (forall ((?v0 |B$|) (?v1 |A$|) )(!(= (|snd$a| (|pair$| ?v0 ?v1)) ?v1) :pattern ( (|pair$| ?v0 ?v1) )))
+))
+(let (($x22 (forall ((?v0 |B$|) (?v1 |A$|) )(= (|snd$a| (|pair$| ?v0 ?v1)) ?v1))
+))
+(let (($x21 (= (|snd$a| (|pair$| ?1 ?0)) ?0)))
+(let ((@x94 (|mp~| (asserted $x22) (|nnf-pos| (refl (|~| $x21 $x21)) (|~| $x22 $x22)) $x22)))
+(let ((@x120 (mp @x94 (|quant-intro| (refl (= $x21 $x21)) (= $x22 $x115)) $x115)))
+(let (($x131 (or (not $x115) $x124)))
+(let ((@x132 ((_ |quant-inst| |y$| |x$|) $x131)))
+(let ((?x128 (|fst$a| ?x26)))
+(let (($x129 (= ?x128 |x$|)))
+(let (($x103 (forall ((?v0 |A$|) (?v1 |B$|) )(!(= (|fst$a| (|pair$a| ?v0 ?v1)) ?v0) :pattern ( (|pair$a| ?v0 ?v1) )))
+))
+(let (($x16 (forall ((?v0 |A$|) (?v1 |B$|) )(= (|fst$a| (|pair$a| ?v0 ?v1)) ?v0))
+))
+(let (($x15 (= (|fst$a| (|pair$a| ?1 ?0)) ?1)))
+(let ((@x84 (|mp~| (asserted $x16) (|nnf-pos| (refl (|~| $x15 $x15)) (|~| $x16 $x16)) $x16)))
+(let ((@x108 (mp @x84 (|quant-intro| (refl (= $x15 $x15)) (= $x16 $x103)) $x103)))
+(let (($x136 (or (not $x103) $x129)))
+(let ((@x137 ((_ |quant-inst| |x$| |y$|) $x136)))
+(let ((@x152 (trans (monotonicity (|and-elim| @x69 $x27) (= ?x32 ?x128)) (|unit-resolution| @x137 @x108 $x129) (= ?x32 |x$|))))
+(let ((@x154 (trans @x152 (symm (|unit-resolution| @x132 @x120 $x124) (= |x$| ?x123)) (= ?x32 ?x123))))
+(let ((@x74 (|not-or-elim| (mp (asserted (not (=> $x31 $x34))) @x68 (not (or (not $x31) $x34))) (not $x34))))
+(|unit-resolution| @x74 (trans @x154 @x150 $x34) false))))))))))))))))))))))))))))))))))))
+
+9e587b4eedc0dc25b019cb54b54b4a4e643bf93e 49 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let ((?x45 (|fun_app$| |f$| |i$|)))
+(let ((?x36 (|fun_upd$| |f$|)))
+(let ((?x37 (|fun_app$b| ?x36 |i1$|)))
+(let ((?x39 (|fun_app$a| ?x37 |v1$|)))
+(let ((?x40 (|fun_upd$| ?x39)))
+(let ((?x41 (|fun_app$b| ?x40 |i2$|)))
+(let ((?x43 (|fun_app$a| ?x41 |v2$|)))
+(let ((?x44 (|fun_app$| ?x43 |i$|)))
+(let (($x46 (= ?x44 ?x45)))
+(let (($x29 (= |i$| |i1$|)))
+(let ((?x178 (ite $x29 |v1$| ?x45)))
+(let (($x185 (= ?x45 ?x178)))
+(let (($x30 (not $x29)))
+(let (($x32 (= |i$| |i2$|)))
+(let (($x33 (not $x32)))
+(let (($x34 (and $x30 $x33)))
+(let ((@x78 (monotonicity (rewrite (= (=> $x34 $x46) (or (not $x34) $x46))) (= (not (=> $x34 $x46)) (not (or (not $x34) $x46))))))
+(let ((@x79 (|not-or-elim| (mp (asserted (not (=> $x34 $x46))) @x78 (not (or (not $x34) $x46))) $x34)))
+(let ((@x80 (|and-elim| @x79 $x30)))
+(let ((@x197 (symm (|unit-resolution| (|def-axiom| (or $x29 $x185)) @x80 $x185) (= ?x178 ?x45))))
+(let ((?x116 (|fun_app$| ?x39 |i$|)))
+(let (($x179 (= ?x116 ?x178)))
+(let (($x103 (forall ((?v0 |A_b_fun$|) (?v1 |A$|) (?v2 |B$|) (?v3 |A$|) )(!(let ((?x21 (|fun_app$| (|fun_app$a| (|fun_app$b| (|fun_upd$| ?v0) ?v1) ?v2) ?v3)))
+(= ?x21 (ite (= ?v3 ?v1) ?v2 (|fun_app$| ?v0 ?v3)))) :pattern ( (|fun_app$| (|fun_app$a| (|fun_app$b| (|fun_upd$| ?v0) ?v1) ?v2) ?v3) )))
+))
+(let (($x26 (forall ((?v0 |A_b_fun$|) (?v1 |A$|) (?v2 |B$|) (?v3 |A$|) )(let ((?x21 (|fun_app$| (|fun_app$a| (|fun_app$b| (|fun_upd$| ?v0) ?v1) ?v2) ?v3)))
+(= ?x21 (ite (= ?v3 ?v1) ?v2 (|fun_app$| ?v0 ?v3)))))
+))
+(let ((?x21 (|fun_app$| (|fun_app$a| (|fun_app$b| (|fun_upd$| ?3) ?2) ?1) ?0)))
+(let (($x25 (= ?x21 (ite (= ?0 ?2) ?1 (|fun_app$| ?3 ?0)))))
+(let ((@x94 (|mp~| (asserted $x26) (|nnf-pos| (refl (|~| $x25 $x25)) (|~| $x26 $x26)) $x26)))
+(let ((@x108 (mp @x94 (|quant-intro| (refl (= $x25 $x25)) (= $x26 $x103)) $x103)))
+(let (($x123 (not $x103)))
+(let (($x182 (or $x123 $x179)))
+(let ((@x183 ((_ |quant-inst| |f$| |i1$| |v1$| |i$|) $x182)))
+(let ((?x117 (ite $x32 |v2$| ?x116)))
+(let (($x127 (= ?x116 ?x117)))
+(let ((@x82 (|and-elim| @x79 $x33)))
+(let ((@x195 (symm (|unit-resolution| (|def-axiom| (or $x32 $x127)) @x82 $x127) (= ?x117 ?x116))))
+(let (($x120 (= ?x44 ?x117)))
+(let (($x124 (or $x123 $x120)))
+(let ((@x125 ((_ |quant-inst| (|fun_app$a| ?x37 |v1$|) |i2$| |v2$| |i$|) $x124)))
+(let ((@x201 (trans (trans (|unit-resolution| @x125 @x108 $x120) @x195 (= ?x44 ?x116)) (|unit-resolution| @x183 @x108 $x179) (= ?x44 ?x178))))
+(let ((@x84 (|not-or-elim| (mp (asserted (not (=> $x34 $x46))) @x78 (not (or (not $x34) $x46))) (not $x46))))
+(|unit-resolution| @x84 (trans @x201 @x197 $x46) false)))))))))))))))))))))))))))))))))))))))))))
+
+b3ae8e1fe1d7b019d0bef97ff09cdb8e0a1cd7dd 25 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let (($x7 (|f$| |g$| |x$|)))
+(let (($x71 (not $x7)))
+(let (($x63 (not (or (= $x7 (|fun_app$| |g$| |x$|)) $x7 (|fun_app$| |g$| |x$|)))))
+(let (($x10 (= $x7 (and (|fun_app$| |g$| |x$|) true))))
+(let (($x15 (not (or $x10 (or (= $x7 true) (= (|fun_app$| |g$| |x$|) true))))))
+(let (($x8 (|fun_app$| |g$| |x$|)))
+(let (($x51 (or $x7 $x8)))
+(let (($x42 (= $x7 $x8)))
+(let (($x54 (or $x42 $x51)))
+(let ((@x65 (monotonicity (rewrite (= $x54 (or $x42 $x7 $x8))) (= (not $x54) $x63))))
+(let ((@x53 (monotonicity (rewrite (= (= $x7 true) $x7)) (rewrite (= (= $x8 true) $x8)) (= (or (= $x7 true) (= $x8 true)) $x51))))
+(let ((@x41 (monotonicity (rewrite (= (and $x8 true) $x8)) (= $x10 (= $x7 $x8)))))
+(let ((@x56 (monotonicity (trans @x41 (rewrite (= (= $x7 $x8) $x42)) (= $x10 $x42)) @x53 (= (or $x10 (or (= $x7 true) (= $x8 true))) $x54))))
+(let ((@x67 (trans (monotonicity @x56 (= $x15 (not $x54))) @x65 (= $x15 $x63))))
+(let ((@x68 (mp (asserted $x15) @x67 $x63)))
+(let ((@x72 (|not-or-elim| @x68 $x71)))
+(let (($x73 (not $x8)))
+(let ((@x74 (|not-or-elim| @x68 $x73)))
+(let (($x75 (= $x71 $x8)))
+(let ((@x77 (mp (|not-or-elim| @x68 (not $x42)) (rewrite (= (not $x42) $x75)) $x75)))
+(|unit-resolution| (|unit-resolution| (|def-axiom| (or $x7 $x8 (not $x75))) @x77 $x51) @x74 @x72 false)))))))))))))))))))))))
+
+d9e693c8b48e2988c493bb1e4e83656e750403bf 14 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let (($x7 (exists ((?v0 |A$|) )(|g$| ?v0))
+))
+(let (($x8 (ite $x7 true false)))
+(let (($x9 (|f$| $x8)))
+(let (($x10 (=> $x9 true)))
+(let (($x11 (not $x10)))
+(let ((@x40 (monotonicity (monotonicity (rewrite (= $x8 $x7)) (= $x9 (|f$| $x7))) (= $x10 (=> (|f$| $x7) true)))))
+(let ((@x44 (trans @x40 (rewrite (= (=> (|f$| $x7) true) true)) (= $x10 true))))
+(let ((@x51 (trans (monotonicity @x44 (= $x11 (not true))) (rewrite (= (not true) false)) (= $x11 false))))
+(mp (asserted $x11) @x51 false)))))))))))
+
+389aa7c628b5a2215f1c34b1b3aea4f4becc378c 14 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let (($x7 (forall ((?v0 |A$|) )(|g$| ?v0))
+))
+(let (($x8 (ite $x7 true false)))
+(let (($x9 (|f$| $x8)))
+(let (($x10 (=> $x9 true)))
+(let (($x11 (not $x10)))
+(let ((@x40 (monotonicity (monotonicity (rewrite (= $x8 $x7)) (= $x9 (|f$| $x7))) (= $x10 (=> (|f$| $x7) true)))))
+(let ((@x44 (trans @x40 (rewrite (= (=> (|f$| $x7) true) true)) (= $x10 true))))
+(let ((@x51 (trans (monotonicity @x44 (= $x11 (not true))) (rewrite (= (not true) false)) (= $x11 false))))
+(mp (asserted $x11) @x51 false)))))))))))
+
+53b477f55537542d72fa148413e684cc3ff42e5b 46 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let ((?x17 (|fun_app$a| |le$| 3)))
+(let (($x19 (|fun_app$| ?x17 42)))
+(let (($x73 (not $x19)))
+(let (($x15 (= |le$| |uu$|)))
+(let ((@x71 (monotonicity (rewrite (= (=> $x15 $x19) (or (not $x15) $x19))) (= (not (=> $x15 $x19)) (not (or (not $x15) $x19))))))
+(let ((@x72 (|not-or-elim| (mp (asserted (not (=> $x15 $x19))) @x71 (not (or (not $x15) $x19))) $x15)))
+(let ((@x126 (monotonicity (symm @x72 (= |uu$| |le$|)) (= (|fun_app$a| |uu$| 3) ?x17))))
+(let ((@x130 (symm (monotonicity @x126 (= (|fun_app$| (|fun_app$a| |uu$| 3) 42) $x19)) (= $x19 (|fun_app$| (|fun_app$a| |uu$| 3) 42)))))
+(let ((@x133 (monotonicity @x130 (= $x73 (not (|fun_app$| (|fun_app$a| |uu$| 3) 42))))))
+(let ((@x75 (|not-or-elim| (mp (asserted (not (=> $x15 $x19))) @x71 (not (or (not $x15) $x19))) $x73)))
+(let ((?x81 (|fun_app$a| |uu$| 3)))
+(let (($x82 (|fun_app$| ?x81 42)))
+(let (($x58 (forall ((?v0 Int) (?v1 Int) )(!(let (($x52 (<= (+ ?v0 (* (~ 1) ?v1)) 0)))
+(let (($x9 (|fun_app$| (|fun_app$a| |uu$| ?v0) ?v1)))
+(= $x9 $x52))) :pattern ( (|fun_app$| (|fun_app$a| |uu$| ?v0) ?v1) )))
+))
+(let (($x52 (<= (+ ?1 (* (~ 1) ?0)) 0)))
+(let (($x9 (|fun_app$| (|fun_app$a| |uu$| ?1) ?0)))
+(let (($x55 (= $x9 $x52)))
+(let (($x13 (forall ((?v0 Int) (?v1 Int) )(!(let (($x10 (<= ?v0 ?v1)))
+(let (($x9 (|fun_app$| (|fun_app$a| |uu$| ?v0) ?v1)))
+(= $x9 $x10))) :pattern ( (|fun_app$| (|fun_app$a| |uu$| ?v0) ?v1) )))
+))
+(let (($x46 (forall ((?v0 Int) (?v1 Int) )(!(let (($x10 (<= ?v0 ?v1)))
+(let (($x9 (|fun_app$| (|fun_app$a| |uu$| ?v0) ?v1)))
+(= $x9 $x10))) :pattern ( (|fun_app$| (|fun_app$a| |uu$| ?v0) ?v1) )))
+))
+(let ((@x57 (monotonicity (rewrite (= (<= ?1 ?0) $x52)) (= (= $x9 (<= ?1 ?0)) $x55))))
+(let ((@x48 (|quant-intro| (rewrite (= (= $x9 (<= ?1 ?0)) (= $x9 (<= ?1 ?0)))) (= $x13 $x46))))
+(let ((@x63 (mp (asserted $x13) (trans @x48 (|quant-intro| @x57 (= $x46 $x58)) (= $x13 $x58)) $x58)))
+(let ((@x80 (|mp~| @x63 (|nnf-pos| (refl (|~| $x55 $x55)) (|~| $x58 $x58)) $x58)))
+(let (($x113 (or (not $x58) $x82)))
+(let (($x116 (= (or (not $x58) (= $x82 (<= (+ 3 (* (~ 1) 42)) 0))) $x113)))
+(let ((?x83 (* (~ 1) 42)))
+(let ((?x84 (+ 3 ?x83)))
+(let (($x85 (<= ?x84 0)))
+(let (($x86 (= $x82 $x85)))
+(let ((@x97 (trans (monotonicity (rewrite (= ?x83 (~ 42))) (= ?x84 (+ 3 (~ 42)))) (rewrite (= (+ 3 (~ 42)) (~ 39))) (= ?x84 (~ 39)))))
+(let ((@x104 (trans (monotonicity @x97 (= $x85 (<= (~ 39) 0))) (rewrite (= (<= (~ 39) 0) true)) (= $x85 true))))
+(let ((@x111 (trans (monotonicity @x104 (= $x86 (= $x82 true))) (rewrite (= (= $x82 true) $x82)) (= $x86 $x82))))
+(let ((@x121 (mp ((_ |quant-inst| 3 42) (or (not $x58) $x86)) (trans (monotonicity @x111 $x116) (rewrite (= $x113 $x113)) $x116) $x113)))
+(|unit-resolution| (|unit-resolution| @x121 @x80 $x82) (mp @x75 @x133 (not $x82)) false)))))))))))))))))))))))))))))))))))
+
+737e9aeb0ce08125531c37a003d0a11fe8c1aa00 189 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let ((?x37 (|nat$| 2)))
+(let ((?x38 (|cons$| ?x37 |nil$|)))
+(let ((?x32 (|nat$| 1)))
+(let ((?x39 (|cons$| ?x32 ?x38)))
+(let ((?x33 (|cons$| ?x32 |nil$|)))
+(let ((?x31 (|nat$| 0)))
+(let ((?x34 (|cons$| ?x31 ?x33)))
+(let ((?x35 (|map$| |uu$| ?x34)))
+(let (($x40 (= ?x35 ?x39)))
+(let ((?x208 (|map$| |uu$| ?x33)))
+(let ((?x326 (|map$| |uu$| |nil$|)))
+(let ((?x325 (|fun_app$| |uu$| ?x32)))
+(let ((?x327 (|cons$| ?x325 ?x326)))
+(let (($x328 (= ?x208 ?x327)))
+(let (($x181 (forall ((?v0 |Nat_nat_fun$|) (?v1 |Nat$|) (?v2 |Nat_list$|) )(!(let ((?x27 (|cons$| (|fun_app$| ?v0 ?v1) (|map$| ?v0 ?v2))))
+(let ((?x24 (|map$| ?v0 (|cons$| ?v1 ?v2))))
+(= ?x24 ?x27))) :pattern ( (|map$| ?v0 (|cons$| ?v1 ?v2)) ) :pattern ( (|cons$| (|fun_app$| ?v0 ?v1) (|map$| ?v0 ?v2)) )))
+))
+(let (($x29 (forall ((?v0 |Nat_nat_fun$|) (?v1 |Nat$|) (?v2 |Nat_list$|) )(let ((?x27 (|cons$| (|fun_app$| ?v0 ?v1) (|map$| ?v0 ?v2))))
+(let ((?x24 (|map$| ?v0 (|cons$| ?v1 ?v2))))
+(= ?x24 ?x27))))
+))
+(let ((?x27 (|cons$| (|fun_app$| ?2 ?1) (|map$| ?2 ?0))))
+(let ((?x24 (|map$| ?2 (|cons$| ?1 ?0))))
+(let (($x28 (= ?x24 ?x27)))
+(let ((@x156 (|mp~| (asserted $x29) (|nnf-pos| (refl (|~| $x28 $x28)) (|~| $x29 $x29)) $x29)))
+(let ((@x186 (mp @x156 (|quant-intro| (refl (= $x28 $x28)) (= $x29 $x181)) $x181)))
+(let (($x213 (not $x181)))
+(let (($x331 (or $x213 $x328)))
+(let ((@x332 ((_ |quant-inst| |uu$| (|nat$| 1) |nil$|) $x331)))
+(let (($x339 (= ?x326 |nil$|)))
+(let (($x173 (forall ((?v0 |Nat_nat_fun$|) )(!(= (|map$| ?v0 |nil$|) |nil$|) :pattern ( (|map$| ?v0 |nil$|) )))
+))
+(let (($x19 (forall ((?v0 |Nat_nat_fun$|) )(= (|map$| ?v0 |nil$|) |nil$|))
+))
+(let ((@x175 (refl (= (= (|map$| ?0 |nil$|) |nil$|) (= (|map$| ?0 |nil$|) |nil$|)))))
+(let ((@x148 (refl (|~| (= (|map$| ?0 |nil$|) |nil$|) (= (|map$| ?0 |nil$|) |nil$|)))))
+(let ((@x178 (mp (|mp~| (asserted $x19) (|nnf-pos| @x148 (|~| $x19 $x19)) $x19) (|quant-intro| @x175 (= $x19 $x173)) $x173)))
+(let (($x343 (or (not $x173) $x339)))
+(let ((@x344 ((_ |quant-inst| |uu$|) $x343)))
+(let ((?x255 (|of_nat$| ?x32)))
+(let ((?x340 (+ 1 ?x255)))
+(let ((?x341 (|nat$| ?x340)))
+(let (($x345 (= ?x325 ?x341)))
+(let (($x85 (forall ((?v0 |Nat$|) )(!(let ((?x7 (|fun_app$| |uu$| ?v0)))
+(= ?x7 (|nat$| (+ 1 (|of_nat$| ?v0))))) :pattern ( (|fun_app$| |uu$| ?v0) )))
+))
+(let ((?x7 (|fun_app$| |uu$| ?0)))
+(let (($x82 (= ?x7 (|nat$| (+ 1 (|of_nat$| ?0))))))
+(let (($x14 (forall ((?v0 |Nat$|) )(!(let ((?x7 (|fun_app$| |uu$| ?v0)))
+(= ?x7 (|nat$| (+ (|of_nat$| ?v0) 1)))) :pattern ( (|fun_app$| |uu$| ?v0) )))
+))
+(let ((@x81 (monotonicity (rewrite (= (+ (|of_nat$| ?0) 1) (+ 1 (|of_nat$| ?0)))) (= (|nat$| (+ (|of_nat$| ?0) 1)) (|nat$| (+ 1 (|of_nat$| ?0)))))))
+(let ((@x84 (monotonicity @x81 (= (= ?x7 (|nat$| (+ (|of_nat$| ?0) 1))) $x82))))
+(let ((@x146 (|mp~| (mp (asserted $x14) (|quant-intro| @x84 (= $x14 $x85)) $x85) (|nnf-pos| (refl (|~| $x82 $x82)) (|~| $x85 $x85)) $x85)))
+(let (($x348 (or (not $x85) $x345)))
+(let ((@x349 ((_ |quant-inst| (|nat$| 1)) $x348)))
+(let ((?x404 (|of_nat$| ?x341)))
+(let ((?x454 (|nat$| ?x404)))
+(let (($x455 (= ?x454 ?x341)))
+(let (($x188 (forall ((?v0 |Nat$|) )(!(= (|nat$| (|of_nat$| ?v0)) ?v0) :pattern ( (|of_nat$| ?v0) )))
+))
+(let (($x44 (forall ((?v0 |Nat$|) )(= (|nat$| (|of_nat$| ?v0)) ?v0))
+))
+(let ((@x190 (refl (= (= (|nat$| (|of_nat$| ?0)) ?0) (= (|nat$| (|of_nat$| ?0)) ?0)))))
+(let ((@x160 (refl (|~| (= (|nat$| (|of_nat$| ?0)) ?0) (= (|nat$| (|of_nat$| ?0)) ?0)))))
+(let ((@x193 (mp (|mp~| (asserted $x44) (|nnf-pos| @x160 (|~| $x44 $x44)) $x44) (|quant-intro| @x190 (= $x44 $x188)) $x188)))
+(let (($x461 (or (not $x188) $x455)))
+(let ((@x462 ((_ |quant-inst| (|nat$| ?x340)) $x461)))
+(let ((?x415 (* (~ 1) ?x404)))
+(let ((?x416 (+ ?x255 ?x415)))
+(let (($x432 (<= ?x416 (~ 1))))
+(let (($x414 (= ?x416 (~ 1))))
+(let (($x407 (>= ?x255 (~ 1))))
+(let (($x401 (>= ?x255 1)))
+(let (($x256 (= ?x255 1)))
+(let (($x195 (forall ((?v0 Int) )(!(let (($x49 (= (|of_nat$| (|nat$| ?v0)) ?v0)))
+(let (($x103 (>= ?v0 0)))
+(let (($x104 (not $x103)))
+(or $x104 $x49)))) :pattern ( (|nat$| ?v0) )))
+))
+(let (($x110 (forall ((?v0 Int) )(let (($x49 (= (|of_nat$| (|nat$| ?v0)) ?v0)))
+(let (($x103 (>= ?v0 0)))
+(let (($x104 (not $x103)))
+(or $x104 $x49)))))
+))
+(let (($x49 (= (|of_nat$| (|nat$| ?0)) ?0)))
+(let (($x103 (>= ?0 0)))
+(let (($x104 (not $x103)))
+(let (($x107 (or $x104 $x49)))
+(let (($x51 (forall ((?v0 Int) )(let (($x49 (= (|of_nat$| (|nat$| ?v0)) ?v0)))
+(let (($x46 (<= 0 ?v0)))
+(=> $x46 $x49))))
+))
+(let (($x98 (forall ((?v0 Int) )(let (($x49 (= (|of_nat$| (|nat$| ?v0)) ?v0)))
+(or (not (<= 0 ?v0)) $x49)))
+))
+(let ((@x106 (monotonicity (rewrite (= (<= 0 ?0) $x103)) (= (not (<= 0 ?0)) $x104))))
+(let ((@x112 (|quant-intro| (monotonicity @x106 (= (or (not (<= 0 ?0)) $x49) $x107)) (= $x98 $x110))))
+(let ((@x97 (rewrite (= (=> (<= 0 ?0) $x49) (or (not (<= 0 ?0)) $x49)))))
+(let ((@x115 (mp (asserted $x51) (trans (|quant-intro| @x97 (= $x51 $x98)) @x112 (= $x51 $x110)) $x110)))
+(let ((@x200 (mp (|mp~| @x115 (|nnf-pos| (refl (|~| $x107 $x107)) (|~| $x110 $x110)) $x110) (|quant-intro| (refl (= $x107 $x107)) (= $x110 $x195)) $x195)))
+(let (($x235 (not $x195)))
+(let (($x271 (or $x235 $x256)))
+(let ((@x225 (rewrite (= (not true) false))))
+(let ((@x259 (rewrite (= (>= 1 0) true))))
+(let ((@x263 (trans (monotonicity @x259 (= (not (>= 1 0)) (not true))) @x225 (= (not (>= 1 0)) false))))
+(let ((@x266 (monotonicity @x263 (= (or (not (>= 1 0)) $x256) (or false $x256)))))
+(let ((@x270 (trans @x266 (rewrite (= (or false $x256) $x256)) (= (or (not (>= 1 0)) $x256) $x256))))
+(let ((@x275 (monotonicity @x270 (= (or $x235 (or (not (>= 1 0)) $x256)) $x271))))
+(let ((@x278 (trans @x275 (rewrite (= $x271 $x271)) (= (or $x235 (or (not (>= 1 0)) $x256)) $x271))))
+(let ((@x279 (mp ((_ |quant-inst| 1) (or $x235 (or (not (>= 1 0)) $x256))) @x278 $x271)))
+(let ((@x477 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x256) $x401)) (|unit-resolution| @x279 @x200 $x256) $x401)))
+(let (($x410 (not $x407)))
+(let (($x421 (or $x235 $x410 $x414)))
+(let (($x405 (= ?x404 ?x340)))
+(let (($x402 (>= ?x340 0)))
+(let (($x403 (not $x402)))
+(let (($x406 (or $x403 $x405)))
+(let (($x422 (or $x235 $x406)))
+(let ((@x420 (monotonicity (monotonicity (rewrite (= $x402 $x407)) (= $x403 $x410)) (rewrite (= $x405 $x414)) (= $x406 (or $x410 $x414)))))
+(let ((@x430 (trans (monotonicity @x420 (= $x422 (or $x235 (or $x410 $x414)))) (rewrite (= (or $x235 (or $x410 $x414)) $x421)) (= $x422 $x421))))
+(let ((@x431 (mp ((_ |quant-inst| (+ 1 ?x255)) $x422) @x430 $x421)))
+(let ((@x482 (|unit-resolution| @x431 @x200 (|unit-resolution| ((_ |th-lemma| arith farkas 1 1) (or (not $x401) $x407)) @x477 $x407) $x414)))
+(let (($x433 (>= ?x416 (~ 1))))
+(let (($x400 (<= ?x255 1)))
+(let ((@x492 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x256) $x400)) (|unit-resolution| @x279 @x200 $x256) $x400)))
+(let ((@x494 ((_ |th-lemma| arith eq-propagate -1 -1 1 1) @x477 @x492 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x414) $x433)) @x482 $x433) (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x414) $x432)) @x482 $x432) (= ?x404 2))))
+(let ((@x502 (trans (monotonicity (symm @x494 (= 2 ?x404)) (= ?x37 ?x454)) (|unit-resolution| @x462 @x193 $x455) (= ?x37 ?x341))))
+(let ((@x504 (trans @x502 (symm (|unit-resolution| @x349 @x146 $x345) (= ?x341 ?x325)) (= ?x37 ?x325))))
+(let ((@x506 (monotonicity @x504 (symm (|unit-resolution| @x344 @x178 $x339) (= |nil$| ?x326)) (= ?x38 ?x327))))
+(let ((@x510 (trans @x506 (symm (|unit-resolution| @x332 @x186 $x328) (= ?x327 ?x208)) (= ?x38 ?x208))))
+(let ((?x216 (|of_nat$| ?x31)))
+(let ((?x329 (+ 1 ?x216)))
+(let ((?x330 (|nat$| ?x329)))
+(let ((?x207 (|fun_app$| |uu$| ?x31)))
+(let (($x333 (= ?x207 ?x330)))
+(let (($x337 (or (not $x85) $x333)))
+(let ((@x338 ((_ |quant-inst| (|nat$| 0)) $x337)))
+(let ((?x350 (|of_nat$| ?x330)))
+(let ((?x452 (|nat$| ?x350)))
+(let (($x453 (= ?x452 ?x330)))
+(let (($x457 (or (not $x188) $x453)))
+(let ((@x458 ((_ |quant-inst| (|nat$| ?x329)) $x457)))
+(let ((?x362 (* (~ 1) ?x350)))
+(let ((?x363 (+ ?x216 ?x362)))
+(let (($x379 (<= ?x363 (~ 1))))
+(let (($x361 (= ?x363 (~ 1))))
+(let (($x354 (>= ?x216 (~ 1))))
+(let (($x335 (>= ?x216 0)))
+(let (($x217 (= ?x216 0)))
+(let (($x236 (or $x235 $x217)))
+(let ((@x220 (rewrite (= (>= 0 0) true))))
+(let ((@x227 (trans (monotonicity @x220 (= (not (>= 0 0)) (not true))) @x225 (= (not (>= 0 0)) false))))
+(let ((@x230 (monotonicity @x227 (= (or (not (>= 0 0)) $x217) (or false $x217)))))
+(let ((@x234 (trans @x230 (rewrite (= (or false $x217) $x217)) (= (or (not (>= 0 0)) $x217) $x217))))
+(let ((@x240 (monotonicity @x234 (= (or $x235 (or (not (>= 0 0)) $x217)) $x236))))
+(let ((@x243 (trans @x240 (rewrite (= $x236 $x236)) (= (or $x235 (or (not (>= 0 0)) $x217)) $x236))))
+(let ((@x244 (mp ((_ |quant-inst| 0) (or $x235 (or (not (>= 0 0)) $x217))) @x243 $x236)))
+(let ((@x517 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x217) $x335)) (|unit-resolution| @x244 @x200 $x217) $x335)))
+(let (($x357 (not $x354)))
+(let (($x368 (or $x235 $x357 $x361)))
+(let (($x351 (= ?x350 ?x329)))
+(let (($x346 (>= ?x329 0)))
+(let (($x347 (not $x346)))
+(let (($x352 (or $x347 $x351)))
+(let (($x369 (or $x235 $x352)))
+(let ((@x367 (monotonicity (monotonicity (rewrite (= $x346 $x354)) (= $x347 $x357)) (rewrite (= $x351 $x361)) (= $x352 (or $x357 $x361)))))
+(let ((@x377 (trans (monotonicity @x367 (= $x369 (or $x235 (or $x357 $x361)))) (rewrite (= (or $x235 (or $x357 $x361)) $x368)) (= $x369 $x368))))
+(let ((@x378 (mp ((_ |quant-inst| (+ 1 ?x216)) $x369) @x377 $x368)))
+(let ((@x522 (|unit-resolution| @x378 @x200 (|unit-resolution| ((_ |th-lemma| arith farkas 1 1) (or (not $x335) $x354)) @x517 $x354) $x361)))
+(let (($x380 (>= ?x363 (~ 1))))
+(let (($x334 (<= ?x216 0)))
+(let ((@x532 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x217) $x334)) (|unit-resolution| @x244 @x200 $x217) $x334)))
+(let ((@x534 ((_ |th-lemma| arith eq-propagate -1 -1 1 1) @x517 @x532 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x361) $x380)) @x522 $x380) (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x361) $x379)) @x522 $x379) (= ?x350 1))))
+(let ((@x542 (trans (monotonicity (symm @x534 (= 1 ?x350)) (= ?x32 ?x452)) (|unit-resolution| @x458 @x193 $x453) (= ?x32 ?x330))))
+(let ((@x544 (trans @x542 (symm (|unit-resolution| @x338 @x146 $x333) (= ?x330 ?x207)) (= ?x32 ?x207))))
+(let ((@x549 (symm (monotonicity @x544 @x510 (= ?x39 (|cons$| ?x207 ?x208))) (= (|cons$| ?x207 ?x208) ?x39))))
+(let ((?x209 (|cons$| ?x207 ?x208)))
+(let (($x210 (= ?x35 ?x209)))
+(let (($x214 (or $x213 $x210)))
+(let ((@x215 ((_ |quant-inst| |uu$| (|nat$| 0) (|cons$| ?x32 |nil$|)) $x214)))
+(let (($x41 (not $x40)))
+(let ((@x91 (asserted $x41)))
+(|unit-resolution| @x91 (trans (|unit-resolution| @x215 @x186 $x210) @x549 $x40) false))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))
+
+516b38db52977db0759f7a2fb4ee2c61d2623ab0 11 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let (($x7 (forall ((?v0 |A$|) )(|p$| ?v0))
+))
+(let (($x8 (not $x7)))
+(let (($x9 (or $x7 $x8)))
+(let (($x10 (not $x9)))
+(let ((@x40 (trans (monotonicity (rewrite (= $x9 true)) (= $x10 (not true))) (rewrite (= (not true) false)) (= $x10 false))))
+(mp (asserted $x10) @x40 false))))))))
+
+120a595aca724e41775e5a997277b8d456a7e9fe 183 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let ((?x24 (|nat$| 6)))
+(let ((?x17 (|nat$| 4)))
+(let ((?x18 (|dec_10$| ?x17)))
+(let ((?x19 (|of_nat$| ?x18)))
+(let ((?x20 (* 4 ?x19)))
+(let ((?x21 (|nat$| ?x20)))
+(let ((?x22 (|dec_10$| ?x21)))
+(let (($x25 (= ?x22 ?x24)))
+(let ((?x348 (|of_nat$| ?x24)))
+(let ((?x393 (+ (~ 10) ?x348)))
+(let ((?x394 (|nat$| ?x393)))
+(let ((?x395 (|dec_10$| ?x394)))
+(let (($x390 (>= ?x348 10)))
+(let ((?x396 (ite $x390 ?x395 ?x24)))
+(let (($x403 (= ?x24 ?x396)))
+(let (($x404 (not $x390)))
+(let (($x398 (<= ?x348 6)))
+(let (($x349 (= ?x348 6)))
+(let (($x188 (forall ((?v0 Int) )(!(let ((?x33 (|nat$| ?v0)))
+(let ((?x34 (|of_nat$| ?x33)))
+(let (($x35 (= ?x34 ?v0)))
+(let (($x114 (>= ?v0 0)))
+(let (($x115 (not $x114)))
+(or $x115 $x35)))))) :pattern ( (|nat$| ?v0) )))
+))
+(let (($x121 (forall ((?v0 Int) )(let ((?x33 (|nat$| ?v0)))
+(let ((?x34 (|of_nat$| ?x33)))
+(let (($x35 (= ?x34 ?v0)))
+(let (($x114 (>= ?v0 0)))
+(let (($x115 (not $x114)))
+(or $x115 $x35)))))))
+))
+(let ((?x33 (|nat$| ?0)))
+(let ((?x34 (|of_nat$| ?x33)))
+(let (($x35 (= ?x34 ?0)))
+(let (($x114 (>= ?0 0)))
+(let (($x115 (not $x114)))
+(let (($x118 (or $x115 $x35)))
+(let (($x37 (forall ((?v0 Int) )(let ((?x33 (|nat$| ?v0)))
+(let ((?x34 (|of_nat$| ?x33)))
+(let (($x35 (= ?x34 ?v0)))
+(let (($x32 (<= 0 ?v0)))
+(=> $x32 $x35))))))
+))
+(let (($x109 (forall ((?v0 Int) )(let ((?x33 (|nat$| ?v0)))
+(let ((?x34 (|of_nat$| ?x33)))
+(let (($x35 (= ?x34 ?v0)))
+(or (not (<= 0 ?v0)) $x35)))))
+))
+(let ((@x117 (monotonicity (rewrite (= (<= 0 ?0) $x114)) (= (not (<= 0 ?0)) $x115))))
+(let ((@x123 (|quant-intro| (monotonicity @x117 (= (or (not (<= 0 ?0)) $x35) $x118)) (= $x109 $x121))))
+(let ((@x108 (rewrite (= (=> (<= 0 ?0) $x35) (or (not (<= 0 ?0)) $x35)))))
+(let ((@x126 (mp (asserted $x37) (trans (|quant-intro| @x108 (= $x37 $x109)) @x123 (= $x37 $x121)) $x121)))
+(let ((@x193 (mp (|mp~| @x126 (|nnf-pos| (refl (|~| $x118 $x118)) (|~| $x121 $x121)) $x121) (|quant-intro| (refl (= $x118 $x118)) (= $x121 $x188)) $x188)))
+(let (($x274 (not $x188)))
+(let (($x364 (or $x274 $x349)))
+(let ((@x352 (rewrite (= (>= 6 0) true))))
+(let ((@x356 (trans (monotonicity @x352 (= (not (>= 6 0)) (not true))) (rewrite (= (not true) false)) (= (not (>= 6 0)) false))))
+(let ((@x359 (monotonicity @x356 (= (or (not (>= 6 0)) $x349) (or false $x349)))))
+(let ((@x363 (trans @x359 (rewrite (= (or false $x349) $x349)) (= (or (not (>= 6 0)) $x349) $x349))))
+(let ((@x368 (monotonicity @x363 (= (or $x274 (or (not (>= 6 0)) $x349)) $x364))))
+(let ((@x371 (trans @x368 (rewrite (= $x364 $x364)) (= (or $x274 (or (not (>= 6 0)) $x349)) $x364))))
+(let ((@x372 (mp ((_ |quant-inst| 6) (or $x274 (or (not (>= 6 0)) $x349))) @x371 $x364)))
+(let ((@x422 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x349) $x398)) (|unit-resolution| @x372 @x193 $x349) $x398)))
+(let ((@x427 (|unit-resolution| (|def-axiom| (or $x390 $x403)) (|unit-resolution| ((_ |th-lemma| arith farkas 1 1) (or (not $x398) $x404)) @x422 $x404) $x403)))
+(let ((?x389 (|dec_10$| ?x24)))
+(let (($x397 (= ?x389 ?x396)))
+(let (($x175 (forall ((?v0 |Nat$|) )(!(let ((?x69 (|dec_10$| (|nat$| (+ (~ 10) (|of_nat$| ?v0))))))
+(let ((?x88 (ite (>= (|of_nat$| ?v0) 10) ?x69 ?v0)))
+(let ((?x6 (|dec_10$| ?v0)))
+(= ?x6 ?x88)))) :pattern ( (|dec_10$| ?v0) ) :pattern ( (|of_nat$| ?v0) )))
+))
+(let (($x96 (forall ((?v0 |Nat$|) )(let ((?x69 (|dec_10$| (|nat$| (+ (~ 10) (|of_nat$| ?v0))))))
+(let ((?x88 (ite (>= (|of_nat$| ?v0) 10) ?x69 ?v0)))
+(let ((?x6 (|dec_10$| ?v0)))
+(= ?x6 ?x88)))))
+))
+(let ((?x69 (|dec_10$| (|nat$| (+ (~ 10) (|of_nat$| ?0))))))
+(let ((?x88 (ite (>= (|of_nat$| ?0) 10) ?x69 ?0)))
+(let ((?x6 (|dec_10$| ?0)))
+(let (($x93 (= ?x6 ?x88)))
+(let (($x15 (forall ((?v0 |Nat$|) )(let ((?x7 (|of_nat$| ?v0)))
+(let (($x9 (< ?x7 10)))
+(let ((?x6 (|dec_10$| ?v0)))
+(= ?x6 (ite $x9 ?v0 (|dec_10$| (|nat$| (- ?x7 10)))))))))
+))
+(let (($x78 (forall ((?v0 |Nat$|) )(let ((?x69 (|dec_10$| (|nat$| (+ (~ 10) (|of_nat$| ?v0))))))
+(let ((?x7 (|of_nat$| ?v0)))
+(let (($x9 (< ?x7 10)))
+(let ((?x72 (ite $x9 ?v0 ?x69)))
+(let ((?x6 (|dec_10$| ?v0)))
+(= ?x6 ?x72)))))))
+))
+(let ((?x7 (|of_nat$| ?0)))
+(let (($x9 (< ?x7 10)))
+(let ((?x72 (ite $x9 ?0 ?x69)))
+(let ((@x87 (monotonicity (rewrite (= $x9 (not (>= ?x7 10)))) (= ?x72 (ite (not (>= ?x7 10)) ?0 ?x69)))))
+(let ((@x92 (trans @x87 (rewrite (= (ite (not (>= ?x7 10)) ?0 ?x69) ?x88)) (= ?x72 ?x88))))
+(let ((@x98 (|quant-intro| (monotonicity @x92 (= (= ?x6 ?x72) $x93)) (= $x78 $x96))))
+(let (($x75 (= ?x6 ?x72)))
+(let ((@x68 (monotonicity (rewrite (= (- ?x7 10) (+ (~ 10) ?x7))) (= (|nat$| (- ?x7 10)) (|nat$| (+ (~ 10) ?x7))))))
+(let ((@x74 (monotonicity (monotonicity @x68 (= (|dec_10$| (|nat$| (- ?x7 10))) ?x69)) (= (ite $x9 ?0 (|dec_10$| (|nat$| (- ?x7 10)))) ?x72))))
+(let ((@x77 (monotonicity @x74 (= (= ?x6 (ite $x9 ?0 (|dec_10$| (|nat$| (- ?x7 10))))) $x75))))
+(let ((@x101 (mp (asserted $x15) (trans (|quant-intro| @x77 (= $x15 $x78)) @x98 (= $x15 $x96)) $x96)))
+(let ((@x180 (mp (|mp~| @x101 (|nnf-pos| (refl (|~| $x93 $x93)) (|~| $x96 $x96)) $x96) (|quant-intro| (refl (= $x93 $x93)) (= $x96 $x175)) $x175)))
+(let (($x209 (not $x175)))
+(let (($x400 (or $x209 $x397)))
+(let ((@x401 ((_ |quant-inst| (|nat$| 6)) $x400)))
+(let ((?x200 (|of_nat$| ?x17)))
+(let ((?x383 (* (~ 1) ?x200)))
+(let ((?x384 (+ ?x19 ?x383)))
+(let (($x385 (<= ?x384 0)))
+(let (($x382 (= ?x19 ?x200)))
+(let ((?x202 (+ (~ 10) ?x200)))
+(let ((?x203 (|nat$| ?x202)))
+(let ((?x204 (|dec_10$| ?x203)))
+(let (($x201 (>= ?x200 10)))
+(let ((?x205 (ite $x201 ?x204 ?x17)))
+(let (($x213 (= ?x17 ?x205)))
+(let (($x214 (not $x201)))
+(let (($x284 (<= ?x200 4)))
+(let (($x256 (= ?x200 4)))
+(let (($x275 (or $x274 $x256)))
+(let ((@x259 (rewrite (= (>= 4 0) true))))
+(let ((@x266 (trans (monotonicity @x259 (= (not (>= 4 0)) (not true))) (rewrite (= (not true) false)) (= (not (>= 4 0)) false))))
+(let ((@x269 (monotonicity @x266 (= (or (not (>= 4 0)) $x256) (or false $x256)))))
+(let ((@x273 (trans @x269 (rewrite (= (or false $x256) $x256)) (= (or (not (>= 4 0)) $x256) $x256))))
+(let ((@x279 (monotonicity @x273 (= (or $x274 (or (not (>= 4 0)) $x256)) $x275))))
+(let ((@x282 (trans @x279 (rewrite (= $x275 $x275)) (= (or $x274 (or (not (>= 4 0)) $x256)) $x275))))
+(let ((@x283 (mp ((_ |quant-inst| 4) (or $x274 (or (not (>= 4 0)) $x256))) @x282 $x275)))
+(let ((@x433 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x256) $x284)) (|unit-resolution| @x283 @x193 $x256) $x284)))
+(let ((@x438 (|unit-resolution| (|def-axiom| (or $x201 $x213)) (|unit-resolution| ((_ |th-lemma| arith farkas 1 1) (or (not $x284) $x214)) @x433 $x214) $x213)))
+(let (($x206 (= ?x18 ?x205)))
+(let (($x210 (or $x209 $x206)))
+(let ((@x211 ((_ |quant-inst| (|nat$| 4)) $x210)))
+(let ((@x443 (trans (|unit-resolution| @x211 @x180 $x206) (symm @x438 (= ?x205 ?x17)) (= ?x18 ?x17))))
+(let ((@x448 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x382) $x385)) (monotonicity @x443 $x382) $x385)))
+(let (($x386 (>= ?x384 0)))
+(let ((@x451 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x382) $x386)) (monotonicity @x443 $x382) $x386)))
+(let (($x285 (>= ?x200 4)))
+(let ((@x454 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x256) $x285)) (|unit-resolution| @x283 @x193 $x256) $x285)))
+(let ((?x207 (|of_nat$| ?x21)))
+(let ((?x309 (* (~ 1) ?x207)))
+(let ((?x310 (+ ?x20 ?x309)))
+(let (($x325 (<= ?x310 0)))
+(let (($x307 (= ?x310 0)))
+(let (($x299 (>= ?x19 0)))
+(let ((@x459 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 1) (or $x299 (not $x285) (not $x386))) @x454 @x451 $x299)))
+(let (($x302 (not $x299)))
+(let (($x311 (or $x302 $x307)))
+(let (($x314 (or $x274 $x302 $x307)))
+(let (($x297 (= ?x207 ?x20)))
+(let (($x295 (>= ?x20 0)))
+(let (($x296 (not $x295)))
+(let (($x298 (or $x296 $x297)))
+(let (($x315 (or $x274 $x298)))
+(let ((@x313 (monotonicity (monotonicity (rewrite (= $x295 $x299)) (= $x296 $x302)) (rewrite (= $x297 $x307)) (= $x298 $x311))))
+(let ((@x323 (trans (monotonicity @x313 (= $x315 (or $x274 $x311))) (rewrite (= (or $x274 $x311) $x314)) (= $x315 $x314))))
+(let ((@x324 (mp ((_ |quant-inst| (* 4 ?x19)) $x315) @x323 $x314)))
+(let ((@x465 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x307) $x325)) (|unit-resolution| (|unit-resolution| @x324 @x193 $x311) @x459 $x307) $x325)))
+(let (($x326 (>= ?x310 0)))
+(let ((@x468 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x307) $x326)) (|unit-resolution| (|unit-resolution| @x324 @x193 $x311) @x459 $x307) $x326)))
+(let ((@x472 (monotonicity ((_ |th-lemma| arith eq-propagate 1 1 -4 -4 -4 -4) @x468 @x465 @x454 @x433 @x451 @x448 (= (+ (~ 10) ?x207) 6)) (= (|nat$| (+ (~ 10) ?x207)) ?x24))))
+(let ((?x219 (+ (~ 10) ?x207)))
+(let ((?x220 (|nat$| ?x219)))
+(let ((?x221 (|dec_10$| ?x220)))
+(let (($x208 (>= ?x207 10)))
+(let ((?x222 (ite $x208 ?x221 ?x21)))
+(let (($x228 (= ?x221 ?x222)))
+(let ((@x476 (|unit-resolution| ((_ |th-lemma| arith assign-bounds 1 4 4) (or $x208 (not $x325) (not $x285) (not $x386))) @x454 @x465 @x451 $x208)))
+(let ((@x480 (symm (|unit-resolution| (|def-axiom| (or (not $x208) $x228)) @x476 $x228) (= ?x222 ?x221))))
+(let (($x223 (= ?x22 ?x222)))
+(let (($x226 (or $x209 $x223)))
+(let ((@x227 ((_ |quant-inst| (|nat$| ?x20)) $x226)))
+(let ((@x488 (trans (trans (|unit-resolution| @x227 @x180 $x223) @x480 (= ?x22 ?x221)) (monotonicity @x472 (= ?x221 ?x389)) (= ?x22 ?x389))))
+(let ((@x491 (trans (trans @x488 (|unit-resolution| @x401 @x180 $x397) (= ?x22 ?x396)) (symm @x427 (= ?x396 ?x24)) $x25)))
+(let (($x26 (not $x25)))
+(let ((@x102 (asserted $x26)))
+(|unit-resolution| @x102 @x491 false))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))
+
+a25b64a8f04cffe57bd9a525d4bad154c19d2b20 310 0
+unsat
+((set-logic <null>)
+(proof
+(let ((?x42 (|map$| |uu$| |xs$|)))
+(let ((?x43 (|eval_dioph$| |ks$| ?x42)))
+(let ((?x145 (* (~ 1) ?x43)))
+(let ((?x146 (+ |l$| ?x145)))
+(let ((?x149 (|div$| ?x146 2)))
+(let ((?x47 (|map$| |uua$| |xs$|)))
+(let ((?x48 (|eval_dioph$| |ks$| ?x47)))
+(let ((?x769 (+ ?x48 (* (~ 1) ?x149))))
+(let (($x771 (>= ?x769 0)))
+(let ((?x452 (* (~ 1) |l$|)))
+(let ((?x39 (|eval_dioph$| |ks$| |xs$|)))
+(let ((?x776 (+ ?x39 ?x452)))
+(let (($x778 (>= ?x776 0)))
+(let (($x41 (= ?x39 |l$|)))
+(let (($x152 (= ?x48 ?x149)))
+(let (($x362 (not $x152)))
+(let ((?x45 (|mod$| |l$| 2)))
+(let ((?x44 (|mod$| ?x43 2)))
+(let (($x46 (= ?x44 ?x45)))
+(let (($x361 (not $x46)))
+(let (($x363 (or $x361 $x362)))
+(let ((?x730 (div ?x43 2)))
+(let ((?x913 (* (~ 1) ?x730)))
+(let ((?x685 (mod ?x43 2)))
+(let ((?x712 (* (~ 1) ?x685)))
+(let ((?x645 (div |l$| 2)))
+(let ((?x671 (* (~ 1) ?x645)))
+(let ((?x599 (mod |l$| 2)))
+(let ((?x626 (* (~ 1) ?x599)))
+(let ((?x656 (* (~ 1) ?x48)))
+(let (($x737 (>= (+ |l$| ?x44 ?x656 ?x626 ?x671 ?x712 ?x913) 1)))
+(let ((?x658 (* (~ 2) ?x645)))
+(let ((?x659 (+ |l$| ?x626 ?x658)))
+(let (($x657 (= ?x659 0)))
+(let ((@x108 (|true-axiom| true)))
+(let ((@x945 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x657) (>= ?x659 0))) (|unit-resolution| ((_ |th-lemma| arith) (or false $x657)) @x108 $x657) (>= ?x659 0))))
+(let ((?x627 (+ ?x45 ?x626)))
+(let (($x628 (= ?x627 0)))
+(let (($x418 (forall ((?v0 Int) (?v1 Int) )(!(let ((?x83 (mod ?v0 ?v1)))
+(let ((?x230 (* (~ 1) ?v1)))
+(let ((?x227 (* (~ 1) ?v0)))
+(let ((?x273 (mod ?x227 ?x230)))
+(let ((?x279 (* (~ 1) ?x273)))
+(let (($x248 (<= ?v1 0)))
+(let ((?x299 (ite $x248 ?x279 ?x83)))
+(let (($x72 (= ?v1 0)))
+(let ((?x304 (ite $x72 ?v0 ?x299)))
+(let ((?x82 (|mod$| ?v0 ?v1)))
+(= ?x82 ?x304))))))))))) :pattern ( (|mod$| ?v0 ?v1) )))
+))
+(let (($x310 (forall ((?v0 Int) (?v1 Int) )(let ((?x83 (mod ?v0 ?v1)))
+(let ((?x230 (* (~ 1) ?v1)))
+(let ((?x227 (* (~ 1) ?v0)))
+(let ((?x273 (mod ?x227 ?x230)))
+(let ((?x279 (* (~ 1) ?x273)))
+(let (($x248 (<= ?v1 0)))
+(let ((?x299 (ite $x248 ?x279 ?x83)))
+(let (($x72 (= ?v1 0)))
+(let ((?x304 (ite $x72 ?v0 ?x299)))
+(let ((?x82 (|mod$| ?v0 ?v1)))
+(= ?x82 ?x304))))))))))))
+))
+(let ((?x83 (mod ?1 ?0)))
+(let ((?x230 (* (~ 1) ?0)))
+(let ((?x227 (* (~ 1) ?1)))
+(let ((?x273 (mod ?x227 ?x230)))
+(let ((?x279 (* (~ 1) ?x273)))
+(let (($x248 (<= ?0 0)))
+(let ((?x299 (ite $x248 ?x279 ?x83)))
+(let (($x72 (= ?0 0)))
+(let ((?x304 (ite $x72 ?1 ?x299)))
+(let ((?x82 (|mod$| ?1 ?0)))
+(let (($x307 (= ?x82 ?x304)))
+(let (($x89 (forall ((?v0 Int) (?v1 Int) )(let (($x72 (= ?v1 0)))
+(let ((?x87 (ite $x72 ?v0 (ite (< 0 ?v1) (mod ?v0 ?v1) (- (mod (- ?v0) (- ?v1)))))))
+(let ((?x82 (|mod$| ?v0 ?v1)))
+(= ?x82 ?x87)))))
+))
+(let (($x293 (forall ((?v0 Int) (?v1 Int) )(let ((?x230 (* (~ 1) ?v1)))
+(let ((?x227 (* (~ 1) ?v0)))
+(let ((?x273 (mod ?x227 ?x230)))
+(let ((?x279 (* (~ 1) ?x273)))
+(let ((?x83 (mod ?v0 ?v1)))
+(let (($x73 (< 0 ?v1)))
+(let ((?x284 (ite $x73 ?x83 ?x279)))
+(let (($x72 (= ?v1 0)))
+(let ((?x287 (ite $x72 ?v0 ?x284)))
+(let ((?x82 (|mod$| ?v0 ?v1)))
+(= ?x82 ?x287))))))))))))
+))
+(let ((@x298 (monotonicity (rewrite (= (< 0 ?0) (not $x248))) (= (ite (< 0 ?0) ?x83 ?x279) (ite (not $x248) ?x83 ?x279)))))
+(let ((@x303 (trans @x298 (rewrite (= (ite (not $x248) ?x83 ?x279) ?x299)) (= (ite (< 0 ?0) ?x83 ?x279) ?x299))))
+(let ((@x306 (monotonicity @x303 (= (ite $x72 ?1 (ite (< 0 ?0) ?x83 ?x279)) ?x304))))
+(let ((@x309 (monotonicity @x306 (= (= ?x82 (ite $x72 ?1 (ite (< 0 ?0) ?x83 ?x279))) $x307))))
+(let (($x73 (< 0 ?0)))
+(let ((?x284 (ite $x73 ?x83 ?x279)))
+(let ((?x287 (ite $x72 ?1 ?x284)))
+(let (($x290 (= ?x82 ?x287)))
+(let (($x291 (= (= ?x82 (ite $x72 ?1 (ite $x73 ?x83 (- (mod (- ?1) (- ?0)))))) $x290)))
+(let ((@x275 (monotonicity (rewrite (= (- ?1) ?x227)) (rewrite (= (- ?0) ?x230)) (= (mod (- ?1) (- ?0)) ?x273))))
+(let ((@x283 (trans (monotonicity @x275 (= (- (mod (- ?1) (- ?0))) (- ?x273))) (rewrite (= (- ?x273) ?x279)) (= (- (mod (- ?1) (- ?0))) ?x279))))
+(let ((@x286 (monotonicity @x283 (= (ite $x73 ?x83 (- (mod (- ?1) (- ?0)))) ?x284))))
+(let ((@x289 (monotonicity @x286 (= (ite $x72 ?1 (ite $x73 ?x83 (- (mod (- ?1) (- ?0))))) ?x287))))
+(let ((@x314 (trans (|quant-intro| (monotonicity @x289 $x291) (= $x89 $x293)) (|quant-intro| @x309 (= $x293 $x310)) (= $x89 $x310))))
+(let ((@x360 (|mp~| (mp (asserted $x89) @x314 $x310) (|nnf-pos| (refl (|~| $x307 $x307)) (|~| $x310 $x310)) $x310)))
+(let ((@x423 (mp @x360 (|quant-intro| (refl (= $x307 $x307)) (= $x310 $x418)) $x418)))
+(let (($x633 (not $x418)))
+(let (($x634 (or $x633 $x628)))
+(let (($x440 (<= 2 0)))
+(let ((?x600 (ite $x440 (* (~ 1) (mod ?x452 (* (~ 1) 2))) ?x599)))
+(let (($x439 (= 2 0)))
+(let ((?x601 (ite $x439 |l$| ?x600)))
+(let (($x602 (= ?x45 ?x601)))
+(let ((@x457 (rewrite (= (* (~ 1) 2) (~ 2)))))
+(let ((@x608 (monotonicity (monotonicity @x457 (= (mod ?x452 (* (~ 1) 2)) (mod ?x452 (~ 2)))) (= (* (~ 1) (mod ?x452 (* (~ 1) 2))) (* (~ 1) (mod ?x452 (~ 2)))))))
+(let ((@x451 (rewrite (= $x440 false))))
+(let ((@x611 (monotonicity @x451 @x608 (= ?x600 (ite false (* (~ 1) (mod ?x452 (~ 2))) ?x599)))))
+(let ((@x615 (trans @x611 (rewrite (= (ite false (* (~ 1) (mod ?x452 (~ 2))) ?x599) ?x599)) (= ?x600 ?x599))))
+(let ((@x449 (rewrite (= $x439 false))))
+(let ((@x622 (trans (monotonicity @x449 @x615 (= ?x601 (ite false |l$| ?x599))) (rewrite (= (ite false |l$| ?x599) ?x599)) (= ?x601 ?x599))))
+(let ((@x632 (trans (monotonicity @x622 (= $x602 (= ?x45 ?x599))) (rewrite (= (= ?x45 ?x599) $x628)) (= $x602 $x628))))
+(let ((@x641 (trans (monotonicity @x632 (= (or $x633 $x602) $x634)) (rewrite (= $x634 $x634)) (= (or $x633 $x602) $x634))))
+(let ((@x950 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x628) (>= ?x627 0))) (|unit-resolution| (mp ((_ |quant-inst| |l$| 2) (or $x633 $x602)) @x641 $x634) @x423 $x628) (>= ?x627 0))))
+(let (($x1021 (not $x778)))
+(let (($x777 (<= ?x776 0)))
+(let (($x770 (<= ?x769 0)))
+(let (($x364 (not $x363)))
+(let ((@x741 (hypothesis $x364)))
+(let ((@x1018 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or $x362 $x770)) (|unit-resolution| (|def-axiom| (or $x363 $x152)) @x741 $x152) $x770)))
+(let ((?x520 (+ |l$| ?x145 (* (~ 2) (div ?x146 2)) (* (~ 1) (mod (+ |l$| ?x43) 2)))))
+(let (($x517 (= ?x520 0)))
+(let ((@x876 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x517) (>= ?x520 0))) (|unit-resolution| ((_ |th-lemma| arith) (or false $x517)) @x108 $x517) (>= ?x520 0))))
+(let ((?x584 (* (~ 2) ?x48)))
+(let ((?x585 (+ ?x39 ?x145 ?x584)))
+(let (($x586 (= ?x585 0)))
+(let (($x384 (forall ((?v0 |Int_list$|) (?v1 |Nat_list$|) )(!(let ((?x24 (|eval_dioph$| ?v0 ?v1)))
+(let ((?x136 (+ ?x24 (* (~ 1) (|eval_dioph$| ?v0 (|map$| |uu$| ?v1))) (* (~ 2) (|eval_dioph$| ?v0 (|map$| |uua$| ?v1))))))
+(= ?x136 0))) :pattern ( (|eval_dioph$| ?v0 (|map$| |uu$| ?v1)) ) :pattern ( (|eval_dioph$| ?v0 (|map$| |uua$| ?v1)) )))
+))
+(let (($x138 (forall ((?v0 |Int_list$|) (?v1 |Nat_list$|) )(let ((?x24 (|eval_dioph$| ?v0 ?v1)))
+(let ((?x136 (+ ?x24 (* (~ 1) (|eval_dioph$| ?v0 (|map$| |uu$| ?v1))) (* (~ 2) (|eval_dioph$| ?v0 (|map$| |uua$| ?v1))))))
+(= ?x136 0))))
+))
+(let ((?x24 (|eval_dioph$| ?1 ?0)))
+(let ((?x136 (+ ?x24 (* (~ 1) (|eval_dioph$| ?1 (|map$| |uu$| ?0))) (* (~ 2) (|eval_dioph$| ?1 (|map$| |uua$| ?0))))))
+(let (($x132 (= ?x136 0)))
+(let (($x36 (forall ((?v0 |Int_list$|) (?v1 |Nat_list$|) )(let ((?x24 (|eval_dioph$| ?v0 ?v1)))
+(let ((?x27 (|eval_dioph$| ?v0 (|map$| |uu$| ?v1))))
+(let ((?x34 (+ (* (|eval_dioph$| ?v0 (|map$| |uua$| ?v1)) 2) ?x27)))
+(= ?x34 ?x24)))))
+))
+(let (($x127 (forall ((?v0 |Int_list$|) (?v1 |Nat_list$|) )(let ((?x24 (|eval_dioph$| ?v0 ?v1)))
+(let ((?x32 (|eval_dioph$| ?v0 (|map$| |uua$| ?v1))))
+(let ((?x113 (* 2 ?x32)))
+(let ((?x27 (|eval_dioph$| ?v0 (|map$| |uu$| ?v1))))
+(let ((?x119 (+ ?x27 ?x113)))
+(= ?x119 ?x24)))))))
+))
+(let ((?x32 (|eval_dioph$| ?1 (|map$| |uua$| ?0))))
+(let ((?x113 (* 2 ?x32)))
+(let ((?x27 (|eval_dioph$| ?1 (|map$| |uu$| ?0))))
+(let ((?x119 (+ ?x27 ?x113)))
+(let (($x124 (= ?x119 ?x24)))
+(let ((@x118 (monotonicity (rewrite (= (* ?x32 2) ?x113)) (= (+ (* ?x32 2) ?x27) (+ ?x113 ?x27)))))
+(let ((@x123 (trans @x118 (rewrite (= (+ ?x113 ?x27) ?x119)) (= (+ (* ?x32 2) ?x27) ?x119))))
+(let ((@x129 (|quant-intro| (monotonicity @x123 (= (= (+ (* ?x32 2) ?x27) ?x24) $x124)) (= $x36 $x127))))
+(let ((@x142 (trans @x129 (|quant-intro| (rewrite (= $x124 $x132)) (= $x127 $x138)) (= $x36 $x138))))
+(let ((@x335 (|mp~| (mp (asserted $x36) @x142 $x138) (|nnf-pos| (refl (|~| $x132 $x132)) (|~| $x138 $x138)) $x138)))
+(let ((@x389 (mp @x335 (|quant-intro| (refl (= $x132 $x132)) (= $x138 $x384)) $x384)))
+(let ((@x883 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x586) (<= ?x585 0))) (|unit-resolution| ((_ |quant-inst| |ks$| |xs$|) (or (not $x384) $x586)) @x389 $x586) (<= ?x585 0))))
+(let ((?x479 (+ ?x149 (* (~ 1) (div ?x146 2)))))
+(let (($x480 (= ?x479 0)))
+(let (($x411 (forall ((?v0 Int) (?v1 Int) )(!(let ((?x74 (div ?v0 ?v1)))
+(let ((?x230 (* (~ 1) ?v1)))
+(let ((?x227 (* (~ 1) ?v0)))
+(let ((?x233 (div ?x227 ?x230)))
+(let (($x248 (<= ?v1 0)))
+(let ((?x255 (ite $x248 ?x233 ?x74)))
+(let (($x72 (= ?v1 0)))
+(let ((?x71 (|div$| ?v0 ?v1)))
+(= ?x71 (ite $x72 0 ?x255)))))))))) :pattern ( (|div$| ?v0 ?v1) )))
+))
+(let (($x266 (forall ((?v0 Int) (?v1 Int) )(let ((?x74 (div ?v0 ?v1)))
+(let ((?x230 (* (~ 1) ?v1)))
+(let ((?x227 (* (~ 1) ?v0)))
+(let ((?x233 (div ?x227 ?x230)))
+(let (($x248 (<= ?v1 0)))
+(let ((?x255 (ite $x248 ?x233 ?x74)))
+(let (($x72 (= ?v1 0)))
+(let ((?x71 (|div$| ?v0 ?v1)))
+(= ?x71 (ite $x72 0 ?x255)))))))))))
+))
+(let ((?x71 (|div$| ?1 ?0)))
+(let (($x263 (= ?x71 (ite $x72 0 (ite $x248 (div ?x227 ?x230) (div ?1 ?0))))))
+(let (($x81 (forall ((?v0 Int) (?v1 Int) )(let (($x72 (= ?v1 0)))
+(let ((?x79 (ite $x72 0 (ite (< 0 ?v1) (div ?v0 ?v1) (div (- ?v0) (- ?v1))))))
+(let ((?x71 (|div$| ?v0 ?v1)))
+(= ?x71 ?x79)))))
+))
+(let (($x245 (forall ((?v0 Int) (?v1 Int) )(let ((?x230 (* (~ 1) ?v1)))
+(let ((?x227 (* (~ 1) ?v0)))
+(let ((?x233 (div ?x227 ?x230)))
+(let ((?x74 (div ?v0 ?v1)))
+(let (($x73 (< 0 ?v1)))
+(let ((?x236 (ite $x73 ?x74 ?x233)))
+(let (($x72 (= ?v1 0)))
+(let ((?x239 (ite $x72 0 ?x236)))
+(let ((?x71 (|div$| ?v0 ?v1)))
+(= ?x71 ?x239)))))))))))
+))
+(let ((?x233 (div ?x227 ?x230)))
+(let ((?x74 (div ?1 ?0)))
+(let ((?x236 (ite $x73 ?x74 ?x233)))
+(let ((?x239 (ite $x72 0 ?x236)))
+(let (($x242 (= ?x71 ?x239)))
+(let ((@x254 (monotonicity (rewrite (= $x73 (not $x248))) (= ?x236 (ite (not $x248) ?x74 ?x233)))))
+(let ((@x259 (trans @x254 (rewrite (= (ite (not $x248) ?x74 ?x233) (ite $x248 ?x233 ?x74))) (= ?x236 (ite $x248 ?x233 ?x74)))))
+(let ((@x265 (monotonicity (monotonicity @x259 (= ?x239 (ite $x72 0 (ite $x248 ?x233 ?x74)))) (= $x242 $x263))))
+(let (($x243 (= (= ?x71 (ite $x72 0 (ite $x73 ?x74 (div (- ?1) (- ?0))))) $x242)))
+(let ((@x235 (monotonicity (rewrite (= (- ?1) ?x227)) (rewrite (= (- ?0) ?x230)) (= (div (- ?1) (- ?0)) ?x233))))
+(let ((@x241 (monotonicity (monotonicity @x235 (= (ite $x73 ?x74 (div (- ?1) (- ?0))) ?x236)) (= (ite $x72 0 (ite $x73 ?x74 (div (- ?1) (- ?0)))) ?x239))))
+(let ((@x270 (trans (|quant-intro| (monotonicity @x241 $x243) (= $x81 $x245)) (|quant-intro| @x265 (= $x245 $x266)) (= $x81 $x266))))
+(let ((@x355 (|mp~| (mp (asserted $x81) @x270 $x266) (|nnf-pos| (refl (|~| $x263 $x263)) (|~| $x266 $x266)) $x266)))
+(let ((@x416 (mp @x355 (|quant-intro| (refl (= $x263 $x263)) (= $x266 $x411)) $x411)))
+(let (($x486 (or (not $x411) $x480)))
+(let ((?x444 (div ?x146 2)))
+(let ((?x445 (ite $x440 (div (* (~ 1) ?x146) (* (~ 1) 2)) ?x444)))
+(let ((?x446 (ite $x439 0 ?x445)))
+(let (($x447 (= ?x149 ?x446)))
+(let ((@x460 (monotonicity (rewrite (= (* (~ 1) ?x146) (+ ?x452 ?x43))) @x457 (= (div (* (~ 1) ?x146) (* (~ 1) 2)) (div (+ ?x452 ?x43) (~ 2))))))
+(let ((@x463 (monotonicity @x451 @x460 (= ?x445 (ite false (div (+ ?x452 ?x43) (~ 2)) ?x444)))))
+(let ((@x467 (trans @x463 (rewrite (= (ite false (div (+ ?x452 ?x43) (~ 2)) ?x444) ?x444)) (= ?x445 ?x444))))
+(let ((@x474 (trans (monotonicity @x449 @x467 (= ?x446 (ite false 0 ?x444))) (rewrite (= (ite false 0 ?x444) ?x444)) (= ?x446 ?x444))))
+(let ((@x484 (trans (monotonicity @x474 (= $x447 (= ?x149 ?x444))) (rewrite (= (= ?x149 ?x444) $x480)) (= $x447 $x480))))
+(let ((@x493 (trans (monotonicity @x484 (= (or (not $x411) $x447) $x486)) (rewrite (= $x486 $x486)) (= (or (not $x411) $x447) $x486))))
+(let ((@x885 (|unit-resolution| (mp ((_ |quant-inst| (+ |l$| ?x145) 2) (or (not $x411) $x447)) @x493 $x486) @x416 $x480)))
+(let ((@x889 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x480) (<= ?x479 0))) @x885 (<= ?x479 0))))
+(let ((@x892 (|unit-resolution| ((_ |th-lemma| arith) (or false (>= (mod (+ |l$| ?x43) 2) 0))) @x108 (>= (mod (+ |l$| ?x43) 2) 0))))
+(let ((@x893 ((_ |th-lemma| arith farkas 1 -2 -2 -1 1 1) @x892 @x889 (hypothesis $x770) @x883 (hypothesis (not $x777)) @x876 false)))
+(let (($x169 (not $x41)))
+(let (($x370 (= $x41 $x363)))
+(let ((@x369 (monotonicity (rewrite (= (and $x46 $x152) $x364)) (= (= $x169 (and $x46 $x152)) (= $x169 $x364)))))
+(let ((@x374 (trans @x369 (rewrite (= (= $x169 $x364) $x370)) (= (= $x169 (and $x46 $x152)) $x370))))
+(let (($x155 (and $x46 $x152)))
+(let (($x170 (= $x169 $x155)))
+(let (($x53 (= $x41 (and $x46 (= ?x48 (|div$| (- |l$| ?x43) 2))))))
+(let (($x54 (not $x53)))
+(let ((@x151 (monotonicity (rewrite (= (- |l$| ?x43) ?x146)) (= (|div$| (- |l$| ?x43) 2) ?x149))))
+(let ((@x157 (monotonicity (monotonicity @x151 (= (= ?x48 (|div$| (- |l$| ?x43) 2)) $x152)) (= (and $x46 (= ?x48 (|div$| (- |l$| ?x43) 2))) $x155))))
+(let ((@x165 (trans (monotonicity @x157 (= $x53 (= $x41 $x155))) (rewrite (= (= $x41 $x155) (= $x41 $x155))) (= $x53 (= $x41 $x155)))))
+(let ((@x174 (trans (monotonicity @x165 (= $x54 (not (= $x41 $x155)))) (rewrite (= (not (= $x41 $x155)) $x170)) (= $x54 $x170))))
+(let ((@x375 (mp (mp (asserted $x54) @x174 $x170) @x374 $x370)))
+(let ((@x438 (|unit-resolution| (|def-axiom| (or $x169 $x363 (not $x370))) @x375 (or $x169 $x363))))
+(let ((@x1025 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or $x41 (not $x777) $x1021)) (|unit-resolution| @x438 @x741 $x169) (or (not $x777) $x1021))))
+(let ((@x1026 (|unit-resolution| @x1025 (|unit-resolution| (lemma @x893 (or $x777 (not $x770))) @x1018 $x777) $x1021)))
+(let ((@x1029 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or $x361 (>= (+ ?x44 (* (~ 1) ?x45)) 0))) (|unit-resolution| (|def-axiom| (or $x363 $x46)) @x741 $x46) (>= (+ ?x44 (* (~ 1) ?x45)) 0))))
+(let ((?x744 (+ ?x43 ?x712 (* (~ 2) ?x730))))
+(let (($x742 (= ?x744 0)))
+(let ((@x1032 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x742) (>= ?x744 0))) (|unit-resolution| ((_ |th-lemma| arith) (or false $x742)) @x108 $x742) (>= ?x744 0))))
+(let ((?x713 (+ ?x44 ?x712)))
+(let (($x714 (= ?x713 0)))
+(let (($x719 (or $x633 $x714)))
+(let ((?x686 (ite $x440 (* (~ 1) (mod ?x145 (* (~ 1) 2))) ?x685)))
+(let ((?x687 (ite $x439 ?x43 ?x686)))
+(let (($x688 (= ?x44 ?x687)))
+(let ((@x694 (monotonicity (monotonicity @x457 (= (mod ?x145 (* (~ 1) 2)) (mod ?x145 (~ 2)))) (= (* (~ 1) (mod ?x145 (* (~ 1) 2))) (* (~ 1) (mod ?x145 (~ 2)))))))
+(let ((@x697 (monotonicity @x451 @x694 (= ?x686 (ite false (* (~ 1) (mod ?x145 (~ 2))) ?x685)))))
+(let ((@x701 (trans @x697 (rewrite (= (ite false (* (~ 1) (mod ?x145 (~ 2))) ?x685) ?x685)) (= ?x686 ?x685))))
+(let ((@x708 (trans (monotonicity @x449 @x701 (= ?x687 (ite false ?x43 ?x685))) (rewrite (= (ite false ?x43 ?x685) ?x685)) (= ?x687 ?x685))))
+(let ((@x718 (trans (monotonicity @x708 (= $x688 (= ?x44 ?x685))) (rewrite (= (= ?x44 ?x685) $x714)) (= $x688 $x714))))
+(let ((@x726 (trans (monotonicity @x718 (= (or $x633 $x688) $x719)) (rewrite (= $x719 $x719)) (= (or $x633 $x688) $x719))))
+(let ((@x1035 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x714) (>= ?x713 0))) (|unit-resolution| (mp ((_ |quant-inst| (|eval_dioph$| |ks$| ?x42) 2) (or $x633 $x688)) @x726 $x719) @x423 $x714) (>= ?x713 0))))
+(let ((@x992 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x586) (>= ?x585 0))) (|unit-resolution| ((_ |quant-inst| |ks$| |xs$|) (or (not $x384) $x586)) @x389 $x586) (>= ?x585 0))))
+(let ((?x773 (+ ?x44 (* (~ 1) ?x45))))
+(let (($x774 (<= ?x773 0)))
+(let ((@x1010 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or $x361 $x774)) (|unit-resolution| (|def-axiom| (or $x363 $x46)) @x741 $x46) $x774)))
+(let ((@x1014 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or $x362 $x771)) (|unit-resolution| (|def-axiom| (or $x363 $x152)) @x741 $x152) $x771)))
+(let ((@x963 (|unit-resolution| ((_ |th-lemma| arith) (or false (not (>= (mod (+ |l$| ?x43) 2) 2)))) @x108 (not (>= (mod (+ |l$| ?x43) 2) 2)))))
+(let ((@x748 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x628) (<= ?x627 0))) (|unit-resolution| (mp ((_ |quant-inst| |l$| 2) (or $x633 $x602)) @x641 $x634) @x423 $x628) (<= ?x627 0))))
+(let ((@x932 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x742) (<= ?x744 0))) (|unit-resolution| ((_ |th-lemma| arith) (or false $x742)) @x108 $x742) (<= ?x744 0))))
+(let ((@x852 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x657) (<= ?x659 0))) (|unit-resolution| ((_ |th-lemma| arith) (or false $x657)) @x108 $x657) (<= ?x659 0))))
+(let ((@x937 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x714) (<= ?x713 0))) (|unit-resolution| (mp ((_ |quant-inst| (|eval_dioph$| |ks$| ?x42) 2) (or $x633 $x688)) @x726 $x719) @x423 $x714) (<= ?x713 0))))
+(let ((@x954 (hypothesis $x771)))
+(let ((@x957 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x480) (>= ?x479 0))) @x885 (>= ?x479 0))))
+(let ((@x960 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x517) (<= ?x520 0))) (|unit-resolution| ((_ |th-lemma| arith) (or false $x517)) @x108 $x517) (<= ?x520 0))))
+(let ((@x515 ((_ |th-lemma| arith farkas 1 -2 -2 -2 1 1 1 1 1 1) @x960 @x957 @x954 (hypothesis $x737) @x937 @x852 @x932 (hypothesis $x774) @x748 @x963 false)))
+(let ((@x1015 (|unit-resolution| (lemma @x515 (or (not $x737) (not $x771) (not $x774))) @x1014 @x1010 (not $x737))))
+(let ((@x1037 (|unit-resolution| @x1015 ((_ |th-lemma| arith) @x992 @x1035 @x1032 @x1029 @x1026 @x950 @x945 $x737) false)))
+(let ((@x1038 (lemma @x1037 $x363)))
+(let ((@x434 (|unit-resolution| (|def-axiom| (or $x41 $x364 (not $x370))) @x375 (or $x41 $x364))))
+(let ((@x1120 (|unit-resolution| @x434 @x1038 $x41)))
+(let ((@x1125 ((_ |th-lemma| arith farkas 2 2 1 1 1 1) (hypothesis (not $x771)) @x889 @x892 @x876 @x883 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or $x169 $x778)) @x1120 $x778) false)))
+(let ((@x1048 ((_ |th-lemma| arith farkas -1 1 1 -2 -2 1) @x992 (|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or $x169 $x777)) @x1120 $x777) @x960 @x957 (hypothesis (not $x770)) @x963 false)))
+(let ((?x587 (|mod$| ?x39 2)))
+(let (($x588 (= ?x587 ?x44)))
+(let (($x377 (forall ((?v0 |Int_list$|) (?v1 |Nat_list$|) )(!(= (|mod$| (|eval_dioph$| ?v0 ?v1) 2) (|mod$| (|eval_dioph$| ?v0 (|map$| |uu$| ?v1)) 2)) :pattern ( (|eval_dioph$| ?v0 (|map$| |uu$| ?v1)) )))
+))
+(let (($x30 (forall ((?v0 |Int_list$|) (?v1 |Nat_list$|) )(= (|mod$| (|eval_dioph$| ?v0 ?v1) 2) (|mod$| (|eval_dioph$| ?v0 (|map$| |uu$| ?v1)) 2)))
+))
+(let (($x29 (= (|mod$| ?x24 2) (|mod$| ?x27 2))))
+(let ((@x330 (|mp~| (asserted $x30) (|nnf-pos| (refl (|~| $x29 $x29)) (|~| $x30 $x30)) $x30)))
+(let ((@x382 (mp @x330 (|quant-intro| (refl (= $x29 $x29)) (= $x30 $x377)) $x377)))
+(let ((@x1104 (symm (|unit-resolution| ((_ |quant-inst| |ks$| |xs$|) (or (not $x377) $x588)) @x382 $x588) (= ?x44 ?x587))))
+(let ((@x763 (|unit-resolution| (hypothesis $x361) (trans @x1104 (monotonicity @x1120 (= ?x587 ?x45)) $x46) false)))
+(let ((@x1050 (|unit-resolution| (|unit-resolution| (|def-axiom| (or $x364 $x361 $x362)) @x1038 $x363) (lemma @x763 $x46) $x362)))
+(|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or $x152 (not $x770) (not $x771))) @x1050 (lemma @x1048 $x770) (lemma @x1125 $x771) false)))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))))
+
+d4d0e08ac1741a77a8448ec3a55e48fb2a240ee9 62 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let ((?x32 (|collect$| |uu$|)))
+(let ((?x33 (|sup$| ?x32)))
+(let (($x38 (|less_eq$| ?x33 ?x33)))
+(let (($x39 (not $x38)))
+(let ((@x117 (asserted $x39)))
+(let ((?x34 (|collect$| |uua$|)))
+(let ((?x35 (|sup$| ?x34)))
+(let (($x37 (|less_eq$| ?x35 ?x33)))
+(let ((@x115 (asserted $x37)))
+(let (($x36 (|less_eq$| ?x33 ?x35)))
+(let ((@x114 (asserted $x36)))
+(let (($x159 (forall ((?v0 |A$|) (?v1 |A$|) (?v2 |A$|) )(!(let (($x29 (|less_eq$| ?v0 ?v2)))
+(let (($x27 (|less_eq$| ?v1 ?v2)))
+(let (($x136 (not $x27)))
+(let (($x26 (|less_eq$| ?v0 ?v1)))
+(let (($x135 (not $x26)))
+(or $x135 $x136 $x29)))))) :pattern ( (|less_eq$| ?v0 ?v1) (|less_eq$| ?v1 ?v2) )))
+))
+(let (($x154 (forall ((?v0 |A$|) (?v1 |A$|) (?v2 |A$|) )(let (($x29 (|less_eq$| ?v0 ?v2)))
+(let (($x27 (|less_eq$| ?v1 ?v2)))
+(let (($x136 (not $x27)))
+(let (($x26 (|less_eq$| ?v0 ?v1)))
+(let (($x135 (not $x26)))
+(or $x135 $x136 $x29)))))))
+))
+(let (($x29 (|less_eq$| ?2 ?0)))
+(let (($x27 (|less_eq$| ?1 ?0)))
+(let (($x136 (not $x27)))
+(let (($x26 (|less_eq$| ?2 ?1)))
+(let (($x135 (not $x26)))
+(let (($x149 (or $x135 $x136 $x29)))
+(let (($x111 (forall ((?v0 |A$|) (?v1 |A$|) (?v2 |A$|) )(let (($x29 (|less_eq$| ?v0 ?v2)))
+(let (($x27 (|less_eq$| ?v1 ?v2)))
+(let (($x26 (|less_eq$| ?v0 ?v1)))
+(let (($x28 (and $x26 $x27)))
+(let (($x106 (not $x28)))
+(or $x106 $x29)))))))
+))
+(let ((@x141 (monotonicity (rewrite (= (and $x26 $x27) (not (or $x135 $x136)))) (= (not (and $x26 $x27)) (not (not (or $x135 $x136)))))))
+(let ((@x145 (trans @x141 (rewrite (= (not (not (or $x135 $x136))) (or $x135 $x136))) (= (not (and $x26 $x27)) (or $x135 $x136)))))
+(let ((@x148 (monotonicity @x145 (= (or (not (and $x26 $x27)) $x29) (or (or $x135 $x136) $x29)))))
+(let ((@x153 (trans @x148 (rewrite (= (or (or $x135 $x136) $x29) $x149)) (= (or (not (and $x26 $x27)) $x29) $x149))))
+(let ((@x129 (refl (|~| (or (not (and $x26 $x27)) $x29) (or (not (and $x26 $x27)) $x29)))))
+(let (($x31 (forall ((?v0 |A$|) (?v1 |A$|) (?v2 |A$|) )(let (($x29 (|less_eq$| ?v0 ?v2)))
+(let (($x27 (|less_eq$| ?v1 ?v2)))
+(let (($x26 (|less_eq$| ?v0 ?v1)))
+(let (($x28 (and $x26 $x27)))
+(=> $x28 $x29))))))
+))
+(let ((@x110 (rewrite (= (=> (and $x26 $x27) $x29) (or (not (and $x26 $x27)) $x29)))))
+(let ((@x132 (|mp~| (mp (asserted $x31) (|quant-intro| @x110 (= $x31 $x111)) $x111) (|nnf-pos| @x129 (|~| $x111 $x111)) $x111)))
+(let ((@x164 (mp (mp @x132 (|quant-intro| @x153 (= $x111 $x154)) $x154) (|quant-intro| (refl (= $x149 $x149)) (= $x154 $x159)) $x159)))
+(let (($x166 (not $x37)))
+(let (($x165 (not $x36)))
+(let (($x170 (not $x159)))
+(let (($x171 (or $x170 $x165 $x166 $x38)))
+(let ((@x176 (mp ((_ |quant-inst| (|sup$| ?x32) (|sup$| ?x34) (|sup$| ?x32)) (or $x170 (or $x165 $x166 $x38))) (rewrite (= (or $x170 (or $x165 $x166 $x38)) $x171)) $x171)))
+(|unit-resolution| @x176 @x164 @x114 @x115 @x117 false)))))))))))))))))))))))))))))))))))))
+
+ce2ba5128c1bc3cef10c9328de9b15558b908319 25 0
+unsat
+((set-logic AUFLIA)
+(proof
+(let (($x51 (|pred$e| 1)))
+(let (($x52 (not $x51)))
+(let ((@x142 (asserted $x52)))
+(let (($x198 (forall ((?v0 Int) )(!(|pred$e| ?v0) :pattern ( (|pred$e| ?v0) )))
+))
+(let (($x139 (forall ((?v0 Int) )(|pred$e| ?v0))
+))
+(let (($x49 (forall ((?v0 Int) )(let (($x47 (or (|pred$d| (|cons$d| ?v0 |nil$d|)) (not (|pred$d| (|cons$d| ?v0 |nil$d|))))))
+(let (($x42 (|pred$e| ?v0)))
+(and $x42 $x47))))
+))
+(let (($x42 (|pred$e| ?0)))
+(let (($x47 (or (|pred$d| (|cons$d| ?0 |nil$d|)) (not (|pred$d| (|cons$d| ?0 |nil$d|))))))
+(let (($x48 (and $x42 $x47)))
+(let ((@x134 (monotonicity (rewrite (= $x47 true)) (= $x48 (and $x42 true)))))
+(let ((@x141 (|quant-intro| (trans @x134 (rewrite (= (and $x42 true) $x42)) (= $x48 $x42)) (= $x49 $x139))))
+(let ((@x168 (|mp~| (mp (asserted $x49) @x141 $x139) (|nnf-pos| (refl (|~| $x42 $x42)) (|~| $x139 $x139)) $x139)))
+(let ((@x203 (mp @x168 (|quant-intro| (refl (= $x42 $x42)) (= $x139 $x198)) $x198)))
+(let (($x207 (or (not $x198) $x51)))
+(let ((@x208 ((_ |quant-inst| 1) $x207)))
+(|unit-resolution| @x208 @x203 @x142 false))))))))))))))))))
+
--- a/src/HOL/SMT_Examples/SMT_Examples.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/SMT_Examples/SMT_Examples.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -11,38 +11,34 @@
 declare [[smt_certificates = "SMT_Examples.certs"]]
 declare [[smt_read_only_certificates = true]]
 
+declare [[smt2_certificates = "SMT_Examples.certs2"]]
+declare [[smt2_read_only_certificates = true]]
 
 
 section {* Propositional and first-order logic *}
 
-lemma "True" by smt
-
-lemma "p \<or> \<not>p" by smt
-
-lemma "(p \<and> True) = p" by smt
-
-lemma "(p \<or> q) \<and> \<not>p \<Longrightarrow> q" by smt
-
-lemma "(a \<and> b) \<or> (c \<and> d) \<Longrightarrow> (a \<and> b) \<or> (c \<and> d)"
-  by smt
-
-lemma "(p1 \<and> p2) \<or> p3 \<longrightarrow> (p1 \<longrightarrow> (p3 \<and> p2) \<or> (p1 \<and> p3)) \<or> p1" by smt
-
-lemma "P=P=P=P=P=P=P=P=P=P" by smt
+lemma "True" by smt2
+lemma "p \<or> \<not>p" by smt2
+lemma "(p \<and> True) = p" by smt2
+lemma "(p \<or> q) \<and> \<not>p \<Longrightarrow> q" by smt2
+lemma "(a \<and> b) \<or> (c \<and> d) \<Longrightarrow> (a \<and> b) \<or> (c \<and> d)" by smt2
+lemma "(p1 \<and> p2) \<or> p3 \<longrightarrow> (p1 \<longrightarrow> (p3 \<and> p2) \<or> (p1 \<and> p3)) \<or> p1" by smt2
+lemma "P = P = P = P = P = P = P = P = P = P" by smt2
 
 lemma
-  assumes "a | b | c | d"
-      and "e | f | (a & d)"
-      and "~(a | (c & ~c)) | b"
-      and "~(b & (x | ~x)) | c"
-      and "~(d | False) | c"
-      and "~(c | (~p & (p | (q & ~q))))"
+  assumes "a \<or> b \<or> c \<or> d"
+      and "e \<or> f \<or> (a \<and> d)"
+      and "\<not> (a \<or> (c \<and> ~c)) \<or> b"
+      and "\<not> (b \<and> (x \<or> \<not> x)) \<or> c"
+      and "\<not> (d \<or> False) \<or> c"
+      and "\<not> (c \<or> (\<not> p \<and> (p \<or> (q \<and> \<not> q))))"
   shows False
-  using assms by smt
+  using assms by smt2
 
 axiomatization symm_f :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" where
   symm_f: "symm_f x y = symm_f y x"
-lemma "a = a \<and> symm_f a b = symm_f b a" by (smt symm_f)
+
+lemma "a = a \<and> symm_f a b = symm_f b a" by (smt2 symm_f)
 
 (*
 Taken from ~~/src/HOL/ex/SAT_Examples.thy.
@@ -53,253 +49,246 @@
   and "~x30"
   and "~x29"
   and "~x59"
-  and "x1 | x31 | x0"
-  and "x2 | x32 | x1"
-  and "x3 | x33 | x2"
-  and "x4 | x34 | x3"
-  and "x35 | x4"
-  and "x5 | x36 | x30"
-  and "x6 | x37 | x5 | x31"
-  and "x7 | x38 | x6 | x32"
-  and "x8 | x39 | x7 | x33"
-  and "x9 | x40 | x8 | x34"
-  and "x41 | x9 | x35"
-  and "x10 | x42 | x36"
-  and "x11 | x43 | x10 | x37"
-  and "x12 | x44 | x11 | x38"
-  and "x13 | x45 | x12 | x39"
-  and "x14 | x46 | x13 | x40"
-  and "x47 | x14 | x41"
-  and "x15 | x48 | x42"
-  and "x16 | x49 | x15 | x43"
-  and "x17 | x50 | x16 | x44"
-  and "x18 | x51 | x17 | x45"
-  and "x19 | x52 | x18 | x46"
-  and "x53 | x19 | x47"
-  and "x20 | x54 | x48"
-  and "x21 | x55 | x20 | x49"
-  and "x22 | x56 | x21 | x50"
-  and "x23 | x57 | x22 | x51"
-  and "x24 | x58 | x23 | x52"
-  and "x59 | x24 | x53"
-  and "x25 | x54"
-  and "x26 | x25 | x55"
-  and "x27 | x26 | x56"
-  and "x28 | x27 | x57"
-  and "x29 | x28 | x58"
-  and "~x1 | ~x31"
-  and "~x1 | ~x0"
-  and "~x31 | ~x0"
-  and "~x2 | ~x32"
-  and "~x2 | ~x1"
-  and "~x32 | ~x1"
-  and "~x3 | ~x33"
-  and "~x3 | ~x2"
-  and "~x33 | ~x2"
-  and "~x4 | ~x34"
-  and "~x4 | ~x3"
-  and "~x34 | ~x3"
-  and "~x35 | ~x4"
-  and "~x5 | ~x36"
-  and "~x5 | ~x30"
-  and "~x36 | ~x30"
-  and "~x6 | ~x37"
-  and "~x6 | ~x5"
-  and "~x6 | ~x31"
-  and "~x37 | ~x5"
-  and "~x37 | ~x31"
-  and "~x5 | ~x31"
-  and "~x7 | ~x38"
-  and "~x7 | ~x6"
-  and "~x7 | ~x32"
-  and "~x38 | ~x6"
-  and "~x38 | ~x32"
-  and "~x6 | ~x32"
-  and "~x8 | ~x39"
-  and "~x8 | ~x7"
-  and "~x8 | ~x33"
-  and "~x39 | ~x7"
-  and "~x39 | ~x33"
-  and "~x7 | ~x33"
-  and "~x9 | ~x40"
-  and "~x9 | ~x8"
-  and "~x9 | ~x34"
-  and "~x40 | ~x8"
-  and "~x40 | ~x34"
-  and "~x8 | ~x34"
-  and "~x41 | ~x9"
-  and "~x41 | ~x35"
-  and "~x9 | ~x35"
-  and "~x10 | ~x42"
-  and "~x10 | ~x36"
-  and "~x42 | ~x36"
-  and "~x11 | ~x43"
-  and "~x11 | ~x10"
-  and "~x11 | ~x37"
-  and "~x43 | ~x10"
-  and "~x43 | ~x37"
-  and "~x10 | ~x37"
-  and "~x12 | ~x44"
-  and "~x12 | ~x11"
-  and "~x12 | ~x38"
-  and "~x44 | ~x11"
-  and "~x44 | ~x38"
-  and "~x11 | ~x38"
-  and "~x13 | ~x45"
-  and "~x13 | ~x12"
-  and "~x13 | ~x39"
-  and "~x45 | ~x12"
-  and "~x45 | ~x39"
-  and "~x12 | ~x39"
-  and "~x14 | ~x46"
-  and "~x14 | ~x13"
-  and "~x14 | ~x40"
-  and "~x46 | ~x13"
-  and "~x46 | ~x40"
-  and "~x13 | ~x40"
-  and "~x47 | ~x14"
-  and "~x47 | ~x41"
-  and "~x14 | ~x41"
-  and "~x15 | ~x48"
-  and "~x15 | ~x42"
-  and "~x48 | ~x42"
-  and "~x16 | ~x49"
-  and "~x16 | ~x15"
-  and "~x16 | ~x43"
-  and "~x49 | ~x15"
-  and "~x49 | ~x43"
-  and "~x15 | ~x43"
-  and "~x17 | ~x50"
-  and "~x17 | ~x16"
-  and "~x17 | ~x44"
-  and "~x50 | ~x16"
-  and "~x50 | ~x44"
-  and "~x16 | ~x44"
-  and "~x18 | ~x51"
-  and "~x18 | ~x17"
-  and "~x18 | ~x45"
-  and "~x51 | ~x17"
-  and "~x51 | ~x45"
-  and "~x17 | ~x45"
-  and "~x19 | ~x52"
-  and "~x19 | ~x18"
-  and "~x19 | ~x46"
-  and "~x52 | ~x18"
-  and "~x52 | ~x46"
-  and "~x18 | ~x46"
-  and "~x53 | ~x19"
-  and "~x53 | ~x47"
-  and "~x19 | ~x47"
-  and "~x20 | ~x54"
-  and "~x20 | ~x48"
-  and "~x54 | ~x48"
-  and "~x21 | ~x55"
-  and "~x21 | ~x20"
-  and "~x21 | ~x49"
-  and "~x55 | ~x20"
-  and "~x55 | ~x49"
-  and "~x20 | ~x49"
-  and "~x22 | ~x56"
-  and "~x22 | ~x21"
-  and "~x22 | ~x50"
-  and "~x56 | ~x21"
-  and "~x56 | ~x50"
-  and "~x21 | ~x50"
-  and "~x23 | ~x57"
-  and "~x23 | ~x22"
-  and "~x23 | ~x51"
-  and "~x57 | ~x22"
-  and "~x57 | ~x51"
-  and "~x22 | ~x51"
-  and "~x24 | ~x58"
-  and "~x24 | ~x23"
-  and "~x24 | ~x52"
-  and "~x58 | ~x23"
-  and "~x58 | ~x52"
-  and "~x23 | ~x52"
-  and "~x59 | ~x24"
-  and "~x59 | ~x53"
-  and "~x24 | ~x53"
-  and "~x25 | ~x54"
-  and "~x26 | ~x25"
-  and "~x26 | ~x55"
-  and "~x25 | ~x55"
-  and "~x27 | ~x26"
-  and "~x27 | ~x56"
-  and "~x26 | ~x56"
-  and "~x28 | ~x27"
-  and "~x28 | ~x57"
-  and "~x27 | ~x57"
-  and "~x29 | ~x28"
-  and "~x29 | ~x58"
-  and "~x28 | ~x58"
+  and "x1 \<or> x31 \<or> x0"
+  and "x2 \<or> x32 \<or> x1"
+  and "x3 \<or> x33 \<or> x2"
+  and "x4 \<or> x34 \<or> x3"
+  and "x35 \<or> x4"
+  and "x5 \<or> x36 \<or> x30"
+  and "x6 \<or> x37 \<or> x5 \<or> x31"
+  and "x7 \<or> x38 \<or> x6 \<or> x32"
+  and "x8 \<or> x39 \<or> x7 \<or> x33"
+  and "x9 \<or> x40 \<or> x8 \<or> x34"
+  and "x41 \<or> x9 \<or> x35"
+  and "x10 \<or> x42 \<or> x36"
+  and "x11 \<or> x43 \<or> x10 \<or> x37"
+  and "x12 \<or> x44 \<or> x11 \<or> x38"
+  and "x13 \<or> x45 \<or> x12 \<or> x39"
+  and "x14 \<or> x46 \<or> x13 \<or> x40"
+  and "x47 \<or> x14 \<or> x41"
+  and "x15 \<or> x48 \<or> x42"
+  and "x16 \<or> x49 \<or> x15 \<or> x43"
+  and "x17 \<or> x50 \<or> x16 \<or> x44"
+  and "x18 \<or> x51 \<or> x17 \<or> x45"
+  and "x19 \<or> x52 \<or> x18 \<or> x46"
+  and "x53 \<or> x19 \<or> x47"
+  and "x20 \<or> x54 \<or> x48"
+  and "x21 \<or> x55 \<or> x20 \<or> x49"
+  and "x22 \<or> x56 \<or> x21 \<or> x50"
+  and "x23 \<or> x57 \<or> x22 \<or> x51"
+  and "x24 \<or> x58 \<or> x23 \<or> x52"
+  and "x59 \<or> x24 \<or> x53"
+  and "x25 \<or> x54"
+  and "x26 \<or> x25 \<or> x55"
+  and "x27 \<or> x26 \<or> x56"
+  and "x28 \<or> x27 \<or> x57"
+  and "x29 \<or> x28 \<or> x58"
+  and "~x1 \<or> ~x31"
+  and "~x1 \<or> ~x0"
+  and "~x31 \<or> ~x0"
+  and "~x2 \<or> ~x32"
+  and "~x2 \<or> ~x1"
+  and "~x32 \<or> ~x1"
+  and "~x3 \<or> ~x33"
+  and "~x3 \<or> ~x2"
+  and "~x33 \<or> ~x2"
+  and "~x4 \<or> ~x34"
+  and "~x4 \<or> ~x3"
+  and "~x34 \<or> ~x3"
+  and "~x35 \<or> ~x4"
+  and "~x5 \<or> ~x36"
+  and "~x5 \<or> ~x30"
+  and "~x36 \<or> ~x30"
+  and "~x6 \<or> ~x37"
+  and "~x6 \<or> ~x5"
+  and "~x6 \<or> ~x31"
+  and "~x37 \<or> ~x5"
+  and "~x37 \<or> ~x31"
+  and "~x5 \<or> ~x31"
+  and "~x7 \<or> ~x38"
+  and "~x7 \<or> ~x6"
+  and "~x7 \<or> ~x32"
+  and "~x38 \<or> ~x6"
+  and "~x38 \<or> ~x32"
+  and "~x6 \<or> ~x32"
+  and "~x8 \<or> ~x39"
+  and "~x8 \<or> ~x7"
+  and "~x8 \<or> ~x33"
+  and "~x39 \<or> ~x7"
+  and "~x39 \<or> ~x33"
+  and "~x7 \<or> ~x33"
+  and "~x9 \<or> ~x40"
+  and "~x9 \<or> ~x8"
+  and "~x9 \<or> ~x34"
+  and "~x40 \<or> ~x8"
+  and "~x40 \<or> ~x34"
+  and "~x8 \<or> ~x34"
+  and "~x41 \<or> ~x9"
+  and "~x41 \<or> ~x35"
+  and "~x9 \<or> ~x35"
+  and "~x10 \<or> ~x42"
+  and "~x10 \<or> ~x36"
+  and "~x42 \<or> ~x36"
+  and "~x11 \<or> ~x43"
+  and "~x11 \<or> ~x10"
+  and "~x11 \<or> ~x37"
+  and "~x43 \<or> ~x10"
+  and "~x43 \<or> ~x37"
+  and "~x10 \<or> ~x37"
+  and "~x12 \<or> ~x44"
+  and "~x12 \<or> ~x11"
+  and "~x12 \<or> ~x38"
+  and "~x44 \<or> ~x11"
+  and "~x44 \<or> ~x38"
+  and "~x11 \<or> ~x38"
+  and "~x13 \<or> ~x45"
+  and "~x13 \<or> ~x12"
+  and "~x13 \<or> ~x39"
+  and "~x45 \<or> ~x12"
+  and "~x45 \<or> ~x39"
+  and "~x12 \<or> ~x39"
+  and "~x14 \<or> ~x46"
+  and "~x14 \<or> ~x13"
+  and "~x14 \<or> ~x40"
+  and "~x46 \<or> ~x13"
+  and "~x46 \<or> ~x40"
+  and "~x13 \<or> ~x40"
+  and "~x47 \<or> ~x14"
+  and "~x47 \<or> ~x41"
+  and "~x14 \<or> ~x41"
+  and "~x15 \<or> ~x48"
+  and "~x15 \<or> ~x42"
+  and "~x48 \<or> ~x42"
+  and "~x16 \<or> ~x49"
+  and "~x16 \<or> ~x15"
+  and "~x16 \<or> ~x43"
+  and "~x49 \<or> ~x15"
+  and "~x49 \<or> ~x43"
+  and "~x15 \<or> ~x43"
+  and "~x17 \<or> ~x50"
+  and "~x17 \<or> ~x16"
+  and "~x17 \<or> ~x44"
+  and "~x50 \<or> ~x16"
+  and "~x50 \<or> ~x44"
+  and "~x16 \<or> ~x44"
+  and "~x18 \<or> ~x51"
+  and "~x18 \<or> ~x17"
+  and "~x18 \<or> ~x45"
+  and "~x51 \<or> ~x17"
+  and "~x51 \<or> ~x45"
+  and "~x17 \<or> ~x45"
+  and "~x19 \<or> ~x52"
+  and "~x19 \<or> ~x18"
+  and "~x19 \<or> ~x46"
+  and "~x52 \<or> ~x18"
+  and "~x52 \<or> ~x46"
+  and "~x18 \<or> ~x46"
+  and "~x53 \<or> ~x19"
+  and "~x53 \<or> ~x47"
+  and "~x19 \<or> ~x47"
+  and "~x20 \<or> ~x54"
+  and "~x20 \<or> ~x48"
+  and "~x54 \<or> ~x48"
+  and "~x21 \<or> ~x55"
+  and "~x21 \<or> ~x20"
+  and "~x21 \<or> ~x49"
+  and "~x55 \<or> ~x20"
+  and "~x55 \<or> ~x49"
+  and "~x20 \<or> ~x49"
+  and "~x22 \<or> ~x56"
+  and "~x22 \<or> ~x21"
+  and "~x22 \<or> ~x50"
+  and "~x56 \<or> ~x21"
+  and "~x56 \<or> ~x50"
+  and "~x21 \<or> ~x50"
+  and "~x23 \<or> ~x57"
+  and "~x23 \<or> ~x22"
+  and "~x23 \<or> ~x51"
+  and "~x57 \<or> ~x22"
+  and "~x57 \<or> ~x51"
+  and "~x22 \<or> ~x51"
+  and "~x24 \<or> ~x58"
+  and "~x24 \<or> ~x23"
+  and "~x24 \<or> ~x52"
+  and "~x58 \<or> ~x23"
+  and "~x58 \<or> ~x52"
+  and "~x23 \<or> ~x52"
+  and "~x59 \<or> ~x24"
+  and "~x59 \<or> ~x53"
+  and "~x24 \<or> ~x53"
+  and "~x25 \<or> ~x54"
+  and "~x26 \<or> ~x25"
+  and "~x26 \<or> ~x55"
+  and "~x25 \<or> ~x55"
+  and "~x27 \<or> ~x26"
+  and "~x27 \<or> ~x56"
+  and "~x26 \<or> ~x56"
+  and "~x28 \<or> ~x27"
+  and "~x28 \<or> ~x57"
+  and "~x27 \<or> ~x57"
+  and "~x29 \<or> ~x28"
+  and "~x29 \<or> ~x58"
+  and "~x28 \<or> ~x58"
   shows False
-  using assms by smt
+  using assms by smt (* smt2 FIXME: THM 0 *)
 
 lemma "\<forall>x::int. P x \<longrightarrow> (\<forall>y::int. P x \<or> P y)"
-  by smt
+  by smt2
 
 lemma
   assumes "(\<forall>x y. P x y = x)"
   shows "(\<exists>y. P x y) = P x c"
-  using assms by smt
+  using assms by smt (* smt2 FIXME: Option *)
 
 lemma
   assumes "(\<forall>x y. P x y = x)"
   and "(\<forall>x. \<exists>y. P x y) = (\<forall>x. P x c)"
   shows "(EX y. P x y) = P x c"
-  using assms by smt
+  using assms by smt (* smt2 FIXME: Option *)
 
 lemma
   assumes "if P x then \<not>(\<exists>y. P y) else (\<forall>y. \<not>P y)"
   shows "P x \<longrightarrow> P y"
-  using assms by smt
+  using assms by smt2
 
 
 section {* Arithmetic *}
 
 subsection {* Linear arithmetic over integers and reals *}
 
-lemma "(3::int) = 3" by smt
-
-lemma "(3::real) = 3" by smt
-
-lemma "(3 :: int) + 1 = 4" by smt
-
-lemma "x + (y + z) = y + (z + (x::int))" by smt
-
-lemma "max (3::int) 8 > 5" by smt
-
-lemma "abs (x :: real) + abs y \<ge> abs (x + y)" by smt
-
-lemma "P ((2::int) < 3) = P True" by smt
-
-lemma "x + 3 \<ge> 4 \<or> x < (1::int)" by smt
+lemma "(3::int) = 3" by smt2
+lemma "(3::real) = 3" by smt2
+lemma "(3 :: int) + 1 = 4" by smt2
+lemma "x + (y + z) = y + (z + (x::int))" by smt2
+lemma "max (3::int) 8 > 5" by smt2
+lemma "abs (x :: real) + abs y \<ge> abs (x + y)" by smt2
+lemma "P ((2::int) < 3) = P True" by smt2
+lemma "x + 3 \<ge> 4 \<or> x < (1::int)" by smt2
 
 lemma
   assumes "x \<ge> (3::int)" and "y = x + 4"
   shows "y - x > 0"
-  using assms by smt
+  using assms by smt2
 
-lemma "let x = (2 :: int) in x + x \<noteq> 5" by smt
+lemma "let x = (2 :: int) in x + x \<noteq> 5" by smt2
 
 lemma
   fixes x :: real
   assumes "3 * x + 7 * a < 4" and "3 < 2 * x"
   shows "a < 0"
-  using assms by smt
+  using assms by smt2
 
-lemma "(0 \<le> y + -1 * x \<or> \<not> 0 \<le> x \<or> 0 \<le> (x::int)) = (\<not> False)" by smt
+lemma "(0 \<le> y + -1 * x \<or> \<not> 0 \<le> x \<or> 0 \<le> (x::int)) = (\<not> False)" by smt2
 
 lemma "
-  (n < m & m < n') | (n < m & m = n') | (n < n' & n' < m) |
-  (n = n' & n' < m) | (n = m & m < n') |
-  (n' < m & m < n) | (n' < m & m = n) |
-  (n' < n & n < m) | (n' = n & n < m) | (n' = m & m < n) |
-  (m < n & n < n') | (m < n & n' = n) | (m < n' & n' < n) |
-  (m = n & n < n') | (m = n' & n' < n) |
-  (n' = m & m = (n::int))"
-  by smt
+  (n < m \<and> m < n') \<or> (n < m \<and> m = n') \<or> (n < n' \<and> n' < m) \<or>
+  (n = n' \<and> n' < m) \<or> (n = m \<and> m < n') \<or>
+  (n' < m \<and> m < n) \<or> (n' < m \<and> m = n) \<or>
+  (n' < n \<and> n < m) \<or> (n' = n \<and> n < m) \<or> (n' = m \<and> m < n) \<or>
+  (m < n \<and> n < n') \<or> (m < n \<and> n' = n) \<or> (m < n' \<and> n' < n) \<or>
+  (m = n \<and> n < n') \<or> (m = n' \<and> n' < n) \<or>
+  (n' = m \<and> m = (n::int))"
+  by smt2
 
 text{*
 The following example was taken from HOL/ex/PresburgerEx.thy, where it says:
@@ -320,220 +309,197 @@
 lemma "\<lbrakk> x3 = abs x2 - x1; x4 = abs x3 - x2; x5 = abs x4 - x3;
          x6 = abs x5 - x4; x7 = abs x6 - x5; x8 = abs x7 - x6;
          x9 = abs x8 - x7; x10 = abs x9 - x8; x11 = abs x10 - x9 \<rbrakk>
- \<Longrightarrow> x1 = x10 & x2 = (x11::int)"
-  by smt
+ \<Longrightarrow> x1 = x10 \<and> x2 = (x11::int)"
+  by smt2
 
 
-lemma "let P = 2 * x + 1 > x + (x::real) in P \<or> False \<or> P" by smt
+lemma "let P = 2 * x + 1 > x + (x::real) in P \<or> False \<or> P" by smt2
 
 lemma "x + (let y = x mod 2 in 2 * y + 1) \<ge> x + (1::int)"
-  using [[z3_with_extensions]]
-  by smt
+  using [[z3_new_extensions]] by smt2
 
 lemma "x + (let y = x mod 2 in y + y) < x + (3::int)"
-  using [[z3_with_extensions]]
-  by smt
+  using [[z3_new_extensions]] by smt2
 
 lemma
   assumes "x \<noteq> (0::real)"
-  shows "x + x \<noteq> (let P = (abs x > 1) in if P \<or> \<not>P then 4 else 2) * x"
-  using assms by smt
+  shows "x + x \<noteq> (let P = (abs x > 1) in if P \<or> \<not> P then 4 else 2) * x"
+  using assms [[z3_new_extensions]] by smt2
 
 lemma
   assumes "(n + m) mod 2 = 0" and "n mod 4 = 3"
-  shows "n mod 2 = 1 & m mod 2 = (1::int)"
-  using assms [[z3_with_extensions]] by smt
-
+  shows "n mod 2 = 1 \<and> m mod 2 = (1::int)"
+  using assms [[z3_new_extensions]] by smt2
 
 
 subsection {* Linear arithmetic with quantifiers *}
 
-lemma "~ (\<exists>x::int. False)" by smt
-
-lemma "~ (\<exists>x::real. False)" by smt
+lemma "~ (\<exists>x::int. False)" by smt2
+lemma "~ (\<exists>x::real. False)" by smt2
 
 lemma "\<exists>x::int. 0 < x"
   using [[smt_oracle=true]] (* no Z3 proof *)
-  by smt
+  by smt (* smt2 FIXME: requires Z3 4.3.1 *)
 
 lemma "\<exists>x::real. 0 < x"
   using [[smt_oracle=true]] (* no Z3 proof *)
-  by smt
+  by smt (* smt2 FIXME: requires Z3 4.3.1 *)
 
 lemma "\<forall>x::int. \<exists>y. y > x"
   using [[smt_oracle=true]] (* no Z3 proof *)
-  by smt
-
-lemma "\<forall>x y::int. (x = 0 \<and> y = 1) \<longrightarrow> x \<noteq> y" by smt
-
-lemma "\<exists>x::int. \<forall>y. x < y \<longrightarrow> y < 0 \<or> y >= 0" by smt
-
-lemma "\<forall>x y::int. x < y \<longrightarrow> (2 * x + 1) < (2 * y)" by smt
-
-lemma "\<forall>x y::int. (2 * x + 1) \<noteq> (2 * y)" by smt
-
-lemma "\<forall>x y::int. x + y > 2 \<or> x + y = 2 \<or> x + y < 2" by smt
-
-lemma "\<forall>x::int. if x > 0 then x + 1 > 0 else 1 > x" by smt
+  by smt (* smt2 FIXME: requires Z3 4.3.1 *)
 
-lemma "if (ALL x::int. x < 0 \<or> x > 0) then False else True" by smt
-
-lemma "(if (ALL x::int. x < 0 \<or> x > 0) then -1 else 3) > (0::int)" by smt
-
-lemma "~ (\<exists>x y z::int. 4 * x + -6 * y = (1::int))" by smt
-
-lemma "\<exists>x::int. \<forall>x y. 0 < x \<and> 0 < y \<longrightarrow> (0::int) < x + y" by smt
-
-lemma "\<exists>u::int. \<forall>(x::int) y::real. 0 < x \<and> 0 < y \<longrightarrow> -1 < x" by smt
-
-lemma "\<exists>x::int. (\<forall>y. y \<ge> x \<longrightarrow> y > 0) \<longrightarrow> x > 0" by smt
-
-lemma "\<forall>x::int. SMT.trigger [[SMT.pat x]] (x < a \<longrightarrow> 2 * x < 2 * a)" by smt
-
-lemma "\<forall>(a::int) b::int. 0 < b \<or> b < 1" by smt
+lemma "\<forall>x y::int. (x = 0 \<and> y = 1) \<longrightarrow> x \<noteq> y" by smt2
+lemma "\<exists>x::int. \<forall>y. x < y \<longrightarrow> y < 0 \<or> y >= 0" by smt2
+lemma "\<forall>x y::int. x < y \<longrightarrow> (2 * x + 1) < (2 * y)" by smt2
+lemma "\<forall>x y::int. (2 * x + 1) \<noteq> (2 * y)" by smt2
+lemma "\<forall>x y::int. x + y > 2 \<or> x + y = 2 \<or> x + y < 2" by smt2
+lemma "\<forall>x::int. if x > 0 then x + 1 > 0 else 1 > x" by smt2
+lemma "if (ALL x::int. x < 0 \<or> x > 0) then False else True"  by smt (* smt2 FIXME: requires Z3 4.3.1 *)
+lemma "(if (ALL x::int. x < 0 \<or> x > 0) then -1 else 3) > (0::int)" by smt (* smt2 FIXME: requires Z3 4.3.1 *)
+lemma "~ (\<exists>x y z::int. 4 * x + -6 * y = (1::int))" by smt2
+lemma "\<exists>x::int. \<forall>x y. 0 < x \<and> 0 < y \<longrightarrow> (0::int) < x + y" by smt2
+lemma "\<exists>u::int. \<forall>(x::int) y::real. 0 < x \<and> 0 < y \<longrightarrow> -1 < x" by smt2
+lemma "\<exists>x::int. (\<forall>y. y \<ge> x \<longrightarrow> y > 0) \<longrightarrow> x > 0" by smt (* smt2 FIXME: requires Z3 4.3.1 *)
+lemma "\<forall>x::int. SMT2.trigger [[SMT2.pat x]] (x < a \<longrightarrow> 2 * x < 2 * a)" by smt2
+lemma "\<forall>(a::int) b::int. 0 < b \<or> b < 1" by smt2
 
 
 subsection {* Non-linear arithmetic over integers and reals *}
 
 lemma "a > (0::int) \<Longrightarrow> a*b > 0 \<Longrightarrow> b > 0"
-  using [[smt_oracle, z3_with_extensions]]
-  by smt
+  using [[smt2_oracle, z3_new_extensions]]
+  by smt2
 
 lemma  "(a::int) * (x + 1 + y) = a * x + a * (y + 1)"
-  using [[z3_with_extensions]]
-  by smt
+  using [[z3_new_extensions]]
+  by smt2
 
 lemma "((x::real) * (1 + y) - x * (1 - y)) = (2 * x * y)"
-  using [[z3_with_extensions]]
-  by smt
+  using [[z3_new_extensions]]
+  by smt2
 
 lemma
   "(U::int) + (1 + p) * (b + e) + p * d =
    U + (2 * (1 + p) * (b + e) + (1 + p) * d + d * p) - (1 + p) * (b + d + e)"
-  using [[z3_with_extensions]]
-  by smt
+  using [[z3_new_extensions]] by smt2
 
-lemma [z3_rule]:
+lemma [z3_rule, z3_new_rule]:
   fixes x :: "int"
   assumes "x * y \<le> 0" and "\<not> y \<le> 0" and "\<not> x \<le> 0"
   shows False
   using assms by (metis mult_le_0_iff)
 
 lemma "x * y \<le> (0 :: int) \<Longrightarrow> x \<le> 0 \<or> y \<le> 0"
-  using [[z3_with_extensions]]
-  by smt
-
+  using [[z3_with_extensions]] [[z3_new_extensions]]
+  by smt (* smt2 FIXME: "th-lemma" tactic fails *)
 
 
 subsection {* Linear arithmetic for natural numbers *}
 
-lemma "2 * (x::nat) ~= 1" by smt
+lemma "2 * (x::nat) ~= 1" by smt2
 
-lemma "a < 3 \<Longrightarrow> (7::nat) > 2 * a" by smt
+lemma "a < 3 \<Longrightarrow> (7::nat) > 2 * a" by smt2
 
-lemma "let x = (1::nat) + y in x - y > 0 * x" by smt
+lemma "let x = (1::nat) + y in x - y > 0 * x" by smt2
 
 lemma
   "let x = (1::nat) + y in
    let P = (if x > 0 then True else False) in
    False \<or> P = (x - 1 = y) \<or> (\<not>P \<longrightarrow> False)"
-  by smt
+  by smt2
 
-lemma "int (nat \<bar>x::int\<bar>) = \<bar>x\<bar>" by smt
+lemma "int (nat \<bar>x::int\<bar>) = \<bar>x\<bar>" by smt2
 
 definition prime_nat :: "nat \<Rightarrow> bool" where
   "prime_nat p = (1 < p \<and> (\<forall>m. m dvd p --> m = 1 \<or> m = p))"
-lemma "prime_nat (4*m + 1) \<Longrightarrow> m \<ge> (1::nat)" by (smt prime_nat_def)
+lemma "prime_nat (4*m + 1) \<Longrightarrow> m \<ge> (1::nat)" by (smt2 prime_nat_def)
 
 
 section {* Pairs *}
 
 lemma "fst (x, y) = a \<Longrightarrow> x = a"
-  using fst_conv
-  by smt
+  using fst_conv by smt2
 
 lemma "p1 = (x, y) \<and> p2 = (y, x) \<Longrightarrow> fst p1 = snd p2"
-  using fst_conv snd_conv
-  by smt
+  using fst_conv snd_conv by smt2
 
 
 section {* Higher-order problems and recursion *}
 
 lemma "i \<noteq> i1 \<and> i \<noteq> i2 \<Longrightarrow> (f (i1 := v1, i2 := v2)) i = f i"
-  using fun_upd_same fun_upd_apply
-  by smt
+  using fun_upd_same fun_upd_apply by smt2
 
 lemma "(f g (x::'a::type) = (g x \<and> True)) \<or> (f g x = True) \<or> (g x = True)"
-  by smt
+  by smt2
 
-lemma "id x = x \<and> id True = True" by (smt id_def)
+lemma "id x = x \<and> id True = True"
+  by (smt id_def) (* smt2 FIXME: Option *)
 
 lemma "i \<noteq> i1 \<and> i \<noteq> i2 \<Longrightarrow> ((f (i1 := v1)) (i2 := v2)) i = f i"
-  using fun_upd_same fun_upd_apply
-  by smt
+  using fun_upd_same fun_upd_apply by smt2
 
 lemma
   "f (\<exists>x. g x) \<Longrightarrow> True"
   "f (\<forall>x. g x) \<Longrightarrow> True"
-  by smt+
-
-lemma True using let_rsp by smt
+  by smt2+
 
-lemma "le = op \<le> \<Longrightarrow> le (3::int) 42" by smt
-
-lemma "map (\<lambda>i::nat. i + 1) [0, 1] = [1, 2]" by (smt list.map)
-
-
-lemma "(ALL x. P x) | ~ All P" by smt
+lemma True using let_rsp by smt2
+lemma "le = op \<le> \<Longrightarrow> le (3::int) 42" by smt2
+lemma "map (\<lambda>i::nat. i + 1) [0, 1] = [1, 2]" by (smt2 list.map)
+lemma "(ALL x. P x) \<or> ~ All P" by smt2
 
 fun dec_10 :: "nat \<Rightarrow> nat" where
   "dec_10 n = (if n < 10 then n else dec_10 (n - 10))"
-lemma "dec_10 (4 * dec_10 4) = 6" by (smt dec_10.simps)
 
+lemma "dec_10 (4 * dec_10 4) = 6" by (smt2 dec_10.simps)
 
 axiomatization
   eval_dioph :: "int list \<Rightarrow> nat list \<Rightarrow> int"
-  where
+where
   eval_dioph_mod:
   "eval_dioph ks xs mod int n = eval_dioph ks (map (\<lambda>x. x mod n) xs) mod int n"
-  and
+and
   eval_dioph_div_mult:
   "eval_dioph ks (map (\<lambda>x. x div n) xs) * int n +
    eval_dioph ks (map (\<lambda>x. x mod n) xs) = eval_dioph ks xs"
+
 lemma
   "(eval_dioph ks xs = l) =
    (eval_dioph ks (map (\<lambda>x. x mod 2) xs) mod 2 = l mod 2 \<and>
     eval_dioph ks (map (\<lambda>x. x div 2) xs) =
       (l - eval_dioph ks (map (\<lambda>x. x mod 2) xs)) div 2)"
-  using [[smt_oracle=true]] (*FIXME*)
-  using [[z3_with_extensions]]
-  by (smt eval_dioph_mod[where n=2] eval_dioph_div_mult[where n=2])
+  using [[smt2_oracle=true]] (*FIXME*)
+  using [[z3_new_extensions]]
+  by (smt2 eval_dioph_mod[where n=2] eval_dioph_div_mult[where n=2])
 
 
 context complete_lattice
 begin
 
 lemma
-  assumes "Sup { a | i::bool . True } \<le> Sup { b | i::bool . True }"
-  and     "Sup { b | i::bool . True } \<le> Sup { a | i::bool . True }"
-  shows   "Sup { a | i::bool . True } \<le> Sup { a | i::bool . True }"
-  using assms by (smt order_trans)
+  assumes "Sup {a | i::bool. True} \<le> Sup {b | i::bool. True}"
+  and "Sup {b | i::bool. True} \<le> Sup {a | i::bool. True}"
+  shows "Sup {a | i::bool. True} \<le> Sup {a | i::bool. True}"
+  using assms by (smt2 order_trans)
 
 end
 
 
-
 section {* Monomorphization examples *}
 
 definition Pred :: "'a \<Rightarrow> bool" where "Pred x = True"
-lemma poly_Pred: "Pred x \<and> (Pred [x] \<or> \<not>Pred[x])" by (simp add: Pred_def)
-lemma "Pred (1::int)" by (smt poly_Pred)
+
+lemma poly_Pred: "Pred x \<and> (Pred [x] \<or> \<not> Pred [x])" by (simp add: Pred_def)
+
+lemma "Pred (1::int)" by (smt2 poly_Pred)
 
 axiomatization g :: "'a \<Rightarrow> nat"
 axiomatization where
   g1: "g (Some x) = g [x]" and
   g2: "g None = g []" and
   g3: "g xs = length xs"
-lemma "g (Some (3::int)) = g (Some True)" by (smt g1 g2 g3 list.size)
+
+lemma "g (Some (3::int)) = g (Some True)" by (smt g1 g2 g3 list.size) (* smt2 FIXME: Option *)
 
 end
--- a/src/HOL/SMT_Examples/SMT_Tests.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/SMT_Examples/SMT_Tests.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -9,11 +9,11 @@
 begin
 
 smt_status
+smt2_status
 
 text {* Most examples are taken from various Isabelle theories and from HOL4. *}
 
 
-
 section {* Propositional logic *}
 
 lemma
@@ -24,7 +24,7 @@
   "True \<or> False"
   "False \<longrightarrow> True"
   "\<not>(False \<longleftrightarrow> True)"
-  by smt+
+  by smt2+
 
 lemma
   "P \<or> \<not>P"
@@ -63,7 +63,7 @@
   "\<not>(P \<longleftrightarrow> \<not>P)"
   "(P \<longrightarrow> Q) \<longleftrightarrow> (\<not>Q \<longrightarrow> \<not>P)"
   "P \<longleftrightarrow> P \<longleftrightarrow> P \<longleftrightarrow> P \<longleftrightarrow> P \<longleftrightarrow> P \<longleftrightarrow> P \<longleftrightarrow> P \<longleftrightarrow> P \<longleftrightarrow> P"
-  by smt+
+  by smt2+
 
 lemma
   "(if P then Q1 else Q2) \<longleftrightarrow> ((P \<longrightarrow> Q1) \<and> (\<not>P \<longrightarrow> Q2))"
@@ -72,15 +72,14 @@
   "(if P1 \<and> P2 then Q1 else Q2) \<longleftrightarrow> (if P1 then if P2 then Q1 else Q2 else Q2)"
   "(P1 \<longrightarrow> (if P2 then Q1 else Q2)) \<longleftrightarrow>
    (if P1 \<longrightarrow> P2 then P1 \<longrightarrow> Q1 else P1 \<longrightarrow> Q2)"
-  by smt+
+  by smt2+
 
 lemma
   "case P of True \<Rightarrow> P | False \<Rightarrow> \<not>P"
   "case P of False \<Rightarrow> \<not>P | True \<Rightarrow> P"
   "case \<not>P of True \<Rightarrow> \<not>P | False \<Rightarrow> P"
   "case P of True \<Rightarrow> (Q \<longrightarrow> P) | False \<Rightarrow> (P \<longrightarrow> Q)"
-  by smt+
-
+  by smt2+
 
 
 section {* First-order logic with equality *}
@@ -93,7 +92,7 @@
   "x = y \<longrightarrow> g x y = g y x"
   "f (f x) = x \<and> f (f (f (f (f x)))) = x \<longrightarrow> f x = x"
   "((if a then b else c) = d) = ((a \<longrightarrow> (b = d)) \<and> (\<not> a \<longrightarrow> (c = d)))"
-  by smt+
+  by smt2+
 
 lemma
   "\<forall>x. x = x"
@@ -106,12 +105,11 @@
   "(\<forall>x. P x \<longrightarrow> P (f x)) \<and> P d \<longrightarrow> P (f(f(f(d))))"
   "(\<forall>x y. s x y = s y x) \<longrightarrow> a = a \<and> s a b = s b a"
   "(\<forall>s. q s \<longrightarrow> r s) \<and> \<not>r s \<and> (\<forall>s. \<not>r s \<and> \<not>q s \<longrightarrow> p t \<or> q t) \<longrightarrow> p t \<or> r t"
-  by smt+
+  by smt2+
 
 lemma
   "(\<forall>x. P x) \<and> R \<longleftrightarrow> (\<forall>x. P x \<and> R)"
-  using [[smt_oracle]] by smt
-  (* BUG: Z3 proof parser (line 34): unknown function symbol: "S2!val!0" *)
+  by smt2
 
 lemma
   "\<exists>x. x = x"
@@ -120,7 +118,7 @@
   "(\<exists>x. P x) \<and> R \<longleftrightarrow> (\<exists>x. P x \<and> R)"
   "(\<exists>x y z. S x z) \<longleftrightarrow> (\<exists>x z. S x z)"
   "\<not>((\<exists>x. \<not>P x) \<and> ((\<exists>x. P x) \<or> (\<exists>x. P x \<and> Q x)) \<and> \<not>(\<exists>x. P x))"
-  by smt+
+  by smt2+
 
 lemma
   "\<exists>x y. x = y"
@@ -129,8 +127,7 @@
   "\<exists>x. P x \<longrightarrow> P a \<and> P b"
   "\<exists>x. (\<exists>y. P y) \<longrightarrow> P x"
   "(\<exists>x. Q \<longrightarrow> P x) \<longleftrightarrow> (Q \<longrightarrow> (\<exists>x. P x))"
-  using [[smt_oracle]] by smt+
-  (* BUG: Z3 proof parser (line 34): unknown function symbol: "S2!val!0" *)
+  oops (* smt2 FIXME: requires Z3 4.3.1 *)
 
 lemma
   "(\<not>(\<exists>x. P x)) \<longleftrightarrow> (\<forall>x. \<not> P x)"
@@ -138,7 +135,7 @@
   "(\<forall>x y. R x y = x) \<longrightarrow> (\<exists>y. R x y) = R x c"
   "(if P x then \<not>(\<exists>y. P y) else (\<forall>y. \<not>P y)) \<longrightarrow> P x \<longrightarrow> P y"
   "(\<forall>x y. R x y = x) \<and> (\<forall>x. \<exists>y. R x y) = (\<forall>x. R x c) \<longrightarrow> (\<exists>y. R x y) = R x c"
-  by smt+
+  by smt+ (* smt2 FIXME: Option *)
 
 lemma
   "\<forall>x. \<exists>y. f x y = f x (g x)"
@@ -149,20 +146,20 @@
   "(\<exists>x. \<forall>y. P x \<longleftrightarrow> P y) \<longrightarrow> ((\<exists>x. P x) \<longleftrightarrow> (\<forall>y. P y))"
   "\<exists>z. P z \<longrightarrow> (\<forall>x. P x)"
   "(\<exists>y. \<forall>x. R x y) \<longrightarrow> (\<forall>x. \<exists>y. R x y)"
-  by smt+
+  by smt+ (* smt2 FIXME: requires Z3 4.3.1 *)
 
 lemma
-  "(\<exists>! x. P x) \<longrightarrow> (\<exists>x. P x)"
+  "(\<exists>!x. P x) \<longrightarrow> (\<exists>x. P x)"
   "(\<exists>!x. P x) \<longleftrightarrow> (\<exists>x. P x \<and> (\<forall>y. y \<noteq> x \<longrightarrow> \<not>P y))"
   "P a \<longrightarrow> (\<forall>x. P x \<longrightarrow> x = a) \<longrightarrow> (\<exists>!x. P x)"
   "(\<exists>x. P x) \<and> (\<forall>x y. P x \<and> P y \<longrightarrow> x = y) \<longrightarrow> (\<exists>!x. P x)"
   "(\<exists>!x. P x) \<and> (\<forall>x. P x \<and> (\<forall>y. P y \<longrightarrow> y = x) \<longrightarrow> R) \<longrightarrow> R"
-  by smt+
+  by smt2+
 
 lemma
   "(\<forall>x\<in>M. P x) \<and> c \<in> M \<longrightarrow> P c"
   "(\<exists>x\<in>M. P x) \<or> \<not>(P c \<and> c \<in> M)"
-  by smt+
+  by smt2+
 
 lemma
   "let P = True in P"
@@ -173,65 +170,64 @@
   "(let x = y1; z = y2 in R x z) \<longleftrightarrow> (let z = y2; x = y1 in R x z)"
   "(let x = y1; z = y2 in R x z) \<longleftrightarrow> (let z = y1; x = y2 in R z x)"
   "let P = (\<forall>x. Q x) in if P then P else \<not>P"
-  by smt+
+  by smt2+
 
 lemma
   "a \<noteq> b \<and> a \<noteq> c \<and> b \<noteq> c \<and> (\<forall>x y. f x = f y \<longrightarrow> y = x) \<longrightarrow> f a \<noteq> f b"
-  by smt
+  by smt2
 
 lemma
   "(\<forall>x y z. f x y = f x z \<longrightarrow> y = z) \<and> b \<noteq> c \<longrightarrow> f a b \<noteq> f a c"
   "(\<forall>x y z. f x y = f z y \<longrightarrow> x = z) \<and> a \<noteq> d \<longrightarrow> f a b \<noteq> f d b"
-  by smt+
+  by smt2+
 
 
 section {* Guidance for quantifier heuristics: patterns and weights *}
 
 lemma
-  assumes "\<forall>x. SMT.trigger [[SMT.pat (f x)]] (f x = x)"
+  assumes "\<forall>x. SMT2.trigger [[SMT2.pat (f x)]] (f x = x)"
   shows "f 1 = 1"
-  using assms by smt
+  using assms using [[smt2_trace]] by smt2
 
 lemma
-  assumes "\<forall>x y. SMT.trigger [[SMT.pat (f x), SMT.pat (g y)]] (f x = g y)"
+  assumes "\<forall>x y. SMT2.trigger [[SMT2.pat (f x), SMT2.pat (g y)]] (f x = g y)"
   shows "f a = g b"
-  using assms by smt
+  using assms by smt2
 
 lemma
-  assumes "ALL x. SMT.trigger [[SMT.pat (P x)]] (P x --> Q x)"
+  assumes "ALL x. SMT2.trigger [[SMT2.pat (P x)]] (P x --> Q x)"
   and "P t"
   shows "Q t"
-  using assms by smt
+  using assms by smt2
 
 lemma
-  assumes "ALL x. SMT.trigger [[SMT.pat (P x), SMT.pat (Q x)]]
+  assumes "ALL x. SMT2.trigger [[SMT2.pat (P x), SMT2.pat (Q x)]]
     (P x & Q x --> R x)"
   and "P t" and "Q t"
   shows "R t"
-  using assms by smt
+  using assms by smt2
 
 lemma
-  assumes "ALL x. SMT.trigger [[SMT.pat (P x)], [SMT.pat (Q x)]]
+  assumes "ALL x. SMT2.trigger [[SMT2.pat (P x)], [SMT2.pat (Q x)]]
     ((P x --> R x) & (Q x --> R x))"
   and "P t | Q t"
   shows "R t"
-  using assms by smt
-
-lemma
-  assumes "ALL x. SMT.trigger [[SMT.pat (P x)]] (SMT.weight 2 (P x --> Q x))"
-  and "P t"
-  shows "Q t"
-  using assms by smt
+  using assms by smt2
 
 lemma
-  assumes "ALL x. SMT.weight 1 (P x --> Q x)"
+  assumes "ALL x. SMT2.trigger [[SMT2.pat (P x)]] (SMT2.weight 2 (P x --> Q x))"
   and "P t"
   shows "Q t"
-  using assms by smt
+  using assms by smt2
+
+lemma
+  assumes "ALL x. SMT2.weight 1 (P x --> Q x)"
+  and "P t"
+  shows "Q t"
+  using assms by smt2
 
 
-
-section {* Meta logical connectives *}
+section {* Meta-logical connectives *}
 
 lemma
   "True \<Longrightarrow> True"
@@ -252,8 +248,7 @@
   "(\<And>x y. h x y \<and> h y x) \<Longrightarrow> \<forall>x. h x x"
   "(p \<or> q) \<and> \<not>p \<Longrightarrow> q"
   "(a \<and> b) \<or> (c \<and> d) \<Longrightarrow> (a \<and> b) \<or> (c \<and> d)"
-  by smt+
-
+  by smt+ (* smt2 FIXME: Option *)
 
 
 section {* Natural numbers *}
@@ -264,7 +259,7 @@
   "(0::nat) < 1"
   "(0::nat) \<le> 1"
   "(123456789::nat) < 2345678901"
-  by smt+
+  by smt2+
 
 lemma
   "Suc 0 = 1"
@@ -272,7 +267,7 @@
   "x < Suc x"
   "(Suc x = Suc y) = (x = y)"
   "Suc (x + y) < Suc x + Suc y"
-  by smt+
+  by smt2+
 
 lemma
   "(x::nat) + 0 = x"
@@ -280,15 +275,15 @@
   "x + y = y + x"
   "x + (y + z) = (x + y) + z"
   "(x + y = 0) = (x = 0 \<and> y = 0)"
-  by smt+
+  by smt2+
 
-lemma 
+lemma
   "(x::nat) - 0 = x"
   "x < y \<longrightarrow> x - y = 0"
   "x - y = 0 \<or> y - x = 0"
   "(x - y) + y = (if x < y then y else x)"
-  "x - y - z = x - (y + z)" 
-  by smt+
+  "x - y - z = x - (y + z)"
+  by smt2+
 
 lemma
   "(x::nat) * 0 = 0"
@@ -296,7 +291,7 @@
   "x * 1 = x"
   "1 * x = x"
   "3 * x = x * 3"
-  by smt+
+  by smt2+
 
 lemma
   "(0::nat) div 0 = 0"
@@ -310,8 +305,8 @@
   "(3::nat) div 3 = 1"
   "(x::nat) div 3 \<le> x"
   "(x div 3 = x) = (x = 0)"
-  using [[z3_with_extensions]]
-  by smt+
+  using [[z3_new_extensions]]
+  by smt2+
 
 lemma
   "(0::nat) mod 0 = 0"
@@ -325,14 +320,14 @@
   "(3::nat) mod 3 = 0"
   "x mod 3 < 3"
   "(x mod 3 = x) = (x < 3)"
-  using [[z3_with_extensions]]
-  by smt+
+  using [[z3_new_extensions]]
+  by smt2+
 
 lemma
   "(x::nat) = x div 1 * 1 + x mod 1"
   "x = x div 3 * 3 + x mod 3"
-  using [[z3_with_extensions]]
-  by smt+
+  using [[z3_new_extensions]]
+  by smt2+
 
 lemma
   "min (x::nat) y \<le> x"
@@ -341,7 +336,7 @@
   "z < x \<and> z < y \<longrightarrow> z < min x y"
   "min x y = min y x"
   "min x 0 = 0"
-  by smt+
+  by smt2+
 
 lemma
   "max (x::nat) y \<ge> x"
@@ -350,7 +345,7 @@
   "z > x \<and> z > y \<longrightarrow> z > max x y"
   "max x y = max y x"
   "max x 0 = x"
-  by smt+
+  by smt2+
 
 lemma
   "0 \<le> (x::nat)"
@@ -366,8 +361,7 @@
   "x \<le> y \<longrightarrow> y < z \<longrightarrow> x \<le> z"
   "x < y \<longrightarrow> y < z \<longrightarrow> x < z"
   "x < y \<and> y < z \<longrightarrow> \<not>(z < x)"
-  by smt+
-
+  by smt2+
 
 
 section {* Integers *}
@@ -382,7 +376,7 @@
   "-123 + 345 < (567::int)"
   "(123456789::int) < 2345678901"
   "(-123456789::int) < 2345678901"
-  by smt+
+  by smt2+
 
 lemma
   "(x::int) + 0 = x"
@@ -390,7 +384,7 @@
   "x + y = y + x"
   "x + (y + z) = (x + y) + z"
   "(x + y = 0) = (x = -y)"
-  by smt+
+  by smt2+
 
 lemma
   "(-1::int) = - 1"
@@ -398,16 +392,16 @@
   "-(x::int) < 0 \<longleftrightarrow> x > 0"
   "x > 0 \<longrightarrow> -x < 0"
   "x < 0 \<longrightarrow> -x > 0"
-  by smt+
+  by smt2+
 
-lemma 
+lemma
   "(x::int) - 0 = x"
   "0 - x = -x"
   "x < y \<longrightarrow> x - y < 0"
   "x - y = -(y - x)"
   "x - y = -y + x"
-  "x - y - z = x - (y + z)" 
-  by smt+
+  "x - y - z = x - (y + z)"
+  by smt2+
 
 lemma
   "(x::int) * 0 = 0"
@@ -417,7 +411,7 @@
   "x * -1 = -x"
   "-1 * x = -x"
   "3 * x = x * 3"
-  by smt+
+  by smt2+
 
 lemma
   "(0::int) div 0 = 0"
@@ -444,8 +438,8 @@
   "(-1::int) div -3 = 0"
   "(-3::int) div -3 = 1"
   "(-5::int) div -3 = 1"
-  using [[z3_with_extensions]]
-  by smt+
+  using [[z3_new_extensions]]
+  by smt2+
 
 lemma
   "(0::int) mod 0 = 0"
@@ -474,14 +468,14 @@
   "(-5::int) mod -3 = -2"
   "x mod 3 < 3"
   "(x mod 3 = x) \<longrightarrow> (x < 3)"
-  using [[z3_with_extensions]]
-  by smt+
+  using [[z3_new_extensions]]
+  by smt2+
 
 lemma
   "(x::int) = x div 1 * 1 + x mod 1"
   "x = x div 3 * 3 + x mod 3"
-  using [[z3_with_extensions]]
-  by smt+
+  using [[z3_new_extensions]]
+  by smt2+
 
 lemma
   "abs (x::int) \<ge> 0"
@@ -489,7 +483,7 @@
   "(x \<ge> 0) = (abs x = x)"
   "(x \<le> 0) = (abs x = -x)"
   "abs (abs x) = abs x"
-  by smt+
+  by smt2+
 
 lemma
   "min (x::int) y \<le> x"
@@ -498,7 +492,7 @@
   "min x y = min y x"
   "x \<ge> 0 \<longrightarrow> min x 0 = 0"
   "min x y \<le> abs (x + y)"
-  by smt+
+  by smt2+
 
 lemma
   "max (x::int) y \<ge> x"
@@ -507,7 +501,7 @@
   "max x y = max y x"
   "x \<ge> 0 \<longrightarrow> max x 0 = x"
   "max x y \<ge> - abs x - abs y"
-  by smt+
+  by smt2+
 
 lemma
   "0 < (x::int) \<and> x \<le> 1 \<longrightarrow> x = 1"
@@ -522,8 +516,7 @@
   "x \<le> y \<longrightarrow> y < z \<longrightarrow> x \<le> z"
   "x < y \<longrightarrow> y < z \<longrightarrow> x < z"
   "x < y \<and> y < z \<longrightarrow> \<not>(z < x)"
-  by smt+
-
+  by smt2+
 
 
 section {* Reals *}
@@ -539,7 +532,7 @@
   "-123 + 345 < (567::real)"
   "(123456789::real) < 2345678901"
   "(-123456789::real) < 2345678901"
-  by smt+
+  by smt2+
 
 lemma
   "(x::real) + 0 = x"
@@ -547,7 +540,7 @@
   "x + y = y + x"
   "x + (y + z) = (x + y) + z"
   "(x + y = 0) = (x = -y)"
-  by smt+
+  by smt2+
 
 lemma
   "(-1::real) = - 1"
@@ -555,7 +548,7 @@
   "-(x::real) < 0 \<longleftrightarrow> x > 0"
   "x > 0 \<longrightarrow> -x < 0"
   "x < 0 \<longrightarrow> -x > 0"
-  by smt+
+  by smt2+
 
 lemma
   "(x::real) - 0 = x"
@@ -563,8 +556,8 @@
   "x < y \<longrightarrow> x - y < 0"
   "x - y = -(y - x)"
   "x - y = -y + x"
-  "x - y - z = x - (y + z)" 
-  by smt+
+  "x - y - z = x - (y + z)"
+  by smt2+
 
 lemma
   "(x::real) * 0 = 0"
@@ -574,7 +567,7 @@
   "x * -1 = -x"
   "-1 * x = -x"
   "3 * x = x * 3"
-  by smt+
+  by smt2+
 
 lemma
   "(1/2 :: real) < 1"
@@ -585,16 +578,16 @@
   "(x::real) / 1 = x"
   "x > 0 \<longrightarrow> x / 3 < x"
   "x < 0 \<longrightarrow> x / 3 > x"
-  using [[z3_with_extensions]]
-  by smt+
+  using [[z3_new_extensions]]
+  by smt2+
 
 lemma
   "(3::real) * (x / 3) = x"
   "(x * 3) / 3 = x"
   "x > 0 \<longrightarrow> 2 * x / 3 < x"
   "x < 0 \<longrightarrow> 2 * x / 3 > x"
-  using [[z3_with_extensions]]
-  by smt+
+  using [[z3_new_extensions]]
+  by smt2+
 
 lemma
   "abs (x::real) \<ge> 0"
@@ -602,7 +595,7 @@
   "(x \<ge> 0) = (abs x = x)"
   "(x \<le> 0) = (abs x = -x)"
   "abs (abs x) = abs x"
-  by smt+
+  by smt2+
 
 lemma
   "min (x::real) y \<le> x"
@@ -611,7 +604,7 @@
   "min x y = min y x"
   "x \<ge> 0 \<longrightarrow> min x 0 = 0"
   "min x y \<le> abs (x + y)"
-  by smt+
+  by smt2+
 
 lemma
   "max (x::real) y \<ge> x"
@@ -620,7 +613,7 @@
   "max x y = max y x"
   "x \<ge> 0 \<longrightarrow> max x 0 = x"
   "max x y \<ge> - abs x - abs y"
-  by smt+
+  by smt2+
 
 lemma
   "x \<le> (x::real)"
@@ -633,8 +626,7 @@
   "x \<le> y \<longrightarrow> y < z \<longrightarrow> x \<le> z"
   "x < y \<longrightarrow> y < z \<longrightarrow> x < z"
   "x < y \<and> y < z \<longrightarrow> \<not>(z < x)"
-  by smt+
-
+  by smt2+
 
 
 section {* Datatypes, Records, and Typedefs *}
@@ -657,7 +649,7 @@
   "(fst (x, y) = snd (x, y)) = (x = y)"
   "(fst p = snd p) = (p = (snd p, fst p))"
   using fst_conv snd_conv pair_collapse
-  by smt+
+  by smt2+
 
 lemma
   "[x] \<noteq> Nil"
@@ -670,13 +662,13 @@
   "hd (tl [x, y, z]) = y"
   "tl (tl [x, y, z]) = [z]"
   using list.sel(1,3) list.simps
-  by smt+
+  by smt2+
 
 lemma
   "fst (hd [(a, b)]) = a"
   "snd (hd [(a, b)]) = b"
   using fst_conv snd_conv pair_collapse list.sel(1,3) list.simps
-  by smt+
+  by smt2+
 
 
 subsubsection {* Records *}
@@ -694,7 +686,7 @@
   "cx p1 \<noteq> cx p2 \<longrightarrow> p1 \<noteq> p2"
   "cy p1 \<noteq> cy p2 \<longrightarrow> p1 \<noteq> p2"
   using point.simps
-  by smt+
+  by smt2+
 
 lemma
   "cx \<lparr> cx = 3, cy = 4 \<rparr> = 3"
@@ -705,8 +697,7 @@
   "p = \<lparr> cx = 3, cy = 4 \<rparr> \<longrightarrow> p \<lparr> cx := 3 \<rparr> \<lparr> cy := 4 \<rparr> = p"
   "p = \<lparr> cx = 3, cy = 4 \<rparr> \<longrightarrow> p \<lparr> cy := 4 \<rparr> \<lparr> cx := 3 \<rparr> = p"
   using point.simps
-  using [[z3_options="AUTO_CONFIG=false"]]
-  by smt+
+  by smt2+
 
 lemma
   "cy (p \<lparr> cx := a \<rparr>) = cy p"
@@ -722,7 +713,7 @@
   "cy p1 \<noteq> cy p2 \<longrightarrow> p1 \<noteq> p2"
   "black p1 \<noteq> black p2 \<longrightarrow> p1 \<noteq> p2"
   using point.simps bw_point.simps
-  by smt+
+  by smt2+
 
 lemma
   "cx \<lparr> cx = 3, cy = 4, black = b \<rparr> = 3"
@@ -738,8 +729,7 @@
   "p = \<lparr> cx = 3, cy = 4, black = True \<rparr> \<longrightarrow>
      p \<lparr> black := True \<rparr> \<lparr> cx := 3 \<rparr> \<lparr> cy := 4 \<rparr> = p"
   using point.simps bw_point.simps
-  using [[z3_options="AUTO_CONFIG=false"]]
-  by smt+
+  by smt+ (* smt2 FIXME: Option *)
 
 lemma
   "\<lparr> cx = 3, cy = 4, black = b \<rparr> \<lparr> black := w \<rparr> = \<lparr> cx = 3, cy = 4, black = w \<rparr>"
@@ -773,8 +763,7 @@
   "nplus n1 n2 = n3"
   using n1_def n2_def n3_def nplus_def
   using three_def' Rep_three Abs_three_inverse
-  using [[z3_options="AUTO_CONFIG=false"]]
-  by smt+
+  by smt2+
 
 
 subsection {* With support by the SMT solver (but without proofs) *}
@@ -795,8 +784,8 @@
   "(fst (x, y) = snd (x, y)) = (x = y)"
   "(fst p = snd p) = (p = (snd p, fst p))"
   using fst_conv snd_conv pair_collapse
-  using [[smt_datatypes, smt_oracle, z3_with_extensions]]
-  by smt+
+  using [[smt2_oracle, z3_new_extensions]]
+  by smt2+
 
 lemma
   "[x] \<noteq> Nil"
@@ -809,15 +798,15 @@
   "hd (tl [x, y, z]) = y"
   "tl (tl [x, y, z]) = [z]"
   using list.sel(1,3)
-  using [[smt_datatypes, smt_oracle, z3_with_extensions]]
-  by smt+
+  using [[smt2_oracle, z3_new_extensions]]
+  by smt2+
 
 lemma
   "fst (hd [(a, b)]) = a"
   "snd (hd [(a, b)]) = b"
   using fst_conv snd_conv pair_collapse list.sel(1,3)
-  using [[smt_datatypes, smt_oracle, z3_with_extensions]]
-  by smt+
+  using [[smt2_oracle, z3_new_extensions]]
+  by smt2+
 
 
 subsubsection {* Records *}
@@ -828,9 +817,8 @@
   "cx p1 \<noteq> cx p2 \<longrightarrow> p1 \<noteq> p2"
   "cy p1 \<noteq> cy p2 \<longrightarrow> p1 \<noteq> p2"
   using point.simps
-  using [[smt_datatypes, smt_oracle, z3_with_extensions]]
-  using [[z3_options="AUTO_CONFIG=false"]]
-  by smt+
+  using [[smt2_oracle, z3_new_extensions]]
+  by smt2+
 
 lemma
   "cx \<lparr> cx = 3, cy = 4 \<rparr> = 3"
@@ -841,18 +829,16 @@
   "p = \<lparr> cx = 3, cy = 4 \<rparr> \<longrightarrow> p \<lparr> cx := 3 \<rparr> \<lparr> cy := 4 \<rparr> = p"
   "p = \<lparr> cx = 3, cy = 4 \<rparr> \<longrightarrow> p \<lparr> cy := 4 \<rparr> \<lparr> cx := 3 \<rparr> = p"
   using point.simps
-  using [[smt_datatypes, smt_oracle, z3_with_extensions]]
-  using [[z3_options="AUTO_CONFIG=false"]]
-  by smt+
+  using [[smt2_oracle, z3_new_extensions]]
+  by smt2+
 
 lemma
   "cy (p \<lparr> cx := a \<rparr>) = cy p"
   "cx (p \<lparr> cy := a \<rparr>) = cx p"
   "p \<lparr> cx := 3 \<rparr> \<lparr> cy := 4 \<rparr> = p \<lparr> cy := 4 \<rparr> \<lparr> cx := 3 \<rparr>"
   using point.simps
-  using [[smt_datatypes, smt_oracle, z3_with_extensions]]
-  using [[z3_options="AUTO_CONFIG=false"]]
-  by smt+
+  using [[smt2_oracle, z3_new_extensions]]
+  by smt2+
 
 lemma
   "p1 = p2 \<longrightarrow> cx p1 = cx p2"
@@ -862,8 +848,8 @@
   "cy p1 \<noteq> cy p2 \<longrightarrow> p1 \<noteq> p2"
   "black p1 \<noteq> black p2 \<longrightarrow> p1 \<noteq> p2"
   using point.simps bw_point.simps
-  using [[smt_datatypes, smt_oracle, z3_with_extensions]]
-  by smt+
+  using [[smt2_oracle, z3_new_extensions]]
+  by smt2+
 
 lemma
   "cx \<lparr> cx = 3, cy = 4, black = b \<rparr> = 3"
@@ -879,9 +865,8 @@
   "p = \<lparr> cx = 3, cy = 4, black = True \<rparr> \<longrightarrow>
      p \<lparr> black := True \<rparr> \<lparr> cx := 3 \<rparr> \<lparr> cy := 4 \<rparr> = p"
   using point.simps bw_point.simps
-  using [[smt_datatypes, smt_oracle, z3_with_extensions]]
-  using [[z3_options="AUTO_CONFIG=false"]]
-  by smt+
+  using [[smt2_oracle, z3_new_extensions]]
+  by smt2+
 
 lemma
   "\<lparr> cx = 3, cy = 4, black = b \<rparr> \<lparr> black := w \<rparr> = \<lparr> cx = 3, cy = 4, black = w \<rparr>"
@@ -893,9 +878,8 @@
   "p \<lparr> cx := 3 \<rparr> \<lparr> cy := 4 \<rparr> \<lparr> black := True \<rparr> =
      p \<lparr> black := True \<rparr> \<lparr> cy := 4 \<rparr> \<lparr> cx := 3 \<rparr>"
   using point.simps bw_point.simps
-  using [[smt_datatypes, smt_oracle, z3_with_extensions]]
-  using [[z3_options="AUTO_CONFIG=false"]]
-  by smt
+  using [[smt2_oracle, z3_new_extensions]]
+  by smt2
 
 
 subsubsection {* Type definitions *}
@@ -907,10 +891,8 @@
   "nplus n1 n1 = n2"
   "nplus n1 n2 = n3"
   using n1_def n2_def n3_def nplus_def
-  using [[smt_datatypes, smt_oracle, z3_with_extensions]]
-  using [[z3_options="AUTO_CONFIG=false"]]
-  by smt+
-
+  using [[smt2_oracle, z3_new_extensions]]
+  by smt2+
 
 
 section {* Function updates *}
@@ -924,7 +906,7 @@
   "i1 = i2 \<longrightarrow> (f (i1 := v1, i2 := v2)) i1 = v2"
   "i1 \<noteq> i2 \<and>i1 \<noteq> i3 \<and>  i2 \<noteq> i3 \<longrightarrow> (f (i1 := v1, i2 := v2)) i3 = f i3"
   using fun_upd_same fun_upd_apply
-  by smt+
+  by smt2+
 
 
 
@@ -932,7 +914,7 @@
 
 lemma Empty: "x \<notin> {}" by simp
 
-lemmas smt_sets = Empty UNIV_I Un_iff Int_iff
+lemmas smt2_sets = Empty UNIV_I Un_iff Int_iff
 
 lemma
   "x \<notin> {}"
@@ -950,6 +932,6 @@
   "x \<in> P \<inter> P \<longleftrightarrow> x \<in> P"
   "x \<in> P \<inter> (Q \<inter> R) \<longleftrightarrow> x \<in> (P \<inter> Q) \<inter> R"
   "{x. x \<in> P} = {y. y \<in> P}"
-  by (smt smt_sets)+
+  by (smt2 smt2_sets)+
 
 end
--- a/src/HOL/SMT_Examples/SMT_Word_Examples.certs	Fri Mar 14 16:54:01 2014 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,58 +0,0 @@
-ef8166854c461e296fe9c596b7a3fe12065b0c65 1 0
-unsat
-d1ec4aa8c4a5f474e8dbb8a8acbdd12fc33b0cda 1 0
-unsat
-03dee604b20d98218bc88a69efcbf0f1c6dc057a 1 0
-unsat
-6a68bdb5b2a92a239021f6188a807622fe7b8213 1 0
-unsat
-9be3195f24c1786963c05e51e63a24efa7c83141 1 0
-unsat
-608ed753221bdf2f1769ea3686a0f970108a7087 1 0
-unsat
-610484e81fc38952a9a2cb0bfc83d424f7f96ca8 1 0
-unsat
-0a06a4c179bec3512f3dc01e338f246fca223ddb 1 0
-unsat
-dd232118189a55ac7fc80599d2738be8bbaa9333 1 0
-unsat
-8426f9081bd693e21cd8b2e13d07cea1e69e8abd 1 0
-unsat
-8d83ab1c5a55640d0165bbd736d2cc217bcc2efd 1 0
-unsat
-542ef8141028455b42a51f60e3981a74373a60b3 1 0
-unsat
-564709a03da50b938c3b1ab2a8a2f0dc8d8a4749 1 0
-unsat
-c4acaeb4324634878481e3faae3beae53a641067 1 0
-unsat
-873ce0289bcfaf43a446c6ed55bec4289eea0ffd 1 0
-unsat
-8383b80b5e8011f2b51c01ea89c14ce766f5a82b 1 0
-unsat
-6694dc1c5420588e5e48281a8835ac019bfb1aa7 1 0
-unsat
-4094196f5d25f48682e6634b4326469abc38d250 1 0
-unsat
-0597f614ff89c7376d01987b4737ab991b5a321c 1 0
-unsat
-44f955a3f3fab3f5203ec29edc7e00e7cb81bedc 1 0
-unsat
-927e5f0e88fadf6d2f604b1d863a37fc682f942b 1 0
-unsat
-818922160b53f843888d258a1ef7e5d5ddf5129f 1 0
-unsat
-afc6dff121c48475665b0ef064826ffa2cad0e85 1 0
-unsat
-b9ab61d9521faeaa45ec63bff4581742c3e6c550 1 0
-unsat
-8e60769fce6622cdca312aa54d4a77329a99dac2 1 0
-unsat
-bd55726cefc783f8e9ef4ad38596e1f24cca3663 1 0
-unsat
-4e48efd5c9874aedf200e06875d5597b31d98075 1 0
-unsat
-e5c27ae0a583eeafeaa4ef3c59b1b4ec53e06b0f 1 0
-unsat
-7d3ef49480d3ed3a7e5f2d7a12e7108cf7fc7819 1 0
-unsat
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SMT_Examples/SMT_Word_Examples.certs2	Fri Mar 14 17:32:11 2014 +0100
@@ -0,0 +1,369 @@
+d47653c43412ab1eb54730f2f5a4f4bdf44fcb5a 8 0
+unsat
+((set-logic <null>)
+(proof
+(let ((@x36 (monotonicity (rewrite (= (bvneg (_ bv5 4)) (_ bv11 4))) (= (= (_ bv11 4) (bvneg (_ bv5 4))) (= (_ bv11 4) (_ bv11 4))))))
+(let ((@x40 (trans @x36 (rewrite (= (= (_ bv11 4) (_ bv11 4)) true)) (= (= (_ bv11 4) (bvneg (_ bv5 4))) true))))
+(let ((@x47 (trans (monotonicity @x40 (= (not (= (_ bv11 4) (bvneg (_ bv5 4)))) (not true))) (rewrite (= (not true) false)) (= (not (= (_ bv11 4) (bvneg (_ bv5 4)))) false))))
+(mp (asserted (not (= (_ bv11 4) (bvneg (_ bv5 4))))) @x47 false))))))
+
+da258c2a22a4a00129b43deac09b90d379043340 7 0
+unsat
+((set-logic <null>)
+(proof
+(let ((@x33 (monotonicity (rewrite (= (= (_ bv11 4) (_ bv11 4)) true)) (= (not (= (_ bv11 4) (_ bv11 4))) (not true)))))
+(let ((@x37 (trans @x33 (rewrite (= (not true) false)) (= (not (= (_ bv11 4) (_ bv11 4))) false))))
+(mp (asserted (not (= (_ bv11 4) (_ bv11 4)))) @x37 false)))))
+
+f2b47b92988d2f0c1404b109b621ba9a6c2b9d1c 7 0
+unsat
+((set-logic <null>)
+(proof
+(let ((@x36 (monotonicity (rewrite (= (bvult (_ bv23 8) (_ bv27 8)) true)) (= (not (bvult (_ bv23 8) (_ bv27 8))) (not true)))))
+(let ((@x40 (trans @x36 (rewrite (= (not true) false)) (= (not (bvult (_ bv23 8) (_ bv27 8))) false))))
+(mp (asserted (not (bvult (_ bv23 8) (_ bv27 8)))) @x40 false)))))
+
+1c22485fb98e3caa4d683df39ead427fc3568432 9 0
+unsat
+((set-logic <null>)
+(proof
+(let ((@x36 (monotonicity (rewrite (= (bvadd (_ bv27 5) (_ bv11 5)) (_ bv6 5))) (= (= (bvadd (_ bv27 5) (_ bv11 5)) (_ bv6 5)) (= (_ bv6 5) (_ bv6 5))))))
+(let ((@x40 (trans @x36 (rewrite (= (= (_ bv6 5) (_ bv6 5)) true)) (= (= (bvadd (_ bv27 5) (_ bv11 5)) (_ bv6 5)) true))))
+(let ((@x43 (monotonicity @x40 (= (not (= (bvadd (_ bv27 5) (_ bv11 5)) (_ bv6 5))) (not true)))))
+(let ((@x47 (trans @x43 (rewrite (= (not true) false)) (= (not (= (bvadd (_ bv27 5) (_ bv11 5)) (_ bv6 5))) false))))
+(mp (asserted (not (= (bvadd (_ bv27 5) (_ bv11 5)) (_ bv6 5)))) @x47 false)))))))
+
+651f74a079a4aa15d5d621208d8e038db0369475 9 0
+unsat
+((set-logic <null>)
+(proof
+(let ((@x36 (monotonicity (rewrite (= (bvmul (_ bv7 8) (_ bv3 8)) (_ bv21 8))) (= (= (bvmul (_ bv7 8) (_ bv3 8)) (_ bv21 8)) (= (_ bv21 8) (_ bv21 8))))))
+(let ((@x40 (trans @x36 (rewrite (= (= (_ bv21 8) (_ bv21 8)) true)) (= (= (bvmul (_ bv7 8) (_ bv3 8)) (_ bv21 8)) true))))
+(let ((@x43 (monotonicity @x40 (= (not (= (bvmul (_ bv7 8) (_ bv3 8)) (_ bv21 8))) (not true)))))
+(let ((@x47 (trans @x43 (rewrite (= (not true) false)) (= (not (= (bvmul (_ bv7 8) (_ bv3 8)) (_ bv21 8))) false))))
+(mp (asserted (not (= (bvmul (_ bv7 8) (_ bv3 8)) (_ bv21 8)))) @x47 false)))))))
+
+eefb4f0a8b690f38fb11c31757b3209b40cfd1c5 9 0
+unsat
+((set-logic <null>)
+(proof
+(let ((@x41 (monotonicity (rewrite (= (bvsub (_ bv11 8) (_ bv27 8)) (_ bv240 8))) (rewrite (= (bvneg (_ bv16 8)) (_ bv240 8))) (= (= (bvsub (_ bv11 8) (_ bv27 8)) (bvneg (_ bv16 8))) (= (_ bv240 8) (_ bv240 8))))))
+(let ((@x45 (trans @x41 (rewrite (= (= (_ bv240 8) (_ bv240 8)) true)) (= (= (bvsub (_ bv11 8) (_ bv27 8)) (bvneg (_ bv16 8))) true))))
+(let ((@x48 (monotonicity @x45 (= (not (= (bvsub (_ bv11 8) (_ bv27 8)) (bvneg (_ bv16 8)))) (not true)))))
+(let ((@x52 (trans @x48 (rewrite (= (not true) false)) (= (not (= (bvsub (_ bv11 8) (_ bv27 8)) (bvneg (_ bv16 8)))) false))))
+(mp (asserted (not (= (bvsub (_ bv11 8) (_ bv27 8)) (bvneg (_ bv16 8))))) @x52 false)))))))
+
+e251dcc0ad168cb65c5bc1d32039c72ca2609bb3 7 0
+unsat
+((set-logic <null>)
+(proof
+(let ((@x33 (monotonicity (rewrite (= (= (_ bv11 5) (_ bv11 5)) true)) (= (not (= (_ bv11 5) (_ bv11 5))) (not true)))))
+(let ((@x37 (trans @x33 (rewrite (= (not true) false)) (= (not (= (_ bv11 5) (_ bv11 5))) false))))
+(mp (asserted (not (= (_ bv11 5) (_ bv11 5)))) @x37 false)))))
+
+4f488dde65b4a70d1d31d589531d3445a9be689f 11 0
+unsat
+((set-logic <null>)
+(proof
+(let ((@x40 (monotonicity (rewrite (= (bvneg (_ bv40 7)) (_ bv88 7))) (= (bvadd (bvneg (_ bv40 7)) (_ bv1 7)) (bvadd (_ bv88 7) (_ bv1 7))))))
+(let ((@x45 (trans @x40 (rewrite (= (bvadd (_ bv88 7) (_ bv1 7)) (_ bv89 7))) (= (bvadd (bvneg (_ bv40 7)) (_ bv1 7)) (_ bv89 7)))))
+(let ((@x50 (monotonicity @x45 (rewrite (= (bvneg (_ bv39 7)) (_ bv89 7))) (= (= (bvadd (bvneg (_ bv40 7)) (_ bv1 7)) (bvneg (_ bv39 7))) (= (_ bv89 7) (_ bv89 7))))))
+(let ((@x54 (trans @x50 (rewrite (= (= (_ bv89 7) (_ bv89 7)) true)) (= (= (bvadd (bvneg (_ bv40 7)) (_ bv1 7)) (bvneg (_ bv39 7))) true))))
+(let ((@x57 (monotonicity @x54 (= (not (= (bvadd (bvneg (_ bv40 7)) (_ bv1 7)) (bvneg (_ bv39 7)))) (not true)))))
+(let ((@x61 (trans @x57 (rewrite (= (not true) false)) (= (not (= (bvadd (bvneg (_ bv40 7)) (_ bv1 7)) (bvneg (_ bv39 7)))) false))))
+(mp (asserted (not (= (bvadd (bvneg (_ bv40 7)) (_ bv1 7)) (bvneg (_ bv39 7))))) @x61 false)))))))))
+
+2e53bd8b513a3dc9dee350d0d8b1c315cf2b2449 19 0
+unsat
+((set-logic <null>)
+(proof
+(let ((?x13 (bvadd |b$| |c$|)))
+(let ((?x14 (bvadd ?x13 |a$|)))
+(let ((?x8 (bvmul (_ bv2 32) |b$|)))
+(let ((?x9 (bvadd |a$| ?x8)))
+(let ((?x11 (bvadd ?x9 |c$|)))
+(let ((?x12 (bvsub ?x11 |b$|)))
+(let (($x15 (= ?x12 ?x14)))
+(let (($x16 (not $x15)))
+(let ((@x56 (rewrite (= (= (bvadd |a$| |b$| |c$|) (bvadd |a$| |b$| |c$|)) true))))
+(let ((@x46 (rewrite (= (bvsub (bvadd |a$| ?x8 |c$|) |b$|) (bvadd |a$| |b$| |c$|)))))
+(let ((@x44 (monotonicity (rewrite (= ?x11 (bvadd |a$| ?x8 |c$|))) (= ?x12 (bvsub (bvadd |a$| ?x8 |c$|) |b$|)))))
+(let ((@x54 (monotonicity (trans @x44 @x46 (= ?x12 (bvadd |a$| |b$| |c$|))) (rewrite (= ?x14 (bvadd |a$| |b$| |c$|))) (= $x15 (= (bvadd |a$| |b$| |c$|) (bvadd |a$| |b$| |c$|))))))
+(let ((@x61 (monotonicity (trans @x54 @x56 (= $x15 true)) (= $x16 (not true)))))
+(let ((@x65 (trans @x61 (rewrite (= (not true) false)) (= $x16 false))))
+(mp (asserted $x16) @x65 false)))))))))))))))))
+
+570be43092e6421d4222501467362afbd680c2e2 18 0
+unsat
+((set-logic <null>)
+(proof
+(let ((?x9 (bvmul (_ bv4 4) |x$|)))
+(let (($x10 (= ?x9 (_ bv4 4))))
+(let (($x41 (= (_ bv5 4) |x$|)))
+(let (($x54 (not (or (not $x41) (= (_ bv4 4) ?x9)))))
+(let ((@x46 (monotonicity (rewrite (= (= |x$| (_ bv5 4)) $x41)) (= (not (= |x$| (_ bv5 4))) (not $x41)))))
+(let ((@x53 (monotonicity @x46 (rewrite (= $x10 (= (_ bv4 4) ?x9))) (= (or (not (= |x$| (_ bv5 4))) $x10) (or (not $x41) (= (_ bv4 4) ?x9))))))
+(let (($x12 (not (=> (= |x$| (_ bv5 4)) $x10))))
+(let ((@x37 (rewrite (= (=> (= |x$| (_ bv5 4)) $x10) (or (not (= |x$| (_ bv5 4))) $x10)))))
+(let ((@x58 (trans (monotonicity @x37 (= $x12 (not (or (not (= |x$| (_ bv5 4))) $x10)))) (monotonicity @x53 (= (not (or (not (= |x$| (_ bv5 4))) $x10)) $x54)) (= $x12 $x54))))
+(let ((@x65 (monotonicity (|not-or-elim| (mp (asserted $x12) @x58 $x54) $x41) (= ?x9 (bvmul (_ bv4 4) (_ bv5 4))))))
+(let ((@x71 (monotonicity (trans @x65 (rewrite (= (bvmul (_ bv4 4) (_ bv5 4)) (_ bv4 4))) $x10) (= (= (_ bv4 4) ?x9) (= (_ bv4 4) (_ bv4 4))))))
+(let ((@x75 (trans @x71 (rewrite (= (= (_ bv4 4) (_ bv4 4)) true)) (= (= (_ bv4 4) ?x9) true))))
+(let ((@x82 (trans (monotonicity @x75 (= (not (= (_ bv4 4) ?x9)) (not true))) (rewrite (= (not true) false)) (= (not (= (_ bv4 4) ?x9)) false))))
+(mp (|not-or-elim| (mp (asserted $x12) @x58 $x54) (not (= (_ bv4 4) ?x9))) @x82 false))))))))))))))))
+
+2538d74409c652fbc39d33a15f25d18c9bb179bf 9 0
+unsat
+((set-logic <null>)
+(proof
+(let ((@x35 (monotonicity (rewrite (= (bvand (_ bv6 32) (_ bv5 32)) (_ bv4 32))) (= (= (bvand (_ bv6 32) (_ bv5 32)) (_ bv4 32)) (= (_ bv4 32) (_ bv4 32))))))
+(let ((@x39 (trans @x35 (rewrite (= (= (_ bv4 32) (_ bv4 32)) true)) (= (= (bvand (_ bv6 32) (_ bv5 32)) (_ bv4 32)) true))))
+(let ((@x42 (monotonicity @x39 (= (not (= (bvand (_ bv6 32) (_ bv5 32)) (_ bv4 32))) (not true)))))
+(let ((@x46 (trans @x42 (rewrite (= (not true) false)) (= (not (= (bvand (_ bv6 32) (_ bv5 32)) (_ bv4 32))) false))))
+(mp (asserted (not (= (bvand (_ bv6 32) (_ bv5 32)) (_ bv4 32)))) @x46 false)))))))
+
+ea3351a199ecbca8cfc913892024d6ba767f41dc 9 0
+unsat
+((set-logic <null>)
+(proof
+(let ((@x35 (monotonicity (rewrite (= (bvor (_ bv6 8) (_ bv3 8)) (_ bv7 8))) (= (= (bvor (_ bv6 8) (_ bv3 8)) (_ bv7 8)) (= (_ bv7 8) (_ bv7 8))))))
+(let ((@x39 (trans @x35 (rewrite (= (= (_ bv7 8) (_ bv7 8)) true)) (= (= (bvor (_ bv6 8) (_ bv3 8)) (_ bv7 8)) true))))
+(let ((@x42 (monotonicity @x39 (= (not (= (bvor (_ bv6 8) (_ bv3 8)) (_ bv7 8))) (not true)))))
+(let ((@x46 (trans @x42 (rewrite (= (not true) false)) (= (not (= (bvor (_ bv6 8) (_ bv3 8)) (_ bv7 8))) false))))
+(mp (asserted (not (= (bvor (_ bv6 8) (_ bv3 8)) (_ bv7 8)))) @x46 false)))))))
+
+a36393e9d24b671ce68aafbd67bbc7bcd4c32a9f 9 0
+unsat
+((set-logic <null>)
+(proof
+(let ((@x35 (monotonicity (rewrite (= (bvxor (_ bv240 8) (_ bv255 8)) (_ bv15 8))) (= (= (bvxor (_ bv240 8) (_ bv255 8)) (_ bv15 8)) (= (_ bv15 8) (_ bv15 8))))))
+(let ((@x39 (trans @x35 (rewrite (= (= (_ bv15 8) (_ bv15 8)) true)) (= (= (bvxor (_ bv240 8) (_ bv255 8)) (_ bv15 8)) true))))
+(let ((@x42 (monotonicity @x39 (= (not (= (bvxor (_ bv240 8) (_ bv255 8)) (_ bv15 8))) (not true)))))
+(let ((@x46 (trans @x42 (rewrite (= (not true) false)) (= (not (= (bvxor (_ bv240 8) (_ bv255 8)) (_ bv15 8))) false))))
+(mp (asserted (not (= (bvxor (_ bv240 8) (_ bv255 8)) (_ bv15 8)))) @x46 false)))))))
+
+48dba82ab628843121b1cc45b6a662d4282a5dfd 8 0
+unsat
+((set-logic <null>)
+(proof
+(let ((@x34 (monotonicity (rewrite (= (bvnot (_ bv240 16)) (_ bv65295 16))) (= (= (bvnot (_ bv240 16)) (_ bv65295 16)) (= (_ bv65295 16) (_ bv65295 16))))))
+(let ((@x38 (trans @x34 (rewrite (= (= (_ bv65295 16) (_ bv65295 16)) true)) (= (= (bvnot (_ bv240 16)) (_ bv65295 16)) true))))
+(let ((@x45 (trans (monotonicity @x38 (= (not (= (bvnot (_ bv240 16)) (_ bv65295 16))) (not true))) (rewrite (= (not true) false)) (= (not (= (bvnot (_ bv240 16)) (_ bv65295 16))) false))))
+(mp (asserted (not (= (bvnot (_ bv240 16)) (_ bv65295 16)))) @x45 false))))))
+
+ffdba93b71ca27a7275f33b87244eb12ceb5e9c2 9 0
+unsat
+((set-logic <null>)
+(proof
+(let ((@x35 (monotonicity (rewrite (= (concat (_ bv11 4) (_ bv27 8)) (_ bv2843 12))) (= (= (concat (_ bv11 4) (_ bv27 8)) (_ bv2843 12)) (= (_ bv2843 12) (_ bv2843 12))))))
+(let ((@x39 (trans @x35 (rewrite (= (= (_ bv2843 12) (_ bv2843 12)) true)) (= (= (concat (_ bv11 4) (_ bv27 8)) (_ bv2843 12)) true))))
+(let ((@x42 (monotonicity @x39 (= (not (= (concat (_ bv11 4) (_ bv27 8)) (_ bv2843 12))) (not true)))))
+(let ((@x46 (trans @x42 (rewrite (= (not true) false)) (= (not (= (concat (_ bv11 4) (_ bv27 8)) (_ bv2843 12))) false))))
+(mp (asserted (not (= (concat (_ bv11 4) (_ bv27 8)) (_ bv2843 12)))) @x46 false)))))))
+
+24677a2d05cd59ffe782bea3654e41f124fc1b93 9 0
+unsat
+((set-logic <null>)
+(proof
+(let ((@x35 (monotonicity (rewrite (= (concat (_ bv3 4) (_ bv15 6)) (_ bv207 10))) (= (= (concat (_ bv3 4) (_ bv15 6)) (_ bv207 10)) (= (_ bv207 10) (_ bv207 10))))))
+(let ((@x39 (trans @x35 (rewrite (= (= (_ bv207 10) (_ bv207 10)) true)) (= (= (concat (_ bv3 4) (_ bv15 6)) (_ bv207 10)) true))))
+(let ((@x42 (monotonicity @x39 (= (not (= (concat (_ bv3 4) (_ bv15 6)) (_ bv207 10))) (not true)))))
+(let ((@x46 (trans @x42 (rewrite (= (not true) false)) (= (not (= (concat (_ bv3 4) (_ bv15 6)) (_ bv207 10))) false))))
+(mp (asserted (not (= (concat (_ bv3 4) (_ bv15 6)) (_ bv207 10)))) @x46 false)))))))
+
+bf709bf2b13e6bf4f9668ad197c5d7a4ca581525 8 0
+unsat
+((set-logic <null>)
+(proof
+(let ((@x50 (monotonicity (rewrite (= ((_ extract 2 1) (_ bv6 4)) (_ bv3 2))) (= (= ((_ extract 2 1) (_ bv6 4)) (_ bv3 2)) (= (_ bv3 2) (_ bv3 2))))))
+(let ((@x54 (trans @x50 (rewrite (= (= (_ bv3 2) (_ bv3 2)) true)) (= (= ((_ extract 2 1) (_ bv6 4)) (_ bv3 2)) true))))
+(let ((@x61 (trans (monotonicity @x54 (= (not (= ((_ extract 2 1) (_ bv6 4)) (_ bv3 2))) (not true))) (rewrite (= (not true) false)) (= (not (= ((_ extract 2 1) (_ bv6 4)) (_ bv3 2))) false))))
+(mp (asserted (not (= ((_ extract 2 1) (_ bv6 4)) (_ bv3 2)))) @x61 false))))))
+
+f487669e8e249c60376443304a5a78c58eddd1cc 9 0
+unsat
+((set-logic <null>)
+(proof
+(let ((@x34 (monotonicity (rewrite (= ((_ zero_extend 6) (_ bv10 4)) (_ bv10 10))) (= (= ((_ zero_extend 6) (_ bv10 4)) (_ bv10 10)) (= (_ bv10 10) (_ bv10 10))))))
+(let ((@x38 (trans @x34 (rewrite (= (= (_ bv10 10) (_ bv10 10)) true)) (= (= ((_ zero_extend 6) (_ bv10 4)) (_ bv10 10)) true))))
+(let ((@x41 (monotonicity @x38 (= (not (= ((_ zero_extend 6) (_ bv10 4)) (_ bv10 10))) (not true)))))
+(let ((@x45 (trans @x41 (rewrite (= (not true) false)) (= (not (= ((_ zero_extend 6) (_ bv10 4)) (_ bv10 10))) false))))
+(mp (asserted (not (= ((_ zero_extend 6) (_ bv10 4)) (_ bv10 10)))) @x45 false)))))))
+
+e6179d5000250fb81646a549216cd9ad7b2619a2 9 0
+unsat
+((set-logic <null>)
+(proof
+(let ((@x34 (monotonicity (rewrite (= ((_ sign_extend 2) (_ bv10 4)) (_ bv58 6))) (= (= ((_ sign_extend 2) (_ bv10 4)) (_ bv58 6)) (= (_ bv58 6) (_ bv58 6))))))
+(let ((@x38 (trans @x34 (rewrite (= (= (_ bv58 6) (_ bv58 6)) true)) (= (= ((_ sign_extend 2) (_ bv10 4)) (_ bv58 6)) true))))
+(let ((@x41 (monotonicity @x38 (= (not (= ((_ sign_extend 2) (_ bv10 4)) (_ bv58 6))) (not true)))))
+(let ((@x45 (trans @x41 (rewrite (= (not true) false)) (= (not (= ((_ sign_extend 2) (_ bv10 4)) (_ bv58 6))) false))))
+(mp (asserted (not (= ((_ sign_extend 2) (_ bv10 4)) (_ bv58 6)))) @x45 false)))))))
+
+5ca573788c44ee26ee19907e7d9d9ec1635c9a5b 9 0
+unsat
+((set-logic <null>)
+(proof
+(let ((@x51 (monotonicity (rewrite (= (bvshl (_ bv19 8) (_ bv2 8)) (_ bv76 8))) (= (= (bvshl (_ bv19 8) (_ bv2 8)) (_ bv76 8)) (= (_ bv76 8) (_ bv76 8))))))
+(let ((@x55 (trans @x51 (rewrite (= (= (_ bv76 8) (_ bv76 8)) true)) (= (= (bvshl (_ bv19 8) (_ bv2 8)) (_ bv76 8)) true))))
+(let ((@x58 (monotonicity @x55 (= (not (= (bvshl (_ bv19 8) (_ bv2 8)) (_ bv76 8))) (not true)))))
+(let ((@x62 (trans @x58 (rewrite (= (not true) false)) (= (not (= (bvshl (_ bv19 8) (_ bv2 8)) (_ bv76 8))) false))))
+(mp (asserted (not (= (bvshl (_ bv19 8) (_ bv2 8)) (_ bv76 8)))) @x62 false)))))))
+
+b4a2032ff1791888567d8f54fa94d95365cb3255 9 0
+unsat
+((set-logic <null>)
+(proof
+(let ((@x51 (monotonicity (rewrite (= (bvlshr (_ bv25 8) (_ bv2 8)) (_ bv6 8))) (= (= (bvlshr (_ bv25 8) (_ bv2 8)) (_ bv6 8)) (= (_ bv6 8) (_ bv6 8))))))
+(let ((@x55 (trans @x51 (rewrite (= (= (_ bv6 8) (_ bv6 8)) true)) (= (= (bvlshr (_ bv25 8) (_ bv2 8)) (_ bv6 8)) true))))
+(let ((@x58 (monotonicity @x55 (= (not (= (bvlshr (_ bv25 8) (_ bv2 8)) (_ bv6 8))) (not true)))))
+(let ((@x62 (trans @x58 (rewrite (= (not true) false)) (= (not (= (bvlshr (_ bv25 8) (_ bv2 8)) (_ bv6 8))) false))))
+(mp (asserted (not (= (bvlshr (_ bv25 8) (_ bv2 8)) (_ bv6 8)))) @x62 false)))))))
+
+7a4f6966fce99a20413ba9068cef650098d5df66 9 0
+unsat
+((set-logic <null>)
+(proof
+(let ((@x51 (monotonicity (rewrite (= (bvashr (_ bv19 8) (_ bv2 8)) (_ bv4 8))) (= (= (bvashr (_ bv19 8) (_ bv2 8)) (_ bv4 8)) (= (_ bv4 8) (_ bv4 8))))))
+(let ((@x55 (trans @x51 (rewrite (= (= (_ bv4 8) (_ bv4 8)) true)) (= (= (bvashr (_ bv19 8) (_ bv2 8)) (_ bv4 8)) true))))
+(let ((@x58 (monotonicity @x55 (= (not (= (bvashr (_ bv19 8) (_ bv2 8)) (_ bv4 8))) (not true)))))
+(let ((@x62 (trans @x58 (rewrite (= (not true) false)) (= (not (= (bvashr (_ bv19 8) (_ bv2 8)) (_ bv4 8))) false))))
+(mp (asserted (not (= (bvashr (_ bv19 8) (_ bv2 8)) (_ bv4 8)))) @x62 false)))))))
+
+17e3ce1c7a7f4469b5b93e126887f9f5cc55a51d 9 0
+unsat
+((set-logic <null>)
+(proof
+(let ((@x50 (monotonicity (rewrite (= ((_ rotate_right 2) (_ bv6 4)) (_ bv9 4))) (= (= ((_ rotate_right 2) (_ bv6 4)) (_ bv9 4)) (= (_ bv9 4) (_ bv9 4))))))
+(let ((@x54 (trans @x50 (rewrite (= (= (_ bv9 4) (_ bv9 4)) true)) (= (= ((_ rotate_right 2) (_ bv6 4)) (_ bv9 4)) true))))
+(let ((@x57 (monotonicity @x54 (= (not (= ((_ rotate_right 2) (_ bv6 4)) (_ bv9 4))) (not true)))))
+(let ((@x61 (trans @x57 (rewrite (= (not true) false)) (= (not (= ((_ rotate_right 2) (_ bv6 4)) (_ bv9 4))) false))))
+(mp (asserted (not (= ((_ rotate_right 2) (_ bv6 4)) (_ bv9 4)))) @x61 false)))))))
+
+070667ff72c73dc8cd9ebb50cc06803d9785fef4 9 0
+unsat
+((set-logic <null>)
+(proof
+(let ((@x50 (monotonicity (rewrite (= ((_ rotate_left 1) (_ bv14 4)) (_ bv13 4))) (= (= ((_ rotate_left 1) (_ bv14 4)) (_ bv13 4)) (= (_ bv13 4) (_ bv13 4))))))
+(let ((@x54 (trans @x50 (rewrite (= (= (_ bv13 4) (_ bv13 4)) true)) (= (= ((_ rotate_left 1) (_ bv14 4)) (_ bv13 4)) true))))
+(let ((@x57 (monotonicity @x54 (= (not (= ((_ rotate_left 1) (_ bv14 4)) (_ bv13 4))) (not true)))))
+(let ((@x61 (trans @x57 (rewrite (= (not true) false)) (= (not (= ((_ rotate_left 1) (_ bv14 4)) (_ bv13 4))) false))))
+(mp (asserted (not (= ((_ rotate_left 1) (_ bv14 4)) (_ bv13 4)))) @x61 false)))))))
+
+222d5f4b74cc91bf83cca8b1b96f9cfe0f7db0f7 17 0
+unsat
+((set-logic <null>)
+(proof
+(let ((?x9 (bvand |x$| (_ bv255 16))))
+(let ((?x7 (bvand |x$| (_ bv65280 16))))
+(let ((?x10 (bvor ?x7 ?x9)))
+(let (($x11 (= ?x10 |x$|)))
+(let (($x12 (not $x11)))
+(let ((@x57 (symm (commutativity (= (= |x$| ?x10) $x11)) (= $x11 (= |x$| ?x10)))))
+(let ((@x33 (asserted $x12)))
+(let ((@x61 (mp @x33 (monotonicity @x57 (= $x12 (not (= |x$| ?x10)))) (not (= |x$| ?x10)))))
+(let (($x50 (= |x$| ?x10)))
+(let ((@x32 (|true-axiom| true)))
+(let (($x51 (or $x50 false false false false false false false false false false false false false false false false)))
+(let ((@x53 (|unit-resolution| ((_ |th-lemma| bv) $x51) @x32 @x32 @x32 @x32 @x32 @x32 @x32 @x32 @x32 @x32 @x32 @x32 @x32 @x32 @x32 @x32 $x50)))
+(|unit-resolution| @x53 @x61 false)))))))))))))))
+
+f856dea62e897c0065d5a1827265d3ff37ee50c8 51 0
+unsat
+((set-logic <null>)
+(proof
+(let ((?x9 (bvand |w$| (_ bv255 16))))
+(let (($x10 (= ?x9 |w$|)))
+(let (($x62 (not $x10)))
+(let ((@x316 (symm (commutativity (= (= |w$| ?x9) $x10)) (= $x10 (= |w$| ?x9)))))
+(let (($x55 (not (or (bvule (_ bv256 16) |w$|) $x10))))
+(let ((@x47 (monotonicity (rewrite (= (bvult |w$| (_ bv256 16)) (not (bvule (_ bv256 16) |w$|)))) (= (not (bvult |w$| (_ bv256 16))) (not (not (bvule (_ bv256 16) |w$|)))))))
+(let ((@x51 (trans @x47 (rewrite (= (not (not (bvule (_ bv256 16) |w$|))) (bvule (_ bv256 16) |w$|))) (= (not (bvult |w$| (_ bv256 16))) (bvule (_ bv256 16) |w$|)))))
+(let ((@x54 (monotonicity @x51 (= (or (not (bvult |w$| (_ bv256 16))) $x10) (or (bvule (_ bv256 16) |w$|) $x10)))))
+(let ((@x57 (monotonicity @x54 (= (not (or (not (bvult |w$| (_ bv256 16))) $x10)) $x55))))
+(let (($x12 (not (=> (bvult |w$| (_ bv256 16)) $x10))))
+(let ((@x37 (rewrite (= (=> (bvult |w$| (_ bv256 16)) $x10) (or (not (bvult |w$| (_ bv256 16))) $x10)))))
+(let ((@x40 (monotonicity @x37 (= $x12 (not (or (not (bvult |w$| (_ bv256 16))) $x10))))))
+(let ((@x63 (|not-or-elim| (mp (asserted $x12) (trans @x40 @x57 (= $x12 $x55)) $x55) $x62)))
+(let ((@x320 (mp @x63 (monotonicity @x316 (= $x62 (not (= |w$| ?x9)))) (not (= |w$| ?x9)))))
+(let (($x298 (= |w$| ?x9)))
+(let (($x79 (= ((_ extract 15 15) |w$|) (_ bv1 1))))
+(let (($x262 (not $x79)))
+(let (($x72 (= ((_ extract 8 8) |w$|) (_ bv1 1))))
+(let (($x73 (= ((_ extract 9 9) |w$|) (_ bv1 1))))
+(let (($x80 (and $x73 $x72)))
+(let (($x81 (or $x73 $x72 $x80)))
+(let (($x74 (= ((_ extract 10 10) |w$|) (_ bv1 1))))
+(let (($x82 (and $x74 $x81)))
+(let (($x83 (or $x74 $x73 $x72 $x80 $x82)))
+(let (($x75 (= ((_ extract 11 11) |w$|) (_ bv1 1))))
+(let (($x84 (and $x75 $x83)))
+(let (($x85 (or $x75 $x74 $x73 $x72 $x80 $x82 $x84)))
+(let (($x76 (= ((_ extract 12 12) |w$|) (_ bv1 1))))
+(let (($x86 (and $x76 $x85)))
+(let (($x87 (or $x76 $x75 $x74 $x73 $x72 $x80 $x82 $x84 $x86)))
+(let (($x77 (= ((_ extract 13 13) |w$|) (_ bv1 1))))
+(let (($x88 (and $x77 $x87)))
+(let (($x89 (or $x77 $x76 $x75 $x74 $x73 $x72 $x80 $x82 $x84 $x86 $x88)))
+(let (($x78 (= ((_ extract 14 14) |w$|) (_ bv1 1))))
+(let (($x90 (and $x78 $x89)))
+(let (($x91 (or $x78 $x77 $x76 $x75 $x74 $x73 $x72 $x80 $x82 $x84 $x86 $x88 $x90)))
+(let (($x92 (and $x79 $x91)))
+(let (($x93 (or $x79 $x78 $x77 $x76 $x75 $x74 $x73 $x72 $x80 $x82 $x84 $x86 $x88 $x90 $x92)))
+(let (($x295 (not $x93)))
+(let (($x41 (bvule (_ bv256 16) |w$|)))
+(let (($x42 (not $x41)))
+(let ((@x61 (|not-or-elim| (mp (asserted $x12) (trans @x40 @x57 (= $x12 $x55)) $x55) $x42)))
+(let ((@x301 (|unit-resolution| ((_ |th-lemma| bv) (or $x41 $x295)) @x61 $x295)))
+(let ((@x32 (|true-axiom| true)))
+(let (($x310 (or $x298 false false false false false false false false $x72 $x73 $x74 $x75 $x76 $x77 $x78 $x79)))
+(let ((@x312 (|unit-resolution| ((_ |th-lemma| bv) $x310) @x32 @x32 @x32 @x32 @x32 @x32 @x32 @x32 (|unit-resolution| (|def-axiom| (or $x93 (not $x72))) @x301 (not $x72)) (|unit-resolution| (|def-axiom| (or $x93 (not $x73))) @x301 (not $x73)) (|unit-resolution| (|def-axiom| (or $x93 (not $x74))) @x301 (not $x74)) (|unit-resolution| (|def-axiom| (or $x93 (not $x75))) @x301 (not $x75)) (|unit-resolution| (|def-axiom| (or $x93 (not $x76))) @x301 (not $x76)) (|unit-resolution| (|def-axiom| (or $x93 (not $x77))) @x301 (not $x77)) (|unit-resolution| (|def-axiom| (or $x93 (not $x78))) @x301 (not $x78)) (|unit-resolution| (|def-axiom| (or $x93 $x262)) @x301 $x262) $x298)))
+(|unit-resolution| @x312 @x320 false)))))))))))))))))))))))))))))))))))))))))))))))))
+
+39d6b3ac211187a764a365cb2d10eb3330116060 29 0
+unsat
+((set-logic <null>)
+(proof
+(let ((?x6 (|bv2int$| (_ bv0 2))))
+(let (($x181 (<= ?x6 0)))
+(let (($x182 (not $x181)))
+(let (($x173 (forall ((?v0 (_ BitVec 2)) )(!(let ((?x22 (|bv2int$| ?v0)))
+(let (($x59 (<= ?x22 0)))
+(not $x59))) :pattern ( (|bv2int$| ?v0) )))
+))
+(let (($x63 (forall ((?v0 (_ BitVec 2)) )(let ((?x22 (|bv2int$| ?v0)))
+(let (($x59 (<= ?x22 0)))
+(not $x59))))
+))
+(let ((@x175 (refl (= (not (<= (|bv2int$| ?0) 0)) (not (<= (|bv2int$| ?0) 0))))))
+(let ((@x110 (refl (|~| (not (<= (|bv2int$| ?0) 0)) (not (<= (|bv2int$| ?0) 0))))))
+(let (($x24 (forall ((?v0 (_ BitVec 2)) )(let ((?x22 (|bv2int$| ?v0)))
+(< 0 ?x22)))
+))
+(let ((@x62 (rewrite (= (< 0 (|bv2int$| ?0)) (not (<= (|bv2int$| ?0) 0))))))
+(let ((@x113 (|mp~| (mp (asserted $x24) (|quant-intro| @x62 (= $x24 $x63)) $x63) (|nnf-pos| @x110 (|~| $x63 $x63)) $x63)))
+(let ((@x178 (mp @x113 (|quant-intro| @x175 (= $x63 $x173)) $x173)))
+(let (($x185 (not $x173)))
+(let (($x186 (or $x185 $x182)))
+(let ((@x187 ((_ |quant-inst| (_ bv0 2)) $x186)))
+(let (($x8 (= ?x6 0)))
+(let ((@x52 (asserted $x8)))
+(|unit-resolution| ((_ |th-lemma| arith triangle-eq) (or (not $x8) $x181)) @x52 (|unit-resolution| @x187 @x178 $x182) false)))))))))))))))))))
+
+f7db43c56c17d090679f2e9727e9eaa7cf84ab8d 16 0
+unsat
+((set-logic <null>)
+(proof
+(let ((?x10 (|p$| true)))
+(let (($x7 (bvule (_ bv0 4) |a$|)))
+(let (($x8 (ite $x7 true false)))
+(let ((?x9 (|p$| $x8)))
+(let (($x11 (= ?x9 ?x10)))
+(let (($x12 (not $x11)))
+(let ((@x50 (monotonicity (monotonicity (rewrite (= $x7 true)) (= (|p$| $x7) ?x10)) (= (= (|p$| $x7) ?x10) (= ?x10 ?x10)))))
+(let ((@x54 (trans @x50 (rewrite (= (= ?x10 ?x10) true)) (= (= (|p$| $x7) ?x10) true))))
+(let ((@x61 (trans (monotonicity @x54 (= (not (= (|p$| $x7) ?x10)) (not true))) (rewrite (= (not true) false)) (= (not (= (|p$| $x7) ?x10)) false))))
+(let ((@x41 (monotonicity (monotonicity (rewrite (= $x8 $x7)) (= ?x9 (|p$| $x7))) (= $x11 (= (|p$| $x7) ?x10)))))
+(let ((@x44 (monotonicity @x41 (= $x12 (not (= (|p$| $x7) ?x10))))))
+(mp (asserted $x12) (trans @x44 @x61 (= $x12 false)) false))))))))))))))
+
--- a/src/HOL/SMT_Examples/SMT_Word_Examples.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/SMT_Examples/SMT_Word_Examples.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -8,11 +8,10 @@
 imports "~~/src/HOL/Word/Word"
 begin
 
-declare [[smt_oracle = true]]
-declare [[smt_certificates = "SMT_Word_Examples.certs"]]
-declare [[smt_read_only_certificates = true]]
-
-
+declare [[smt2_oracle = true]]
+declare [[z3_new_extensions = true]]
+declare [[smt2_certificates = "SMT_Word_Examples.certs2"]]
+declare [[smt2_read_only_certificates = true]]
 
 text {*
 Currently, there is no proof reconstruction for words.
@@ -20,66 +19,38 @@
 *}
 
 
-
 section {* Bitvector numbers *}
 
-lemma "(27 :: 4 word) = -5" by smt
-
-lemma "(27 :: 4 word) = 11" by smt
-
-lemma "23 < (27::8 word)" by smt
-
-lemma "27 + 11 = (6::5 word)" by smt
-
-lemma "7 * 3 = (21::8 word)" by smt
-
-lemma "11 - 27 = (-16::8 word)" by smt
-
-lemma "- -11 = (11::5 word)" by smt
-
-lemma "-40 + 1 = (-39::7 word)" by smt
-
-lemma "a + 2 * b + c - b = (b + c) + (a :: 32 word)" by smt
-
-lemma "x = (5 :: 4 word) \<Longrightarrow> 4 * x = 4" by smt
-
+lemma "(27 :: 4 word) = -5" by smt2
+lemma "(27 :: 4 word) = 11" by smt2
+lemma "23 < (27::8 word)" by smt2
+lemma "27 + 11 = (6::5 word)" by smt2
+lemma "7 * 3 = (21::8 word)" by smt2
+lemma "11 - 27 = (-16::8 word)" by smt2
+lemma "- -11 = (11::5 word)" by smt2
+lemma "-40 + 1 = (-39::7 word)" by smt2
+lemma "a + 2 * b + c - b = (b + c) + (a :: 32 word)" by smt2
+lemma "x = (5 :: 4 word) \<Longrightarrow> 4 * x = 4" by smt2
 
 
 section {* Bit-level logic *}
 
-lemma "0b110 AND 0b101 = (0b100 :: 32 word)" by smt
-
-lemma "0b110 OR 0b011 = (0b111 :: 8 word)" by smt
-
-lemma "0xF0 XOR 0xFF = (0x0F :: 8 word)" by smt
-
-lemma "NOT (0xF0 :: 16 word) = 0xFF0F" by smt
-
-lemma "word_cat (27::4 word) (27::8 word) = (2843::12 word)" by smt
-
-lemma "word_cat (0b0011::4 word) (0b1111::6word) = (0b0011001111 :: 10 word)"
-  by smt
-
-lemma "slice 1 (0b10110 :: 4 word) = (0b11 :: 2 word)" by smt
-
-lemma "ucast (0b1010 :: 4 word) = (0b1010 :: 10 word)" by smt
-
-lemma "scast (0b1010 :: 4 word) = (0b111010 :: 6 word)" by smt
-
-lemma "0b10011 << 2 = (0b1001100::8 word)" by smt
-
-lemma "0b11001 >> 2 = (0b110::8 word)" by smt
-
-lemma "0b10011 >>> 2 = (0b100::8 word)" by smt
-
-lemma "word_rotr 2 0b0110 = (0b1001::4 word)" by smt
-
-lemma "word_rotl 1 0b1110 = (0b1101::4 word)" by smt
-
-lemma "(x AND 0xff00) OR (x AND 0x00ff) = (x::16 word)" by smt
-
-lemma "w < 256 \<Longrightarrow> (w :: 16 word) AND 0x00FF = w" by smt
-
+lemma "0b110 AND 0b101 = (0b100 :: 32 word)" by smt2
+lemma "0b110 OR 0b011 = (0b111 :: 8 word)" by smt2
+lemma "0xF0 XOR 0xFF = (0x0F :: 8 word)" by smt2
+lemma "NOT (0xF0 :: 16 word) = 0xFF0F" by smt2
+lemma "word_cat (27::4 word) (27::8 word) = (2843::12 word)" by smt2
+lemma "word_cat (0b0011::4 word) (0b1111::6word) = (0b0011001111 :: 10 word)" by smt2
+lemma "slice 1 (0b10110 :: 4 word) = (0b11 :: 2 word)" by smt2
+lemma "ucast (0b1010 :: 4 word) = (0b1010 :: 10 word)" by smt2
+lemma "scast (0b1010 :: 4 word) = (0b111010 :: 6 word)" by smt2
+lemma "0b10011 << 2 = (0b1001100::8 word)" by smt2
+lemma "0b11001 >> 2 = (0b110::8 word)" by smt2
+lemma "0b10011 >>> 2 = (0b100::8 word)" by smt2
+lemma "word_rotr 2 0b0110 = (0b1001::4 word)" by smt2
+lemma "word_rotl 1 0b1110 = (0b1101::4 word)" by smt2
+lemma "(x AND 0xff00) OR (x AND 0x00ff) = (x::16 word)" by smt2
+lemma "w < 256 \<Longrightarrow> (w :: 16 word) AND 0x00FF = w" by smt2
 
 
 section {* Combined integer-bitvector properties *}
@@ -91,10 +62,8 @@
       and "bv2int 3 = 3"
       and "\<forall>x::2 word. bv2int x > 0"
   shows "\<forall>i::int. i < 0 \<longrightarrow> (\<forall>x::2 word. bv2int x > i)"
-  using assms
-  using [[z3_options="AUTO_CONFIG=false"]]
-  by smt
+  using assms by smt2
 
-lemma "P (0 \<le> (a :: 4 word)) = P True" by smt
+lemma "P (0 \<le> (a :: 4 word)) = P True" by smt2
 
 end
--- a/src/HOL/Set.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Set.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -1003,7 +1003,7 @@
 lemma if_image_distrib [simp]:
   "(\<lambda>x. if P x then f x else g x) ` S
     = (f ` (S \<inter> {x. P x})) \<union> (g ` (S \<inter> {x. \<not> P x}))"
-  by (auto simp add: image_def)
+  by auto
 
 lemma image_cong:
   "M = N \<Longrightarrow> (\<And>x. x \<in> N \<Longrightarrow> f x = g x) \<Longrightarrow> f ` M = g ` N"
@@ -1054,7 +1054,7 @@
 
 lemma range_composition: 
   "range (\<lambda>x. f (g x)) = f ` range g"
-  by (subst image_image) simp
+  by auto
 
 
 subsubsection {* Some rules with @{text "if"} *}
--- a/src/HOL/Sledgehammer.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Sledgehammer.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -7,7 +7,7 @@
 header {* Sledgehammer: Isabelle--ATP Linkup *}
 
 theory Sledgehammer
-imports ATP SMT
+imports ATP SMT SMT2
 keywords "sledgehammer" :: diag and "sledgehammer_params" :: thy_decl
 begin
 
@@ -27,6 +27,7 @@
 ML_file "Tools/Sledgehammer/sledgehammer_prover.ML"
 ML_file "Tools/Sledgehammer/sledgehammer_prover_atp.ML"
 ML_file "Tools/Sledgehammer/sledgehammer_prover_smt.ML"
+ML_file "Tools/Sledgehammer/sledgehammer_prover_smt2.ML"
 ML_file "Tools/Sledgehammer/sledgehammer_prover_minimize.ML"
 ML_file "Tools/Sledgehammer/sledgehammer_mepo.ML"
 ML_file "Tools/Sledgehammer/sledgehammer_mash.ML"
--- a/src/HOL/Tools/ATP/atp_proof_reconstruct.ML	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Tools/ATP/atp_proof_reconstruct.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -377,22 +377,12 @@
       union (op =) (filter (fn (_, (_, status)) => status = Non_Rec_Def) facts) facts')
     accum fact_names
 
-val isa_ext = Thm.get_name_hint @{thm ext}
-val isa_short_ext = Long_Name.base_name isa_ext
-
-fun ext_name ctxt =
-  if Thm.eq_thm_prop (@{thm ext},
-       singleton (Attrib.eval_thms ctxt) (Facts.named isa_short_ext, [])) then
-    isa_short_ext
-  else
-    isa_ext
-
 val leo2_extcnf_equal_neg_rule = "extcnf_equal_neg"
 val leo2_unfold_def_rule = "unfold_def"
 
 fun add_fact ctxt fact_names ((_, ss), _, _, rule, deps) =
   (if rule = leo2_extcnf_equal_neg_rule then
-     insert (op =) (ext_name ctxt, (Global, General))
+     insert (op =) (short_thm_name ctxt ext, (Global, General))
    else if rule = leo2_unfold_def_rule then
      (* LEO 1.3.3 does not record definitions properly, leading to missing dependencies in the TSTP
         proof. Remove the next line once this is fixed. *)
@@ -401,7 +391,7 @@
      (fn [] =>
          (* agsyHOL and Satallax don't include definitions in their unsatisfiable cores, so we
             assume the worst and include them all here. *)
-         [(ext_name ctxt, (Global, General))] |> add_non_rec_defs fact_names
+         [(short_thm_name ctxt ext, (Global, General))] |> add_non_rec_defs fact_names
        | facts => facts)
    else
      I)
--- a/src/HOL/Tools/ATP/atp_util.ML	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Tools/ATP/atp_util.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -48,8 +48,8 @@
   val is_legitimate_tptp_def : term -> bool
   val transform_elim_prop : term -> term
   val specialize_type : theory -> (string * typ) -> term -> term
-  val strip_subgoal :
-    thm -> int -> Proof.context -> (string * typ) list * term list * term
+  val strip_subgoal : thm -> int -> Proof.context -> (string * typ) list * term list * term
+  val short_thm_name : Proof.context -> thm -> string
 end;
 
 structure ATP_Util : ATP_UTIL =
@@ -305,7 +305,7 @@
   | s_iff (t1, @{const True}) = t1
   | s_iff (@{const False}, t2) = s_not t2
   | s_iff (t1, @{const False}) = s_not t1
-  | s_iff (t1, t2) = HOLogic.eq_const HOLogic.boolT $ t1 $ t2
+  | s_iff (t1, t2) = if t1 aconv t2 then @{const True} else HOLogic.eq_const HOLogic.boolT $ t1 $ t2
 
 fun close_form t =
   fold (fn ((s, i), T) => fn t' =>
@@ -425,4 +425,13 @@
     val concl_t = t |> Logic.strip_assums_concl |> curry subst_bounds frees
   in (rev params, hyp_ts, concl_t) end
 
+fun short_thm_name ctxt th =
+  let
+    val long = Thm.get_name_hint th
+    val short = Long_Name.base_name long
+  in
+    if Thm.eq_thm_prop (th, singleton (Attrib.eval_thms ctxt) (Facts.named short, [])) then short
+    else long
+  end
+
 end;
--- a/src/HOL/Tools/BNF/bnf_fp_util.ML	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Tools/BNF/bnf_fp_util.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -85,8 +85,7 @@
   val dtor_unfold_transferN: string
   val dtor_unfold_uniqueN: string
   val exhaustN: string
-  val hsetN: string
-  val hset_recN: string
+  val colN: string
   val inductN: string
   val injectN: string
   val isNodeN: string
@@ -307,8 +306,7 @@
 val strong_coinductN = "strong_" ^ coinductN
 val dtor_map_strong_coinductN = dtor_mapN ^ "_" ^ strong_coinductN
 val dtor_strong_coinductN = dtorN ^ "_" ^ strong_coinductN
-val hsetN = "Hset"
-val hset_recN = hsetN ^ "_rec"
+val colN = "col"
 val set_inclN = "set_incl"
 val ctor_set_inclN = ctorN ^ "_" ^ set_inclN
 val dtor_set_inclN = dtorN ^ "_" ^ set_inclN
--- a/src/HOL/Tools/BNF/bnf_gfp.ML	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Tools/BNF/bnf_gfp.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -158,15 +158,10 @@
     val sum_bd = Library.foldr1 (uncurry mk_csum) bds;
     val sum_bdT = fst (dest_relT (fastype_of sum_bd));
 
-    val emptys = map (fn T => HOLogic.mk_set T []) passiveAs;
-    val Zeros = map (fn empty =>
-     HOLogic.mk_tuple (map (fn U => absdummy U empty) activeAs)) emptys;
-    val hrecTs = map fastype_of Zeros;
-
-    val ((((((((((((((((((((((((((((((((((zs, zs'), zs_copy), zs_copy2), z's), (ys, ys')),
+    val ((((((((((((((((((((((((((((((((zs, zs'), zs_copy), zs_copy2), z's), (ys, ys')),
       Bs), Bs_copy), B's), B''s), ss), sum_ss), s's), s''s), fs), fs_copy),
-      self_fs), gs), all_gs), xFs), yFs), yFs_copy), RFs), (Rtuple, Rtuple')), (hrecs, hrecs')),
-      (nat, nat')), Rs), Rs_copy), R's), sRs), (idx, idx')), Idx), Ris), Kss), names_lthy) = lthy
+      self_fs), gs), all_gs), xFs), yFs), yFs_copy), RFs), (Rtuple, Rtuple')),
+      (nat, nat')), Rs), Rs_copy), R's), sRs), (idx, idx')), Idx), Ris), names_lthy) = lthy
       |> mk_Frees' "b" activeAs
       ||>> mk_Frees "b" activeAs
       ||>> mk_Frees "b" activeAs
@@ -190,7 +185,6 @@
       ||>> mk_Frees "y" FTsBs
       ||>> mk_Frees "x" FRTs
       ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "Rtuple") all_sbisT
-      ||>> mk_Frees' "rec" hrecTs
       ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "n") HOLogic.natT
       ||>> mk_Frees "R" setRTs
       ||>> mk_Frees "R" setRTs
@@ -198,8 +192,7 @@
       ||>> mk_Frees "R" setsRTs
       ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "i") idxT
       ||>> yield_singleton (mk_Frees "I") (HOLogic.mk_setT idxT)
-      ||>> mk_Frees "Ri" (map (fn T => idxT --> T) setRTs)
-      ||>> mk_Freess "K" (map (fn AT => map (fn T => T --> AT) activeAs) ATs);
+      ||>> mk_Frees "Ri" (map (fn T => idxT --> T) setRTs);
 
     val passive_UNIVs = map HOLogic.mk_UNIV passiveAs;
     val passive_eqs = map HOLogic.eq_const passiveAs;
@@ -355,7 +348,7 @@
           (HOLogic.mk_Trueprop (mk_tcoalg activeAs ss))
       in
         Goal.prove_sorry lthy [] [] goal
-          (K (stac coalg_def 1 THEN CONJ_WRAP
+          (K (rtac (coalg_def RS iffD2) 1 THEN CONJ_WRAP
             (K (EVERY' [rtac ballI, rtac CollectI,
               CONJ_WRAP' (K (EVERY' [rtac @{thm subset_UNIV}])) allAs] 1)) ss))
         |> Thm.close_derivation
@@ -506,203 +499,6 @@
 
     val timer = time (timer "Morphism definition & thms");
 
-    fun hset_rec_bind j = mk_internal_b (hset_recN ^ (if m = 1 then "" else string_of_int j));
-    val hset_rec_def_bind = rpair [] o Thm.def_binding o hset_rec_bind;
-
-    fun hset_rec_spec j Zero hrec hrec' =
-      let
-        fun mk_Suc s setsAs z z' =
-          let
-            val (set, sets) = apfst (fn xs => nth xs (j - 1)) (chop m setsAs);
-            fun mk_UN set k = mk_UNION (set $ (s $ z)) (mk_nthN n hrec k);
-          in
-            Term.absfree z'
-              (mk_union (set $ (s $ z), Library.foldl1 mk_union (map2 mk_UN sets ks)))
-          end;
-
-        val Suc = Term.absdummy HOLogic.natT (Term.absfree hrec'
-          (HOLogic.mk_tuple (map4 mk_Suc ss setssAs zs zs')));
-
-        val rhs = mk_rec_nat Zero Suc;
-      in
-        fold_rev (Term.absfree o Term.dest_Free) ss rhs
-      end;
-
-    val ((hset_rec_frees, (_, hset_rec_def_frees)), (lthy, lthy_old)) =
-      lthy
-      |> fold_map4 (fn j => fn Zero => fn hrec => fn hrec' => Local_Theory.define
-        ((hset_rec_bind j, NoSyn), (hset_rec_def_bind j, hset_rec_spec j Zero hrec hrec')))
-        ls Zeros hrecs hrecs'
-      |>> apsnd split_list o split_list
-      ||> `Local_Theory.restore;
-
-    val phi = Proof_Context.export_morphism lthy_old lthy;
-
-    val hset_rec_defs = map (fn def =>
-      mk_unabs_def n (Morphism.thm phi def RS meta_eq_to_obj_eq)) hset_rec_def_frees;
-    val hset_recs = map (fst o Term.dest_Const o Morphism.term phi) hset_rec_frees;
-
-    fun mk_hset_rec ss nat i j T =
-      let
-        val args = ss @ [nat];
-        val Ts = map fastype_of ss;
-        val bTs = map domain_type Ts;
-        val hrecT = HOLogic.mk_tupleT (map (fn U => U --> HOLogic.mk_setT T) bTs)
-        val hset_recT = Library.foldr (op -->) (Ts, HOLogic.natT --> hrecT);
-      in
-        mk_nthN n (Term.list_comb (Const (nth hset_recs (j - 1), hset_recT), args)) i
-      end;
-
-    val hset_rec_0ss = mk_rec_simps n @{thm rec_nat_0_imp} hset_rec_defs;
-    val hset_rec_Sucss = mk_rec_simps n @{thm rec_nat_Suc_imp} hset_rec_defs;
-    val hset_rec_0ss' = transpose hset_rec_0ss;
-    val hset_rec_Sucss' = transpose hset_rec_Sucss;
-
-    fun hset_binds j = mk_internal_bs (hsetN ^ (if m = 1 then "" else string_of_int j))
-    fun hset_bind i j = nth (hset_binds j) (i - 1);
-    val hset_def_bind = rpair [] o Thm.def_binding oo hset_bind;
-
-    fun hset_spec i j =
-      let
-        val z = nth zs (i - 1);
-        val T = nth passiveAs (j - 1);
-
-        val rhs = mk_UNION (HOLogic.mk_UNIV HOLogic.natT)
-          (Term.absfree nat' (mk_hset_rec ss nat i j T $ z));
-      in
-        fold_rev (Term.absfree o Term.dest_Free) (ss @ [z]) rhs
-      end;
-
-    val ((hset_frees, (_, hset_def_frees)), (lthy, lthy_old)) =
-      lthy
-      |> fold_map (fn i => fold_map (fn j => Local_Theory.define
-        ((hset_bind i j, NoSyn), (hset_def_bind i j, hset_spec i j))) ls) ks
-      |>> map (apsnd split_list o split_list)
-      |>> apsnd split_list o split_list
-      ||> `Local_Theory.restore;
-
-    val phi = Proof_Context.export_morphism lthy_old lthy;
-
-    val hset_defss = map (map (fn def =>
-      mk_unabs_def (n + 1) (Morphism.thm phi def RS meta_eq_to_obj_eq))) hset_def_frees;
-    val hset_defss' = transpose hset_defss;
-    val hset_namess = map (map (fst o Term.dest_Const o Morphism.term phi)) hset_frees;
-
-    fun mk_hset ss i j T =
-      let
-        val Ts = map fastype_of ss;
-        val bTs = map domain_type Ts;
-        val hsetT = Library.foldr (op -->) (Ts, nth bTs (i - 1) --> HOLogic.mk_setT T);
-      in
-        Term.list_comb (Const (nth (nth hset_namess (i - 1)) (j - 1), hsetT), ss)
-      end;
-
-    val hsetssAs = map (fn i => map2 (mk_hset ss i) ls passiveAs) ks;
-
-    val (set_incl_hset_thmss, set_hset_incl_hset_thmsss) =
-      let
-        fun mk_set_incl_hset s x set hset = fold_rev Logic.all (x :: ss)
-          (HOLogic.mk_Trueprop (mk_leq (set $ (s $ x)) (hset $ x)));
-
-        fun mk_set_hset_incl_hset s x y set hset1 hset2 =
-          fold_rev Logic.all (x :: y :: ss)
-            (Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_mem (x, set $ (s $ y))),
-            HOLogic.mk_Trueprop (mk_leq (hset1 $ x) (hset2 $ y))));
-
-        val set_incl_hset_goalss =
-          map4 (fn s => fn x => fn sets => fn hsets =>
-            map2 (mk_set_incl_hset s x) (take m sets) hsets)
-          ss zs setssAs hsetssAs;
-
-        (*xk : F(i)set(m+k) (si yi) ==> F(k)_hset(j) s1 ... sn xk <= F(i)_hset(j) s1 ... sn yi*)
-        val set_hset_incl_hset_goalsss =
-          map4 (fn si => fn yi => fn sets => fn hsetsi =>
-            map3 (fn xk => fn set => fn hsetsk =>
-              map2 (mk_set_hset_incl_hset si xk yi set) hsetsk hsetsi)
-            zs_copy (drop m sets) hsetssAs)
-          ss zs setssAs hsetssAs;
-      in
-        (map3 (fn goals => fn defs => fn rec_Sucs =>
-          map3 (fn goal => fn def => fn rec_Suc =>
-            Goal.prove_sorry lthy [] [] goal (K (mk_set_incl_hset_tac def rec_Suc))
-            |> Thm.close_derivation)
-          goals defs rec_Sucs)
-        set_incl_hset_goalss hset_defss hset_rec_Sucss,
-        map3 (fn goalss => fn defsi => fn rec_Sucs =>
-          map3 (fn k => fn goals => fn defsk =>
-            map4 (fn goal => fn defk => fn defi => fn rec_Suc =>
-              Goal.prove_sorry lthy [] [] goal
-                (K (mk_set_hset_incl_hset_tac n [defk, defi] rec_Suc k))
-              |> Thm.close_derivation)
-            goals defsk defsi rec_Sucs)
-          ks goalss hset_defss)
-        set_hset_incl_hset_goalsss hset_defss hset_rec_Sucss)
-      end;
-
-    val set_incl_hset_thmss' = transpose set_incl_hset_thmss;
-    val set_hset_incl_hset_thmsss' = transpose (map transpose set_hset_incl_hset_thmsss);
-    val set_hset_thmss = map (map (fn thm => thm RS @{thm set_mp})) set_incl_hset_thmss;
-    val set_hset_hset_thmsss = map (map (map (fn thm => thm RS @{thm set_mp})))
-      set_hset_incl_hset_thmsss;
-    val set_hset_thmss' = transpose set_hset_thmss;
-    val set_hset_hset_thmsss' = transpose (map transpose set_hset_hset_thmsss);
-
-    val hset_minimal_thms =
-      let
-        fun mk_passive_prem set s x K =
-          Logic.all x (HOLogic.mk_Trueprop (mk_leq (set $ (s $ x)) (K $ x)));
-
-        fun mk_active_prem s x1 K1 set x2 K2 =
-          fold_rev Logic.all [x1, x2]
-            (Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_mem (x2, set $ (s $ x1))),
-              HOLogic.mk_Trueprop (mk_leq (K2 $ x2) (K1 $ x1))));
-
-        val premss = map2 (fn j => fn Ks =>
-          map4 mk_passive_prem (map (fn xs => nth xs (j - 1)) setssAs) ss zs Ks @
-            flat (map4 (fn sets => fn s => fn x1 => fn K1 =>
-              map3 (mk_active_prem s x1 K1) (drop m sets) zs_copy Ks) setssAs ss zs Ks))
-          ls Kss;
-
-        val hset_rec_minimal_thms =
-          let
-            fun mk_conjunct j T i K x = mk_leq (mk_hset_rec ss nat i j T $ x) (K $ x);
-            fun mk_concl j T Ks = list_all_free zs
-              (Library.foldr1 HOLogic.mk_conj (map3 (mk_conjunct j T) ks Ks zs));
-            val concls = map3 mk_concl ls passiveAs Kss;
-
-            val goals = map2 (fn prems => fn concl =>
-              Logic.list_implies (prems, HOLogic.mk_Trueprop concl)) premss concls
-
-            val ctss =
-              map (fn phi => map (SOME o certify lthy) [Term.absfree nat' phi, nat]) concls;
-          in
-            map4 (fn goal => fn cts => fn hset_rec_0s => fn hset_rec_Sucs =>
-              singleton (Proof_Context.export names_lthy lthy)
-                (Goal.prove_sorry lthy [] [] goal
-                  (fn {context = ctxt, prems = _} => mk_hset_rec_minimal_tac ctxt m cts hset_rec_0s
-                    hset_rec_Sucs))
-              |> Thm.close_derivation)
-            goals ctss hset_rec_0ss' hset_rec_Sucss'
-          end;
-
-        fun mk_conjunct j T i K x = mk_leq (mk_hset ss i j T $ x) (K $ x);
-        fun mk_concl j T Ks = Library.foldr1 HOLogic.mk_conj (map3 (mk_conjunct j T) ks Ks zs);
-        val concls = map3 mk_concl ls passiveAs Kss;
-
-        val goals = map3 (fn Ks => fn prems => fn concl =>
-          fold_rev Logic.all (Ks @ ss @ zs)
-            (Logic.list_implies (prems, HOLogic.mk_Trueprop concl))) Kss premss concls;
-      in
-        map3 (fn goal => fn hset_defs => fn hset_rec_minimal =>
-          Goal.prove_sorry lthy [] [] goal
-            (fn {context = ctxt, prems = _} => mk_hset_minimal_tac ctxt n hset_defs
-              hset_rec_minimal)
-          |> Thm.close_derivation)
-        goals hset_defss' hset_rec_minimal_thms
-      end;
-
-    val timer = time (timer "Hereditary sets");
-
     (* bisimulation *)
 
     val bis_bind = mk_internal_b bisN;
@@ -1500,8 +1296,13 @@
     val corec_Inls = map (Term.subst_atomic_types (activeBs ~~ Ts)) Inls;
     val corec_UNIVs = map2 (HOLogic.mk_UNIV oo curry mk_sumT) Ts activeAs;
 
-    val ((((((((((((Jzs, Jzs'), Jz's), Jzs_copy), Jz's_copy), Jzs1), Jzs2),
-      TRs), unfold_fs), corec_ss), phis), dtor_set_induct_phiss),
+    val emptys = map (fn T => HOLogic.mk_set T []) passiveAs;
+    val Zeros = map (fn empty =>
+     HOLogic.mk_tuple (map (fn U => absdummy U empty) Ts)) emptys;
+    val hrecTs = map fastype_of Zeros;
+
+    val (((((((((((((Jzs, Jzs'), Jz's), Jzs_copy), Jz's_copy), Jzs1), Jzs2),
+      TRs), unfold_fs), corec_ss), phis), dtor_set_induct_phiss), (hrecs, hrecs')),
       names_lthy) = names_lthy
       |> mk_Frees' "z" Ts
       ||>> mk_Frees "y" Ts'
@@ -1513,7 +1314,8 @@
       ||>> mk_Frees "f" unfold_fTs
       ||>> mk_Frees "s" corec_sTs
       ||>> mk_Frees "P" (map2 mk_pred2T Ts Ts)
-      ||>> mk_Freess "P" (map (fn A => map (mk_pred2T A) Ts) passiveAs);
+      ||>> mk_Freess "P" (map (fn A => map (mk_pred2T A) Ts) passiveAs)
+      ||>> mk_Frees' "rec" hrecTs;
 
     fun dtor_bind i = nth external_bs (i - 1) |> Binding.prefix_name (dtorN ^ "_");
     val dtor_def_bind = rpair [] o Binding.conceal o Thm.def_binding o dtor_bind;
@@ -1831,59 +1633,6 @@
 
     val timer = time (timer "coinduction");
 
-    val setss_by_bnf = map (fn i => map2 (mk_hset dtors i) ls passiveAs) ks;
-    val setss_by_range = transpose setss_by_bnf;
-
-    val (hset_dtor_incl_thmss, hset_hset_dtor_incl_thmsss, dtor_hset_induct_thms) =
-      let
-        fun tinst_of dtor =
-          map (SOME o certify lthy) (dtor :: remove (op =) dtor dtors);
-        fun tinst_of' dtor = case tinst_of dtor of t :: ts => t :: NONE :: ts;
-        val Tinst = map (pairself (certifyT lthy))
-          (map Logic.varifyT_global (deads @ allAs) ~~ (deads @ passiveAs @ Ts));
-        val set_incl_thmss =
-          map2 (fn dtor => map (singleton (Proof_Context.export names_lthy lthy) o
-            Drule.instantiate' [] (tinst_of' dtor) o
-            Thm.instantiate (Tinst, []) o Drule.zero_var_indexes))
-          dtors set_incl_hset_thmss;
-
-        val tinst = splice (map (SOME o certify lthy) dtors) (replicate n NONE)
-        val set_minimal_thms =
-          map (Drule.instantiate' [] tinst o Thm.instantiate (Tinst, []) o
-            Drule.zero_var_indexes)
-          hset_minimal_thms;
-
-        val set_set_incl_thmsss =
-          map2 (fn dtor => map (map (singleton (Proof_Context.export names_lthy lthy) o
-            Drule.instantiate' [] (NONE :: tinst_of' dtor) o
-            Thm.instantiate (Tinst, []) o Drule.zero_var_indexes)))
-          dtors set_hset_incl_hset_thmsss;
-
-        val set_set_incl_thmsss' = transpose (map transpose set_set_incl_thmsss);
-
-        val incls =
-          maps (map (fn thm => thm RS @{thm subset_Collect_iff})) set_incl_thmss @
-            @{thms subset_Collect_iff[OF subset_refl]};
-
-        fun mk_induct_tinst phis jsets y y' =
-          map4 (fn phi => fn jset => fn Jz => fn Jz' =>
-            SOME (certify lthy (Term.absfree Jz' (HOLogic.mk_Collect (fst y', snd y',
-              HOLogic.mk_conj (HOLogic.mk_mem (y, jset $ Jz), phi $ y $ Jz))))))
-          phis jsets Jzs Jzs';
-        val dtor_set_induct_thms =
-          map6 (fn set_minimal => fn set_set_inclss => fn jsets => fn y => fn y' => fn phis =>
-            ((set_minimal
-              |> Drule.instantiate' [] (mk_induct_tinst phis jsets y y')
-              |> unfold_thms lthy incls) OF
-              (replicate n ballI @
-                maps (map (fn thm => thm RS @{thm subset_CollectI})) set_set_inclss))
-            |> singleton (Proof_Context.export names_lthy lthy)
-            |> rule_by_tactic lthy (ALLGOALS (TRY o etac asm_rl)))
-          set_minimal_thms set_set_incl_thmsss' setss_by_range ys ys' dtor_set_induct_phiss
-      in
-        (set_incl_thmss, set_set_incl_thmsss, dtor_set_induct_thms)
-      end;
-
     fun mk_dtor_map_DEADID_thm dtor_inject map_id0 =
       trans OF [iffD2 OF [dtor_inject, id_apply], map_id0 RS sym];
 
@@ -1925,15 +1674,16 @@
         val gTs = map2 (curry op -->) passiveBs passiveCs;
         val uTs = map2 (curry op -->) Ts Ts';
 
-        val ((((((((fs, fs'), fs_copy), gs), us), (Jys, Jys')), (Jys_copy, Jys'_copy)),
-          (ys_copy, ys'_copy)), names_lthy) = names_lthy
+        val (((((((((fs, fs'), fs_copy), gs), us), (Jys, Jys')), (Jys_copy, Jys'_copy)),
+          (ys_copy, ys'_copy)), Kss), names_lthy) = names_lthy
           |> mk_Frees' "f" fTs
           ||>> mk_Frees "f" fTs
           ||>> mk_Frees "g" gTs
           ||>> mk_Frees "u" uTs
           ||>> mk_Frees' "b" Ts'
           ||>> mk_Frees' "b" Ts'
-          ||>> mk_Frees' "y" passiveAs;
+          ||>> mk_Frees' "y" passiveAs
+          ||>> mk_Freess "K" (map (fn AT => map (fn T => T --> AT) Ts) ATs);;
 
         val map_FTFT's = map2 (fn Ds =>
           mk_map_of_bnf Ds (passiveAs @ Ts) (passiveBs @ Ts')) Dss bnfs;
@@ -1951,6 +1701,187 @@
         val set_bss =
           map (flat o map2 (fn B => fn b =>
             if member (op =) resDs (TFree B) then [] else [b]) resBs) set_bss0;
+    
+        fun col_bind j = mk_internal_b (colN ^ (if m = 1 then "" else string_of_int j));
+        val col_def_bind = rpair [] o Thm.def_binding o col_bind;
+    
+        fun col_spec j Zero hrec hrec' =
+          let
+            fun mk_Suc dtor setss z z' =
+              let
+                val (set, sets) = apfst (fn xs => nth xs (j - 1)) (chop m setss);
+                fun mk_UN set k = mk_UNION (set $ (dtor $ z)) (mk_nthN n hrec k);
+              in
+                Term.absfree z'
+                  (mk_union (set $ (dtor $ z), Library.foldl1 mk_union (map2 mk_UN sets ks)))
+              end;
+    
+            val Suc = Term.absdummy HOLogic.natT (Term.absfree hrec'
+              (HOLogic.mk_tuple (map4 mk_Suc dtors FTs_setss Jzs Jzs')));
+          in
+            mk_rec_nat Zero Suc
+          end;
+    
+        val ((col_frees, (_, col_def_frees)), (lthy, lthy_old)) =
+          lthy
+          |> fold_map4 (fn j => fn Zero => fn hrec => fn hrec' => Local_Theory.define
+            ((col_bind j, NoSyn), (col_def_bind j, col_spec j Zero hrec hrec')))
+            ls Zeros hrecs hrecs'
+          |>> apsnd split_list o split_list
+          ||> `Local_Theory.restore;
+    
+        val phi = Proof_Context.export_morphism lthy_old lthy;
+    
+        val col_defs = map (fn def => Morphism.thm phi def RS meta_eq_to_obj_eq) col_def_frees;
+        val cols = map (fst o Term.dest_Const o Morphism.term phi) col_frees;
+    
+        fun mk_col Ts nat i j T =
+          let
+            val hrecT = HOLogic.mk_tupleT (map (fn U => U --> HOLogic.mk_setT T) Ts)
+            val colT = HOLogic.natT --> hrecT;
+          in
+            mk_nthN n (Term.list_comb (Const (nth cols (j - 1), colT), [nat])) i
+          end;
+    
+        val col_0ss = mk_rec_simps n @{thm rec_nat_0_imp} col_defs;
+        val col_Sucss = mk_rec_simps n @{thm rec_nat_Suc_imp} col_defs;
+        val col_0ss' = transpose col_0ss;
+        val col_Sucss' = transpose col_Sucss;
+    
+        fun mk_hset Ts i j T =
+          Abs (Name.uu, nth Ts (i - 1), mk_UNION (HOLogic.mk_UNIV HOLogic.natT)
+            (Term.absfree nat' (mk_col Ts nat i j T $ Bound 1)));
+
+        val setss_by_bnf = map (fn i => map2 (mk_hset Ts i) ls passiveAs) ks;
+        val setss_by_range = transpose setss_by_bnf;
+
+        val hset_minimal_thms =
+          let
+            fun mk_passive_prem set dtor x K =
+              Logic.all x (HOLogic.mk_Trueprop (mk_leq (set $ (dtor $ x)) (K $ x)));
+    
+            fun mk_active_prem dtor x1 K1 set x2 K2 =
+              fold_rev Logic.all [x1, x2]
+                (Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_mem (x2, set $ (dtor $ x1))),
+                  HOLogic.mk_Trueprop (mk_leq (K2 $ x2) (K1 $ x1))));
+    
+            val premss = map2 (fn j => fn Ks =>
+              map4 mk_passive_prem (map (fn xs => nth xs (j - 1)) FTs_setss) dtors Jzs Ks @
+                flat (map4 (fn sets => fn s => fn x1 => fn K1 =>
+                  map3 (mk_active_prem s x1 K1) (drop m sets) Jzs_copy Ks) FTs_setss dtors Jzs Ks))
+              ls Kss;
+    
+            val col_minimal_thms =
+              let
+                fun mk_conjunct j T i K x = mk_leq (mk_col Ts nat i j T $ x) (K $ x);
+                fun mk_concl j T Ks = list_all_free Jzs
+                  (Library.foldr1 HOLogic.mk_conj (map3 (mk_conjunct j T) ks Ks Jzs));
+                val concls = map3 mk_concl ls passiveAs Kss;
+    
+                val goals = map2 (fn prems => fn concl =>
+                  Logic.list_implies (prems, HOLogic.mk_Trueprop concl)) premss concls
+    
+                val ctss =
+                  map (fn phi => map (SOME o certify lthy) [Term.absfree nat' phi, nat]) concls;
+              in
+                map4 (fn goal => fn cts => fn col_0s => fn col_Sucs =>
+                  singleton (Proof_Context.export names_lthy lthy)
+                    (Goal.prove_sorry lthy [] [] goal
+                      (fn {context = ctxt, prems = _} => mk_col_minimal_tac ctxt m cts col_0s
+                        col_Sucs))
+                  |> Thm.close_derivation)
+                goals ctss col_0ss' col_Sucss'
+              end;
+    
+            fun mk_conjunct j T i K x = mk_leq (mk_hset Ts i j T $ x) (K $ x);
+            fun mk_concl j T Ks = Library.foldr1 HOLogic.mk_conj (map3 (mk_conjunct j T) ks Ks Jzs);
+            val concls = map3 mk_concl ls passiveAs Kss;
+    
+            val goals = map3 (fn Ks => fn prems => fn concl =>
+              fold_rev Logic.all (Ks @ Jzs)
+                (Logic.list_implies (prems, HOLogic.mk_Trueprop concl))) Kss premss concls;
+          in
+            map2 (fn goal => fn col_minimal =>
+              Goal.prove_sorry lthy [] [] goal
+                (fn {context = ctxt, prems = _} => mk_hset_minimal_tac ctxt n col_minimal)
+              |> Thm.close_derivation)
+            goals col_minimal_thms
+          end;
+
+        val (set_incl_hset_thmss, set_hset_incl_hset_thmsss) =
+          let
+            fun mk_set_incl_hset dtor x set hset = fold_rev Logic.all (x :: ss)
+              (HOLogic.mk_Trueprop (mk_leq (set $ (dtor $ x)) (hset $ x)));
+    
+            fun mk_set_hset_incl_hset dtor x y set hset1 hset2 =
+              fold_rev Logic.all [x, y]
+                (Logic.mk_implies (HOLogic.mk_Trueprop (HOLogic.mk_mem (x, set $ (dtor $ y))),
+                HOLogic.mk_Trueprop (mk_leq (hset1 $ x) (hset2 $ y))));
+    
+            val set_incl_hset_goalss =
+              map4 (fn dtor => fn x => fn sets => fn hsets =>
+                map2 (mk_set_incl_hset dtor x) (take m sets) hsets)
+              dtors Jzs FTs_setss setss_by_bnf;
+    
+            (*xk : F(i)set(m+k) (si yi) ==> F(k)_hset(j) s1 ... sn xk <= F(i)_hset(j) s1 ... sn yi*)
+            val set_hset_incl_hset_goalsss =
+              map4 (fn dtori => fn yi => fn sets => fn hsetsi =>
+                map3 (fn xk => fn set => fn hsetsk =>
+                  map2 (mk_set_hset_incl_hset dtori xk yi set) hsetsk hsetsi)
+                Jzs_copy (drop m sets) setss_by_bnf)
+              dtors Jzs FTs_setss setss_by_bnf;
+          in
+            (map2 (fn goals => fn rec_Sucs =>
+              map2 (fn goal => fn rec_Suc =>
+                Goal.prove_sorry lthy [] [] goal (K (mk_set_incl_hset_tac rec_Suc))
+                |> Thm.close_derivation)
+              goals rec_Sucs)
+            set_incl_hset_goalss col_Sucss,
+            map2 (fn goalss => fn rec_Sucs =>
+              map2 (fn k => fn goals =>
+                map2 (fn goal => fn rec_Suc =>
+                  Goal.prove_sorry lthy [] [] goal
+                    (K (mk_set_hset_incl_hset_tac n rec_Suc k))
+                  |> Thm.close_derivation)
+                goals rec_Sucs)
+              ks goalss)
+            set_hset_incl_hset_goalsss col_Sucss)
+          end;
+    
+        val set_incl_hset_thmss' = transpose set_incl_hset_thmss;
+        val set_hset_incl_hset_thmsss' = transpose (map transpose set_hset_incl_hset_thmsss);
+        val set_hset_thmss = map (map (fn thm => thm RS @{thm set_mp})) set_incl_hset_thmss;
+        val set_hset_hset_thmsss = map (map (map (fn thm => thm RS @{thm set_mp})))
+          set_hset_incl_hset_thmsss;
+        val set_hset_thmss' = transpose set_hset_thmss;
+        val set_hset_hset_thmsss' = transpose (map transpose set_hset_hset_thmsss);
+
+
+        val timer = time (timer "Hereditary sets");
+    
+        val dtor_hset_induct_thms =
+          let
+            val incls =
+              maps (map (fn thm => thm RS @{thm subset_Collect_iff})) set_incl_hset_thmss @
+                @{thms subset_Collect_iff[OF subset_refl]};
+
+            val cTs = map (SOME o certifyT lthy) params';    
+            fun mk_induct_tinst phis jsets y y' =
+              map4 (fn phi => fn jset => fn Jz => fn Jz' =>
+                SOME (certify lthy (Term.absfree Jz' (HOLogic.mk_Collect (fst y', snd y',
+                  HOLogic.mk_conj (HOLogic.mk_mem (y, jset $ Jz), phi $ y $ Jz))))))
+              phis jsets Jzs Jzs';
+          in
+            map6 (fn set_minimal => fn set_set_inclss => fn jsets => fn y => fn y' => fn phis =>
+              ((set_minimal
+                |> Drule.instantiate' cTs (mk_induct_tinst phis jsets y y')
+                |> unfold_thms lthy incls) OF
+                (replicate n ballI @
+                  maps (map (fn thm => thm RS @{thm subset_CollectI})) set_set_inclss))
+              |> singleton (Proof_Context.export names_lthy lthy)
+              |> rule_by_tactic lthy (ALLGOALS (TRY o etac asm_rl)))
+            hset_minimal_thms set_hset_incl_hset_thmsss' setss_by_range ys ys' dtor_set_induct_phiss
+          end;
 
         fun close_wit I wit = (I, fold_rev Term.absfree (map (nth ys') I) wit);
 
@@ -2209,9 +2140,9 @@
         val timer = time (timer "set functions for the new codatatypes");
 
         val colss = map2 (fn j => fn T =>
-          map (fn i => mk_hset_rec dtors nat i j T) ks) ls passiveAs;
+          map (fn i => mk_col Ts nat i j T) ks) ls passiveAs;
         val colss' = map2 (fn j => fn T =>
-          map (fn i => mk_hset_rec dtor's nat i j T) ks) ls passiveBs;
+          map (fn i => mk_col Ts' nat i j T) ks) ls passiveBs;
 
         val col_natural_thmss =
           let
@@ -2233,7 +2164,7 @@
                     (fn {context = ctxt, prems = _} => mk_col_natural_tac ctxt cts rec_0s rec_Sucs
                       dtor_Jmap_thms set_mapss))
                 |> Thm.close_derivation)
-              goals ctss hset_rec_0ss' hset_rec_Sucss';
+              goals ctss col_0ss' col_Sucss';
           in
             map (split_conj_thm o mk_specN n) thms
           end;
@@ -2257,7 +2188,7 @@
                     (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt Jbd_defs THEN
                       mk_col_bd_tac m j cts rec_0s rec_Sucs sbd_Card_order sbd_Cinfinite set_sbdss))
                 |> Thm.close_derivation)
-              ls goals ctss hset_rec_0ss' hset_rec_Sucss';
+              ls goals ctss col_0ss' col_Sucss';
           in
             map (split_conj_thm o mk_specN n) thms
           end;
@@ -2312,8 +2243,8 @@
             Jrel_unabs_defs;
 
         val fold_Jsets = fold_thms lthy Jset_unabs_defs;
-        val dtor_Jset_incl_thmss = map (map fold_Jsets) hset_dtor_incl_thmss;
-        val dtor_set_Jset_incl_thmsss = map (map (map fold_Jsets)) hset_hset_dtor_incl_thmsss;
+        val dtor_Jset_incl_thmss = map (map fold_Jsets) set_incl_hset_thmss;
+        val dtor_set_Jset_incl_thmsss = map (map (map fold_Jsets)) set_hset_incl_hset_thmsss;
         val dtor_Jset_induct_thms = map fold_Jsets dtor_hset_induct_thms;
         val Jwit_thmss = map (map fold_Jsets) wit_thmss;
 
@@ -2487,9 +2418,8 @@
         val map_comp0_tacs = map (fn thm => K (rtac (thm RS sym) 1)) Jmap_comp0_thms;
         val map_cong0_tacs = map (fn thm => fn ctxt => mk_map_cong0_tac ctxt m thm) map_cong0_thms;
         val set_map0_tacss =
-          map2 (map2 (fn def => fn col => fn ctxt =>
-              unfold_thms_tac ctxt Jset_defs THEN mk_set_map0_tac def col))
-            hset_defss (transpose col_natural_thmss);
+          map (map (fn col => fn ctxt => unfold_thms_tac ctxt Jset_defs THEN mk_set_map0_tac col))
+            (transpose col_natural_thmss);
 
         val Jbd_card_orders = map (fn def => fold_thms lthy [def] sbd_card_order) Jbd_defs;
         val Jbd_Cinfinites = map (fn def => fold_thms lthy [def] sbd_Cinfinite) Jbd_defs;
@@ -2498,9 +2428,9 @@
         val bd_cinf_tacs = map (fn thm => K (rtac (thm RS conjunct1) 1)) Jbd_Cinfinites;
 
         val set_bd_tacss =
-          map3 (fn Cinf => map2 (fn def => fn col => fn ctxt =>
-            unfold_thms_tac ctxt Jset_defs THEN mk_set_bd_tac Cinf def col))
-          Jbd_Cinfinites hset_defss (transpose col_bd_thmss);
+          map2 (fn Cinf => map (fn col => fn ctxt =>
+            unfold_thms_tac ctxt Jset_defs THEN mk_set_bd_tac Cinf col))
+          Jbd_Cinfinites (transpose col_bd_thmss);
 
         val le_rel_OO_tacs = map (fn i => K (rtac (le_Jrel_OO_thm RS mk_conjunctN n i) 1)) ks;
 
--- a/src/HOL/Tools/BNF/bnf_gfp_tactics.ML	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Tools/BNF/bnf_gfp_tactics.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -32,8 +32,8 @@
     thm -> thm -> thm list -> thm list -> thm list list -> tactic
   val mk_dtor_o_ctor_tac: Proof.context -> thm -> thm -> thm -> thm -> thm list -> tactic
   val mk_equiv_lsbis_tac: thm -> thm -> thm -> thm -> thm -> thm -> tactic
-  val mk_hset_minimal_tac: Proof.context -> int -> thm list -> thm -> tactic
-  val mk_hset_rec_minimal_tac: Proof.context -> int -> cterm option list -> thm list -> thm list ->
+  val mk_hset_minimal_tac: Proof.context -> int -> thm -> tactic
+  val mk_col_minimal_tac: Proof.context -> int -> cterm option list -> thm list -> thm list ->
     tactic
   val mk_incl_lsbis_tac: int -> int -> thm -> tactic
   val mk_length_Lev'_tac: thm -> tactic
@@ -56,7 +56,7 @@
   val mk_mor_case_sum_tac: 'a list -> thm -> tactic
   val mk_mor_comp_tac: thm -> thm list -> thm list -> thm list -> tactic
   val mk_mor_elim_tac: thm -> tactic
-  val mk_mor_hset_rec_tac: int -> int -> cterm option list -> int -> thm list -> thm list ->
+  val mk_mor_col_tac: int -> int -> cterm option list -> int -> thm list -> thm list ->
     thm list -> thm list list -> thm list list -> tactic
   val mk_mor_incl_tac: thm -> thm list -> tactic
   val mk_mor_str_tac: 'a list -> thm -> tactic
@@ -74,14 +74,14 @@
   val mk_sbis_lsbis_tac: Proof.context -> thm list -> thm -> thm -> tactic
   val mk_set_Lev_tac: Proof.context -> cterm option list -> thm list -> thm list -> thm list ->
     thm list -> thm list list -> tactic
-  val mk_set_bd_tac: thm -> thm -> thm -> tactic
-  val mk_set_hset_incl_hset_tac: int -> thm list -> thm -> int -> tactic
+  val mk_set_bd_tac: thm -> thm -> tactic
+  val mk_set_hset_incl_hset_tac: int -> thm -> int -> tactic
   val mk_set_image_Lev_tac: Proof.context -> cterm option list -> thm list -> thm list ->
     thm list -> thm list -> thm list list -> thm list list -> tactic
-  val mk_set_incl_hset_tac: thm -> thm -> tactic
+  val mk_set_incl_hset_tac: thm -> tactic
   val mk_set_ge_tac: int  -> thm -> thm list -> tactic
   val mk_set_le_tac: int -> thm -> thm list -> thm list list -> tactic
-  val mk_set_map0_tac: thm -> thm -> tactic
+  val mk_set_map0_tac: thm -> tactic
   val mk_unfold_unique_mor_tac: thm list -> thm -> thm -> thm list -> tactic
   val mk_unfold_transfer_tac: Proof.context -> int -> thm -> thm list -> thm list -> tactic
   val mk_wit_tac: Proof.context -> int -> thm list -> thm list -> thm list -> thm list -> tactic
@@ -117,27 +117,29 @@
   REPEAT_DETERM (eresolve_tac [CollectE, conjE] 1) THEN atac 1;
 
 fun mk_mor_elim_tac mor_def =
-  (dtac (subst OF [mor_def]) THEN'
+  (dtac (mor_def RS iffD1) THEN'
   REPEAT o etac conjE THEN'
   TRY o rtac @{thm image_subsetI} THEN'
   etac bspec THEN'
   atac) 1;
 
 fun mk_mor_incl_tac mor_def map_ids =
-  (stac mor_def THEN'
+  (rtac (mor_def RS iffD2) THEN'
   rtac conjI THEN'
-  CONJ_WRAP' (K (EVERY' [rtac ballI, etac set_mp, stac id_apply, atac])) map_ids THEN'
+  CONJ_WRAP' (K (EVERY' [rtac ballI, etac set_mp, etac (id_apply RS @{thm ssubst_mem})]))
+    map_ids THEN'
   CONJ_WRAP' (fn thm =>
-   (EVERY' [rtac ballI, rtac (thm RS trans), rtac sym, rtac (id_apply RS arg_cong)])) map_ids) 1;
+    (EVERY' [rtac ballI, rtac (thm RS trans), rtac sym, rtac (id_apply RS arg_cong)])) map_ids) 1;
 
 fun mk_mor_comp_tac mor_def mor_images morEs map_comp_ids =
   let
-    fun fbetw_tac image = EVERY' [rtac ballI, stac o_apply, etac image, etac image, atac];
+    fun fbetw_tac image = EVERY' [rtac ballI, rtac (o_apply RS @{thm ssubst_mem}), etac image,
+      etac image, atac];
     fun mor_tac ((mor_image, morE), map_comp_id) =
       EVERY' [rtac ballI, stac o_apply, rtac trans, rtac (map_comp_id RS sym), rtac trans,
         etac (morE RS arg_cong), atac, etac morE, etac mor_image, atac];
   in
-    (stac mor_def THEN' rtac conjI THEN'
+    (rtac (mor_def RS iffD2) THEN' rtac conjI THEN'
     CONJ_WRAP' fbetw_tac mor_images THEN'
     CONJ_WRAP' mor_tac ((mor_images ~~ morEs) ~~ map_comp_ids)) 1
   end;
@@ -149,29 +151,26 @@
       rtac UNIV_I, rtac sym, rtac o_apply];
   in
     EVERY' [rtac iffI, CONJ_WRAP' mor_tac morEs,
-    stac mor_def, rtac conjI, CONJ_WRAP' (K (rtac ballI THEN' rtac UNIV_I)) morEs,
+    rtac (mor_def RS iffD2), rtac conjI, CONJ_WRAP' (K (rtac ballI THEN' rtac UNIV_I)) morEs,
     CONJ_WRAP' (fn i =>
       EVERY' [dtac (mk_conjunctN n i), rtac ballI, etac @{thm comp_eq_dest}]) (1 upto n)] 1
   end;
 
 fun mk_mor_str_tac ks mor_UNIV =
-  (stac mor_UNIV THEN' CONJ_WRAP' (K (rtac refl)) ks) 1;
+  (rtac (mor_UNIV RS iffD2) THEN' CONJ_WRAP' (K (rtac refl)) ks) 1;
 
 fun mk_mor_case_sum_tac ks mor_UNIV =
-  (stac mor_UNIV THEN' CONJ_WRAP' (K (rtac @{thm case_sum_o_inj(1)[symmetric]})) ks) 1;
+  (rtac (mor_UNIV RS iffD2) THEN' CONJ_WRAP' (K (rtac @{thm case_sum_o_inj(1)[symmetric]})) ks) 1;
 
-fun mk_set_incl_hset_tac def rec_Suc =
-  EVERY' (stac def ::
-    map rtac [@{thm SUP_upper2}, UNIV_I, @{thm ord_le_eq_trans}, @{thm Un_upper1},
-      sym, rec_Suc]) 1;
+fun mk_set_incl_hset_tac rec_Suc =
+  EVERY' (map rtac [@{thm SUP_upper2}, UNIV_I, @{thm ord_le_eq_trans}, @{thm Un_upper1}, sym,
+    rec_Suc]) 1;
 
-fun mk_set_hset_incl_hset_tac n defs rec_Suc i =
-  EVERY' (map (TRY oo stac) defs @
-    map rtac [@{thm UN_least}, subsetI, @{thm UN_I}, UNIV_I, set_mp, equalityD2, rec_Suc, UnI2,
-      mk_UnIN n i] @
-    [etac @{thm UN_I}, atac]) 1;
+fun mk_set_hset_incl_hset_tac n rec_Suc i =
+  EVERY' (map rtac [@{thm UN_least}, subsetI, @{thm UN_I}, UNIV_I, set_mp, equalityD2, rec_Suc,
+      UnI2, mk_UnIN n i] @ [etac @{thm UN_I}, atac]) 1;
 
-fun mk_hset_rec_minimal_tac ctxt m cts rec_0s rec_Sucs =
+fun mk_col_minimal_tac ctxt m cts rec_0s rec_Sucs =
   EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
     REPEAT_DETERM o rtac allI,
     CONJ_WRAP' (fn thm => EVERY'
@@ -186,13 +185,12 @@
             rtac subset_trans, atac, Goal.assume_rule_tac ctxt])) rec_0s])
       rec_Sucs] 1;
 
-fun mk_hset_minimal_tac ctxt n hset_defs hset_rec_minimal =
-  (CONJ_WRAP' (fn def => (EVERY' [rtac ord_eq_le_trans, rtac def,
-    rtac @{thm UN_least}, rtac rev_mp, rtac hset_rec_minimal,
+fun mk_hset_minimal_tac ctxt n col_minimal =
+  (CONJ_WRAP' (K (EVERY' [rtac @{thm UN_least}, rtac rev_mp, rtac col_minimal,
     EVERY' (replicate ((n + 1) * n) (Goal.assume_rule_tac ctxt)), rtac impI,
-    REPEAT_DETERM o eresolve_tac [allE, conjE], atac])) hset_defs) 1
+    REPEAT_DETERM o eresolve_tac [allE, conjE], atac])) (1 upto n)) 1
 
-fun mk_mor_hset_rec_tac m n cts j rec_0s rec_Sucs morEs set_map0ss coalg_setss =
+fun mk_mor_col_tac m n cts j rec_0s rec_Sucs morEs set_map0ss coalg_setss =
   EVERY' [rtac (Drule.instantiate' [] cts nat_induct),
     REPEAT_DETERM o rtac allI,
     CONJ_WRAP' (fn thm => EVERY' (map rtac [impI, thm RS trans, thm RS sym])) rec_0s,
@@ -266,7 +264,7 @@
   end;
 
 fun mk_bis_converse_tac m bis_rel rel_congs rel_converseps =
-  EVERY' [stac bis_rel, dtac (bis_rel RS iffD1),
+  EVERY' [rtac (bis_rel RS iffD2), dtac (bis_rel RS iffD1),
     REPEAT_DETERM o etac conjE, rtac conjI,
     CONJ_WRAP' (K (EVERY' [rtac converse_shift, etac subset_trans,
       rtac equalityD2, rtac @{thm converse_Times}])) rel_congs,
@@ -280,7 +278,7 @@
         rtac @{thm conversepI}, etac mp, etac @{thm converseD}]) (rel_congs ~~ rel_converseps)] 1;
 
 fun mk_bis_O_tac ctxt m bis_rel rel_congs rel_OOs =
-  EVERY' [stac bis_rel, REPEAT_DETERM o dtac (bis_rel RS iffD1),
+  EVERY' [rtac (bis_rel RS iffD2), REPEAT_DETERM o dtac (bis_rel RS iffD1),
     REPEAT_DETERM o etac conjE, rtac conjI,
     CONJ_WRAP' (K (EVERY' [etac @{thm relcomp_subset_Sigma}, atac])) rel_congs,
     CONJ_WRAP' (fn (rel_cong, rel_OO) =>
@@ -298,7 +296,7 @@
 fun mk_bis_Gr_tac ctxt bis_rel rel_Grps mor_images morEs coalg_ins =
   unfold_thms_tac ctxt (bis_rel :: @{thm eq_alt} :: @{thm in_rel_Gr} :: rel_Grps) THEN
   EVERY' [rtac conjI,
-    CONJ_WRAP' (fn thm => rtac (@{thm Gr_incl} RS ssubst) THEN' etac thm) mor_images,
+    CONJ_WRAP' (fn thm => rtac (@{thm Gr_incl} RS iffD2) THEN' etac thm) mor_images,
     CONJ_WRAP' (fn (coalg_in, morE) =>
       EVERY' [rtac allI, rtac allI, rtac impI, rtac @{thm GrpI}, etac (morE RS trans),
         etac @{thm GrD1}, etac (@{thm GrD2} RS arg_cong), etac coalg_in, etac @{thm GrD1}])
@@ -564,7 +562,7 @@
             (set_Levs, set_image_Levs))))) =>
             EVERY' [rtac ballI,
               REPEAT_DETERM o eresolve_tac [CollectE, @{thm SuccE}, @{thm UN_E}],
-              rtac (rev_mp OF [rv_last, impI]), etac exE, rtac (isNode_def RS ssubst),
+              rtac (rev_mp OF [rv_last, impI]), etac exE, rtac (isNode_def RS iffD2),
               rtac exI, rtac conjI,
               if n = 1 then rtac refl
               else etac (sum_weak_case_cong RS trans) THEN' rtac (mk_sum_caseN n i),
@@ -582,7 +580,7 @@
           (ks ~~ (rv_lasts ~~ (isNode_defs ~~ (set_map0ss ~~
             (set_Levss ~~ set_image_Levss))))),
         (*root isNode*)
-          rtac (isNode_def RS ssubst), rtac exI, rtac conjI,
+          rtac (isNode_def RS iffD2), rtac exI, rtac conjI,
           CONVERSION (Conv.top_conv
             (K (Conv.try_conv (Conv.rewr_conv (rv_Nil RS eq_reflection)))) ctxt),
           if n = 1 then rtac refl else rtac (mk_sum_caseN n i),
@@ -626,7 +624,7 @@
             rtac @{thm UN_least}, rtac subsetI, rtac CollectI, rtac @{thm UN_I[OF UNIV_I]},
             rtac (Lev_Suc RS equalityD2 RS set_mp), rtac (mk_UnIN n i'), rtac CollectI,
             REPEAT_DETERM o rtac exI, rtac conjI, rtac refl, etac conjI, atac,
-            rtac trans, rtac @{thm shift_def}, rtac ssubst, rtac @{thm fun_eq_iff}, rtac allI,
+            rtac trans, rtac @{thm shift_def}, rtac iffD2, rtac @{thm fun_eq_iff}, rtac allI,
             CONVERSION (Conv.top_conv
               (K (Conv.try_conv (Conv.rewr_conv (rv_Cons RS eq_reflection)))) ctxt),
             if n = 1 then K all_tac
@@ -636,7 +634,7 @@
   in
     (rtac mor_cong THEN'
     EVERY' (map (fn thm => rtac (thm RS @{thm ext})) beh_defs) THEN'
-    stac mor_def THEN' rtac conjI THEN'
+    rtac (mor_def RS iffD2) THEN' rtac conjI THEN'
     CONJ_WRAP' fbetw_tac
       (ks ~~ (carT_defs ~~ (isNode_defs ~~ (Lev_0s ~~ (rv_Nils ~~ ((length_Levs ~~ length_Lev's) ~~
         (rv_lastss ~~ (set_map0ss ~~ (set_Levsss ~~ set_image_Levsss))))))))) THEN'
@@ -657,22 +655,22 @@
     etac (sym RS arg_cong RS trans), rtac map_comp_id] 1;
 
 fun mk_coalg_final_tac m coalg_def congruent_str_finals equiv_LSBISs set_map0ss coalgT_setss =
-  EVERY' [stac coalg_def,
+  EVERY' [rtac (coalg_def RS iffD2),
     CONJ_WRAP' (fn ((set_map0s, coalgT_sets), (equiv_LSBIS, congruent_str_final)) =>
       EVERY' [rtac @{thm univ_preserves}, rtac equiv_LSBIS, rtac congruent_str_final,
         rtac ballI, rtac @{thm ssubst_mem}, rtac o_apply, rtac CollectI,
         REPEAT_DETERM_N m o EVERY' [rtac conjI, rtac subset_UNIV],
         CONJ_WRAP' (fn (equiv_LSBIS, (set_map0, coalgT_set)) =>
           EVERY' [rtac (set_map0 RS ord_eq_le_trans),
-            rtac @{thm image_subsetI}, rtac ssubst, rtac @{thm proj_in_iff},
+            rtac @{thm image_subsetI}, rtac iffD2, rtac @{thm proj_in_iff},
             rtac equiv_LSBIS, etac set_rev_mp, etac coalgT_set])
         (equiv_LSBISs ~~ (drop m set_map0s ~~ coalgT_sets))])
     ((set_map0ss ~~ coalgT_setss) ~~ (equiv_LSBISs ~~ congruent_str_finals))] 1;
 
 fun mk_mor_T_final_tac mor_def congruent_str_finals equiv_LSBISs =
-  EVERY' [stac mor_def, rtac conjI,
+  EVERY' [rtac (mor_def RS iffD2), rtac conjI,
     CONJ_WRAP' (fn equiv_LSBIS =>
-      EVERY' [rtac ballI, rtac ssubst, rtac @{thm proj_in_iff}, rtac equiv_LSBIS, atac])
+      EVERY' [rtac ballI, rtac iffD2, rtac @{thm proj_in_iff}, rtac equiv_LSBIS, atac])
     equiv_LSBISs,
     CONJ_WRAP' (fn (equiv_LSBIS, congruent_str_final) =>
       EVERY' [rtac ballI, rtac sym, rtac trans, rtac @{thm univ_commute}, rtac equiv_LSBIS,
@@ -758,7 +756,7 @@
   etac unfold_unique_mor 1;
 
 fun mk_dtor_coinduct_tac m raw_coind bis_rel rel_congs =
-  EVERY' [rtac rev_mp, rtac raw_coind, rtac ssubst, rtac bis_rel, rtac conjI,
+  EVERY' [rtac rev_mp, rtac raw_coind, rtac iffD2, rtac bis_rel, rtac conjI,
     CONJ_WRAP' (K (rtac @{thm ord_le_eq_trans[OF subset_UNIV UNIV_Times_UNIV[THEN sym]]}))
     rel_congs,
     CONJ_WRAP' (fn rel_cong => EVERY' [rtac allI, rtac allI, rtac impI,
@@ -873,10 +871,9 @@
       (rec_Sucs ~~ (dtor_maps ~~ set_map0ss))] 1
   end;
 
-fun mk_set_map0_tac hset_def col_natural =
-  EVERY' (map rtac [@{thm ext}, o_apply RS trans, hset_def RS trans, sym,
-    o_apply RS trans, @{thm image_cong} OF [hset_def, refl] RS trans,
-    @{thm image_UN} RS trans, refl RS @{thm UN_cong}, col_natural]) 1;
+fun mk_set_map0_tac col_natural =
+  EVERY' (map rtac [@{thm ext}, o_apply RS trans, sym, o_apply RS trans, @{thm image_UN} RS trans,
+    refl RS @{thm UN_cong}, col_natural]) 1;
 
 fun mk_col_bd_tac m j cts rec_0s rec_Sucs sbd_Card_order sbd_Cinfinite set_sbdss =
   let
@@ -897,9 +894,8 @@
       (rec_Sucs ~~ set_sbdss)] 1
   end;
 
-fun mk_set_bd_tac sbd_Cinfinite hset_def col_bd =
-  EVERY' (map rtac [ordIso_ordLeq_trans, @{thm card_of_ordIso_subst}, hset_def,
-    @{thm UNION_Cinfinite_bound}, ordIso_ordLeq_trans, @{thm card_of_nat},
+fun mk_set_bd_tac sbd_Cinfinite col_bd =
+  EVERY' (map rtac [@{thm UNION_Cinfinite_bound}, ordIso_ordLeq_trans, @{thm card_of_nat},
     @{thm natLeq_ordLeq_cinfinite}, sbd_Cinfinite, ballI, col_bd, sbd_Cinfinite]) 1;
 
 fun mk_le_rel_OO_tac coinduct rel_Jrels rel_OOs =
--- a/src/HOL/Tools/BNF/bnf_lfp.ML	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Tools/BNF/bnf_lfp.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -296,7 +296,7 @@
           (HOLogic.mk_Trueprop (mk_talg activeAs ss))
       in
         Goal.prove_sorry lthy [] [] goal
-          (K (stac alg_def 1 THEN CONJ_WRAP (K (EVERY' [rtac ballI, rtac UNIV_I] 1)) ss))
+          (K (rtac (alg_def RS iffD2) 1 THEN CONJ_WRAP (K (EVERY' [rtac ballI, rtac UNIV_I] 1)) ss))
         |> Thm.close_derivation
       end;
 
@@ -985,7 +985,7 @@
       lthy
       |> fold_map3 (fn b => fn mx => fn car_init =>
         typedef (Binding.conceal b, params, mx) car_init NONE
-          (EVERY' [rtac ssubst, rtac @{thm ex_in_conv}, resolve_tac alg_not_empty_thms,
+          (EVERY' [rtac iffD2, rtac @{thm ex_in_conv}, resolve_tac alg_not_empty_thms,
             rtac alg_init_thm] 1)) bs mixfixes car_inits
       |>> apsnd split_list o split_list;
 
@@ -1005,7 +1005,7 @@
     fun mk_inver_thm mk_tac rep abs X thm =
       Goal.prove_sorry lthy [] []
         (HOLogic.mk_Trueprop (mk_inver rep abs X))
-        (K (EVERY' [rtac ssubst, rtac @{thm inver_def}, rtac ballI, mk_tac thm] 1))
+        (K (EVERY' [rtac iffD2, rtac @{thm inver_def}, rtac ballI, mk_tac thm] 1))
       |> Thm.close_derivation;
 
     val inver_Reps = map4 (mk_inver_thm rtac) Abs_Ts Rep_Ts (map HOLogic.mk_UNIV Ts) Rep_inverses;
--- a/src/HOL/Tools/BNF/bnf_lfp_rec_sugar.ML	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Tools/BNF/bnf_lfp_rec_sugar.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -8,6 +8,8 @@
 
 signature BNF_LFP_REC_SUGAR =
 sig
+  datatype primrec_option = Nonexhaustive_Option
+
   type basic_lfp_sugar =
     {T: typ,
      fp_res_index: int,
@@ -33,7 +35,7 @@
 
   val add_primrec: (binding * typ option * mixfix) list ->
     (Attrib.binding * term) list -> local_theory -> (term list * thm list list) * local_theory
-  val add_primrec_cmd: (binding * string option * mixfix) list ->
+  val add_primrec_cmd: primrec_option list -> (binding * string option * mixfix) list ->
     (Attrib.binding * string) list -> local_theory -> (term list * thm list list) * local_theory
   val add_primrec_global: (binding * typ option * mixfix) list ->
     (Attrib.binding * term) list -> theory -> (term list * thm list list) * theory
@@ -62,6 +64,8 @@
 exception OLD_PRIMREC of unit;
 exception PRIMREC of string * term list;
 
+datatype primrec_option = Nonexhaustive_Option;
+
 datatype rec_call =
   No_Rec of int * typ |
   Mutual_Rec of (int * typ) * (int * typ) |
@@ -346,7 +350,8 @@
       |> fold_rev lambda (args @ left_args @ right_args)
     end);
 
-fun build_defs lthy bs mxs (funs_data : eqn_data list list) (rec_specs : rec_spec list) has_call =
+fun build_defs lthy nonexhaustive bs mxs (funs_data : eqn_data list list)
+    (rec_specs : rec_spec list) has_call =
   let
     val n_funs = length funs_data;
 
@@ -355,8 +360,10 @@
       |> maps (uncurry (finds (fn (x, y) => #ctr x = #ctr y))
           ##> (fn x => null x orelse
             raise PRIMREC ("excess equations in definition", map #rhs_term x)) #> fst);
-    val _ = ctr_spec_eqn_data_list' |> map (fn (_, x) => length x <= 1 orelse
-      raise PRIMREC ("multiple equations for constructor", map #user_eqn x));
+    val _ = ctr_spec_eqn_data_list' |> map (fn ({ctr, ...}, x) =>
+        if length x > 1 then raise PRIMREC ("multiple equations for constructor", map #user_eqn x)
+        else if length x = 1 orelse nonexhaustive then ()
+        else warning ("no equation for constructor " ^ Syntax.string_of_term lthy ctr));
 
     val ctr_spec_eqn_data_list =
       ctr_spec_eqn_data_list' @ (drop n_funs rec_specs |> maps #ctr_specs |> map (rpair []));
@@ -414,7 +421,7 @@
   unfold_thms_tac ctxt (nested_simps ctxt @ map_comps @ map_idents) THEN
   HEADGOAL (rtac refl);
 
-fun prepare_primrec fixes specs lthy0 =
+fun prepare_primrec nonexhaustive fixes specs lthy0 =
   let
     val thy = Proof_Context.theory_of lthy0;
 
@@ -455,7 +462,7 @@
         raise PRIMREC ("argument " ^ quote (Syntax.string_of_term lthy ctr) ^
           " is not a constructor in left-hand side", [user_eqn])) eqns_data;
 
-    val defs = build_defs lthy bs mxs funs_data rec_specs has_call;
+    val defs = build_defs lthy nonexhaustive bs mxs funs_data rec_specs has_call;
 
     fun prove lthy' def_thms' ({ctr_specs, nested_map_idents, nested_map_comps, ...} : rec_spec)
         (fun_data : eqn_data list) =
@@ -505,9 +512,10 @@
       lthy |> Local_Theory.notes (notes @ common_notes) |> snd)
   end;
 
-fun add_primrec_simple fixes ts lthy =
+fun add_primrec_simple' opts fixes ts lthy =
   let
-    val (((names, defs), prove), lthy') = prepare_primrec fixes ts lthy
+    val nonexhaustive = member (op =) opts Nonexhaustive_Option;
+    val (((names, defs), prove), lthy') = prepare_primrec nonexhaustive fixes ts lthy
       handle ERROR str => raise PRIMREC (str, []);
   in
     lthy'
@@ -521,8 +529,10 @@
            error ("primrec error:\n  " ^ str ^ "\nin\n  " ^
              space_implode "\n  " (map (quote o Syntax.string_of_term lthy) eqns));
 
-fun gen_primrec old_primrec prep_spec (raw_fixes : (binding * 'a option * mixfix) list) raw_spec
-    lthy =
+val add_primrec_simple = add_primrec_simple' [];
+
+fun gen_primrec old_primrec prep_spec opts
+    (raw_fixes : (binding * 'a option * mixfix) list) raw_spec lthy =
   let
     val d = duplicates (op =) (map (Binding.name_of o #1) raw_fixes)
     val _ = null d orelse raise PRIMREC ("duplicate function name(s): " ^ commas d, []);
@@ -543,7 +553,7 @@
         end);
   in
     lthy
-    |> add_primrec_simple fixes (map snd specs)
+    |> add_primrec_simple' opts fixes (map snd specs)
     |-> (fn (names, (ts, (jss, simpss))) =>
       Spec_Rules.add Spec_Rules.Equational (ts, flat simpss)
       #> Local_Theory.notes (mk_notes jss names simpss)
@@ -551,7 +561,7 @@
   end
   handle OLD_PRIMREC () => old_primrec raw_fixes raw_spec lthy |>> apsnd single;
 
-val add_primrec = gen_primrec Primrec.add_primrec Specification.check_spec;
+val add_primrec = gen_primrec Primrec.add_primrec Specification.check_spec [];
 val add_primrec_cmd = gen_primrec Primrec.add_primrec_cmd Specification.read_spec;
 
 fun add_primrec_global fixes specs =
@@ -564,8 +574,14 @@
   #> add_primrec fixes specs
   ##> Local_Theory.exit_global;
 
+val primrec_option_parser = Parse.group (fn () => "option")
+  (Parse.reserved "nonexhaustive" >> K Nonexhaustive_Option)
+
 val _ = Outer_Syntax.local_theory @{command_spec "primrec"}
   "define primitive recursive functions"
-  (Parse.fixes -- Parse_Spec.where_alt_specs >> (snd oo uncurry add_primrec_cmd));
+  ((Scan.optional (@{keyword "("} |--
+      Parse.!!! (Parse.list1 primrec_option_parser) --| @{keyword ")"}) []) --
+    (Parse.fixes -- Parse_Spec.where_alt_specs)
+    >> (fn (opts, (fixes, spec)) => snd o add_primrec_cmd opts fixes spec));
 
 end;
--- a/src/HOL/Tools/BNF/bnf_lfp_tactics.ML	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Tools/BNF/bnf_lfp_tactics.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -86,7 +86,8 @@
 val subset_trans = @{thm subset_trans};
 val trans_fun_cong_image_id_id_apply = @{thm trans[OF fun_cong[OF image_id] id_apply]};
 val rev_bspec = Drule.rotate_prems 1 bspec;
-val Un_cong = @{thm arg_cong2[of _ _ _ _ "op \<union>"]}
+val Un_cong = @{thm arg_cong2[of _ _ _ _ "op \<union>"]};
+val relChainD = @{thm iffD2[OF meta_eq_to_obj_eq[OF relChain_def]]};
 
 fun mk_alg_set_tac alg_def =
   EVERY' [dtac (alg_def RS iffD1), REPEAT_DETERM o etac conjE, etac bspec, rtac CollectI,
@@ -102,24 +103,26 @@
   etac @{thm emptyE}) 1;
 
 fun mk_mor_elim_tac mor_def =
-  (dtac (subst OF [mor_def]) THEN'
+  (dtac (mor_def RS iffD1) THEN'
   REPEAT o etac conjE THEN'
   TRY o rtac @{thm image_subsetI} THEN'
   etac bspec THEN'
   atac) 1;
 
 fun mk_mor_incl_tac mor_def map_ids =
-  (stac mor_def THEN'
+  (rtac (mor_def RS iffD2) THEN'
   rtac conjI THEN'
-  CONJ_WRAP' (K (EVERY' [rtac ballI, etac set_mp, stac id_apply, atac])) map_ids THEN'
+  CONJ_WRAP' (K (EVERY' [rtac ballI, etac set_mp, etac (id_apply RS @{thm ssubst_mem})]))
+    map_ids THEN'
   CONJ_WRAP' (fn thm =>
-   (EVERY' [rtac ballI, rtac trans, rtac id_apply, stac thm, rtac refl])) map_ids) 1;
+    (EVERY' [rtac ballI, rtac trans, rtac id_apply, stac thm, rtac refl])) map_ids) 1;
 
 fun mk_mor_comp_tac mor_def set_maps map_comp_ids =
   let
-    val fbetw_tac = EVERY' [rtac ballI, stac o_apply, etac bspec, etac bspec, atac];
+    val fbetw_tac =
+      EVERY' [rtac ballI, rtac (o_apply RS @{thm ssubst_mem}), etac bspec, etac bspec, atac];
     fun mor_tac (set_map, map_comp_id) =
-      EVERY' [rtac ballI, stac o_apply, rtac trans,
+      EVERY' [rtac ballI, rtac (o_apply RS trans), rtac trans,
         rtac trans, dtac rev_bspec, atac, etac arg_cong,
          REPEAT o eresolve_tac [CollectE, conjE], etac bspec, rtac CollectI] THEN'
       CONJ_WRAP' (fn thm =>
@@ -128,7 +131,7 @@
             etac bspec, etac set_mp, atac])]) set_map THEN'
       rtac (map_comp_id RS arg_cong);
   in
-    (dtac (mor_def RS subst) THEN' dtac (mor_def RS subst) THEN' stac mor_def THEN'
+    (dtac (mor_def RS iffD1) THEN' dtac (mor_def RS iffD1) THEN' rtac (mor_def RS iffD2) THEN'
     REPEAT o etac conjE THEN'
     rtac conjI THEN'
     CONJ_WRAP' (K fbetw_tac) set_maps THEN'
@@ -150,9 +153,9 @@
          rtac trans, etac (morE RS arg_cong), rtac CollectI, Collect_tac set_map,
          rtac trans, rtac (map_comp_id RS arg_cong), rtac (map_cong0L RS arg_cong),
          REPEAT_DETERM_N (length morEs) o
-           (EVERY' [rtac subst, rtac @{thm inver_pointfree}, etac @{thm inver_mono}, atac])];
+           (EVERY' [rtac iffD1, rtac @{thm inver_pointfree}, etac @{thm inver_mono}, atac])];
   in
-    (stac mor_def THEN'
+    (rtac (mor_def RS iffD2) THEN'
     dtac (alg_def RS iffD1) THEN'
     dtac (alg_def RS iffD1) THEN'
     REPEAT o etac conjE THEN'
@@ -162,12 +165,12 @@
   end;
 
 fun mk_mor_str_tac ks mor_def =
-  (stac mor_def THEN' rtac conjI THEN'
+  (rtac (mor_def RS iffD2) THEN' rtac conjI THEN'
   CONJ_WRAP' (K (EVERY' [rtac ballI, rtac UNIV_I])) ks THEN'
   CONJ_WRAP' (K (EVERY' [rtac ballI, rtac refl])) ks) 1;
 
 fun mk_mor_convol_tac ks mor_def =
-  (stac mor_def THEN' rtac conjI THEN'
+  (rtac (mor_def RS iffD2) THEN' rtac conjI THEN'
   CONJ_WRAP' (K (EVERY' [rtac ballI, rtac UNIV_I])) ks THEN'
   CONJ_WRAP' (K (EVERY' [rtac ballI, rtac trans, rtac @{thm fst_convol'}, rtac o_apply])) ks) 1;
 
@@ -179,17 +182,17 @@
       rtac sym, rtac o_apply];
   in
     EVERY' [rtac iffI, CONJ_WRAP' mor_tac morEs,
-    stac mor_def, rtac conjI, CONJ_WRAP' (K (rtac ballI THEN' rtac UNIV_I)) morEs,
-    REPEAT_DETERM o etac conjE, REPEAT_DETERM_N n o dtac (@{thm fun_eq_iff} RS subst),
+    rtac (mor_def RS iffD2), rtac conjI, CONJ_WRAP' (K (rtac ballI THEN' rtac UNIV_I)) morEs,
+    REPEAT_DETERM o etac conjE, REPEAT_DETERM_N n o dtac (@{thm fun_eq_iff} RS iffD1),
     CONJ_WRAP' (K (EVERY' [rtac ballI, REPEAT_DETERM o etac allE, rtac trans,
-      etac (o_apply RS subst), rtac o_apply])) morEs] 1
+      etac (o_apply RS sym RS trans), rtac o_apply])) morEs] 1
   end;
 
 fun mk_iso_alt_tac mor_images mor_inv =
   let
     val n = length mor_images;
     fun if_wrap_tac thm =
-      EVERY' [rtac ssubst, rtac @{thm bij_betw_iff_ex}, rtac exI, rtac conjI,
+      EVERY' [rtac iffD2, rtac @{thm bij_betw_iff_ex}, rtac exI, rtac conjI,
         rtac @{thm inver_surj}, etac thm, etac thm, atac, etac conjI, atac]
     val if_tac =
       EVERY' [etac thin_rl, etac thin_rl, REPEAT o eresolve_tac [conjE, exE],
@@ -222,7 +225,7 @@
       (set_maps ~~ alg_sets);
   in
     (rtac rev_mp THEN' DETERM o bij_betw_inv_tac THEN' rtac impI THEN'
-    stac alg_def THEN' copy_str_tac) 1
+    rtac (alg_def RS iffD2) THEN' copy_str_tac) 1
   end;
 
 fun mk_copy_alg_tac ctxt set_maps alg_sets mor_def iso_alt copy_str =
@@ -240,7 +243,7 @@
       (set_maps ~~ alg_sets);
   in
     (rtac (iso_alt RS iffD2) THEN' etac copy_str THEN' REPEAT_DETERM o atac THEN'
-    rtac conjI THEN' stac mor_def THEN' rtac conjI THEN' fbetw_tac THEN' mor_tac THEN'
+    rtac conjI THEN' rtac (mor_def RS iffD2) THEN' rtac conjI THEN' fbetw_tac THEN' mor_tac THEN'
     CONJ_WRAP' (K atac) alg_sets) 1
   end;
 
@@ -270,16 +273,16 @@
       EVERY' [rtac Un_cong, minG_tac, rtac @{thm image_cong}, rtac thm,
         REPEAT_DETERM_N (length in_congs) o minG_tac, rtac refl];
   in
-    (rtac (worel RS (@{thm wo_rel.worec_fixpoint} RS fun_cong)) THEN' rtac ssubst THEN'
+    (rtac (worel RS (@{thm wo_rel.worec_fixpoint} RS fun_cong)) THEN' rtac iffD2 THEN'
     rtac meta_eq_to_obj_eq THEN' rtac (worel RS @{thm wo_rel.adm_wo_def}) THEN'
     REPEAT_DETERM_N 3 o rtac allI THEN' rtac impI THEN'
     CONJ_WRAP_GEN' (EVERY' [rtac Pair_eqI, rtac conjI]) minH_tac in_congs) 1
   end;
 
-fun mk_min_algs_mono_tac ctxt min_algs = EVERY' [stac @{thm relChain_def}, rtac allI, rtac allI,
-  rtac impI, rtac @{thm case_split}, rtac @{thm xt1(3)}, rtac min_algs, etac @{thm FieldI2},
-  rtac subsetI, rtac UnI1, rtac @{thm UN_I}, etac @{thm underS_I}, atac, atac,
-  rtac equalityD1, dtac @{thm notnotD}, hyp_subst_tac ctxt, rtac refl] 1;
+fun mk_min_algs_mono_tac ctxt min_algs = EVERY' [rtac relChainD, rtac allI, rtac allI, rtac impI,
+  rtac @{thm case_split}, rtac @{thm xt1(3)}, rtac min_algs, etac @{thm FieldI2}, rtac subsetI,
+  rtac UnI1, rtac @{thm UN_I}, etac @{thm underS_I}, atac, atac, rtac equalityD1,
+  dtac @{thm notnotD}, hyp_subst_tac ctxt, rtac refl] 1;
 
 fun mk_min_algs_card_of_tac cT ct m worel min_algs in_bds bd_Card_order bd_Cnotzero
   suc_Card_order suc_Cinfinite suc_Cnotzero suc_Asuc Asuc_Cinfinite =
@@ -370,13 +373,14 @@
   end;
 
 fun mk_card_of_min_alg_tac min_alg_def card_of suc_Card_order suc_Asuc Asuc_Cinfinite =
-  EVERY' [stac min_alg_def, rtac @{thm UNION_Cinfinite_bound},
-    rtac @{thm ordIso_ordLeq_trans}, rtac @{thm card_of_Field_ordIso}, rtac suc_Card_order,
-    rtac @{thm ordLess_imp_ordLeq}, rtac suc_Asuc, rtac ballI, dtac rev_mp, rtac card_of,
-    REPEAT_DETERM o etac conjE, atac, rtac Asuc_Cinfinite] 1;
+  EVERY' [rtac @{thm ordIso_ordLeq_trans}, rtac (min_alg_def RS @{thm card_of_ordIso_subst}),
+    rtac @{thm UNION_Cinfinite_bound}, rtac @{thm ordIso_ordLeq_trans},
+    rtac @{thm card_of_Field_ordIso}, rtac suc_Card_order, rtac @{thm ordLess_imp_ordLeq},
+    rtac suc_Asuc, rtac ballI, dtac rev_mp, rtac card_of,  REPEAT_DETERM o etac conjE, atac,
+    rtac Asuc_Cinfinite] 1;
 
 fun mk_least_min_alg_tac min_alg_def least =
-  EVERY' [stac min_alg_def, rtac @{thm UN_least}, dtac least, dtac mp, atac,
+  EVERY' [rtac (min_alg_def RS ord_eq_le_trans), rtac @{thm UN_least}, dtac least, dtac mp, atac,
     REPEAT_DETERM o etac conjE, atac] 1;
 
 fun mk_alg_select_tac ctxt Abs_inverse =
@@ -393,24 +397,25 @@
       CONJ_WRAP' (fn thm => EVERY' [rtac ballI, rtac thm]) str_init_defs;
     fun alg_epi_tac ((alg_set, str_init_def), set_map) =
       EVERY' [rtac ballI, REPEAT_DETERM o eresolve_tac [CollectE, conjE], rtac CollectI,
-        rtac ballI, ftac (alg_select RS bspec), stac str_init_def, etac alg_set,
-        REPEAT_DETERM o EVERY' [rtac ord_eq_le_trans, resolve_tac set_map, rtac subset_trans,
-            etac @{thm image_mono}, rtac @{thm image_Collect_subsetI}, etac bspec, atac]];
+        rtac ballI, ftac (alg_select RS bspec), rtac (str_init_def RS @{thm ssubst_mem}),
+        etac alg_set, REPEAT_DETERM o EVERY' [rtac ord_eq_le_trans, resolve_tac set_map,
+          rtac subset_trans, etac @{thm image_mono}, rtac @{thm image_Collect_subsetI}, etac bspec,
+          atac]];
   in
-    (rtac mor_cong THEN' REPEAT_DETERM_N n o (rtac sym THEN' rtac @{thm comp_id}) THEN'
-    rtac (Thm.permute_prems 0 1 mor_comp) THEN' etac (Thm.permute_prems 0 1 mor_comp) THEN'
-    stac mor_def THEN' rtac conjI THEN' fbetw_tac THEN' mor_tac THEN' rtac mor_incl_min_alg THEN'
-    stac alg_def THEN' CONJ_WRAP' alg_epi_tac ((alg_sets ~~ str_init_defs) ~~ set_maps)) 1
+    EVERY' [rtac mor_cong, REPEAT_DETERM_N n o (rtac sym THEN' rtac @{thm comp_id}),
+      rtac (Thm.permute_prems 0 1 mor_comp), etac (Thm.permute_prems 0 1 mor_comp),
+      rtac (mor_def RS iffD2), rtac conjI, fbetw_tac, mor_tac, rtac mor_incl_min_alg,
+      rtac (alg_def RS iffD2), CONJ_WRAP' alg_epi_tac ((alg_sets ~~ str_init_defs) ~~ set_maps)] 1
   end;
 
 fun mk_init_ex_mor_tac ctxt Abs_inverse copy_alg_ex alg_min_alg card_of_min_algs
     mor_comp mor_select mor_incl_min_alg =
   let
     val n = length card_of_min_algs;
-    val card_of_ordIso_tac = EVERY' [rtac ssubst, rtac @{thm card_of_ordIso},
+    val card_of_ordIso_tac = EVERY' [rtac iffD2, rtac @{thm card_of_ordIso},
       rtac @{thm ordIso_symmetric}, rtac conjunct1, rtac conjunct2, atac];
-    fun internalize_tac card_of = EVERY' [rtac subst, rtac @{thm internalize_card_of_ordLeq2},
-      rtac @{thm ordLeq_ordIso_trans}, rtac card_of, rtac subst,
+    fun internalize_tac card_of = EVERY' [rtac iffD1, rtac @{thm internalize_card_of_ordLeq2},
+      rtac @{thm ordLeq_ordIso_trans}, rtac card_of, rtac iffD1,
       rtac @{thm Card_order_iff_ordIso_card_of}, rtac @{thm Card_order_cexp}];
   in
     (rtac rev_mp THEN'
@@ -451,7 +456,7 @@
 
     fun mk_unique_tac (k, least_min_alg) =
       select_prem_tac n (etac @{thm prop_restrict}) k THEN' rtac least_min_alg THEN'
-      stac alg_def THEN'
+      rtac (alg_def RS iffD2) THEN'
       CONJ_WRAP' mk_alg_tac (alg_sets ~~ (in_monos ~~ (morEs ~~ map_cong0s)));
   in
     CONJ_WRAP' mk_unique_tac (ks ~~ least_min_algs) 1
@@ -471,7 +476,7 @@
 
     fun mk_induct_tac least_min_alg =
       rtac ballI THEN' etac @{thm prop_restrict} THEN' rtac least_min_alg THEN'
-      stac alg_def THEN'
+      rtac (alg_def RS iffD2) THEN'
       CONJ_WRAP' mk_alg_tac alg_sets;
   in
     CONJ_WRAP' mk_induct_tac least_min_algs 1
@@ -507,8 +512,8 @@
   end;
 
 fun mk_dtor_o_ctor_tac dtor_def foldx map_comp_id map_cong0L ctor_o_folds =
-  EVERY' [stac dtor_def, rtac @{thm ext}, rtac trans, rtac o_apply, rtac trans, rtac foldx,
-    rtac trans, rtac map_comp_id, rtac trans, rtac map_cong0L,
+  EVERY' [rtac @{thm ext}, rtac trans, rtac o_apply, rtac (dtor_def RS fun_cong RS trans),
+    rtac trans, rtac foldx, rtac trans, rtac map_comp_id, rtac trans, rtac map_cong0L,
     EVERY' (map (fn thm => rtac ballI THEN' rtac (trans OF [thm RS fun_cong, id_apply]))
       ctor_o_folds),
     rtac sym, rtac id_apply] 1;
@@ -530,18 +535,18 @@
     val ks = 1 upto n;
 
     fun mk_IH_tac Rep_inv Abs_inv set_map =
-      DETERM o EVERY' [dtac meta_mp, rtac (Rep_inv RS arg_cong RS subst), etac bspec,
+      DETERM o EVERY' [dtac meta_mp, rtac (Rep_inv RS arg_cong RS iffD1), etac bspec,
         dtac set_rev_mp, rtac equalityD1, rtac set_map, etac imageE,
-        hyp_subst_tac ctxt, rtac (Abs_inv RS ssubst), etac set_mp, atac, atac];
+        hyp_subst_tac ctxt, rtac (Abs_inv RS @{thm ssubst_mem}), etac set_mp, atac, atac];
 
     fun mk_closed_tac (k, (morE, set_maps)) =
       EVERY' [select_prem_tac n (dtac asm_rl) k, rtac ballI, rtac impI,
-        rtac (mor_Abs RS morE RS arg_cong RS ssubst), atac,
+        rtac (mor_Abs RS morE RS arg_cong RS iffD2), atac,
         REPEAT_DETERM o eresolve_tac [CollectE, conjE], dtac @{thm meta_spec},
         EVERY' (map3 mk_IH_tac Rep_invs Abs_invs (drop m set_maps)), atac];
 
     fun mk_induct_tac (Rep, Rep_inv) =
-      EVERY' [rtac (Rep_inv RS arg_cong RS subst), etac (Rep RSN (2, bspec))];
+      EVERY' [rtac (Rep_inv RS arg_cong RS iffD1), etac (Rep RSN (2, bspec))];
   in
     (rtac mp THEN' rtac impI THEN'
     DETERM o CONJ_WRAP_GEN' (etac conjE THEN' rtac conjI) mk_induct_tac (Reps ~~ Rep_invs) THEN'
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/SMT2/smt2_builtin.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -0,0 +1,222 @@
+(*  Title:      HOL/Tools/SMT2/smt2_builtin.ML
+    Author:     Sascha Boehme, TU Muenchen
+
+Tables of types and terms directly supported by SMT solvers.
+*)
+
+signature SMT2_BUILTIN =
+sig
+  (*for experiments*)
+  val filter_builtins: (typ -> bool) -> Proof.context -> Proof.context
+
+  (*built-in types*)
+  val add_builtin_typ: SMT2_Util.class ->
+    typ * (typ -> string option) * (typ -> int -> string option) -> Context.generic ->
+    Context.generic
+  val add_builtin_typ_ext: typ * (typ -> bool) -> Context.generic ->
+    Context.generic
+  val dest_builtin_typ: Proof.context -> typ -> string option
+  val is_builtin_typ_ext: Proof.context -> typ -> bool
+
+  (*built-in numbers*)
+  val dest_builtin_num: Proof.context -> term -> (string * typ) option
+  val is_builtin_num: Proof.context -> term -> bool
+  val is_builtin_num_ext: Proof.context -> term -> bool
+
+  (*built-in functions*)
+  type 'a bfun = Proof.context -> typ -> term list -> 'a
+  type bfunr = string * int * term list * (term list -> term)
+  val add_builtin_fun: SMT2_Util.class -> (string * typ) * bfunr option bfun -> Context.generic ->
+    Context.generic
+  val add_builtin_fun': SMT2_Util.class -> term * string -> Context.generic -> Context.generic
+  val add_builtin_fun_ext: (string * typ) * term list bfun -> Context.generic -> Context.generic
+  val add_builtin_fun_ext': string * typ -> Context.generic -> Context.generic
+  val add_builtin_fun_ext'': string -> Context.generic -> Context.generic
+  val dest_builtin_fun: Proof.context -> string * typ -> term list -> bfunr option
+  val dest_builtin_eq: Proof.context -> term -> term -> bfunr option
+  val dest_builtin_pred: Proof.context -> string * typ -> term list -> bfunr option
+  val dest_builtin_conn: Proof.context -> string * typ -> term list -> bfunr option
+  val dest_builtin: Proof.context -> string * typ -> term list -> bfunr option
+  val dest_builtin_ext: Proof.context -> string * typ -> term list -> term list option
+  val is_builtin_fun: Proof.context -> string * typ -> term list -> bool
+  val is_builtin_fun_ext: Proof.context -> string * typ -> term list -> bool
+end
+
+structure SMT2_Builtin: SMT2_BUILTIN =
+struct
+
+
+(* built-in tables *)
+
+datatype ('a, 'b) kind = Ext of 'a | Int of 'b
+
+type ('a, 'b) ttab = ((typ * ('a, 'b) kind) Ord_List.T) SMT2_Util.dict 
+
+fun typ_ord ((T, _), (U, _)) =
+  let
+    fun tord (TVar _, Type _) = GREATER
+      | tord (Type _, TVar _) = LESS
+      | tord (Type (n, Ts), Type (m, Us)) =
+          if n = m then list_ord tord (Ts, Us)
+          else Term_Ord.typ_ord (T, U)
+      | tord TU = Term_Ord.typ_ord TU
+  in tord (T, U) end
+
+fun insert_ttab cs T f =
+  SMT2_Util.dict_map_default (cs, [])
+    (Ord_List.insert typ_ord (perhaps (try Logic.varifyT_global) T, f))
+
+fun merge_ttab ttabp = SMT2_Util.dict_merge (Ord_List.merge typ_ord) ttabp
+
+fun lookup_ttab ctxt ttab T =
+  let fun match (U, _) = Sign.typ_instance (Proof_Context.theory_of ctxt) (T, U)
+  in
+    get_first (find_first match) (SMT2_Util.dict_lookup ttab (SMT2_Config.solver_class_of ctxt))
+  end
+
+type ('a, 'b) btab = ('a, 'b) ttab Symtab.table
+
+fun insert_btab cs n T f =
+  Symtab.map_default (n, []) (insert_ttab cs T f)
+
+fun merge_btab btabp = Symtab.join (K merge_ttab) btabp
+
+fun lookup_btab ctxt btab (n, T) =
+  (case Symtab.lookup btab n of
+    NONE => NONE
+  | SOME ttab => lookup_ttab ctxt ttab T)
+
+type 'a bfun = Proof.context -> typ -> term list -> 'a
+
+type bfunr = string * int * term list * (term list -> term)
+
+structure Builtins = Generic_Data
+(
+  type T =
+    (typ -> bool, (typ -> string option) * (typ -> int -> string option)) ttab *
+    (term list bfun, bfunr option bfun) btab
+  val empty = ([], Symtab.empty)
+  val extend = I
+  fun merge ((t1, b1), (t2, b2)) = (merge_ttab (t1, t2), merge_btab (b1, b2))
+)
+
+fun filter_ttab keep_T = map (apsnd (filter (keep_T o fst)))
+
+fun filter_builtins keep_T =
+  Context.proof_map (Builtins.map (fn (ttab, btab) =>
+    (filter_ttab keep_T ttab, Symtab.map (K (filter_ttab keep_T)) btab)))
+
+
+(* built-in types *)
+
+fun add_builtin_typ cs (T, f, g) =
+  Builtins.map (apfst (insert_ttab cs T (Int (f, g))))
+
+fun add_builtin_typ_ext (T, f) = Builtins.map (apfst (insert_ttab SMT2_Util.basicC T (Ext f)))
+
+fun lookup_builtin_typ ctxt =
+  lookup_ttab ctxt (fst (Builtins.get (Context.Proof ctxt)))
+
+fun dest_builtin_typ ctxt T =
+  (case lookup_builtin_typ ctxt T of
+    SOME (_, Int (f, _)) => f T
+  | _ => NONE) 
+
+fun is_builtin_typ_ext ctxt T =
+  (case lookup_builtin_typ ctxt T of
+    SOME (_, Int (f, _)) => is_some (f T)
+  | SOME (_, Ext f) => f T
+  | NONE => false)
+
+
+(* built-in numbers *)
+
+fun dest_builtin_num ctxt t =
+  (case try HOLogic.dest_number t of
+    NONE => NONE
+  | SOME (T, i) =>
+      if i < 0 then NONE else
+        (case lookup_builtin_typ ctxt T of
+          SOME (_, Int (_, g)) => g T i |> Option.map (rpair T)
+        | _ => NONE))
+
+val is_builtin_num = is_some oo dest_builtin_num
+
+fun is_builtin_num_ext ctxt t =
+  (case try HOLogic.dest_number t of
+    NONE => false
+  | SOME (T, _) => is_builtin_typ_ext ctxt T)
+
+
+(* built-in functions *)
+
+fun add_builtin_fun cs ((n, T), f) =
+  Builtins.map (apsnd (insert_btab cs n T (Int f)))
+
+fun add_builtin_fun' cs (t, n) =
+  let
+    val c as (m, T) = Term.dest_Const t
+    fun app U ts = Term.list_comb (Const (m, U), ts)
+    fun bfun _ U ts = SOME (n, length (Term.binder_types T), ts, app U)
+  in add_builtin_fun cs (c, bfun) end
+
+fun add_builtin_fun_ext ((n, T), f) =
+  Builtins.map (apsnd (insert_btab SMT2_Util.basicC n T (Ext f)))
+
+fun add_builtin_fun_ext' c = add_builtin_fun_ext (c, fn _ => fn _ => I)
+
+fun add_builtin_fun_ext'' n context =
+  let val thy = Context.theory_of context
+  in add_builtin_fun_ext' (n, Sign.the_const_type thy n) context end
+
+fun lookup_builtin_fun ctxt =
+  lookup_btab ctxt (snd (Builtins.get (Context.Proof ctxt)))
+
+fun dest_builtin_fun ctxt (c as (_, T)) ts =
+  (case lookup_builtin_fun ctxt c of
+    SOME (_, Int f) => f ctxt T ts
+  | _ => NONE)
+
+fun dest_builtin_eq ctxt t u =
+  let
+    val aT = TFree (Name.aT, @{sort type})
+    val c = (@{const_name HOL.eq}, aT --> aT --> @{typ bool})
+    fun mk ts = Term.list_comb (HOLogic.eq_const (Term.fastype_of (hd ts)), ts)
+  in
+    dest_builtin_fun ctxt c []
+    |> Option.map (fn (n, i, _, _) => (n, i, [t, u], mk))
+  end
+
+fun special_builtin_fun pred ctxt (c as (_, T)) ts =
+  if pred (Term.body_type T, Term.binder_types T) then
+    dest_builtin_fun ctxt c ts
+  else NONE
+
+fun dest_builtin_pred ctxt = special_builtin_fun (equal @{typ bool} o fst) ctxt
+
+fun dest_builtin_conn ctxt =
+  special_builtin_fun (forall (equal @{typ bool}) o (op ::)) ctxt
+
+fun dest_builtin ctxt c ts =
+  let val t = Term.list_comb (Const c, ts)
+  in
+    (case dest_builtin_num ctxt t of
+      SOME (n, _) => SOME (n, 0, [], K t)
+    | NONE => dest_builtin_fun ctxt c ts)
+  end
+
+fun dest_builtin_fun_ext ctxt (c as (_, T)) ts =    
+  (case lookup_builtin_fun ctxt c of
+    SOME (_, Int f) => f ctxt T ts |> Option.map (fn (_, _, us, _) => us)
+  | SOME (_, Ext f) => SOME (f ctxt T ts)
+  | NONE => NONE)
+
+fun dest_builtin_ext ctxt c ts =
+  if is_builtin_num_ext ctxt (Term.list_comb (Const c, ts)) then SOME []
+  else dest_builtin_fun_ext ctxt c ts
+
+fun is_builtin_fun ctxt c ts = is_some (dest_builtin_fun ctxt c ts)
+
+fun is_builtin_fun_ext ctxt c ts = is_some (dest_builtin_fun_ext ctxt c ts)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/SMT2/smt2_config.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -0,0 +1,250 @@
+(*  Title:      HOL/Tools/SMT2/smt2_config.ML
+    Author:     Sascha Boehme, TU Muenchen
+
+Configuration options and diagnostic tools for SMT.
+*)
+
+signature SMT2_CONFIG =
+sig
+  (*solver*)
+  type solver_info = {
+    name: string,
+    class: Proof.context -> SMT2_Util.class,
+    avail: unit -> bool,
+    options: Proof.context -> string list }
+  val add_solver: solver_info -> Context.generic -> Context.generic
+  val set_solver_options: string * string -> Context.generic -> Context.generic
+  val is_available: Proof.context -> string -> bool
+  val available_solvers_of: Proof.context -> string list
+  val select_solver: string -> Context.generic -> Context.generic
+  val solver_of: Proof.context -> string
+  val solver_class_of: Proof.context -> SMT2_Util.class
+  val solver_options_of: Proof.context -> string list
+
+  (*options*)
+  val oracle: bool Config.T
+  val timeout: real Config.T
+  val random_seed: int Config.T
+  val read_only_certificates: bool Config.T
+  val verbose: bool Config.T
+  val trace: bool Config.T
+  val trace_used_facts: bool Config.T
+  val monomorph_limit: int Config.T
+  val monomorph_instances: int Config.T
+  val infer_triggers: bool Config.T
+  val filter_only_facts: bool Config.T
+  val debug_files: string Config.T
+
+  (*tools*)
+  val with_timeout: Proof.context -> ('a -> 'b) -> 'a -> 'b
+
+  (*diagnostics*)
+  val trace_msg: Proof.context -> ('a -> string) -> 'a -> unit
+  val verbose_msg: Proof.context -> ('a -> string) -> 'a -> unit
+
+  (*certificates*)
+  val select_certificates: string -> Context.generic -> Context.generic
+  val certificates_of: Proof.context -> Cache_IO.cache option
+
+  (*setup*)
+  val print_setup: Proof.context -> unit
+end
+
+structure SMT2_Config: SMT2_CONFIG =
+struct
+
+(* solver *)
+
+type solver_info = {
+  name: string,
+  class: Proof.context -> SMT2_Util.class,
+  avail: unit -> bool,
+  options: Proof.context -> string list}
+
+(* FIXME just one data slot (record) per program unit *)
+structure Solvers = Generic_Data
+(
+  type T = (solver_info * string list) Symtab.table * string option
+  val empty = (Symtab.empty, NONE)
+  val extend = I
+  fun merge ((ss1, s1), (ss2, s2)) =
+    (Symtab.merge (K true) (ss1, ss2), merge_options (s1, s2))
+)
+
+fun set_solver_options (name, options) =
+  let val opts = String.tokens (Symbol.is_ascii_blank o str) options
+  in Solvers.map (apfst (Symtab.map_entry name (apsnd (K opts)))) end
+
+fun add_solver (info as {name, ...} : solver_info) context =
+  if Symtab.defined (fst (Solvers.get context)) name then
+    error ("Solver already registered: " ^ quote name)
+  else
+    context
+    |> Solvers.map (apfst (Symtab.update (name, (info, []))))
+    |> Context.map_theory (Attrib.setup (Binding.name (name ^ "_options"))
+        (Scan.lift (@{keyword "="} |-- Args.name) >>
+          (Thm.declaration_attribute o K o set_solver_options o pair name))
+        ("Additional command line options for SMT solver " ^ quote name))
+
+fun all_solvers_of ctxt = Symtab.keys (fst (Solvers.get (Context.Proof ctxt)))
+
+fun solver_name_of ctxt = snd (Solvers.get (Context.Proof ctxt))
+
+fun is_available ctxt name =
+  (case Symtab.lookup (fst (Solvers.get (Context.Proof ctxt))) name of
+    SOME ({avail, ...}, _) => avail ()
+  | NONE => false)
+
+fun available_solvers_of ctxt =
+  filter (is_available ctxt) (all_solvers_of ctxt)
+
+fun warn_solver (Context.Proof ctxt) name =
+      Context_Position.if_visible ctxt
+        warning ("The SMT solver " ^ quote name ^ " is not installed.")
+  | warn_solver _ _ = ();
+
+fun select_solver name context =
+  let
+    val ctxt = Context.proof_of context
+    val upd = Solvers.map (apsnd (K (SOME name)))
+  in
+    if not (member (op =) (all_solvers_of ctxt) name) then
+      error ("Trying to select unknown solver: " ^ quote name)
+    else if not (is_available ctxt name) then
+      (warn_solver context name; upd context)
+    else upd context
+  end
+
+fun no_solver_err () = error "No SMT solver selected"
+
+fun solver_of ctxt =
+  (case solver_name_of ctxt of
+    SOME name => name
+  | NONE => no_solver_err ())
+
+fun solver_info_of default select ctxt =
+  (case Solvers.get (Context.Proof ctxt) of
+    (solvers, SOME name) => select (Symtab.lookup solvers name)
+  | (_, NONE) => default ())
+
+fun solver_class_of ctxt =
+  let fun class_of ({class, ...}: solver_info, _) = class ctxt
+  in solver_info_of no_solver_err (class_of o the) ctxt end
+
+fun solver_options_of ctxt =
+  let
+    fun all_options NONE = []
+      | all_options (SOME ({options, ...} : solver_info, opts)) =
+          opts @ options ctxt
+  in solver_info_of (K []) all_options ctxt end
+
+val setup_solver =
+  Attrib.setup @{binding smt2_solver}
+    (Scan.lift (@{keyword "="} |-- Args.name) >>
+      (Thm.declaration_attribute o K o select_solver))
+    "SMT solver configuration"
+
+
+(* options *)
+
+val oracle = Attrib.setup_config_bool @{binding smt2_oracle} (K true)
+val timeout = Attrib.setup_config_real @{binding smt2_timeout} (K 30.0)
+val random_seed = Attrib.setup_config_int @{binding smt2_random_seed} (K 1)
+val read_only_certificates = Attrib.setup_config_bool @{binding smt2_read_only_certificates} (K false)
+val verbose = Attrib.setup_config_bool @{binding smt2_verbose} (K true)
+val trace = Attrib.setup_config_bool @{binding smt2_trace} (K false)
+val trace_used_facts = Attrib.setup_config_bool @{binding smt2_trace_used_facts} (K false)
+val monomorph_limit = Attrib.setup_config_int @{binding smt2_monomorph_limit} (K 10)
+val monomorph_instances = Attrib.setup_config_int @{binding smt2_monomorph_instances} (K 500)
+val infer_triggers = Attrib.setup_config_bool @{binding smt2_infer_triggers} (K false)
+val filter_only_facts = Attrib.setup_config_bool @{binding smt2_filter_only_facts} (K false)
+val debug_files = Attrib.setup_config_string @{binding smt2_debug_files} (K "")
+
+
+(* diagnostics *)
+
+fun cond_trace flag f x = if flag then tracing ("SMT: " ^ f x) else ()
+
+fun verbose_msg ctxt = cond_trace (Config.get ctxt verbose)
+
+fun trace_msg ctxt = cond_trace (Config.get ctxt trace)
+
+
+(* tools *)
+
+fun with_timeout ctxt f x =
+  TimeLimit.timeLimit (seconds (Config.get ctxt timeout)) f x
+  handle TimeLimit.TimeOut => raise SMT2_Failure.SMT SMT2_Failure.Time_Out
+
+
+(* certificates *)
+
+(* FIXME just one data slot (record) per program unit *)
+structure Certificates = Generic_Data
+(
+  type T = Cache_IO.cache option
+  val empty = NONE
+  val extend = I
+  fun merge (s, _) = s  (* FIXME merge options!? *)
+)
+
+val get_certificates_path =
+  Option.map (Cache_IO.cache_path_of) o Certificates.get o Context.Proof
+
+fun select_certificates name context = context |> Certificates.put (
+  if name = "" then NONE
+  else
+    Path.explode name
+    |> Path.append (Thy_Load.master_directory (Context.theory_of context))
+    |> SOME o Cache_IO.unsynchronized_init)
+
+val certificates_of = Certificates.get o Context.Proof
+
+val setup_certificates =
+  Attrib.setup @{binding smt2_certificates}
+    (Scan.lift (@{keyword "="} |-- Args.name) >>
+      (Thm.declaration_attribute o K o select_certificates))
+    "SMT certificates configuration"
+
+
+(* setup *)
+
+val _ = Theory.setup (
+  setup_solver #>
+  setup_certificates)
+
+fun print_setup ctxt =
+  let
+    fun string_of_bool b = if b then "true" else "false"
+
+    val names = available_solvers_of ctxt
+    val ns = if null names then ["(none)"] else sort_strings names
+    val n = the_default "(none)" (solver_name_of ctxt)
+    val opts = solver_options_of ctxt
+    
+    val t = string_of_real (Config.get ctxt timeout)
+
+    val certs_filename =
+      (case get_certificates_path ctxt of
+        SOME path => Path.print path
+      | NONE => "(disabled)")
+  in
+    Pretty.writeln (Pretty.big_list "SMT setup:" [
+      Pretty.str ("Current SMT solver: " ^ n),
+      Pretty.str ("Current SMT solver options: " ^ space_implode " " opts),
+      Pretty.str_list "Available SMT solvers: "  "" ns,
+      Pretty.str ("Current timeout: " ^ t ^ " seconds"),
+      Pretty.str ("With proofs: " ^
+        string_of_bool (not (Config.get ctxt oracle))),
+      Pretty.str ("Certificates cache: " ^ certs_filename),
+      Pretty.str ("Fixed certificates: " ^
+        string_of_bool (Config.get ctxt read_only_certificates))])
+  end
+
+val _ =
+  Outer_Syntax.improper_command @{command_spec "smt2_status"}
+    "show the available SMT solvers, the currently selected SMT solver, \
+    \and the values of SMT configuration options"
+    (Scan.succeed (Toplevel.keep (print_setup o Toplevel.context_of)))
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/SMT2/smt2_datatypes.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -0,0 +1,133 @@
+(*  Title:      HOL/Tools/SMT2/smt2_datatypes.ML
+    Author:     Sascha Boehme, TU Muenchen
+
+Collector functions for common type declarations and their representation
+as algebraic datatypes.
+*)
+
+signature SMT2_DATATYPES =
+sig
+  val add_decls: typ ->
+    (typ * (term * term list) list) list list * Proof.context ->
+    (typ * (term * term list) list) list list * Proof.context
+end
+
+structure SMT2_Datatypes: SMT2_DATATYPES =
+struct
+
+val lhs_head_of = Term.head_of o fst o Logic.dest_equals o Thm.prop_of
+
+fun mk_selectors T Ts ctxt =
+  let
+    val (sels, ctxt') =
+      Variable.variant_fixes (replicate (length Ts) "select") ctxt
+  in (map2 (fn n => fn U => Free (n, T --> U)) sels Ts, ctxt') end
+
+
+(* datatype declarations *)
+
+fun get_datatype_decl ({descr, ...} : Datatype.info) n Ts ctxt =
+  let
+    fun get_vars (_, (m, vs, _)) = if m = n then SOME vs else NONE
+    val vars = the (get_first get_vars descr) ~~ Ts
+    val lookup_var = the o AList.lookup (op =) vars
+
+    fun typ_of (dt as Datatype.DtTFree _) = lookup_var dt
+      | typ_of (Datatype.DtType (m, dts)) = Type (m, map typ_of dts)
+      | typ_of (Datatype.DtRec i) =
+          the (AList.lookup (op =) descr i)
+          |> (fn (m, dts, _) => Type (m, map typ_of dts))
+
+    fun mk_constr T (m, dts) ctxt =
+      let
+        val Ts = map typ_of dts
+        val constr = Const (m, Ts ---> T)
+        val (selects, ctxt') = mk_selectors T Ts ctxt
+      in ((constr, selects), ctxt') end
+
+    fun mk_decl (i, (_, _, constrs)) ctxt =
+      let
+        val T = typ_of (Datatype.DtRec i)
+        val (css, ctxt') = fold_map (mk_constr T) constrs ctxt
+      in ((T, css), ctxt') end
+
+  in fold_map mk_decl descr ctxt end
+
+
+(* record declarations *)
+
+val record_name_of = Long_Name.implode o fst o split_last o Long_Name.explode
+
+fun get_record_decl ({ext_def, ...} : Record.info) T ctxt =
+  let
+    val (con, _) = Term.dest_Const (lhs_head_of ext_def)
+    val (fields, more) = Record.get_extT_fields (Proof_Context.theory_of ctxt) T
+    val fieldTs = map snd fields @ [snd more]
+
+    val constr = Const (con, fieldTs ---> T)
+    val (selects, ctxt') = mk_selectors T fieldTs ctxt
+  in ((T, [(constr, selects)]), ctxt') end
+
+
+(* typedef declarations *)
+
+fun get_typedef_decl (info : Typedef.info) T Ts =
+  let
+    val ({Abs_name, Rep_name, abs_type, rep_type, ...}, _) = info
+
+    val env = snd (Term.dest_Type abs_type) ~~ Ts
+    val instT = Term.map_atyps (perhaps (AList.lookup (op =) env))
+
+    val constr = Const (Abs_name, instT (rep_type --> abs_type))
+    val select = Const (Rep_name, instT (abs_type --> rep_type))
+  in (T, [(constr, [select])]) end
+
+
+(* collection of declarations *)
+
+fun declared declss T = exists (exists (equal T o fst)) declss
+fun declared' dss T = exists (exists (equal T o fst) o snd) dss
+
+fun get_decls T n Ts ctxt =
+  let val thy = Proof_Context.theory_of ctxt
+  in
+    (case Datatype.get_info thy n of
+      SOME info => get_datatype_decl info n Ts ctxt
+    | NONE =>
+        (case Record.get_info thy (record_name_of n) of
+          SOME info => get_record_decl info T ctxt |>> single
+        | NONE =>
+            (case Typedef.get_info ctxt n of
+              [] => ([], ctxt)
+            | info :: _ => ([get_typedef_decl info T Ts], ctxt))))
+  end
+
+fun add_decls T (declss, ctxt) =
+  let
+    fun depends Ts ds = exists (member (op =) (map fst ds)) Ts
+
+    fun add (TFree _) = I
+      | add (TVar _) = I
+      | add (T as Type (@{type_name fun}, _)) =
+          fold add (Term.body_type T :: Term.binder_types T)
+      | add @{typ bool} = I
+      | add (T as Type (n, Ts)) = (fn (dss, ctxt1) =>
+          if declared declss T orelse declared' dss T then (dss, ctxt1)
+          else if SMT2_Builtin.is_builtin_typ_ext ctxt1 T then (dss, ctxt1)
+          else
+            (case get_decls T n Ts ctxt1 of
+              ([], _) => (dss, ctxt1)
+            | (ds, ctxt2) =>
+                let
+                  val constrTs =
+                    maps (map (snd o Term.dest_Const o fst) o snd) ds
+                  val Us = fold (union (op =) o Term.binder_types) constrTs []
+
+                  fun ins [] = [(Us, ds)]
+                    | ins ((Uds as (Us', _)) :: Udss) =
+                        if depends Us' ds then (Us, ds) :: Uds :: Udss
+                        else Uds :: ins Udss
+            in fold add Us (ins dss, ctxt2) end))
+  in add T ([], ctxt) |>> append declss o map snd end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/SMT2/smt2_failure.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -0,0 +1,61 @@
+(*  Title:      HOL/Tools/SMT2/smt2_failure.ML
+    Author:     Sascha Boehme, TU Muenchen
+
+Failures and exception of SMT.
+*)
+
+signature SMT2_FAILURE =
+sig
+  type counterexample = {
+    is_real_cex: bool,
+    free_constraints: term list,
+    const_defs: term list}
+  datatype failure =
+    Counterexample of counterexample |
+    Time_Out |
+    Out_Of_Memory |
+    Abnormal_Termination of int |
+    Other_Failure of string
+  val pretty_counterexample: Proof.context -> counterexample -> Pretty.T
+  val string_of_failure: Proof.context -> failure -> string
+  exception SMT of failure
+end
+
+structure SMT2_Failure: SMT2_FAILURE =
+struct
+
+type counterexample = {
+  is_real_cex: bool,
+  free_constraints: term list,
+  const_defs: term list}
+
+datatype failure =
+  Counterexample of counterexample |
+  Time_Out |
+  Out_Of_Memory |
+  Abnormal_Termination of int |
+  Other_Failure of string
+
+fun pretty_counterexample ctxt {is_real_cex, free_constraints, const_defs} =
+  let
+    val msg =
+      if is_real_cex then "Counterexample found (possibly spurious)"
+      else "Potential counterexample found"
+  in
+    if null free_constraints andalso null const_defs then Pretty.str msg
+    else
+      Pretty.big_list (msg ^ ":")
+        (map (Syntax.pretty_term ctxt) (free_constraints @ const_defs))
+  end
+
+fun string_of_failure ctxt (Counterexample cex) =
+      Pretty.string_of (pretty_counterexample ctxt cex)
+  | string_of_failure _ Time_Out = "Timed out"
+  | string_of_failure _ Out_Of_Memory = "Ran out of memory"
+  | string_of_failure _ (Abnormal_Termination err) =
+      "Solver terminated abnormally with error code " ^ string_of_int err
+  | string_of_failure _ (Other_Failure msg) = msg
+
+exception SMT of failure
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/SMT2/smt2_normalize.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -0,0 +1,577 @@
+(*  Title:      HOL/Tools/SMT2/smt2_normalize.ML
+    Author:     Sascha Boehme, TU Muenchen
+
+Normalization steps on theorems required by SMT solvers.
+*)
+
+signature SMT2_NORMALIZE =
+sig
+  val drop_fact_warning: Proof.context -> thm -> unit
+  val atomize_conv: Proof.context -> conv
+  type extra_norm = Proof.context -> thm list * thm list -> thm list * thm list
+  val add_extra_norm: SMT2_Util.class * extra_norm -> Context.generic -> Context.generic
+  val normalize: Proof.context -> (int option * thm) list -> (int * thm) list
+end
+
+structure SMT2_Normalize: SMT2_NORMALIZE =
+struct
+
+fun drop_fact_warning ctxt =
+  SMT2_Config.verbose_msg ctxt (prefix "Warning: dropping assumption: " o
+    Display.string_of_thm ctxt)
+
+
+(* general theorem normalizations *)
+
+(** instantiate elimination rules **)
+ 
+local
+  val (cpfalse, cfalse) = `SMT2_Util.mk_cprop (Thm.cterm_of @{theory} @{const False})
+
+  fun inst f ct thm =
+    let val cv = f (Drule.strip_imp_concl (Thm.cprop_of thm))
+    in Thm.instantiate ([], [(cv, ct)]) thm end
+in
+
+fun instantiate_elim thm =
+  (case Thm.concl_of thm of
+    @{const Trueprop} $ Var (_, @{typ bool}) => inst Thm.dest_arg cfalse thm
+  | Var _ => inst I cpfalse thm
+  | _ => thm)
+
+end
+
+
+(** normalize definitions **)
+
+fun norm_def thm =
+  (case Thm.prop_of thm of
+    @{const Trueprop} $ (Const (@{const_name HOL.eq}, _) $ _ $ Abs _) =>
+      norm_def (thm RS @{thm fun_cong})
+  | Const (@{const_name "=="}, _) $ _ $ Abs _ => norm_def (thm RS @{thm meta_eq_to_obj_eq})
+  | _ => thm)
+
+
+(** atomization **)
+
+fun atomize_conv ctxt ct =
+  (case Thm.term_of ct of
+    @{const "==>"} $ _ $ _ =>
+      Conv.binop_conv (atomize_conv ctxt) then_conv Conv.rewr_conv @{thm atomize_imp}
+  | Const (@{const_name "=="}, _) $ _ $ _ =>
+      Conv.binop_conv (atomize_conv ctxt) then_conv Conv.rewr_conv @{thm atomize_eq}
+  | Const (@{const_name all}, _) $ Abs _ =>
+      Conv.binder_conv (atomize_conv o snd) ctxt then_conv Conv.rewr_conv @{thm atomize_all}
+  | _ => Conv.all_conv) ct
+
+val setup_atomize =
+  fold SMT2_Builtin.add_builtin_fun_ext'' [@{const_name "==>"}, @{const_name "=="},
+    @{const_name all}, @{const_name Trueprop}]
+
+
+(** unfold special quantifiers **)
+
+local
+  val special_quants = [
+    (@{const_name Ex1}, @{thm Ex1_def_raw}),
+    (@{const_name Ball}, @{thm Ball_def_raw}),
+    (@{const_name Bex}, @{thm Bex_def_raw})]
+  
+  fun special_quant (Const (n, _)) = AList.lookup (op =) special_quants n
+    | special_quant _ = NONE
+
+  fun special_quant_conv _ ct =
+    (case special_quant (Thm.term_of ct) of
+      SOME thm => Conv.rewr_conv thm
+    | NONE => Conv.all_conv) ct
+in
+
+fun unfold_special_quants_conv ctxt =
+  SMT2_Util.if_exists_conv (is_some o special_quant) (Conv.top_conv special_quant_conv ctxt)
+
+val setup_unfolded_quants = fold (SMT2_Builtin.add_builtin_fun_ext'' o fst) special_quants
+
+end
+
+
+(** trigger inference **)
+
+local
+  (*** check trigger syntax ***)
+
+  fun dest_trigger (Const (@{const_name pat}, _) $ _) = SOME true
+    | dest_trigger (Const (@{const_name nopat}, _) $ _) = SOME false
+    | dest_trigger _ = NONE
+
+  fun eq_list [] = false
+    | eq_list (b :: bs) = forall (equal b) bs
+
+  fun proper_trigger t =
+    t
+    |> these o try HOLogic.dest_list
+    |> map (map_filter dest_trigger o these o try HOLogic.dest_list)
+    |> (fn [] => false | bss => forall eq_list bss)
+
+  fun proper_quant inside f t =
+    (case t of
+      Const (@{const_name All}, _) $ Abs (_, _, u) => proper_quant true f u
+    | Const (@{const_name Ex}, _) $ Abs (_, _, u) => proper_quant true f u
+    | @{const trigger} $ p $ u =>
+        (if inside then f p else false) andalso proper_quant false f u
+    | Abs (_, _, u) => proper_quant false f u
+    | u1 $ u2 => proper_quant false f u1 andalso proper_quant false f u2
+    | _ => true)
+
+  fun check_trigger_error ctxt t =
+    error ("SMT triggers must only occur under quantifier and multipatterns " ^
+      "must have the same kind: " ^ Syntax.string_of_term ctxt t)
+
+  fun check_trigger_conv ctxt ct =
+    if proper_quant false proper_trigger (SMT2_Util.term_of ct) then Conv.all_conv ct
+    else check_trigger_error ctxt (Thm.term_of ct)
+
+
+  (*** infer simple triggers ***)
+
+  fun dest_cond_eq ct =
+    (case Thm.term_of ct of
+      Const (@{const_name HOL.eq}, _) $ _ $ _ => Thm.dest_binop ct
+    | @{const HOL.implies} $ _ $ _ => dest_cond_eq (Thm.dest_arg ct)
+    | _ => raise CTERM ("no equation", [ct]))
+
+  fun get_constrs thy (Type (n, _)) = these (Datatype.get_constrs thy n)
+    | get_constrs _ _ = []
+
+  fun is_constr thy (n, T) =
+    let fun match (m, U) = m = n andalso Sign.typ_instance thy (T, U)
+    in can (the o find_first match o get_constrs thy o Term.body_type) T end
+
+  fun is_constr_pat thy t =
+    (case Term.strip_comb t of
+      (Free _, []) => true
+    | (Const c, ts) => is_constr thy c andalso forall (is_constr_pat thy) ts
+    | _ => false)
+
+  fun is_simp_lhs ctxt t =
+    (case Term.strip_comb t of
+      (Const c, ts as _ :: _) =>
+        not (SMT2_Builtin.is_builtin_fun_ext ctxt c ts) andalso
+        forall (is_constr_pat (Proof_Context.theory_of ctxt)) ts
+    | _ => false)
+
+  fun has_all_vars vs t =
+    subset (op aconv) (vs, map Free (Term.add_frees t []))
+
+  fun minimal_pats vs ct =
+    if has_all_vars vs (Thm.term_of ct) then
+      (case Thm.term_of ct of
+        _ $ _ =>
+          (case pairself (minimal_pats vs) (Thm.dest_comb ct) of
+            ([], []) => [[ct]]
+          | (ctss, ctss') => union (eq_set (op aconvc)) ctss ctss')
+      | _ => [])
+    else []
+
+  fun proper_mpat _ _ _ [] = false
+    | proper_mpat thy gen u cts =
+        let
+          val tps = (op ~~) (`gen (map Thm.term_of cts))
+          fun some_match u = tps |> exists (fn (t', t) =>
+            Pattern.matches thy (t', u) andalso not (t aconv u))
+        in not (Term.exists_subterm some_match u) end
+
+  val pat = SMT2_Util.mk_const_pat @{theory} @{const_name SMT2.pat} SMT2_Util.destT1
+  fun mk_pat ct = Thm.apply (SMT2_Util.instT' ct pat) ct
+
+  fun mk_clist T = pairself (Thm.cterm_of @{theory}) (HOLogic.cons_const T, HOLogic.nil_const T)
+  fun mk_list (ccons, cnil) f cts = fold_rev (Thm.mk_binop ccons o f) cts cnil
+  val mk_pat_list = mk_list (mk_clist @{typ SMT2.pattern})
+  val mk_mpat_list = mk_list (mk_clist @{typ "SMT2.pattern list"})  
+  fun mk_trigger ctss = mk_mpat_list (mk_pat_list mk_pat) ctss
+
+  val trigger_eq = mk_meta_eq @{lemma "p = SMT2.trigger t p" by (simp add: trigger_def)}
+
+  fun insert_trigger_conv [] ct = Conv.all_conv ct
+    | insert_trigger_conv ctss ct =
+        let val (ctr, cp) = Thm.dest_binop (Thm.rhs_of trigger_eq) ||> rpair ct
+        in Thm.instantiate ([], [cp, (ctr, mk_trigger ctss)]) trigger_eq end
+
+  fun infer_trigger_eq_conv outer_ctxt (ctxt, cvs) ct =
+    let
+      val (lhs, rhs) = dest_cond_eq ct
+
+      val vs = map Thm.term_of cvs
+      val thy = Proof_Context.theory_of ctxt
+
+      fun get_mpats ct =
+        if is_simp_lhs ctxt (Thm.term_of ct) then minimal_pats vs ct
+        else []
+      val gen = Variable.export_terms ctxt outer_ctxt
+      val filter_mpats = filter (proper_mpat thy gen (Thm.term_of rhs))
+
+    in insert_trigger_conv (filter_mpats (get_mpats lhs)) ct end
+
+  fun has_trigger (@{const SMT2.trigger} $ _ $ _) = true
+    | has_trigger _ = false
+
+  fun try_trigger_conv cv ct =
+    if SMT2_Util.under_quant has_trigger (SMT2_Util.term_of ct) then Conv.all_conv ct
+    else Conv.try_conv cv ct
+
+  fun infer_trigger_conv ctxt =
+    if Config.get ctxt SMT2_Config.infer_triggers then
+      try_trigger_conv (SMT2_Util.under_quant_conv (infer_trigger_eq_conv ctxt) ctxt)
+    else Conv.all_conv
+in
+
+fun trigger_conv ctxt =
+  SMT2_Util.prop_conv (check_trigger_conv ctxt then_conv infer_trigger_conv ctxt)
+
+val setup_trigger =
+  fold SMT2_Builtin.add_builtin_fun_ext''
+    [@{const_name SMT2.pat}, @{const_name SMT2.nopat}, @{const_name SMT2.trigger}]
+
+end
+
+
+(** adding quantifier weights **)
+
+local
+  (*** check weight syntax ***)
+
+  val has_no_weight =
+    not o Term.exists_subterm (fn @{const SMT2.weight} => true | _ => false)
+
+  fun is_weight (@{const SMT2.weight} $ w $ t) =
+        (case try HOLogic.dest_number w of
+          SOME (_, i) => i >= 0 andalso has_no_weight t
+        | _ => false)
+    | is_weight t = has_no_weight t
+
+  fun proper_trigger (@{const SMT2.trigger} $ _ $ t) = is_weight t
+    | proper_trigger t = is_weight t 
+
+  fun check_weight_error ctxt t =
+    error ("SMT weight must be a non-negative number and must only occur " ^
+      "under the top-most quantifier and an optional trigger: " ^
+      Syntax.string_of_term ctxt t)
+
+  fun check_weight_conv ctxt ct =
+    if SMT2_Util.under_quant proper_trigger (SMT2_Util.term_of ct) then Conv.all_conv ct
+    else check_weight_error ctxt (Thm.term_of ct)
+
+
+  (*** insertion of weights ***)
+
+  fun under_trigger_conv cv ct =
+    (case Thm.term_of ct of
+      @{const SMT2.trigger} $ _ $ _ => Conv.arg_conv cv
+    | _ => cv) ct
+
+  val weight_eq = mk_meta_eq @{lemma "p = SMT2.weight i p" by (simp add: weight_def)}
+  fun mk_weight_eq w =
+    let val cv = Thm.dest_arg1 (Thm.rhs_of weight_eq)
+    in Thm.instantiate ([], [(cv, Numeral.mk_cnumber @{ctyp int} w)]) weight_eq end
+
+  fun add_weight_conv NONE _ = Conv.all_conv
+    | add_weight_conv (SOME weight) ctxt =
+        let val cv = Conv.rewr_conv (mk_weight_eq weight)
+        in SMT2_Util.under_quant_conv (K (under_trigger_conv cv)) ctxt end
+in
+
+fun weight_conv weight ctxt = 
+  SMT2_Util.prop_conv (check_weight_conv ctxt then_conv add_weight_conv weight ctxt)
+
+val setup_weight = SMT2_Builtin.add_builtin_fun_ext'' @{const_name SMT2.weight}
+
+end
+
+
+(** combined general normalizations **)
+
+fun gen_normalize1_conv ctxt weight =
+  atomize_conv ctxt then_conv
+  unfold_special_quants_conv ctxt then_conv
+  Thm.beta_conversion true then_conv
+  trigger_conv ctxt then_conv
+  weight_conv weight ctxt
+
+fun gen_normalize1 ctxt weight =
+  instantiate_elim #>
+  norm_def #>
+  Conv.fconv_rule (Thm.beta_conversion true then_conv Thm.eta_conversion) #>
+  Drule.forall_intr_vars #>
+  Conv.fconv_rule (gen_normalize1_conv ctxt weight) #>
+  (* Z3 4.3.1 silently normalizes "P --> Q --> R" to "P & Q --> R" *)
+  Raw_Simplifier.rewrite_rule ctxt @{thms HOL.imp_conjL[symmetric, THEN eq_reflection]}
+
+fun gen_norm1_safe ctxt (i, (weight, thm)) =
+  (case try (gen_normalize1 ctxt weight) thm of
+    SOME thm' => SOME (i, thm')
+  | NONE => (drop_fact_warning ctxt thm; NONE))
+
+fun gen_normalize ctxt iwthms = map_filter (gen_norm1_safe ctxt) iwthms
+
+
+
+(* unfolding of definitions and theory-specific rewritings *)
+
+fun expand_head_conv cv ct =
+  (case Thm.term_of ct of
+    _ $ _ =>
+      Conv.fun_conv (expand_head_conv cv) then_conv
+      Conv.try_conv (Thm.beta_conversion false)
+  | _ => cv) ct
+
+
+(** rewrite bool case expressions as if expressions **)
+
+local
+  fun is_case_bool (Const (@{const_name "bool.case_bool"}, _)) = true
+    | is_case_bool _ = false
+
+  fun unfold_conv _ =
+    SMT2_Util.if_true_conv (is_case_bool o Term.head_of)
+      (expand_head_conv (Conv.rewr_conv @{thm case_bool_if}))
+in
+
+fun rewrite_case_bool_conv ctxt =
+  SMT2_Util.if_exists_conv is_case_bool (Conv.top_conv unfold_conv ctxt)
+
+val setup_case_bool = SMT2_Builtin.add_builtin_fun_ext'' @{const_name "bool.case_bool"}
+
+end
+
+
+(** unfold abs, min and max **)
+
+local
+  val defs = [
+    (@{const_name min}, @{thm min_def_raw}),
+    (@{const_name max}, @{thm max_def_raw}),
+    (@{const_name abs}, @{thm abs_if_raw})]
+
+  fun abs_min_max ctxt (Const (n, Type (@{type_name fun}, [T, _]))) =
+        (case AList.lookup (op =) defs n of
+          NONE => NONE
+        | SOME thm => if SMT2_Builtin.is_builtin_typ_ext ctxt T then SOME thm else NONE)
+    | abs_min_max _ _ = NONE
+
+  fun unfold_amm_conv ctxt ct =
+    (case abs_min_max ctxt (Term.head_of (Thm.term_of ct)) of
+      SOME thm => expand_head_conv (Conv.rewr_conv thm)
+    | NONE => Conv.all_conv) ct
+in
+
+fun unfold_abs_min_max_conv ctxt =
+  SMT2_Util.if_exists_conv (is_some o abs_min_max ctxt) (Conv.top_conv unfold_amm_conv ctxt)
+  
+val setup_abs_min_max = fold (SMT2_Builtin.add_builtin_fun_ext'' o fst) defs
+
+end
+
+
+(** embedding of standard natural number operations into integer operations **)
+
+local
+  val nat_embedding = @{thms nat_int' int_nat_nneg int_nat_neg}
+
+  val simple_nat_ops = [
+    @{const less (nat)}, @{const less_eq (nat)},
+    @{const Suc}, @{const plus (nat)}, @{const minus (nat)}]
+
+  val mult_nat_ops =
+    [@{const times (nat)}, @{const div (nat)}, @{const mod (nat)}]
+
+  val nat_ops = simple_nat_ops @ mult_nat_ops
+
+  val nat_consts = nat_ops @ [@{const numeral (nat)},
+    @{const zero_class.zero (nat)}, @{const one_class.one (nat)}]
+
+  val nat_int_coercions = [@{const of_nat (int)}, @{const nat}]
+
+  val builtin_nat_ops = nat_int_coercions @ simple_nat_ops
+
+  val is_nat_const = member (op aconv) nat_consts
+
+  fun is_nat_const' @{const of_nat (int)} = true
+    | is_nat_const' t = is_nat_const t
+
+  val expands = map mk_meta_eq @{thms nat_zero_as_int nat_one_as_int nat_numeral_as_int
+    nat_less_as_int nat_leq_as_int Suc_as_int nat_plus_as_int nat_minus_as_int nat_times_as_int
+    nat_div_as_int nat_mod_as_int}
+
+  val ints = map mk_meta_eq @{thms int_0 int_1 int_Suc int_plus int_minus int_mult zdiv_int
+    zmod_int}
+  val int_if = mk_meta_eq @{lemma "int (if P then n else m) = (if P then int n else int m)" by simp}
+
+  fun mk_number_eq ctxt i lhs =
+    let
+      val eq = SMT2_Util.mk_cequals lhs (Numeral.mk_cnumber @{ctyp int} i)
+      val ctxt' = put_simpset HOL_ss ctxt addsimps @{thms Int.int_numeral}
+      val tac = HEADGOAL (Simplifier.simp_tac ctxt')
+    in Goal.norm_result ctxt (Goal.prove_internal ctxt [] eq (K tac)) end
+
+  fun ite_conv cv1 cv2 =
+    Conv.combination_conv (Conv.combination_conv (Conv.arg_conv cv1) cv2) cv2
+
+  fun int_conv ctxt ct =
+    (case Thm.term_of ct of
+      @{const of_nat (int)} $ (n as (@{const numeral (nat)} $ _)) =>
+        Conv.rewr_conv (mk_number_eq ctxt (snd (HOLogic.dest_number n)) ct)
+    | @{const of_nat (int)} $ _ =>
+        (Conv.rewrs_conv ints then_conv Conv.sub_conv ints_conv ctxt) else_conv
+        (Conv.rewr_conv int_if then_conv
+          ite_conv (nat_conv ctxt) (int_conv ctxt)) else_conv
+        Conv.sub_conv (Conv.top_sweep_conv nat_conv) ctxt
+    | _ => Conv.no_conv) ct
+
+  and ints_conv ctxt = Conv.top_sweep_conv int_conv ctxt
+
+  and expand_conv ctxt =
+    SMT2_Util.if_conv (is_nat_const o Term.head_of)
+      (expand_head_conv (Conv.rewrs_conv expands) then_conv ints_conv ctxt) (int_conv ctxt)
+
+  and nat_conv ctxt = SMT2_Util.if_exists_conv is_nat_const' (Conv.top_sweep_conv expand_conv ctxt)
+
+  val uses_nat_int = Term.exists_subterm (member (op aconv) nat_int_coercions)
+in
+
+val nat_as_int_conv = nat_conv
+
+fun add_nat_embedding thms =
+  if exists (uses_nat_int o Thm.prop_of) thms then (thms, nat_embedding) else (thms, [])
+
+val setup_nat_as_int =
+  SMT2_Builtin.add_builtin_typ_ext (@{typ nat}, K true) #>
+  fold (SMT2_Builtin.add_builtin_fun_ext' o Term.dest_Const) builtin_nat_ops
+
+end
+
+
+(** normalize numerals **)
+
+local
+  (*
+    rewrite Numeral1 into 1
+    rewrite - 0 into 0
+  *)
+
+  fun is_irregular_number (Const (@{const_name numeral}, _) $ Const (@{const_name num.One}, _)) =
+        true
+    | is_irregular_number (Const (@{const_name uminus}, _) $ Const (@{const_name Groups.zero}, _)) =
+        true
+    | is_irregular_number _ =
+        false;
+
+  fun is_strange_number ctxt t = is_irregular_number t andalso SMT2_Builtin.is_builtin_num ctxt t;
+
+  val proper_num_ss =
+    simpset_of (put_simpset HOL_ss @{context}
+      addsimps @{thms Num.numeral_One minus_zero})
+
+  fun norm_num_conv ctxt =
+    SMT2_Util.if_conv (is_strange_number ctxt) (Simplifier.rewrite (put_simpset proper_num_ss ctxt))
+      Conv.no_conv
+in
+
+fun normalize_numerals_conv ctxt =
+  SMT2_Util.if_exists_conv (is_strange_number ctxt) (Conv.top_sweep_conv norm_num_conv ctxt)
+
+end
+
+
+(** combined unfoldings and rewritings **)
+
+fun unfold_conv ctxt =
+  rewrite_case_bool_conv ctxt then_conv
+  unfold_abs_min_max_conv ctxt then_conv
+  nat_as_int_conv ctxt then_conv
+  Thm.beta_conversion true
+
+fun unfold1 ctxt = map (apsnd (Conv.fconv_rule (unfold_conv ctxt)))
+
+fun burrow_ids f ithms =
+  let
+    val (is, thms) = split_list ithms
+    val (thms', extra_thms) = f thms
+  in (is ~~ thms') @ map (pair ~1) extra_thms end
+
+fun unfold2 ctxt ithms =
+  ithms
+  |> map (apsnd (Conv.fconv_rule (normalize_numerals_conv ctxt)))
+  |> burrow_ids add_nat_embedding
+
+
+
+(* overall normalization *)
+
+type extra_norm = Proof.context -> thm list * thm list -> thm list * thm list
+
+structure Extra_Norms = Generic_Data
+(
+  type T = extra_norm SMT2_Util.dict
+  val empty = []
+  val extend = I
+  fun merge data = SMT2_Util.dict_merge fst data
+)
+
+fun add_extra_norm (cs, norm) = Extra_Norms.map (SMT2_Util.dict_update (cs, norm))
+
+fun apply_extra_norms ctxt ithms =
+  let
+    val cs = SMT2_Config.solver_class_of ctxt
+    val es = SMT2_Util.dict_lookup (Extra_Norms.get (Context.Proof ctxt)) cs
+  in burrow_ids (fold (fn e => e ctxt) es o rpair []) ithms end
+
+local
+  val ignored = member (op =) [@{const_name All}, @{const_name Ex},
+    @{const_name Let}, @{const_name If}, @{const_name HOL.eq}]
+
+  val schematic_consts_of =
+    let
+      fun collect (@{const SMT2.trigger} $ p $ t) =
+            collect_trigger p #> collect t
+        | collect (t $ u) = collect t #> collect u
+        | collect (Abs (_, _, t)) = collect t
+        | collect (t as Const (n, _)) = 
+            if not (ignored n) then Monomorph.add_schematic_consts_of t else I
+        | collect _ = I
+      and collect_trigger t =
+        let val dest = these o try HOLogic.dest_list 
+        in fold (fold collect_pat o dest) (dest t) end
+      and collect_pat (Const (@{const_name SMT2.pat}, _) $ t) = collect t
+        | collect_pat (Const (@{const_name SMT2.nopat}, _) $ t) = collect t
+        | collect_pat _ = I
+    in (fn t => collect t Symtab.empty) end
+in
+
+fun monomorph ctxt xthms =
+  let val (xs, thms) = split_list xthms
+  in
+    map (pair 1) thms
+    |> Monomorph.monomorph schematic_consts_of ctxt
+    |> maps (uncurry (map o pair)) o map2 pair xs o map (map snd)
+  end
+
+end
+
+fun normalize ctxt wthms =
+  wthms
+  |> map_index I
+  |> gen_normalize ctxt
+  |> unfold1 ctxt
+  |> monomorph ctxt
+  |> unfold2 ctxt
+  |> apply_extra_norms ctxt
+
+val _ = Theory.setup (Context.theory_map (
+  setup_atomize #>
+  setup_unfolded_quants #>
+  setup_trigger #>
+  setup_weight #>
+  setup_case_bool #>
+  setup_abs_min_max #>
+  setup_nat_as_int))
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/SMT2/smt2_real.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -0,0 +1,121 @@
+(*  Title:      HOL/Tools/SMT2/smt2_real.ML
+    Author:     Sascha Boehme, TU Muenchen
+
+SMT setup for reals.
+*)
+
+structure SMT2_Real: sig end =
+struct
+
+
+(* SMT-LIB logic *)
+
+fun smtlib_logic ts =
+  if exists (Term.exists_type (Term.exists_subtype (equal @{typ real}))) ts
+  then SOME "AUFLIRA"
+  else NONE
+
+
+(* SMT-LIB and Z3 built-ins *)
+
+local
+  fun real_num _ i = SOME (string_of_int i ^ ".0")
+
+  fun is_linear [t] = SMT2_Util.is_number t
+    | is_linear [t, u] = SMT2_Util.is_number t orelse SMT2_Util.is_number u
+    | is_linear _ = false
+
+  fun mk_times ts = Term.list_comb (@{const times (real)}, ts)
+
+  fun times _ _ ts = if is_linear ts then SOME ("*", 2, ts, mk_times) else NONE
+in
+
+val setup_builtins =
+  SMT2_Builtin.add_builtin_typ SMTLIB2_Interface.smtlib2C
+    (@{typ real}, K (SOME "Real"), real_num) #>
+  fold (SMT2_Builtin.add_builtin_fun' SMTLIB2_Interface.smtlib2C) [
+    (@{const less (real)}, "<"),
+    (@{const less_eq (real)}, "<="),
+    (@{const uminus (real)}, "~"),
+    (@{const plus (real)}, "+"),
+    (@{const minus (real)}, "-") ] #>
+  SMT2_Builtin.add_builtin_fun SMTLIB2_Interface.smtlib2C
+    (Term.dest_Const @{const times (real)}, times) #>
+  SMT2_Builtin.add_builtin_fun' Z3_New_Interface.smtlib2_z3C
+    (@{const times (real)}, "*") #>
+  SMT2_Builtin.add_builtin_fun' Z3_New_Interface.smtlib2_z3C
+    (@{const divide (real)}, "/")
+
+end
+
+
+(* Z3 constructors *)
+
+local
+  fun z3_mk_builtin_typ (Z3_New_Interface.Sym ("Real", _)) = SOME @{typ real}
+    | z3_mk_builtin_typ (Z3_New_Interface.Sym ("real", _)) = SOME @{typ real}
+        (*FIXME: delete*)
+    | z3_mk_builtin_typ _ = NONE
+
+  fun z3_mk_builtin_num _ i T =
+    if T = @{typ real} then SOME (Numeral.mk_cnumber @{ctyp real} i)
+    else NONE
+
+  fun mk_nary _ cu [] = cu
+    | mk_nary ct _ cts = uncurry (fold_rev (Thm.mk_binop ct)) (split_last cts)
+
+  val mk_uminus = Thm.apply (Thm.cterm_of @{theory} @{const uminus (real)})
+  val add = Thm.cterm_of @{theory} @{const plus (real)}
+  val real0 = Numeral.mk_cnumber @{ctyp real} 0
+  val mk_sub = Thm.mk_binop (Thm.cterm_of @{theory} @{const minus (real)})
+  val mk_mul = Thm.mk_binop (Thm.cterm_of @{theory} @{const times (real)})
+  val mk_div = Thm.mk_binop (Thm.cterm_of @{theory} @{const divide (real)})
+  val mk_lt = Thm.mk_binop (Thm.cterm_of @{theory} @{const less (real)})
+  val mk_le = Thm.mk_binop (Thm.cterm_of @{theory} @{const less_eq (real)})
+
+  fun z3_mk_builtin_fun (Z3_New_Interface.Sym ("-", _)) [ct] = SOME (mk_uminus ct)
+    | z3_mk_builtin_fun (Z3_New_Interface.Sym ("+", _)) cts =
+        SOME (mk_nary add real0 cts)
+    | z3_mk_builtin_fun (Z3_New_Interface.Sym ("-", _)) [ct, cu] =
+        SOME (mk_sub ct cu)
+    | z3_mk_builtin_fun (Z3_New_Interface.Sym ("*", _)) [ct, cu] =
+        SOME (mk_mul ct cu)
+    | z3_mk_builtin_fun (Z3_New_Interface.Sym ("/", _)) [ct, cu] =
+        SOME (mk_div ct cu)
+    | z3_mk_builtin_fun (Z3_New_Interface.Sym ("<", _)) [ct, cu] =
+        SOME (mk_lt ct cu)
+    | z3_mk_builtin_fun (Z3_New_Interface.Sym ("<=", _)) [ct, cu] =
+        SOME (mk_le ct cu)
+    | z3_mk_builtin_fun (Z3_New_Interface.Sym (">", _)) [ct, cu] =
+        SOME (mk_lt cu ct)
+    | z3_mk_builtin_fun (Z3_New_Interface.Sym (">=", _)) [ct, cu] =
+        SOME (mk_le cu ct)
+    | z3_mk_builtin_fun _ _ = NONE
+in
+
+val z3_mk_builtins = {
+  mk_builtin_typ = z3_mk_builtin_typ,
+  mk_builtin_num = z3_mk_builtin_num,
+  mk_builtin_fun = (fn _ => fn sym => fn cts =>
+    (case try (#T o Thm.rep_cterm o hd) cts of
+      SOME @{typ real} => z3_mk_builtin_fun sym cts
+    | _ => NONE)) }
+
+end
+
+
+(* Z3 proof replay *)
+
+val real_linarith_proc = Simplifier.simproc_global @{theory} "fast_real_arith" [
+  "(m::real) < n", "(m::real) <= n", "(m::real) = n"] Lin_Arith.simproc
+
+
+(* setup *)
+
+val _ = Theory.setup (Context.theory_map (
+  SMTLIB2_Interface.add_logic (10, smtlib_logic) #>
+  setup_builtins #>
+  Z3_New_Interface.add_mk_builtins z3_mk_builtins #>
+  Z3_New_Replay_Util.add_simproc real_linarith_proc))
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/SMT2/smt2_solver.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -0,0 +1,330 @@
+(*  Title:      HOL/Tools/SMT2/smt2_solver.ML
+    Author:     Sascha Boehme, TU Muenchen
+
+SMT solvers registry and SMT tactic.
+*)
+
+signature SMT2_SOLVER =
+sig
+  (*configuration*)
+  datatype outcome = Unsat | Sat | Unknown
+  type solver_config = {
+    name: string,
+    class: Proof.context -> SMT2_Util.class,
+    avail: unit -> bool,
+    command: unit -> string list,
+    options: Proof.context -> string list,
+    default_max_relevant: int,
+    can_filter: bool,
+    outcome: string -> string list -> outcome * string list,
+    cex_parser: (Proof.context -> SMT2_Translate.replay_data -> string list ->
+      term list * term list) option,
+    replay: (Proof.context -> SMT2_Translate.replay_data -> string list ->
+      ((int * (int * thm)) list * Z3_New_Proof.z3_step list) * thm) option}
+
+  (*registry*)
+  val add_solver: solver_config -> theory -> theory
+  val default_max_relevant: Proof.context -> string -> int
+
+  (*filter*)
+  val smt2_filter: Proof.context -> thm -> ('a * (int option * thm)) list -> int -> Time.time ->
+    {outcome: SMT2_Failure.failure option, rewrite_rules: thm list, conjecture_id: int,
+     helper_ids: (int * thm) list, fact_ids: (int * ('a * thm)) list,
+     z3_proof: Z3_New_Proof.z3_step list}
+
+  (*tactic*)
+  val smt2_tac: Proof.context -> thm list -> int -> tactic
+  val smt2_tac': Proof.context -> thm list -> int -> tactic
+end
+
+structure SMT2_Solver: SMT2_SOLVER =
+struct
+
+
+(* interface to external solvers *)
+
+local
+
+fun make_command command options problem_path proof_path =
+  "(exec 2>&1;" :: map File.shell_quote (command () @ options) @
+  [File.shell_path problem_path, ")", ">", File.shell_path proof_path]
+  |> space_implode " "
+
+fun with_trace ctxt msg f x =
+  let val _ = SMT2_Config.trace_msg ctxt (fn () => msg) ()
+  in f x end
+
+fun run ctxt name mk_cmd input =
+  (case SMT2_Config.certificates_of ctxt of
+    NONE =>
+      if not (SMT2_Config.is_available ctxt name) then
+        error ("The SMT solver " ^ quote name ^ " is not installed")
+      else if Config.get ctxt SMT2_Config.debug_files = "" then
+        with_trace ctxt ("Invoking SMT solver " ^ quote name ^ " ...") (Cache_IO.run mk_cmd) input
+      else
+        let
+          val base_path = Path.explode (Config.get ctxt SMT2_Config.debug_files)
+          val in_path = Path.ext "smt2_in" base_path
+          val out_path = Path.ext "smt2_out" base_path
+        in Cache_IO.raw_run mk_cmd input in_path out_path end
+  | SOME certs =>
+      (case Cache_IO.lookup certs input of
+        (NONE, key) =>
+          if Config.get ctxt SMT2_Config.read_only_certificates then
+            error ("Bad certificate cache: missing certificate")
+          else
+            Cache_IO.run_and_cache certs key mk_cmd input
+      | (SOME output, _) =>
+          with_trace ctxt ("Using cached certificate from " ^
+            File.shell_path (Cache_IO.cache_path_of certs) ^ " ...") I output))
+
+(* Z3 returns 1 if "get-model" or "get-model" fails *)
+val normal_return_codes = [0, 1]
+
+fun run_solver ctxt name mk_cmd input =
+  let
+    fun pretty tag ls = Pretty.string_of (Pretty.big_list tag (map Pretty.str ls))
+
+    val _ = SMT2_Config.trace_msg ctxt (pretty "Problem:" o split_lines) input
+
+    val {redirected_output = res, output = err, return_code} =
+      SMT2_Config.with_timeout ctxt (run ctxt name mk_cmd) input
+    val _ = SMT2_Config.trace_msg ctxt (pretty "Solver:") err
+
+    val output = fst (take_suffix (equal "") res)
+    val _ = SMT2_Config.trace_msg ctxt (pretty "Result:") output
+
+    val _ = member (op =) normal_return_codes return_code orelse
+      raise SMT2_Failure.SMT (SMT2_Failure.Abnormal_Termination return_code)
+  in output end
+
+fun trace_assms ctxt =
+  SMT2_Config.trace_msg ctxt (Pretty.string_of o
+    Pretty.big_list "Assertions:" o map (Display.pretty_thm ctxt o snd))
+
+fun trace_replay_data ({context = ctxt, typs, terms, ...} : SMT2_Translate.replay_data) =
+  let
+    fun pretty_eq n p = Pretty.block [Pretty.str n, Pretty.str " = ", p]
+    fun p_typ (n, T) = pretty_eq n (Syntax.pretty_typ ctxt T)
+    fun p_term (n, t) = pretty_eq n (Syntax.pretty_term ctxt t)
+  in
+    SMT2_Config.trace_msg ctxt (fn () =>
+      Pretty.string_of (Pretty.big_list "Names:" [
+        Pretty.big_list "sorts:" (map p_typ (Symtab.dest typs)),
+        Pretty.big_list "functions:" (map p_term (Symtab.dest terms))])) ()
+  end
+
+in
+
+fun invoke name command ithms ctxt =
+  let
+    val options = SMT2_Config.solver_options_of ctxt
+    val comments = [space_implode " " options]
+
+    val (str, replay_data as {context = ctxt', ...}) =
+      ithms
+      |> tap (trace_assms ctxt)
+      |> SMT2_Translate.translate ctxt comments
+      ||> tap trace_replay_data
+  in (run_solver ctxt' name (make_command command options) str, replay_data) end
+
+end
+
+
+(* configuration *)
+
+datatype outcome = Unsat | Sat | Unknown
+
+type solver_config = {
+  name: string,
+  class: Proof.context -> SMT2_Util.class,
+  avail: unit -> bool,
+  command: unit -> string list,
+  options: Proof.context -> string list,
+  default_max_relevant: int,
+  can_filter: bool,
+  outcome: string -> string list -> outcome * string list,
+  cex_parser: (Proof.context -> SMT2_Translate.replay_data -> string list ->
+    term list * term list) option,
+  replay: (Proof.context -> SMT2_Translate.replay_data -> string list ->
+    ((int * (int * thm)) list * Z3_New_Proof.z3_step list) * thm) option}
+
+
+(* check well-sortedness *)
+
+val has_topsort = Term.exists_type (Term.exists_subtype (fn
+    TFree (_, []) => true
+  | TVar (_, []) => true
+  | _ => false))
+
+(* top sorts cause problems with atomization *)
+fun check_topsort ctxt thm =
+  if has_topsort (Thm.prop_of thm) then (SMT2_Normalize.drop_fact_warning ctxt thm; TrueI) else thm
+
+
+(* registry *)
+
+type solver_info = {
+  command: unit -> string list,
+  default_max_relevant: int,
+  can_filter: bool,
+  finish: Proof.context -> SMT2_Translate.replay_data -> string list ->
+    ((int * (int * thm)) list * Z3_New_Proof.z3_step list) * thm}
+
+structure Solvers = Generic_Data
+(
+  type T = solver_info Symtab.table
+  val empty = Symtab.empty
+  val extend = I
+  fun merge data = Symtab.merge (K true) data
+)
+
+local
+  fun finish outcome cex_parser replay ocl outer_ctxt
+      (replay_data as {context = ctxt, ...} : SMT2_Translate.replay_data) output =
+    (case outcome output of
+      (Unsat, ls) =>
+        if not (Config.get ctxt SMT2_Config.oracle) andalso is_some replay
+        then the replay outer_ctxt replay_data ls
+        else (([], []), ocl ())
+    | (result, ls) =>
+        let val (ts, us) = (case cex_parser of SOME f => f ctxt replay_data ls | _ => ([], []))
+        in
+          raise SMT2_Failure.SMT (SMT2_Failure.Counterexample {
+            is_real_cex = (result = Sat),
+            free_constraints = ts,
+            const_defs = us})
+        end)
+
+  val cfalse = Thm.cterm_of @{theory} (@{const Trueprop} $ @{const False})
+in
+
+fun add_solver ({name, class, avail, command, options, default_max_relevant, can_filter,
+    outcome, cex_parser, replay} : solver_config) =
+  let
+    fun solver ocl = {
+      command = command,
+      default_max_relevant = default_max_relevant,
+      can_filter = can_filter,
+      finish = finish (outcome name) cex_parser replay ocl}
+
+    val info = {name = name, class = class, avail = avail, options = options}
+  in
+    Thm.add_oracle (Binding.name name, K cfalse) #-> (fn (_, ocl) =>
+    Context.theory_map (Solvers.map (Symtab.update_new (name, solver ocl)))) #>
+    Context.theory_map (SMT2_Config.add_solver info)
+  end
+
+end
+
+fun get_info ctxt name = the (Symtab.lookup (Solvers.get (Context.Proof ctxt)) name)
+
+fun name_and_info_of ctxt =
+  let val name = SMT2_Config.solver_of ctxt
+  in (name, get_info ctxt name) end
+
+fun apply_solver ctxt wthms0 =
+  let
+    val wthms = map (apsnd (check_topsort ctxt)) wthms0
+    val (name, {command, finish, ...}) = name_and_info_of ctxt
+    val (output, replay_data) = invoke name command (SMT2_Normalize.normalize ctxt wthms) ctxt
+  in (finish ctxt replay_data output, replay_data) end
+
+val default_max_relevant = #default_max_relevant oo get_info
+val can_filter = #can_filter o snd o name_and_info_of 
+
+
+(* filter *)
+
+val no_id = ~1
+
+fun smt2_filter ctxt goal xwfacts i time_limit =
+  let
+    val ctxt =
+      ctxt
+      |> Config.put SMT2_Config.oracle false
+      |> Config.put SMT2_Config.filter_only_facts true
+      |> Config.put SMT2_Config.timeout (Time.toReal time_limit)
+
+    val ({context = ctxt, prems, concl, ...}, _) = Subgoal.focus ctxt i goal
+    fun negate ct = Thm.dest_comb ct ||> Thm.apply @{cterm Not} |-> Thm.apply
+    val cprop =
+      (case try negate (Thm.rhs_of (SMT2_Normalize.atomize_conv ctxt concl)) of
+        SOME ct => ct
+      | NONE => raise SMT2_Failure.SMT (SMT2_Failure.Other_Failure "goal is not a HOL term"))
+
+    val wconjecture = (NONE, Thm.assume cprop)
+    val wprems = map (pair NONE) prems
+    val wfacts = map snd xwfacts
+    val wthms = wconjecture :: wprems @ wfacts
+    val iwthms = map_index I wthms
+
+    val conjecture_i = 0
+    val facts_i = 1 + length wprems
+  in
+    wthms
+    |> apply_solver ctxt
+    |> (fn (((iidths0, z3_proof), _), {rewrite_rules, ...}) =>
+      let val iidths = if can_filter ctxt then iidths0 else map (apsnd (apfst (K no_id))) iwthms
+      in
+        {outcome = NONE,
+         rewrite_rules = rewrite_rules,
+         conjecture_id =
+           the_default no_id (Option.map fst (AList.lookup (op =) iidths conjecture_i)),
+         helper_ids = map_filter (try (fn (~1, idth) => idth)) iidths,
+         fact_ids = map_filter (fn (i, (id, _)) =>
+           try (apsnd (apsnd snd o nth xwfacts)) (id, i - facts_i)) iidths,
+         z3_proof = z3_proof}
+      end)
+  end
+  handle SMT2_Failure.SMT fail => {outcome = SOME fail, rewrite_rules = [], conjecture_id = no_id,
+    helper_ids = [], fact_ids = [], z3_proof = []}
+
+
+(* SMT tactic *)
+
+local
+  fun trace_assumptions ctxt wfacts iidths =
+    let val used = map_filter (try (snd o nth wfacts) o fst) iidths in
+      if Config.get ctxt SMT2_Config.trace_used_facts andalso length wfacts > 0 then
+        tracing (Pretty.string_of (Pretty.big_list "SMT used facts:"
+          (map (Display.pretty_thm ctxt) used)))
+      else ()
+    end
+
+  fun solve ctxt wfacts =
+    wfacts
+    |> apply_solver ctxt
+    |> fst
+    |>> apfst (trace_assumptions ctxt wfacts)
+    |> snd
+
+  fun str_of ctxt fail =
+    "Solver " ^ SMT2_Config.solver_of ctxt ^ ": " ^ SMT2_Failure.string_of_failure ctxt fail
+
+  fun safe_solve ctxt wfacts = SOME (solve ctxt wfacts)
+    handle
+      SMT2_Failure.SMT (fail as SMT2_Failure.Counterexample _) =>
+        (SMT2_Config.verbose_msg ctxt (str_of ctxt) fail; NONE)
+    | SMT2_Failure.SMT (fail as SMT2_Failure.Time_Out) =>
+        error ("SMT: Solver " ^ quote (SMT2_Config.solver_of ctxt) ^ ": " ^
+          SMT2_Failure.string_of_failure ctxt fail ^ " (setting the " ^
+          "configuration option " ^ quote (Config.name_of SMT2_Config.timeout) ^ " might help)")
+    | SMT2_Failure.SMT fail => error (str_of ctxt fail)
+
+  fun resolve (SOME thm) = rtac thm 1
+    | resolve NONE = no_tac
+
+  fun tac prove ctxt rules =
+    CONVERSION (SMT2_Normalize.atomize_conv ctxt)
+    THEN' rtac @{thm ccontr}
+    THEN' SUBPROOF (fn {context = ctxt, prems, ...} =>
+      resolve (prove ctxt (map (pair NONE) (rules @ prems)))) ctxt
+in
+
+val smt2_tac = tac safe_solve
+val smt2_tac' = tac (SOME oo solve)
+
+end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/SMT2/smt2_systems.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -0,0 +1,162 @@
+(*  Title:      HOL/Tools/SMT2/smt2_systems.ML
+    Author:     Sascha Boehme, TU Muenchen
+
+Setup SMT solvers.
+*)
+
+signature SMT2_SYSTEMS =
+sig
+  datatype z3_non_commercial =
+    Z3_Non_Commercial_Unknown |
+    Z3_Non_Commercial_Accepted |
+    Z3_Non_Commercial_Declined
+  val z3_non_commercial: unit -> z3_non_commercial
+  val z3_extensions: bool Config.T
+end
+
+structure SMT2_Systems: SMT2_SYSTEMS =
+struct
+
+(* helper functions *)
+
+fun make_avail name () = getenv (name ^ "_SOLVER") <> ""
+
+fun make_command name () = [getenv (name ^ "_SOLVER")]
+
+fun outcome_of unsat sat unknown solver_name line =
+  if String.isPrefix unsat line then SMT2_Solver.Unsat
+  else if String.isPrefix sat line then SMT2_Solver.Sat
+  else if String.isPrefix unknown line then SMT2_Solver.Unknown
+  else raise SMT2_Failure.SMT (SMT2_Failure.Other_Failure ("Solver " ^ quote solver_name ^
+    " failed -- enable tracing using the " ^ quote (Config.name_of SMT2_Config.trace) ^
+    " option for details"))
+
+fun on_first_line test_outcome solver_name lines =
+  let
+    val empty_line = (fn "" => true | _ => false)
+    val split_first = (fn [] => ("", []) | l :: ls => (l, ls))
+    val (l, ls) = split_first (snd (take_prefix empty_line lines))
+  in (test_outcome solver_name l, ls) end
+
+
+(* CVC3 *)
+
+local
+  fun cvc3_options ctxt = [
+    "-seed", string_of_int (Config.get ctxt SMT2_Config.random_seed),
+    "-lang", "smtlib", "-output-lang", "presentation",
+    "-timeout", string_of_int (Real.ceil (Config.get ctxt SMT2_Config.timeout))]
+in
+
+val cvc3: SMT2_Solver.solver_config = {
+  name = "cvc3_new",
+  class = K SMTLIB2_Interface.smtlib2C,
+  avail = make_avail "CVC3_NEW",
+  command = make_command "CVC3_NEW",
+  options = cvc3_options,
+  default_max_relevant = 400 (* FUDGE *),
+  can_filter = false,
+  outcome =
+    on_first_line (outcome_of "Unsatisfiable." "Satisfiable." "Unknown."),
+  cex_parser = NONE,
+  replay = NONE }
+
+end
+
+
+(* Yices *)
+
+val yices: SMT2_Solver.solver_config = {
+  name = "yices_new",
+  class = K SMTLIB2_Interface.smtlib2C,
+  avail = make_avail "YICES_NEW",
+  command = make_command "YICES_NEW",
+  options = (fn ctxt => [
+    "--rand-seed=" ^ string_of_int (Config.get ctxt SMT2_Config.random_seed),
+    "--timeout=" ^
+      string_of_int (Real.ceil (Config.get ctxt SMT2_Config.timeout)),
+    "--smtlib"]),
+  default_max_relevant = 350 (* FUDGE *),
+  can_filter = false,
+  outcome = on_first_line (outcome_of "unsat" "sat" "unknown"),
+  cex_parser = NONE,
+  replay = NONE }
+
+
+(* Z3 *)
+
+datatype z3_non_commercial =
+  Z3_Non_Commercial_Unknown |
+  Z3_Non_Commercial_Accepted |
+  Z3_Non_Commercial_Declined
+
+local
+  val accepted = member (op =) ["yes", "Yes", "YES"]
+  val declined = member (op =) ["no", "No", "NO"]
+in
+
+fun z3_non_commercial () =
+  let
+    val flag1 = Options.default_string @{option z3_non_commercial}
+    val flag2 = getenv "Z3_NON_COMMERCIAL"
+  in
+    if accepted flag1 then Z3_Non_Commercial_Accepted
+    else if declined flag1 then Z3_Non_Commercial_Declined
+    else if accepted flag2 then Z3_Non_Commercial_Accepted
+    else if declined flag2 then Z3_Non_Commercial_Declined
+    else Z3_Non_Commercial_Unknown
+  end
+
+fun if_z3_non_commercial f =
+  (case z3_non_commercial () of
+    Z3_Non_Commercial_Accepted => f ()
+  | Z3_Non_Commercial_Declined =>
+      error (Pretty.string_of (Pretty.para
+        "The SMT solver Z3 may be used only for non-commercial applications."))
+  | Z3_Non_Commercial_Unknown =>
+      error (Pretty.string_of (Pretty.para
+        ("The SMT solver Z3 is not activated. To activate it, set the Isabelle \
+         \system option \"z3_non_commercial\" to \"yes\" (e.g. via \
+         \the Isabelle/jEdit menu Plugin Options / Isabelle / General)."))))
+
+end
+
+val z3_extensions = Attrib.setup_config_bool @{binding z3_new_extensions} (K false)
+
+local
+  fun z3_make_command name () = if_z3_non_commercial (make_command name)
+
+  fun z3_options ctxt =
+    ["REFINE_INJ_AXIOM=false" (* not supported by replay *),
+     "-rs:" ^ string_of_int (Config.get ctxt SMT2_Config.random_seed),
+     "-T:" ^ string_of_int (Real.ceil (Config.get ctxt SMT2_Config.timeout)),
+     "-smt2"]
+
+  fun select_class ctxt =
+    if Config.get ctxt z3_extensions then Z3_New_Interface.smtlib2_z3C
+    else SMTLIB2_Interface.smtlib2C
+in
+
+val z3: SMT2_Solver.solver_config = {
+  name = "z3_new",
+  class = select_class,
+  avail = make_avail "Z3_NEW",
+  command = z3_make_command "Z3_NEW",
+  options = z3_options,
+  default_max_relevant = 350 (* FUDGE *),
+  can_filter = true,
+  outcome = on_first_line (outcome_of "unsat" "sat" "unknown"),
+  cex_parser = SOME Z3_New_Model.parse_counterex,
+  replay = SOME Z3_New_Replay.replay }
+
+end
+
+
+(* overall setup *)
+
+val _ = Theory.setup (
+  SMT2_Solver.add_solver cvc3 #>
+  SMT2_Solver.add_solver yices #>
+  SMT2_Solver.add_solver z3)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/SMT2/smt2_translate.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -0,0 +1,534 @@
+(*  Title:      HOL/Tools/SMT2/smt2_translate.ML
+    Author:     Sascha Boehme, TU Muenchen
+
+Translate theorems into an SMT intermediate format and serialize them.
+*)
+
+signature SMT2_TRANSLATE =
+sig
+  (*intermediate term structure*)
+  datatype squant = SForall | SExists
+  datatype 'a spattern = SPat of 'a list | SNoPat of 'a list
+  datatype sterm =
+    SVar of int |
+    SApp of string * sterm list |
+    SLet of string * sterm * sterm |
+    SQua of squant * string list * sterm spattern list * int option * sterm
+
+  (*translation configuration*)
+  type sign = {
+    header: string,
+    sorts: string list,
+    dtyps: (string * (string * (string * string) list) list) list list,
+    funcs: (string * (string list * string)) list }
+  type config = {
+    header: term list -> string,
+    has_datatypes: bool,
+    serialize: string list -> sign -> sterm list -> string }
+  type replay_data = {
+    context: Proof.context,
+    typs: typ Symtab.table,
+    terms: term Symtab.table,
+    rewrite_rules: thm list,
+    assms: (int * thm) list }
+
+  (*translation*)
+  val add_config: SMT2_Util.class * (Proof.context -> config) -> Context.generic -> Context.generic
+  val translate: Proof.context -> string list -> (int * thm) list -> string * replay_data
+end
+
+structure SMT2_Translate: SMT2_TRANSLATE =
+struct
+
+
+(* intermediate term structure *)
+
+datatype squant = SForall | SExists
+
+datatype 'a spattern = SPat of 'a list | SNoPat of 'a list
+
+datatype sterm =
+  SVar of int |
+  SApp of string * sterm list |
+  SLet of string * sterm * sterm |
+  SQua of squant * string list * sterm spattern list * int option * sterm
+
+
+
+(* translation configuration *)
+
+type sign = {
+  header: string,
+  sorts: string list,
+  dtyps: (string * (string * (string * string) list) list) list list,
+  funcs: (string * (string list * string)) list }
+
+type config = {
+  header: term list -> string,
+  has_datatypes: bool,
+  serialize: string list -> sign -> sterm list -> string }
+
+type replay_data = {
+  context: Proof.context,
+  typs: typ Symtab.table,
+  terms: term Symtab.table,
+  rewrite_rules: thm list,
+  assms: (int * thm) list }
+
+
+
+(* translation context *)
+
+fun add_components_of_typ (Type (s, Ts)) =
+    cons (Long_Name.base_name s) #> fold_rev add_components_of_typ Ts
+  | add_components_of_typ (TFree (s, _)) = cons (perhaps (try (unprefix "'")) s)
+  | add_components_of_typ _ = I;
+
+fun suggested_name_of_typ T = space_implode "_" (add_components_of_typ T []);
+
+fun suggested_name_of_term (Const (s, _)) = Long_Name.base_name s
+  | suggested_name_of_term (Free (s, _)) = s
+  | suggested_name_of_term _ = Name.uu
+
+val empty_tr_context = (Name.context, Typtab.empty, Termtab.empty)
+val safe_suffix = "$"
+
+fun add_typ T proper (cx as (names, typs, terms)) =
+  (case Typtab.lookup typs T of
+    SOME (name, _) => (name, cx)
+  | NONE =>
+      let
+        val sugg = Name.desymbolize true (suggested_name_of_typ T) ^ safe_suffix
+        val (name, names') = Name.variant sugg names
+        val typs' = Typtab.update (T, (name, proper)) typs
+      in (name, (names', typs', terms)) end)
+
+fun add_fun t sort (cx as (names, typs, terms)) =
+  (case Termtab.lookup terms t of
+    SOME (name, _) => (name, cx)
+  | NONE => 
+      let
+        val sugg = Name.desymbolize false (suggested_name_of_term t) ^ safe_suffix
+        val (name, names') = Name.variant sugg names
+        val terms' = Termtab.update (t, (name, sort)) terms
+      in (name, (names', typs, terms')) end)
+
+fun sign_of header dtyps (_, typs, terms) = {
+  header = header,
+  sorts = Typtab.fold (fn (_, (n, true)) => cons n | _ => I) typs [],
+  dtyps = dtyps,
+  funcs = Termtab.fold (fn (_, (n, SOME ss)) => cons (n,ss) | _ => I) terms []}
+
+fun replay_data_of ctxt rules assms (_, typs, terms) =
+  let
+    fun add_typ (T, (n, _)) = Symtab.update (n, T)
+    val typs' = Typtab.fold add_typ typs Symtab.empty
+
+    fun add_fun (t, (n, _)) = Symtab.update (n, t)
+    val terms' = Termtab.fold add_fun terms Symtab.empty
+  in
+    {context=ctxt, typs=typs', terms=terms', rewrite_rules=rules, assms=assms}
+  end
+
+
+
+(* preprocessing *)
+
+(** datatype declarations **)
+
+fun collect_datatypes_and_records (tr_context, ctxt) ts =
+  let
+    val (declss, ctxt') = fold (Term.fold_types SMT2_Datatypes.add_decls) ts ([], ctxt)
+
+    fun is_decl_typ T = exists (exists (equal T o fst)) declss
+
+    fun add_typ' T proper =
+      (case SMT2_Builtin.dest_builtin_typ ctxt' T of
+        SOME n => pair n
+      | NONE => add_typ T proper)
+
+    fun tr_select sel =
+      let val T = Term.range_type (Term.fastype_of sel)
+      in add_fun sel NONE ##>> add_typ' T (not (is_decl_typ T)) end
+    fun tr_constr (constr, selects) =
+      add_fun constr NONE ##>> fold_map tr_select selects
+    fun tr_typ (T, cases) = add_typ' T false ##>> fold_map tr_constr cases
+    val (declss', tr_context') = fold_map (fold_map tr_typ) declss tr_context
+
+    fun add (constr, selects) =
+      Termtab.update (constr, length selects) #>
+      fold (Termtab.update o rpair 1) selects
+    val funcs = fold (fold (fold add o snd)) declss Termtab.empty
+  in ((funcs, declss', tr_context', ctxt'), ts) end
+    (* FIXME: also return necessary datatype and record theorems *)
+
+
+(** eta-expand quantifiers, let expressions and built-ins *)
+
+local
+  fun eta f T t = Abs (Name.uu, T, f (Term.incr_boundvars 1 t $ Bound 0))
+
+  fun exp f T = eta f (Term.domain_type (Term.domain_type T))
+
+  fun exp2 T q =
+    let val U = Term.domain_type T
+    in Abs (Name.uu, U, q $ eta I (Term.domain_type U) (Bound 0)) end
+
+  fun expf k i T t =
+    let val Ts = drop i (fst (SMT2_Util.dest_funT k T))
+    in
+      Term.incr_boundvars (length Ts) t
+      |> fold_rev (fn i => fn u => u $ Bound i) (0 upto length Ts - 1)
+      |> fold_rev (fn T => fn u => Abs (Name.uu, T, u)) Ts
+    end
+in
+
+fun eta_expand ctxt funcs =
+  let
+    fun exp_func t T ts =
+      (case Termtab.lookup funcs t of
+        SOME k => Term.list_comb (t, ts) |> k <> length ts ? expf k (length ts) T
+      | NONE => Term.list_comb (t, ts))
+
+    fun expand ((q as Const (@{const_name All}, _)) $ Abs a) = q $ abs_expand a
+      | expand ((q as Const (@{const_name All}, T)) $ t) = q $ exp expand T t
+      | expand (q as Const (@{const_name All}, T)) = exp2 T q
+      | expand ((q as Const (@{const_name Ex}, _)) $ Abs a) = q $ abs_expand a
+      | expand ((q as Const (@{const_name Ex}, T)) $ t) = q $ exp expand T t
+      | expand (q as Const (@{const_name Ex}, T)) = exp2 T q
+      | expand ((l as Const (@{const_name Let}, _)) $ t $ Abs a) = expand (Term.betapply (Abs a, t))
+      | expand ((l as Const (@{const_name Let}, T)) $ t $ u) = expand (u $ t)
+      | expand ((l as Const (@{const_name Let}, T)) $ t) =
+          let val U = Term.domain_type (Term.range_type T)
+          in Abs (Name.uu, U, Bound 0 $ Term.incr_boundvars 1 t) end
+      | expand (Const (@{const_name Let}, T)) =
+          let val U = Term.domain_type (Term.range_type T)
+          in Abs (Name.uu, Term.domain_type T, Abs (Name.uu, U, Bound 0 $ Bound 1)) end
+      | expand t =
+          (case Term.strip_comb t of
+            (u as Const (c as (_, T)), ts) =>
+              (case SMT2_Builtin.dest_builtin ctxt c ts of
+                SOME (_, k, us, mk) =>
+                  if k = length us then mk (map expand us)
+                  else if k < length us then chop k (map expand us) |>> mk |> Term.list_comb
+                  else expf k (length ts) T (mk (map expand us))
+              | NONE => exp_func u T (map expand ts))
+          | (u as Free (_, T), ts) => exp_func u T (map expand ts)
+          | (Abs a, ts) => Term.list_comb (abs_expand a, map expand ts)
+          | (u, ts) => Term.list_comb (u, map expand ts))
+
+    and abs_expand (n, T, t) = Abs (n, T, expand t)
+  
+  in map expand end
+
+end
+
+
+(** introduce explicit applications **)
+
+local
+  (*
+    Make application explicit for functions with varying number of arguments.
+  *)
+
+  fun add t i = apfst (Termtab.map_default (t, i) (Integer.min i))
+  fun add_type T = apsnd (Typtab.update (T, ()))
+
+  fun min_arities t =
+    (case Term.strip_comb t of
+      (u as Const _, ts) => add u (length ts) #> fold min_arities ts
+    | (u as Free _, ts) => add u (length ts) #> fold min_arities ts
+    | (Abs (_, T, u), ts) => (can dest_funT T ? add_type T) #> min_arities u #> fold min_arities ts
+    | (_, ts) => fold min_arities ts)
+
+  fun minimize types t i =
+    let
+      fun find_min j [] _ = j
+        | find_min j (U :: Us) T =
+            if Typtab.defined types T then j else find_min (j + 1) Us (U --> T)
+
+      val (Ts, T) = Term.strip_type (Term.type_of t)
+    in find_min 0 (take i (rev Ts)) T end
+
+  fun app u (t, T) = (Const (@{const_name SMT2.fun_app}, T --> T) $ t $ u, Term.range_type T)
+
+  fun apply i t T ts =
+    let
+      val (ts1, ts2) = chop i ts
+      val (_, U) = SMT2_Util.dest_funT i T
+    in fst (fold app ts2 (Term.list_comb (t, ts1), U)) end
+in
+
+fun intro_explicit_application ctxt funcs ts =
+  let
+    val (arities, types) = fold min_arities ts (Termtab.empty, Typtab.empty)
+    val arities' = Termtab.map (minimize types) arities (* FIXME: highly suspicious *)
+
+    fun app_func t T ts =
+      if is_some (Termtab.lookup funcs t) then Term.list_comb (t, ts)
+      else apply (the (Termtab.lookup arities' t)) t T ts
+
+    fun in_list T f t = HOLogic.mk_list T (map f (HOLogic.dest_list t))
+
+    fun traverse Ts t =
+      (case Term.strip_comb t of
+        (q as Const (@{const_name All}, _), [Abs (x, T, u)]) =>
+          q $ Abs (x, T, in_trigger (T :: Ts) u)
+      | (q as Const (@{const_name Ex}, _), [Abs (x, T, u)]) =>
+          q $ Abs (x, T, in_trigger (T :: Ts) u)
+      | (q as Const (@{const_name Let}, _), [u1, u2 as Abs _]) =>
+          q $ traverse Ts u1 $ traverse Ts u2
+      | (u as Const (c as (_, T)), ts) =>
+          (case SMT2_Builtin.dest_builtin ctxt c ts of
+            SOME (_, k, us, mk) =>
+              let
+                val (ts1, ts2) = chop k (map (traverse Ts) us)
+                val U = Term.strip_type T |>> snd o chop k |> (op --->)
+              in apply 0 (mk ts1) U ts2 end
+          | NONE => app_func u T (map (traverse Ts) ts))
+      | (u as Free (_, T), ts) => app_func u T (map (traverse Ts) ts)
+      | (u as Bound i, ts) => apply 0 u (nth Ts i) (map (traverse Ts) ts)
+      | (Abs (n, T, u), ts) => traverses Ts (Abs (n, T, traverse (T::Ts) u)) ts
+      | (u, ts) => traverses Ts u ts)
+    and in_trigger Ts ((c as @{const SMT2.trigger}) $ p $ t) = c $ in_pats Ts p $ in_weight Ts t
+      | in_trigger Ts t = in_weight Ts t
+    and in_pats Ts ps =
+      in_list @{typ "SMT2.pattern list"} (in_list @{typ SMT2.pattern} (in_pat Ts)) ps
+    and in_pat Ts ((p as Const (@{const_name SMT2.pat}, _)) $ t) = p $ traverse Ts t
+      | in_pat Ts ((p as Const (@{const_name SMT2.nopat}, _)) $ t) = p $ traverse Ts t
+      | in_pat _ t = raise TERM ("bad pattern", [t])
+    and in_weight Ts ((c as @{const SMT2.weight}) $ w $ t) = c $ w $ traverse Ts t
+      | in_weight Ts t = traverse Ts t 
+    and traverses Ts t ts = Term.list_comb (t, map (traverse Ts) ts)
+  in map (traverse []) ts end
+
+val fun_app_eq = mk_meta_eq @{thm SMT2.fun_app_def}
+
+end
+
+
+(** map HOL formulas to FOL formulas (i.e., separate formulas froms terms) **)
+
+local
+  val is_quant = member (op =) [@{const_name All}, @{const_name Ex}]
+
+  val fol_rules = [
+    Let_def,
+    @{lemma "P = True == P" by (rule eq_reflection) simp},
+    @{lemma "if P then True else False == P" by (rule eq_reflection) simp}]
+
+  exception BAD_PATTERN of unit
+
+  fun wrap_in_if pat t =
+    if pat then raise BAD_PATTERN () else @{const If (bool)} $ t $ @{const True} $ @{const False}
+
+  fun is_builtin_conn_or_pred ctxt c ts =
+    is_some (SMT2_Builtin.dest_builtin_conn ctxt c ts) orelse
+    is_some (SMT2_Builtin.dest_builtin_pred ctxt c ts)
+in
+
+fun folify ctxt =
+  let
+    fun in_list T f t = HOLogic.mk_list T (map_filter f (HOLogic.dest_list t))
+
+    fun in_term pat t =
+      (case Term.strip_comb t of
+        (@{const True}, []) => t
+      | (@{const False}, []) => t
+      | (u as Const (@{const_name If}, _), [t1, t2, t3]) =>
+          if pat then raise BAD_PATTERN () else u $ in_form t1 $ in_term pat t2 $ in_term pat t3
+      | (Const (c as (n, _)), ts) =>
+          if is_builtin_conn_or_pred ctxt c ts then wrap_in_if pat (in_form t)
+          else if is_quant n then wrap_in_if pat (in_form t)
+          else Term.list_comb (Const c, map (in_term pat) ts)
+      | (Free c, ts) => Term.list_comb (Free c, map (in_term pat) ts)
+      | _ => t)
+
+    and in_weight ((c as @{const SMT2.weight}) $ w $ t) = c $ w $ in_form t
+      | in_weight t = in_form t 
+
+    and in_pat ((p as Const (@{const_name SMT2.pat}, _)) $ t) =
+          p $ in_term true t
+      | in_pat ((p as Const (@{const_name SMT2.nopat}, _)) $ t) =
+          p $ in_term true t
+      | in_pat t = raise TERM ("bad pattern", [t])
+
+    and in_pats ps =
+      in_list @{typ "SMT2.pattern list"} (SOME o in_list @{typ SMT2.pattern} (try in_pat)) ps
+
+    and in_trigger ((c as @{const SMT2.trigger}) $ p $ t) = c $ in_pats p $ in_weight t
+      | in_trigger t = in_weight t
+
+    and in_form t =
+      (case Term.strip_comb t of
+        (q as Const (qn, _), [Abs (n, T, u)]) =>
+          if is_quant qn then q $ Abs (n, T, in_trigger u)
+          else in_term false t
+      | (Const c, ts) =>
+          (case SMT2_Builtin.dest_builtin_conn ctxt c ts of
+            SOME (_, _, us, mk) => mk (map in_form us)
+          | NONE =>
+              (case SMT2_Builtin.dest_builtin_pred ctxt c ts of
+                SOME (_, _, us, mk) => mk (map (in_term false) us)
+              | NONE => in_term false t))
+      | _ => in_term false t)
+  in
+    map in_form #>
+    pair (fol_rules, I)
+  end
+
+end
+
+
+(* translation into intermediate format *)
+
+(** utility functions **)
+
+val quantifier = (fn
+    @{const_name All} => SOME SForall
+  | @{const_name Ex} => SOME SExists
+  | _ => NONE)
+
+fun group_quant qname Ts (t as Const (q, _) $ Abs (_, T, u)) =
+      if q = qname then group_quant qname (T :: Ts) u else (Ts, t)
+  | group_quant _ Ts t = (Ts, t)
+
+fun dest_weight (@{const SMT2.weight} $ w $ t) = (SOME (snd (HOLogic.dest_number w)), t)
+  | dest_weight t = (NONE, t)
+
+fun dest_pat (Const (@{const_name SMT2.pat}, _) $ t) = (t, true)
+  | dest_pat (Const (@{const_name SMT2.nopat}, _) $ t) = (t, false)
+  | dest_pat t = raise TERM ("bad pattern", [t])
+
+fun dest_pats [] = I
+  | dest_pats ts =
+      (case map dest_pat ts |> split_list ||> distinct (op =) of
+        (ps, [true]) => cons (SPat ps)
+      | (ps, [false]) => cons (SNoPat ps)
+      | _ => raise TERM ("bad multi-pattern", ts))
+
+fun dest_trigger (@{const SMT2.trigger} $ tl $ t) =
+      (rev (fold (dest_pats o HOLogic.dest_list) (HOLogic.dest_list tl) []), t)
+  | dest_trigger t = ([], t)
+
+fun dest_quant qn T t = quantifier qn |> Option.map (fn q =>
+  let
+    val (Ts, u) = group_quant qn [T] t
+    val (ps, p) = dest_trigger u
+    val (w, b) = dest_weight p
+  in (q, rev Ts, ps, w, b) end)
+
+fun fold_map_pat f (SPat ts) = fold_map f ts #>> SPat
+  | fold_map_pat f (SNoPat ts) = fold_map f ts #>> SNoPat
+
+
+(** translation from Isabelle terms into SMT intermediate terms **)
+
+fun intermediate header dtyps builtin ctxt ts trx =
+  let
+    fun transT (T as TFree _) = add_typ T true
+      | transT (T as TVar _) = (fn _ => raise TYPE ("bad SMT type", [T], []))
+      | transT (T as Type _) =
+          (case SMT2_Builtin.dest_builtin_typ ctxt T of
+            SOME n => pair n
+          | NONE => add_typ T true)
+
+    fun app n ts = SApp (n, ts)
+
+    fun trans t =
+      (case Term.strip_comb t of
+        (Const (qn, _), [Abs (_, T, t1)]) =>
+          (case dest_quant qn T t1 of
+            SOME (q, Ts, ps, w, b) =>
+              fold_map transT Ts ##>> fold_map (fold_map_pat trans) ps ##>>
+              trans b #>> (fn ((Ts', ps'), b') => SQua (q, Ts', ps', w, b'))
+          | NONE => raise TERM ("unsupported quantifier", [t]))
+      | (Const (@{const_name Let}, _), [t1, Abs (_, T, t2)]) =>
+          transT T ##>> trans t1 ##>> trans t2 #>> (fn ((U, u1), u2) => SLet (U, u1, u2))
+      | (u as Const (c as (_, T)), ts) =>
+          (case builtin ctxt c ts of
+            SOME (n, _, us, _) => fold_map trans us #>> app n
+          | NONE => transs u T ts)
+      | (u as Free (_, T), ts) => transs u T ts
+      | (Bound i, []) => pair (SVar i)
+      | _ => raise TERM ("bad SMT term", [t]))
+ 
+    and transs t T ts =
+      let val (Us, U) = SMT2_Util.dest_funT (length ts) T
+      in
+        fold_map transT Us ##>> transT U #-> (fn Up =>
+          add_fun t (SOME Up) ##>> fold_map trans ts #>> SApp)
+      end
+
+    val (us, trx') = fold_map trans ts trx
+  in ((sign_of (header ts) dtyps trx', us), trx') end
+
+
+
+(* translation *)
+
+structure Configs = Generic_Data
+(
+  type T = (Proof.context -> config) SMT2_Util.dict
+  val empty = []
+  val extend = I
+  fun merge data = SMT2_Util.dict_merge fst data
+)
+
+fun add_config (cs, cfg) = Configs.map (SMT2_Util.dict_update (cs, cfg))
+
+fun get_config ctxt = 
+  let val cs = SMT2_Config.solver_class_of ctxt
+  in
+    (case SMT2_Util.dict_get (Configs.get (Context.Proof ctxt)) cs of
+      SOME cfg => cfg ctxt
+    | NONE => error ("SMT: no translation configuration found " ^
+        "for solver class " ^ quote (SMT2_Util.string_of_class cs)))
+  end
+
+fun translate ctxt comments ithms =
+  let
+    val {header, has_datatypes, serialize} = get_config ctxt
+
+    fun no_dtyps (tr_context, ctxt) ts =
+      ((Termtab.empty, [], tr_context, ctxt), ts)
+
+    val ts1 = map (Envir.beta_eta_contract o SMT2_Util.prop_of o snd) ithms
+
+    val ((funcs, dtyps, tr_context, ctxt1), ts2) =
+      ((empty_tr_context, ctxt), ts1)
+      |-> (if has_datatypes then collect_datatypes_and_records else no_dtyps)
+
+    fun is_binder (Const (@{const_name Let}, _) $ _) = true
+      | is_binder t = Lambda_Lifting.is_quantifier t
+
+    fun mk_trigger ((q as Const (@{const_name All}, _)) $ Abs (n, T, t)) =
+          q $ Abs (n, T, mk_trigger t)
+      | mk_trigger (eq as (Const (@{const_name HOL.eq}, T) $ lhs $ _)) =
+          Term.domain_type T --> @{typ SMT2.pattern}
+          |> (fn T => Const (@{const_name SMT2.pat}, T) $ lhs)
+          |> HOLogic.mk_list @{typ SMT2.pattern} o single
+          |> HOLogic.mk_list @{typ "SMT2.pattern list"} o single
+          |> (fn t => @{const SMT2.trigger} $ t $ eq)
+      | mk_trigger t = t
+
+    val (ctxt2, ts3) =
+      ts2
+      |> eta_expand ctxt1 funcs
+      |> rpair ctxt1
+      |-> Lambda_Lifting.lift_lambdas NONE is_binder
+      |-> (fn (ts', defs) => fn ctxt' =>
+          map mk_trigger defs @ ts'
+          |> intro_explicit_application ctxt' funcs 
+          |> pair ctxt')
+
+    val ((rewrite_rules, builtin), ts4) = folify ctxt2 ts3
+      |>> apfst (cons fun_app_eq)
+  in
+    (ts4, tr_context)
+    |-> intermediate header dtyps (builtin SMT2_Builtin.dest_builtin) ctxt2
+    |>> uncurry (serialize comments)
+    ||> replay_data_of ctxt2 rewrite_rules ithms
+  end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/SMT2/smt2_util.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -0,0 +1,224 @@
+(*  Title:      HOL/Tools/SMT2/smt2_util.ML
+    Author:     Sascha Boehme, TU Muenchen
+
+General utility functions.
+*)
+
+signature SMT2_UTIL =
+sig
+  (*basic combinators*)
+  val repeat: ('a -> 'a option) -> 'a -> 'a
+  val repeat_yield: ('a -> 'b -> ('a * 'b) option) -> 'a -> 'b -> 'a * 'b
+
+  (*class dictionaries*)
+  type class = string list
+  val basicC: class
+  val string_of_class: class -> string
+  type 'a dict = (class * 'a) Ord_List.T
+  val dict_map_default: class * 'a -> ('a -> 'a) -> 'a dict -> 'a dict
+  val dict_update: class * 'a -> 'a dict -> 'a dict
+  val dict_merge: ('a * 'a -> 'a) -> 'a dict * 'a dict -> 'a dict
+  val dict_lookup: 'a dict -> class -> 'a list
+  val dict_get: 'a dict -> class -> 'a option
+
+  (*types*)
+  val dest_funT: int -> typ -> typ list * typ
+
+  (*terms*)
+  val dest_conj: term -> term * term
+  val dest_disj: term -> term * term
+  val under_quant: (term -> 'a) -> term -> 'a
+  val is_number: term -> bool
+
+  (*patterns and instantiations*)
+  val mk_const_pat: theory -> string -> (ctyp -> 'a) -> 'a * cterm
+  val destT1: ctyp -> ctyp
+  val destT2: ctyp -> ctyp
+  val instTs: ctyp list -> ctyp list * cterm -> cterm
+  val instT: ctyp -> ctyp * cterm -> cterm
+  val instT': cterm -> ctyp * cterm -> cterm
+
+  (*certified terms*)
+  val certify: Proof.context -> term -> cterm
+  val typ_of: cterm -> typ
+  val dest_cabs: cterm -> Proof.context -> cterm * Proof.context
+  val dest_all_cabs: cterm -> Proof.context -> cterm * Proof.context
+  val dest_cbinder: cterm -> Proof.context -> cterm * Proof.context
+  val dest_all_cbinders: cterm -> Proof.context -> cterm * Proof.context
+  val mk_cprop: cterm -> cterm
+  val dest_cprop: cterm -> cterm
+  val mk_cequals: cterm -> cterm -> cterm
+  val term_of: cterm -> term
+  val prop_of: thm -> term
+
+  (*conversions*)
+  val if_conv: (term -> bool) -> conv -> conv -> conv
+  val if_true_conv: (term -> bool) -> conv -> conv
+  val if_exists_conv: (term -> bool) -> conv -> conv
+  val binders_conv: (Proof.context -> conv) -> Proof.context -> conv
+  val under_quant_conv: (Proof.context * cterm list -> conv) ->
+    Proof.context -> conv
+  val prop_conv: conv -> conv
+end
+
+structure SMT2_Util: SMT2_UTIL =
+struct
+
+(* basic combinators *)
+
+fun repeat f =
+  let fun rep x = (case f x of SOME y => rep y | NONE => x)
+  in rep end
+
+fun repeat_yield f =
+  let fun rep x y = (case f x y of SOME (x', y') => rep x' y' | NONE => (x, y))
+  in rep end
+
+
+(* class dictionaries *)
+
+type class = string list
+
+val basicC = []
+
+fun string_of_class [] = "basic"
+  | string_of_class cs = "basic." ^ space_implode "." cs
+
+type 'a dict = (class * 'a) Ord_List.T
+
+fun class_ord ((cs1, _), (cs2, _)) =
+  rev_order (list_ord fast_string_ord (cs1, cs2))
+
+fun dict_insert (cs, x) d =
+  if AList.defined (op =) d cs then d
+  else Ord_List.insert class_ord (cs, x) d
+
+fun dict_map_default (cs, x) f =
+  dict_insert (cs, x) #> AList.map_entry (op =) cs f
+
+fun dict_update (e as (_, x)) = dict_map_default e (K x)
+
+fun dict_merge val_merge = sort class_ord o AList.join (op =) (K val_merge)
+
+fun dict_lookup d cs =
+  let fun match (cs', x) = if is_prefix (op =) cs' cs then SOME x else NONE
+  in map_filter match d end
+
+fun dict_get d cs =
+  (case AList.lookup (op =) d cs of
+    NONE => (case cs of [] => NONE | _ => dict_get d (take (length cs - 1) cs))
+  | SOME x => SOME x)
+
+
+(* types *)
+
+val dest_funT =
+  let
+    fun dest Ts 0 T = (rev Ts, T)
+      | dest Ts i (Type ("fun", [T, U])) = dest (T::Ts) (i-1) U
+      | dest _ _ T = raise TYPE ("not a function type", [T], [])
+  in dest [] end
+
+
+(* terms *)
+
+fun dest_conj (@{const HOL.conj} $ t $ u) = (t, u)
+  | dest_conj t = raise TERM ("not a conjunction", [t])
+
+fun dest_disj (@{const HOL.disj} $ t $ u) = (t, u)
+  | dest_disj t = raise TERM ("not a disjunction", [t])
+
+fun under_quant f t =
+  (case t of
+    Const (@{const_name All}, _) $ Abs (_, _, u) => under_quant f u
+  | Const (@{const_name Ex}, _) $ Abs (_, _, u) => under_quant f u
+  | _ => f t)
+
+val is_number =
+  let
+    fun is_num env (Const (@{const_name Let}, _) $ t $ Abs (_, _, u)) = is_num (t :: env) u
+      | is_num env (Bound i) = i < length env andalso is_num env (nth env i)
+      | is_num _ t = can HOLogic.dest_number t
+  in is_num [] end
+
+
+(* patterns and instantiations *)
+
+fun mk_const_pat thy name destT =
+  let val cpat = Thm.cterm_of thy (Const (name, Sign.the_const_type thy name))
+  in (destT (Thm.ctyp_of_term cpat), cpat) end
+
+val destT1 = hd o Thm.dest_ctyp
+val destT2 = hd o tl o Thm.dest_ctyp
+
+fun instTs cUs (cTs, ct) = Thm.instantiate_cterm (cTs ~~ cUs, []) ct
+fun instT cU (cT, ct) = instTs [cU] ([cT], ct)
+fun instT' ct = instT (Thm.ctyp_of_term ct)
+
+
+(* certified terms *)
+
+fun certify ctxt = Thm.cterm_of (Proof_Context.theory_of ctxt)
+
+fun typ_of ct = #T (Thm.rep_cterm ct) 
+
+fun dest_cabs ct ctxt =
+  (case Thm.term_of ct of
+    Abs _ =>
+      let val (n, ctxt') = yield_singleton Variable.variant_fixes Name.uu ctxt
+      in (snd (Thm.dest_abs (SOME n) ct), ctxt') end
+  | _ => raise CTERM ("no abstraction", [ct]))
+
+val dest_all_cabs = repeat_yield (try o dest_cabs) 
+
+fun dest_cbinder ct ctxt =
+  (case Thm.term_of ct of
+    Const _ $ Abs _ => dest_cabs (Thm.dest_arg ct) ctxt
+  | _ => raise CTERM ("not a binder", [ct]))
+
+val dest_all_cbinders = repeat_yield (try o dest_cbinder)
+
+val mk_cprop = Thm.apply (Thm.cterm_of @{theory} @{const Trueprop})
+
+fun dest_cprop ct =
+  (case Thm.term_of ct of
+    @{const Trueprop} $ _ => Thm.dest_arg ct
+  | _ => raise CTERM ("not a property", [ct]))
+
+val equals = mk_const_pat @{theory} @{const_name "=="} destT1
+fun mk_cequals ct cu = Thm.mk_binop (instT' ct equals) ct cu
+
+val dest_prop = (fn @{const Trueprop} $ t => t | t => t)
+fun term_of ct = dest_prop (Thm.term_of ct)
+fun prop_of thm = dest_prop (Thm.prop_of thm)
+
+
+(* conversions *)
+
+fun if_conv pred cv1 cv2 ct = if pred (Thm.term_of ct) then cv1 ct else cv2 ct
+
+fun if_true_conv pred cv = if_conv pred cv Conv.all_conv
+
+fun if_exists_conv pred = if_true_conv (Term.exists_subterm pred)
+
+fun binders_conv cv ctxt =
+  Conv.binder_conv (binders_conv cv o snd) ctxt else_conv cv ctxt
+
+fun under_quant_conv cv ctxt =
+  let
+    fun quant_conv inside ctxt cvs ct =
+      (case Thm.term_of ct of
+        Const (@{const_name All}, _) $ Abs _ =>
+          Conv.binder_conv (under_conv cvs) ctxt
+      | Const (@{const_name Ex}, _) $ Abs _ =>
+          Conv.binder_conv (under_conv cvs) ctxt
+      | _ => if inside then cv (ctxt, cvs) else Conv.all_conv) ct
+    and under_conv cvs (cv, ctxt) = quant_conv true ctxt (cv :: cvs)
+  in quant_conv false ctxt [] end
+
+fun prop_conv cv ct =
+  (case Thm.term_of ct of
+    @{const Trueprop} $ _ => Conv.arg_conv cv ct
+  | _ => raise CTERM ("not a property", [ct]))
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/SMT2/smtlib2.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -0,0 +1,199 @@
+(*  Title:      HOL/Tools/SMT2/smtlib2.ML
+    Author:     Sascha Boehme, TU Muenchen
+
+Parsing and generating SMT-LIB 2.
+*)
+
+signature SMTLIB2 =
+sig
+  exception PARSE of int * string
+  datatype tree = 
+    Num of int |
+    Dec of int * int |
+    Str of string |
+    Sym of string |
+    Key of string |
+    S of tree list
+  val parse: string list -> tree
+  val pretty_tree: tree -> Pretty.T
+  val str_of: tree -> string
+end
+
+structure SMTLIB2: SMTLIB2 =
+struct
+
+(* data structures *)
+
+exception PARSE of int * string
+
+datatype tree = 
+  Num of int |
+  Dec of int * int |
+  Str of string |
+  Sym of string |
+  Key of string |
+  S of tree list
+
+datatype unfinished = None | String of string | Symbol of string
+
+
+
+(* utilities *)
+
+fun read_raw pred l cs =
+  (case take_prefix pred cs of
+    ([], []) => raise PARSE (l, "empty token")
+  | ([], c :: _) => raise PARSE (l, "unexpected character " ^ quote c)
+  | x => x)
+
+
+
+(* numerals and decimals *)
+
+fun int_of cs = fst (read_int cs)
+
+fun read_num l cs =
+  (case read_raw Symbol.is_ascii_digit l cs of
+    (cs1, "." :: cs') =>
+      let val (cs2, cs'') = read_raw Symbol.is_ascii_digit l cs'
+      in (Dec (int_of cs1, int_of cs2), cs'') end
+  | (cs1, cs2) => (Num (int_of cs1), cs2))
+
+
+
+(* binary numbers *)
+
+fun is_bin c = (c = "0" orelse c = "1")
+
+fun read_bin l cs = read_raw is_bin l cs |>> Num o fst o read_radix_int 2
+
+
+
+(* hex numbers *)
+
+val is_hex = member (op =) (raw_explode "0123456789abcdefABCDEF")
+
+fun within c1 c2 c = (ord c1 <= ord c andalso ord c <= ord c2)
+
+fun unhex i [] = i
+  | unhex i (c :: cs) =
+      if within "0" "9" c then unhex (i * 16 + (ord c - ord "0")) cs
+      else if within "a" "f" c then unhex (i * 16 + (ord c - ord "a" + 10)) cs
+      else if within "A" "F" c then unhex (i * 16 + (ord c - ord "A" + 10)) cs
+      else raise Fail ("bad hex character " ^ quote c)
+
+fun read_hex l cs = read_raw is_hex l cs |>> Num o unhex 0
+
+
+
+(* symbols *)
+
+val symbol_chars = raw_explode "~!@$%^&*_+=<>.?/-" 
+
+fun is_sym c =
+  Symbol.is_ascii_letter c orelse
+  Symbol.is_ascii_digit c orelse
+  member (op =) symbol_chars c
+
+fun read_sym f l cs = read_raw is_sym l cs |>> f o implode
+
+
+
+(* quoted tokens *)
+
+fun read_quoted stop (escape, replacement) cs =
+  let
+    fun read _ [] = NONE
+      | read rs (cs as (c :: cs')) =
+          if is_prefix (op =) stop cs then
+            SOME (implode (rev rs), drop (length stop) cs)
+          else if not (null escape) andalso is_prefix (op =) escape cs then
+            read (replacement :: rs) (drop (length escape) cs)
+          else read (c :: rs) cs'
+  in read [] cs end
+
+fun read_string cs = read_quoted ["\\", "\""] (["\\", "\\"], "\\") cs
+fun read_symbol cs = read_quoted ["|"] ([], "") cs
+
+
+
+(* core parser *)
+
+fun read _ [] rest tss = (rest, tss)
+  | read l ("(" :: cs) None tss = read l cs None ([] :: tss)
+  | read l (")" :: cs) None (ts1 :: ts2 :: tss) =
+      read l cs None ((S (rev ts1) :: ts2) :: tss)
+  | read l ("#" :: "x" :: cs) None (ts :: tss) =
+      token read_hex l cs ts tss
+  | read l ("#" :: cs) None (ts :: tss) =
+      token read_bin l cs ts tss
+  | read l (":" :: cs) None (ts :: tss) =
+      token (read_sym Key) l cs ts tss
+  | read l ("\"" :: cs) None (ts :: tss) =
+      quoted read_string String Str l "" cs ts tss
+  | read l ("|" :: cs) None (ts :: tss) =
+      quoted read_symbol Symbol Sym l "" cs ts tss
+  | read l ((c as "!") :: cs) None (ts :: tss) =
+      token (fn _ => pair (Sym c)) l cs ts tss
+  | read l (c :: cs) None (ts :: tss) =
+      if Symbol.is_ascii_blank c then read l cs None (ts :: tss)
+      else if Symbol.is_digit c then token read_num l (c :: cs) ts tss
+      else token (read_sym Sym) l (c :: cs) ts tss
+  | read l cs (String s) (ts :: tss) =
+      quoted read_string String Str l s cs ts tss
+  | read l cs (Symbol s) (ts :: tss) =
+      quoted read_symbol Symbol Sym l s cs ts tss
+  | read l _ _ [] = raise PARSE (l, "bad parser state")
+
+and token f l cs ts tss =
+  let val (t, cs') = f l cs
+  in read l cs' None ((t :: ts) :: tss) end
+
+and quoted r f g l s cs ts tss =
+  (case r cs of
+    NONE => (f (s ^ implode cs), ts :: tss)
+  | SOME (s', cs') => read l cs' None ((g (s ^ s') :: ts) :: tss))
+  
+
+
+(* overall parser *)
+
+fun read_line l line = read l (raw_explode line)
+
+fun add_line line (l, (None, tss)) =
+      if size line = 0 orelse nth_string line 0 = ";" then (l + 1, (None, tss))
+      else (l + 1, read_line l line None tss)
+  | add_line line (l, (unfinished, tss)) =
+      (l + 1, read_line l line unfinished tss)
+
+fun finish (_, (None, [[t]])) = t
+  | finish (l, _) = raise PARSE (l, "bad nesting")
+
+fun parse lines = finish (fold add_line lines (1, (None, [[]])))
+
+
+
+(* pretty printer *)
+
+fun pretty_tree (Num i) = Pretty.str (string_of_int i)
+  | pretty_tree (Dec (i, j)) =
+      Pretty.str (string_of_int i ^ "." ^ string_of_int j)
+  | pretty_tree (Str s) =
+      raw_explode s
+      |> maps (fn "\"" => ["\\", "\""] | "\\" => ["\\", "\\"] | c => [c])
+      |> implode
+      |> enclose "\"" "\""
+      |> Pretty.str
+  | pretty_tree (Sym s) =
+      if String.isPrefix "(" s (* for bit vector functions *) orelse
+         forall is_sym (raw_explode s) then
+        Pretty.str s
+      else
+        Pretty.str ("|" ^ s ^ "|")
+  | pretty_tree (Key s) = Pretty.str (":" ^ s)
+  | pretty_tree (S trees) =
+      Pretty.enclose "(" ")" (Pretty.separate "" (map pretty_tree trees))
+
+val str_of = Pretty.str_of o pretty_tree
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/SMT2/smtlib2_interface.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -0,0 +1,154 @@
+(*  Title:      HOL/Tools/SMT2/smtlib2_interface.ML
+    Author:     Sascha Boehme, TU Muenchen
+    Author:     Jasmin Blanchette, TU Muenchen
+
+Interface to SMT solvers based on the SMT-LIB 2 format.
+*)
+
+signature SMTLIB2_INTERFACE =
+sig
+  val smtlib2C: SMT2_Util.class
+  val add_logic: int * (term list -> string option) -> Context.generic -> Context.generic
+  val translate_config: Proof.context -> SMT2_Translate.config
+end
+
+structure SMTLIB2_Interface: SMTLIB2_INTERFACE =
+struct
+
+val smtlib2C = ["smtlib2"]
+
+
+(* builtins *)
+
+local
+  fun int_num _ i = SOME (string_of_int i)
+
+  fun is_linear [t] = SMT2_Util.is_number t
+    | is_linear [t, u] = SMT2_Util.is_number t orelse SMT2_Util.is_number u
+    | is_linear _ = false
+
+  fun times _ _ ts =
+    let val mk = Term.list_comb o pair @{const times (int)}
+    in if is_linear ts then SOME ("*", 2, ts, mk) else NONE end
+in
+
+val setup_builtins =
+  fold (SMT2_Builtin.add_builtin_typ smtlib2C) [
+    (@{typ bool}, K (SOME "Bool"), K (K NONE)),
+    (@{typ int}, K (SOME "Int"), int_num)] #>
+  fold (SMT2_Builtin.add_builtin_fun' smtlib2C) [
+    (@{const True}, "true"),
+    (@{const False}, "false"),
+    (@{const Not}, "not"),
+    (@{const HOL.conj}, "and"),
+    (@{const HOL.disj}, "or"),
+    (@{const HOL.implies}, "=>"),
+    (@{const HOL.eq ('a)}, "="),
+    (@{const If ('a)}, "ite"),
+    (@{const less (int)}, "<"),
+    (@{const less_eq (int)}, "<="),
+    (@{const uminus (int)}, "~"),
+    (@{const plus (int)}, "+"),
+    (@{const minus (int)}, "-")] #>
+  SMT2_Builtin.add_builtin_fun smtlib2C
+    (Term.dest_Const @{const times (int)}, times)
+
+end
+
+
+(* serialization *)
+
+(** header **)
+
+fun fst_int_ord ((i1, _), (i2, _)) = int_ord (i1, i2)
+
+structure Logics = Generic_Data
+(
+  type T = (int * (term list -> string option)) list
+  val empty = []
+  val extend = I
+  fun merge data = Ord_List.merge fst_int_ord data
+)
+
+fun add_logic pf = Logics.map (Ord_List.insert fst_int_ord pf)
+
+fun choose_logic ctxt ts =
+  let
+    fun choose [] = "AUFLIA"
+      | choose ((_, f) :: fs) = (case f ts of SOME s => s | NONE => choose fs)
+  in "(set-logic " ^ choose (Logics.get (Context.Proof ctxt)) ^ ")\n" end
+
+
+(** serialization **)
+
+fun var i = "?v" ^ string_of_int i
+
+fun tree_of_sterm l (SMT2_Translate.SVar i) = SMTLIB2.Sym (var (l - i - 1))
+  | tree_of_sterm _ (SMT2_Translate.SApp (n, [])) = SMTLIB2.Sym n
+  | tree_of_sterm l (SMT2_Translate.SApp (n, ts)) =
+      SMTLIB2.S (SMTLIB2.Sym n :: map (tree_of_sterm l) ts)
+  | tree_of_sterm _ (SMT2_Translate.SLet _) =
+      raise Fail "SMT-LIB: unsupported let expression"
+  | tree_of_sterm l (SMT2_Translate.SQua (q, ss, pats, w, t)) =
+      let
+        val l' = l + length ss
+
+        fun quant_name SMT2_Translate.SForall = "forall"
+          | quant_name SMT2_Translate.SExists = "exists"
+
+        fun gen_trees_of_pat keyword ps =
+          [SMTLIB2.Key keyword, (case map (tree_of_sterm l') ps of [t] => t | ts => SMTLIB2.S ts)]
+        fun trees_of_pat (SMT2_Translate.SPat ps) = gen_trees_of_pat "pattern" ps
+          | trees_of_pat (SMT2_Translate.SNoPat ps) = gen_trees_of_pat "no-pattern" ps
+        fun trees_of_weight NONE = []
+          | trees_of_weight (SOME i) = [SMTLIB2.Key "weight", SMTLIB2.Num i]
+        fun tree_of_pats_weight [] NONE t = t
+          | tree_of_pats_weight pats w t =
+            SMTLIB2.S (SMTLIB2.Sym "!" :: t :: maps trees_of_pat pats @ trees_of_weight w)
+
+        val vs = map_index (fn (i, ty) =>
+          SMTLIB2.S [SMTLIB2.Sym (var (l + i)), SMTLIB2.Sym ty]) ss
+
+        val body = t
+          |> tree_of_sterm l'
+          |> tree_of_pats_weight pats w
+      in
+        SMTLIB2.S [SMTLIB2.Sym (quant_name q), SMTLIB2.S vs, body]
+      end
+
+
+fun sctrarg (sel, typ) = "(" ^ sel ^ " " ^ typ ^ ")"
+fun sctr (name, args) = enclose "(" ")" (space_implode " " (name :: map sctrarg args))
+fun sdatatype (name, ctrs) = enclose "(" ")" (space_implode " " (name :: map sctr ctrs))
+
+fun string_of_fun (f, (ss, s)) = f ^ " (" ^ space_implode " " ss ^ ") " ^ s
+
+fun serialize comments {header, sorts, dtyps, funcs} ts =
+  Buffer.empty
+  |> fold (Buffer.add o enclose "; " "\n") comments
+  |> Buffer.add "(set-option :produce-proofs true)\n"
+  |> Buffer.add header
+  |> fold (Buffer.add o enclose "(declare-sort " " 0)\n")
+       (sort fast_string_ord sorts)
+  |> (if null dtyps then I
+    else Buffer.add (enclose "(declare-datatypes () (" "))\n"
+      (space_implode "\n  " (maps (map sdatatype) dtyps))))
+  |> fold (Buffer.add o enclose "(declare-fun " ")\n" o string_of_fun)
+       (sort (fast_string_ord o pairself fst) funcs)
+  |> fold (Buffer.add o enclose "(assert " ")\n" o SMTLIB2.str_of o
+       tree_of_sterm 0) ts
+  |> Buffer.add "(check-sat)\n(get-proof)\n" (* FIXME: (get-model)\n" *)
+  |> Buffer.content
+
+(* interface *)
+
+fun translate_config ctxt = {
+  header = choose_logic ctxt,
+  has_datatypes = false,
+  serialize = serialize}
+
+val _ = Theory.setup (Context.theory_map
+  (setup_builtins #>
+   SMT2_Translate.add_config (smtlib2C, translate_config)))
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/SMT2/z3_new_interface.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -0,0 +1,198 @@
+(*  Title:      HOL/Tools/SMT2/z3_new_interface.ML
+    Author:     Sascha Boehme, TU Muenchen
+
+Interface to Z3 based on a relaxed version of SMT-LIB.
+*)
+
+signature Z3_NEW_INTERFACE =
+sig
+  val smtlib2_z3C: SMT2_Util.class
+
+  datatype sym = Sym of string * sym list
+  type mk_builtins = {
+    mk_builtin_typ: sym -> typ option,
+    mk_builtin_num: theory -> int -> typ -> cterm option,
+    mk_builtin_fun: theory -> sym -> cterm list -> cterm option }
+  val add_mk_builtins: mk_builtins -> Context.generic -> Context.generic
+  val mk_builtin_typ: Proof.context -> sym -> typ option
+  val mk_builtin_num: Proof.context -> int -> typ -> cterm option
+  val mk_builtin_fun: Proof.context -> sym -> cterm list -> cterm option
+
+  val is_builtin_theory_term: Proof.context -> term -> bool
+end
+
+structure Z3_New_Interface: Z3_NEW_INTERFACE =
+struct
+
+val smtlib2_z3C = SMTLIB2_Interface.smtlib2C @ ["z3"]
+
+
+
+(* interface *)
+
+local
+  fun translate_config ctxt =
+    let
+      val {serialize, ...} = SMTLIB2_Interface.translate_config ctxt
+    in
+      {header=K "", serialize=serialize, has_datatypes=true}
+    end
+
+  fun is_div_mod @{const div (int)} = true
+    | is_div_mod @{const mod (int)} = true
+    | is_div_mod _ = false
+
+  val have_int_div_mod = exists (Term.exists_subterm is_div_mod o Thm.prop_of)
+
+  fun add_div_mod _ (thms, extra_thms) =
+    if have_int_div_mod thms orelse have_int_div_mod extra_thms then
+      (thms, @{thms div_as_z3div mod_as_z3mod} @ extra_thms)
+    else (thms, extra_thms)
+
+  val setup_builtins =
+    SMT2_Builtin.add_builtin_fun' smtlib2_z3C (@{const times (int)}, "*") #>
+    SMT2_Builtin.add_builtin_fun' smtlib2_z3C (@{const z3div}, "div") #>
+    SMT2_Builtin.add_builtin_fun' smtlib2_z3C (@{const z3mod}, "mod")
+in
+
+val _ = Theory.setup (Context.theory_map (
+  setup_builtins #>
+  SMT2_Normalize.add_extra_norm (smtlib2_z3C, add_div_mod) #>
+  SMT2_Translate.add_config (smtlib2_z3C, translate_config)))
+
+end
+
+
+
+(* constructors *)
+
+datatype sym = Sym of string * sym list
+
+
+(** additional constructors **)
+
+type mk_builtins = {
+  mk_builtin_typ: sym -> typ option,
+  mk_builtin_num: theory -> int -> typ -> cterm option,
+  mk_builtin_fun: theory -> sym -> cterm list -> cterm option }
+
+fun chained _ [] = NONE
+  | chained f (b :: bs) = (case f b of SOME y => SOME y | NONE => chained f bs)
+
+fun chained_mk_builtin_typ bs sym =
+  chained (fn {mk_builtin_typ=mk, ...} : mk_builtins => mk sym) bs
+
+fun chained_mk_builtin_num ctxt bs i T =
+  let val thy = Proof_Context.theory_of ctxt
+  in chained (fn {mk_builtin_num=mk, ...} : mk_builtins => mk thy i T) bs end
+
+fun chained_mk_builtin_fun ctxt bs s cts =
+  let val thy = Proof_Context.theory_of ctxt
+  in chained (fn {mk_builtin_fun=mk, ...} : mk_builtins => mk thy s cts) bs end
+
+fun fst_int_ord ((i1, _), (i2, _)) = int_ord (i1, i2)
+
+structure Mk_Builtins = Generic_Data
+(
+  type T = (int * mk_builtins) list
+  val empty = []
+  val extend = I
+  fun merge data = Ord_List.merge fst_int_ord data
+)
+
+fun add_mk_builtins mk = Mk_Builtins.map (Ord_List.insert fst_int_ord (serial (), mk))
+
+fun get_mk_builtins ctxt = map snd (Mk_Builtins.get (Context.Proof ctxt))
+
+
+(** basic and additional constructors **)
+
+fun mk_builtin_typ _ (Sym ("Bool", _)) = SOME @{typ bool}
+  | mk_builtin_typ _ (Sym ("Int", _)) = SOME @{typ int}
+  | mk_builtin_typ _ (Sym ("bool", _)) = SOME @{typ bool}  (*FIXME: legacy*)
+  | mk_builtin_typ _ (Sym ("int", _)) = SOME @{typ int}  (*FIXME: legacy*)
+  | mk_builtin_typ ctxt sym = chained_mk_builtin_typ (get_mk_builtins ctxt) sym
+
+fun mk_builtin_num _ i @{typ int} = SOME (Numeral.mk_cnumber @{ctyp int} i)
+  | mk_builtin_num ctxt i T =
+      chained_mk_builtin_num ctxt (get_mk_builtins ctxt) i T
+
+val mk_true = Thm.cterm_of @{theory} (@{const Not} $ @{const False})
+val mk_false = Thm.cterm_of @{theory} @{const False}
+val mk_not = Thm.apply (Thm.cterm_of @{theory} @{const Not})
+val mk_implies = Thm.mk_binop (Thm.cterm_of @{theory} @{const HOL.implies})
+val mk_iff = Thm.mk_binop (Thm.cterm_of @{theory} @{const HOL.eq (bool)})
+val conj = Thm.cterm_of @{theory} @{const HOL.conj}
+val disj = Thm.cterm_of @{theory} @{const HOL.disj}
+
+fun mk_nary _ cu [] = cu
+  | mk_nary ct _ cts = uncurry (fold_rev (Thm.mk_binop ct)) (split_last cts)
+
+val eq = SMT2_Util.mk_const_pat @{theory} @{const_name HOL.eq} SMT2_Util.destT1
+fun mk_eq ct cu = Thm.mk_binop (SMT2_Util.instT' ct eq) ct cu
+
+val if_term =
+  SMT2_Util.mk_const_pat @{theory} @{const_name If} (SMT2_Util.destT1 o SMT2_Util.destT2)
+fun mk_if cc ct = Thm.mk_binop (Thm.apply (SMT2_Util.instT' ct if_term) cc) ct
+
+val access = SMT2_Util.mk_const_pat @{theory} @{const_name fun_app} SMT2_Util.destT1
+fun mk_access array = Thm.apply (SMT2_Util.instT' array access) array
+
+val update =
+  SMT2_Util.mk_const_pat @{theory} @{const_name fun_upd} (Thm.dest_ctyp o SMT2_Util.destT1)
+fun mk_update array index value =
+  let val cTs = Thm.dest_ctyp (Thm.ctyp_of_term array)
+  in Thm.apply (Thm.mk_binop (SMT2_Util.instTs cTs update) array index) value end
+
+val mk_uminus = Thm.apply (Thm.cterm_of @{theory} @{const uminus (int)})
+val add = Thm.cterm_of @{theory} @{const plus (int)}
+val int0 = Numeral.mk_cnumber @{ctyp int} 0
+val mk_sub = Thm.mk_binop (Thm.cterm_of @{theory} @{const minus (int)})
+val mk_mul = Thm.mk_binop (Thm.cterm_of @{theory} @{const times (int)})
+val mk_div = Thm.mk_binop (Thm.cterm_of @{theory} @{const z3div})
+val mk_mod = Thm.mk_binop (Thm.cterm_of @{theory} @{const z3mod})
+val mk_lt = Thm.mk_binop (Thm.cterm_of @{theory} @{const less (int)})
+val mk_le = Thm.mk_binop (Thm.cterm_of @{theory} @{const less_eq (int)})
+
+fun mk_builtin_fun ctxt sym cts =
+  (case (sym, cts) of
+    (Sym ("true", _), []) => SOME mk_true
+  | (Sym ("false", _), []) => SOME mk_false
+  | (Sym ("not", _), [ct]) => SOME (mk_not ct)
+  | (Sym ("and", _), _) => SOME (mk_nary conj mk_true cts)
+  | (Sym ("or", _), _) => SOME (mk_nary disj mk_false cts)
+  | (Sym ("implies", _), [ct, cu]) => SOME (mk_implies ct cu)
+  | (Sym ("iff", _), [ct, cu]) => SOME (mk_iff ct cu)
+  | (Sym ("~", _), [ct, cu]) => SOME (mk_iff ct cu)
+  | (Sym ("xor", _), [ct, cu]) => SOME (mk_not (mk_iff ct cu))
+  | (Sym ("if", _), [ct1, ct2, ct3]) => SOME (mk_if ct1 ct2 ct3)
+  | (Sym ("ite", _), [ct1, ct2, ct3]) => SOME (mk_if ct1 ct2 ct3) (* FIXME: remove *)
+  | (Sym ("=", _), [ct, cu]) => SOME (mk_eq ct cu)
+  | (Sym ("select", _), [ca, ck]) => SOME (Thm.apply (mk_access ca) ck)
+  | (Sym ("store", _), [ca, ck, cv]) => SOME (mk_update ca ck cv)
+  | _ =>
+    (case (sym, try (#T o Thm.rep_cterm o hd) cts, cts) of
+      (Sym ("+", _), SOME @{typ int}, _) => SOME (mk_nary add int0 cts)
+    | (Sym ("-", _), SOME @{typ int}, [ct]) => SOME (mk_uminus ct)
+    | (Sym ("-", _), SOME @{typ int}, [ct, cu]) => SOME (mk_sub ct cu)
+    | (Sym ("*", _), SOME @{typ int}, [ct, cu]) => SOME (mk_mul ct cu)
+    | (Sym ("div", _), SOME @{typ int}, [ct, cu]) => SOME (mk_div ct cu)
+    | (Sym ("mod", _), SOME @{typ int}, [ct, cu]) => SOME (mk_mod ct cu)
+    | (Sym ("<", _), SOME @{typ int}, [ct, cu]) => SOME (mk_lt ct cu)
+    | (Sym ("<=", _), SOME @{typ int}, [ct, cu]) => SOME (mk_le ct cu)
+    | (Sym (">", _), SOME @{typ int}, [ct, cu]) => SOME (mk_lt cu ct)
+    | (Sym (">=", _), SOME @{typ int}, [ct, cu]) => SOME (mk_le cu ct)
+    | _ => chained_mk_builtin_fun ctxt (get_mk_builtins ctxt) sym cts))
+
+
+
+(* abstraction *)
+
+fun is_builtin_theory_term ctxt t =
+  if SMT2_Builtin.is_builtin_num ctxt t then true
+  else
+    (case Term.strip_comb t of
+      (Const c, ts) => SMT2_Builtin.is_builtin_fun ctxt c ts
+    | _ => false)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/SMT2/z3_new_isar.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -0,0 +1,118 @@
+(*  Title:      HOL/Tools/SMT2/z3_new_isar.ML
+    Author:     Jasmin Blanchette, TU Muenchen
+
+Z3 proofs as generic ATP proofs for Isar proof reconstruction.
+*)
+
+signature Z3_NEW_ISAR =
+sig
+  type ('a, 'b) atp_step = ('a, 'b) ATP_Proof.atp_step
+
+  val atp_proof_of_z3_proof: theory -> thm list -> int -> (int * string) list ->
+    Z3_New_Proof.z3_step list -> (term, string) atp_step list
+end;
+
+structure Z3_New_Isar: Z3_NEW_ISAR =
+struct
+
+open ATP_Util
+open ATP_Problem
+open ATP_Proof
+open ATP_Proof_Reconstruct
+
+val z3_apply_def_rule = Z3_New_Proof.string_of_rule Z3_New_Proof.Apply_Def
+val z3_hypothesis_rule = Z3_New_Proof.string_of_rule Z3_New_Proof.Hypothesis
+val z3_intro_def_rule = Z3_New_Proof.string_of_rule Z3_New_Proof.Intro_Def
+val z3_lemma_rule = Z3_New_Proof.string_of_rule Z3_New_Proof.Lemma
+
+fun inline_z3_defs _ [] = []
+  | inline_z3_defs defs ((name, role, t, rule, deps) :: lines) =
+    if rule = z3_intro_def_rule then
+      let val def = t |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> swap in
+        inline_z3_defs (insert (op =) def defs)
+          (map (replace_dependencies_in_line (name, [])) lines)
+      end
+    else if rule = z3_apply_def_rule then
+      inline_z3_defs defs (map (replace_dependencies_in_line (name, [])) lines)
+    else
+      (name, role, Term.subst_atomic defs t, rule, deps) :: inline_z3_defs defs lines
+
+fun add_z3_hypotheses [] = I
+  | add_z3_hypotheses hyps =
+    HOLogic.dest_Trueprop
+    #> curry s_imp (Library.foldr1 s_conj (map HOLogic.dest_Trueprop hyps))
+    #> HOLogic.mk_Trueprop
+
+fun inline_z3_hypotheses _ _ [] = []
+  | inline_z3_hypotheses hyp_names hyps ((name, role, t, rule, deps) :: lines) =
+    if rule = z3_hypothesis_rule then
+      inline_z3_hypotheses (name :: hyp_names) (AList.map_default (op =) (t, []) (cons name) hyps)
+        lines
+    else
+      let val deps' = subtract (op =) hyp_names deps in
+        if rule = z3_lemma_rule then
+          (name, role, t, rule, deps') :: inline_z3_hypotheses hyp_names hyps lines
+        else
+          let
+            val add_hyps = filter_out (null o inter (op =) deps o snd) hyps
+            val t' = add_z3_hypotheses (map fst add_hyps) t
+            val deps' = subtract (op =) hyp_names deps
+            val hyps' = fold (AList.update (op =) o apsnd (insert (op =) name)) add_hyps hyps
+          in
+            (name, role, t', rule, deps') :: inline_z3_hypotheses hyp_names hyps' lines
+          end
+      end
+
+fun simplify_bool (@{const Not} $ t) = s_not (simplify_bool t)
+  | simplify_bool (@{const conj} $ t $ u) = s_conj (simplify_bool t, simplify_bool u)
+  | simplify_bool (@{const disj} $ t $ u) = s_disj (simplify_bool t, simplify_bool u)
+  | simplify_bool (@{const implies} $ t $ u) = s_imp (simplify_bool t, simplify_bool u)
+  | simplify_bool (@{const HOL.eq (bool)} $ t $ u) = s_iff (simplify_bool t, simplify_bool u)
+  | simplify_bool (t as Const (@{const_name HOL.eq}, _) $ u $ v) =
+    if u aconv v then @{const True} else t
+  | simplify_bool (t $ u) = simplify_bool t $ simplify_bool u
+  | simplify_bool (Abs (s, T, t)) = Abs (s, T, simplify_bool t)
+  | simplify_bool t = t
+
+(* It is not entirely clear why this should be necessary, especially for abstractions variables. *)
+val unskolemize_names =
+  Term.map_abs_vars (perhaps (try Name.dest_skolem))
+  #> Term.map_aterms (perhaps (try (fn Free (s, T) => Free (Name.dest_skolem s, T))))
+
+fun atp_proof_of_z3_proof thy rewrite_rules conjecture_id fact_ids proof =
+  let
+    fun step_of (Z3_New_Proof.Z3_Step {id, rule, prems, concl, ...}) =
+      let
+        fun step_name_of id = (string_of_int id, the_list (AList.lookup (op =) fact_ids id))
+
+        val name as (_, ss) = step_name_of id
+
+        val role =
+          (case rule of
+            Z3_New_Proof.Asserted =>
+              if not (null ss) then Axiom
+              else if id = conjecture_id then Negated_Conjecture
+              else Hypothesis
+          | Z3_New_Proof.Rewrite => Lemma
+          | Z3_New_Proof.Rewrite_Star => Lemma
+          | Z3_New_Proof.Skolemize => Lemma
+          | Z3_New_Proof.Th_Lemma _ => Lemma
+          | _ => Plain)
+
+        val concl' = concl
+          |> Raw_Simplifier.rewrite_term thy rewrite_rules []
+          |> Object_Logic.atomize_term thy
+          |> simplify_bool
+          |> unskolemize_names
+          |> HOLogic.mk_Trueprop
+      in
+        (name, role, concl', Z3_New_Proof.string_of_rule rule, map step_name_of prems)
+      end
+  in
+    proof
+    |> map step_of
+    |> inline_z3_defs []
+    |> inline_z3_hypotheses [] []
+  end
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/SMT2/z3_new_model.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -0,0 +1,334 @@
+(*  Title:      HOL/Tools/SMT2/z3_new_model.ML
+    Author:     Sascha Boehme and Philipp Meyer, TU Muenchen
+
+Parser for counterexamples generated by Z3.
+*)
+
+signature Z3_NEW_MODEL =
+sig
+  val parse_counterex: Proof.context -> SMT2_Translate.replay_data -> string list ->
+    term list * term list
+end
+
+structure Z3_New_Model: Z3_NEW_MODEL =
+struct
+
+(* counterexample expressions *)
+
+datatype expr = True | False | Number of int * int option | Value of int |
+  Array of array | App of string * expr list
+and array = Fresh of expr | Store of (array * expr) * expr
+
+
+(* parsing *)
+
+val space = Scan.many Symbol.is_ascii_blank
+fun spaced p = p --| space
+fun in_parens p = spaced (Scan.$$ "(") |-- p --| spaced (Scan.$$ ")")
+fun in_braces p = spaced (Scan.$$ "{") |-- p --| spaced (Scan.$$ "}")
+
+val digit = (fn
+  "0" => SOME 0 | "1" => SOME 1 | "2" => SOME 2 | "3" => SOME 3 |
+  "4" => SOME 4 | "5" => SOME 5 | "6" => SOME 6 | "7" => SOME 7 |
+  "8" => SOME 8 | "9" => SOME 9 | _ => NONE)
+
+val nat_num = spaced (Scan.repeat1 (Scan.some digit) >>
+  (fn ds => fold (fn d => fn i => i * 10 + d) ds 0))
+val int_num = spaced (Scan.optional ($$ "-" >> K (fn i => ~i)) I :|--
+  (fn sign => nat_num >> sign))
+
+val is_char = Symbol.is_ascii_letter orf Symbol.is_ascii_digit orf
+  member (op =) (raw_explode "_+*-/%~=<>$&|?!.@^#")
+val name = spaced (Scan.many1 is_char >> implode)
+
+fun $$$ s = spaced (Scan.this_string s)
+
+fun array_expr st = st |> in_parens (
+  $$$ "const" |-- expr >> Fresh ||
+  $$$ "store" |-- array_expr -- expr -- expr >> Store)
+
+and expr st = st |> (
+  $$$ "true" >> K True ||
+  $$$ "false" >> K False ||
+  int_num -- Scan.option ($$$ "/" |-- int_num) >> Number ||
+  $$$ "val!" |-- nat_num >> Value ||
+  name >> (App o rpair []) ||
+  array_expr >> Array ||
+  in_parens (name -- Scan.repeat1 expr) >> App)
+
+fun args st = ($$$ "->" >> K [] || expr ::: args) st
+val args_case = args -- expr
+val else_case = $$$ "else" -- $$$ "->" |-- expr >> pair ([] : expr list)
+
+val func =
+  let fun cases st = (else_case >> single || args_case ::: cases) st
+  in in_braces cases end
+
+val cex = space |--
+  Scan.repeat (name --| $$$ "->" -- (func || expr >> (single o pair [])))
+
+fun resolve terms ((n, k), cases) =
+  (case Symtab.lookup terms n of
+    NONE => NONE
+  | SOME t => SOME ((t, k), cases))
+
+fun annotate _ (_, []) = NONE
+  | annotate terms (n, [([], c)]) = resolve terms ((n, 0), (c, []))
+  | annotate _ (_, [_]) = NONE
+  | annotate terms (n, cases as (args, _) :: _) =
+      let val (cases', (_, else_case)) = split_last cases
+      in resolve terms ((n, length args), (else_case, cases')) end
+
+fun read_cex terms ls =
+  maps (cons "\n" o raw_explode) ls
+  |> try (fst o Scan.finite Symbol.stopper cex)
+  |> the_default []
+  |> map_filter (annotate terms)
+
+
+(* translation into terms *)
+
+fun max_value vs =
+  let
+    fun max_val_expr (Value i) = Integer.max i
+      | max_val_expr (App (_, es)) = fold max_val_expr es
+      | max_val_expr (Array a) = max_val_array a
+      | max_val_expr _ = I
+
+    and max_val_array (Fresh e) = max_val_expr e
+      | max_val_array (Store ((a, e1), e2)) =
+          max_val_array a #> max_val_expr e1 #> max_val_expr e2
+
+    fun max_val (_, (ec, cs)) =
+      max_val_expr ec #> fold (fn (es, e) => fold max_val_expr (e :: es)) cs
+
+  in fold max_val vs ~1 end
+
+fun with_context terms f vs = fst (fold_map f vs (terms, max_value vs + 1))
+
+fun get_term n T es (cx as (terms, next_val)) =
+  (case Symtab.lookup terms n of
+    SOME t => ((t, es), cx)
+  | NONE =>
+      let val t = Var (("skolem", next_val), T)
+      in ((t, []), (Symtab.update (n, t) terms, next_val + 1)) end)
+
+fun trans_expr _ True = pair @{const True}
+  | trans_expr _ False = pair @{const False}
+  | trans_expr T (Number (i, NONE)) = pair (HOLogic.mk_number T i)
+  | trans_expr T (Number (i, SOME j)) =
+      pair (Const (@{const_name divide}, [T, T] ---> T) $
+        HOLogic.mk_number T i $ HOLogic.mk_number T j)
+  | trans_expr T (Value i) = pair (Var (("value", i), T))
+  | trans_expr T (Array a) = trans_array T a
+  | trans_expr T (App (n, es)) = get_term n T es #-> (fn (t, es') =>
+      let val Ts = fst (SMT2_Util.dest_funT (length es') (Term.fastype_of t))
+      in fold_map (uncurry trans_expr) (Ts ~~ es') #>> Term.list_comb o pair t end)
+
+and trans_array T a =
+  let val (dT, rT) = Term.dest_funT T
+  in
+    (case a of
+      Fresh e => trans_expr rT e #>> (fn t => Abs ("x", dT, t))
+    | Store ((a', e1), e2) =>
+        trans_array T a' ##>> trans_expr dT e1 ##>> trans_expr rT e2 #>>
+        (fn ((m, k), v) =>
+          Const (@{const_name fun_upd}, [T, dT, rT] ---> T) $ m $ k $ v))
+  end
+
+fun trans_pattern T ([], e) = trans_expr T e #>> pair []
+  | trans_pattern T (arg :: args, e) =
+      trans_expr (Term.domain_type T) arg ##>>
+      trans_pattern (Term.range_type T) (args, e) #>>
+      (fn (arg', (args', e')) => (arg' :: args', e'))
+
+fun mk_fun_upd T U = Const (@{const_name fun_upd}, [T --> U, T, U, T] ---> U)
+
+fun mk_update ([], u) _ = u
+  | mk_update ([t], u) f =
+      uncurry mk_fun_upd (Term.dest_funT (Term.fastype_of f)) $ f $ t $ u
+  | mk_update (t :: ts, u) f =
+      let
+        val (dT, rT) = Term.dest_funT (Term.fastype_of f)
+        val (dT', rT') = Term.dest_funT rT
+      in
+        mk_fun_upd dT rT $ f $ t $
+          mk_update (ts, u) (absdummy dT' (Const ("_", rT')))
+      end
+
+fun mk_lambda Ts (t, pats) =
+  fold_rev absdummy Ts t |> fold mk_update pats
+
+fun translate ((t, k), (e, cs)) =
+  let
+    val T = Term.fastype_of t
+    val (Us, U) = SMT2_Util.dest_funT k (Term.fastype_of t)
+
+    fun mk_full_def u' pats =
+      pats
+      |> filter_out (fn (_, u) => u aconv u')
+      |> HOLogic.mk_eq o pair t o mk_lambda Us o pair u'
+
+    fun mk_eq (us, u) = HOLogic.mk_eq (Term.list_comb (t, us), u)
+    fun mk_eqs u' [] = [HOLogic.mk_eq (t, u')]
+      | mk_eqs _ pats = map mk_eq pats
+  in
+    trans_expr U e ##>>
+    (if k = 0 then pair [] else fold_map (trans_pattern T) cs) #>>
+    (fn (u', pats) => (mk_eqs u' pats, mk_full_def u' pats))
+  end
+
+
+(* normalization *)
+
+fun partition_eqs f =
+  let
+    fun part t (xs, ts) =
+      (case try HOLogic.dest_eq t of
+        SOME (l, r) => (case f l r of SOME x => (x::xs, ts) | _ => (xs, t::ts))
+      | NONE => (xs, t :: ts))
+  in (fn ts => fold part ts ([], [])) end
+
+fun first_eq pred =
+  let
+    fun part _ [] = NONE
+      | part us (t :: ts) =
+          (case try (pred o HOLogic.dest_eq) t of
+            SOME (SOME lr) => SOME (lr, fold cons us ts)
+          | _ => part (t :: us) ts)
+  in (fn ts => part [] ts) end
+
+fun replace_vars tab =
+  let
+    fun repl v = the_default v (AList.lookup (op aconv) tab v)
+    fun replace (v as Var _) = repl v
+      | replace (v as Free _) = repl v
+      | replace t = t
+  in map (Term.map_aterms replace) end
+
+fun remove_int_nat_coercions (eqs, defs) =
+  let
+    fun mk_nat_num t i =
+      (case try HOLogic.dest_number i of
+        SOME (_, n) => SOME (t, HOLogic.mk_number @{typ nat} n)
+      | NONE => NONE)
+    fun nat_of (@{const of_nat (int)} $ (t as Var _)) i = mk_nat_num t i
+      | nat_of (@{const nat} $ i) (t as Var _) = mk_nat_num t i
+      | nat_of _ _ = NONE
+    val (nats, eqs') = partition_eqs nat_of eqs
+
+    fun is_coercion t =
+      (case try HOLogic.dest_eq t of
+        SOME (@{const of_nat (int)}, _) => true
+      | SOME (@{const nat}, _) => true
+      | _ => false)
+  in pairself (replace_vars nats) (eqs', filter_out is_coercion defs) end
+
+fun unfold_funapp (eqs, defs) =
+  let
+    fun unfold_app (Const (@{const_name SMT2.fun_app}, _) $ f $ t) = f $ t
+      | unfold_app t = t
+    fun unfold_eq ((eq as Const (@{const_name HOL.eq}, _)) $ t $ u) =
+          eq $ unfold_app t $ u
+      | unfold_eq t = t
+
+    fun is_fun_app t =
+      (case try HOLogic.dest_eq t of
+        SOME (Const (@{const_name SMT2.fun_app}, _), _) => true
+      | _ => false)
+
+  in (map unfold_eq eqs, filter_out is_fun_app defs) end
+
+val unfold_eqs =
+  let
+    val is_ground = not o Term.exists_subterm Term.is_Var
+    fun is_non_rec (v, t) = not (Term.exists_subterm (equal v) t)
+
+    fun rewr_var (l as Var _, r) = if is_ground r then SOME (l, r) else NONE
+      | rewr_var (r, l as Var _) = if is_ground r then SOME (l, r) else NONE
+      | rewr_var _ = NONE
+
+    fun rewr_free' e = if is_non_rec e then SOME e else NONE
+    fun rewr_free (e as (Free _, _)) = rewr_free' e
+      | rewr_free (e as (_, Free _)) = rewr_free' (swap e)
+      | rewr_free _ = NONE
+
+    fun is_trivial (Const (@{const_name HOL.eq}, _) $ t $ u) = t aconv u
+      | is_trivial _ = false
+
+    fun replace r = replace_vars [r] #> filter_out is_trivial
+
+    fun unfold_vars (es, ds) =
+      (case first_eq rewr_var es of
+        SOME (lr, es') => unfold_vars (pairself (replace lr) (es', ds))
+      | NONE => (es, ds))
+
+    fun unfold_frees ues (es, ds) =
+      (case first_eq rewr_free es of
+        SOME (lr, es') =>
+          pairself (replace lr) (es', ds)
+          |> unfold_frees (HOLogic.mk_eq lr :: replace lr ues)
+      | NONE => (ues @ es, ds))
+
+  in unfold_vars #> unfold_frees [] end
+
+fun swap_free ((eq as Const (@{const_name HOL.eq}, _)) $ t $ (u as Free _)) =
+      eq $ u $ t
+  | swap_free t = t
+
+fun frees_for_vars ctxt (eqs, defs) =
+  let
+    fun fresh_free i T (cx as (frees, ctxt)) =
+      (case Inttab.lookup frees i of
+        SOME t => (t, cx)
+      | NONE =>
+          let
+            val (n, ctxt') = yield_singleton Variable.variant_fixes "" ctxt
+            val t = Free (n, T)
+          in (t, (Inttab.update (i, t) frees, ctxt')) end)
+
+    fun repl_var (Var ((_, i), T)) = fresh_free i T
+      | repl_var (t $ u) = repl_var t ##>> repl_var u #>> op $
+      | repl_var (Abs (n, T, t)) = repl_var t #>> (fn t' => Abs (n, T, t'))
+      | repl_var t = pair t
+  in
+    (Inttab.empty, ctxt)
+    |> fold_map repl_var eqs
+    ||>> fold_map repl_var defs
+    |> fst
+  end
+
+
+(* overall procedure *)
+
+val is_free_constraint = Term.exists_subterm (fn Free _ => true | _ => false)
+
+fun is_free_def (Const (@{const_name HOL.eq}, _) $ Free _ $ _) = true
+  | is_free_def _ = false
+
+fun defined tp =
+  try (pairself (fst o HOLogic.dest_eq)) tp
+  |> the_default false o Option.map (op aconv)
+
+fun add_free_defs free_cs defs =
+  let val (free_defs, defs') = List.partition is_free_def defs
+  in (free_cs @ filter_out (member defined free_cs) free_defs, defs') end
+
+fun is_const_def (Const (@{const_name HOL.eq}, _) $ Const _ $ _) = true
+  | is_const_def _ = false
+
+(* TODO: Adapt parser to SMT-LIB 2 format for models *)
+fun parse_counterex ctxt ({terms, ...} : SMT2_Translate.replay_data) ls =
+  read_cex terms ls
+  |> with_context terms translate
+  |> apfst flat o split_list
+  |> remove_int_nat_coercions
+  |> unfold_funapp
+  |> unfold_eqs
+  |>> map swap_free
+  |>> filter is_free_constraint
+  |-> add_free_defs
+  |> frees_for_vars ctxt
+  ||> filter is_const_def
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/SMT2/z3_new_proof.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -0,0 +1,563 @@
+(*  Title:      HOL/Tools/SMT2/z3_new_proof.ML
+    Author:     Sascha Boehme, TU Muenchen
+
+Z3 proofs: parsing and abstract syntax tree.
+*)
+
+signature Z3_NEW_PROOF =
+sig
+  (*proof rules*)
+  datatype z3_rule = True_Axiom | Asserted | Goal | Modus_Ponens | Reflexivity |
+    Symmetry | Transitivity | Transitivity_Star | Monotonicity | Quant_Intro |
+    Distributivity | And_Elim | Not_Or_Elim | Rewrite | Rewrite_Star |
+    Pull_Quant | Pull_Quant_Star | Push_Quant | Elim_Unused_Vars |
+    Dest_Eq_Res | Quant_Inst | Hypothesis | Lemma | Unit_Resolution |
+    Iff_True | Iff_False | Commutativity | Def_Axiom | Intro_Def | Apply_Def |
+    Iff_Oeq | Nnf_Pos | Nnf_Neg | Nnf_Star | Cnf_Star | Skolemize |
+    Modus_Ponens_Oeq | Th_Lemma of string
+  val string_of_rule: z3_rule -> string
+
+  (*proofs*)
+  datatype z3_step = Z3_Step of {
+    id: int,
+    rule: z3_rule,
+    prems: int list,
+    concl: term,
+    fixes: string list,
+    is_fix_step: bool}
+
+  (*type and term parsers*)
+  type type_parser = SMTLIB2.tree * typ list -> typ option
+  type term_parser = SMTLIB2.tree * term list -> term option
+  val add_type_parser: type_parser -> Context.generic -> Context.generic
+  val add_term_parser: term_parser -> Context.generic -> Context.generic
+
+  (*proof parser*)
+  val parse: typ Symtab.table -> term Symtab.table -> string list ->
+    Proof.context -> z3_step list * Proof.context
+end
+
+structure Z3_New_Proof: Z3_NEW_PROOF =
+struct
+
+(* proof rules *)
+
+datatype z3_rule = True_Axiom | Asserted | Goal | Modus_Ponens | Reflexivity |
+  Symmetry | Transitivity | Transitivity_Star | Monotonicity | Quant_Intro |
+  Distributivity | And_Elim | Not_Or_Elim | Rewrite | Rewrite_Star |
+  Pull_Quant | Pull_Quant_Star | Push_Quant | Elim_Unused_Vars | Dest_Eq_Res |
+  Quant_Inst | Hypothesis | Lemma | Unit_Resolution | Iff_True | Iff_False |
+  Commutativity | Def_Axiom | Intro_Def | Apply_Def | Iff_Oeq | Nnf_Pos |
+  Nnf_Neg | Nnf_Star | Cnf_Star | Skolemize | Modus_Ponens_Oeq |
+  Th_Lemma of string
+  (* TODO: some proof rules come with further information
+     that is currently dropped by the parser *)
+
+val rule_names = Symtab.make [
+  ("true-axiom", True_Axiom),
+  ("asserted", Asserted),
+  ("goal", Goal),
+  ("mp", Modus_Ponens),
+  ("refl", Reflexivity),
+  ("symm", Symmetry),
+  ("trans", Transitivity),
+  ("trans*", Transitivity_Star),
+  ("monotonicity", Monotonicity),
+  ("quant-intro", Quant_Intro),
+  ("distributivity", Distributivity),
+  ("and-elim", And_Elim),
+  ("not-or-elim", Not_Or_Elim),
+  ("rewrite", Rewrite),
+  ("rewrite*", Rewrite_Star),
+  ("pull-quant", Pull_Quant),
+  ("pull-quant*", Pull_Quant_Star),
+  ("push-quant", Push_Quant),
+  ("elim-unused", Elim_Unused_Vars),
+  ("der", Dest_Eq_Res),
+  ("quant-inst", Quant_Inst),
+  ("hypothesis", Hypothesis),
+  ("lemma", Lemma),
+  ("unit-resolution", Unit_Resolution),
+  ("iff-true", Iff_True),
+  ("iff-false", Iff_False),
+  ("commutativity", Commutativity),
+  ("def-axiom", Def_Axiom),
+  ("intro-def", Intro_Def),
+  ("apply-def", Apply_Def),
+  ("iff~", Iff_Oeq),
+  ("nnf-pos", Nnf_Pos),
+  ("nnf-neg", Nnf_Neg),
+  ("nnf*", Nnf_Star),
+  ("cnf*", Cnf_Star),
+  ("sk", Skolemize),
+  ("mp~", Modus_Ponens_Oeq)]
+
+fun rule_of_string name =
+  (case Symtab.lookup rule_names name of
+    SOME rule => rule
+  | NONE => error ("unknown Z3 proof rule " ^ quote name))
+
+fun string_of_rule (Th_Lemma kind) = "th-lemma " ^ kind
+  | string_of_rule r =
+      let fun eq_rule (s, r') = if r = r' then SOME s else NONE 
+      in the (Symtab.get_first eq_rule rule_names) end
+
+
+
+(* proofs *)
+
+datatype z3_node = Z3_Node of {
+  id: int,
+  rule: z3_rule,
+  prems: z3_node list,
+  concl: term,
+  bounds: string list}
+
+fun mk_node id rule prems concl bounds =
+  Z3_Node {id=id, rule=rule, prems=prems, concl=concl, bounds=bounds}
+
+datatype z3_step = Z3_Step of {
+  id: int,
+  rule: z3_rule,
+  prems: int list,
+  concl: term,
+  fixes: string list,
+  is_fix_step: bool}
+
+fun mk_step id rule prems concl fixes is_fix_step =
+  Z3_Step {id=id, rule=rule, prems=prems, concl=concl, fixes=fixes,
+    is_fix_step=is_fix_step}
+
+
+
+(* core type and term parser *)
+
+fun core_type_parser (SMTLIB2.Sym "Bool", []) = SOME @{typ HOL.bool}
+  | core_type_parser (SMTLIB2.Sym "Int", []) = SOME @{typ Int.int}
+  | core_type_parser _ = NONE
+
+fun mk_unary n t =
+  let val T = fastype_of t
+  in Const (n, T --> T) $ t end
+
+fun mk_binary' n T U t1 t2 = Const (n, [T, T] ---> U) $ t1 $ t2
+
+fun mk_binary n t1 t2 =
+  let val T = fastype_of t1
+  in mk_binary' n T T t1 t2 end
+
+fun mk_rassoc f t ts =
+  let val us = rev (t :: ts)
+  in fold f (tl us) (hd us) end
+
+fun mk_lassoc f t ts = fold (fn u1 => fn u2 => f u2 u1) ts t
+
+fun mk_lassoc' n = mk_lassoc (mk_binary n)
+
+fun mk_binary_pred n S t1 t2 =
+  let
+    val T1 = fastype_of t1
+    val T2 = fastype_of t2
+    val T =
+      if T1 <> Term.dummyT then T1
+      else if T2 <> Term.dummyT then T2
+      else TVar (("?a", serial ()), S)
+  in mk_binary' n T @{typ HOL.bool} t1 t2 end
+
+fun mk_less t1 t2 = mk_binary_pred @{const_name ord_class.less} @{sort linorder} t1 t2
+fun mk_less_eq t1 t2 = mk_binary_pred @{const_name ord_class.less_eq} @{sort linorder} t1 t2
+
+fun core_term_parser (SMTLIB2.Sym "true", _) = SOME @{const HOL.True}
+  | core_term_parser (SMTLIB2.Sym "false", _) = SOME @{const HOL.False}
+  | core_term_parser (SMTLIB2.Sym "not", [t]) = SOME (HOLogic.mk_not t)
+  | core_term_parser (SMTLIB2.Sym "and", t :: ts) = SOME (mk_rassoc (curry HOLogic.mk_conj) t ts)
+  | core_term_parser (SMTLIB2.Sym "or", t :: ts) = SOME (mk_rassoc (curry HOLogic.mk_disj) t ts)
+  | core_term_parser (SMTLIB2.Sym "=>", [t1, t2]) = SOME (HOLogic.mk_imp (t1, t2))
+  | core_term_parser (SMTLIB2.Sym "implies", [t1, t2]) = SOME (HOLogic.mk_imp (t1, t2))
+  | core_term_parser (SMTLIB2.Sym "=", [t1, t2]) = SOME (HOLogic.mk_eq (t1, t2))
+  | core_term_parser (SMTLIB2.Sym "~", [t1, t2]) = SOME (HOLogic.mk_eq (t1, t2))
+  | core_term_parser (SMTLIB2.Sym "ite", [t1, t2, t3]) =
+      let
+        val T = fastype_of t2
+        val c = Const (@{const_name HOL.If}, [@{typ HOL.bool}, T, T] ---> T)
+      in SOME (c $ t1 $ t2 $ t3) end
+  | core_term_parser (SMTLIB2.Num i, []) = SOME (HOLogic.mk_number @{typ Int.int} i)
+  | core_term_parser (SMTLIB2.Sym "-", [t]) = SOME (mk_unary @{const_name uminus_class.uminus} t)
+  | core_term_parser (SMTLIB2.Sym "~", [t]) = SOME (mk_unary @{const_name uminus_class.uminus} t)
+  | core_term_parser (SMTLIB2.Sym "+", t :: ts) =
+      SOME (mk_lassoc' @{const_name plus_class.plus} t ts)
+  | core_term_parser (SMTLIB2.Sym "-", t :: ts) =
+      SOME (mk_lassoc' @{const_name minus_class.minus} t ts)
+  | core_term_parser (SMTLIB2.Sym "*", t :: ts) =
+      SOME (mk_lassoc' @{const_name times_class.times} t ts)
+  | core_term_parser (SMTLIB2.Sym "div", [t1, t2]) = SOME (mk_binary @{const_name SMT2.z3div} t1 t2)
+  | core_term_parser (SMTLIB2.Sym "mod", [t1, t2]) = SOME (mk_binary @{const_name SMT2.z3mod} t1 t2)
+  | core_term_parser (SMTLIB2.Sym "<", [t1, t2]) = SOME (mk_less t1 t2)
+  | core_term_parser (SMTLIB2.Sym ">", [t1, t2]) = SOME (mk_less t2 t1)
+  | core_term_parser (SMTLIB2.Sym "<=", [t1, t2]) = SOME (mk_less_eq t1 t2)
+  | core_term_parser (SMTLIB2.Sym ">=", [t1, t2]) = SOME (mk_less_eq t2 t1)
+  | core_term_parser _ = NONE
+
+
+
+(* type and term parsers *)
+
+type type_parser = SMTLIB2.tree * typ list -> typ option
+
+type term_parser = SMTLIB2.tree * term list -> term option
+
+fun id_ord ((id1, _), (id2, _)) = int_ord (id1, id2)
+
+structure Parsers = Generic_Data
+(
+  type T = (int * type_parser) list * (int * term_parser) list
+  val empty : T = ([(serial (), core_type_parser)], [(serial (), core_term_parser)])
+  val extend = I
+  fun merge ((tys1, ts1), (tys2, ts2)) =
+    (Ord_List.merge id_ord (tys1, tys2), Ord_List.merge id_ord (ts1, ts2))
+)
+
+fun add_type_parser type_parser =
+  Parsers.map (apfst (Ord_List.insert id_ord (serial (), type_parser)))
+
+fun add_term_parser term_parser =
+  Parsers.map (apsnd (Ord_List.insert id_ord (serial (), term_parser)))
+
+fun get_type_parsers ctxt = map snd (fst (Parsers.get (Context.Proof ctxt)))
+fun get_term_parsers ctxt = map snd (snd (Parsers.get (Context.Proof ctxt)))
+
+fun apply_parsers parsers x =
+  let
+    fun apply [] = NONE
+      | apply (parser :: parsers) =
+          (case parser x of
+            SOME y => SOME y
+          | NONE => apply parsers)
+  in apply parsers end
+
+
+
+(* proof parser context *)
+
+datatype shared = Tree of SMTLIB2.tree | Term of term | Proof of z3_node | None
+
+type 'a context = {
+  ctxt: Proof.context,
+  id: int,
+  syms: shared Symtab.table,
+  typs: typ Symtab.table,
+  funs: term Symtab.table,
+  extra: 'a}
+
+fun mk_context ctxt id syms typs funs extra: 'a context =
+  {ctxt=ctxt, id=id, syms=syms, typs=typs, funs=funs, extra=extra}
+
+fun empty_context ctxt typs funs = mk_context ctxt 1 Symtab.empty typs funs []
+
+fun ctxt_of ({ctxt, ...}: 'a context) = ctxt
+
+fun next_id ({ctxt, id, syms, typs, funs, extra}: 'a context) =
+  (id, mk_context ctxt (id + 1) syms typs funs extra)
+
+fun lookup_binding ({syms, ...}: 'a context) =
+  the_default None o Symtab.lookup syms
+
+fun map_syms f ({ctxt, id, syms, typs, funs, extra}: 'a context) =
+  mk_context ctxt id (f syms) typs funs extra
+
+fun update_binding b = map_syms (Symtab.update b)
+
+fun with_bindings bs f cx =
+  let val bs' = map (lookup_binding cx o fst) bs
+  in
+    cx
+    |> fold update_binding bs
+    |> f
+    ||> fold2 (fn (name, _) => update_binding o pair name) bs bs'
+  end
+
+fun lookup_typ ({typs, ...}: 'a context) = Symtab.lookup typs
+fun lookup_fun ({funs, ...}: 'a context) = Symtab.lookup funs
+
+fun fresh_fun add name n T ({ctxt, id, syms, typs, funs, extra}: 'a context) =
+  let
+    val (n', ctxt') = yield_singleton Variable.variant_fixes n ctxt
+    val t = Free (n', T)
+    val funs' = Symtab.update (name, t) funs
+  in (t, mk_context ctxt' id syms typs funs' (add (n', T) extra)) end
+
+fun declare_fun name n T = snd o fresh_fun cons name n T
+fun declare_free name n T = fresh_fun (cons o pair name) name n T
+
+fun with_fresh_names f ({ctxt, id, syms, typs, funs, extra}: 'a context) =
+  let
+    fun bind (_, v as (_, T)) t = Logic.all_const T $ Term.absfree v t
+
+    val needs_inferT = equal Term.dummyT orf Term.is_TVar
+    val needs_infer = Term.exists_type (Term.exists_subtype needs_inferT)
+    fun infer_types ctxt =
+      singleton (Type_Infer_Context.infer_types ctxt) #>
+      singleton (Proof_Context.standard_term_check_finish ctxt)
+    fun infer ctxt t = if needs_infer t then infer_types ctxt t else t
+
+    type bindings = (string * (string * typ)) list
+    val (t, {ctxt=ctxt', extra=names, ...}: bindings context) =
+      f (mk_context ctxt id syms typs funs [])
+    val t' = infer ctxt' (fold_rev bind names (HOLogic.mk_Trueprop t))
+   
+  in ((t', map fst names), mk_context ctxt id syms typs funs extra) end
+
+
+
+(* proof parser *)
+
+exception Z3_PARSE of string * SMTLIB2.tree
+
+val desymbolize = Name.desymbolize false o perhaps (try (unprefix "?"))
+
+fun parse_type cx ty Ts =
+  (case apply_parsers (get_type_parsers (ctxt_of cx)) (ty, Ts) of
+    SOME T => T
+  | NONE =>
+      (case ty of
+        SMTLIB2.Sym name =>
+          (case lookup_typ cx name of
+            SOME T => T
+          | NONE => raise Z3_PARSE ("unknown Z3 type", ty))
+      | _ => raise Z3_PARSE ("bad Z3 type format", ty)))
+
+fun parse_term t ts cx =
+  (case apply_parsers (get_term_parsers (ctxt_of cx)) (t, ts) of
+    SOME u => (u, cx)
+  | NONE =>
+      (case t of
+        SMTLIB2.Sym name =>
+          (case lookup_fun cx name of
+            SOME u => (Term.list_comb (u, ts), cx)
+          | NONE =>
+              if null ts then declare_free name (desymbolize name) Term.dummyT cx
+              else raise Z3_PARSE ("bad Z3 term", t))
+      | _ => raise Z3_PARSE ("bad Z3 term format", t)))
+
+fun type_of cx ty =
+  (case try (parse_type cx ty) [] of
+    SOME T => T
+  | NONE =>
+      (case ty of
+        SMTLIB2.S (ty' :: tys) => parse_type cx ty' (map (type_of cx) tys)
+      | _ => raise Z3_PARSE ("bad Z3 type", ty)))
+
+fun dest_var cx (SMTLIB2.S [SMTLIB2.Sym name, ty]) = (name, (desymbolize name, type_of cx ty))
+  | dest_var _ v = raise Z3_PARSE ("bad Z3 quantifier variable format", v)
+
+fun dest_body (SMTLIB2.S (SMTLIB2.Sym "!" :: body :: _)) = dest_body body
+  | dest_body body = body
+
+fun dest_binding (SMTLIB2.S [SMTLIB2.Sym name, t]) = (name, Tree t)
+  | dest_binding b = raise Z3_PARSE ("bad Z3 let binding format", b)
+
+fun term_of t cx =
+  (case t of
+    SMTLIB2.S [SMTLIB2.Sym "forall", SMTLIB2.S vars, body] =>
+      quant HOLogic.mk_all vars body cx
+  | SMTLIB2.S [SMTLIB2.Sym "exists", SMTLIB2.S vars, body] =>
+      quant HOLogic.mk_exists vars body cx
+  | SMTLIB2.S [SMTLIB2.Sym "let", SMTLIB2.S bindings, body] =>
+      with_bindings (map dest_binding bindings) (term_of body) cx
+  | SMTLIB2.S (SMTLIB2.Sym "!" :: t :: _) => term_of t cx
+  | SMTLIB2.S (f :: args) =>
+      cx
+      |> fold_map term_of args
+      |-> parse_term f
+  | SMTLIB2.Sym name =>
+      (case lookup_binding cx name of
+        Tree u =>
+          cx
+          |> term_of u
+          |-> (fn u' => pair u' o update_binding (name, Term u'))
+      | Term u => (u, cx)
+      | None => parse_term t [] cx
+      | _ => raise Z3_PARSE ("bad Z3 term format", t))
+  | _ => parse_term t [] cx)
+
+and quant q vars body cx =
+  let val vs = map (dest_var cx) vars
+  in
+    cx
+    |> with_bindings (map (apsnd (Term o Free)) vs) (term_of (dest_body body))
+    |>> fold_rev (fn (_, (n, T)) => fn t => q (n, T, t)) vs
+  end
+
+fun rule_of (SMTLIB2.Sym name) = rule_of_string name
+  | rule_of (SMTLIB2.S (SMTLIB2.Sym "_" :: SMTLIB2.Sym name :: args)) =
+      (case (name, args) of
+        ("th-lemma", SMTLIB2.Sym kind :: _) => Th_Lemma kind
+      | _ => rule_of_string name)
+  | rule_of r = raise Z3_PARSE ("bad Z3 proof rule format", r)
+
+fun node_of p cx =
+  (case p of
+    SMTLIB2.Sym name =>
+      (case lookup_binding cx name of
+        Proof node => (node, cx)
+      | Tree p' =>
+          cx
+          |> node_of p'
+          |-> (fn node => pair node o update_binding (name, Proof node))
+      | _ => raise Z3_PARSE ("bad Z3 proof format", p))
+  | SMTLIB2.S [SMTLIB2.Sym "let", SMTLIB2.S bindings, p] =>
+      with_bindings (map dest_binding bindings) (node_of p) cx
+  | SMTLIB2.S (name :: parts) =>
+      let
+        val (ps, p) = split_last parts
+        val r = rule_of name
+      in
+        cx
+        |> fold_map node_of ps
+        ||>> with_fresh_names (term_of p)
+        ||>> next_id
+        |>> (fn ((prems, (t, ns)), id) => mk_node id r prems t ns)
+      end
+  | _ => raise Z3_PARSE ("bad Z3 proof format", p))
+
+fun dest_name (SMTLIB2.Sym name) = name
+  | dest_name t = raise Z3_PARSE ("bad name", t)
+
+fun dest_seq (SMTLIB2.S ts) = ts
+  | dest_seq t = raise Z3_PARSE ("bad Z3 proof format", t)
+
+fun parse' (SMTLIB2.S (SMTLIB2.Sym "set-logic" :: _) :: ts) cx = parse' ts cx
+  | parse' (SMTLIB2.S [SMTLIB2.Sym "declare-fun", n, tys, ty] :: ts) cx =
+      let
+        val name = dest_name n
+        val Ts = map (type_of cx) (dest_seq tys)
+        val T = type_of cx ty
+      in parse' ts (declare_fun name (desymbolize name) (Ts ---> T) cx) end
+  | parse' (SMTLIB2.S [SMTLIB2.Sym "proof", p] :: _) cx = node_of p cx
+  | parse' ts _ = raise Z3_PARSE ("bad Z3 proof declarations", SMTLIB2.S ts)
+
+fun parse_proof typs funs lines ctxt =
+  let
+    val ts = dest_seq (SMTLIB2.parse lines)
+    val (node, cx) = parse' ts (empty_context ctxt typs funs)
+  in (node, ctxt_of cx) end
+  handle SMTLIB2.PARSE (l, msg) =>
+           error ("parsing error at line " ^ string_of_int l ^ ": " ^ msg)
+       | Z3_PARSE (msg, t) =>
+           error (msg ^ ": " ^ SMTLIB2.str_of t)
+
+
+
+(* handling of bound variables *)
+
+fun subst_of tyenv =
+  let fun add (ix, (S, T)) = cons (TVar (ix, S), T)
+  in Vartab.fold add tyenv [] end
+
+fun substTs_same subst = 
+  let val applyT = Same.function (AList.lookup (op =) subst)
+  in Term_Subst.map_atypsT_same applyT end
+
+fun subst_types ctxt env bounds t =
+  let
+    val match = Sign.typ_match (Proof_Context.theory_of ctxt)
+
+    val t' = singleton (Variable.polymorphic ctxt) t
+    val patTs = map snd (Term.strip_qnt_vars @{const_name all} t')
+    val objTs = map (the o Symtab.lookup env) bounds
+    val subst = subst_of (fold match (patTs ~~ objTs) Vartab.empty)
+  in Same.commit (Term_Subst.map_types_same (substTs_same subst)) t' end
+
+fun eq_quant (@{const_name HOL.All}, _) (@{const_name HOL.All}, _) = true
+  | eq_quant (@{const_name HOL.Ex}, _) (@{const_name HOL.Ex}, _) = true
+  | eq_quant _ _ = false
+
+fun opp_quant (@{const_name HOL.All}, _) (@{const_name HOL.Ex}, _) = true
+  | opp_quant (@{const_name HOL.Ex}, _) (@{const_name HOL.All}, _) = true
+  | opp_quant _ _ = false
+
+fun with_quant pred i (Const q1 $ Abs (_, T1, t1), Const q2 $ Abs (_, T2, t2)) =
+      if pred q1 q2 andalso T1 = T2 then
+        let val t = Var (("", i), T1)
+        in SOME (pairself Term.subst_bound ((t, t1), (t, t2))) end
+      else NONE
+  | with_quant _ _ _ = NONE
+
+fun dest_quant_pair i (@{term HOL.Not} $ t1, t2) =
+      Option.map (apfst HOLogic.mk_not) (with_quant opp_quant i (t1, t2))
+  | dest_quant_pair i (t1, t2) = with_quant eq_quant i (t1, t2)
+
+fun dest_quant i t =
+  (case dest_quant_pair i (HOLogic.dest_eq (HOLogic.dest_Trueprop t)) of
+    SOME (t1, t2) => HOLogic.mk_Trueprop (HOLogic.mk_eq (t1, t2))
+  | NONE => raise TERM ("lift_quant", [t]))
+
+fun match_types ctxt pat obj =
+  (Vartab.empty, Vartab.empty)
+  |> Pattern.first_order_match (Proof_Context.theory_of ctxt) (pat, obj)
+
+fun strip_match ctxt pat i obj =
+  (case try (match_types ctxt pat) obj of
+    SOME (tyenv, _) => subst_of tyenv
+  | NONE => strip_match ctxt pat (i + 1) (dest_quant i obj))
+
+fun dest_all i (Const (@{const_name all}, _) $ (a as Abs (_, T, _))) =
+      dest_all (i + 1) (Term.betapply (a, Var (("", i), T)))
+  | dest_all i t = (i, t)
+
+fun dest_alls t = dest_all (Term.maxidx_of_term t + 1) t
+
+fun match_rule ctxt env (Z3_Node {bounds=bs', concl=t', ...}) bs t =
+  let
+    val t'' = singleton (Variable.polymorphic ctxt) t'
+    val (i, obj) = dest_alls (subst_types ctxt env bs t)
+  in
+    (case try (strip_match ctxt (snd (dest_alls t'')) i) obj of
+      NONE => NONE
+    | SOME subst =>
+        let
+          val applyT = Same.commit (substTs_same subst)
+          val patTs = map snd (Term.strip_qnt_vars @{const_name all} t'')
+        in SOME (Symtab.make (bs' ~~ map applyT patTs)) end)
+  end
+
+
+
+(* linearizing proofs and resolving types of bound variables *)
+
+fun has_step (tab, _) = Inttab.defined tab
+
+fun add_step id rule bounds concl is_fix_step ids (tab, sts) =
+  let val step = mk_step id rule ids concl bounds is_fix_step
+  in (id, (Inttab.update (id, ()) tab, step :: sts)) end
+
+fun is_fix_rule rule prems =
+  member (op =) [Quant_Intro, Nnf_Pos, Nnf_Neg] rule andalso length prems = 1
+
+fun lin_proof ctxt env (Z3_Node {id, rule, prems, concl, bounds}) steps =
+  if has_step steps id then (id, steps)
+  else
+    let
+      val t = subst_types ctxt env bounds concl
+      val add = add_step id rule bounds t
+      fun rec_apply e b = fold_map (lin_proof ctxt e) prems #-> add b
+    in
+      if is_fix_rule rule prems then
+        (case match_rule ctxt env (hd prems) bounds t of
+          NONE => rec_apply env false steps
+        | SOME env' => rec_apply env' true steps)
+      else rec_apply env false steps
+    end
+
+fun linearize ctxt node =
+  rev (snd (snd (lin_proof ctxt Symtab.empty node (Inttab.empty, []))))
+
+
+
+(* overall proof parser *)
+
+fun parse typs funs lines ctxt =
+  let val (node, ctxt') = parse_proof typs funs lines ctxt
+  in (linearize ctxt' node, ctxt') end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/SMT2/z3_new_real.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -0,0 +1,32 @@
+(*  Title:      HOL/Tools/SMT2/z3_new_real.ML
+    Author:     Sascha Boehme, TU Muenchen
+
+Z3 setup for reals.
+*)
+
+structure Z3_New_Real: sig end =
+struct
+
+fun real_type_parser (SMTLIB2.Sym "Real", []) = SOME @{typ Real.real}
+  | real_type_parser _ = NONE
+
+fun real_term_parser (SMTLIB2.Dec (i, 0), []) = SOME (HOLogic.mk_number @{typ Real.real} i)
+  | real_term_parser (SMTLIB2.Sym "/", [t1, t2]) =
+      SOME (@{term "inverse_class.divide :: real => _"} $ t1 $ t2)
+  | real_term_parser (SMTLIB2.Sym "to_real", [t]) = SOME (@{term "Real.real :: int => _"} $ t)
+  | real_term_parser _ = NONE
+
+fun abstract abs t =
+  (case t of
+    (c as @{term "inverse_class.divide :: real => _"}) $ t1 $ t2 =>
+      abs t1 ##>> abs t2 #>> (fn (u1, u2) => SOME (c $ u1 $ u2))
+  | (c as @{term "Real.real :: int => _"}) $ t =>
+      abs t #>> (fn u => SOME (c $ u))
+  | _ => pair NONE)
+
+val _ = Theory.setup (Context.theory_map (
+  Z3_New_Proof.add_type_parser real_type_parser #>
+  Z3_New_Proof.add_term_parser real_term_parser #>
+  Z3_New_Replay_Methods.add_arith_abstracter abstract))
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/SMT2/z3_new_replay.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -0,0 +1,192 @@
+(*  Title:      HOL/Tools/SMT2/z3_new_replay.ML
+    Author:     Sascha Boehme, TU Muenchen
+    Author:     Jasmin Blanchette, TU Muenchen
+
+Z3 proof replay.
+*)
+
+signature Z3_NEW_REPLAY =
+sig
+  val replay: Proof.context -> SMT2_Translate.replay_data -> string list ->
+    ((int * (int * thm)) list * Z3_New_Proof.z3_step list) * thm
+end
+
+structure Z3_New_Replay: Z3_NEW_REPLAY =
+struct
+
+fun params_of t = Term.strip_qnt_vars @{const_name all} t
+
+fun varify ctxt thm =
+  let
+    val maxidx = Thm.maxidx_of thm + 1
+    val vs = params_of (Thm.prop_of thm)
+    val vars = map_index (fn (i, (n, T)) => Var ((n, i + maxidx), T)) vs
+  in Drule.forall_elim_list (map (SMT2_Util.certify ctxt) vars) thm end
+
+fun add_paramTs names t =
+  fold2 (fn n => fn (_, T) => AList.update (op =) (n, T)) names (params_of t)
+
+fun new_fixes ctxt nTs =
+  let
+    val (ns, ctxt') = Variable.variant_fixes (replicate (length nTs) "") ctxt
+    fun mk (n, T) n' = (n, SMT2_Util.certify ctxt' (Free (n', T)))
+  in (ctxt', Symtab.make (map2 mk nTs ns)) end
+
+fun forall_elim_term ct (Const (@{const_name all}, _) $ (a as Abs _)) =
+      Term.betapply (a, Thm.term_of ct)
+  | forall_elim_term _ qt = raise TERM ("forall_elim'", [qt])
+
+fun apply_fixes elim env = fold (elim o the o Symtab.lookup env)
+
+val apply_fixes_prem = uncurry o apply_fixes Thm.forall_elim
+val apply_fixes_concl = apply_fixes forall_elim_term
+
+fun export_fixes env names = Drule.forall_intr_list (map (the o Symtab.lookup env) names)
+
+fun under_fixes f ctxt (prems, nthms) names concl =
+  let
+    val thms1 = map (varify ctxt) prems
+    val (ctxt', env) =
+      add_paramTs names concl []
+      |> fold (uncurry add_paramTs o apsnd Thm.prop_of) nthms
+      |> new_fixes ctxt
+    val thms2 = map (apply_fixes_prem env) nthms
+    val t = apply_fixes_concl env names concl
+  in export_fixes env names (f ctxt' (thms1 @ thms2) t) end
+
+fun replay_thm ctxt assumed nthms
+    (Z3_New_Proof.Z3_Step {id, rule, concl, fixes, is_fix_step, ...}) =
+  if Z3_New_Replay_Methods.is_assumption rule then
+    (case Inttab.lookup assumed id of
+      SOME (_, thm) => thm
+    | NONE => Thm.assume (SMT2_Util.certify ctxt concl))
+  else
+    under_fixes (Z3_New_Replay_Methods.method_for rule) ctxt
+      (if is_fix_step then (map snd nthms, []) else ([], nthms)) fixes concl
+
+fun replay_step ctxt assumed (step as Z3_New_Proof.Z3_Step {id, prems, fixes, ...}) proofs =
+  let val nthms = map (the o Inttab.lookup proofs) prems
+  in Inttab.update (id, (fixes, replay_thm ctxt assumed nthms step)) proofs end
+
+local
+  val remove_trigger = mk_meta_eq @{thm SMT2.trigger_def}
+  val remove_weight = mk_meta_eq @{thm SMT2.weight_def}
+  val remove_fun_app = mk_meta_eq @{thm SMT2.fun_app_def}
+
+  fun rewrite_conv _ [] = Conv.all_conv
+    | rewrite_conv ctxt eqs = Simplifier.full_rewrite (empty_simpset ctxt addsimps eqs)
+
+  val prep_rules = [@{thm Let_def}, remove_trigger, remove_weight,
+    remove_fun_app, Z3_New_Replay_Literals.rewrite_true]
+
+  fun rewrite _ [] = I
+    | rewrite ctxt eqs = Conv.fconv_rule (rewrite_conv ctxt eqs)
+
+  fun lookup_assm assms_net ct =
+    Z3_New_Replay_Util.net_instances assms_net ct
+    |> map (fn ithm as (_, thm) => (ithm, Thm.cprop_of thm aconvc ct))
+in
+
+fun add_asserted outer_ctxt rewrite_rules assms steps ctxt =
+  let
+    val eqs = map (rewrite ctxt [Z3_New_Replay_Literals.rewrite_true]) rewrite_rules
+    val eqs' = union Thm.eq_thm eqs prep_rules
+
+    val assms_net =
+      assms
+      |> map (apsnd (rewrite ctxt eqs'))
+      |> map (apsnd (Conv.fconv_rule Thm.eta_conversion))
+      |> Z3_New_Replay_Util.thm_net_of snd 
+
+    fun revert_conv ctxt = rewrite_conv ctxt eqs' then_conv Thm.eta_conversion
+
+    fun assume thm ctxt =
+      let
+        val ct = Thm.cprem_of thm 1
+        val (thm', ctxt') = yield_singleton Assumption.add_assumes ct ctxt
+      in (thm' RS thm, ctxt') end
+
+    fun add1 id fixes thm1 ((i, th), exact) ((iidths, thms), (ctxt, ptab)) =
+      let
+        val (thm, ctxt') = if exact then (Thm.implies_elim thm1 th, ctxt) else assume thm1 ctxt
+        val thms' = if exact then thms else th :: thms
+      in (((i, (id, th)) :: iidths, thms'), (ctxt', Inttab.update (id, (fixes, thm)) ptab)) end
+
+    fun add (Z3_New_Proof.Z3_Step {id, rule, concl, fixes, ...})
+        (cx as ((iidths, thms), (ctxt, ptab))) =
+      if Z3_New_Replay_Methods.is_assumption rule andalso rule <> Z3_New_Proof.Hypothesis then
+        let
+          val ct = SMT2_Util.certify ctxt concl
+          val thm1 = Thm.trivial ct |> Conv.fconv_rule (Conv.arg1_conv (revert_conv outer_ctxt))
+          val thm2 = singleton (Variable.export ctxt outer_ctxt) thm1
+        in
+          (case lookup_assm assms_net (Thm.cprem_of thm2 1) of
+            [] =>
+              let val (thm, ctxt') = assume thm1 ctxt
+              in ((iidths, thms), (ctxt', Inttab.update (id, (fixes, thm)) ptab)) end
+          | ithms => fold (add1 id fixes thm1) ithms cx)
+        end
+      else
+        cx
+  in fold add steps (([], []), (ctxt, Inttab.empty)) end
+
+end
+
+(* |- (EX x. P x) = P c     |- ~ (ALL x. P x) = ~ P c *)
+local
+  val sk_rules = @{lemma
+    "c = (SOME x. P x) ==> (EX x. P x) = P c"
+    "c = (SOME x. ~ P x) ==> (~ (ALL x. P x)) = (~ P c)"
+    by (metis someI_ex)+}
+in
+
+fun discharge_sk_tac i st =
+  (rtac @{thm trans} i
+   THEN resolve_tac sk_rules i
+   THEN (rtac @{thm refl} ORELSE' discharge_sk_tac) (i+1)
+   THEN rtac @{thm refl} i) st
+
+end
+
+fun make_discharge_rules rules = rules @ [@{thm allI}, @{thm refl},
+  @{thm reflexive}, Z3_New_Replay_Literals.true_thm]
+
+val intro_def_rules = @{lemma
+  "(~ P | P) & (P | ~ P)"
+  "(P | ~ P) & (~ P | P)"
+  by fast+}
+
+fun discharge_assms_tac rules =
+  REPEAT (HEADGOAL (resolve_tac (intro_def_rules @ rules) ORELSE' SOLVED' discharge_sk_tac))
+  
+fun discharge_assms ctxt rules thm =
+  (if Thm.nprems_of thm = 0 then
+     thm
+   else
+     (case Seq.pull (discharge_assms_tac rules thm) of
+       SOME (thm', _) => thm'
+     | NONE => raise THM ("failed to discharge premise", 1, [thm])))
+  |> Goal.norm_result ctxt
+
+fun discharge rules outer_ctxt inner_ctxt =
+  singleton (Proof_Context.export inner_ctxt outer_ctxt)
+  #> discharge_assms outer_ctxt (make_discharge_rules rules)
+
+fun replay outer_ctxt
+    ({context=ctxt, typs, terms, rewrite_rules, assms} : SMT2_Translate.replay_data) output =
+  let
+    val (steps, ctxt2) = Z3_New_Proof.parse typs terms output ctxt
+    val ((iidths, rules), (ctxt3, assumed)) = add_asserted outer_ctxt rewrite_rules assms steps ctxt2
+  in
+    if Config.get ctxt3 SMT2_Config.filter_only_facts then
+      ((iidths, steps), TrueI)
+    else
+      let
+        val ctxt4 = put_simpset (Z3_New_Replay_Util.make_simpset ctxt3 []) ctxt3
+        val proofs = fold (replay_step ctxt4 assumed) steps assumed
+        val (_, Z3_New_Proof.Z3_Step {id, ...}) = split_last steps
+        val thm = Inttab.lookup proofs id |> the |> snd |> discharge rules outer_ctxt ctxt4
+      in (([], steps), thm) end
+  end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/SMT2/z3_new_replay_literals.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -0,0 +1,356 @@
+(*  Title:      HOL/Tools/SMT2/z3_new_replay_literals.ML
+    Author:     Sascha Boehme, TU Muenchen
+
+Proof tools related to conjunctions and disjunctions.
+*)
+
+signature Z3_NEW_REPLAY_LITERALS =
+sig
+  (*literal table*)
+  type littab = thm Termtab.table
+  val make_littab: thm list -> littab
+  val insert_lit: thm -> littab -> littab
+  val delete_lit: thm -> littab -> littab
+  val lookup_lit: littab -> term -> thm option
+  val get_first_lit: (term -> bool) -> littab -> thm option
+
+  (*rules*)
+  val true_thm: thm
+  val rewrite_true: thm
+
+  (*properties*)
+  val is_conj: term -> bool
+  val is_disj: term -> bool
+  val exists_lit: bool -> (term -> bool) -> term -> bool
+  val negate: cterm -> cterm
+
+  (*proof tools*)
+  val explode: bool -> bool -> bool -> term list -> thm -> thm list
+  val join: bool -> littab -> term -> thm
+  val prove_conj_disj_eq: cterm -> thm
+end
+
+structure Z3_New_Replay_Literals: Z3_NEW_REPLAY_LITERALS =
+struct
+
+
+
+(* literal table *)
+
+type littab = thm Termtab.table
+
+fun make_littab thms = fold (Termtab.update o `SMT2_Util.prop_of) thms Termtab.empty
+
+fun insert_lit thm = Termtab.update (`SMT2_Util.prop_of thm)
+fun delete_lit thm = Termtab.delete (SMT2_Util.prop_of thm)
+fun lookup_lit lits = Termtab.lookup lits
+fun get_first_lit f =
+  Termtab.get_first (fn (t, thm) => if f t then SOME thm else NONE)
+
+
+
+(* rules *)
+
+val true_thm = @{lemma "~False" by simp}
+val rewrite_true = @{lemma "True == ~ False" by simp}
+
+
+
+(* properties and term operations *)
+
+val is_neg = (fn @{const Not} $ _ => true | _ => false)
+fun is_neg' f = (fn @{const Not} $ t => f t | _ => false)
+val is_dneg = is_neg' is_neg
+val is_conj = (fn @{const HOL.conj} $ _ $ _ => true | _ => false)
+val is_disj = (fn @{const HOL.disj} $ _ $ _ => true | _ => false)
+
+fun dest_disj_term' f = (fn
+    @{const Not} $ (@{const HOL.disj} $ t $ u) => SOME (f t, f u)
+  | _ => NONE)
+
+val dest_conj_term = (fn @{const HOL.conj} $ t $ u => SOME (t, u) | _ => NONE)
+val dest_disj_term =
+  dest_disj_term' (fn @{const Not} $ t => t | t => @{const Not} $ t)
+
+fun exists_lit is_conj P =
+  let
+    val dest = if is_conj then dest_conj_term else dest_disj_term
+    fun exists t = P t orelse
+      (case dest t of
+        SOME (t1, t2) => exists t1 orelse exists t2
+      | NONE => false)
+  in exists end
+
+val negate = Thm.apply (Thm.cterm_of @{theory} @{const Not})
+
+
+
+(* proof tools *)
+
+(** explosion of conjunctions and disjunctions **)
+
+local
+  val precomp = Z3_New_Replay_Util.precompose2
+
+  fun destc ct = Thm.dest_binop (Thm.dest_arg ct)
+  val dest_conj1 = precomp destc @{thm conjunct1}
+  val dest_conj2 = precomp destc @{thm conjunct2}
+  fun dest_conj_rules t =
+    dest_conj_term t |> Option.map (K (dest_conj1, dest_conj2))
+    
+  fun destd f ct = f (Thm.dest_binop (Thm.dest_arg (Thm.dest_arg ct)))
+  val dn1 = apfst Thm.dest_arg and dn2 = apsnd Thm.dest_arg
+  val dest_disj1 = precomp (destd I) @{lemma "~(P | Q) ==> ~P" by fast}
+  val dest_disj2 = precomp (destd dn1) @{lemma "~(~P | Q) ==> P" by fast}
+  val dest_disj3 = precomp (destd I) @{lemma "~(P | Q) ==> ~Q" by fast}
+  val dest_disj4 = precomp (destd dn2) @{lemma "~(P | ~Q) ==> Q" by fast}
+
+  fun dest_disj_rules t =
+    (case dest_disj_term' is_neg t of
+      SOME (true, true) => SOME (dest_disj2, dest_disj4)
+    | SOME (true, false) => SOME (dest_disj2, dest_disj3)
+    | SOME (false, true) => SOME (dest_disj1, dest_disj4)
+    | SOME (false, false) => SOME (dest_disj1, dest_disj3)
+    | NONE => NONE)
+
+  fun destn ct = [Thm.dest_arg (Thm.dest_arg (Thm.dest_arg ct))]
+  val dneg_rule = Z3_New_Replay_Util.precompose destn @{thm notnotD}
+in
+
+(*
+  explode a term into literals and collect all rules to be able to deduce
+  particular literals afterwards
+*)
+fun explode_term is_conj =
+  let
+    val dest = if is_conj then dest_conj_term else dest_disj_term
+    val dest_rules = if is_conj then dest_conj_rules else dest_disj_rules
+
+    fun add (t, rs) = Termtab.map_default (t, rs)
+      (fn rs' => if length rs' < length rs then rs' else rs)
+
+    fun explode1 rules t =
+      (case dest t of
+        SOME (t1, t2) =>
+          let val (rule1, rule2) = the (dest_rules t)
+          in
+            explode1 (rule1 :: rules) t1 #>
+            explode1 (rule2 :: rules) t2 #>
+            add (t, rev rules)
+          end
+      | NONE => add (t, rev rules))
+
+    fun explode0 (@{const Not} $ (@{const Not} $ t)) =
+          Termtab.make [(t, [dneg_rule])]
+      | explode0 t = explode1 [] t Termtab.empty
+
+  in explode0 end
+
+(*
+  extract a literal by applying previously collected rules
+*)
+fun extract_lit thm rules = fold Z3_New_Replay_Util.compose rules thm
+
+
+(*
+  explode a theorem into its literals
+*)
+fun explode is_conj full keep_intermediate stop_lits =
+  let
+    val dest_rules = if is_conj then dest_conj_rules else dest_disj_rules
+    val tab = fold (Termtab.update o rpair ()) stop_lits Termtab.empty
+
+    fun explode1 thm =
+      if Termtab.defined tab (SMT2_Util.prop_of thm) then cons thm
+      else
+        (case dest_rules (SMT2_Util.prop_of thm) of
+          SOME (rule1, rule2) =>
+            explode2 rule1 thm #>
+            explode2 rule2 thm #>
+            keep_intermediate ? cons thm
+        | NONE => cons thm)
+
+    and explode2 dest_rule thm =
+      if full orelse
+        exists_lit is_conj (Termtab.defined tab) (SMT2_Util.prop_of thm)
+      then explode1 (Z3_New_Replay_Util.compose dest_rule thm)
+      else cons (Z3_New_Replay_Util.compose dest_rule thm)
+
+    fun explode0 thm =
+      if not is_conj andalso is_dneg (SMT2_Util.prop_of thm)
+      then [Z3_New_Replay_Util.compose dneg_rule thm]
+      else explode1 thm []
+
+  in explode0 end
+
+end
+
+
+(** joining of literals to conjunctions or disjunctions **)
+
+local
+  fun on_cprem i f thm = f (Thm.cprem_of thm i)
+  fun on_cprop f thm = f (Thm.cprop_of thm)
+  fun precomp2 f g thm = (on_cprem 1 f thm, on_cprem 2 g thm, f, g, thm)
+  fun comp2 (cv1, cv2, f, g, rule) thm1 thm2 =
+    Thm.instantiate ([], [(cv1, on_cprop f thm1), (cv2, on_cprop g thm2)]) rule
+    |> Z3_New_Replay_Util.discharge thm1 |> Z3_New_Replay_Util.discharge thm2
+
+  fun d1 ct = Thm.dest_arg ct and d2 ct = Thm.dest_arg (Thm.dest_arg ct)
+
+  val conj_rule = precomp2 d1 d1 @{thm conjI}
+  fun comp_conj ((_, thm1), (_, thm2)) = comp2 conj_rule thm1 thm2
+
+  val disj1 = precomp2 d2 d2 @{lemma "~P ==> ~Q ==> ~(P | Q)" by fast}
+  val disj2 = precomp2 d2 d1 @{lemma "~P ==> Q ==> ~(P | ~Q)" by fast}
+  val disj3 = precomp2 d1 d2 @{lemma "P ==> ~Q ==> ~(~P | Q)" by fast}
+  val disj4 = precomp2 d1 d1 @{lemma "P ==> Q ==> ~(~P | ~Q)" by fast}
+
+  fun comp_disj ((false, thm1), (false, thm2)) = comp2 disj1 thm1 thm2
+    | comp_disj ((false, thm1), (true, thm2)) = comp2 disj2 thm1 thm2
+    | comp_disj ((true, thm1), (false, thm2)) = comp2 disj3 thm1 thm2
+    | comp_disj ((true, thm1), (true, thm2)) = comp2 disj4 thm1 thm2
+
+  fun dest_conj (@{const HOL.conj} $ t $ u) = ((false, t), (false, u))
+    | dest_conj t = raise TERM ("dest_conj", [t])
+
+  val neg = (fn @{const Not} $ t => (true, t) | t => (false, @{const Not} $ t))
+  fun dest_disj (@{const Not} $ (@{const HOL.disj} $ t $ u)) = (neg t, neg u)
+    | dest_disj t = raise TERM ("dest_disj", [t])
+
+  val precomp = Z3_New_Replay_Util.precompose
+  val dnegE = precomp (single o d2 o d1) @{thm notnotD}
+  val dnegI = precomp (single o d1) @{lemma "P ==> ~~P" by fast}
+  fun as_dneg f t = f (@{const Not} $ (@{const Not} $ t))
+
+  val precomp2 = Z3_New_Replay_Util.precompose2
+  fun dni f = apsnd f o Thm.dest_binop o f o d1
+  val negIffE = precomp2 (dni d1) @{lemma "~(P = (~Q)) ==> Q = P" by fast}
+  val negIffI = precomp2 (dni I) @{lemma "P = Q ==> ~(Q = (~P))" by fast}
+  val iff_const = @{const HOL.eq (bool)}
+  fun as_negIff f (@{const HOL.eq (bool)} $ t $ u) =
+        f (@{const Not} $ (iff_const $ u $ (@{const Not} $ t)))
+    | as_negIff _ _ = NONE
+in
+
+fun join is_conj littab t =
+  let
+    val comp = if is_conj then comp_conj else comp_disj
+    val dest = if is_conj then dest_conj else dest_disj
+
+    val lookup = lookup_lit littab
+
+    fun lookup_rule t =
+      (case t of
+        @{const Not} $ (@{const Not} $ t) => (Z3_New_Replay_Util.compose dnegI, lookup t)
+      | @{const Not} $ (@{const HOL.eq (bool)} $ t $ (@{const Not} $ u)) =>
+          (Z3_New_Replay_Util.compose negIffI, lookup (iff_const $ u $ t))
+      | @{const Not} $ ((eq as Const (@{const_name HOL.eq}, _)) $ t $ u) =>
+          let fun rewr lit = lit COMP @{thm not_sym}
+          in (rewr, lookup (@{const Not} $ (eq $ u $ t))) end
+      | _ =>
+          (case as_dneg lookup t of
+            NONE => (Z3_New_Replay_Util.compose negIffE, as_negIff lookup t)
+          | x => (Z3_New_Replay_Util.compose dnegE, x)))
+
+    fun join1 (s, t) =
+      (case lookup t of
+        SOME lit => (s, lit)
+      | NONE => 
+          (case lookup_rule t of
+            (rewrite, SOME lit) => (s, rewrite lit)
+          | (_, NONE) => (s, comp (pairself join1 (dest t)))))
+
+  in snd (join1 (if is_conj then (false, t) else (true, t))) end
+
+end
+
+
+(** proving equality of conjunctions or disjunctions **)
+
+fun iff_intro thm1 thm2 = thm2 COMP (thm1 COMP @{thm iffI})
+
+local
+  val cp1 = @{lemma "(~P) = (~Q) ==> P = Q" by simp}
+  val cp2 = @{lemma "(~P) = Q ==> P = (~Q)" by fastforce}
+  val cp3 = @{lemma "P = (~Q) ==> (~P) = Q" by simp}
+in
+fun contrapos1 prove (ct, cu) = prove (negate ct, negate cu) COMP cp1
+fun contrapos2 prove (ct, cu) = prove (negate ct, Thm.dest_arg cu) COMP cp2
+fun contrapos3 prove (ct, cu) = prove (Thm.dest_arg ct, negate cu) COMP cp3
+end
+
+local
+  val contra_rule = @{lemma "P ==> ~P ==> False" by (rule notE)}
+  fun contra_left conj thm =
+    let
+      val rules = explode_term conj (SMT2_Util.prop_of thm)
+      fun contra_lits (t, rs) =
+        (case t of
+          @{const Not} $ u => Termtab.lookup rules u |> Option.map (pair rs)
+        | _ => NONE)
+    in
+      (case Termtab.lookup rules @{const False} of
+        SOME rs => extract_lit thm rs
+      | NONE =>
+          the (Termtab.get_first contra_lits rules)
+          |> pairself (extract_lit thm)
+          |> (fn (nlit, plit) => nlit COMP (plit COMP contra_rule)))
+    end
+
+  val falseE_v = Thm.dest_arg (Thm.dest_arg (Thm.cprop_of @{thm FalseE}))
+  fun contra_right ct = Thm.instantiate ([], [(falseE_v, ct)]) @{thm FalseE}
+in
+
+fun contradict conj ct =
+  iff_intro (Z3_New_Replay_Util.under_assumption (contra_left conj) ct) (contra_right ct)
+
+end
+
+local
+  fun prove_eq l r (cl, cr) =
+    let
+      fun explode' is_conj = explode is_conj true (l <> r) []
+      fun make_tab is_conj thm = make_littab (true_thm :: explode' is_conj thm)
+      fun prove is_conj ct tab = join is_conj tab (Thm.term_of ct)
+
+      val thm1 = Z3_New_Replay_Util.under_assumption (prove r cr o make_tab l) cl
+      val thm2 = Z3_New_Replay_Util.under_assumption (prove l cl o make_tab r) cr
+    in iff_intro thm1 thm2 end
+
+  datatype conj_disj = CONJ | DISJ | NCON | NDIS
+  fun kind_of t =
+    if is_conj t then SOME CONJ
+    else if is_disj t then SOME DISJ
+    else if is_neg' is_conj t then SOME NCON
+    else if is_neg' is_disj t then SOME NDIS
+    else NONE
+in
+
+fun prove_conj_disj_eq ct =
+  let val cp as (cl, cr) = Thm.dest_binop (Thm.dest_arg ct)
+  in
+    (case (kind_of (Thm.term_of cl), Thm.term_of cr) of
+      (SOME CONJ, @{const False}) => contradict true cl
+    | (SOME DISJ, @{const Not} $ @{const False}) =>
+        contrapos2 (contradict false o fst) cp
+    | (kl, _) =>
+        (case (kl, kind_of (Thm.term_of cr)) of
+          (SOME CONJ, SOME CONJ) => prove_eq true true cp
+        | (SOME CONJ, SOME NDIS) => prove_eq true false cp
+        | (SOME CONJ, _) => prove_eq true true cp
+        | (SOME DISJ, SOME DISJ) => contrapos1 (prove_eq false false) cp
+        | (SOME DISJ, SOME NCON) => contrapos2 (prove_eq false true) cp
+        | (SOME DISJ, _) => contrapos1 (prove_eq false false) cp
+        | (SOME NCON, SOME NCON) => contrapos1 (prove_eq true true) cp
+        | (SOME NCON, SOME DISJ) => contrapos3 (prove_eq true false) cp
+        | (SOME NCON, NONE) => contrapos3 (prove_eq true false) cp
+        | (SOME NDIS, SOME NDIS) => prove_eq false false cp
+        | (SOME NDIS, SOME CONJ) => prove_eq false true cp
+        | (SOME NDIS, NONE) => prove_eq false true cp
+        | _ => raise CTERM ("prove_conj_disj_eq", [ct])))
+  end
+
+end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/SMT2/z3_new_replay_methods.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -0,0 +1,667 @@
+(*  Title:      HOL/Tools/SMT2/z3_new_replay_methods.ML
+    Author:     Sascha Boehme, TU Muenchen
+    Author:     Jasmin Blanchette, TU Muenchen
+
+Proof methods for replaying Z3 proofs.
+*)
+
+signature Z3_NEW_REPLAY_METHODS =
+sig
+  (*abstraction*)
+  type abs_context = int * term Termtab.table
+  type 'a abstracter = term -> abs_context -> 'a * abs_context
+  val add_arith_abstracter: (term abstracter -> term option abstracter) ->
+    Context.generic -> Context.generic
+
+  (*theory lemma methods*)
+  type th_lemma_method = Proof.context -> thm list -> term -> thm
+  val add_th_lemma_method: string * th_lemma_method -> Context.generic ->
+    Context.generic
+
+  (*methods for Z3 proof rules*)
+  type z3_method = Proof.context -> thm list -> term -> thm
+  val true_axiom: z3_method
+  val mp: z3_method
+  val refl: z3_method
+  val symm: z3_method
+  val trans: z3_method
+  val cong: z3_method
+  val quant_intro: z3_method
+  val distrib: z3_method
+  val and_elim: z3_method
+  val not_or_elim: z3_method
+  val rewrite: z3_method
+  val rewrite_star: z3_method
+  val pull_quant: z3_method
+  val push_quant: z3_method
+  val elim_unused: z3_method
+  val dest_eq_res: z3_method
+  val quant_inst: z3_method
+  val lemma: z3_method
+  val unit_res: z3_method
+  val iff_true: z3_method
+  val iff_false: z3_method
+  val comm: z3_method
+  val def_axiom: z3_method
+  val apply_def: z3_method
+  val iff_oeq: z3_method
+  val nnf_pos: z3_method
+  val nnf_neg: z3_method
+  val mp_oeq: z3_method
+  val th_lemma: string -> z3_method
+  val is_assumption: Z3_New_Proof.z3_rule -> bool
+  val method_for: Z3_New_Proof.z3_rule -> z3_method
+end
+
+structure Z3_New_Replay_Methods: Z3_NEW_REPLAY_METHODS =
+struct
+
+type z3_method = Proof.context -> thm list -> term -> thm
+
+
+
+(* utility functions *)
+
+val trace = SMT2_Config.trace_msg
+
+fun pretty_thm ctxt thm = Syntax.pretty_term ctxt (Thm.concl_of thm)
+
+fun pretty_goal ctxt msg rule thms t =
+  let
+    val full_msg = msg ^ ": " ^ quote (Z3_New_Proof.string_of_rule rule)
+    val assms =
+      if null thms then []
+      else [Pretty.big_list "assumptions:" (map (pretty_thm ctxt) thms)]
+    val concl = Pretty.big_list "proposition:" [Syntax.pretty_term ctxt t]
+  in Pretty.big_list full_msg (assms @ [concl]) end
+
+fun replay_error ctxt msg rule thms t = error (Pretty.string_of (pretty_goal ctxt msg rule thms t))
+
+fun replay_rule_error ctxt = replay_error ctxt "Failed to replay Z3 proof step"
+
+fun trace_goal ctxt rule thms t =
+  trace ctxt (fn () => Pretty.string_of (pretty_goal ctxt "Goal" rule thms t))
+
+fun as_prop (t as Const (@{const_name Trueprop}, _) $ _) = t
+  | as_prop t = HOLogic.mk_Trueprop t
+
+fun dest_prop (Const (@{const_name Trueprop}, _) $ t) = t
+  | dest_prop t = t
+
+fun dest_thm thm = dest_prop (Thm.concl_of thm)
+
+fun certify_prop ctxt t = SMT2_Util.certify ctxt (as_prop t)
+
+fun try_provers ctxt rule [] thms t = replay_rule_error ctxt rule thms t
+  | try_provers ctxt rule ((name, prover) :: named_provers) thms t =
+      (case (trace ctxt (K ("Trying prover " ^ quote name)); try prover t) of
+        SOME thm => thm
+      | NONE => try_provers ctxt rule named_provers thms t)
+
+fun match ctxt pat t =
+  (Vartab.empty, Vartab.empty)
+  |> Pattern.first_order_match (Proof_Context.theory_of ctxt) (pat, t)
+
+fun gen_certify_inst sel mk cert ctxt thm t =
+  let
+    val inst = match ctxt (dest_thm thm) (dest_prop t)
+    fun cert_inst (ix, (a, b)) = (cert (mk (ix, a)), cert b)
+  in Vartab.fold (cons o cert_inst) (sel inst) [] end
+
+fun match_instantiateT ctxt t thm =
+  if Term.exists_type (Term.exists_subtype Term.is_TVar) (dest_thm thm) then
+    let val certT = Thm.ctyp_of (Proof_Context.theory_of ctxt)
+    in Thm.instantiate (gen_certify_inst fst TVar certT ctxt thm t, []) thm end
+  else thm
+
+fun match_instantiate ctxt t thm =
+  let
+    val cert = SMT2_Util.certify ctxt
+    val thm' = match_instantiateT ctxt t thm
+  in Thm.instantiate ([], gen_certify_inst snd Var cert ctxt thm' t) thm' end
+
+fun apply_rule ctxt t =
+  (case Z3_New_Replay_Rules.apply ctxt (certify_prop ctxt t) of
+    SOME thm => thm
+  | NONE => raise Fail "apply_rule")
+
+fun discharge _ [] thm = thm
+  | discharge i (rule :: rules) thm = discharge (i + Thm.nprems_of rule) rules (rule RSN (i, thm))
+
+fun by_tac ctxt thms ns ts t tac =
+  Goal.prove ctxt [] (map as_prop ts) (as_prop t)
+    (fn {context, prems} => HEADGOAL (tac context prems))
+  |> Drule.generalize ([], ns)
+  |> discharge 1 thms
+
+fun prove ctxt t tac = by_tac ctxt [] [] [] t (K o tac)
+
+fun prop_tac ctxt prems =
+  Method.insert_tac prems THEN' (Classical.fast_tac ctxt ORELSE' Clasimp.force_tac ctxt)
+
+fun quant_tac ctxt = Blast.blast_tac ctxt
+
+
+
+(* plug-ins *)
+
+type abs_context = int * term Termtab.table
+
+type 'a abstracter = term -> abs_context -> 'a * abs_context
+
+type th_lemma_method = Proof.context -> thm list -> term -> thm
+
+fun id_ord ((id1, _), (id2, _)) = int_ord (id1, id2)
+
+structure Plugins = Generic_Data
+(
+  type T =
+    (int * (term abstracter -> term option abstracter)) list *
+    th_lemma_method Symtab.table
+  val empty = ([], Symtab.empty)
+  val extend = I
+  fun merge ((abss1, ths1), (abss2, ths2)) = (
+    Ord_List.merge id_ord (abss1, abss2),
+    Symtab.merge (K true) (ths1, ths2))
+)
+
+fun add_arith_abstracter abs = Plugins.map (apfst (Ord_List.insert id_ord (serial (), abs)))
+fun get_arith_abstracters ctxt = map snd (fst (Plugins.get (Context.Proof ctxt)))
+
+fun add_th_lemma_method method = Plugins.map (apsnd (Symtab.update_new method))
+fun get_th_lemma_method ctxt = snd (Plugins.get (Context.Proof ctxt))
+
+
+
+(* abstraction *)
+
+fun prove_abstract ctxt thms t tac f =
+  let
+    val ((prems, concl), (_, ts)) = f (1, Termtab.empty)
+    val ns = Termtab.fold (fn (_, v) => cons (fst (Term.dest_Free v))) ts []
+  in
+    by_tac ctxt [] ns prems concl tac
+    |> match_instantiate ctxt t
+    |> discharge 1 thms
+  end
+
+fun prove_abstract' ctxt t tac f =
+  prove_abstract ctxt [] t tac (f #>> pair [])
+
+fun lookup_term (_, terms) t = Termtab.lookup terms t
+
+fun abstract_sub t f cx =
+  (case lookup_term cx t of
+    SOME v => (v, cx)
+  | NONE => f cx)
+
+fun mk_fresh_free t (i, terms) =
+  let val v = Free ("t" ^ string_of_int i, fastype_of t)
+  in (v, (i + 1, Termtab.update (t, v) terms)) end
+
+fun apply_abstracters _ [] _ cx = (NONE, cx)
+  | apply_abstracters abs (abstracter :: abstracters) t cx =
+      (case abstracter abs t cx of
+        (NONE, _) => apply_abstracters abs abstracters t cx
+      | x as (SOME _, _) => x)
+
+fun abstract_term (t as _ $ _) = abstract_sub t (mk_fresh_free t)
+  | abstract_term (t as Abs _) = abstract_sub t (mk_fresh_free t)
+  | abstract_term t = pair t
+
+fun abstract_bin abs f t t1 t2 = abstract_sub t (abs t1 ##>> abs t2 #>> f)
+
+fun abstract_ter abs f t t1 t2 t3 =
+  abstract_sub t (abs t1 ##>> abs t2 ##>> abs t3 #>> (Parse.triple1 #> f))
+
+fun abstract_lit (@{const HOL.Not} $ t) = abstract_term t #>> HOLogic.mk_not
+  | abstract_lit t = abstract_term t
+
+fun abstract_not abs (t as @{const HOL.Not} $ t1) =
+      abstract_sub t (abs t1 #>> HOLogic.mk_not)
+  | abstract_not _ t = abstract_lit t
+
+fun abstract_conj (t as @{const HOL.conj} $ t1 $ t2) =
+      abstract_bin abstract_conj HOLogic.mk_conj t t1 t2
+  | abstract_conj t = abstract_lit t
+
+fun abstract_disj (t as @{const HOL.disj} $ t1 $ t2) =
+      abstract_bin abstract_disj HOLogic.mk_disj t t1 t2
+  | abstract_disj t = abstract_lit t
+
+fun abstract_prop (t as (c as @{const If (bool)}) $ t1 $ t2 $ t3) =
+      abstract_ter abstract_prop (fn (t1, t2, t3) => c $ t1 $ t2 $ t3) t t1 t2 t3
+  | abstract_prop (t as @{const HOL.disj} $ t1 $ t2) =
+      abstract_bin abstract_prop HOLogic.mk_disj t t1 t2
+  | abstract_prop (t as @{const HOL.conj} $ t1 $ t2) =
+      abstract_bin abstract_prop HOLogic.mk_conj t t1 t2
+  | abstract_prop (t as @{const HOL.implies} $ t1 $ t2) =
+      abstract_bin abstract_prop HOLogic.mk_imp t t1 t2
+  | abstract_prop (t as @{term "HOL.eq :: bool => _"} $ t1 $ t2) =
+      abstract_bin abstract_prop HOLogic.mk_eq t t1 t2
+  | abstract_prop t = abstract_not abstract_prop t
+
+fun abstract_arith ctxt u =
+  let
+    fun abs (t as (c as Const _) $ Abs (s, T, t')) =
+          abstract_sub t (abs t' #>> (fn u' => c $ Abs (s, T, u')))
+      | abs (t as (c as Const (@{const_name If}, _)) $ t1 $ t2 $ t3) =
+          abstract_ter abs (fn (t1, t2, t3) => c $ t1 $ t2 $ t3) t t1 t2 t3
+      | abs (t as @{const HOL.Not} $ t1) = abstract_sub t (abs t1 #>> HOLogic.mk_not)
+      | abs (t as @{const HOL.disj} $ t1 $ t2) =
+          abstract_sub t (abs t1 ##>> abs t2 #>> HOLogic.mk_disj)
+      | abs (t as (c as Const (@{const_name uminus_class.uminus}, _)) $ t1) =
+          abstract_sub t (abs t1 #>> (fn u => c $ u))
+      | abs (t as (c as Const (@{const_name plus_class.plus}, _)) $ t1 $ t2) =
+          abstract_sub t (abs t1 ##>> abs t2 #>> (fn (u1, u2) => c $ u1 $ u2))
+      | abs (t as (c as Const (@{const_name minus_class.minus}, _)) $ t1 $ t2) =
+          abstract_sub t (abs t1 ##>> abs t2 #>> (fn (u1, u2) => c $ u1 $ u2))
+      | abs (t as (c as Const (@{const_name times_class.times}, _)) $ t1 $ t2) =
+          abstract_sub t (abs t1 ##>> abs t2 #>> (fn (u1, u2) => c $ u1 $ u2))
+      | abs (t as (c as Const (@{const_name z3div}, _)) $ t1 $ t2) =
+          abstract_sub t (abs t1 ##>> abs t2 #>> (fn (u1, u2) => c $ u1 $ u2))
+      | abs (t as (c as Const (@{const_name z3mod}, _)) $ t1 $ t2) =
+          abstract_sub t (abs t1 ##>> abs t2 #>> (fn (u1, u2) => c $ u1 $ u2))
+      | abs (t as (c as Const (@{const_name HOL.eq}, _)) $ t1 $ t2) =
+          abstract_sub t (abs t1 ##>> abs t2 #>> (fn (u1, u2) => c $ u1 $ u2))
+      | abs (t as (c as Const (@{const_name ord_class.less}, _)) $ t1 $ t2) =
+          abstract_sub t (abs t1 ##>> abs t2 #>> (fn (u1, u2) => c $ u1 $ u2))
+      | abs (t as (c as Const (@{const_name ord_class.less_eq}, _)) $ t1 $ t2) =
+          abstract_sub t (abs t1 ##>> abs t2 #>> (fn (u1, u2) => c $ u1 $ u2))
+      | abs t = abstract_sub t (fn cx =>
+          if can HOLogic.dest_number t then (t, cx)
+          else
+            (case apply_abstracters abs (get_arith_abstracters ctxt) t cx of
+              (SOME u, cx') => (u, cx')
+            | (NONE, _) => abstract_term t cx))
+  in abs u end
+
+
+
+(* truth axiom *)
+
+fun true_axiom _ _ _ = @{thm TrueI}
+
+
+
+(* modus ponens *)
+
+fun mp _ [p, p_eq_q] _ = discharge 1 [p_eq_q, p] iffD1
+  | mp ctxt thms t = replay_rule_error ctxt Z3_New_Proof.Modus_Ponens thms t
+
+val mp_oeq = mp
+
+
+
+(* reflexivity *)
+
+fun refl ctxt _ t = match_instantiate ctxt t @{thm refl}
+
+
+
+(* symmetry *)
+
+fun symm _ [thm] _ = thm RS @{thm sym}
+  | symm ctxt thms t = replay_rule_error ctxt Z3_New_Proof.Reflexivity thms t
+
+
+
+(* transitivity *)
+
+fun trans _ [thm1, thm2] _ = thm1 RSN (1, thm2 RSN (2, @{thm trans}))
+  | trans ctxt thms t = replay_rule_error ctxt Z3_New_Proof.Transitivity thms t
+
+
+
+(* congruence *)
+
+fun ctac prems i st = st |> (
+  resolve_tac (@{thm refl} :: prems) i
+  ORELSE (cong_tac i THEN ctac prems (i + 1) THEN ctac prems i))
+
+fun cong_basic ctxt thms t =
+  let val st = Thm.trivial (certify_prop ctxt t)
+  in
+    (case Seq.pull (ctac thms 1 st) of
+      SOME (thm, _) => thm
+    | NONE => raise THM ("cong", 0, thms @ [st]))
+  end
+
+val cong_dest_rules = @{lemma
+  "(~ P | Q) & (P | ~ Q) ==> P = Q"
+  "(P | ~ Q) & (~ P | Q) ==> P = Q"
+  by fast+}
+
+fun cong_full ctxt thms t = prove ctxt t (fn ctxt' =>
+  Method.insert_tac thms
+  THEN' (Classical.fast_tac ctxt'
+    ORELSE' dresolve_tac cong_dest_rules
+    THEN' Classical.fast_tac ctxt'))
+
+fun cong ctxt thms = try_provers ctxt Z3_New_Proof.Monotonicity [
+  ("basic", cong_basic ctxt thms),
+  ("full", cong_full ctxt thms)] thms
+
+
+
+(* quantifier introduction *)
+
+val quant_intro_rules = @{lemma
+  "(!!x. P x = Q x) ==> (ALL x. P x) = (ALL x. Q x)"
+  "(!!x. P x = Q x) ==> (EX x. P x) = (EX x. Q x)"
+  "(!!x. (~ P x) = Q x) ==> (~ (EX x. P x)) = (ALL x. Q x)"
+  "(!!x. (~ P x) = Q x) ==> (~ (ALL x. P x)) = (EX x. Q x)"
+  by fast+}
+
+fun quant_intro ctxt [thm] t =
+    prove ctxt t (K (REPEAT_ALL_NEW (resolve_tac (thm :: quant_intro_rules))))
+  | quant_intro ctxt thms t = replay_rule_error ctxt Z3_New_Proof.Quant_Intro thms t
+
+
+
+(* distributivity of conjunctions and disjunctions *)
+
+(* TODO: there are no tests with this proof rule *)
+fun distrib ctxt _ t =
+  prove_abstract' ctxt t prop_tac (abstract_prop (dest_prop t))
+
+
+
+(* elimination of conjunctions *)
+
+fun and_elim ctxt [thm] t =
+      prove_abstract ctxt [thm] t prop_tac (
+        abstract_lit (dest_prop t) ##>>
+        abstract_conj (dest_thm thm) #>>
+        apfst single o swap)
+  | and_elim ctxt thms t = replay_rule_error ctxt Z3_New_Proof.And_Elim thms t
+
+
+
+(* elimination of negated disjunctions *)
+
+fun not_or_elim ctxt [thm] t =
+      prove_abstract ctxt [thm] t prop_tac (
+        abstract_lit (dest_prop t) ##>>
+        abstract_not abstract_disj (dest_thm thm) #>>
+        apfst single o swap)
+  | not_or_elim ctxt thms t =
+      replay_rule_error ctxt Z3_New_Proof.Not_Or_Elim thms t
+
+
+
+(* rewriting *)
+
+fun abstract_eq f1 f2 (Const (@{const_name HOL.eq}, _) $ t1 $ t2) =
+      f1 t1 ##>> f2 t2 #>> HOLogic.mk_eq
+  | abstract_eq _ _ t = abstract_term t
+
+fun prove_prop_rewrite ctxt t =
+  prove_abstract' ctxt t prop_tac (
+    abstract_eq abstract_prop abstract_prop (dest_prop t))
+
+fun arith_rewrite_tac ctxt _ =
+  TRY o Simplifier.simp_tac ctxt
+  THEN_ALL_NEW (Arith_Data.arith_tac ctxt ORELSE' Clasimp.force_tac ctxt)
+
+fun prove_arith_rewrite ctxt t =
+  prove_abstract' ctxt t arith_rewrite_tac (
+    abstract_eq (abstract_arith ctxt) (abstract_arith ctxt) (dest_prop t))
+
+fun rewrite ctxt _ = try_provers ctxt Z3_New_Proof.Rewrite [
+  ("rules", apply_rule ctxt),
+  ("prop_rewrite", prove_prop_rewrite ctxt),
+  ("arith_rewrite", prove_arith_rewrite ctxt)] []
+
+fun rewrite_star ctxt = rewrite ctxt
+
+
+
+(* pulling quantifiers *)
+
+fun pull_quant ctxt _ t = prove ctxt t quant_tac
+
+
+
+(* pushing quantifiers *)
+
+fun push_quant _ _ _ = raise Fail "unsupported" (* FIXME *)
+
+
+
+(* elimination of unused bound variables *)
+
+val elim_all = @{lemma "P = Q ==> (ALL x. P) = Q" by fast}
+val elim_ex = @{lemma "P = Q ==> (EX x. P) = Q" by fast}
+
+fun elim_unused_tac i st = (
+  match_tac [@{thm refl}]
+  ORELSE' (match_tac [elim_all, elim_ex] THEN' elim_unused_tac)
+  ORELSE' (
+    match_tac [@{thm iff_allI}, @{thm iff_exI}]
+    THEN' elim_unused_tac)) i st
+
+fun elim_unused ctxt _ t = prove ctxt t (fn _ => elim_unused_tac)
+
+
+
+(* destructive equality resolution *)
+
+fun dest_eq_res _ _ _ = raise Fail "dest_eq_res" (* FIXME *)
+
+
+
+(* quantifier instantiation *)
+
+val quant_inst_rule = @{lemma "~P x | Q ==> ~(ALL x. P x) | Q" by fast}
+
+fun quant_inst ctxt _ t = prove ctxt t (fn _ => 
+  REPEAT_ALL_NEW (rtac quant_inst_rule)
+  THEN' rtac @{thm excluded_middle})
+
+
+
+(* propositional lemma *)
+
+exception LEMMA of unit
+
+val intro_hyp_rule1 = @{lemma "(~P ==> Q) ==> P | Q" by fast}
+val intro_hyp_rule2 = @{lemma "(P ==> Q) ==> ~P | Q" by fast}
+
+fun norm_lemma thm =
+  (thm COMP_INCR intro_hyp_rule1)
+  handle THM _ => thm COMP_INCR intro_hyp_rule2
+
+fun negated_prop (@{const HOL.Not} $ t) = HOLogic.mk_Trueprop t
+  | negated_prop t = HOLogic.mk_Trueprop (HOLogic.mk_not t)
+
+fun intro_hyps tab (t as @{const HOL.disj} $ t1 $ t2) cx =
+      lookup_intro_hyps tab t (fold (intro_hyps tab) [t1, t2]) cx
+  | intro_hyps tab t cx =
+      lookup_intro_hyps tab t (fn _ => raise LEMMA ()) cx
+
+and lookup_intro_hyps tab t f (cx as (thm, terms)) =
+  (case Termtab.lookup tab (negated_prop t) of
+    NONE => f cx
+  | SOME hyp => (norm_lemma (Thm.implies_intr hyp thm), t :: terms))
+
+fun lemma ctxt (thms as [thm]) t =
+    (let
+       val tab = Termtab.make (map (`Thm.term_of) (#hyps (Thm.crep_thm thm)))
+       val (thm', terms) = intro_hyps tab (dest_prop t) (thm, [])
+     in
+       prove_abstract ctxt [thm'] t prop_tac (
+         fold (snd oo abstract_lit) terms #>
+         abstract_disj (dest_thm thm') #>> single ##>>
+         abstract_disj (dest_prop t))
+     end
+     handle LEMMA () => replay_error ctxt "Bad proof state" Z3_New_Proof.Lemma thms t)
+  | lemma ctxt thms t = replay_rule_error ctxt Z3_New_Proof.Lemma thms t
+
+
+
+(* unit resolution *)
+
+fun abstract_unit (t as (@{const HOL.Not} $ (@{const HOL.disj} $ t1 $ t2))) =
+      abstract_sub t (abstract_unit t1 ##>> abstract_unit t2 #>>
+        HOLogic.mk_not o HOLogic.mk_disj)
+  | abstract_unit (t as (@{const HOL.disj} $ t1 $ t2)) =
+      abstract_sub t (abstract_unit t1 ##>> abstract_unit t2 #>>
+        HOLogic.mk_disj)
+  | abstract_unit t = abstract_lit t
+
+fun unit_res ctxt thms t =
+  prove_abstract ctxt thms t prop_tac (
+    fold_map (abstract_unit o dest_thm) thms ##>>
+    abstract_unit (dest_prop t) #>>
+    (fn (prems, concl) => (prems, concl)))
+
+
+
+(* iff-true *)
+
+val iff_true_rule = @{lemma "P ==> P = True" by fast}
+
+fun iff_true _ [thm] _ = thm RS iff_true_rule
+  | iff_true ctxt thms t = replay_rule_error ctxt Z3_New_Proof.Iff_True thms t
+
+
+
+(* iff-false *)
+
+val iff_false_rule = @{lemma "~P ==> P = False" by fast}
+
+fun iff_false _ [thm] _ = thm RS iff_false_rule
+  | iff_false ctxt thms t = replay_rule_error ctxt Z3_New_Proof.Iff_False thms t
+
+
+
+(* commutativity *)
+
+fun comm ctxt _ t = match_instantiate ctxt t @{thm eq_commute}
+
+
+
+(* definitional axioms *)
+
+fun def_axiom_disj ctxt t =
+  (case dest_prop t of
+    @{const HOL.disj} $ u1 $ u2 =>
+      prove_abstract' ctxt t prop_tac (
+        abstract_prop u2 ##>> abstract_prop u1 #>> HOLogic.mk_disj o swap)
+  | u => prove_abstract' ctxt t prop_tac (abstract_prop u))
+
+fun def_axiom ctxt _ = try_provers ctxt Z3_New_Proof.Def_Axiom [
+  ("rules", apply_rule ctxt),
+  ("disj", def_axiom_disj ctxt)] []
+
+
+
+(* application of definitions *)
+
+fun apply_def _ [thm] _ = thm (* TODO: cover also the missing cases *)
+  | apply_def ctxt thms t = replay_rule_error ctxt Z3_New_Proof.Apply_Def thms t
+
+
+
+(* iff-oeq *)
+
+fun iff_oeq _ _ _ = raise Fail "iff_oeq" (* FIXME *)
+
+
+
+(* negation normal form *)
+
+fun nnf_prop ctxt thms t =
+  prove_abstract ctxt thms t prop_tac (
+    fold_map (abstract_prop o dest_thm) thms ##>>
+    abstract_prop (dest_prop t))
+
+fun nnf ctxt rule thms = try_provers ctxt rule [
+  ("prop", nnf_prop ctxt thms),
+  ("quant", quant_intro ctxt [hd thms])] thms
+
+fun nnf_pos ctxt = nnf ctxt Z3_New_Proof.Nnf_Pos
+fun nnf_neg ctxt = nnf ctxt Z3_New_Proof.Nnf_Neg
+
+
+
+(* theory lemmas *)
+
+fun arith_th_lemma_tac ctxt prems =
+  Method.insert_tac prems
+  THEN' SELECT_GOAL (Local_Defs.unfold_tac ctxt @{thms z3div_def z3mod_def})
+  THEN' Arith_Data.arith_tac ctxt
+
+fun arith_th_lemma ctxt thms t =
+  prove_abstract ctxt thms t arith_th_lemma_tac (
+    fold_map (abstract_arith ctxt o dest_thm) thms ##>>
+    abstract_arith ctxt (dest_prop t))
+
+val _ = Theory.setup (Context.theory_map (add_th_lemma_method ("arith", arith_th_lemma)))
+
+fun th_lemma name ctxt thms =
+  (case Symtab.lookup (get_th_lemma_method ctxt) name of
+    SOME method => method ctxt thms
+  | NONE => replay_error ctxt "Bad theory" (Z3_New_Proof.Th_Lemma name) thms)
+
+
+
+(* mapping of rules to methods *)
+
+fun is_assumption Z3_New_Proof.Asserted = true
+  | is_assumption Z3_New_Proof.Goal = true
+  | is_assumption Z3_New_Proof.Hypothesis = true
+  | is_assumption Z3_New_Proof.Intro_Def = true
+  | is_assumption Z3_New_Proof.Skolemize = true
+  | is_assumption _ = false
+
+fun unsupported rule ctxt = replay_error ctxt "Unsupported" rule
+fun assumed rule ctxt = replay_error ctxt "Assumed" rule
+
+fun choose Z3_New_Proof.True_Axiom = true_axiom
+  | choose (r as Z3_New_Proof.Asserted) = assumed r
+  | choose (r as Z3_New_Proof.Goal) = assumed r
+  | choose Z3_New_Proof.Modus_Ponens = mp
+  | choose Z3_New_Proof.Reflexivity = refl
+  | choose Z3_New_Proof.Symmetry = symm
+  | choose Z3_New_Proof.Transitivity = trans
+  | choose (r as Z3_New_Proof.Transitivity_Star) = unsupported r
+  | choose Z3_New_Proof.Monotonicity = cong
+  | choose Z3_New_Proof.Quant_Intro = quant_intro
+  | choose Z3_New_Proof.Distributivity = distrib
+  | choose Z3_New_Proof.And_Elim = and_elim
+  | choose Z3_New_Proof.Not_Or_Elim = not_or_elim
+  | choose Z3_New_Proof.Rewrite = rewrite
+  | choose Z3_New_Proof.Rewrite_Star = rewrite_star
+  | choose Z3_New_Proof.Pull_Quant = pull_quant
+  | choose (r as Z3_New_Proof.Pull_Quant_Star) = unsupported r
+  | choose Z3_New_Proof.Push_Quant = push_quant
+  | choose Z3_New_Proof.Elim_Unused_Vars = elim_unused
+  | choose Z3_New_Proof.Dest_Eq_Res = dest_eq_res
+  | choose Z3_New_Proof.Quant_Inst = quant_inst
+  | choose (r as Z3_New_Proof.Hypothesis) = assumed r
+  | choose Z3_New_Proof.Lemma = lemma
+  | choose Z3_New_Proof.Unit_Resolution = unit_res
+  | choose Z3_New_Proof.Iff_True = iff_true
+  | choose Z3_New_Proof.Iff_False = iff_false
+  | choose Z3_New_Proof.Commutativity = comm
+  | choose Z3_New_Proof.Def_Axiom = def_axiom
+  | choose (r as Z3_New_Proof.Intro_Def) = assumed r
+  | choose Z3_New_Proof.Apply_Def = apply_def
+  | choose Z3_New_Proof.Iff_Oeq = iff_oeq
+  | choose Z3_New_Proof.Nnf_Pos = nnf_pos
+  | choose Z3_New_Proof.Nnf_Neg = nnf_neg
+  | choose (r as Z3_New_Proof.Nnf_Star) = unsupported r
+  | choose (r as Z3_New_Proof.Cnf_Star) = unsupported r
+  | choose (r as Z3_New_Proof.Skolemize) = assumed r
+  | choose Z3_New_Proof.Modus_Ponens_Oeq = mp_oeq
+  | choose (Z3_New_Proof.Th_Lemma name) = th_lemma name
+
+fun with_tracing rule method ctxt thms t =
+  let val _ = trace_goal ctxt rule thms t
+  in method ctxt thms t end
+
+fun method_for rule = with_tracing rule (choose rule)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/SMT2/z3_new_replay_rules.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -0,0 +1,54 @@
+(*  Title:      HOL/Tools/SMT2/z3_new_replay_rules.ML
+    Author:     Sascha Boehme, TU Muenchen
+
+Custom rules for Z3 proof replay.
+*)
+
+signature Z3_NEW_REPLAY_RULES =
+sig
+  val apply: Proof.context -> cterm -> thm option
+end
+
+structure Z3_New_Replay_Rules: Z3_NEW_REPLAY_RULES =
+struct
+
+structure Data = Generic_Data
+(
+  type T = thm Net.net
+  val empty = Net.empty
+  val extend = I
+  val merge = Net.merge Thm.eq_thm
+)
+
+fun maybe_instantiate ct thm =
+  try Thm.first_order_match (Thm.cprop_of thm, ct)
+  |> Option.map (fn inst => Thm.instantiate inst thm)
+
+fun apply ctxt ct =
+  let
+    val net = Data.get (Context.Proof ctxt)
+    val xthms = Net.match_term net (Thm.term_of ct)
+
+    fun select ct = map_filter (maybe_instantiate ct) xthms 
+    fun select' ct =
+      let val thm = Thm.trivial ct
+      in map_filter (try (fn rule => rule COMP thm)) xthms end
+
+  in try hd (case select ct of [] => select' ct | xthms' => xthms') end
+
+val prep = `Thm.prop_of
+
+fun ins thm net = Net.insert_term Thm.eq_thm (prep thm) net handle Net.INSERT => net
+fun del thm net = Net.delete_term Thm.eq_thm (prep thm) net handle Net.DELETE => net
+
+val add = Thm.declaration_attribute (Data.map o ins)
+val del = Thm.declaration_attribute (Data.map o del)
+
+val name = Binding.name "z3_new_rule"
+
+val description = "declaration of Z3 proof rules"
+
+val _ = Theory.setup (Attrib.setup name (Attrib.add_del add del) description #>
+  Global_Theory.add_thms_dynamic (name, Net.content o Data.get))
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/SMT2/z3_new_replay_util.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -0,0 +1,155 @@
+(*  Title:      HOL/Tools/SMT2/z3_new_replay_util.ML
+    Author:     Sascha Boehme, TU Muenchen
+
+Helper functions required for Z3 proof replay.
+*)
+
+signature Z3_NEW_REPLAY_UTIL =
+sig
+  (*theorem nets*)
+  val thm_net_of: ('a -> thm) -> 'a list -> 'a Net.net
+  val net_instances: (int * thm) Net.net -> cterm -> (int * thm) list
+
+  (*proof combinators*)
+  val under_assumption: (thm -> thm) -> cterm -> thm
+  val discharge: thm -> thm -> thm
+
+  (*a faster COMP*)
+  type compose_data
+  val precompose: (cterm -> cterm list) -> thm -> compose_data
+  val precompose2: (cterm -> cterm * cterm) -> thm -> compose_data
+  val compose: compose_data -> thm -> thm
+
+  (*simpset*)
+  val add_simproc: Simplifier.simproc -> Context.generic -> Context.generic
+  val make_simpset: Proof.context -> thm list -> simpset
+end
+
+structure Z3_New_Replay_Util: Z3_NEW_REPLAY_UTIL =
+struct
+
+
+
+(* theorem nets *)
+
+fun thm_net_of f xthms =
+  let fun insert xthm = Net.insert_term (K false) (Thm.prop_of (f xthm), xthm)
+  in fold insert xthms Net.empty end
+
+fun maybe_instantiate ct thm =
+  try Thm.first_order_match (Thm.cprop_of thm, ct)
+  |> Option.map (fn inst => Thm.instantiate inst thm)
+
+local
+  fun instances_from_net match f net ct =
+    let
+      val lookup = if match then Net.match_term else Net.unify_term
+      val xthms = lookup net (Thm.term_of ct)
+      fun select ct = map_filter (f (maybe_instantiate ct)) xthms 
+      fun select' ct =
+        let val thm = Thm.trivial ct
+        in map_filter (f (try (fn rule => rule COMP thm))) xthms end
+    in (case select ct of [] => select' ct | xthms' => xthms') end
+in
+
+fun net_instances net =
+  instances_from_net false (fn f => fn (i, thm) => Option.map (pair i) (f thm))
+    net
+
+end
+
+
+
+(* proof combinators *)
+
+fun under_assumption f ct =
+  let val ct' = SMT2_Util.mk_cprop ct in Thm.implies_intr ct' (f (Thm.assume ct')) end
+
+fun discharge p pq = Thm.implies_elim pq p
+
+
+
+(* a faster COMP *)
+
+type compose_data = cterm list * (cterm -> cterm list) * thm
+
+fun list2 (x, y) = [x, y]
+
+fun precompose f rule = (f (Thm.cprem_of rule 1), f, rule)
+fun precompose2 f rule = precompose (list2 o f) rule
+
+fun compose (cvs, f, rule) thm =
+  discharge thm (Thm.instantiate ([], cvs ~~ f (Thm.cprop_of thm)) rule)
+
+
+
+(* 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 ctxt t =
+    let
+      val (le, r, s) = dest_binop t
+      val less = Const (@{const_name less}, Term.fastype_of le)
+      val prems = Simplifier.prems_of ctxt
+    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 ctxt 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 = Simplifier.prems_of ctxt
+    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
+
+  val basic_simpset =
+    simpset_of (put_simpset HOL_ss @{context}
+      addsimps @{thms field_simps times_divide_eq_right times_divide_eq_left arith_special
+        arith_simps rel_simps array_rules z3div_def z3mod_def}
+      addsimprocs [@{simproc binary_int_div}, @{simproc binary_int_mod},
+        Simplifier.simproc_global @{theory} "fast_int_arith" [
+          "(m::int) < n", "(m::int) <= n", "(m::int) = n"] Lin_Arith.simproc,
+        Simplifier.simproc_global @{theory} "antisym_le" ["(x::'a::order) <= y"] prove_antisym_le,
+        Simplifier.simproc_global @{theory} "antisym_less" ["~ (x::'a::linorder) < y"]
+          prove_antisym_less])
+
+  structure Simpset = Generic_Data
+  (
+    type T = simpset
+    val empty = basic_simpset
+    val extend = I
+    val merge = Simplifier.merge_ss
+  )
+in
+
+fun add_simproc simproc context =
+  Simpset.map (simpset_map (Context.proof_of context)
+    (fn ctxt => ctxt addsimprocs [simproc])) context
+
+fun make_simpset ctxt rules =
+  simpset_of (put_simpset (Simpset.get (Context.Proof ctxt)) ctxt addsimps rules)
+
+end
+
+end
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_commands.ML	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_commands.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -29,7 +29,7 @@
 
 (* val cvc3N = "cvc3" *)
 val yicesN = "yices"
-val z3N = "z3"
+val z3_newN = "z3_new"
 
 val runN = "run"
 val minN = "min"
@@ -195,14 +195,6 @@
   fun merge data : T = AList.merge (op =) (K true) data
 )
 
-fun is_prover_supported ctxt =
-  let val thy = Proof_Context.theory_of ctxt in
-    is_proof_method orf is_atp thy orf is_smt_prover ctxt
-  end
-
-fun is_prover_installed ctxt =
-  is_proof_method orf is_smt_prover ctxt orf is_atp_installed (Proof_Context.theory_of ctxt)
-
 fun remotify_prover_if_supported_and_not_already_remote ctxt name =
   if String.isPrefix remote_prefix name then
     SOME name
@@ -220,7 +212,7 @@
 (* The first ATP of the list is used by Auto Sledgehammer. Because of the low
    timeout, it makes sense to put E first. *)
 fun default_provers_param_value mode ctxt =
-  [eN, spassN, z3N, e_sineN, vampireN, yicesN]
+  [eN, spassN, z3_newN, e_sineN, vampireN, yicesN]
   |> map_filter (remotify_prover_if_not_installed ctxt)
   (* In "try" mode, leave at least one thread to another slow tool (e.g. Nitpick) *)
   |> take (Multithreading.max_threads_value () - (if mode = Try then 1 else 0))
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -68,32 +68,31 @@
 fun add_fact_of_dependencies [(_, ss as _ :: _)] = apsnd (union (op =) ss)
   | add_fact_of_dependencies names = apfst (insert (op =) (label_of_clause names))
 
-(* No "real" literals means only type information (tfree_tcs, clsrel, or clsarity). *)
-fun is_only_type_information t = t aconv @{prop True}
+fun is_True_prop t = t aconv @{prop True}
 
-(* Discard facts; consolidate adjacent lines that prove the same formula, since they differ only in
-   type information. *)
 fun add_line_pass1 (line as (name, role, t, rule, [])) lines =
     (* No dependencies: lemma (for Z3), fact, conjecture, or (for Vampire) internal facts or
        definitions. *)
-    if role = Lemma orelse role = Conjecture orelse role = Negated_Conjecture orelse
-       role = Hypothesis orelse is_arith_rule rule then
+    if role = Conjecture orelse role = Negated_Conjecture then
+      line :: lines
+    else if is_True_prop t then
+      map (replace_dependencies_in_line (name, [])) lines
+    else if role = Lemma orelse role = Hypothesis orelse is_arith_rule rule then
       line :: lines
     else if role = Axiom then
-      (* Facts are not proof lines. *)
-      lines |> is_only_type_information t ? map (replace_dependencies_in_line (name, []))
+      lines (* axioms (facts) need no proof lines *)
     else
       map (replace_dependencies_in_line (name, [])) lines
   | add_line_pass1 line lines = line :: lines
 
-fun add_lines_pass2 res [] = rev res
-  | add_lines_pass2 res ((name, role, t, rule, deps) :: lines) =
+fun add_lines_pass2 res _ [] = rev res
+  | add_lines_pass2 res prev_t ((line as (name, role, t, rule, deps)) :: lines) =
     let
       val is_last_line = null lines
 
       fun looks_interesting () =
-        not (is_only_type_information t) andalso null (Term.add_tvars t [])
-        andalso length deps >= 2 andalso not (can the_single lines)
+        not (is_True_prop t) andalso not (t aconv prev_t) andalso null (Term.add_tvars t []) andalso
+        length deps >= 2 andalso not (can the_single lines)
 
       fun is_skolemizing_line (_, _, _, rule', deps') =
         is_skolemize_rule rule' andalso member (op =) deps' name
@@ -102,9 +101,9 @@
       if role <> Plain orelse is_skolemize_rule rule orelse is_arith_rule rule orelse
          is_datatype_rule rule orelse is_last_line orelse looks_interesting () orelse
          is_before_skolemize_rule () then
-        add_lines_pass2 ((name, role, t, rule, deps) :: res) lines
+        add_lines_pass2 (line :: res) t lines
       else
-        add_lines_pass2 res (map (replace_dependencies_in_line (name, deps)) lines)
+        add_lines_pass2 res t (map (replace_dependencies_in_line (name, deps)) lines)
     end
 
 type isar_params =
@@ -112,7 +111,7 @@
   * (term, string) atp_step list * thm
 
 val basic_systematic_methods = [Metis_Method (NONE, NONE), Meson_Method, Blast_Method]
-val simp_based_methods = [Simp_Method, Auto_Method, Fastforce_Method, Force_Method]
+val simp_based_methods = [Auto_Method, Simp_Method, Fastforce_Method, Force_Method]
 val basic_arith_methods = [Linarith_Method, Presburger_Method, Algebra_Method]
 
 val arith_methods = basic_arith_methods @ simp_based_methods @ basic_systematic_methods
@@ -125,6 +124,8 @@
 fun isar_proof_text ctxt debug isar_proofs smt_proofs isar_params
     (one_line_params as (_, _, _, _, subgoal, subgoal_count)) =
   let
+    val _ = if debug then Output.urgent_message "Constructing Isar proof..." else ()
+
     fun isar_proof_of () =
       let
         val SOME (verbose, alt_metis_args, preplay_timeout, compress_isar, try0_isar, minimize,
@@ -134,7 +135,7 @@
 
         fun massage_methods (meths as meth :: _) =
           if not try0_isar then [meth]
-          else if smt_proofs = SOME true then SMT_Method :: meths
+          else if smt_proofs = SOME true then SMT2_Method :: meths
           else meths
 
         val (params, _, concl_t) = strip_subgoal goal subgoal ctxt
@@ -153,7 +154,7 @@
         val atp_proof =
           atp_proof
           |> rpair [] |-> fold_rev add_line_pass1
-          |> add_lines_pass2 []
+          |> add_lines_pass2 [] Term.dummy
 
         val conjs =
           map_filter (fn (name, role, _, _, _) =>
@@ -272,8 +273,6 @@
         and isar_proof outer fix assms lems infs =
           Proof (fix, assms, lems @ isar_steps outer NONE [] infs)
 
-        val string_of_isar_proof = string_of_isar_proof ctxt subgoal subgoal_count
-
         val trace = Config.get ctxt trace
 
         val canonical_isar_proof =
@@ -303,7 +302,8 @@
         fun trace_isar_proof label proof =
           if trace then
             tracing (timestamp () ^ "\n" ^ label ^ ":\n\n" ^
-              string_of_isar_proof (comment_isar_proof comment_of proof) ^ "\n")
+              string_of_isar_proof ctxt subgoal subgoal_count
+                (comment_isar_proof comment_of proof) ^ "\n")
           else
             ()
 
@@ -335,7 +335,7 @@
                #> kill_useless_labels_in_isar_proof
                #> relabel_isar_proof_nicely)
       in
-        (case string_of_isar_proof isar_proof of
+        (case string_of_isar_proof ctxt subgoal subgoal_count isar_proof of
           "" =>
           if isar_proofs = SOME true then "\nNo structured proof available (proof too simple)."
           else ""
@@ -363,15 +363,14 @@
         (case try isar_proof_of () of
           SOME s => s
         | NONE =>
-          if isar_proofs = SOME true then "\nWarning: The Isar proof construction failed." else "")
+          if isar_proofs = SOME true then "\nWarning: Isar proof construction failed." else "")
   in one_line_proof ^ isar_proof end
 
 fun isar_proof_would_be_a_good_idea smt_proofs (meth, play) =
   (case play of
-    Played _ => meth = SMT_Method andalso smt_proofs <> SOME true
-  | Play_Timed_Out _ => true
-  | Play_Failed => true
-  | Not_Played => false)
+    Played _ => meth = SMT2_Method andalso smt_proofs <> SOME true
+  | Play_Timed_Out time => Time.> (time, Time.zeroTime)
+  | Play_Failed => true)
 
 fun proof_text ctxt debug isar_proofs smt_proofs isar_params num_chained
     (one_line_params as (preplay, _, _, _, _, _)) =
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_isar_preplay.ML	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_isar_preplay.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -178,7 +178,8 @@
     end
   | set_preplay_outcomes_of_isar_step _ _ _ _ _ _ = ()
 
-fun peek_at_outcome outcome = if Lazy.is_finished outcome then Lazy.force outcome else Not_Played
+fun peek_at_outcome outcome =
+  if Lazy.is_finished outcome then Lazy.force outcome else Play_Timed_Out Time.zeroTime
 
 fun get_best_method_outcome force =
   tap (List.app (K () o Lazy.future Future.default_params o snd)) (* optional parallelism *)
@@ -195,14 +196,17 @@
         outcome as Played _ => outcome
       | _ => snd (get_best_method_outcome Lazy.force meths_outcomes))
     else
-      (case get_best_method_outcome peek_at_outcome meths_outcomes of
-        (_, Not_Played) => snd (get_best_method_outcome Lazy.force meths_outcomes)
-      | (_, outcome) => outcome)
+      let val outcome = snd (get_best_method_outcome peek_at_outcome meths_outcomes) in
+        if outcome = Play_Timed_Out Time.zeroTime then
+          snd (get_best_method_outcome Lazy.force meths_outcomes)
+        else
+          outcome
+      end
   end
 
 fun preplay_outcome_of_isar_step_for_method preplay_data l =
   AList.lookup (op =) (the (Canonical_Label_Tab.lookup preplay_data l))
-  #> the_default (Lazy.value Not_Played)
+  #> the_default (Lazy.value (Play_Timed_Out Time.zeroTime))
 
 fun fastest_method_of_isar_step preplay_data =
   the o Canonical_Label_Tab.lookup preplay_data
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_isar_proof.ML	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_isar_proof.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -259,7 +259,7 @@
 
     fun is_direct_method (Metis_Method _) = true
       | is_direct_method Meson_Method = true
-      | is_direct_method SMT_Method = true
+      | is_direct_method SMT2_Method = true
       | is_direct_method _ = false
 
     (* Local facts are always passed via "using", which affects "meson" and "metis". This is
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_proof_methods.ML	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_proof_methods.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -12,7 +12,7 @@
   datatype proof_method =
     Metis_Method of string option * string option |
     Meson_Method |
-    SMT_Method |
+    SMT2_Method |
     Blast_Method |
     Simp_Method |
     Simp_Size_Method |
@@ -26,8 +26,7 @@
   datatype play_outcome =
     Played of Time.time |
     Play_Timed_Out of Time.time |
-    Play_Failed |
-    Not_Played
+    Play_Failed
 
   type minimize_command = string list -> string
   type one_line_params =
@@ -51,7 +50,7 @@
 datatype proof_method =
   Metis_Method of string option * string option |
   Meson_Method |
-  SMT_Method |
+  SMT2_Method |
   Blast_Method |
   Simp_Method |
   Simp_Size_Method |
@@ -65,8 +64,7 @@
 datatype play_outcome =
   Played of Time.time |
   Play_Timed_Out of Time.time |
-  Play_Failed |
-  Not_Played
+  Play_Failed
 
 type minimize_command = string list -> string
 type one_line_params =
@@ -78,7 +76,7 @@
   | Metis_Method (type_enc_opt, lam_trans_opt) =>
     "metis (" ^ commas (map_filter I [type_enc_opt, lam_trans_opt]) ^ ")"
   | Meson_Method => "meson"
-  | SMT_Method => "smt"
+  | SMT2_Method => "smt2"
   | Blast_Method => "blast"
   | Simp_Method => "simp"
   | Simp_Size_Method => "simp add: size_ne_size_imp_ne"
@@ -93,7 +91,7 @@
    exceeded" warnings and the like. *)
 fun silence_proof_methods debug =
   Try0.silence_methods debug
-  #> Config.put SMT_Config.verbose debug
+  #> Config.put SMT2_Config.verbose debug
 
 fun tac_of_proof_method ctxt0 debug (local_facts, global_facts) meth =
   let val ctxt = silence_proof_methods debug ctxt0 in
@@ -103,8 +101,7 @@
       Metis_Tactic.metis_tac [type_enc_opt |> the_default (hd partial_type_encs)]
         (lam_trans_opt |> the_default default_metis_lam_trans) ctxt global_facts
     | Meson_Method => Meson.meson_tac ctxt global_facts
-
-    | SMT_Method => SMT_Solver.smt_tac ctxt global_facts
+    | SMT2_Method => SMT2_Solver.smt2_tac ctxt global_facts
     | _ =>
       Method.insert_tac global_facts THEN'
       (case meth of
@@ -121,17 +118,14 @@
   end
 
 fun string_of_play_outcome (Played time) = string_of_ext_time (false, time)
-  | string_of_play_outcome (Play_Timed_Out time) = string_of_ext_time (true, time) ^ ", timed out"
+  | string_of_play_outcome (Play_Timed_Out time) =
+    if time = Time.zeroTime then "" else string_of_ext_time (true, time) ^ ", timed out"
   | string_of_play_outcome Play_Failed = "failed"
-  | string_of_play_outcome _ = "unknown"
 
 fun play_outcome_ord (Played time1, Played time2) =
     int_ord (pairself Time.toMilliseconds (time1, time2))
   | play_outcome_ord (Played _, _) = LESS
   | play_outcome_ord (_, Played _) = GREATER
-  | play_outcome_ord (Not_Played, Not_Played) = EQUAL
-  | play_outcome_ord (Not_Played, _) = LESS
-  | play_outcome_ord (_, Not_Played) = GREATER
   | play_outcome_ord (Play_Timed_Out time1, Play_Timed_Out time2) =
     int_ord (pairself Time.toMilliseconds (time1, time2))
   | play_outcome_ord (Play_Timed_Out _, _) = LESS
@@ -160,11 +154,11 @@
   (* unusing_chained_facts used_chaineds num_chained ^ *)
   apply_on_subgoal i n ^ command_call (string_of_proof_method meth) ss
 
-fun show_time NONE = ""
-  | show_time (SOME ext_time) = " (" ^ string_of_ext_time ext_time ^ ")"
-
-fun try_command_line banner time command =
-  banner ^ ": " ^ Active.sendback_markup [Markup.padding_command] command ^ show_time time ^ "."
+fun try_command_line banner play command =
+  let val s = string_of_play_outcome play in
+    banner ^ ": " ^ Active.sendback_markup [Markup.padding_command] command ^
+    (s |> s <> "" ? enclose " (" ")") ^ "."
+  end
 
 fun minimize_line _ [] = ""
   | minimize_line minimize_command ss =
@@ -182,18 +176,11 @@
   let
     val (chained, extra) = split_used_facts used_facts
 
-    val (failed, ext_time) =
-      (case play of
-        Played time => (false, (SOME (false, time)))
-      | Play_Timed_Out time => (false, SOME (true, time))
-      | Play_Failed => (true, NONE)
-      | Not_Played => (false, NONE))
-
     val try_line =
       map fst extra
       |> proof_method_command meth subgoal subgoal_count (map fst chained) num_chained
-      |> (if failed then enclose "One-line proof reconstruction failed: " "."
-          else try_command_line banner ext_time)
+      |> (if play = Play_Failed then enclose "One-line proof reconstruction failed: " "."
+          else try_command_line banner play)
   in
     try_line ^ minimize_line minimize_command (map fst (extra @ chained))
   end
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_prover.ML	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_prover.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -184,7 +184,7 @@
       Metis_Method (SOME full_typesN, NONE),
       Metis_Method (SOME no_typesN, SOME desperate_lam_trans),
       Metis_Method (SOME really_full_type_enc, SOME desperate_lam_trans)]) @
-  (if smt_proofs then [SMT_Method] else [])
+  (if smt_proofs then [SMT2_Method] else [])
 
 fun extract_proof_method ({type_enc, lam_trans, ...} : params)
       (Metis_Method (type_enc', lam_trans')) =
@@ -195,7 +195,7 @@
         (if is_none lam_trans' andalso is_none lam_trans then []
          else [("lam_trans", [lam_trans' |> the_default default_metis_lam_trans])])
     in (metisN, override_params) end
-  | extract_proof_method _ SMT_Method = (smtN, [])
+  | extract_proof_method _ SMT2_Method = (smtN, [])
 
 (* based on "Mirabelle.can_apply" and generalized *)
 fun timed_apply timeout tac state i =
@@ -219,7 +219,7 @@
     fun get_preferred meths = if member (op =) meths preferred then preferred else meth
   in
     if timeout = Time.zeroTime then
-      (get_preferred meths, Not_Played)
+      (get_preferred meths, Play_Timed_Out Time.zeroTime)
     else
       let
         val _ = if mode = Minimize then Output.urgent_message "Preplaying proof..." else ()
@@ -247,9 +247,10 @@
   end
 
 val canonical_isar_supported_prover = eN
+val z3_newN = "z3_new"
 
 fun isar_supported_prover_of thy name =
-  if is_atp thy name then name
+  if is_atp thy name orelse name = z3_newN then name
   else if is_atp_installed thy canonical_isar_supported_prover then canonical_isar_supported_prover
   else name
 
@@ -260,7 +261,7 @@
     val maybe_isar_name = name |> isar_proofs = SOME true ? isar_supported_prover_of thy
     val (min_name, override_params) =
       (case preplay of
-        (meth, Played _) =>
+        (meth as Metis_Method _, Played _) =>
         if isar_proofs = SOME true then (maybe_isar_name, []) else extract_proof_method params meth
       | _ => (maybe_isar_name, []))
   in minimize_command override_params min_name end
@@ -279,7 +280,8 @@
     val (remote_provers, local_provers) =
       proof_method_names @
       sort_strings (supported_atps thy) @
-      sort_strings (SMT_Solver.available_solvers_of ctxt)
+      sort_strings (SMT_Solver.available_solvers_of ctxt) @
+      sort_strings (SMT2_Config.available_solvers_of ctxt)
       |> List.partition (String.isPrefix remote_prefix)
   in
     Output.urgent_message ("Supported provers: " ^
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_prover_minimize.ML	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_prover_minimize.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -49,6 +49,7 @@
 open Sledgehammer_Prover
 open Sledgehammer_Prover_ATP
 open Sledgehammer_Prover_SMT
+open Sledgehammer_Prover_SMT2
 
 fun run_proof_method mode name (params as {debug, verbose, timeout, type_enc, lam_trans, ...})
     minimize_command
@@ -56,7 +57,7 @@
   let
     val meth =
       if name = metisN then Metis_Method (type_enc, lam_trans)
-      else if name = smtN then SMT_Method
+      else if name = smtN then SMT2_Method
       else raise Fail ("unknown proof_method: " ^ quote name)
     val used_facts = facts |> map fst
   in
@@ -89,11 +90,12 @@
 
 fun is_prover_supported ctxt =
   let val thy = Proof_Context.theory_of ctxt in
-    is_proof_method orf is_atp thy orf is_smt_prover ctxt
+    is_proof_method orf is_atp thy orf is_smt_prover ctxt orf is_smt2_prover ctxt
   end
 
 fun is_prover_installed ctxt =
-  is_proof_method orf is_smt_prover ctxt orf is_atp_installed (Proof_Context.theory_of ctxt)
+  is_proof_method orf is_smt_prover ctxt orf is_smt2_prover ctxt orf
+  is_atp_installed (Proof_Context.theory_of ctxt)
 
 val proof_method_default_max_facts = 20
 
@@ -103,8 +105,12 @@
       proof_method_default_max_facts
     else if is_atp thy name then
       fold (Integer.max o fst o #1 o fst o snd) (#best_slices (get_atp thy name ()) ctxt) 0
-    else (* is_smt_prover ctxt name *)
+    else if is_smt_prover ctxt name then
       SMT_Solver.default_max_relevant ctxt name
+    else if is_smt2_prover ctxt name then
+      SMT2_Solver.default_max_relevant ctxt name
+    else
+      error ("No such prover: " ^ name ^ ".")
   end
 
 fun get_prover ctxt mode name =
@@ -112,6 +118,7 @@
     if is_proof_method name then run_proof_method mode name
     else if is_atp thy name then run_atp mode name
     else if is_smt_prover ctxt name then run_smt_solver mode name
+    else if is_smt2_prover ctxt name then run_smt2_solver mode name
     else error ("No such prover: " ^ name ^ ".")
   end
 
@@ -344,7 +351,7 @@
                                   adjust_proof_method_params override_params params))
                  else
                    (prover_fast_enough (), (name, params))
-               | (SMT_Method, Played timeout) =>
+               | (SMT2_Method, Played timeout) =>
                  (* Cheat: Assume the original prover is as fast as "smt" for the goal it proved
                     itself. *)
                  (can_min_fast_enough timeout, (name, params))
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_prover_smt.ML	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_prover_smt.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -234,7 +234,7 @@
         NONE =>
         (Lazy.lazy (fn () =>
            play_one_line_proof mode debug verbose preplay_timeout used_pairs state subgoal
-             SMT_Method (bunch_of_proof_methods (smt_proofs <> SOME false) false liftingN)),
+             SMT2_Method (bunch_of_proof_methods (smt_proofs <> SOME false) false liftingN)),
          fn preplay =>
             let
               val one_line_params =
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_prover_smt2.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -0,0 +1,270 @@
+(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_prover_smt2.ML
+    Author:     Fabian Immler, TU Muenchen
+    Author:     Makarius
+    Author:     Jasmin Blanchette, TU Muenchen
+
+SMT solvers as Sledgehammer provers.
+*)
+
+signature SLEDGEHAMMER_PROVER_SMT2 =
+sig
+  type stature = ATP_Problem_Generate.stature
+  type mode = Sledgehammer_Prover.mode
+  type prover = Sledgehammer_Prover.prover
+
+  val smt2_builtins : bool Config.T
+  val smt2_triggers : bool Config.T
+  val smt2_weights : bool Config.T
+  val smt2_weight_min_facts : int Config.T
+  val smt2_min_weight : int Config.T
+  val smt2_max_weight : int Config.T
+  val smt2_max_weight_index : int Config.T
+  val smt2_weight_curve : (int -> int) Unsynchronized.ref
+  val smt2_max_slices : int Config.T
+  val smt2_slice_fact_frac : real Config.T
+  val smt2_slice_time_frac : real Config.T
+  val smt2_slice_min_secs : int Config.T
+
+  val is_smt2_prover : Proof.context -> string -> bool
+  val run_smt2_solver : mode -> string -> prover
+end;
+
+structure Sledgehammer_Prover_SMT2 : SLEDGEHAMMER_PROVER_SMT2 =
+struct
+
+open ATP_Util
+open ATP_Proof
+open ATP_Systems
+open ATP_Problem_Generate
+open ATP_Proof_Reconstruct
+open Sledgehammer_Util
+open Sledgehammer_Proof_Methods
+open Sledgehammer_Isar
+open Sledgehammer_Prover
+
+val smt2_builtins = Attrib.setup_config_bool @{binding sledgehammer_smt2_builtins} (K true)
+val smt2_triggers = Attrib.setup_config_bool @{binding sledgehammer_smt2_triggers} (K true)
+val smt2_weights = Attrib.setup_config_bool @{binding sledgehammer_smt2_weights} (K true)
+val smt2_weight_min_facts =
+  Attrib.setup_config_int @{binding sledgehammer_smt2_weight_min_facts} (K 20)
+
+val is_smt2_prover = member (op =) o SMT2_Config.available_solvers_of
+
+(* FUDGE *)
+val smt2_min_weight = Attrib.setup_config_int @{binding sledgehammer_smt2_min_weight} (K 0)
+val smt2_max_weight = Attrib.setup_config_int @{binding sledgehammer_smt2_max_weight} (K 10)
+val smt2_max_weight_index =
+  Attrib.setup_config_int @{binding sledgehammer_smt2_max_weight_index} (K 200)
+val smt2_weight_curve = Unsynchronized.ref (fn x : int => x * x)
+
+fun smt2_fact_weight ctxt j num_facts =
+  if Config.get ctxt smt2_weights andalso num_facts >= Config.get ctxt smt2_weight_min_facts then
+    let
+      val min = Config.get ctxt smt2_min_weight
+      val max = Config.get ctxt smt2_max_weight
+      val max_index = Config.get ctxt smt2_max_weight_index
+      val curve = !smt2_weight_curve
+    in
+      SOME (max - (max - min + 1) * curve (Int.max (0, max_index - j - 1)) div curve max_index)
+    end
+  else
+    NONE
+
+fun weight_smt2_fact ctxt num_facts ((info, th), j) =
+  let val thy = Proof_Context.theory_of ctxt in
+    (info, (smt2_fact_weight ctxt j num_facts, Thm.transfer thy th (* TODO: needed? *)))
+  end
+
+(* "SMT2_Failure.Abnormal_Termination" carries the solver's return code. Until these are sorted out
+   properly in the SMT module, we must interpret these here. *)
+val z3_failures =
+  [(101, OutOfResources),
+   (103, MalformedInput),
+   (110, MalformedInput),
+   (112, TimedOut)]
+val unix_failures =
+  [(138, Crashed),
+   (139, Crashed)]
+val smt2_failures = z3_failures @ unix_failures
+
+fun failure_of_smt2_failure (SMT2_Failure.Counterexample {is_real_cex, ...}) =
+    if is_real_cex then Unprovable else GaveUp
+  | failure_of_smt2_failure SMT2_Failure.Time_Out = TimedOut
+  | failure_of_smt2_failure (SMT2_Failure.Abnormal_Termination code) =
+    (case AList.lookup (op =) smt2_failures code of
+      SOME failure => failure
+    | NONE => UnknownError ("Abnormal termination with exit code " ^ string_of_int code ^ "."))
+  | failure_of_smt2_failure SMT2_Failure.Out_Of_Memory = OutOfResources
+  | failure_of_smt2_failure (SMT2_Failure.Other_Failure s) = UnknownError s
+
+(* FUDGE *)
+val smt2_max_slices = Attrib.setup_config_int @{binding sledgehammer_smt2_max_slices} (K 8)
+val smt2_slice_fact_frac =
+  Attrib.setup_config_real @{binding sledgehammer_smt2_slice_fact_frac} (K 0.667)
+val smt2_slice_time_frac =
+  Attrib.setup_config_real @{binding sledgehammer_smt2_slice_time_frac} (K 0.333)
+val smt2_slice_min_secs = Attrib.setup_config_int @{binding sledgehammer_smt2_slice_min_secs} (K 3)
+
+val is_boring_builtin_typ =
+  not o exists_subtype (member (op =) [@{typ nat}, @{typ int}, HOLogic.realT])
+
+fun smt2_filter_loop name ({debug, overlord, max_mono_iters, max_new_mono_instances, timeout, slice,
+      ...} : params) state goal i =
+  let
+    fun repair_context ctxt =
+      ctxt |> Context.proof_map (SMT2_Config.select_solver name)
+           |> Config.put SMT2_Config.verbose debug
+           |> (if overlord then
+                 Config.put SMT2_Config.debug_files
+                   (overlord_file_location_of_prover name |> (fn (path, name) => path ^ "/" ^ name))
+               else
+                 I)
+           |> Config.put SMT2_Config.infer_triggers (Config.get ctxt smt2_triggers)
+           |> not (Config.get ctxt smt2_builtins)
+              ? (SMT2_Builtin.filter_builtins is_boring_builtin_typ
+                 #> Config.put SMT2_Systems.z3_extensions false)
+           |> repair_monomorph_context max_mono_iters default_max_mono_iters max_new_mono_instances
+                default_max_new_mono_instances
+
+    val state = Proof.map_context (repair_context) state
+    val ctxt = Proof.context_of state
+    val max_slices = if slice then Config.get ctxt smt2_max_slices else 1
+
+    fun do_slice timeout slice outcome0 time_so_far
+        (weighted_factss as (fact_filter, weighted_facts) :: _) =
+      let
+        val timer = Timer.startRealTimer ()
+        val slice_timeout =
+          if slice < max_slices then
+            let val ms = Time.toMilliseconds timeout in
+              Int.min (ms, Int.max (1000 * Config.get ctxt smt2_slice_min_secs,
+                Real.ceil (Config.get ctxt smt2_slice_time_frac * Real.fromInt ms)))
+              |> Time.fromMilliseconds
+            end
+          else
+            timeout
+        val num_facts = length weighted_facts
+        val _ =
+          if debug then
+            quote name ^ " slice " ^ string_of_int slice ^ " with " ^ string_of_int num_facts ^
+            " fact" ^ plural_s num_facts ^ " for " ^ string_of_time slice_timeout
+            |> Output.urgent_message
+          else
+            ()
+        val birth = Timer.checkRealTimer timer
+        val _ = if debug then Output.urgent_message "Invoking SMT solver..." else ()
+
+        val filter_result as {outcome, ...} =
+          SMT2_Solver.smt2_filter ctxt goal weighted_facts i slice_timeout
+          handle exn =>
+            if Exn.is_interrupt exn orelse debug then
+              reraise exn
+            else
+              {outcome = SOME (SMT2_Failure.Other_Failure (ML_Compiler.exn_message exn)),
+               rewrite_rules = [], conjecture_id = ~1, helper_ids = [], fact_ids = [],
+               z3_proof = []}
+
+        val death = Timer.checkRealTimer timer
+        val outcome0 = if is_none outcome0 then SOME outcome else outcome0
+        val time_so_far = Time.+ (time_so_far, Time.- (death, birth))
+        val timeout = Time.- (timeout, Timer.checkRealTimer timer)
+
+        val too_many_facts_perhaps =
+          (case outcome of
+            NONE => false
+          | SOME (SMT2_Failure.Counterexample _) => false
+          | SOME SMT2_Failure.Time_Out => slice_timeout <> timeout
+          | SOME (SMT2_Failure.Abnormal_Termination _) => true (* kind of *)
+          | SOME SMT2_Failure.Out_Of_Memory => true
+          | SOME (SMT2_Failure.Other_Failure _) => true)
+      in
+        if too_many_facts_perhaps andalso slice < max_slices andalso num_facts > 0 andalso
+           Time.> (timeout, Time.zeroTime) then
+          let
+            val new_num_facts =
+              Real.ceil (Config.get ctxt smt2_slice_fact_frac * Real.fromInt num_facts)
+            val weighted_factss as (new_fact_filter, _) :: _ =
+              weighted_factss
+              |> (fn (x :: xs) => xs @ [x])
+              |> app_hd (apsnd (take new_num_facts))
+            val show_filter = fact_filter <> new_fact_filter
+
+            fun num_of_facts fact_filter num_facts =
+              string_of_int num_facts ^ (if show_filter then " " ^ quote fact_filter else "") ^
+              " fact" ^ plural_s num_facts
+
+            val _ =
+              if debug then
+                quote name ^ " invoked with " ^
+                num_of_facts fact_filter num_facts ^ ": " ^
+                string_of_atp_failure (failure_of_smt2_failure (the outcome)) ^
+                " Retrying with " ^ num_of_facts new_fact_filter new_num_facts ^
+                "..."
+                |> Output.urgent_message
+              else
+                ()
+          in
+            do_slice timeout (slice + 1) outcome0 time_so_far weighted_factss
+          end
+        else
+          {outcome = if is_none outcome then NONE else the outcome0, filter_result = filter_result,
+           used_from = map (apsnd snd) weighted_facts, run_time = time_so_far}
+      end
+  in
+    do_slice timeout 1 NONE Time.zeroTime
+  end
+
+fun run_smt2_solver mode name (params as {debug, verbose, isar_proofs, compress_isar,
+      try0_isar, smt_proofs, minimize, preplay_timeout, ...})
+    minimize_command ({state, goal, subgoal, subgoal_count, factss, ...} : prover_problem) =
+  let
+    val thy = Proof.theory_of state
+    val ctxt = Proof.context_of state
+
+    fun weight_facts facts =
+      let val num_facts = length facts in
+        map (weight_smt2_fact ctxt num_facts) (facts ~~ (0 upto num_facts - 1))
+      end
+
+    val weighted_factss = map (apsnd weight_facts) factss
+    val {outcome, filter_result = {conjecture_id, rewrite_rules, helper_ids, fact_ids, z3_proof,
+           ...}, used_from, run_time} =
+      smt2_filter_loop name params state goal subgoal weighted_factss
+    val used_named_facts = map snd fact_ids
+    val used_facts = map fst used_named_facts
+    val outcome = Option.map failure_of_smt2_failure outcome
+
+    val (preplay, message, message_tail) =
+      (case outcome of
+        NONE =>
+        (Lazy.lazy (fn () =>
+           play_one_line_proof mode debug verbose preplay_timeout used_named_facts state subgoal
+             SMT2_Method (bunch_of_proof_methods (smt_proofs <> SOME false) false liftingN)),
+         fn preplay =>
+            let
+              val fact_ids =
+                map (fn (id, th) => (id, short_thm_name ctxt th)) helper_ids @
+                map (fn (id, ((name, _), _)) => (id, name)) fact_ids
+              val atp_proof = Z3_New_Isar.atp_proof_of_z3_proof thy rewrite_rules conjecture_id
+                fact_ids z3_proof
+              val isar_params =
+                K (verbose, (NONE, NONE), preplay_timeout, compress_isar, try0_isar,
+                   minimize <> SOME false, atp_proof, goal)
+              val one_line_params =
+                (preplay, proof_banner mode name, used_facts,
+                 choose_minimize_command thy params minimize_command name preplay, subgoal,
+                 subgoal_count)
+              val num_chained = length (#facts (Proof.goal state))
+            in
+              proof_text ctxt debug isar_proofs smt_proofs isar_params num_chained one_line_params
+            end,
+         if verbose then "\nSMT solver real CPU time: " ^ string_of_time run_time ^ "." else "")
+      | SOME failure =>
+        (Lazy.value (Metis_Method (NONE, NONE), Play_Failed),
+         fn _ => string_of_atp_failure failure, ""))
+  in
+    {outcome = outcome, used_facts = used_facts, used_from = used_from, run_time = run_time,
+     preplay = preplay, message = message, message_tail = message_tail}
+  end
+
+end;
--- a/src/HOL/Tools/simpdata.ML	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Tools/simpdata.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -109,10 +109,21 @@
 fun mksimps pairs (_: Proof.context) =
   map_filter (try mk_eq) o mk_atomize pairs o gen_all;
 
+val simp_legacy_precond =
+  Attrib.setup_config_bool @{binding "simp_legacy_precond"} (K false)
+
 fun unsafe_solver_tac ctxt =
-  (fn i => REPEAT_DETERM (match_tac @{thms simp_impliesI} i)) THEN'
-  FIRST' [resolve_tac (reflexive_thm :: @{thm TrueI} :: @{thm refl} :: Simplifier.prems_of ctxt),
-    atac, etac @{thm FalseE}];
+  let
+    val intros =
+      if Config.get ctxt simp_legacy_precond then [] else [@{thm conjI}]
+    val sol_thms =
+      reflexive_thm :: @{thm TrueI} :: @{thm refl} :: Simplifier.prems_of ctxt;
+    fun sol_tac i =
+      FIRST [resolve_tac sol_thms i, atac i , etac @{thm FalseE} i] ORELSE
+      (match_tac intros THEN_ALL_NEW sol_tac) i
+  in
+    (fn i => REPEAT_DETERM (match_tac @{thms simp_impliesI} i)) THEN' sol_tac
+  end;
 
 val unsafe_solver = mk_solver "HOL unsafe" unsafe_solver_tac;
 
--- a/src/HOL/Transfer.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Transfer.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -156,10 +156,10 @@
 by(simp add: bi_unique_def)
 
 lemma right_uniqueI: "(\<And>x y z. \<lbrakk> A x y; A x z \<rbrakk> \<Longrightarrow> y = z) \<Longrightarrow> right_unique A"
-unfolding right_unique_def by blast
+unfolding right_unique_def by fast
 
 lemma right_uniqueD: "\<lbrakk> right_unique A; A x y; A x z \<rbrakk> \<Longrightarrow> y = z"
-unfolding right_unique_def by blast
+unfolding right_unique_def by fast
 
 lemma right_total_alt_def:
   "right_total R \<longleftrightarrow> ((R ===> op \<longrightarrow>) ===> op \<longrightarrow>) All All"
@@ -204,18 +204,18 @@
   by auto
 
 lemma bi_total_OO: "\<lbrakk>bi_total A; bi_total B\<rbrakk> \<Longrightarrow> bi_total (A OO B)"
-  unfolding bi_total_def OO_def by metis
+  unfolding bi_total_def OO_def by fast
 
 lemma bi_unique_OO: "\<lbrakk>bi_unique A; bi_unique B\<rbrakk> \<Longrightarrow> bi_unique (A OO B)"
-  unfolding bi_unique_def OO_def by metis
+  unfolding bi_unique_def OO_def by blast
 
 lemma right_total_OO:
   "\<lbrakk>right_total A; right_total B\<rbrakk> \<Longrightarrow> right_total (A OO B)"
-  unfolding right_total_def OO_def by metis
+  unfolding right_total_def OO_def by fast
 
 lemma right_unique_OO:
   "\<lbrakk>right_unique A; right_unique B\<rbrakk> \<Longrightarrow> right_unique (A OO B)"
-  unfolding right_unique_def OO_def by metis
+  unfolding right_unique_def OO_def by fast
 
 
 subsection {* Properties of relators *}
@@ -278,7 +278,7 @@
 lemma bi_unique_fun [transfer_rule]:
   "\<lbrakk>bi_total A; bi_unique B\<rbrakk> \<Longrightarrow> bi_unique (A ===> B)"
   unfolding bi_total_def bi_unique_def rel_fun_def fun_eq_iff
-  by (safe, metis, fast)
+  by fast+
 
 
 subsection {* Transfer rules *}
@@ -289,7 +289,7 @@
     (transfer_bforall (Domainp A)) transfer_forall"
   using assms unfolding right_total_def
   unfolding transfer_forall_def transfer_bforall_def rel_fun_def Domainp_iff
-  by metis
+  by fast
 
 text {* Transfer rules using implication instead of equality on booleans. *}
 
@@ -300,7 +300,7 @@
   "bi_total A \<Longrightarrow> ((A ===> op =) ===> rev_implies) transfer_forall transfer_forall"
   "bi_total A \<Longrightarrow> ((A ===> rev_implies) ===> rev_implies) transfer_forall transfer_forall"
   unfolding transfer_forall_def rev_implies_def rel_fun_def right_total_def bi_total_def
-  by metis+
+  by fast+
 
 lemma transfer_implies_transfer [transfer_rule]:
   "(op =        ===> op =        ===> op =       ) transfer_implies transfer_implies"
@@ -327,13 +327,13 @@
   assumes "right_total A"
   shows "((A ===> op=) ===> op=) (Bex (Collect (Domainp A))) Ex"
 using assms unfolding right_total_def Bex_def rel_fun_def Domainp_iff[abs_def]
-by blast
+by fast
 
 lemma right_total_All_transfer[transfer_rule]:
   assumes "right_total A"
   shows "((A ===> op =) ===> op =) (Ball (Collect (Domainp A))) All"
 using assms unfolding right_total_def Ball_def rel_fun_def Domainp_iff[abs_def]
-by blast
+by fast
 
 lemma All_transfer [transfer_rule]:
   assumes "bi_total A"
--- a/src/HOL/Word/Bool_List_Representation.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Word/Bool_List_Representation.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -941,7 +941,6 @@
   apply (subst bin_rsplit_aux.simps)
   apply (subst bin_rsplit_aux.simps)
   apply (clarsimp split: prod.split)
-  apply auto
   done
 
 lemma bin_rsplitl_aux_append:
@@ -950,7 +949,6 @@
   apply (subst bin_rsplitl_aux.simps)
   apply (subst bin_rsplitl_aux.simps)
   apply (clarsimp split: prod.split)
-  apply auto
   done
 
 lemmas rsplit_aux_apps [where bs = "[]"] =
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Word/Tools/smt2_word.ML	Fri Mar 14 17:32:11 2014 +0100
@@ -0,0 +1,144 @@
+(*  Title:      HOL/Word/Tools/smt2_word.ML
+    Author:     Sascha Boehme, TU Muenchen
+
+SMT setup for words.
+*)
+
+structure SMT2_Word: sig end =
+struct
+
+open Word_Lib
+
+(* SMT-LIB logic *)
+
+fun smtlib_logic ts =
+  if exists (Term.exists_type (Term.exists_subtype is_wordT)) ts
+  then SOME "QF_AUFBV"
+  else NONE
+
+
+(* SMT-LIB builtins *)
+
+local
+  val smtlib2C = SMTLIB2_Interface.smtlib2C
+
+  val wordT = @{typ "'a::len word"}
+
+  fun index1 s i = "(_ " ^ s ^ " " ^ string_of_int i ^ ")"
+  fun index2 s i j = "(_ " ^ s ^ " " ^ string_of_int i ^ " " ^ string_of_int j ^ ")"
+
+  fun word_typ (Type (@{type_name word}, [T])) = Option.map (index1 "BitVec") (try dest_binT T)
+    | word_typ _ = NONE
+
+  fun word_num (Type (@{type_name word}, [T])) k =
+        Option.map (index1 ("bv" ^ string_of_int k)) (try dest_binT T)
+    | word_num _ _ = NONE
+
+  fun if_fixed pred m n T ts =
+    let val (Us, U) = Term.strip_type T
+    in
+      if pred (U, Us) then
+        SOME (n, length Us, ts, Term.list_comb o pair (Const (m, T)))
+      else NONE
+    end
+
+  fun if_fixed_all m = if_fixed (forall (can dest_wordT) o (op ::)) m
+  fun if_fixed_args m = if_fixed (forall (can dest_wordT) o snd) m
+
+  fun add_word_fun f (t, n) =
+    let val (m, _) = Term.dest_Const t
+    in SMT2_Builtin.add_builtin_fun smtlib2C (Term.dest_Const t, K (f m n)) end
+
+  fun hd2 xs = hd (tl xs)
+
+  fun mk_nat i = @{const nat} $ HOLogic.mk_number @{typ nat} i
+
+  fun dest_nat (@{const nat} $ n) = snd (HOLogic.dest_number n)
+    | dest_nat t = raise TERM ("not a natural number", [t])
+
+  fun mk_shift c [t, u] = Const c $ t $ mk_nat (snd (HOLogic.dest_number u))
+    | mk_shift c ts = raise TERM ("bad arguments", Const c :: ts)
+
+  fun shift m n T ts =
+    let val U = Term.domain_type T
+    in
+      (case (can dest_wordT U, try (dest_nat o hd2) ts) of
+        (true, SOME i) =>
+          SOME (n, 2, [hd ts, HOLogic.mk_number U i], mk_shift (m, T))
+      | _ => NONE)   (* FIXME: also support non-numerical shifts *)
+    end
+
+  fun mk_extract c i ts = Term.list_comb (Const c, mk_nat i :: ts)
+
+  fun extract m n T ts =
+    let val U = Term.range_type (Term.range_type T)
+    in
+      (case (try (dest_nat o hd) ts, try dest_wordT U) of
+        (SOME lb, SOME i) =>
+          SOME (index2 n (i + lb - 1) lb, 1, tl ts, mk_extract (m, T) lb)
+      | _ => NONE)
+    end
+
+  fun mk_extend c ts = Term.list_comb (Const c, ts)
+
+  fun extend m n T ts =
+    let val (U1, U2) = Term.dest_funT T
+    in
+      (case (try dest_wordT U1, try dest_wordT U2) of
+        (SOME i, SOME j) =>
+          if j-i >= 0 then SOME (index1 n (j-i), 1, ts, mk_extend (m, T))
+          else NONE
+      | _ => NONE)
+    end
+
+  fun mk_rotate c i ts = Term.list_comb (Const c, mk_nat i :: ts)
+
+  fun rotate m n T ts =
+    let val U = Term.domain_type (Term.range_type T)
+    in
+      (case (can dest_wordT U, try (dest_nat o hd) ts) of
+        (true, SOME i) => SOME (index1 n i, 1, tl ts, mk_rotate (m, T) i)
+      | _ => NONE)
+    end
+in
+
+val setup_builtins =
+  SMT2_Builtin.add_builtin_typ smtlib2C (wordT, word_typ, word_num) #>
+  fold (add_word_fun if_fixed_all) [
+    (@{term "uminus :: 'a::len word => _"}, "bvneg"),
+    (@{term "plus :: 'a::len word => _"}, "bvadd"),
+    (@{term "minus :: 'a::len word => _"}, "bvsub"),
+    (@{term "times :: 'a::len word => _"}, "bvmul"),
+    (@{term "bitNOT :: 'a::len word => _"}, "bvnot"),
+    (@{term "bitAND :: 'a::len word => _"}, "bvand"),
+    (@{term "bitOR :: 'a::len word => _"}, "bvor"),
+    (@{term "bitXOR :: 'a::len word => _"}, "bvxor"),
+    (@{term "word_cat :: 'a::len word => _"}, "concat") ] #>
+  fold (add_word_fun shift) [
+    (@{term "shiftl :: 'a::len word => _ "}, "bvshl"),
+    (@{term "shiftr :: 'a::len word => _"}, "bvlshr"),
+    (@{term "sshiftr :: 'a::len word => _"}, "bvashr") ] #>
+  add_word_fun extract
+    (@{term "slice :: _ => 'a::len word => _"}, "extract") #>
+  fold (add_word_fun extend) [
+    (@{term "ucast :: 'a::len word => _"}, "zero_extend"),
+    (@{term "scast :: 'a::len word => _"}, "sign_extend") ] #>
+  fold (add_word_fun rotate) [
+    (@{term word_rotl}, "rotate_left"),
+    (@{term word_rotr}, "rotate_right") ] #>
+  fold (add_word_fun if_fixed_args) [
+    (@{term "less :: 'a::len word => _"}, "bvult"),
+    (@{term "less_eq :: 'a::len word => _"}, "bvule"),
+    (@{term word_sless}, "bvslt"),
+    (@{term word_sle}, "bvsle") ]
+
+end
+
+
+(* setup *)
+
+val _ = Theory.setup (Context.theory_map (
+  SMTLIB2_Interface.add_logic (20, smtlib_logic) #>
+  setup_builtins))
+
+end
--- a/src/HOL/Word/Word.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/Word/Word.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -4738,6 +4738,7 @@
 ML_file "Tools/word_lib.ML"
 ML_file "Tools/smt_word.ML"
 setup SMT_Word.setup
+ML_file "Tools/smt2_word.ML"
 
 hide_const (open) Word
 
--- a/src/HOL/ZF/HOLZF.thy	Fri Mar 14 16:54:01 2014 +0100
+++ b/src/HOL/ZF/HOLZF.thy	Fri Mar 14 17:32:11 2014 +0100
@@ -261,7 +261,7 @@
   apply (frule Elem_Elem_PFun[where p=y], simp)
   apply (subgoal_tac "isFun F")
   apply (simp add: isFun_def isOpair_def)  
-  apply (auto simp add: Fst Snd, blast)
+  apply (auto simp add: Fst Snd)
   apply (auto simp add: PFun_def Sep)
   done