merged
authorhuffman
Wed, 27 Oct 2010 11:11:35 -0700
changeset 40219 b283680d8044
parent 40218 f7d4d023a899 (current diff)
parent 40211 43916ac560a4 (diff)
child 40226 dfa0d94991e4
child 40321 d065b195ec89
merged
src/HOL/Tools/SMT/cvc3_solver.ML
src/HOL/Tools/SMT/yices_solver.ML
src/HOL/Tools/SMT/z3_solver.ML
src/Tools/auto_solve.ML
--- a/CONTRIBUTORS	Wed Oct 27 11:10:36 2010 -0700
+++ b/CONTRIBUTORS	Wed Oct 27 11:11:35 2010 -0700
@@ -7,6 +7,9 @@
 --------------------------------------
 
 * September 2010: Florian Haftmann, TUM
+  Refined concepts for evaluation, i.e. normalisation of terms using different techniques.
+
+* September 2010: Florian Haftmann, TUM
   Code generation for Scala.
 
 * August 2010: Johannes Hoelzl, Armin Heller, and Robert Himmelmann, TUM
--- a/NEWS	Wed Oct 27 11:10:36 2010 -0700
+++ b/NEWS	Wed Oct 27 11:11:35 2010 -0700
@@ -6,6 +6,8 @@
 
 *** General ***
 
+* Significantly improved Isabelle/Isar implementation manual.
+
 * Explicit treatment of UTF8 sequences as Isabelle symbols, such that
 a Unicode character is treated as a single symbol, not a sequence of
 non-ASCII bytes as before.  Since Isabelle/ML string literals may
@@ -74,6 +76,13 @@
 
 *** HOL ***
 
+* New command 'partial_function' provides basic support for recursive
+function definitions over complete partial orders. Concrete instances
+are provided for i) the option type, ii) tail recursion on arbitrary
+types, and iii) the heap monad of Imperative_HOL. See
+HOL/ex/Fundefs.thy and HOL/Imperative_HOL/ex/Linked_Lists.thy for
+examples.
+
 * Predicates `distinct` and `sorted` now defined inductively, with nice
 induction rules.  INCOMPATIBILITY: former distinct.simps and sorted.simps
 now named distinct_simps and sorted_simps.
@@ -278,6 +287,9 @@
   meson_disj_FalseD2 ~> Meson.disj_FalseD2
 INCOMPATIBILITY.
 
+* Auto Solve: Renamed "Auto Solve Direct". The tool is now available manually as
+  "solve_direct".
+
 * Sledgehammer:
   - Renamed lemmas:
     COMBI_def ~> Meson.COMBI_def
@@ -301,6 +313,15 @@
     sledgehammer [atp = ...] ~> sledgehammer [prover = ...]
     INCOMPATIBILITY.
 
+* Changed SMT configuration options:
+  - Renamed:
+    z3_proofs ~> smt_oracle (with swapped semantics)
+    z3_trace_assms ~> smt_trace_used_facts
+    INCOMPATIBILITY.
+  - Added:
+    cvc3_options
+    yices_options
+    smt_datatypes
 
 *** FOL ***
 
@@ -314,6 +335,10 @@
 
 *** ML ***
 
+* Antiquotation @{assert} inlines a function bool -> unit that raises
+Fail if the argument is false.  Due to inlining the source position of
+failed assertions is included in the error output.
+
 * Renamed setmp_noncritical to Unsynchronized.setmp to emphasize its
 meaning.
 
--- a/doc-src/IsarImplementation/Thy/Base.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/Base.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -3,10 +3,4 @@
 uses "../../antiquote_setup.ML"
 begin
 
-(* FIXME move to src/Pure/ML/ml_antiquote.ML *)
-ML {*
-  ML_Antiquote.inline "assert"
-    (Scan.succeed "(fn b => if b then () else raise General.Fail \"Assertion failed\")")
-*}
-
 end
--- a/doc-src/IsarImplementation/Thy/Integration.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/Integration.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -92,7 +92,7 @@
   \item @{ML "Toplevel.timing := true"} makes the toplevel print timing
   information for each Isar command being executed.
 
-  \item @{ML Toplevel.profiling}~@{verbatim ":="}~@{text "n"} controls
+  \item @{ML Toplevel.profiling}~@{ML_text ":="}~@{text "n"} controls
   low-level profiling of the underlying ML runtime system.  For
   Poly/ML, @{text "n = 1"} means time and @{text "n = 2"} space
   profiling.
@@ -260,7 +260,7 @@
   @{index_ML Thy_Info.get_theory: "string -> theory"} \\
   @{index_ML Thy_Info.remove_thy: "string -> unit"} \\[1ex]
   @{index_ML Thy_Info.register_thy: "theory -> unit"} \\[1ex]
-  @{verbatim "datatype action = Update | Remove"} \\
+  @{ML_text "datatype action = Update | Remove"} \\
   @{index_ML Thy_Info.add_hook: "(Thy_Info.action -> string -> unit) -> unit"} \\
   \end{mldecls}
 
--- a/doc-src/IsarImplementation/Thy/Isar.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/Isar.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -4,9 +4,9 @@
 
 chapter {* Isar language elements *}
 
-text {*
-  The Isar proof language (see also \cite[\S2]{isabelle-isar-ref})
-  consists of three main categories of language elements:
+text {* The Isar proof language (see also
+  \cite[\S2]{isabelle-isar-ref}) consists of three main categories of
+  language elements as follows.
 
   \begin{enumerate}
 
@@ -44,12 +44,12 @@
 text {* A \emph{proof command} is state transition of the Isar/VM
   proof interpreter.
 
-  In principle, Isar proof commands could be defined in
-  user-space as well.  The system is built like that in the first
-  place: part of the commands are primitive, the other part is defined
-  as derived elements.  Adding to the genuine structured proof
-  language requires profound understanding of the Isar/VM machinery,
-  though, so this is beyond the scope of this manual.
+  In principle, Isar proof commands could be defined in user-space as
+  well.  The system is built like that in the first place: one part of
+  the commands are primitive, the other part is defined as derived
+  elements.  Adding to the genuine structured proof language requires
+  profound understanding of the Isar/VM machinery, though, so this is
+  beyond the scope of this manual.
 
   What can be done realistically is to define some diagnostic commands
   that inspect the general state of the Isar/VM, and report some
@@ -58,11 +58,10 @@
   goals (if available).
 
   Another common application is to define a toplevel command that
-  poses a problem to the user as Isar proof state and stores the final
-  result in a suitable context data slot.  Thus a proof can be
+  poses a problem to the user as Isar proof state and processes the
+  final result relatively to the context.  Thus a proof can be
   incorporated into the context of some user-space tool, without
-  modifying the Isar proof language itself.
-*}
+  modifying the Isar proof language itself.  *}
 
 text %mlref {*
   \begin{mldecls}
@@ -131,13 +130,14 @@
   storing named facts).  Note that at this generic level the target
   context is specified as @{ML_type Proof.context}, but the usual
   wrapping of toplevel proofs into command transactions will provide a
-  @{ML_type local_theory} here (see also \chref{ch:local-theory}).
-  This usually affects the way how results are stored.
+  @{ML_type local_theory} here (\chref{ch:local-theory}).  This
+  affects the way how results are stored.
 
   The @{text "statement"} is given as a nested list of terms, each
   associated with optional @{keyword "is"} patterns as usual in the
-  Isar source language.  The original list structure over terms is
-  turned into one over theorems when @{text "after_qed"} is invoked.
+  Isar source language.  The original nested list structure over terms
+  is turned into one over theorems when @{text "after_qed"} is
+  invoked.
 
   \end{description}
 *}
@@ -212,15 +212,15 @@
 
   \end{itemize}
 
-  \medskip Syntactically, the language of proof methods is embedded
-  into Isar as arguments to certain commands, e.g.\ @{command "by"} or
-  @{command apply}.  User-space additions are reasonably easy by
-  plugging suitable method-valued parser functions into the framework,
-  using the @{command method_setup} command, for example.
+  \medskip Syntactically, the language of proof methods appears as
+  arguments to Isar commands like @{command "by"} or @{command apply}.
+  User-space additions are reasonably easy by plugging suitable
+  method-valued parser functions into the framework, using the
+  @{command method_setup} command, for example.
 
   To get a better idea about the range of possibilities, consider the
-  following Isar proof schemes.  First some general form of structured
-  proof text:
+  following Isar proof schemes.  This is the general form of
+  structured proof text:
 
   \medskip
   \begin{tabular}{l}
@@ -231,15 +231,15 @@
   \end{tabular}
   \medskip
 
-  The goal configuration consists of @{text "facts\<^sub>1"} and @{text
-  "facts\<^sub>2"} appended in that order, and various @{text "props"} being
-  claimed here.  The @{text "initial_method"} is invoked with facts
-  and goals together and refines the problem to something that is
-  handled recursively in the proof @{text "body"}.  The @{text
-  "terminal_method"} has another chance to finish-off any remaining
+  The goal configuration consists of @{text "facts\<^sub>1"} and
+  @{text "facts\<^sub>2"} appended in that order, and various @{text
+  "props"} being claimed.  The @{text "initial_method"} is invoked
+  with facts and goals together and refines the problem to something
+  that is handled recursively in the proof @{text "body"}.  The @{text
+  "terminal_method"} has another chance to finish any remaining
   subgoals, but it does not see the facts of the initial step.
 
-  \medskip The second pattern illustrates unstructured proof scripts:
+  \medskip This pattern illustrates unstructured proof scripts:
 
   \medskip
   \begin{tabular}{l}
@@ -251,14 +251,14 @@
   \end{tabular}
   \medskip
 
-  The @{text "method\<^sub>1"} operates on the original claim together while
+  The @{text "method\<^sub>1"} operates on the original claim while
   using @{text "facts\<^sub>1"}.  Since the @{command apply} command
-  structurally resets the facts, the @{text "method\<^sub>2"} will operate on
-  the remaining goal state without facts.  The @{text "method\<^sub>3"} will
-  see again a collection of @{text "facts\<^sub>3"} that has been inserted
-  into the script explicitly.
+  structurally resets the facts, the @{text "method\<^sub>2"} will
+  operate on the remaining goal state without facts.  The @{text
+  "method\<^sub>3"} will see again a collection of @{text
+  "facts\<^sub>3"} that has been inserted into the script explicitly.
 
-  \medskip Empirically, Isar proof methods can be categorized as
+  \medskip Empirically, any Isar proof method can be categorized as
   follows.
 
   \begin{enumerate}
@@ -268,7 +268,8 @@
 
   Example: @{method "induct"}, which is also a ``raw'' method since it
   operates on the internal representation of simultaneous claims as
-  Pure conjunction, instead of separate subgoals.
+  Pure conjunction (\secref{{sec:logic-aux}}), instead of separate
+  subgoals (\secref{sec::tactical-goals}).
 
   \item \emph{Structured method} with strong emphasis on facts outside
   the goal state.
@@ -285,8 +286,8 @@
   \item \emph{Old-style tactic emulation} with detailed numeric goal
   addressing and explicit references to entities of the internal goal
   state (which are otherwise invisible from proper Isar proof text).
-  To make the special non-standard status clear, the naming convention
-  @{text "foo_tac"} needs to be observed.
+  The naming convention @{text "foo_tac"} makes this special
+  non-standard status clear.
 
   Example: @{method "rule_tac"}.
 
@@ -360,7 +361,7 @@
   \medskip The following toy examples illustrate how the goal facts
   and state are passed to proof methods.  The pre-defined proof method
   called ``@{method tactic}'' wraps ML source of type @{ML_type
-  tactic} (abstracted over @{verbatim facts}).  This allows immediate
+  tactic} (abstracted over @{ML_text facts}).  This allows immediate
   experimentation without parsing of concrete syntax. *}
 
 example_proof
@@ -396,9 +397,9 @@
 
   The @{ML Attrib.thms} parser produces a list of theorems from the
   usual Isar syntax involving attribute expressions etc.\ (syntax
-  category @{syntax thmrefs}).  The resulting @{verbatim thms} are
-  added to @{ML HOL_basic_ss} which already contains the basic
-  Simplifier setup for HOL.
+  category @{syntax thmrefs}) \cite{isabelle-isar-ref}.  The resulting
+  @{ML_text thms} are added to @{ML HOL_basic_ss} which already
+  contains the basic Simplifier setup for HOL.
 
   The tactic @{ML asm_full_simp_tac} is the one that is also used in
   method @{method simp} by default.  The extra wrapping by the @{ML
@@ -478,8 +479,8 @@
   ``simple'' methods, i.e.\ the goal facts are merely inserted as goal
   premises by the @{ML SIMPLE_METHOD'} or @{ML SIMPLE_METHOD} wrapper.
   For proof methods that are similar to the standard collection of
-  @{method simp}, @{method blast}, @{method auto} little more can be
-  done here.
+  @{method simp}, @{method blast}, @{method fast}, @{method auto}
+  there is little more that can be done.
 
   Note that using the primary goal facts in the same manner as the
   method arguments obtained via concrete syntax or the context does
@@ -503,8 +504,8 @@
   manner, such as a discrimination net that is indexed by the
   left-hand sides of rewrite rules.  For variations on the Simplifier,
   re-use of the existing type @{ML_type simpset} is adequate, but
-  scalability requires it be maintained statically within the context
-  data, not dynamically on each tool invocation.  *}
+  scalability would require it be maintained statically within the
+  context data, not dynamically on each tool invocation.  *}
 
 
 section {* Attributes \label{sec:attributes} *}
--- a/doc-src/IsarImplementation/Thy/Logic.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/Logic.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -784,7 +784,7 @@
 *}
 
 
-subsection {* Auxiliary definitions *}
+subsection {* Auxiliary definitions \label{sec:logic-aux} *}
 
 text {*
   Theory @{text "Pure"} provides a few auxiliary definitions, see
--- a/doc-src/IsarImplementation/Thy/ML.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/ML.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -10,12 +10,12 @@
   environment.  This covers a variety of aspects that are geared
   towards an efficient and robust platform for applications of formal
   logic with fully foundational proof construction --- according to
-  the well-known \emph{LCF principle}.  There are specific library
-  modules and infrastructure to address the needs for such difficult
-  tasks.  For example, the raw parallel programming model of Poly/ML
-  is presented as considerably more abstract concept of \emph{future
-  values}, which is then used to augment the inference kernel, proof
-  interpreter, and theory loader accordingly.
+  the well-known \emph{LCF principle}.  There is specific
+  infrastructure with library modules to address the needs of this
+  difficult task.  For example, the raw parallel programming model of
+  Poly/ML is presented as considerably more abstract concept of
+  \emph{future values}, which is then used to augment the inference
+  kernel, proof interpreter, and theory loader accordingly.
 
   The main aspects of Isabelle/ML are introduced below.  These
   first-hand explanations should help to understand how proper
@@ -47,14 +47,14 @@
   standardization.  A coding style that is changed every few years or
   with every new contributor is no style at all, because consistency
   is quickly lost.  Global consistency is hard to achieve, though.
-  One should always strife at least for local consistency of modules
-  and sub-systems, without deviating from some general principles how
-  to write Isabelle/ML.
+  Nonetheless, one should always strive at least for local consistency
+  of modules and sub-systems, without deviating from some general
+  principles how to write Isabelle/ML.
 
-  In a sense, a common coding style is like an \emph{orthography} for
-  the sources: it helps to read quickly over the text and see through
-  the main points, without getting distracted by accidental
-  presentation of free-style source.
+  In a sense, good coding style is like an \emph{orthography} for the
+  sources: it helps to read quickly over the text and see through the
+  main points, without getting distracted by accidental presentation
+  of free-style code.
 *}
 
 
@@ -63,7 +63,7 @@
 text {* Isabelle source files have a certain standardized header
   format (with precise spacing) that follows ancient traditions
   reaching back to the earliest versions of the system by Larry
-  Paulson.  E.g.\ see @{"file" "~~/src/Pure/thm.ML"}.
+  Paulson.  See @{"file" "~~/src/Pure/thm.ML"}, for example.
 
   The header includes at least @{verbatim Title} and @{verbatim
   Author} entries, followed by a prose description of the purpose of
@@ -71,8 +71,8 @@
   paragraphs of explanations.
 
   The rest of the file is divided into sections, subsections,
-  subsubsections, paragraphs etc.\ using some a layout via ML comments
-  as follows.
+  subsubsections, paragraphs etc.\ using a simple layout via ML
+  comments as follows.
 
 \begin{verbatim}
 (*** section ***)
@@ -84,8 +84,8 @@
 (*short paragraph*)
 
 (*
-  long paragraph
-  more text
+  long paragraph,
+  with more text
 *)
 \end{verbatim}
 
@@ -94,7 +94,7 @@
   as in the example above).
 
   \medskip The precise wording of the prose text given in these
-  headings is chosen carefully to indicate the main theme of the
+  headings is chosen carefully to introduce the main theme of the
   subsequent formal ML text.
 *}
 
@@ -106,50 +106,48 @@
 
   \paragraph{Notation.}  A name consists of 1--3 \emph{words} (rarely
   4, but not more) that are separated by underscore.  There are three
-  variants concerning upper or lower case letters, which are just for
+  variants concerning upper or lower case letters, which are used for
   certain ML categories as follows:
 
   \medskip
   \begin{tabular}{lll}
   variant & example & ML categories \\\hline
-  lower-case & @{verbatim foo_bar} & values, types, record fields \\
-  capitalized & @{verbatim Foo_Bar} & datatype constructors, \\
-    & & structures, functors \\
-  upper-case & @{verbatim FOO_BAR} & special values (combinators), \\
-    & & exception constructors, signatures \\
+  lower-case & @{ML_text foo_bar} & values, types, record fields \\
+  capitalized & @{ML_text Foo_Bar} & datatype constructors, structures, functors \\
+  upper-case & @{ML_text FOO_BAR} & special values, exception constructors, signatures \\
   \end{tabular}
   \medskip
 
   For historical reasons, many capitalized names omit underscores,
-  e.g.\ old-style @{verbatim FooBar} instead of @{verbatim Foo_Bar}.
-  Genuine mixed-case names are \emph{not} used --- clear division of
-  words is essential for readability.\footnote{Camel-case was invented
-  to workaround the lack of underscore in some early non-ASCII
-  character sets and keywords.  Later it became a habitual in some
-  language communities that are now strong in numbers.}
+  e.g.\ old-style @{ML_text FooBar} instead of @{ML_text Foo_Bar}.
+  Genuine mixed-case names are \emph{not} used, bacause clear division
+  of words is essential for readability.\footnote{Camel-case was
+  invented to workaround the lack of underscore in some early
+  non-ASCII character sets.  Later it became habitual in some language
+  communities that are now strong in numbers.}
 
   A single (capital) character does not count as ``word'' in this
-  respect: some Isabelle/ML are suffixed by extra markers like this:
-  @{verbatim foo_barT}.
+  respect: some Isabelle/ML names are suffixed by extra markers like
+  this: @{ML_text foo_barT}.
 
-  Name variants are produced by adding 1--3 primes, e.g.\ @{verbatim
-  foo'}, @{verbatim foo''}, or @{verbatim foo'''}, but not @{verbatim
+  Name variants are produced by adding 1--3 primes, e.g.\ @{ML_text
+  foo'}, @{ML_text foo''}, or @{ML_text foo'''}, but not @{ML_text
   foo''''} or more.  Decimal digits scale better to larger numbers,
-  e.g.\ @{verbatim foo0}, @{verbatim foo1}, @{verbatim foo42}.
+  e.g.\ @{ML_text foo0}, @{ML_text foo1}, @{ML_text foo42}.
 
   \paragraph{Scopes.}  Apart from very basic library modules, ML
   structures are not ``opened'', but names are referenced with
   explicit qualification, as in @{ML Syntax.string_of_term} for
   example.  When devising names for structures and their components it
   is important aim at eye-catching compositions of both parts, because
-  this is how they are always seen in the sources and documentation.
-  For the same reasons, aliases of well-known library functions should
-  be avoided.
+  this is how they are seen in the sources and documentation.  For the
+  same reasons, aliases of well-known library functions should be
+  avoided.
 
-  Local names of function abstraction or case/lets bindings are
+  Local names of function abstraction or case/let bindings are
   typically shorter, sometimes using only rudiments of ``words'',
   while still avoiding cryptic shorthands.  An auxiliary function
-  called @{verbatim helper}, @{verbatim aux}, or @{verbatim f} is
+  called @{ML_text helper}, @{ML_text aux}, or @{ML_text f} is
   considered bad style.
 
   Example:
@@ -184,23 +182,23 @@
   \end{verbatim}
 
 
-  \paragraph{Specific conventions.} Here are some important specific
-  name forms that occur frequently in the sources.
+  \paragraph{Specific conventions.} Here are some specific name forms
+  that occur frequently in the sources.
 
   \begin{itemize}
 
-  \item A function that maps @{verbatim foo} to @{verbatim bar} is
-  called @{verbatim foo_to_bar} or @{verbatim bar_of_foo} (never
-  @{verbatim foo2bar}, @{verbatim bar_from_foo}, @{verbatim
-  bar_for_foo}, or @{verbatim bar4foo}).
+  \item A function that maps @{ML_text foo} to @{ML_text bar} is
+  called @{ML_text foo_to_bar} or @{ML_text bar_of_foo} (never
+  @{ML_text foo2bar}, @{ML_text bar_from_foo}, @{ML_text
+  bar_for_foo}, or @{ML_text bar4foo}).
 
-  \item The name component @{verbatim legacy} means that the operation
+  \item The name component @{ML_text legacy} means that the operation
   is about to be discontinued soon.
 
-  \item The name component @{verbatim old} means that this is historic
+  \item The name component @{ML_text old} means that this is historic
   material that might disappear at some later stage.
 
-  \item The name component @{verbatim global} means that this works
+  \item The name component @{ML_text global} means that this works
   with the background theory instead of the regular local context
   (\secref{sec:context}), sometimes for historical reasons, sometimes
   due a genuine lack of locality of the concept involved, sometimes as
@@ -211,57 +209,58 @@
 
   \item Variables of the main context types of the Isabelle/Isar
   framework (\secref{sec:context} and \chref{ch:local-theory}) have
-  firm naming conventions to allow to visualize the their data flow
-  via plain regular expressions in the editor.  In particular:
+  firm naming conventions as follows:
 
   \begin{itemize}
 
-  \item theories are called @{verbatim thy}, rarely @{verbatim theory}
-  (never @{verbatim thry})
+  \item theories are called @{ML_text thy}, rarely @{ML_text theory}
+  (never @{ML_text thry})
 
-  \item proof contexts are called @{verbatim ctxt}, rarely @{verbatim
-  context} (never @{verbatim ctx})
+  \item proof contexts are called @{ML_text ctxt}, rarely @{ML_text
+  context} (never @{ML_text ctx})
 
-  \item generic contexts are called @{verbatim context}, rarely
-  @{verbatim ctxt}
+  \item generic contexts are called @{ML_text context}, rarely
+  @{ML_text ctxt}
 
-  \item local theories are called @{verbatim lthy}, unless there is an
-  implicit conversion to a general proof context (semantic super-type)
+  \item local theories are called @{ML_text lthy}, except for local
+  theories that are treated as proof context (which is a semantic
+  super-type)
 
   \end{itemize}
 
   Variations with primed or decimal numbers are always possible, as
-  well as ``sematic prefixes'' like @{verbatim foo_thy} or @{verbatim
+  well as sematic prefixes like @{ML_text foo_thy} or @{ML_text
   bar_ctxt}, but the base conventions above need to be preserved.
+  This allows to visualize the their data flow via plain regular
+  expressions in the editor.
 
-  \item The main logical entities (\secref{ch:logic}) also have
-  established naming convention:
+  \item The main logical entities (\secref{ch:logic}) have established
+  naming convention as follows:
 
   \begin{itemize}
 
-  \item sorts are called @{verbatim S}
+  \item sorts are called @{ML_text S}
 
-  \item types are called @{verbatim T}, @{verbatim U}, or @{verbatim
-  ty} (never @{verbatim t})
+  \item types are called @{ML_text T}, @{ML_text U}, or @{ML_text
+  ty} (never @{ML_text t})
 
-  \item terms are called @{verbatim t}, @{verbatim u}, or @{verbatim
-  tm} (never @{verbatim trm})
+  \item terms are called @{ML_text t}, @{ML_text u}, or @{ML_text
+  tm} (never @{ML_text trm})
 
-  \item certified types are called @{verbatim cT}, rarely @{verbatim
+  \item certified types are called @{ML_text cT}, rarely @{ML_text
   T}, with variants as for types
 
-  \item certified terms are called @{verbatim ct}, rarely @{verbatim
+  \item certified terms are called @{ML_text ct}, rarely @{ML_text
   t}, with variants as for terms
 
-  \item theorems are called @{verbatim th}, or @{verbatim thm}
+  \item theorems are called @{ML_text th}, or @{ML_text thm}
 
   \end{itemize}
 
   Proper semantic names override these conventions completely.  For
   example, the left-hand side of an equation (as a term) can be called
-  @{verbatim lhs} (not @{verbatim lhs_tm}).  Or a term that is treated
-  specifically as a variable can be called @{verbatim v} or @{verbatim
-  x}.
+  @{ML_text lhs} (not @{ML_text lhs_tm}).  Or a term that is known
+  to be a variable can be called @{ML_text v} or @{ML_text x}.
 
   \end{itemize}
 *}
@@ -275,11 +274,11 @@
   programming.
 
   \paragraph{Line length} is 80 characters according to ancient
-  standards, but we allow as much as 100 characters, not
-  more.\footnote{Readability requires to keep the beginning of a line
+  standards, but we allow as much as 100 characters (not
+  more).\footnote{Readability requires to keep the beginning of a line
   in view while watching its end.  Modern wide-screen displays do not
-  change the way how the human brain works.  As a rule of thumb,
-  sources need to be printable on plain paper.} The extra 20
+  change the way how the human brain works.  Sources also need to be
+  printable on plain paper with reasonable font-size.} The extra 20
   characters acknowledge the space requirements due to qualified
   library references in Isabelle/ML.
 
@@ -310,16 +309,17 @@
     c);
   \end{verbatim}
 
-  Some special infixes (e.g.\ @{verbatim "|>"}) work better at the
+  Some special infixes (e.g.\ @{ML_text "|>"}) work better at the
   start of the line, but punctuation is always at the end.
 
   Function application follows the tradition of @{text "\<lambda>"}-calculus,
-  not informal mathematics.  For example: @{verbatim "f a b"} for a
-  curried function, or @{verbatim "g (a, b)"} for a tupled function.
-  Note that the space between @{verbatim g} and the pair @{verbatim
+  not informal mathematics.  For example: @{ML_text "f a b"} for a
+  curried function, or @{ML_text "g (a, b)"} for a tupled function.
+  Note that the space between @{ML_text g} and the pair @{ML_text
   "(a, b)"} follows the important principle of
-  \emph{compositionality}: the layout of @{verbatim "g p"} does not
-  change when @{verbatim "p"} is refined to a concrete pair.
+  \emph{compositionality}: the layout of @{ML_text "g p"} does not
+  change when @{ML_text "p"} is refined to the concrete pair
+  @{ML_text "(a, b)"}.
 
   \paragraph{Indentation} uses plain spaces, never hard
   tabulators.\footnote{Tabulators were invented to move the carriage
@@ -329,7 +329,7 @@
   editor configuration.}
 
   Each level of nesting is indented by 2 spaces, sometimes 1, very
-  rarely 4, never 8.
+  rarely 4, never 8 or any other odd number.
 
   Indentation follows a simple logical format that only depends on the
   nesting depth, not the accidental length of the text that initiates
@@ -362,18 +362,18 @@
   \emph{maintainability}) etc.
 
   \medskip For similar reasons, any kind of two-dimensional or tabular
-  layouts, ASCII-art with lines or boxes of stars etc. should be
+  layouts, ASCII-art with lines or boxes of asterisks etc.\ should be
   avoided.
 
-  \paragraph{Complex expressions} consisting of multi-clausal function
-  definitions, @{verbatim case}, @{verbatim handle}, @{verbatim let},
-  or combinations require special attention.  The syntax of Standard
-  ML is a bit too ambitious in admitting too much variance that can
-  distort the meaning of the text.
+  \paragraph{Complex expressions} that consist of multi-clausal
+  function definitions, @{ML_text handle}, @{ML_text case},
+  @{ML_text let} (and combinations) require special attention.  The
+  syntax of Standard ML is quite ambitious and admits a lot of
+  variance that can distort the meaning of the text.
 
-  Clauses of @{verbatim fun}, @{verbatim fn}, @{verbatim case},
-  @{verbatim handle} get extra indentation to indicate the nesting
-  clearly.  For example:
+  Clauses of @{ML_text fun}, @{ML_text fn}, @{ML_text handle},
+  @{ML_text case} get extra indentation to indicate the nesting
+  clearly.  Example:
 
   \begin{verbatim}
   (* RIGHT *)
@@ -392,10 +392,10 @@
     expr2
   \end{verbatim}
 
-  Body expressions consisting of @{verbatim case} or @{verbatim let}
+  Body expressions consisting of @{ML_text case} or @{ML_text let}
   require care to maintain compositionality, to prevent loss of
-  logical indentation where it is particularly imporant to see the
-  structure of the text later on.  Example:
+  logical indentation where it is especially important to see the
+  structure of the text.  Example:
 
   \begin{verbatim}
   (* RIGHT *)
@@ -425,16 +425,16 @@
     end
   \end{verbatim}
 
-  Extra parentheses around @{verbatim case} expressions are optional,
-  but help to analyse the nesting easily based on simple string
-  matching in the editor.
+  Extra parentheses around @{ML_text case} expressions are optional,
+  but help to analyse the nesting based on character matching in the
+  editor.
 
   \medskip There are two main exceptions to the overall principle of
   compositionality in the layout of complex expressions.
 
   \begin{enumerate}
 
-  \item @{verbatim "if"} expressions are iterated as if there would be
+  \item @{ML_text "if"} expressions are iterated as if there would be
   a multi-branch conditional in SML, e.g.
 
   \begin{verbatim}
@@ -445,10 +445,10 @@
   else e3
   \end{verbatim}
 
-  \item @{verbatim fn} abstractions are often layed-out as if they
+  \item @{ML_text fn} abstractions are often layed-out as if they
   would lack any structure by themselves.  This traditional form is
   motivated by the possibility to shift function arguments back and
-  forth wrt.\ additional combinators.  For example:
+  forth wrt.\ additional combinators.  Example:
 
   \begin{verbatim}
   (* RIGHT *)
@@ -457,13 +457,13 @@
     expr)
   \end{verbatim}
 
-  Here the visual appearance is that of three arguments @{verbatim x},
-  @{verbatim y}, @{verbatim z}.
+  Here the visual appearance is that of three arguments @{ML_text x},
+  @{ML_text y}, @{ML_text z}.
 
   \end{enumerate}
 
   Such weakly structured layout should be use with great care.  Here
-  is a counter-example involving @{verbatim let} expressions:
+  is a counter-example involving @{ML_text let} expressions:
 
   \begin{verbatim}
   (* WRONG *)
@@ -493,34 +493,35 @@
   (mainly Poly/ML, but also SML/NJ).
 
   Isabelle/Pure marks the point where the original ML toplevel is
-  superseded by the Isar toplevel that maintains a uniform environment
-  for arbitrary ML values (see also \secref{sec:context}).  This
-  formal context holds logical entities as well as ML compiler
-  bindings, among many other things.  Raw Standard ML is never
-  encountered again after the initial bootstrap of Isabelle/Pure.
+  superseded by the Isar toplevel that maintains a uniform context for
+  arbitrary ML values (see also \secref{sec:context}).  This formal
+  environment holds ML compiler bindings, logical entities, and many
+  other things.  Raw SML is never encountered again after the initial
+  bootstrap of Isabelle/Pure.
 
-  Object-logics such as Isabelle/HOL are built within the
-  Isabelle/ML/Isar environment of Pure by introducing suitable
-  theories with associated ML text, either inlined or as separate
-  files.  Thus Isabelle/HOL is defined as a regular user-space
-  application within the Isabelle framework.  Further add-on tools can
-  be implemented in ML within the Isar context in the same manner: ML
-  is part of the regular user-space repertoire of Isabelle.
+  Object-logics like Isabelle/HOL are built within the
+  Isabelle/ML/Isar environment by introducing suitable theories with
+  associated ML modules, either inlined or as separate files.  Thus
+  Isabelle/HOL is defined as a regular user-space application within
+  the Isabelle framework.  Further add-on tools can be implemented in
+  ML within the Isar context in the same manner: ML is part of the
+  standard repertoire of Isabelle, and there is no distinction between
+  ``user'' and ``developer'' in this respect.
 *}
 
 
 subsection {* Isar ML commands *}
 
-text {* The primary Isar source language provides various facilities
-  to open a ``window'' to the underlying ML compiler.  Especially see
-  @{command_ref "use"} and @{command_ref "ML"}, which work exactly the
+text {* The primary Isar source language provides facilities to ``open
+  a window'' to the underlying ML compiler.  Especially see the Isar
+  commands @{command_ref "use"} and @{command_ref "ML"}: both work the
   same way, only the source text is provided via a file vs.\ inlined,
   respectively.  Apart from embedding ML into the main theory
   definition like that, there are many more commands that refer to ML
   source, such as @{command_ref setup} or @{command_ref declaration}.
-  An example of even more fine-grained embedding of ML into Isar is
-  the proof method @{method_ref tactic}, which refines the pending goal state
-  via a given expression of type @{ML_type tactic}.
+  Even more fine-grained embedding of ML into Isar is encountered in
+  the proof method @{method_ref tactic}, which refines the pending
+  goal state via a given expression of type @{ML_type tactic}.
 *}
 
 text %mlex {* The following artificial example demonstrates some ML
@@ -534,7 +535,7 @@
     | factorial n = n * factorial (n - 1)
 *}
 
-text {* Here the ML environment is really managed by Isabelle, i.e.\
+text {* Here the ML environment is already managed by Isabelle, i.e.\
   the @{ML factorial} function is not yet accessible in the preceding
   paragraph, nor in a different theory that is independent from the
   current one in the import hierarchy.
@@ -542,33 +543,32 @@
   Removing the above ML declaration from the source text will remove
   any trace of this definition as expected.  The Isabelle/ML toplevel
   environment is managed in a \emph{stateless} way: unlike the raw ML
-  toplevel or similar command loops of Computer Algebra systems, there
-  are no global side-effects involved here.\footnote{Such a stateless
-  compilation environment is also a prerequisite for robust parallel
-  compilation within independent nodes of the implicit theory
-  development graph.}
+  toplevel there are no global side-effects involved
+  here.\footnote{Such a stateless compilation environment is also a
+  prerequisite for robust parallel compilation within independent
+  nodes of the implicit theory development graph.}
 
-  \medskip The next example shows how to embed ML into Isar proofs.
-  Instead of @{command_ref "ML"} for theory mode, we use @{command_ref
-  "ML_prf"} for proof mode.  As illustrated below, its effect on the
-  ML environment is local to the whole proof body, while ignoring its
-  internal block structure.  *}
+  \medskip The next example shows how to embed ML into Isar proofs, using
+ @{command_ref "ML_prf"} instead of Instead of @{command_ref "ML"}.
+  As illustrated below, the effect on the ML environment is local to
+  the whole proof body, ignoring the block structure.
+*}
 
 example_proof
   ML_prf %"ML" {* val a = 1 *}
-  { -- {* Isar block structure ignored by ML environment *}
+  {
     ML_prf %"ML" {* val b = a + 1 *}
   } -- {* Isar block structure ignored by ML environment *}
   ML_prf %"ML" {* val c = b + 1 *}
 qed
 
 text {* By side-stepping the normal scoping rules for Isar proof
-  blocks, embedded ML code can refer to the different contexts
-  explicitly, and manipulate corresponding entities, e.g.\ export a
-  fact from a block context.
+  blocks, embedded ML code can refer to the different contexts and
+  manipulate corresponding entities, e.g.\ export a fact from a block
+  context.
 
   \medskip Two further ML commands are useful in certain situations:
-  @{command_ref ML_val} and @{command_ref ML_command} are both
+  @{command_ref ML_val} and @{command_ref ML_command} are
   \emph{diagnostic} in the sense that there is no effect on the
   underlying environment, and can thus used anywhere (even outside a
   theory).  The examples below produce long strings of digits by
@@ -591,14 +591,14 @@
   formal context is passed as a thread-local reference variable.  Thus
   ML code may access the theory context during compilation, by reading
   or writing the (local) theory under construction.  Note that such
-  direct access to the compile-time context is rare; in practice it is
-  typically via some derived ML functions.
+  direct access to the compile-time context is rare.  In practice it
+  is typically done via some derived ML functions instead.
 *}
 
 text %mlref {*
   \begin{mldecls}
   @{index_ML ML_Context.the_generic_context: "unit -> Context.generic"} \\
-  @{index_ML "Context.>> ": "(Context.generic -> Context.generic) -> unit"} \\
+  @{index_ML "Context.>>": "(Context.generic -> Context.generic) -> unit"} \\
   @{index_ML bind_thms: "string * thm list -> unit"} \\
   @{index_ML bind_thm: "string * thm -> unit"} \\
   \end{mldecls}
@@ -620,11 +620,11 @@
   put into a global ``standard'' format before being stored.
 
   \item @{ML bind_thm} is similar to @{ML bind_thms} but refers to a
-  singleton theorem.
+  singleton fact.
 
   \end{description}
 
-  It is very important to note that the above functions are really
+  It is important to note that the above functions are really
   restricted to the compile time, even though the ML compiler is
   invoked at run-time.  The majority of ML code either uses static
   antiquotations (\secref{sec:ML-antiq}) or refers to the theory or
@@ -635,9 +635,8 @@
 subsection {* Antiquotations \label{sec:ML-antiq} *}
 
 text {* A very important consequence of embedding SML into Isar is the
-  concept of \emph{ML antiquotation}.  First, the standard token
-  language of ML is augmented by special syntactic entities of the
-  following form:
+  concept of \emph{ML antiquotation}.  The standard token language of
+  ML is augmented by special syntactic entities of the following form:
 
   \indexouternonterm{antiquote}
   \begin{rail}
@@ -646,8 +645,8 @@
   \end{rail}
 
   Here @{syntax nameref} and @{syntax args} are regular outer syntax
-  categories.  Note that attributes and proof methods use similar
-  syntax.
+  categories \cite{isabelle-isar-ref}.  Attributes and proof methods
+  use similar syntax.
 
   \medskip A regular antiquotation @{text "@{name args}"} processes
   its arguments by the usual means of the Isar source language, and
@@ -664,7 +663,7 @@
   effect by introducing local blocks within the pre-compilation
   environment.
 
-  \medskip See also \cite{Wenzel-Chaieb:2007b} for a slightly broader
+  \medskip See also \cite{Wenzel-Chaieb:2007b} for a broader
   perspective on Isabelle/ML antiquotations.  *}
 
 text %mlantiq {*
@@ -698,10 +697,9 @@
   \end{description}
 *}
 
-text %mlex {* The following artificial example gives a first
-  impression about using the antiquotation elements introduced so far,
-  together with the basic @{text "@{thm}"} antiquotation defined
-  later.
+text %mlex {* The following artificial example gives some impression
+  about the antiquotation elements introduced so far, together with
+  the important @{text "@{thm}"} antiquotation defined later.
 *}
 
 ML {*
@@ -712,9 +710,9 @@
   \<rbrace>
 *}
 
-text {*
-  The extra block delimiters do not affect the compiled code itself, i.e.\
-  function @{ML foo} is available in the present context.
+text {* The extra block delimiters do not affect the compiled code
+  itself, i.e.\ function @{ML foo} is available in the present context
+  of this paragraph.
 *}
 
 
@@ -727,19 +725,19 @@
   functions in that setting can save a lot of extra boiler-plate of
   redundant shuffling of arguments, auxiliary abstractions etc.
 
-  First of all, functions are usually \emph{curried}: the idea of
-  @{text "f"} taking @{text "n"} arguments of type @{text "\<tau>\<^sub>i"} (for
-  @{text "i \<in> {1, \<dots> n}"}) with result @{text "\<tau>"} is represented by
-  the iterated function space @{text "\<tau>\<^sub>1 \<rightarrow> \<dots> \<rightarrow> \<tau>\<^sub>n \<rightarrow> \<tau>"}.  This is
-  isomorphic to the encoding via tuples @{text "\<tau>\<^sub>1 \<times> \<dots> \<times> \<tau>\<^sub>n \<rightarrow> \<tau>"}, but
-  the curried version fits more smoothly into the basic
-  calculus.\footnote{The difference is even more significant in
-  higher-order logic, because additional the redundant tuple structure
-  needs to be accommodated by formal reasoning.}
+  Functions are usually \emph{curried}: the idea of turning arguments
+  of type @{text "\<tau>\<^sub>i"} (for @{text "i \<in> {1, \<dots> n}"}) into a result of
+  type @{text "\<tau>"} is represented by the iterated function space
+  @{text "\<tau>\<^sub>1 \<rightarrow> \<dots> \<rightarrow> \<tau>\<^sub>n \<rightarrow> \<tau>"}.  This is isomorphic to the well-known
+  encoding via tuples @{text "\<tau>\<^sub>1 \<times> \<dots> \<times> \<tau>\<^sub>n \<rightarrow> \<tau>"}, but the curried
+  version fits more smoothly into the basic calculus.\footnote{The
+  difference is even more significant in higher-order logic, because
+  the redundant tuple structure needs to be accommodated by formal
+  reasoning.}
 
   Currying gives some flexiblity due to \emph{partial application}.  A
   function @{text "f: \<tau>\<^sub>1 \<rightarrow> \<tau>\<^bsub>2\<^esub> \<rightarrow> \<tau>"} can be applied to @{text "x: \<tau>\<^sub>1"}
-  and the remaining @{text "(f x): \<tau>\<^sub>2 \<rightarrow> \<tau>"} passed to other functions
+  and the remaining @{text "(f x): \<tau>\<^sub>2 \<rightarrow> \<tau>"} passed to another function
   etc.  How well this works in practice depends on the order of
   arguments.  In the worst case, arguments are arranged erratically,
   and using a function in a certain situation always requires some
@@ -747,11 +745,10 @@
   decorate the code with meaningless permutations of arguments.
 
   This can be avoided by \emph{canonical argument order}, which
-  observes certain standard patterns of function definitions and
-  minimizes adhoc permutations in their application.  In Isabelle/ML,
-  large portions of non-trivial source code are written without ever
-  using the infamous function @{text "swap: \<alpha> \<times> \<beta> \<rightarrow> \<beta> \<times> \<alpha>"} or the
-  combinator @{text "C: (\<alpha> \<rightarrow> \<beta> \<rightarrow> \<gamma>) \<rightarrow> (\<beta> \<rightarrow> \<alpha> \<rightarrow> \<gamma>)"}, which is not even
+  observes certain standard patterns and minimizes adhoc permutations
+  in their application.  In Isabelle/ML, large portions text can be
+  written without ever using @{text "swap: \<alpha> \<times> \<beta> \<rightarrow> \<beta> \<times> \<alpha>"}, or the
+  combinator @{text "C: (\<alpha> \<rightarrow> \<beta> \<rightarrow> \<gamma>) \<rightarrow> (\<beta> \<rightarrow> \<alpha> \<rightarrow> \<gamma>)"} that is not even
   defined in our library.
 
   \medskip The basic idea is that arguments that vary less are moved
@@ -779,15 +776,16 @@
   concrete container.  The argument order makes it easy to use other
   combinators: @{text "forall (member B) list"} will check a list of
   elements for membership in @{text "B"} etc. Often the explicit
-  @{text "list"} is pointless and can be contracted in the application
-  @{text "forall (member B)"} to get directly a predicate again.
+  @{text "list"} is pointless and can be contracted to @{text "forall
+  (member B)"} to get directly a predicate again.
 
-  The update operation naturally ``varies'' the container, so it moves
+  In contrast, an update operation varies the container, so it moves
   to the right: @{text "insert a"} is a function @{text "\<beta> \<rightarrow> \<beta>"} to
   insert a value @{text "a"}.  These can be composed naturally as
-  follows: @{text "insert c \<circ> insert b \<circ> insert a"}.  This works well,
-  apart from some awkwardness due to conventional mathematical
-  function composition, which can be easily amended as follows.
+  @{text "insert c \<circ> insert b \<circ> insert a"}.  The slightly awkward
+  inversion if the composition order is due to conventional
+  mathematical notation, which can be easily amended as explained
+  below.
 *}
 
 
@@ -795,10 +793,10 @@
 
 text {* Regular function application and infix notation works best for
   relatively deeply structured expressions, e.g.\ @{text "h (f x y + g
-  z)"}.  The important special case if \emph{linear transformation}
-  applies a cascade of functions as follows: @{text "f\<^sub>n (\<dots> (f\<^sub>1 x))"}.
-  This becomes hard to read and maintain if the functions are
-  themselves complex expressions.  The notation can be significantly
+  z)"}.  The important special case of \emph{linear transformation}
+  applies a cascade of functions @{text "f\<^sub>n (\<dots> (f\<^sub>1 x))"}.  This
+  becomes hard to read and maintain if the functions are themselves
+  given as complex expressions.  The notation can be significantly
   improved by introducing \emph{forward} versions of application and
   composition as follows:
 
@@ -840,7 +838,7 @@
 subsection {* Canonical iteration *}
 
 text {* As explained above, a function @{text "f: \<alpha> \<rightarrow> \<beta> \<rightarrow> \<beta>"} can be
-  understood as update on a configuration of type @{text "\<beta>"} that is
+  understood as update on a configuration of type @{text "\<beta>"},
   parametrized by arguments of type @{text "\<alpha>"}.  Given @{text "a: \<alpha>"}
   the partial application @{text "(f a): \<beta> \<rightarrow> \<beta>"} operates
   homogeneously on @{text "\<beta>"}.  This can be iterated naturally over a
@@ -890,10 +888,10 @@
   The literature on functional programming provides a multitude of
   combinators called @{text "foldl"}, @{text "foldr"} etc.  SML97
   provides its own variations as @{ML List.foldl} and @{ML
-  List.foldr}, while the classic Isabelle library still has the
-  slightly more convenient historic @{ML Library.foldl} and @{ML
-  Library.foldr}.  To avoid further confusion, all of this should be
-  ignored, and @{ML fold} (or @{ML fold_rev}) used exclusively.
+  List.foldr}, while the classic Isabelle library also has the
+  historic @{ML Library.foldl} and @{ML Library.foldr}.  To avoid
+  further confusion, all of this should be ignored, and @{ML fold} (or
+  @{ML fold_rev}) used exclusively.
   \end{warn}
 *}
 
@@ -918,8 +916,9 @@
   memory (``deforestation''), but requires some practice to read and
   write it fluently.
 
-  \medskip The next example elaborates this idea and demonstrates fast
-  accumulation of tree content using a text buffer.
+  \medskip The next example elaborates the idea of canonical
+  iteration, demonstrating fast accumulation of tree content using a
+  text buffer.
 *}
 
 ML {*
@@ -943,31 +942,28 @@
 
 text {* The slow part of @{ML slow_content} is the @{ML implode} of
   the recursive results, because it copies previously produced strings
-  afresh.
+  again.
 
   The incremental @{ML add_content} avoids this by operating on a
-  buffer that is passed-through in a linear fashion.  Using @{verbatim
-  "#>"} and contraction over the actual @{ML_type "Buffer.T"} argument
-  saves some additional boiler-plate, but requires again some
-  practice.  Of course, the two @{ML "Buffer.add"} invocations with
-  concatenated strings could have been split into smaller parts, but
-  this would have obfuscated the source without making a big
-  difference in allocations.  Here we have done some
+  buffer that is passed through in a linear fashion.  Using @{ML_text
+  "#>"} and contraction over the actual buffer argument saves some
+  additional boiler-plate.  Of course, the two @{ML "Buffer.add"}
+  invocations with concatenated strings could have been split into
+  smaller parts, but this would have obfuscated the source without
+  making a big difference in allocations.  Here we have done some
   peephole-optimization for the sake of readability.
 
   Another benefit of @{ML add_content} is its ``open'' form as a
-  function on @{ML_type Buffer.T} that can be continued in further
-  linear transformations, folding etc.  Thus it is more compositional
-  than the naive @{ML slow_content}.  As a realistic example, compare
-  the slightly old-style @{ML "Term.maxidx_of_term: term -> int"} with
-  the newer @{ML "Term.maxidx_term: term -> int -> int"} in
-  Isabelle/Pure.
+  function on buffers that can be continued in further linear
+  transformations, folding etc.  Thus it is more compositional than
+  the naive @{ML slow_content}.  As realistic example, compare the
+  old-style @{ML "Term.maxidx_of_term: term -> int"} with the newer
+  @{ML "Term.maxidx_term: term -> int -> int"} in Isabelle/Pure.
 
-  Note that @{ML fast_content} above is only defined for completeness
-  of the example.  In many practical situations, it is customary to
-  defined the incremental @{ML add_content} only and leave the
-  initialization and termination to the concrete application by the
-  user.
+  Note that @{ML fast_content} above is only defined as example.  In
+  many practical situations, it is customary to provide the
+  incremental @{ML add_content} only and leave the initialization and
+  termination to the concrete application by the user.
 *}
 
 
@@ -1016,12 +1012,12 @@
 
   \item @{ML warning}~@{text "text"} outputs @{text "text"} as
   warning, which typically means some extra emphasis on the front-end
-  side (color highlighting, icon).
+  side (color highlighting, icons, etc.).
 
   \item @{ML error}~@{text "text"} raises exception @{ML ERROR}~@{text
   "text"} and thus lets the Isar toplevel print @{text "text"} on the
   error channel, which typically means some extra emphasis on the
-  front-end side (color highlighting, icon).
+  front-end side (color highlighting, icons, etc.).
 
   This assumes that the exception is not handled before the command
   terminates.  Handling exception @{ML ERROR}~@{text "text"} is a
@@ -1032,7 +1028,7 @@
   The actual error channel is accessed via @{ML Output.error_msg}, but
   the interaction protocol of Proof~General \emph{crashes} if that
   function is used in regular ML code: error output and toplevel
-  command failure always need to coincide here.
+  command failure always need to coincide.
   \end{warn}
 
   \end{description}
@@ -1046,14 +1042,14 @@
   some system console log, with a low chance that the user will ever
   see it.  Moreover, as a genuine side-effect on global process
   channels, there is no proper way to retract output when Isar command
-  transactions are reset.
+  transactions are reset by the system.
   \end{warn}
 
   \begin{warn}
   The message channels should be used in a message-oriented manner.
-  This means that multi-line output that logically belongs together
-  needs to be issued by a \emph{single} invocation of @{ML writeln}
-  etc.  with the functional concatenation of all message constituents.
+  This means that multi-line output that logically belongs together is
+  issued by a \emph{single} invocation of @{ML writeln} etc.\ with the
+  functional concatenation of all message constituents.
   \end{warn}
 *}
 
@@ -1108,16 +1104,16 @@
   These exceptions indicate a genuine breakdown of the program, so the
   main purpose is to determine quickly what has happened where.
   Traditionally, the (short) exception message would include the name
-  of an ML function, although this no longer necessary, because the ML
-  runtime system prints a detailed source position of the
-  corresponding @{verbatim raise} keyword.
+  of an ML function, although this is no longer necessary, because the
+  ML runtime system prints a detailed source position of the
+  corresponding @{ML_text raise} keyword.
 
   \medskip User modules can always introduce their own custom
   exceptions locally, e.g.\ to organize internal failures robustly
   without overlapping with existing exceptions.  Exceptions that are
   exposed in module signatures require extra care, though, and should
-  \emph{not} be introduced by default.  Surprise by end-users or ML
-  users of a module can be often minimized by using plain user errors.
+  \emph{not} be introduced by default.  Surprise by users of a module
+  can be often minimized by using plain user errors instead.
 
   \paragraph{Interrupts.}  These indicate arbitrary system events:
   both the ML runtime system and the Isabelle/ML infrastructure signal
@@ -1131,7 +1127,7 @@
   that intercepts interrupts becomes dependent on physical effects of
   the environment.  Even worse, exception handling patterns that are
   too general by accident, e.g.\ by mispelled exception constructors,
-  will cover interrupts unintentionally, and thus render the program
+  will cover interrupts unintentionally and thus render the program
   semantics ill-defined.
 
   Note that the Interrupt exception dates back to the original SML90
@@ -1139,8 +1135,8 @@
   avoid its malign impact on ML program semantics, but without
   providing a viable alternative.  Isabelle/ML recovers physical
   interruptibility (which an indispensable tool to implement managed
-  evaluation of Isar command transactions), but requires user code to
-  be strictly transparent wrt.\ interrupts.
+  evaluation of command transactions), but requires user code to be
+  strictly transparent wrt.\ interrupts.
 
   \begin{warn}
   Isabelle/ML user code needs to terminate promptly on interruption,
@@ -1167,15 +1163,15 @@
   \item @{ML try}~@{text "f x"} makes the partiality of evaluating
   @{text "f x"} explicit via the option datatype.  Interrupts are
   \emph{not} handled here, i.e.\ this form serves as safe replacement
-  for the \emph{unsafe} version @{verbatim "(SOME"}~@{text "f
-  x"}~@{verbatim "handle _ => NONE)"} that is occasionally seen in
+  for the \emph{unsafe} version @{ML_text "(SOME"}~@{text "f
+  x"}~@{ML_text "handle _ => NONE)"} that is occasionally seen in
   books about SML.
 
   \item @{ML can} is similar to @{ML try} with more abstract result.
 
   \item @{ML ERROR}~@{text "msg"} represents user errors; this
-  exception is always raised via the @{ML error} function (see
-  \secref{sec:message-channels}).
+  exception is normally raised indirectly via the @{ML error} function
+  (see \secref{sec:message-channels}).
 
   \item @{ML Fail}~@{text "msg"} represents general program failures.
 
@@ -1187,12 +1183,11 @@
   while preserving its implicit position information (if possible,
   depending on the ML platform).
 
-  \item @{ML exception_trace}~@{verbatim "(fn _ =>"}~@{text
-  "e"}@{verbatim ")"} evaluates expression @{text "e"} while printing
+  \item @{ML exception_trace}~@{ML_text "(fn () =>"}~@{text
+  "e"}@{ML_text ")"} evaluates expression @{text "e"} while printing
   a full trace of its stack of nested exceptions (if possible,
-  depending on the ML platform).\footnote{In various versions of
-  Poly/ML the trace will appear on raw stdout of the Isabelle
-  process.}
+  depending on the ML platform).\footnote{In versions of Poly/ML the
+  trace will appear on raw stdout of the Isabelle process.}
 
   Inserting @{ML exception_trace} into ML code temporarily is useful
   for debugging, but not suitable for production code.
@@ -1207,10 +1202,10 @@
 
   \begin{description}
 
-  \item @{text "@{assert}"} inlines a function @{text "bool \<Rightarrow> unit"}
-  that raises @{ML Fail} if the argument is @{ML false}.  Due to
-  inlining the source position of failed assertions is included in the
-  error output.
+  \item @{text "@{assert}"} inlines a function
+  @{ML_type "bool -> unit"} that raises @{ML Fail} if the argument is
+  @{ML false}.  Due to inlining the source position of failed
+  assertions is included in the error output.
 
   \end{description}
 *}
@@ -1218,12 +1213,12 @@
 
 section {* Basic data types *}
 
-text {* The basis library proposal of SML97 need to be treated with
+text {* The basis library proposal of SML97 needs to be treated with
   caution.  Many of its operations simply do not fit with important
   Isabelle/ML conventions (like ``canonical argument order'', see
-  \secref{sec:canonical-argument-order}), others can cause serious
-  problems with the parallel evaluation model of Isabelle/ML (such as
-  @{ML TextIO.print} or @{ML OS.Process.system}).
+  \secref{sec:canonical-argument-order}), others cause problems with
+  the parallel evaluation model of Isabelle/ML (such as @{ML
+  TextIO.print} or @{ML OS.Process.system}).
 
   Subsequently we give a brief overview of important operations on
   basic ML data types.
@@ -1240,7 +1235,7 @@
   \begin{description}
 
   \item Type @{ML_type char} is \emph{not} used.  The smallest textual
-  unit in Isabelle is represented a ``symbol'' (see
+  unit in Isabelle is represented as a ``symbol'' (see
   \secref{sec:symbols}).
 
   \end{description}
@@ -1263,11 +1258,10 @@
   This works uniformly for all supported ML platforms (Poly/ML and
   SML/NJ).
 
-  Literal integers in ML text (e.g.\ @{ML
-  123456789012345678901234567890}) are forced to be of this one true
+  Literal integers in ML text are forced to be of this one true
   integer type --- overloading of SML97 is disabled.
 
-  Structure @{ML_struct IntInf} of SML97 is obsolete, always use
+  Structure @{ML_struct IntInf} of SML97 is obsolete and superseded by
   @{ML_struct Int}.  Structure @{ML_struct Integer} in @{"file"
   "~~/src/Pure/General/integer.ML"} provides some additional
   operations.
@@ -1328,7 +1322,7 @@
 
   Note that @{ML insert} is conservative about elements that are
   already a @{ML member} of the list, while @{ML update} ensures that
-  the last entry is always put in front.  The latter discipline is
+  the latest entry is always put in front.  The latter discipline is
   often more appropriate in declarations of context data
   (\secref{sec:context-data}) that are issued by the user in Isar
   source: more recent declarations normally take precedence over
@@ -1337,12 +1331,12 @@
   \end{description}
 *}
 
-text %mlex {* Using canonical @{ML fold} together with canonical @{ML
-  cons}, or similar standard operations, alternates the orientation of
-  data.  The is quite natural and should not altered forcible by
-  inserting extra applications @{ML rev}.  The alternative @{ML
-  fold_rev} can be used in the relatively few situations, where
-  alternation should be prevented.
+text %mlex {* Using canonical @{ML fold} together with @{ML cons}, or
+  similar standard operations, alternates the orientation of data.
+  The is quite natural and should not be altered forcible by inserting
+  extra applications of @{ML rev}.  The alternative @{ML fold_rev} can
+  be used in the few situations, where alternation should be
+  prevented.
 *}
 
 ML {*
@@ -1393,7 +1387,8 @@
 
   \item @{ML AList.lookup}, @{ML AList.defined}, @{ML AList.update}
   implement the main ``framework operations'' for mappings in
-  Isabelle/ML, with standard conventions for names and types.
+  Isabelle/ML, following standard conventions for their names and
+  types.
 
   Note that a function called @{text lookup} is obliged to express its
   partiality via an explicit option element.  There is no choice to
@@ -1473,7 +1468,7 @@
 
   Isabelle/Isar exploits the inherent structure of theories and proofs
   to support \emph{implicit parallelism} to a large extent.  LCF-style
-  theorem provides almost ideal conditions for that; see also
+  theorem provides almost ideal conditions for that, see also
   \cite{Wenzel:2009}.  This means, significant parts of theory and
   proof checking is parallelized by default.  A maximum speedup-factor
   of 3.0 on 4 cores and 5.0 on 8 cores can be
@@ -1486,11 +1481,10 @@
 
   \medskip ML threads lack the memory protection of separate
   processes, and operate concurrently on shared heap memory.  This has
-  the advantage that results of independent computations are
-  immediately available to other threads: abstract values can be
-  passed between threads without copying or awkward serialization that
-  is typically required for explicit message passing between separate
-  processes.
+  the advantage that results of independent computations are directly
+  available to other threads: abstract values can be passed without
+  copying or awkward serialization that is typically required for
+  separate processes.
 
   To make shared-memory multi-threading work robustly and efficiently,
   some programming guidelines need to be observed.  While the ML
@@ -1535,15 +1529,15 @@
   channels, environment variables, current working directory.
 
   \item Writable resources in the file-system that are shared among
-  different threads or other processes.
+  different threads or external processes.
 
   \end{itemize}
 
   Isabelle/ML provides various mechanisms to avoid critical shared
-  resources in most practical situations.  As last resort there are
-  some mechanisms for explicit synchronization.  The following
-  guidelines help to make Isabelle/ML programs work smoothly in a
-  concurrent environment.
+  resources in most situations.  As last resort there are some
+  mechanisms for explicit synchronization.  The following guidelines
+  help to make Isabelle/ML programs work smoothly in a concurrent
+  environment.
 
   \begin{itemize}
 
@@ -1553,7 +1547,7 @@
   plain value and user tools can get/map their own data in a purely
   functional manner.  Configuration options within the context
   (\secref{sec:config-options}) provide simple drop-in replacements
-  for formerly writable flags.
+  for historic reference variables.
 
   \item Keep components with local state information re-entrant.
   Instead of poking initial values into (private) global references, a
@@ -1669,8 +1663,8 @@
 
   Entering the critical section without contention is very fast, and
   several basic system operations do so frequently.  Each thread
-  should leave the critical section quickly, otherwise parallel
-  performance may degrade.
+  should stay within the critical section quickly only very briefly,
+  otherwise parallel performance may degrade.
 
   \item @{ML CRITICAL} is the same as @{ML NAMED_CRITICAL} with empty
   name argument.
@@ -1684,26 +1678,22 @@
 
   \item @{ML Synchronized.guarded_access}~@{text "var f"} lets the
   function @{text "f"} operate within a critical section on the state
-  @{text "x"} as follows: if @{text "f x"} produces @{ML NONE}, we
-  continue to wait on the internal condition variable, expecting that
+  @{text "x"} as follows: if @{text "f x"} produces @{ML NONE}, it
+  continues to wait on the internal condition variable, expecting that
   some other thread will eventually change the content in a suitable
-  manner; if @{text "f x"} produces @{ML SOME}~@{text "(y, x')"} we
-  are satisfied and assign the new state value @{text "x'"}, broadcast
-  a signal to all waiting threads on the associated condition
-  variable, and return the result @{text "y"}.
+  manner; if @{text "f x"} produces @{ML SOME}~@{text "(y, x')"} it is
+  satisfied and assigns the new state value @{text "x'"}, broadcasts a
+  signal to all waiting threads on the associated condition variable,
+  and returns the result @{text "y"}.
 
   \end{description}
 
-  There are some further variants of the general @{ML
+  There are some further variants of the @{ML
   Synchronized.guarded_access} combinator, see @{"file"
   "~~/src/Pure/Concurrent/synchronized.ML"} for details.
 *}
 
-text %mlex {* See @{"file" "~~/src/Pure/Concurrent/mailbox.ML"} how to
-  implement a mailbox as synchronized variable over a purely
-  functional queue.
-
-  \medskip The following example implements a counter that produces
+text %mlex {* The following example implements a counter that produces
   positive integers that are unique over the runtime of the Isabelle
   process:
 *}
@@ -1726,4 +1716,8 @@
   @{assert} (a <> b);
 *}
 
+text {* \medskip See @{"file" "~~/src/Pure/Concurrent/mailbox.ML"} how
+  to implement a mailbox as synchronized variable over a purely
+  functional queue. *}
+
 end
\ No newline at end of file
--- a/doc-src/IsarImplementation/Thy/Prelim.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/Prelim.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -640,7 +640,7 @@
 setup my_flag_setup
 
 text {* Now the user can refer to @{attribute my_flag} in
-  declarations, while we can retrieve the current value from the
+  declarations, while ML tools can retrieve the current value from the
   context via @{ML Config.get}.  *}
 
 ML_val {* @{assert} (Config.get @{context} my_flag = false) *}
@@ -897,7 +897,7 @@
   @{assert} (list2 = ["x", "xa", "a", "aa", "'a", "'aa"]);
 *}
 
-text {* \medskip The same works reletively to the formal context as
+text {* \medskip The same works relatively to the formal context as
   follows. *}
 
 locale ex = fixes a b c :: 'a
@@ -1210,13 +1210,13 @@
 *}
 
 text %mlex {* The following example yields the source position of some
-  concrete binding inlined into the text.
+  concrete binding inlined into the text:
 *}
 
 ML {* Binding.pos_of @{binding here} *}
 
 text {* \medskip That position can be also printed in a message as
-  follows. *}
+  follows: *}
 
 ML_command {*
   writeln
@@ -1225,6 +1225,7 @@
 
 text {* This illustrates a key virtue of formalized bindings as
   opposed to raw specifications of base names: the system can use this
-  additional information for advanced feedback given to the user. *}
+  additional information for feedback given to the user (error
+  messages etc.). *}
 
 end
--- a/doc-src/IsarImplementation/Thy/Proof.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/Proof.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -180,13 +180,12 @@
   val t2 = Syntax.read_term ctxt2 "x";  (*x::nat is enforced*)
 *}
 
-text {* In the above example, the starting context had been derived
-  from the toplevel theory, which means that fixed variables are
-  internalized literally: @{verbatim "x"} is mapped again to
-  @{verbatim "x"}, and attempting to fix it again in the subsequent
-  context is an error.  Alternatively, fixed parameters can be renamed
-  explicitly as follows:
-*}
+text {* In the above example, the starting context is derived from the
+  toplevel theory, which means that fixed variables are internalized
+  literally: @{verbatim "x"} is mapped again to @{verbatim "x"}, and
+  attempting to fix it again in the subsequent context is an error.
+  Alternatively, fixed parameters can be renamed explicitly as
+  follows: *}
 
 ML {*
   val ctxt0 = @{context};
--- a/doc-src/IsarImplementation/Thy/document/Base.tex	Wed Oct 27 11:10:36 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/document/Base.tex	Wed Oct 27 11:11:35 2010 -0700
@@ -11,41 +11,8 @@
 \ Base\isanewline
 \isakeyword{imports}\ Main\isanewline
 \isakeyword{uses}\ {\isachardoublequoteopen}{\isachardot}{\isachardot}{\isacharslash}{\isachardot}{\isachardot}{\isacharslash}antiquote{\isacharunderscore}setup{\isachardot}ML{\isachardoublequoteclose}\isanewline
-\isakeyword{begin}%
-\endisatagtheory
-{\isafoldtheory}%
-%
-\isadelimtheory
-\isanewline
-%
-\endisadelimtheory
-\isanewline
-%
-\isadelimML
+\isakeyword{begin}\isanewline
 \isanewline
-%
-\endisadelimML
-%
-\isatagML
-\isacommand{ML}\isamarkupfalse%
-\ {\isacharverbatimopen}\isanewline
-\ \ ML{\isacharunderscore}Antiquote{\isachardot}inline\ {\isachardoublequote}assert{\isachardoublequote}\isanewline
-\ \ \ \ {\isacharparenleft}Scan{\isachardot}succeed\ {\isachardoublequote}{\isacharparenleft}fn\ b\ {\isacharequal}{\isachargreater}\ if\ b\ then\ {\isacharparenleft}{\isacharparenright}\ else\ raise\ General{\isachardot}Fail\ {\isacharbackslash}{\isachardoublequote}Assertion\ failed{\isacharbackslash}{\isachardoublequote}{\isacharparenright}{\isachardoublequote}{\isacharparenright}\isanewline
-{\isacharverbatimclose}%
-\endisatagML
-{\isafoldML}%
-%
-\isadelimML
-\isanewline
-%
-\endisadelimML
-%
-\isadelimtheory
-\isanewline
-%
-\endisadelimtheory
-%
-\isatagtheory
 \isacommand{end}\isamarkupfalse%
 %
 \endisatagtheory
--- a/doc-src/IsarImplementation/Thy/document/Isar.tex	Wed Oct 27 11:10:36 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/document/Isar.tex	Wed Oct 27 11:11:35 2010 -0700
@@ -23,8 +23,9 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-The Isar proof language (see also \cite[\S2]{isabelle-isar-ref})
-  consists of three main categories of language elements:
+The Isar proof language (see also
+  \cite[\S2]{isabelle-isar-ref}) consists of three main categories of
+  language elements as follows.
 
   \begin{enumerate}
 
@@ -63,12 +64,12 @@
 A \emph{proof command} is state transition of the Isar/VM
   proof interpreter.
 
-  In principle, Isar proof commands could be defined in
-  user-space as well.  The system is built like that in the first
-  place: part of the commands are primitive, the other part is defined
-  as derived elements.  Adding to the genuine structured proof
-  language requires profound understanding of the Isar/VM machinery,
-  though, so this is beyond the scope of this manual.
+  In principle, Isar proof commands could be defined in user-space as
+  well.  The system is built like that in the first place: one part of
+  the commands are primitive, the other part is defined as derived
+  elements.  Adding to the genuine structured proof language requires
+  profound understanding of the Isar/VM machinery, though, so this is
+  beyond the scope of this manual.
 
   What can be done realistically is to define some diagnostic commands
   that inspect the general state of the Isar/VM, and report some
@@ -77,8 +78,8 @@
   goals (if available).
 
   Another common application is to define a toplevel command that
-  poses a problem to the user as Isar proof state and stores the final
-  result in a suitable context data slot.  Thus a proof can be
+  poses a problem to the user as Isar proof state and processes the
+  final result relatively to the context.  Thus a proof can be
   incorporated into the context of some user-space tool, without
   modifying the Isar proof language itself.%
 \end{isamarkuptext}%
@@ -154,13 +155,14 @@
   storing named facts).  Note that at this generic level the target
   context is specified as \verb|Proof.context|, but the usual
   wrapping of toplevel proofs into command transactions will provide a
-  \verb|local_theory| here (see also \chref{ch:local-theory}).
-  This usually affects the way how results are stored.
+  \verb|local_theory| here (\chref{ch:local-theory}).  This
+  affects the way how results are stored.
 
   The \isa{statement} is given as a nested list of terms, each
   associated with optional \hyperlink{keyword.is}{\mbox{\isa{\isakeyword{is}}}} patterns as usual in the
-  Isar source language.  The original list structure over terms is
-  turned into one over theorems when \isa{after{\isacharunderscore}qed} is invoked.
+  Isar source language.  The original nested list structure over terms
+  is turned into one over theorems when \isa{after{\isacharunderscore}qed} is
+  invoked.
 
   \end{description}%
 \end{isamarkuptext}%
@@ -324,15 +326,15 @@
 
   \end{itemize}
 
-  \medskip Syntactically, the language of proof methods is embedded
-  into Isar as arguments to certain commands, e.g.\ \hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}} or
-  \hyperlink{command.apply}{\mbox{\isa{\isacommand{apply}}}}.  User-space additions are reasonably easy by
-  plugging suitable method-valued parser functions into the framework,
-  using the \hyperlink{command.method-setup}{\mbox{\isa{\isacommand{method{\isacharunderscore}setup}}}} command, for example.
+  \medskip Syntactically, the language of proof methods appears as
+  arguments to Isar commands like \hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}} or \hyperlink{command.apply}{\mbox{\isa{\isacommand{apply}}}}.
+  User-space additions are reasonably easy by plugging suitable
+  method-valued parser functions into the framework, using the
+  \hyperlink{command.method-setup}{\mbox{\isa{\isacommand{method{\isacharunderscore}setup}}}} command, for example.
 
   To get a better idea about the range of possibilities, consider the
-  following Isar proof schemes.  First some general form of structured
-  proof text:
+  following Isar proof schemes.  This is the general form of
+  structured proof text:
 
   \medskip
   \begin{tabular}{l}
@@ -343,13 +345,13 @@
   \end{tabular}
   \medskip
 
-  The goal configuration consists of \isa{facts\isactrlsub {\isadigit{1}}} and \isa{facts\isactrlsub {\isadigit{2}}} appended in that order, and various \isa{props} being
-  claimed here.  The \isa{initial{\isacharunderscore}method} is invoked with facts
-  and goals together and refines the problem to something that is
-  handled recursively in the proof \isa{body}.  The \isa{terminal{\isacharunderscore}method} has another chance to finish-off any remaining
+  The goal configuration consists of \isa{facts\isactrlsub {\isadigit{1}}} and
+  \isa{facts\isactrlsub {\isadigit{2}}} appended in that order, and various \isa{props} being claimed.  The \isa{initial{\isacharunderscore}method} is invoked
+  with facts and goals together and refines the problem to something
+  that is handled recursively in the proof \isa{body}.  The \isa{terminal{\isacharunderscore}method} has another chance to finish any remaining
   subgoals, but it does not see the facts of the initial step.
 
-  \medskip The second pattern illustrates unstructured proof scripts:
+  \medskip This pattern illustrates unstructured proof scripts:
 
   \medskip
   \begin{tabular}{l}
@@ -361,14 +363,12 @@
   \end{tabular}
   \medskip
 
-  The \isa{method\isactrlsub {\isadigit{1}}} operates on the original claim together while
+  The \isa{method\isactrlsub {\isadigit{1}}} operates on the original claim while
   using \isa{facts\isactrlsub {\isadigit{1}}}.  Since the \hyperlink{command.apply}{\mbox{\isa{\isacommand{apply}}}} command
-  structurally resets the facts, the \isa{method\isactrlsub {\isadigit{2}}} will operate on
-  the remaining goal state without facts.  The \isa{method\isactrlsub {\isadigit{3}}} will
-  see again a collection of \isa{facts\isactrlsub {\isadigit{3}}} that has been inserted
-  into the script explicitly.
+  structurally resets the facts, the \isa{method\isactrlsub {\isadigit{2}}} will
+  operate on the remaining goal state without facts.  The \isa{method\isactrlsub {\isadigit{3}}} will see again a collection of \isa{facts\isactrlsub {\isadigit{3}}} that has been inserted into the script explicitly.
 
-  \medskip Empirically, Isar proof methods can be categorized as
+  \medskip Empirically, any Isar proof method can be categorized as
   follows.
 
   \begin{enumerate}
@@ -378,7 +378,8 @@
 
   Example: \hyperlink{method.induct}{\mbox{\isa{induct}}}, which is also a ``raw'' method since it
   operates on the internal representation of simultaneous claims as
-  Pure conjunction, instead of separate subgoals.
+  Pure conjunction (\secref{{sec:logic-aux}}), instead of separate
+  subgoals (\secref{sec::tactical-goals}).
 
   \item \emph{Structured method} with strong emphasis on facts outside
   the goal state.
@@ -395,8 +396,8 @@
   \item \emph{Old-style tactic emulation} with detailed numeric goal
   addressing and explicit references to entities of the internal goal
   state (which are otherwise invisible from proper Isar proof text).
-  To make the special non-standard status clear, the naming convention
-  \isa{foo{\isacharunderscore}tac} needs to be observed.
+  The naming convention \isa{foo{\isacharunderscore}tac} makes this special
+  non-standard status clear.
 
   Example: \hyperlink{method.rule-tac}{\mbox{\isa{rule{\isacharunderscore}tac}}}.
 
@@ -610,9 +611,9 @@
 
   The \verb|Attrib.thms| parser produces a list of theorems from the
   usual Isar syntax involving attribute expressions etc.\ (syntax
-  category \hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}).  The resulting \verb|thms| are
-  added to \verb|HOL_basic_ss| which already contains the basic
-  Simplifier setup for HOL.
+  category \hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}) \cite{isabelle-isar-ref}.  The resulting
+  \verb|thms| are added to \verb|HOL_basic_ss| which already
+  contains the basic Simplifier setup for HOL.
 
   The tactic \verb|asm_full_simp_tac| is the one that is also used in
   method \hyperlink{method.simp}{\mbox{\isa{simp}}} by default.  The extra wrapping by the \verb|CHANGED| tactical ensures progress of simplification: identical goal
@@ -795,8 +796,8 @@
   ``simple'' methods, i.e.\ the goal facts are merely inserted as goal
   premises by the \verb|SIMPLE_METHOD'| or \verb|SIMPLE_METHOD| wrapper.
   For proof methods that are similar to the standard collection of
-  \hyperlink{method.simp}{\mbox{\isa{simp}}}, \hyperlink{method.blast}{\mbox{\isa{blast}}}, \hyperlink{method.auto}{\mbox{\isa{auto}}} little more can be
-  done here.
+  \hyperlink{method.simp}{\mbox{\isa{simp}}}, \hyperlink{method.blast}{\mbox{\isa{blast}}}, \hyperlink{method.fast}{\mbox{\isa{fast}}}, \hyperlink{method.auto}{\mbox{\isa{auto}}}
+  there is little more that can be done.
 
   Note that using the primary goal facts in the same manner as the
   method arguments obtained via concrete syntax or the context does
@@ -819,8 +820,8 @@
   manner, such as a discrimination net that is indexed by the
   left-hand sides of rewrite rules.  For variations on the Simplifier,
   re-use of the existing type \verb|simpset| is adequate, but
-  scalability requires it be maintained statically within the context
-  data, not dynamically on each tool invocation.%
+  scalability would require it be maintained statically within the
+  context data, not dynamically on each tool invocation.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
--- a/doc-src/IsarImplementation/Thy/document/Logic.tex	Wed Oct 27 11:10:36 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/document/Logic.tex	Wed Oct 27 11:11:35 2010 -0700
@@ -829,7 +829,7 @@
 %
 \endisadelimmlantiq
 %
-\isamarkupsubsection{Auxiliary definitions%
+\isamarkupsubsection{Auxiliary definitions \label{sec:logic-aux}%
 }
 \isamarkuptrue%
 %
--- a/doc-src/IsarImplementation/Thy/document/ML.tex	Wed Oct 27 11:10:36 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/document/ML.tex	Wed Oct 27 11:11:35 2010 -0700
@@ -29,12 +29,12 @@
   environment.  This covers a variety of aspects that are geared
   towards an efficient and robust platform for applications of formal
   logic with fully foundational proof construction --- according to
-  the well-known \emph{LCF principle}.  There are specific library
-  modules and infrastructure to address the needs for such difficult
-  tasks.  For example, the raw parallel programming model of Poly/ML
-  is presented as considerably more abstract concept of \emph{future
-  values}, which is then used to augment the inference kernel, proof
-  interpreter, and theory loader accordingly.
+  the well-known \emph{LCF principle}.  There is specific
+  infrastructure with library modules to address the needs of this
+  difficult task.  For example, the raw parallel programming model of
+  Poly/ML is presented as considerably more abstract concept of
+  \emph{future values}, which is then used to augment the inference
+  kernel, proof interpreter, and theory loader accordingly.
 
   The main aspects of Isabelle/ML are introduced below.  These
   first-hand explanations should help to understand how proper
@@ -70,14 +70,14 @@
   standardization.  A coding style that is changed every few years or
   with every new contributor is no style at all, because consistency
   is quickly lost.  Global consistency is hard to achieve, though.
-  One should always strife at least for local consistency of modules
-  and sub-systems, without deviating from some general principles how
-  to write Isabelle/ML.
+  Nonetheless, one should always strive at least for local consistency
+  of modules and sub-systems, without deviating from some general
+  principles how to write Isabelle/ML.
 
-  In a sense, a common coding style is like an \emph{orthography} for
-  the sources: it helps to read quickly over the text and see through
-  the main points, without getting distracted by accidental
-  presentation of free-style source.%
+  In a sense, good coding style is like an \emph{orthography} for the
+  sources: it helps to read quickly over the text and see through the
+  main points, without getting distracted by accidental presentation
+  of free-style code.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -89,15 +89,15 @@
 Isabelle source files have a certain standardized header
   format (with precise spacing) that follows ancient traditions
   reaching back to the earliest versions of the system by Larry
-  Paulson.  E.g.\ see \hyperlink{file.~~/src/Pure/thm.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}thm{\isachardot}ML}}}}.
+  Paulson.  See \hyperlink{file.~~/src/Pure/thm.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}thm{\isachardot}ML}}}}, for example.
 
   The header includes at least \verb|Title| and \verb|Author| entries, followed by a prose description of the purpose of
   the module.  The latter can range from a single line to several
   paragraphs of explanations.
 
   The rest of the file is divided into sections, subsections,
-  subsubsections, paragraphs etc.\ using some a layout via ML comments
-  as follows.
+  subsubsections, paragraphs etc.\ using a simple layout via ML
+  comments as follows.
 
 \begin{verbatim}
 (*** section ***)
@@ -109,8 +109,8 @@
 (*short paragraph*)
 
 (*
-  long paragraph
-  more text
+  long paragraph,
+  with more text
 *)
 \end{verbatim}
 
@@ -119,7 +119,7 @@
   as in the example above).
 
   \medskip The precise wording of the prose text given in these
-  headings is chosen carefully to indicate the main theme of the
+  headings is chosen carefully to introduce the main theme of the
   subsequent formal ML text.%
 \end{isamarkuptext}%
 \isamarkuptrue%
@@ -134,31 +134,29 @@
 
   \paragraph{Notation.}  A name consists of 1--3 \emph{words} (rarely
   4, but not more) that are separated by underscore.  There are three
-  variants concerning upper or lower case letters, which are just for
+  variants concerning upper or lower case letters, which are used for
   certain ML categories as follows:
 
   \medskip
   \begin{tabular}{lll}
   variant & example & ML categories \\\hline
   lower-case & \verb|foo_bar| & values, types, record fields \\
-  capitalized & \verb|Foo_Bar| & datatype constructors, \\
-    & & structures, functors \\
-  upper-case & \verb|FOO_BAR| & special values (combinators), \\
-    & & exception constructors, signatures \\
+  capitalized & \verb|Foo_Bar| & datatype constructors, structures, functors \\
+  upper-case & \verb|FOO_BAR| & special values, exception constructors, signatures \\
   \end{tabular}
   \medskip
 
   For historical reasons, many capitalized names omit underscores,
   e.g.\ old-style \verb|FooBar| instead of \verb|Foo_Bar|.
-  Genuine mixed-case names are \emph{not} used --- clear division of
-  words is essential for readability.\footnote{Camel-case was invented
-  to workaround the lack of underscore in some early non-ASCII
-  character sets and keywords.  Later it became a habitual in some
-  language communities that are now strong in numbers.}
+  Genuine mixed-case names are \emph{not} used, bacause clear division
+  of words is essential for readability.\footnote{Camel-case was
+  invented to workaround the lack of underscore in some early
+  non-ASCII character sets.  Later it became habitual in some language
+  communities that are now strong in numbers.}
 
   A single (capital) character does not count as ``word'' in this
-  respect: some Isabelle/ML are suffixed by extra markers like this:
-  \verb|foo_barT|.
+  respect: some Isabelle/ML names are suffixed by extra markers like
+  this: \verb|foo_barT|.
 
   Name variants are produced by adding 1--3 primes, e.g.\ \verb|foo'|, \verb|foo''|, or \verb|foo'''|, but not \verb|foo''''| or more.  Decimal digits scale better to larger numbers,
   e.g.\ \verb|foo0|, \verb|foo1|, \verb|foo42|.
@@ -168,11 +166,11 @@
   explicit qualification, as in \verb|Syntax.string_of_term| for
   example.  When devising names for structures and their components it
   is important aim at eye-catching compositions of both parts, because
-  this is how they are always seen in the sources and documentation.
-  For the same reasons, aliases of well-known library functions should
-  be avoided.
+  this is how they are seen in the sources and documentation.  For the
+  same reasons, aliases of well-known library functions should be
+  avoided.
 
-  Local names of function abstraction or case/lets bindings are
+  Local names of function abstraction or case/let bindings are
   typically shorter, sometimes using only rudiments of ``words'',
   while still avoiding cryptic shorthands.  An auxiliary function
   called \verb|helper|, \verb|aux|, or \verb|f| is
@@ -210,8 +208,8 @@
   \end{verbatim}
 
 
-  \paragraph{Specific conventions.} Here are some important specific
-  name forms that occur frequently in the sources.
+  \paragraph{Specific conventions.} Here are some specific name forms
+  that occur frequently in the sources.
 
   \begin{itemize}
 
@@ -236,8 +234,7 @@
 
   \item Variables of the main context types of the Isabelle/Isar
   framework (\secref{sec:context} and \chref{ch:local-theory}) have
-  firm naming conventions to allow to visualize the their data flow
-  via plain regular expressions in the editor.  In particular:
+  firm naming conventions as follows:
 
   \begin{itemize}
 
@@ -249,16 +246,19 @@
   \item generic contexts are called \verb|context|, rarely
   \verb|ctxt|
 
-  \item local theories are called \verb|lthy|, unless there is an
-  implicit conversion to a general proof context (semantic super-type)
+  \item local theories are called \verb|lthy|, except for local
+  theories that are treated as proof context (which is a semantic
+  super-type)
 
   \end{itemize}
 
   Variations with primed or decimal numbers are always possible, as
-  well as ``sematic prefixes'' like \verb|foo_thy| or \verb|bar_ctxt|, but the base conventions above need to be preserved.
+  well as sematic prefixes like \verb|foo_thy| or \verb|bar_ctxt|, but the base conventions above need to be preserved.
+  This allows to visualize the their data flow via plain regular
+  expressions in the editor.
 
-  \item The main logical entities (\secref{ch:logic}) also have
-  established naming convention:
+  \item The main logical entities (\secref{ch:logic}) have established
+  naming convention as follows:
 
   \begin{itemize}
 
@@ -278,8 +278,8 @@
 
   Proper semantic names override these conventions completely.  For
   example, the left-hand side of an equation (as a term) can be called
-  \verb|lhs| (not \verb|lhs_tm|).  Or a term that is treated
-  specifically as a variable can be called \verb|v| or \verb|x|.
+  \verb|lhs| (not \verb|lhs_tm|).  Or a term that is known
+  to be a variable can be called \verb|v| or \verb|x|.
 
   \end{itemize}%
 \end{isamarkuptext}%
@@ -296,11 +296,11 @@
   programming.
 
   \paragraph{Line length} is 80 characters according to ancient
-  standards, but we allow as much as 100 characters, not
-  more.\footnote{Readability requires to keep the beginning of a line
+  standards, but we allow as much as 100 characters (not
+  more).\footnote{Readability requires to keep the beginning of a line
   in view while watching its end.  Modern wide-screen displays do not
-  change the way how the human brain works.  As a rule of thumb,
-  sources need to be printable on plain paper.} The extra 20
+  change the way how the human brain works.  Sources also need to be
+  printable on plain paper with reasonable font-size.} The extra 20
   characters acknowledge the space requirements due to qualified
   library references in Isabelle/ML.
 
@@ -339,7 +339,8 @@
   curried function, or \verb|g (a, b)| for a tupled function.
   Note that the space between \verb|g| and the pair \verb|(a, b)| follows the important principle of
   \emph{compositionality}: the layout of \verb|g p| does not
-  change when \verb|p| is refined to a concrete pair.
+  change when \verb|p| is refined to the concrete pair
+  \verb|(a, b)|.
 
   \paragraph{Indentation} uses plain spaces, never hard
   tabulators.\footnote{Tabulators were invented to move the carriage
@@ -349,7 +350,7 @@
   editor configuration.}
 
   Each level of nesting is indented by 2 spaces, sometimes 1, very
-  rarely 4, never 8.
+  rarely 4, never 8 or any other odd number.
 
   Indentation follows a simple logical format that only depends on the
   nesting depth, not the accidental length of the text that initiates
@@ -382,18 +383,18 @@
   \emph{maintainability}) etc.
 
   \medskip For similar reasons, any kind of two-dimensional or tabular
-  layouts, ASCII-art with lines or boxes of stars etc. should be
+  layouts, ASCII-art with lines or boxes of asterisks etc.\ should be
   avoided.
 
-  \paragraph{Complex expressions} consisting of multi-clausal function
-  definitions, \verb|case|, \verb|handle|, \verb|let|,
-  or combinations require special attention.  The syntax of Standard
-  ML is a bit too ambitious in admitting too much variance that can
-  distort the meaning of the text.
+  \paragraph{Complex expressions} that consist of multi-clausal
+  function definitions, \verb|handle|, \verb|case|,
+  \verb|let| (and combinations) require special attention.  The
+  syntax of Standard ML is quite ambitious and admits a lot of
+  variance that can distort the meaning of the text.
 
-  Clauses of \verb|fun|, \verb|fn|, \verb|case|,
-  \verb|handle| get extra indentation to indicate the nesting
-  clearly.  For example:
+  Clauses of \verb|fun|, \verb|fn|, \verb|handle|,
+  \verb|case| get extra indentation to indicate the nesting
+  clearly.  Example:
 
   \begin{verbatim}
   (* RIGHT *)
@@ -414,8 +415,8 @@
 
   Body expressions consisting of \verb|case| or \verb|let|
   require care to maintain compositionality, to prevent loss of
-  logical indentation where it is particularly imporant to see the
-  structure of the text later on.  Example:
+  logical indentation where it is especially important to see the
+  structure of the text.  Example:
 
   \begin{verbatim}
   (* RIGHT *)
@@ -446,8 +447,8 @@
   \end{verbatim}
 
   Extra parentheses around \verb|case| expressions are optional,
-  but help to analyse the nesting easily based on simple string
-  matching in the editor.
+  but help to analyse the nesting based on character matching in the
+  editor.
 
   \medskip There are two main exceptions to the overall principle of
   compositionality in the layout of complex expressions.
@@ -468,7 +469,7 @@
   \item \verb|fn| abstractions are often layed-out as if they
   would lack any structure by themselves.  This traditional form is
   motivated by the possibility to shift function arguments back and
-  forth wrt.\ additional combinators.  For example:
+  forth wrt.\ additional combinators.  Example:
 
   \begin{verbatim}
   (* RIGHT *)
@@ -516,19 +517,20 @@
   (mainly Poly/ML, but also SML/NJ).
 
   Isabelle/Pure marks the point where the original ML toplevel is
-  superseded by the Isar toplevel that maintains a uniform environment
-  for arbitrary ML values (see also \secref{sec:context}).  This
-  formal context holds logical entities as well as ML compiler
-  bindings, among many other things.  Raw Standard ML is never
-  encountered again after the initial bootstrap of Isabelle/Pure.
+  superseded by the Isar toplevel that maintains a uniform context for
+  arbitrary ML values (see also \secref{sec:context}).  This formal
+  environment holds ML compiler bindings, logical entities, and many
+  other things.  Raw SML is never encountered again after the initial
+  bootstrap of Isabelle/Pure.
 
-  Object-logics such as Isabelle/HOL are built within the
-  Isabelle/ML/Isar environment of Pure by introducing suitable
-  theories with associated ML text, either inlined or as separate
-  files.  Thus Isabelle/HOL is defined as a regular user-space
-  application within the Isabelle framework.  Further add-on tools can
-  be implemented in ML within the Isar context in the same manner: ML
-  is part of the regular user-space repertoire of Isabelle.%
+  Object-logics like Isabelle/HOL are built within the
+  Isabelle/ML/Isar environment by introducing suitable theories with
+  associated ML modules, either inlined or as separate files.  Thus
+  Isabelle/HOL is defined as a regular user-space application within
+  the Isabelle framework.  Further add-on tools can be implemented in
+  ML within the Isar context in the same manner: ML is part of the
+  standard repertoire of Isabelle, and there is no distinction between
+  ``user'' and ``developer'' in this respect.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -537,16 +539,16 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-The primary Isar source language provides various facilities
-  to open a ``window'' to the underlying ML compiler.  Especially see
-  \indexref{}{command}{use}\hyperlink{command.use}{\mbox{\isa{\isacommand{use}}}} and \indexref{}{command}{ML}\hyperlink{command.ML}{\mbox{\isa{\isacommand{ML}}}}, which work exactly the
+The primary Isar source language provides facilities to ``open
+  a window'' to the underlying ML compiler.  Especially see the Isar
+  commands \indexref{}{command}{use}\hyperlink{command.use}{\mbox{\isa{\isacommand{use}}}} and \indexref{}{command}{ML}\hyperlink{command.ML}{\mbox{\isa{\isacommand{ML}}}}: both work the
   same way, only the source text is provided via a file vs.\ inlined,
   respectively.  Apart from embedding ML into the main theory
   definition like that, there are many more commands that refer to ML
   source, such as \indexref{}{command}{setup}\hyperlink{command.setup}{\mbox{\isa{\isacommand{setup}}}} or \indexref{}{command}{declaration}\hyperlink{command.declaration}{\mbox{\isa{\isacommand{declaration}}}}.
-  An example of even more fine-grained embedding of ML into Isar is
-  the proof method \indexref{}{method}{tactic}\hyperlink{method.tactic}{\mbox{\isa{tactic}}}, which refines the pending goal state
-  via a given expression of type \verb|tactic|.%
+  Even more fine-grained embedding of ML into Isar is encountered in
+  the proof method \indexref{}{method}{tactic}\hyperlink{method.tactic}{\mbox{\isa{tactic}}}, which refines the pending
+  goal state via a given expression of type \verb|tactic|.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -589,7 +591,7 @@
 \endisadelimML
 %
 \begin{isamarkuptext}%
-Here the ML environment is really managed by Isabelle, i.e.\
+Here the ML environment is already managed by Isabelle, i.e.\
   the \verb|factorial| function is not yet accessible in the preceding
   paragraph, nor in a different theory that is independent from the
   current one in the import hierarchy.
@@ -597,16 +599,15 @@
   Removing the above ML declaration from the source text will remove
   any trace of this definition as expected.  The Isabelle/ML toplevel
   environment is managed in a \emph{stateless} way: unlike the raw ML
-  toplevel or similar command loops of Computer Algebra systems, there
-  are no global side-effects involved here.\footnote{Such a stateless
-  compilation environment is also a prerequisite for robust parallel
-  compilation within independent nodes of the implicit theory
-  development graph.}
+  toplevel there are no global side-effects involved
+  here.\footnote{Such a stateless compilation environment is also a
+  prerequisite for robust parallel compilation within independent
+  nodes of the implicit theory development graph.}
 
-  \medskip The next example shows how to embed ML into Isar proofs.
-  Instead of \indexref{}{command}{ML}\hyperlink{command.ML}{\mbox{\isa{\isacommand{ML}}}} for theory mode, we use \indexref{}{command}{ML\_prf}\hyperlink{command.ML-prf}{\mbox{\isa{\isacommand{ML{\isacharunderscore}prf}}}} for proof mode.  As illustrated below, its effect on the
-  ML environment is local to the whole proof body, while ignoring its
-  internal block structure.%
+  \medskip The next example shows how to embed ML into Isar proofs, using
+ \indexref{}{command}{ML\_prf}\hyperlink{command.ML-prf}{\mbox{\isa{\isacommand{ML{\isacharunderscore}prf}}}} instead of Instead of \indexref{}{command}{ML}\hyperlink{command.ML}{\mbox{\isa{\isacommand{ML}}}}.
+  As illustrated below, the effect on the ML environment is local to
+  the whole proof body, ignoring the block structure.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 \isacommand{example{\isacharunderscore}proof}\isamarkupfalse%
@@ -620,9 +621,6 @@
 \isacommand{ML{\isacharunderscore}prf}\isamarkupfalse%
 \ {\isacharverbatimopen}\ val\ a\ {\isacharequal}\ {\isadigit{1}}\ {\isacharverbatimclose}\isanewline
 \ \ \isacommand{{\isacharbraceleft}}\isamarkupfalse%
-\ %
-\isamarkupcmt{Isar block structure ignored by ML environment%
-}
 \isanewline
 \ \ \ \ \isacommand{ML{\isacharunderscore}prf}\isamarkupfalse%
 \ {\isacharverbatimopen}\ val\ b\ {\isacharequal}\ a\ {\isacharplus}\ {\isadigit{1}}\ {\isacharverbatimclose}\isanewline
@@ -644,12 +642,12 @@
 %
 \begin{isamarkuptext}%
 By side-stepping the normal scoping rules for Isar proof
-  blocks, embedded ML code can refer to the different contexts
-  explicitly, and manipulate corresponding entities, e.g.\ export a
-  fact from a block context.
+  blocks, embedded ML code can refer to the different contexts and
+  manipulate corresponding entities, e.g.\ export a fact from a block
+  context.
 
   \medskip Two further ML commands are useful in certain situations:
-  \indexref{}{command}{ML\_val}\hyperlink{command.ML-val}{\mbox{\isa{\isacommand{ML{\isacharunderscore}val}}}} and \indexref{}{command}{ML\_command}\hyperlink{command.ML-command}{\mbox{\isa{\isacommand{ML{\isacharunderscore}command}}}} are both
+  \indexref{}{command}{ML\_val}\hyperlink{command.ML-val}{\mbox{\isa{\isacommand{ML{\isacharunderscore}val}}}} and \indexref{}{command}{ML\_command}\hyperlink{command.ML-command}{\mbox{\isa{\isacommand{ML{\isacharunderscore}command}}}} are
   \emph{diagnostic} in the sense that there is no effect on the
   underlying environment, and can thus used anywhere (even outside a
   theory).  The examples below produce long strings of digits by
@@ -719,8 +717,8 @@
   formal context is passed as a thread-local reference variable.  Thus
   ML code may access the theory context during compilation, by reading
   or writing the (local) theory under construction.  Note that such
-  direct access to the compile-time context is rare; in practice it is
-  typically via some derived ML functions.%
+  direct access to the compile-time context is rare.  In practice it
+  is typically done via some derived ML functions instead.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -733,7 +731,7 @@
 \begin{isamarkuptext}%
 \begin{mldecls}
   \indexdef{}{ML}{ML\_Context.the\_generic\_context}\verb|ML_Context.the_generic_context: unit -> Context.generic| \\
-  \indexdef{}{ML}{Context.$>$$>$}\verb|Context.>>  : (Context.generic -> Context.generic) -> unit| \\
+  \indexdef{}{ML}{Context.$>$$>$}\verb|Context.>> : (Context.generic -> Context.generic) -> unit| \\
   \indexdef{}{ML}{bind\_thms}\verb|bind_thms: string * thm list -> unit| \\
   \indexdef{}{ML}{bind\_thm}\verb|bind_thm: string * thm -> unit| \\
   \end{mldecls}
@@ -755,11 +753,11 @@
   put into a global ``standard'' format before being stored.
 
   \item \verb|bind_thm| is similar to \verb|bind_thms| but refers to a
-  singleton theorem.
+  singleton fact.
 
   \end{description}
 
-  It is very important to note that the above functions are really
+  It is important to note that the above functions are really
   restricted to the compile time, even though the ML compiler is
   invoked at run-time.  The majority of ML code either uses static
   antiquotations (\secref{sec:ML-antiq}) or refers to the theory or
@@ -780,9 +778,8 @@
 %
 \begin{isamarkuptext}%
 A very important consequence of embedding SML into Isar is the
-  concept of \emph{ML antiquotation}.  First, the standard token
-  language of ML is augmented by special syntactic entities of the
-  following form:
+  concept of \emph{ML antiquotation}.  The standard token language of
+  ML is augmented by special syntactic entities of the following form:
 
   \indexouternonterm{antiquote}
   \begin{rail}
@@ -791,8 +788,8 @@
   \end{rail}
 
   Here \hyperlink{syntax.nameref}{\mbox{\isa{nameref}}} and \hyperlink{syntax.args}{\mbox{\isa{args}}} are regular outer syntax
-  categories.  Note that attributes and proof methods use similar
-  syntax.
+  categories \cite{isabelle-isar-ref}.  Attributes and proof methods
+  use similar syntax.
 
   \medskip A regular antiquotation \isa{{\isacharat}{\isacharbraceleft}name\ args{\isacharbraceright}} processes
   its arguments by the usual means of the Isar source language, and
@@ -808,7 +805,7 @@
   effect by introducing local blocks within the pre-compilation
   environment.
 
-  \medskip See also \cite{Wenzel-Chaieb:2007b} for a slightly broader
+  \medskip See also \cite{Wenzel-Chaieb:2007b} for a broader
   perspective on Isabelle/ML antiquotations.%
 \end{isamarkuptext}%
 \isamarkuptrue%
@@ -863,10 +860,9 @@
 \isatagmlex
 %
 \begin{isamarkuptext}%
-The following artificial example gives a first
-  impression about using the antiquotation elements introduced so far,
-  together with the basic \isa{{\isacharat}{\isacharbraceleft}thm{\isacharbraceright}} antiquotation defined
-  later.%
+The following artificial example gives some impression
+  about the antiquotation elements introduced so far, together with
+  the important \isa{{\isacharat}{\isacharbraceleft}thm{\isacharbraceright}} antiquotation defined later.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -910,8 +906,9 @@
 \endisadelimML
 %
 \begin{isamarkuptext}%
-The extra block delimiters do not affect the compiled code itself, i.e.\
-  function \verb|foo| is available in the present context.%
+The extra block delimiters do not affect the compiled code
+  itself, i.e.\ function \verb|foo| is available in the present context
+  of this paragraph.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -926,19 +923,19 @@
   functions in that setting can save a lot of extra boiler-plate of
   redundant shuffling of arguments, auxiliary abstractions etc.
 
-  First of all, functions are usually \emph{curried}: the idea of
-  \isa{f} taking \isa{n} arguments of type \isa{{\isasymtau}\isactrlsub i} (for
-  \isa{i\ {\isasymin}\ {\isacharbraceleft}{\isadigit{1}}{\isacharcomma}\ {\isasymdots}\ n{\isacharbraceright}}) with result \isa{{\isasymtau}} is represented by
-  the iterated function space \isa{{\isasymtau}\isactrlsub {\isadigit{1}}\ {\isasymrightarrow}\ {\isasymdots}\ {\isasymrightarrow}\ {\isasymtau}\isactrlsub n\ {\isasymrightarrow}\ {\isasymtau}}.  This is
-  isomorphic to the encoding via tuples \isa{{\isasymtau}\isactrlsub {\isadigit{1}}\ {\isasymtimes}\ {\isasymdots}\ {\isasymtimes}\ {\isasymtau}\isactrlsub n\ {\isasymrightarrow}\ {\isasymtau}}, but
-  the curried version fits more smoothly into the basic
-  calculus.\footnote{The difference is even more significant in
-  higher-order logic, because additional the redundant tuple structure
-  needs to be accommodated by formal reasoning.}
+  Functions are usually \emph{curried}: the idea of turning arguments
+  of type \isa{{\isasymtau}\isactrlsub i} (for \isa{i\ {\isasymin}\ {\isacharbraceleft}{\isadigit{1}}{\isacharcomma}\ {\isasymdots}\ n{\isacharbraceright}}) into a result of
+  type \isa{{\isasymtau}} is represented by the iterated function space
+  \isa{{\isasymtau}\isactrlsub {\isadigit{1}}\ {\isasymrightarrow}\ {\isasymdots}\ {\isasymrightarrow}\ {\isasymtau}\isactrlsub n\ {\isasymrightarrow}\ {\isasymtau}}.  This is isomorphic to the well-known
+  encoding via tuples \isa{{\isasymtau}\isactrlsub {\isadigit{1}}\ {\isasymtimes}\ {\isasymdots}\ {\isasymtimes}\ {\isasymtau}\isactrlsub n\ {\isasymrightarrow}\ {\isasymtau}}, but the curried
+  version fits more smoothly into the basic calculus.\footnote{The
+  difference is even more significant in higher-order logic, because
+  the redundant tuple structure needs to be accommodated by formal
+  reasoning.}
 
   Currying gives some flexiblity due to \emph{partial application}.  A
   function \isa{f{\isacharcolon}\ {\isasymtau}\isactrlsub {\isadigit{1}}\ {\isasymrightarrow}\ {\isasymtau}\isactrlbsub {\isadigit{2}}\isactrlesub \ {\isasymrightarrow}\ {\isasymtau}} can be applied to \isa{x{\isacharcolon}\ {\isasymtau}\isactrlsub {\isadigit{1}}}
-  and the remaining \isa{{\isacharparenleft}f\ x{\isacharparenright}{\isacharcolon}\ {\isasymtau}\isactrlsub {\isadigit{2}}\ {\isasymrightarrow}\ {\isasymtau}} passed to other functions
+  and the remaining \isa{{\isacharparenleft}f\ x{\isacharparenright}{\isacharcolon}\ {\isasymtau}\isactrlsub {\isadigit{2}}\ {\isasymrightarrow}\ {\isasymtau}} passed to another function
   etc.  How well this works in practice depends on the order of
   arguments.  In the worst case, arguments are arranged erratically,
   and using a function in a certain situation always requires some
@@ -946,11 +943,10 @@
   decorate the code with meaningless permutations of arguments.
 
   This can be avoided by \emph{canonical argument order}, which
-  observes certain standard patterns of function definitions and
-  minimizes adhoc permutations in their application.  In Isabelle/ML,
-  large portions of non-trivial source code are written without ever
-  using the infamous function \isa{swap{\isacharcolon}\ {\isasymalpha}\ {\isasymtimes}\ {\isasymbeta}\ {\isasymrightarrow}\ {\isasymbeta}\ {\isasymtimes}\ {\isasymalpha}} or the
-  combinator \isa{C{\isacharcolon}\ {\isacharparenleft}{\isasymalpha}\ {\isasymrightarrow}\ {\isasymbeta}\ {\isasymrightarrow}\ {\isasymgamma}{\isacharparenright}\ {\isasymrightarrow}\ {\isacharparenleft}{\isasymbeta}\ {\isasymrightarrow}\ {\isasymalpha}\ {\isasymrightarrow}\ {\isasymgamma}{\isacharparenright}}, which is not even
+  observes certain standard patterns and minimizes adhoc permutations
+  in their application.  In Isabelle/ML, large portions text can be
+  written without ever using \isa{swap{\isacharcolon}\ {\isasymalpha}\ {\isasymtimes}\ {\isasymbeta}\ {\isasymrightarrow}\ {\isasymbeta}\ {\isasymtimes}\ {\isasymalpha}}, or the
+  combinator \isa{C{\isacharcolon}\ {\isacharparenleft}{\isasymalpha}\ {\isasymrightarrow}\ {\isasymbeta}\ {\isasymrightarrow}\ {\isasymgamma}{\isacharparenright}\ {\isasymrightarrow}\ {\isacharparenleft}{\isasymbeta}\ {\isasymrightarrow}\ {\isasymalpha}\ {\isasymrightarrow}\ {\isasymgamma}{\isacharparenright}} that is not even
   defined in our library.
 
   \medskip The basic idea is that arguments that vary less are moved
@@ -977,15 +973,15 @@
   concrete container.  The argument order makes it easy to use other
   combinators: \isa{forall\ {\isacharparenleft}member\ B{\isacharparenright}\ list} will check a list of
   elements for membership in \isa{B} etc. Often the explicit
-  \isa{list} is pointless and can be contracted in the application
-  \isa{forall\ {\isacharparenleft}member\ B{\isacharparenright}} to get directly a predicate again.
+  \isa{list} is pointless and can be contracted to \isa{forall\ {\isacharparenleft}member\ B{\isacharparenright}} to get directly a predicate again.
 
-  The update operation naturally ``varies'' the container, so it moves
+  In contrast, an update operation varies the container, so it moves
   to the right: \isa{insert\ a} is a function \isa{{\isasymbeta}\ {\isasymrightarrow}\ {\isasymbeta}} to
   insert a value \isa{a}.  These can be composed naturally as
-  follows: \isa{insert\ c\ {\isasymcirc}\ insert\ b\ {\isasymcirc}\ insert\ a}.  This works well,
-  apart from some awkwardness due to conventional mathematical
-  function composition, which can be easily amended as follows.%
+  \isa{insert\ c\ {\isasymcirc}\ insert\ b\ {\isasymcirc}\ insert\ a}.  The slightly awkward
+  inversion if the composition order is due to conventional
+  mathematical notation, which can be easily amended as explained
+  below.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -995,10 +991,10 @@
 %
 \begin{isamarkuptext}%
 Regular function application and infix notation works best for
-  relatively deeply structured expressions, e.g.\ \isa{h\ {\isacharparenleft}f\ x\ y\ {\isacharplus}\ g\ z{\isacharparenright}}.  The important special case if \emph{linear transformation}
-  applies a cascade of functions as follows: \isa{f\isactrlsub n\ {\isacharparenleft}{\isasymdots}\ {\isacharparenleft}f\isactrlsub {\isadigit{1}}\ x{\isacharparenright}{\isacharparenright}}.
-  This becomes hard to read and maintain if the functions are
-  themselves complex expressions.  The notation can be significantly
+  relatively deeply structured expressions, e.g.\ \isa{h\ {\isacharparenleft}f\ x\ y\ {\isacharplus}\ g\ z{\isacharparenright}}.  The important special case of \emph{linear transformation}
+  applies a cascade of functions \isa{f\isactrlsub n\ {\isacharparenleft}{\isasymdots}\ {\isacharparenleft}f\isactrlsub {\isadigit{1}}\ x{\isacharparenright}{\isacharparenright}}.  This
+  becomes hard to read and maintain if the functions are themselves
+  given as complex expressions.  The notation can be significantly
   improved by introducing \emph{forward} versions of application and
   composition as follows:
 
@@ -1056,7 +1052,7 @@
 %
 \begin{isamarkuptext}%
 As explained above, a function \isa{f{\isacharcolon}\ {\isasymalpha}\ {\isasymrightarrow}\ {\isasymbeta}\ {\isasymrightarrow}\ {\isasymbeta}} can be
-  understood as update on a configuration of type \isa{{\isasymbeta}} that is
+  understood as update on a configuration of type \isa{{\isasymbeta}},
   parametrized by arguments of type \isa{{\isasymalpha}}.  Given \isa{a{\isacharcolon}\ {\isasymalpha}}
   the partial application \isa{{\isacharparenleft}f\ a{\isacharparenright}{\isacharcolon}\ {\isasymbeta}\ {\isasymrightarrow}\ {\isasymbeta}} operates
   homogeneously on \isa{{\isasymbeta}}.  This can be iterated naturally over a
@@ -1106,9 +1102,10 @@
   \begin{warn}
   The literature on functional programming provides a multitude of
   combinators called \isa{foldl}, \isa{foldr} etc.  SML97
-  provides its own variations as \verb|List.foldl| and \verb|List.foldr|, while the classic Isabelle library still has the
-  slightly more convenient historic \verb|Library.foldl| and \verb|Library.foldr|.  To avoid further confusion, all of this should be
-  ignored, and \verb|fold| (or \verb|fold_rev|) used exclusively.
+  provides its own variations as \verb|List.foldl| and \verb|List.foldr|, while the classic Isabelle library also has the
+  historic \verb|Library.foldl| and \verb|Library.foldr|.  To avoid
+  further confusion, all of this should be ignored, and \verb|fold| (or
+  \verb|fold_rev|) used exclusively.
   \end{warn}%
 \end{isamarkuptext}%
 \isamarkuptrue%
@@ -1173,8 +1170,9 @@
   memory (``deforestation''), but requires some practice to read and
   write it fluently.
 
-  \medskip The next example elaborates this idea and demonstrates fast
-  accumulation of tree content using a text buffer.%
+  \medskip The next example elaborates the idea of canonical
+  iteration, demonstrating fast accumulation of tree content using a
+  text buffer.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -1212,30 +1210,27 @@
 \begin{isamarkuptext}%
 The slow part of \verb|slow_content| is the \verb|implode| of
   the recursive results, because it copies previously produced strings
-  afresh.
+  again.
 
   The incremental \verb|add_content| avoids this by operating on a
-  buffer that is passed-through in a linear fashion.  Using \verb|#>| and contraction over the actual \verb|Buffer.T| argument
-  saves some additional boiler-plate, but requires again some
-  practice.  Of course, the two \verb|Buffer.add| invocations with
-  concatenated strings could have been split into smaller parts, but
-  this would have obfuscated the source without making a big
-  difference in allocations.  Here we have done some
+  buffer that is passed through in a linear fashion.  Using \verb|#>| and contraction over the actual buffer argument saves some
+  additional boiler-plate.  Of course, the two \verb|Buffer.add|
+  invocations with concatenated strings could have been split into
+  smaller parts, but this would have obfuscated the source without
+  making a big difference in allocations.  Here we have done some
   peephole-optimization for the sake of readability.
 
   Another benefit of \verb|add_content| is its ``open'' form as a
-  function on \verb|Buffer.T| that can be continued in further
-  linear transformations, folding etc.  Thus it is more compositional
-  than the naive \verb|slow_content|.  As a realistic example, compare
-  the slightly old-style \verb|Term.maxidx_of_term: term -> int| with
-  the newer \verb|Term.maxidx_term: term -> int -> int| in
-  Isabelle/Pure.
+  function on buffers that can be continued in further linear
+  transformations, folding etc.  Thus it is more compositional than
+  the naive \verb|slow_content|.  As realistic example, compare the
+  old-style \verb|Term.maxidx_of_term: term -> int| with the newer
+  \verb|Term.maxidx_term: term -> int -> int| in Isabelle/Pure.
 
-  Note that \verb|fast_content| above is only defined for completeness
-  of the example.  In many practical situations, it is customary to
-  defined the incremental \verb|add_content| only and leave the
-  initialization and termination to the concrete application by the
-  user.%
+  Note that \verb|fast_content| above is only defined as example.  In
+  many practical situations, it is customary to provide the
+  incremental \verb|add_content| only and leave the initialization and
+  termination to the concrete application by the user.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -1293,11 +1288,11 @@
 
   \item \verb|warning|~\isa{text} outputs \isa{text} as
   warning, which typically means some extra emphasis on the front-end
-  side (color highlighting, icon).
+  side (color highlighting, icons, etc.).
 
   \item \verb|error|~\isa{text} raises exception \verb|ERROR|~\isa{text} and thus lets the Isar toplevel print \isa{text} on the
   error channel, which typically means some extra emphasis on the
-  front-end side (color highlighting, icon).
+  front-end side (color highlighting, icons, etc.).
 
   This assumes that the exception is not handled before the command
   terminates.  Handling exception \verb|ERROR|~\isa{text} is a
@@ -1308,7 +1303,7 @@
   The actual error channel is accessed via \verb|Output.error_msg|, but
   the interaction protocol of Proof~General \emph{crashes} if that
   function is used in regular ML code: error output and toplevel
-  command failure always need to coincide here.
+  command failure always need to coincide.
   \end{warn}
 
   \end{description}
@@ -1322,14 +1317,14 @@
   some system console log, with a low chance that the user will ever
   see it.  Moreover, as a genuine side-effect on global process
   channels, there is no proper way to retract output when Isar command
-  transactions are reset.
+  transactions are reset by the system.
   \end{warn}
 
   \begin{warn}
   The message channels should be used in a message-oriented manner.
-  This means that multi-line output that logically belongs together
-  needs to be issued by a \emph{single} invocation of \verb|writeln|
-  etc.  with the functional concatenation of all message constituents.
+  This means that multi-line output that logically belongs together is
+  issued by a \emph{single} invocation of \verb|writeln| etc.\ with the
+  functional concatenation of all message constituents.
   \end{warn}%
 \end{isamarkuptext}%
 \isamarkuptrue%
@@ -1421,16 +1416,16 @@
   These exceptions indicate a genuine breakdown of the program, so the
   main purpose is to determine quickly what has happened where.
   Traditionally, the (short) exception message would include the name
-  of an ML function, although this no longer necessary, because the ML
-  runtime system prints a detailed source position of the
+  of an ML function, although this is no longer necessary, because the
+  ML runtime system prints a detailed source position of the
   corresponding \verb|raise| keyword.
 
   \medskip User modules can always introduce their own custom
   exceptions locally, e.g.\ to organize internal failures robustly
   without overlapping with existing exceptions.  Exceptions that are
   exposed in module signatures require extra care, though, and should
-  \emph{not} be introduced by default.  Surprise by end-users or ML
-  users of a module can be often minimized by using plain user errors.
+  \emph{not} be introduced by default.  Surprise by users of a module
+  can be often minimized by using plain user errors instead.
 
   \paragraph{Interrupts.}  These indicate arbitrary system events:
   both the ML runtime system and the Isabelle/ML infrastructure signal
@@ -1444,7 +1439,7 @@
   that intercepts interrupts becomes dependent on physical effects of
   the environment.  Even worse, exception handling patterns that are
   too general by accident, e.g.\ by mispelled exception constructors,
-  will cover interrupts unintentionally, and thus render the program
+  will cover interrupts unintentionally and thus render the program
   semantics ill-defined.
 
   Note that the Interrupt exception dates back to the original SML90
@@ -1452,8 +1447,8 @@
   avoid its malign impact on ML program semantics, but without
   providing a viable alternative.  Isabelle/ML recovers physical
   interruptibility (which an indispensable tool to implement managed
-  evaluation of Isar command transactions), but requires user code to
-  be strictly transparent wrt.\ interrupts.
+  evaluation of command transactions), but requires user code to be
+  strictly transparent wrt.\ interrupts.
 
   \begin{warn}
   Isabelle/ML user code needs to terminate promptly on interruption,
@@ -1493,8 +1488,8 @@
   \item \verb|can| is similar to \verb|try| with more abstract result.
 
   \item \verb|ERROR|~\isa{msg} represents user errors; this
-  exception is always raised via the \verb|error| function (see
-  \secref{sec:message-channels}).
+  exception is normally raised indirectly via the \verb|error| function
+  (see \secref{sec:message-channels}).
 
   \item \verb|Fail|~\isa{msg} represents general program failures.
 
@@ -1506,11 +1501,10 @@
   while preserving its implicit position information (if possible,
   depending on the ML platform).
 
-  \item \verb|exception_trace|~\verb|(fn _ =>|~\isa{e}\verb|)| evaluates expression \isa{e} while printing
+  \item \verb|exception_trace|~\verb|(fn () =>|~\isa{e}\verb|)| evaluates expression \isa{e} while printing
   a full trace of its stack of nested exceptions (if possible,
-  depending on the ML platform).\footnote{In various versions of
-  Poly/ML the trace will appear on raw stdout of the Isabelle
-  process.}
+  depending on the ML platform).\footnote{In versions of Poly/ML the
+  trace will appear on raw stdout of the Isabelle process.}
 
   Inserting \verb|exception_trace| into ML code temporarily is useful
   for debugging, but not suitable for production code.
@@ -1539,10 +1533,10 @@
 
   \begin{description}
 
-  \item \isa{{\isacharat}{\isacharbraceleft}assert{\isacharbraceright}} inlines a function \isa{bool\ {\isasymRightarrow}\ unit}
-  that raises \verb|Fail| if the argument is \verb|false|.  Due to
-  inlining the source position of failed assertions is included in the
-  error output.
+  \item \isa{{\isacharat}{\isacharbraceleft}assert{\isacharbraceright}} inlines a function
+  \verb|bool -> unit| that raises \verb|Fail| if the argument is
+  \verb|false|.  Due to inlining the source position of failed
+  assertions is included in the error output.
 
   \end{description}%
 \end{isamarkuptext}%
@@ -1560,12 +1554,11 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-The basis library proposal of SML97 need to be treated with
+The basis library proposal of SML97 needs to be treated with
   caution.  Many of its operations simply do not fit with important
   Isabelle/ML conventions (like ``canonical argument order'', see
-  \secref{sec:canonical-argument-order}), others can cause serious
-  problems with the parallel evaluation model of Isabelle/ML (such as
-  \verb|TextIO.print| or \verb|OS.Process.system|).
+  \secref{sec:canonical-argument-order}), others cause problems with
+  the parallel evaluation model of Isabelle/ML (such as \verb|TextIO.print| or \verb|OS.Process.system|).
 
   Subsequently we give a brief overview of important operations on
   basic ML data types.%
@@ -1590,7 +1583,7 @@
   \begin{description}
 
   \item Type \verb|char| is \emph{not} used.  The smallest textual
-  unit in Isabelle is represented a ``symbol'' (see
+  unit in Isabelle is represented as a ``symbol'' (see
   \secref{sec:symbols}).
 
   \end{description}%
@@ -1628,10 +1621,10 @@
   This works uniformly for all supported ML platforms (Poly/ML and
   SML/NJ).
 
-  Literal integers in ML text (e.g.\ \verb|123456789012345678901234567890|) are forced to be of this one true
+  Literal integers in ML text are forced to be of this one true
   integer type --- overloading of SML97 is disabled.
 
-  Structure \verb|IntInf| of SML97 is obsolete, always use
+  Structure \verb|IntInf| of SML97 is obsolete and superseded by
   \verb|Int|.  Structure \verb|Integer| in \hyperlink{file.~~/src/Pure/General/integer.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}General{\isacharslash}integer{\isachardot}ML}}}} provides some additional
   operations.
 
@@ -1726,7 +1719,7 @@
 
   Note that \verb|insert| is conservative about elements that are
   already a \verb|member| of the list, while \verb|update| ensures that
-  the last entry is always put in front.  The latter discipline is
+  the latest entry is always put in front.  The latter discipline is
   often more appropriate in declarations of context data
   (\secref{sec:context-data}) that are issued by the user in Isar
   source: more recent declarations normally take precedence over
@@ -1750,10 +1743,12 @@
 \isatagmlex
 %
 \begin{isamarkuptext}%
-Using canonical \verb|fold| together with canonical \verb|cons|, or similar standard operations, alternates the orientation of
-  data.  The is quite natural and should not altered forcible by
-  inserting extra applications \verb|rev|.  The alternative \verb|fold_rev| can be used in the relatively few situations, where
-  alternation should be prevented.%
+Using canonical \verb|fold| together with \verb|cons|, or
+  similar standard operations, alternates the orientation of data.
+  The is quite natural and should not be altered forcible by inserting
+  extra applications of \verb|rev|.  The alternative \verb|fold_rev| can
+  be used in the few situations, where alternation should be
+  prevented.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -1852,7 +1847,8 @@
 
   \item \verb|AList.lookup|, \verb|AList.defined|, \verb|AList.update|
   implement the main ``framework operations'' for mappings in
-  Isabelle/ML, with standard conventions for names and types.
+  Isabelle/ML, following standard conventions for their names and
+  types.
 
   Note that a function called \isa{lookup} is obliged to express its
   partiality via an explicit option element.  There is no choice to
@@ -1954,7 +1950,7 @@
 
   Isabelle/Isar exploits the inherent structure of theories and proofs
   to support \emph{implicit parallelism} to a large extent.  LCF-style
-  theorem provides almost ideal conditions for that; see also
+  theorem provides almost ideal conditions for that, see also
   \cite{Wenzel:2009}.  This means, significant parts of theory and
   proof checking is parallelized by default.  A maximum speedup-factor
   of 3.0 on 4 cores and 5.0 on 8 cores can be
@@ -1967,11 +1963,10 @@
 
   \medskip ML threads lack the memory protection of separate
   processes, and operate concurrently on shared heap memory.  This has
-  the advantage that results of independent computations are
-  immediately available to other threads: abstract values can be
-  passed between threads without copying or awkward serialization that
-  is typically required for explicit message passing between separate
-  processes.
+  the advantage that results of independent computations are directly
+  available to other threads: abstract values can be passed without
+  copying or awkward serialization that is typically required for
+  separate processes.
 
   To make shared-memory multi-threading work robustly and efficiently,
   some programming guidelines need to be observed.  While the ML
@@ -2019,15 +2014,15 @@
   channels, environment variables, current working directory.
 
   \item Writable resources in the file-system that are shared among
-  different threads or other processes.
+  different threads or external processes.
 
   \end{itemize}
 
   Isabelle/ML provides various mechanisms to avoid critical shared
-  resources in most practical situations.  As last resort there are
-  some mechanisms for explicit synchronization.  The following
-  guidelines help to make Isabelle/ML programs work smoothly in a
-  concurrent environment.
+  resources in most situations.  As last resort there are some
+  mechanisms for explicit synchronization.  The following guidelines
+  help to make Isabelle/ML programs work smoothly in a concurrent
+  environment.
 
   \begin{itemize}
 
@@ -2037,7 +2032,7 @@
   plain value and user tools can get/map their own data in a purely
   functional manner.  Configuration options within the context
   (\secref{sec:config-options}) provide simple drop-in replacements
-  for formerly writable flags.
+  for historic reference variables.
 
   \item Keep components with local state information re-entrant.
   Instead of poking initial values into (private) global references, a
@@ -2208,8 +2203,8 @@
 
   Entering the critical section without contention is very fast, and
   several basic system operations do so frequently.  Each thread
-  should leave the critical section quickly, otherwise parallel
-  performance may degrade.
+  should stay within the critical section quickly only very briefly,
+  otherwise parallel performance may degrade.
 
   \item \verb|CRITICAL| is the same as \verb|NAMED_CRITICAL| with empty
   name argument.
@@ -2222,17 +2217,17 @@
 
   \item \verb|Synchronized.guarded_access|~\isa{var\ f} lets the
   function \isa{f} operate within a critical section on the state
-  \isa{x} as follows: if \isa{f\ x} produces \verb|NONE|, we
-  continue to wait on the internal condition variable, expecting that
+  \isa{x} as follows: if \isa{f\ x} produces \verb|NONE|, it
+  continues to wait on the internal condition variable, expecting that
   some other thread will eventually change the content in a suitable
-  manner; if \isa{f\ x} produces \verb|SOME|~\isa{{\isacharparenleft}y{\isacharcomma}\ x{\isacharprime}{\isacharparenright}} we
-  are satisfied and assign the new state value \isa{x{\isacharprime}}, broadcast
-  a signal to all waiting threads on the associated condition
-  variable, and return the result \isa{y}.
+  manner; if \isa{f\ x} produces \verb|SOME|~\isa{{\isacharparenleft}y{\isacharcomma}\ x{\isacharprime}{\isacharparenright}} it is
+  satisfied and assigns the new state value \isa{x{\isacharprime}}, broadcasts a
+  signal to all waiting threads on the associated condition variable,
+  and returns the result \isa{y}.
 
   \end{description}
 
-  There are some further variants of the general \verb|Synchronized.guarded_access| combinator, see \hyperlink{file.~~/src/Pure/Concurrent/synchronized.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}Concurrent{\isacharslash}synchronized{\isachardot}ML}}}} for details.%
+  There are some further variants of the \verb|Synchronized.guarded_access| combinator, see \hyperlink{file.~~/src/Pure/Concurrent/synchronized.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}Concurrent{\isacharslash}synchronized{\isachardot}ML}}}} for details.%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -2250,11 +2245,7 @@
 \isatagmlex
 %
 \begin{isamarkuptext}%
-See \hyperlink{file.~~/src/Pure/Concurrent/mailbox.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}Concurrent{\isacharslash}mailbox{\isachardot}ML}}}} how to
-  implement a mailbox as synchronized variable over a purely
-  functional queue.
-
-  \medskip The following example implements a counter that produces
+The following example implements a counter that produces
   positive integers that are unique over the runtime of the Isabelle
   process:%
 \end{isamarkuptext}%
@@ -2301,10 +2292,15 @@
 \isadelimML
 %
 \endisadelimML
-\isanewline
+%
+\begin{isamarkuptext}%
+\medskip See \hyperlink{file.~~/src/Pure/Concurrent/mailbox.ML}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}Pure{\isacharslash}Concurrent{\isacharslash}mailbox{\isachardot}ML}}}} how
+  to implement a mailbox as synchronized variable over a purely
+  functional queue.%
+\end{isamarkuptext}%
+\isamarkuptrue%
 %
 \isadelimtheory
-\isanewline
 %
 \endisadelimtheory
 %
--- a/doc-src/IsarImplementation/Thy/document/Prelim.tex	Wed Oct 27 11:10:36 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/document/Prelim.tex	Wed Oct 27 11:11:35 2010 -0700
@@ -874,7 +874,7 @@
 %
 \begin{isamarkuptext}%
 Now the user can refer to \hyperlink{attribute.my-flag}{\mbox{\isa{my{\isacharunderscore}flag}}} in
-  declarations, while we can retrieve the current value from the
+  declarations, while ML tools can retrieve the current value from the
   context via \verb|Config.get|.%
 \end{isamarkuptext}%
 \isamarkuptrue%
@@ -1323,7 +1323,7 @@
 \endisadelimML
 %
 \begin{isamarkuptext}%
-\medskip The same works reletively to the formal context as
+\medskip The same works relatively to the formal context as
   follows.%
 \end{isamarkuptext}%
 \isamarkuptrue%
@@ -1726,7 +1726,7 @@
 %
 \begin{isamarkuptext}%
 The following example yields the source position of some
-  concrete binding inlined into the text.%
+  concrete binding inlined into the text:%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -1757,7 +1757,7 @@
 %
 \begin{isamarkuptext}%
 \medskip That position can be also printed in a message as
-  follows.%
+  follows:%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
@@ -1785,7 +1785,8 @@
 \begin{isamarkuptext}%
 This illustrates a key virtue of formalized bindings as
   opposed to raw specifications of base names: the system can use this
-  additional information for advanced feedback given to the user.%
+  additional information for feedback given to the user (error
+  messages etc.).%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
--- a/doc-src/IsarImplementation/Thy/document/Proof.tex	Wed Oct 27 11:10:36 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/document/Proof.tex	Wed Oct 27 11:11:35 2010 -0700
@@ -299,12 +299,12 @@
 \endisadelimML
 %
 \begin{isamarkuptext}%
-In the above example, the starting context had been derived
-  from the toplevel theory, which means that fixed variables are
-  internalized literally: \verb|x| is mapped again to
-  \verb|x|, and attempting to fix it again in the subsequent
-  context is an error.  Alternatively, fixed parameters can be renamed
-  explicitly as follows:%
+In the above example, the starting context is derived from the
+  toplevel theory, which means that fixed variables are internalized
+  literally: \verb|x| is mapped again to \verb|x|, and
+  attempting to fix it again in the subsequent context is an error.
+  Alternatively, fixed parameters can be renamed explicitly as
+  follows:%
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
--- a/doc-src/IsarRef/Thy/HOL_Specific.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/doc-src/IsarRef/Thy/HOL_Specific.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -401,11 +401,9 @@
   \begin{rail}
     'primrec' target? fixes 'where' equations
     ;
-    equations: (thmdecl? prop + '|')
+    ('fun' | 'function') target? functionopts? fixes \\ 'where' equations
     ;
-    ('fun' | 'function') target? functionopts? fixes 'where' clauses
-    ;
-    clauses: (thmdecl? prop ('(' 'otherwise' ')')? + '|')
+    equations: (thmdecl? prop + '|')
     ;
     functionopts: '(' (('sequential' | 'domintros' | 'tailrec' | 'default' term) + ',') ')'
     ;
@@ -550,6 +548,71 @@
   \end{description}
 *}
 
+subsection {* Functions with explicit partiality *}
+
+text {*
+  \begin{matharray}{rcl}
+    @{command_def (HOL) "partial_function"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
+    @{attribute_def (HOL) "partial_function_mono"} & : & @{text attribute} \\
+  \end{matharray}
+
+  \begin{rail}
+    'partial_function' target? '(' mode ')' fixes \\ 'where' thmdecl? prop
+  \end{rail}
+
+  \begin{description}
+
+  \item @{command (HOL) "partial_function"} defines recursive
+  functions based on fixpoints in complete partial orders. No
+  termination proof is required from the user or constructed
+  internally. Instead, the possibility of non-termination is modelled
+  explicitly in the result type, which contains an explicit bottom
+  element.
+
+  Pattern matching and mutual recursion are currently not supported.
+  Thus, the specification consists of a single function described by a
+  single recursive equation.
+
+  There are no fixed syntactic restrictions on the body of the
+  function, but the induced functional must be provably monotonic
+  wrt.\ the underlying order.  The monotonicitity proof is performed
+  internally, and the definition is rejected when it fails. The proof
+  can be influenced by declaring hints using the
+  @{attribute (HOL) partial_function_mono} attribute.
+
+  The mandatory @{text mode} argument specifies the mode of operation
+  of the command, which directly corresponds to a complete partial
+  order on the result type. By default, the following modes are
+  defined: 
+
+  \begin{description}
+  \item @{text option} defines functions that map into the @{type
+  option} type. Here, the value @{term None} is used to model a
+  non-terminating computation. Monotonicity requires that if @{term
+  None} is returned by a recursive call, then the overall result
+  must also be @{term None}. This is best achieved through the use of
+  the monadic operator @{const "Option.bind"}.
+  
+  \item @{text tailrec} defines functions with an arbitrary result
+  type and uses the slightly degenerated partial order where @{term
+  "undefined"} is the bottom element.  Now, monotonicity requires that
+  if @{term undefined} is returned by a recursive call, then the
+  overall result must also be @{term undefined}. In practice, this is
+  only satisfied when each recursive call is a tail call, whose result
+  is directly returned. Thus, this mode of operation allows the
+  definition of arbitrary tail-recursive functions.
+  \end{description}
+
+  Experienced users may define new modes by instantiating the locale
+  @{const "partial_function_definitions"} appropriately.
+
+  \item @{attribute (HOL) partial_function_mono} declares rules for
+  use in the internal monononicity proofs of partial function
+  definitions.
+
+  \end{description}
+
+*}
 
 subsection {* Old-style recursive function definitions (TFL) *}
 
--- a/doc-src/IsarRef/Thy/document/HOL_Specific.tex	Wed Oct 27 11:10:36 2010 -0700
+++ b/doc-src/IsarRef/Thy/document/HOL_Specific.tex	Wed Oct 27 11:11:35 2010 -0700
@@ -411,11 +411,9 @@
   \begin{rail}
     'primrec' target? fixes 'where' equations
     ;
-    equations: (thmdecl? prop + '|')
+    ('fun' | 'function') target? functionopts? fixes \\ 'where' equations
     ;
-    ('fun' | 'function') target? functionopts? fixes 'where' clauses
-    ;
-    clauses: (thmdecl? prop ('(' 'otherwise' ')')? + '|')
+    equations: (thmdecl? prop + '|')
     ;
     functionopts: '(' (('sequential' | 'domintros' | 'tailrec' | 'default' term) + ',') ')'
     ;
@@ -560,6 +558,71 @@
 \end{isamarkuptext}%
 \isamarkuptrue%
 %
+\isamarkupsubsection{Functions with explicit partiality%
+}
+\isamarkuptrue%
+%
+\begin{isamarkuptext}%
+\begin{matharray}{rcl}
+    \indexdef{HOL}{command}{partial\_function}\hypertarget{command.HOL.partial-function}{\hyperlink{command.HOL.partial-function}{\mbox{\isa{\isacommand{partial{\isacharunderscore}function}}}}} & : & \isa{{\isachardoublequote}local{\isacharunderscore}theory\ {\isasymrightarrow}\ local{\isacharunderscore}theory{\isachardoublequote}} \\
+    \indexdef{HOL}{attribute}{partial\_function\_mono}\hypertarget{attribute.HOL.partial-function-mono}{\hyperlink{attribute.HOL.partial-function-mono}{\mbox{\isa{partial{\isacharunderscore}function{\isacharunderscore}mono}}}} & : & \isa{attribute} \\
+  \end{matharray}
+
+  \begin{rail}
+    'partial_function' target? '(' mode ')' fixes \\ 'where' thmdecl? prop
+  \end{rail}
+
+  \begin{description}
+
+  \item \hyperlink{command.HOL.partial-function}{\mbox{\isa{\isacommand{partial{\isacharunderscore}function}}}} defines recursive
+  functions based on fixpoints in complete partial orders. No
+  termination proof is required from the user or constructed
+  internally. Instead, the possibility of non-termination is modelled
+  explicitly in the result type, which contains an explicit bottom
+  element.
+
+  Pattern matching and mutual recursion are currently not supported.
+  Thus, the specification consists of a single function described by a
+  single recursive equation.
+
+  There are no fixed syntactic restrictions on the body of the
+  function, but the induced functional must be provably monotonic
+  wrt.\ the underlying order.  The monotonicitity proof is performed
+  internally, and the definition is rejected when it fails. The proof
+  can be influenced by declaring hints using the
+  \hyperlink{attribute.HOL.partial-function-mono}{\mbox{\isa{partial{\isacharunderscore}function{\isacharunderscore}mono}}} attribute.
+
+  The mandatory \isa{mode} argument specifies the mode of operation
+  of the command, which directly corresponds to a complete partial
+  order on the result type. By default, the following modes are
+  defined: 
+
+  \begin{description}
+  \item \isa{option} defines functions that map into the \isa{option} type. Here, the value \isa{None} is used to model a
+  non-terminating computation. Monotonicity requires that if \isa{None} is returned by a recursive call, then the overall result
+  must also be \isa{None}. This is best achieved through the use of
+  the monadic operator \isa{{\isachardoublequote}Option{\isachardot}bind{\isachardoublequote}}.
+  
+  \item \isa{tailrec} defines functions with an arbitrary result
+  type and uses the slightly degenerated partial order where \isa{{\isachardoublequote}undefined{\isachardoublequote}} is the bottom element.  Now, monotonicity requires that
+  if \isa{undefined} is returned by a recursive call, then the
+  overall result must also be \isa{undefined}. In practice, this is
+  only satisfied when each recursive call is a tail call, whose result
+  is directly returned. Thus, this mode of operation allows the
+  definition of arbitrary tail-recursive functions.
+  \end{description}
+
+  Experienced users may define new modes by instantiating the locale
+  \isa{{\isachardoublequote}partial{\isacharunderscore}function{\isacharunderscore}definitions{\isachardoublequote}} appropriately.
+
+  \item \hyperlink{attribute.HOL.partial-function-mono}{\mbox{\isa{partial{\isacharunderscore}function{\isacharunderscore}mono}}} declares rules for
+  use in the internal monononicity proofs of partial function
+  definitions.
+
+  \end{description}%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
 \isamarkupsubsection{Old-style recursive function definitions (TFL)%
 }
 \isamarkuptrue%
--- a/doc-src/Nitpick/nitpick.tex	Wed Oct 27 11:10:36 2010 -0700
+++ b/doc-src/Nitpick/nitpick.tex	Wed Oct 27 11:11:35 2010 -0700
@@ -1751,7 +1751,7 @@
 {\slshape Nitpick ran out of time after checking 9 of 10 scopes.}
 \postw
 
-Furthermore, applying \textit{skew} or \textit{split} to a well-formed tree
+Furthermore, applying \textit{skew} or \textit{split} on a well-formed tree
 should not alter the tree:
 
 \prew
@@ -1802,10 +1802,10 @@
 \postw
 
 Nitpick's output reveals that the element $0$ was added as a left child of $1$,
-where both have a level of 1. This violates the second AA tree invariant, which
-states that a left child's level must be less than its parent's. This shouldn't
-come as a surprise, considering that we commented out the tree rebalancing code.
-Reintroducing the code seems to solve the problem:
+where both nodes have a level of 1. This violates the second AA tree invariant,
+which states that a left child's level must be less than its parent's. This
+shouldn't come as a surprise, considering that we commented out the tree
+rebalancing code. Reintroducing the code seems to solve the problem:
 
 \prew
 \textbf{theorem}~\textit{wf\_insort\/}:\kern.4em ``$\textit{wf}~t\,\Longrightarrow\, \textit{wf}~(\textit{insort}~t~x)$'' \\
--- a/doc-src/Sledgehammer/sledgehammer.tex	Wed Oct 27 11:10:36 2010 -0700
+++ b/doc-src/Sledgehammer/sledgehammer.tex	Wed Oct 27 11:11:35 2010 -0700
@@ -141,10 +141,10 @@
 
 \item[$\bullet$] Alternatively, you can download the Isabelle-aware E and SPASS
 binary packages from Isabelle's download page. Extract the archives, then add a
-line to your \texttt{\char`\~/.isabelle/etc/components} file with the absolute path to
-E or SPASS. For example, if the \texttt{components} does not exist
-yet and you extracted SPASS to \texttt{/usr/local/spass-3.7}, create
-the \texttt{components} file with the single line
+line to your \texttt{\char`\~/.isabelle/etc/components} file with the absolute
+path to E or SPASS. For example, if the \texttt{components} does not exist yet
+and you extracted SPASS to \texttt{/usr/local/spass-3.7}, create the
+\texttt{components} file with the single line
 
 \prew
 \texttt{/usr/local/spass-3.7}
@@ -223,13 +223,7 @@
 $([a] = [b]) = (a = b)$ \\
 Try this command: \textbf{by} (\textit{metis hd.simps}) \\
 To minimize the number of lemmas, try this: \\
-\textbf{sledgehammer} \textit{minimize} [\textit{prover} = \textit{remote\_sine\_e}]~(\textit{hd.simps}). \\[3\smallskipamount]
-%
-Sledgehammer: ``\textit{smt}'' for subgoal 1: \\
-$([a] = [b]) = (a = b)$ \\
-Try this command: \textbf{by} (\textit{smt hd.simps}) \\
-To minimize the number of lemmas, try this: \\
-\textbf{sledgehammer} \textit{minimize} [\textit{prover} = \textit{smt}]~(\textit{hd.simps}).
+\textbf{sledgehammer} \textit{minimize} [\textit{prover} = \textit{remote\_sine\_e}]~(\textit{hd.simps}).
 \postw
 
 Sledgehammer ran E, SPASS, Vampire, SInE-E, and the \textit{smt} proof method in
@@ -290,21 +284,21 @@
 In the general syntax, the \textit{subcommand} may be any of the following:
 
 \begin{enum}
-\item[$\bullet$] \textbf{\textit{run} (the default):} Runs Sledgehammer on subgoal number
-\textit{num} (1 by default), with the given options and facts.
+\item[$\bullet$] \textbf{\textit{run} (the default):} Runs Sledgehammer on
+subgoal number \textit{num} (1 by default), with the given options and facts.
 
 \item[$\bullet$] \textbf{\textit{minimize}:} Attempts to minimize the provided facts
 (specified in the \textit{facts\_override} argument) to obtain a simpler proof
 involving fewer facts. The options and goal number are as for \textit{run}.
 
-\item[$\bullet$] \textbf{\textit{messages}:} Redisplays recent messages issued by
-Sledgehammer. This allows you to examine results that might have been lost due
-to Sledgehammer's asynchronous nature. The \textit{num} argument specifies a
+\item[$\bullet$] \textbf{\textit{messages}:} Redisplays recent messages issued
+by Sledgehammer. This allows you to examine results that might have been lost
+due to Sledgehammer's asynchronous nature. The \textit{num} argument specifies a
 limit on the number of messages to display (5 by default).
 
-\item[$\bullet$] \textbf{\textit{available\_provers}:} Prints the list of installed provers.
-See \S\ref{installation} and \S\ref{mode-of-operation} for more information on
-how to install automatic provers.
+\item[$\bullet$] \textbf{\textit{available\_provers}:} Prints the list of
+installed provers. See \S\ref{installation} and \S\ref{mode-of-operation} for
+more information on how to install automatic provers.
 
 \item[$\bullet$] \textbf{\textit{running\_provers}:} Prints information about
 currently running automatic provers, including elapsed runtime and remaining
@@ -380,10 +374,12 @@
 \begin{enum}
 \item[$\bullet$] \qtybf{string}: A string.
 \item[$\bullet$] \qtybf{bool\/}: \textit{true} or \textit{false}.
-\item[$\bullet$] \qtybf{bool\_or\_smart\/}: \textit{true}, \textit{false}, or \textit{smart}.
+\item[$\bullet$] \qtybf{bool\_or\_smart\/}: \textit{true}, \textit{false}, or
+\textit{smart}.
 \item[$\bullet$] \qtybf{int\/}: An integer.
 \item[$\bullet$] \qtybf{int\_or\_smart\/}: An integer or \textit{smart}.
-\item[$\bullet$] \qtybf{time}: An integer followed by $\textit{min}$ (minutes), $s$ (seconds), or \textit{ms}
+\item[$\bullet$] \qtybf{time}: An integer followed by $\textit{min}$ (minutes),
+$s$ (seconds), or \textit{ms}
 (milliseconds), or the keyword \textit{none} ($\infty$ years).
 \end{enum}
 
@@ -531,9 +527,9 @@
 Specifies whether the \textbf{sledgehammer} command should explain what it does.
 
 \opfalse{debug}{no\_debug}
-Specifies whether Nitpick should display additional debugging information beyond
-what \textit{verbose} already displays. Enabling \textit{debug} also enables
-\textit{verbose} behind the scenes.
+Specifies whether Sledgehammer should display additional debugging information
+beyond what \textit{verbose} already displays. Enabling \textit{debug} also
+enables \textit{verbose} behind the scenes.
 
 \nopagebreak
 {\small See also \textit{overlord} (\S\ref{mode-of-operation}).}
@@ -559,9 +555,11 @@
 Specifies the expected outcome, which must be one of the following:
 
 \begin{enum}
-\item[$\bullet$] \textbf{\textit{some}:} Sledgehammer found a (potentially unsound) proof.
+\item[$\bullet$] \textbf{\textit{some}:} Sledgehammer found a (potentially
+unsound) proof.
 \item[$\bullet$] \textbf{\textit{none}:} Sledgehammer found no proof.
-\item[$\bullet$] \textbf{\textit{unknown}:} Sledgehammer encountered some problem.
+\item[$\bullet$] \textbf{\textit{unknown}:} Sledgehammer encountered some
+problem.
 \end{enum}
 
 Sledgehammer emits an error (if \textit{blocking} is enabled) or a warning
--- a/etc/isar-keywords.el	Wed Oct 27 11:10:36 2010 -0700
+++ b/etc/isar-keywords.el	Wed Oct 27 11:11:35 2010 -0700
@@ -95,7 +95,6 @@
     "find_consts"
     "find_theorems"
     "fix"
-    "fixpat"
     "fixrec"
     "from"
     "full_prf"
@@ -151,6 +150,7 @@
     "overloading"
     "parse_ast_translation"
     "parse_translation"
+    "partial_function"
     "pcpodef"
     "pr"
     "prefer"
@@ -220,6 +220,7 @@
     "sledgehammer"
     "sledgehammer_params"
     "smt_status"
+    "solve_direct"
     "sorry"
     "specification"
     "statespace"
@@ -392,6 +393,7 @@
     "refute"
     "sledgehammer"
     "smt_status"
+    "solve_direct"
     "term"
     "thm"
     "thm_deps"
@@ -462,7 +464,6 @@
     "extract"
     "extract_type"
     "finalconsts"
-    "fixpat"
     "fixrec"
     "fun"
     "hide_class"
@@ -490,6 +491,7 @@
     "overloading"
     "parse_ast_translation"
     "parse_translation"
+    "partial_function"
     "primrec"
     "print_ast_translation"
     "print_translation"
--- a/src/HOL/Boogie/Examples/Boogie_Dijkstra.certs	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Boogie/Examples/Boogie_Dijkstra.certs	Wed Oct 27 11:11:35 2010 -0700
@@ -1,4 +1,4 @@
-585d6a826013e58a18631a689e131cc7e15b8538 6889 0
+a959f1dced399224414912ad7379705a61749faa 6889 0
 #2 := false
 decl f11 :: (-> S5 S2 S1)
 decl ?v1!7 :: (-> S2 S2)
--- a/src/HOL/Boogie/Examples/Boogie_Dijkstra.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Boogie/Examples/Boogie_Dijkstra.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -83,9 +83,9 @@
 
 declare [[smt_certificates="~~/src/HOL/Boogie/Examples/Boogie_Dijkstra.certs"]]
 declare [[smt_fixed=true]]
+declare [[smt_oracle=false]]
 
 boogie_vc Dijkstra
-  using [[z3_proofs=true]]
   by boogie
 
 boogie_end 
--- a/src/HOL/Boogie/Examples/Boogie_Max.certs	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Boogie/Examples/Boogie_Max.certs	Wed Oct 27 11:11:35 2010 -0700
@@ -1,4 +1,4 @@
-3e8da1e086888bc525192984192051f1e39c6752 2224 0
+16ee32e23503be38842c16c90807e66120a960ee 2224 0
 #2 := false
 #8 := 0::int
 decl f5 :: (-> int int)
--- a/src/HOL/Boogie/Examples/Boogie_Max.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Boogie/Examples/Boogie_Max.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -40,9 +40,9 @@
 
 declare [[smt_certificates="~~/src/HOL/Boogie/Examples/Boogie_Max.certs"]]
 declare [[smt_fixed=true]]
+declare [[smt_oracle=false]]
 
 boogie_vc max
-  using [[z3_proofs=true]]
   by boogie
 
 boogie_end
--- a/src/HOL/Boogie/Examples/VCC_Max.certs	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Boogie/Examples/VCC_Max.certs	Wed Oct 27 11:11:35 2010 -0700
@@ -1,4 +1,4 @@
-a2095afb5c01aff7b980eece0ecc2b89ecf5e46c 7862 0
+5815fb85125ba62bb3b4a150604c8978e5a08a22 7862 0
 #2 := false
 decl f111 :: (-> S4 S5 int)
 decl f67 :: (-> S5 int S3 S5)
--- a/src/HOL/Boogie/Examples/VCC_Max.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Boogie/Examples/VCC_Max.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -48,11 +48,11 @@
 
 declare [[smt_certificates="~~/src/HOL/Boogie/Examples/VCC_Max.certs"]]
 declare [[smt_fixed=true]]
+declare [[smt_oracle=false]]
 
 boogie_status
 
 boogie_vc maximum
-  using [[z3_proofs=true]]
   by boogie
 
 boogie_end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Complete_Partial_Order.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -0,0 +1,263 @@
+(* Title:    HOL/Complete_Partial_Order.thy
+   Author:   Brian Huffman, Portland State University
+   Author:   Alexander Krauss, TU Muenchen
+*)
+
+header {* Chain-complete partial orders and their fixpoints *}
+
+theory Complete_Partial_Order
+imports Product_Type
+begin
+
+subsection {* Monotone functions *}
+
+text {* Dictionary-passing version of @{const Orderings.mono}. *}
+
+definition monotone :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
+where "monotone orda ordb f \<longleftrightarrow> (\<forall>x y. orda x y \<longrightarrow> ordb (f x) (f y))"
+
+lemma monotoneI[intro?]: "(\<And>x y. orda x y \<Longrightarrow> ordb (f x) (f y))
+ \<Longrightarrow> monotone orda ordb f"
+unfolding monotone_def by iprover
+
+lemma monotoneD[dest?]: "monotone orda ordb f \<Longrightarrow> orda x y \<Longrightarrow> ordb (f x) (f y)"
+unfolding monotone_def by iprover
+
+
+subsection {* Chains *}
+
+text {* A chain is a totally-ordered set. Chains are parameterized over
+  the order for maximal flexibility, since type classes are not enough.
+*}
+
+definition
+  chain :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool"
+where
+  "chain ord S \<longleftrightarrow> (\<forall>x\<in>S. \<forall>y\<in>S. ord x y \<or> ord y x)"
+
+lemma chainI:
+  assumes "\<And>x y. x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> ord x y \<or> ord y x"
+  shows "chain ord S"
+using assms unfolding chain_def by fast
+
+lemma chainD:
+  assumes "chain ord S" and "x \<in> S" and "y \<in> S"
+  shows "ord x y \<or> ord y x"
+using assms unfolding chain_def by fast
+
+lemma chainE:
+  assumes "chain ord S" and "x \<in> S" and "y \<in> S"
+  obtains "ord x y" | "ord y x"
+using assms unfolding chain_def by fast
+
+subsection {* Chain-complete partial orders *}
+
+text {*
+  A ccpo has a least upper bound for any chain.  In particular, the
+  empty set is a chain, so every ccpo must have a bottom element.
+*}
+
+class ccpo = order +
+  fixes lub :: "'a set \<Rightarrow> 'a"
+  assumes lub_upper: "chain (op \<le>) A \<Longrightarrow> x \<in> A \<Longrightarrow> x \<le> lub A"
+  assumes lub_least: "chain (op \<le>) A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> x \<le> z) \<Longrightarrow> lub A \<le> z"
+begin
+
+subsection {* Transfinite iteration of a function *}
+
+inductive_set iterates :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a set"
+for f :: "'a \<Rightarrow> 'a"
+where
+  step: "x \<in> iterates f \<Longrightarrow> f x \<in> iterates f"
+| lub: "chain (op \<le>) M \<Longrightarrow> \<forall>x\<in>M. x \<in> iterates f \<Longrightarrow> lub M \<in> iterates f"
+
+lemma iterates_le_f:
+  "x \<in> iterates f \<Longrightarrow> monotone (op \<le>) (op \<le>) f \<Longrightarrow> x \<le> f x"
+by (induct x rule: iterates.induct)
+  (force dest: monotoneD intro!: lub_upper lub_least)+
+
+lemma chain_iterates:
+  assumes f: "monotone (op \<le>) (op \<le>) f"
+  shows "chain (op \<le>) (iterates f)" (is "chain _ ?C")
+proof (rule chainI)
+  fix x y assume "x \<in> ?C" "y \<in> ?C"
+  then show "x \<le> y \<or> y \<le> x"
+  proof (induct x arbitrary: y rule: iterates.induct)
+    fix x y assume y: "y \<in> ?C"
+    and IH: "\<And>z. z \<in> ?C \<Longrightarrow> x \<le> z \<or> z \<le> x"
+    from y show "f x \<le> y \<or> y \<le> f x"
+    proof (induct y rule: iterates.induct)
+      case (step y) with IH f show ?case by (auto dest: monotoneD)
+    next
+      case (lub M)
+      then have chM: "chain (op \<le>) M"
+        and IH': "\<And>z. z \<in> M \<Longrightarrow> f x \<le> z \<or> z \<le> f x" by auto
+      show "f x \<le> lub M \<or> lub M \<le> f x"
+      proof (cases "\<exists>z\<in>M. f x \<le> z")
+        case True then have "f x \<le> lub M"
+          apply rule
+          apply (erule order_trans)
+          by (rule lub_upper[OF chM])
+        thus ?thesis ..
+      next
+        case False with IH'
+        show ?thesis by (auto intro: lub_least[OF chM])
+      qed
+    qed
+  next
+    case (lub M y)
+    show ?case
+    proof (cases "\<exists>x\<in>M. y \<le> x")
+      case True then have "y \<le> lub M"
+        apply rule
+        apply (erule order_trans)
+        by (rule lub_upper[OF lub(1)])
+      thus ?thesis ..
+    next
+      case False with lub
+      show ?thesis by (auto intro: lub_least)
+    qed
+  qed
+qed
+
+subsection {* Fixpoint combinator *}
+
+definition
+  fixp :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a"
+where
+  "fixp f = lub (iterates f)"
+
+lemma iterates_fixp:
+  assumes f: "monotone (op \<le>) (op \<le>) f" shows "fixp f \<in> iterates f"
+unfolding fixp_def
+by (simp add: iterates.lub chain_iterates f)
+
+lemma fixp_unfold:
+  assumes f: "monotone (op \<le>) (op \<le>) f"
+  shows "fixp f = f (fixp f)"
+proof (rule antisym)
+  show "fixp f \<le> f (fixp f)"
+    by (intro iterates_le_f iterates_fixp f)
+  have "f (fixp f) \<le> lub (iterates f)"
+    by (intro lub_upper chain_iterates f iterates.step iterates_fixp)
+  thus "f (fixp f) \<le> fixp f"
+    unfolding fixp_def .
+qed
+
+lemma fixp_lowerbound:
+  assumes f: "monotone (op \<le>) (op \<le>) f" and z: "f z \<le> z" shows "fixp f \<le> z"
+unfolding fixp_def
+proof (rule lub_least[OF chain_iterates[OF f]])
+  fix x assume "x \<in> iterates f"
+  thus "x \<le> z"
+  proof (induct x rule: iterates.induct)
+    fix x assume "x \<le> z" with f have "f x \<le> f z" by (rule monotoneD)
+    also note z finally show "f x \<le> z" .
+  qed (auto intro: lub_least)
+qed
+
+
+subsection {* Fixpoint induction *}
+
+definition
+  admissible :: "('a \<Rightarrow> bool) \<Rightarrow> bool"
+where
+  "admissible P = (\<forall>A. chain (op \<le>) A \<longrightarrow> (\<forall>x\<in>A. P x) \<longrightarrow> P (lub A))"
+
+lemma admissibleI:
+  assumes "\<And>A. chain (op \<le>) A \<Longrightarrow> \<forall>x\<in>A. P x \<Longrightarrow> P (lub A)"
+  shows "admissible P"
+using assms unfolding admissible_def by fast
+
+lemma admissibleD:
+  assumes "admissible P"
+  assumes "chain (op \<le>) A"
+  assumes "\<And>x. x \<in> A \<Longrightarrow> P x"
+  shows "P (lub A)"
+using assms by (auto simp: admissible_def)
+
+lemma fixp_induct:
+  assumes adm: "admissible P"
+  assumes mono: "monotone (op \<le>) (op \<le>) f"
+  assumes step: "\<And>x. P x \<Longrightarrow> P (f x)"
+  shows "P (fixp f)"
+unfolding fixp_def using adm chain_iterates[OF mono]
+proof (rule admissibleD)
+  fix x assume "x \<in> iterates f"
+  thus "P x"
+    by (induct rule: iterates.induct)
+      (auto intro: step admissibleD adm)
+qed
+
+lemma admissible_True: "admissible (\<lambda>x. True)"
+unfolding admissible_def by simp
+
+lemma admissible_False: "\<not> admissible (\<lambda>x. False)"
+unfolding admissible_def chain_def by simp
+
+lemma admissible_const: "admissible (\<lambda>x. t) = t"
+by (cases t, simp_all add: admissible_True admissible_False)
+
+lemma admissible_conj:
+  assumes "admissible (\<lambda>x. P x)"
+  assumes "admissible (\<lambda>x. Q x)"
+  shows "admissible (\<lambda>x. P x \<and> Q x)"
+using assms unfolding admissible_def by simp
+
+lemma admissible_all:
+  assumes "\<And>y. admissible (\<lambda>x. P x y)"
+  shows "admissible (\<lambda>x. \<forall>y. P x y)"
+using assms unfolding admissible_def by fast
+
+lemma admissible_ball:
+  assumes "\<And>y. y \<in> A \<Longrightarrow> admissible (\<lambda>x. P x y)"
+  shows "admissible (\<lambda>x. \<forall>y\<in>A. P x y)"
+using assms unfolding admissible_def by fast
+
+lemma chain_compr: "chain (op \<le>) A \<Longrightarrow> chain (op \<le>) {x \<in> A. P x}"
+unfolding chain_def by fast
+
+lemma admissible_disj_lemma:
+  assumes A: "chain (op \<le>)A"
+  assumes P: "\<forall>x\<in>A. \<exists>y\<in>A. x \<le> y \<and> P y"
+  shows "lub A = lub {x \<in> A. P x}"
+proof (rule antisym)
+  have *: "chain (op \<le>) {x \<in> A. P x}"
+    by (rule chain_compr [OF A])
+  show "lub A \<le> lub {x \<in> A. P x}"
+    apply (rule lub_least [OF A])
+    apply (drule P [rule_format], clarify)
+    apply (erule order_trans)
+    apply (simp add: lub_upper [OF *])
+    done
+  show "lub {x \<in> A. P x} \<le> lub A"
+    apply (rule lub_least [OF *])
+    apply clarify
+    apply (simp add: lub_upper [OF A])
+    done
+qed
+
+lemma admissible_disj:
+  fixes P Q :: "'a \<Rightarrow> bool"
+  assumes P: "admissible (\<lambda>x. P x)"
+  assumes Q: "admissible (\<lambda>x. Q x)"
+  shows "admissible (\<lambda>x. P x \<or> Q x)"
+proof (rule admissibleI)
+  fix A :: "'a set" assume A: "chain (op \<le>) A"
+  assume "\<forall>x\<in>A. P x \<or> Q x"
+  hence "(\<forall>x\<in>A. \<exists>y\<in>A. x \<le> y \<and> P y) \<or> (\<forall>x\<in>A. \<exists>y\<in>A. x \<le> y \<and> Q y)"
+    using chainD[OF A] by blast
+  hence "lub A = lub {x \<in> A. P x} \<or> lub A = lub {x \<in> A. Q x}"
+    using admissible_disj_lemma [OF A] by fast
+  thus "P (lub A) \<or> Q (lub A)"
+    apply (rule disjE, simp_all)
+    apply (rule disjI1, rule admissibleD [OF P chain_compr [OF A]], simp)
+    apply (rule disjI2, rule admissibleD [OF Q chain_compr [OF A]], simp)
+    done
+qed
+
+end
+
+
+
+end
--- a/src/HOL/FunDef.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/FunDef.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -5,11 +5,10 @@
 header {* Function Definitions and Termination Proofs *}
 
 theory FunDef
-imports Wellfounded
+imports Partial_Function Wellfounded
 uses
   "Tools/prop_logic.ML"
   "Tools/sat_solver.ML"
-  ("Tools/Function/function_lib.ML")
   ("Tools/Function/function_common.ML")
   ("Tools/Function/context_tree.ML")
   ("Tools/Function/function_core.ML")
@@ -101,7 +100,6 @@
   "wf R \<Longrightarrow> wfP (in_rel R)"
   by (simp add: wfP_def)
 
-use "Tools/Function/function_lib.ML"
 use "Tools/Function/function_common.ML"
 use "Tools/Function/context_tree.ML"
 use "Tools/Function/function_core.ML"
--- a/src/HOL/Imperative_HOL/Heap_Monad.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Imperative_HOL/Heap_Monad.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -594,6 +594,75 @@
 *}
 
 
+section {* Partial function definition setup *}
+
+definition "Heap_ord = img_ord execute (fun_ord option_ord)"
+definition "Heap_lub = img_lub execute Heap (fun_lub (flat_lub None))"
+
+interpretation heap!:
+  partial_function_definitions Heap_ord Heap_lub
+unfolding Heap_ord_def Heap_lub_def
+apply (rule partial_function_image)
+apply (rule partial_function_lift)
+apply (rule flat_interpretation)
+by (auto intro: Heap_eqI)
+
+abbreviation "mono_Heap \<equiv> monotone (fun_ord Heap_ord) Heap_ord"
+
+lemma Heap_ordI:
+  assumes "\<And>h. execute x h = None \<or> execute x h = execute y h"
+  shows "Heap_ord x y"
+using assms unfolding Heap_ord_def img_ord_def fun_ord_def flat_ord_def
+by blast
+
+lemma Heap_ordE:
+  assumes "Heap_ord x y"
+  obtains "execute x h = None" | "execute x h = execute y h"
+using assms unfolding Heap_ord_def img_ord_def fun_ord_def flat_ord_def
+by atomize_elim blast
+
+
+lemma bind_mono[partial_function_mono]:
+assumes mf: "mono_Heap B" and mg: "\<And>y. mono_Heap (\<lambda>f. C y f)"
+shows "mono_Heap (\<lambda>f. B f \<guillemotright>= (\<lambda>y. C y f))"
+proof (rule monotoneI)
+  fix f g :: "'a \<Rightarrow> 'b Heap" assume fg: "fun_ord Heap_ord f g"
+  from mf
+  have 1: "Heap_ord (B f) (B g)" by (rule monotoneD) (rule fg)
+  from mg
+  have 2: "\<And>y'. Heap_ord (C y' f) (C y' g)" by (rule monotoneD) (rule fg)
+
+  have "Heap_ord (B f \<guillemotright>= (\<lambda>y. C y f)) (B g \<guillemotright>= (\<lambda>y. C y f))"
+    (is "Heap_ord ?L ?R")
+  proof (rule Heap_ordI)
+    fix h
+    from 1 show "execute ?L h = None \<or> execute ?L h = execute ?R h"
+      by (rule Heap_ordE[where h = h]) (auto simp: execute_bind_case)
+  qed
+  also
+  have "Heap_ord (B g \<guillemotright>= (\<lambda>y'. C y' f)) (B g \<guillemotright>= (\<lambda>y'. C y' g))"
+    (is "Heap_ord ?L ?R")
+  proof (rule Heap_ordI)
+    fix h
+    show "execute ?L h = None \<or> execute ?L h = execute ?R h"
+    proof (cases "execute (B g) h")
+      case None
+      then have "execute ?L h = None" by (auto simp: execute_bind_case)
+      thus ?thesis ..
+    next
+      case Some
+      then obtain r h' where "execute (B g) h = Some (r, h')"
+        by (metis surjective_pairing)
+      then have "execute ?L h = execute (C r f) h'"
+        "execute ?R h = execute (C r g) h'"
+        by (auto simp: execute_bind_case)
+      with 2[of r] show ?thesis by (auto elim: Heap_ordE)
+    qed
+  qed
+  finally (heap.leq_trans)
+  show "Heap_ord (B f \<guillemotright>= (\<lambda>y. C y f)) (B g \<guillemotright>= (\<lambda>y'. C y' g))" .
+qed
+
 hide_const (open) Heap heap guard raise' fold_map
 
 end
--- a/src/HOL/Imperative_HOL/ex/Linked_Lists.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Imperative_HOL/ex/Linked_Lists.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -37,17 +37,14 @@
                                    }"
 
 
-text {* define traverse using the MREC combinator *}
-
-definition
-  traverse :: "'a\<Colon>heap node \<Rightarrow> 'a list Heap"
+partial_function (heap) traverse :: "'a\<Colon>heap node \<Rightarrow> 'a list Heap"
 where
-[code del]: "traverse = MREC (\<lambda>n. case n of Empty \<Rightarrow> return (Inl [])
-                                | Node x r \<Rightarrow> do { tl \<leftarrow> Ref.lookup r;
-                                                  return (Inr tl) })
-                   (\<lambda>n tl xs. case n of Empty \<Rightarrow> undefined
-                                      | Node x r \<Rightarrow> return (x # xs))"
-
+  [code del]: "traverse l =
+    (case l of Empty \<Rightarrow> return []
+     | Node x r \<Rightarrow> do { tl \<leftarrow> Ref.lookup r;
+                              xs \<leftarrow> traverse tl;
+                              return (x#xs)
+                         })"
 
 lemma traverse_simps[code, simp]:
   "traverse Empty      = return []"
@@ -55,8 +52,7 @@
                               xs \<leftarrow> traverse tl;
                               return (x#xs)
                          }"
-unfolding traverse_def
-by (auto simp: traverse_def MREC_rule)
+by (simp_all add: traverse.simps[of "Empty"] traverse.simps[of "Node x r"])
 
 
 section {* Proving correctness with relational abstraction *}
@@ -528,16 +524,9 @@
 
 subsection {* Definition of in-place reversal *}
 
-definition rev' :: "(('a::heap) node ref \<times> 'a node ref) \<Rightarrow> 'a node ref Heap"
-where "rev' = MREC (\<lambda>(q, p). do { v \<leftarrow> !p; (case v of Empty \<Rightarrow> (return (Inl q))
-                            | Node x next \<Rightarrow> do {
-                                    p := Node x q;
-                                    return (Inr (p, next))
-                                  })})
-             (\<lambda>x s z. return z)"
-
-lemma rev'_simps [code]:
-  "rev' (q, p) =
+partial_function (heap) rev' :: "('a::heap) node ref \<Rightarrow> 'a node ref \<Rightarrow> 'a node ref Heap"
+where
+  [code]: "rev' q p =
    do {
      v \<leftarrow> !p;
      (case v of
@@ -545,22 +534,19 @@
       | Node x next \<Rightarrow>
         do {
           p := Node x q;
-          rev' (p, next)
+          rev' p next
         })
-  }"
-  unfolding rev'_def MREC_rule[of _ _ "(q, p)"] unfolding rev'_def[symmetric]
-thm arg_cong2
-  by (auto simp add: fun_eq_iff intro: arg_cong2[where f = bind] split: node.split)
-
+    }"
+  
 primrec rev :: "('a:: heap) node \<Rightarrow> 'a node Heap" 
 where
   "rev Empty = return Empty"
-| "rev (Node x n) = do { q \<leftarrow> ref Empty; p \<leftarrow> ref (Node x n); v \<leftarrow> rev' (q, p); !v }"
+| "rev (Node x n) = do { q \<leftarrow> ref Empty; p \<leftarrow> ref (Node x n); v \<leftarrow> rev' q p; !v }"
 
 subsection {* Correctness Proof *}
 
 lemma rev'_invariant:
-  assumes "crel (rev' (q, p)) h h' v"
+  assumes "crel (rev' q p) h h' v"
   assumes "list_of' h q qs"
   assumes "list_of' h p ps"
   assumes "\<forall>qrs prs. refs_of' h q qrs \<and> refs_of' h p prs \<longrightarrow> set prs \<inter> set qrs = {}"
@@ -569,7 +555,7 @@
 proof (induct ps arbitrary: qs p q h)
   case Nil
   thus ?case
-    unfolding rev'_simps[of q p] list_of'_def
+    unfolding rev'.simps[of q p] list_of'_def
     by (auto elim!: crel_bindE crel_lookupE crel_returnE)
 next
   case (Cons x xs)
@@ -579,8 +565,8 @@
     (*and "ref_present ref h"*)
     and list_of'_ref: "list_of' h ref xs"
     unfolding list_of'_def by (cases "Ref.get h p", auto)
-  from p_is_Node Cons(2) have crel_rev': "crel (rev' (p, ref)) (Ref.set p (Node x q) h) h' v"
-    by (auto simp add: rev'_simps [of q p] elim!: crel_bindE crel_lookupE crel_updateE)
+  from p_is_Node Cons(2) have crel_rev': "crel (rev' p ref) (Ref.set p (Node x q) h) h' v"
+    by (auto simp add: rev'.simps [of q p] elim!: crel_bindE crel_lookupE crel_updateE)
   from Cons(3) obtain qrs where qrs_def: "refs_of' h q qrs" by (elim list_of'_refs_of')
   from Cons(4) obtain prs where prs_def: "refs_of' h p prs" by (elim list_of'_refs_of')
   from qrs_def prs_def Cons(5) have distinct_pointers: "set qrs \<inter> set prs = {}" by fastsimp
@@ -627,7 +613,7 @@
   with crel_rev obtain p q h1 h2 h3 v where
     init: "crel (ref Empty) h h1 q"
     "crel (ref (Node x ps)) h1 h2 p"
-    and crel_rev':"crel (rev' (q, p)) h2 h3 v"
+    and crel_rev':"crel (rev' q p) h2 h3 v"
     and lookup: "crel (!v) h3 h' r'"
     using rev.simps
     by (auto elim!: crel_bindE)
--- a/src/HOL/IsaMakefile	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/IsaMakefile	Wed Oct 27 11:11:35 2010 -0700
@@ -122,7 +122,6 @@
   $(SRC)/Tools/IsaPlanner/rw_tools.ML \
   $(SRC)/Tools/IsaPlanner/zipper.ML \
   $(SRC)/Tools/atomize_elim.ML \
-  $(SRC)/Tools/auto_solve.ML \
   $(SRC)/Tools/auto_tools.ML \
   $(SRC)/Tools/coherent.ML \
   $(SRC)/Tools/cong_tac.ML \
@@ -134,6 +133,7 @@
   $(SRC)/Tools/nbe.ML \
   $(SRC)/Tools/project_rule.ML \
   $(SRC)/Tools/quickcheck.ML \
+  $(SRC)/Tools/solve_direct.ML \
   $(SRC)/Tools/value.ML \
   HOL.thy \
   Tools/hologic.ML \
@@ -145,6 +145,7 @@
 
 PLAIN_DEPENDENCIES = $(BASE_DEPENDENCIES) \
   Complete_Lattice.thy \
+  Complete_Partial_Order.thy \
   Datatype.thy \
   Extraction.thy \
   Fields.thy \
@@ -159,6 +160,7 @@
   Nat.thy \
   Option.thy \
   Orderings.thy \
+  Partial_Function.thy \
   Plain.thy \
   Power.thy \
   Predicate.thy \
@@ -190,6 +192,7 @@
   Tools/Function/lexicographic_order.ML \
   Tools/Function/measure_functions.ML \
   Tools/Function/mutual.ML \
+  Tools/Function/partial_function.ML \
   Tools/Function/pat_completeness.ML \
   Tools/Function/pattern_split.ML \
   Tools/Function/relation.ML \
@@ -303,12 +306,15 @@
   Tools/numeral.ML \
   Tools/numeral_simprocs.ML \
   Tools/numeral_syntax.ML \
+  Tools/Predicate_Compile/core_data.ML \
+  Tools/Predicate_Compile/mode_inference.ML \
   Tools/Predicate_Compile/predicate_compile_aux.ML \
   Tools/Predicate_Compile/predicate_compile_compilations.ML \
   Tools/Predicate_Compile/predicate_compile_core.ML \
   Tools/Predicate_Compile/predicate_compile_data.ML \
   Tools/Predicate_Compile/predicate_compile_fun.ML \
   Tools/Predicate_Compile/predicate_compile.ML \
+  Tools/Predicate_Compile/predicate_compile_proof.ML \
   Tools/Predicate_Compile/predicate_compile_specialisation.ML \
   Tools/Predicate_Compile/predicate_compile_pred.ML \
   Tools/quickcheck_generators.ML \
@@ -330,20 +336,18 @@
   Tools/Sledgehammer/sledgehammer_atp_reconstruct.ML \
   Tools/Sledgehammer/sledgehammer_atp_translate.ML \
   Tools/Sledgehammer/sledgehammer_util.ML \
-  Tools/SMT/cvc3_solver.ML \
   Tools/SMT/smtlib_interface.ML \
   Tools/SMT/smt_monomorph.ML \
   Tools/SMT/smt_normalize.ML \
+  Tools/SMT/smt_setup_solvers.ML \
   Tools/SMT/smt_solver.ML \
   Tools/SMT/smt_translate.ML \
-  Tools/SMT/yices_solver.ML \
   Tools/SMT/z3_interface.ML \
   Tools/SMT/z3_model.ML \
   Tools/SMT/z3_proof_literals.ML \
   Tools/SMT/z3_proof_parser.ML \
   Tools/SMT/z3_proof_reconstruction.ML \
   Tools/SMT/z3_proof_tools.ML \
-  Tools/SMT/z3_solver.ML \
   Tools/string_code.ML \
   Tools/string_syntax.ML \
   Tools/transfer.ML \
@@ -1354,6 +1358,8 @@
   Predicate_Compile_Examples/Examples.thy				\
   Predicate_Compile_Examples/Context_Free_Grammar_Example.thy 		\
   Predicate_Compile_Examples/Hotel_Example.thy 				\
+  Predicate_Compile_Examples/Hotel_Example_Prolog.thy 			\
+  Predicate_Compile_Examples/Hotel_Example_Small_Generator.thy 		\
   Predicate_Compile_Examples/IMP_1.thy 					\
   Predicate_Compile_Examples/IMP_2.thy 					\
   Predicate_Compile_Examples/IMP_3.thy 					\
--- a/src/HOL/Library/Dlist.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Library/Dlist.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -143,11 +143,11 @@
   proof (induct xs)
     case Nil from empty show ?case by (simp add: empty_def)
   next
-    case (insert x xs)
+    case (Cons x xs)
     then have "\<not> member (Dlist xs) x" and "P (Dlist xs)"
       by (simp_all add: member_def List.member_def)
     with insrt have "P (insert x (Dlist xs))" .
-    with insert show ?case by (simp add: insert_def distinct_remdups_id)
+    with Cons show ?case by (simp add: insert_def distinct_remdups_id)
   qed
   with dxs show "P dxs" by simp
 qed
--- a/src/HOL/Library/Multiset.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Library/Multiset.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -837,11 +837,11 @@
 context linorder
 begin
 
-lemma multiset_of_insort_key [simp]:
+lemma multiset_of_insort [simp]:
   "multiset_of (insort_key k x xs) = {#x#} + multiset_of xs"
   by (induct xs) (simp_all add: ac_simps)
  
-lemma multiset_of_sort_key [simp]:
+lemma multiset_of_sort [simp]:
   "multiset_of (sort_key k xs) = multiset_of xs"
   by (induct xs) (simp_all add: ac_simps)
 
--- a/src/HOL/Library/Permutation.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Library/Permutation.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -168,7 +168,7 @@
     apply simp
     apply (subgoal_tac "a#list <~~> a#ysa@zs")
      apply (metis Cons_eq_appendI perm_append_Cons trans)
-    apply (metis Cons Cons_eq_appendI distinct_simps(2)
+    apply (metis Cons Cons_eq_appendI distinct.simps(2)
       distinct_remdups distinct_remdups_id perm_append_swap perm_distinct_iff)
    apply (subgoal_tac "set (a#list) = set (ysa@a#zs) & distinct (a#list) & distinct (ysa@a#zs)")
     apply (fastsimp simp add: insert_ident)
--- a/src/HOL/List.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/List.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -5,7 +5,7 @@
 header {* The datatype of finite lists *}
 
 theory List
-imports Plain Quotient Presburger Code_Numeral Recdef
+imports Plain Presburger Recdef Code_Numeral Quotient ATP
 uses ("Tools/list_code.ML")
 begin
 
@@ -174,15 +174,10 @@
     "removeAll x [] = []"
   | "removeAll x (y # xs) = (if x = y then removeAll x xs else y # removeAll x xs)"
 
-inductive
+primrec
   distinct :: "'a list \<Rightarrow> bool" where
-    Nil: "distinct []"
-  | insert: "x \<notin> set xs \<Longrightarrow> distinct xs \<Longrightarrow> distinct (x # xs)"
-
-lemma distinct_simps [simp, code]:
-  "distinct [] \<longleftrightarrow> True"
-  "distinct (x # xs) \<longleftrightarrow> x \<notin> set xs \<and> distinct xs"
-  by (auto intro: distinct.intros elim: distinct.cases)
+    "distinct [] \<longleftrightarrow> True"
+  | "distinct (x # xs) \<longleftrightarrow> x \<notin> set xs \<and> distinct xs"
 
 primrec
   remdups :: "'a list \<Rightarrow> 'a list" where
@@ -308,11 +303,12 @@
 definition sort_key :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b list \<Rightarrow> 'b list" where
 "sort_key f xs = foldr (insort_key f) xs []"
 
+definition insort_insert_key :: "('b \<Rightarrow> 'a) \<Rightarrow> 'b \<Rightarrow> 'b list \<Rightarrow> 'b list" where
+  "insort_insert_key f x xs = (if f x \<in> f ` set xs then xs else insort_key f x xs)"
+
 abbreviation "sort \<equiv> sort_key (\<lambda>x. x)"
 abbreviation "insort \<equiv> insort_key (\<lambda>x. x)"
-
-definition insort_insert :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
-  "insort_insert x xs = (if x \<in> set xs then xs else insort x xs)"
+abbreviation "insort_insert \<equiv> insort_insert_key (\<lambda>x. x)"
 
 end
 
@@ -762,6 +758,14 @@
 
 subsubsection {* @{text map} *}
 
+lemma hd_map:
+  "xs \<noteq> [] \<Longrightarrow> hd (map f xs) = f (hd xs)"
+  by (cases xs) simp_all
+
+lemma map_tl:
+  "map f (tl xs) = tl (map f xs)"
+  by (cases xs) simp_all
+
 lemma map_ext: "(!!x. x : set xs --> f x = g x) ==> map f xs = map g xs"
 by (induct xs) simp_all
 
@@ -786,9 +790,8 @@
 by (induct xs) auto
 
 lemma map_cong [fundef_cong, recdef_cong]:
-"xs = ys ==> (!!x. x : set ys ==> f x = g x) ==> map f xs = map g ys"
--- {* a congruence rule for @{text map} *}
-by simp
+  "xs = ys \<Longrightarrow> (\<And>x. x \<in> set ys \<Longrightarrow> f x = g x) \<Longrightarrow> map f xs = map g ys"
+  by simp
 
 lemma map_is_Nil_conv [iff]: "(map f xs = []) = (xs = [])"
 by (cases xs) auto
@@ -990,14 +993,14 @@
   "(x : set xs) = (\<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> set ys)"
   by (auto dest!: split_list_first)
 
-lemma split_list_last: "x : set xs \<Longrightarrow> \<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> set zs"
-proof (induct xs rule:rev_induct)
+lemma split_list_last: "x \<in> set xs \<Longrightarrow> \<exists>ys zs. xs = ys @ x # zs \<and> x \<notin> set zs"
+proof (induct xs rule: rev_induct)
   case Nil thus ?case by simp
 next
   case (snoc a xs)
   show ?case
   proof cases
-    assume "x = a" thus ?case using snoc by simp (metis ex_in_conv set_empty2)
+    assume "x = a" thus ?case using snoc by (metis List.set.simps(1) emptyE)
   next
     assume "x \<noteq> a" thus ?case using snoc by fastsimp
   qed
@@ -1030,8 +1033,7 @@
   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)
+    thus ?thesis by simp (metis Un_upper1 contra_subsetD in_set_conv_decomp_first self_append_conv2 set_append)
   next
     assume "\<not> P x"
     hence "\<exists>x\<in>set xs. P x" using Cons(2) by simp
@@ -2676,6 +2678,10 @@
 
 subsubsection {* @{text "distinct"} and @{text remdups} *}
 
+lemma distinct_tl:
+  "distinct xs \<Longrightarrow> distinct (tl xs)"
+  by (cases xs) simp_all
+
 lemma distinct_append [simp]:
 "distinct (xs @ ys) = (distinct xs \<and> distinct ys \<and> set xs \<inter> set ys = {})"
 by (induct xs) auto
@@ -3636,12 +3642,18 @@
 context linorder
 begin
 
-lemma length_insert[simp] : "length (insort_key f x xs) = Suc (length xs)"
-by (induct xs, auto)
+lemma length_insort [simp]:
+  "length (insort_key f x xs) = Suc (length xs)"
+  by (induct xs) simp_all
+
+lemma insort_key_left_comm:
+  assumes "f x \<noteq> f y"
+  shows "insort_key f y (insort_key f x xs) = insort_key f x (insort_key f y xs)"
+  by (induct xs) (auto simp add: assms dest: antisym)
 
 lemma insort_left_comm:
   "insort x (insort y xs) = insort y (insort x xs)"
-  by (induct xs) auto
+  by (cases "x = y") (auto intro: insort_key_left_comm)
 
 lemma fun_left_comm_insort:
   "fun_left_comm insort"
@@ -3664,6 +3676,10 @@
 apply(induct xs arbitrary: x) apply simp
 by simp (blast intro: order_trans)
 
+lemma sorted_tl:
+  "sorted xs \<Longrightarrow> sorted (tl xs)"
+  by (cases xs) (simp_all add: sorted_Cons)
+
 lemma sorted_append:
   "sorted (xs@ys) = (sorted xs & sorted ys & (\<forall>x \<in> set xs. \<forall>y \<in> set ys. x\<le>y))"
 by (induct xs) (auto simp add:sorted_Cons)
@@ -3719,16 +3735,16 @@
 by(induct xs)(simp_all add:distinct_insort set_sort)
 
 lemma sorted_insort_key: "sorted (map f (insort_key f x xs)) = sorted (map f xs)"
-by(induct xs)(auto simp:sorted_Cons set_insort)
+  by (induct xs) (auto simp:sorted_Cons set_insort)
 
 lemma sorted_insort: "sorted (insort x xs) = sorted xs"
-using sorted_insort_key[where f="\<lambda>x. x"] by simp
-
-theorem sorted_sort_key[simp]: "sorted (map f (sort_key f xs))"
-by(induct xs)(auto simp:sorted_insort_key)
-
-theorem sorted_sort[simp]: "sorted (sort xs)"
-by(induct xs)(auto simp:sorted_insort)
+  using sorted_insort_key [where f="\<lambda>x. x"] by simp
+
+theorem sorted_sort_key [simp]: "sorted (map f (sort_key f xs))"
+  by (induct xs) (auto simp:sorted_insort_key)
+
+theorem sorted_sort [simp]: "sorted (sort xs)"
+  using sorted_sort_key [where f="\<lambda>x. x"] by simp
 
 lemma sorted_butlast:
   assumes "xs \<noteq> []" and "sorted xs"
@@ -3877,32 +3893,53 @@
   finally show "\<not> t < f x" by simp
 qed
 
+lemma insort_insert_key_triv:
+  "f x \<in> f ` set xs \<Longrightarrow> insort_insert_key f x xs = xs"
+  by (simp add: insort_insert_key_def)
+
+lemma insort_insert_triv:
+  "x \<in> set xs \<Longrightarrow> insort_insert x xs = xs"
+  using insort_insert_key_triv [of "\<lambda>x. x"] by simp
+
+lemma insort_insert_insort_key:
+  "f x \<notin> f ` set xs \<Longrightarrow> insort_insert_key f x xs = insort_key f x xs"
+  by (simp add: insort_insert_key_def)
+
+lemma insort_insert_insort:
+  "x \<notin> set xs \<Longrightarrow> insort_insert x xs = insort x xs"
+  using insort_insert_insort_key [of "\<lambda>x. x"] by simp
+
 lemma set_insort_insert:
   "set (insort_insert x xs) = insert x (set xs)"
-  by (auto simp add: insort_insert_def set_insort)
+  by (auto simp add: insort_insert_key_def set_insort)
 
 lemma distinct_insort_insert:
   assumes "distinct xs"
-  shows "distinct (insort_insert x xs)"
-  using assms by (induct xs) (auto simp add: insort_insert_def set_insort)
+  shows "distinct (insort_insert_key f x xs)"
+  using assms by (induct xs) (auto simp add: insort_insert_key_def set_insort)
+
+lemma sorted_insort_insert_key:
+  assumes "sorted (map f xs)"
+  shows "sorted (map f (insort_insert_key f x xs))"
+  using assms by (simp add: insort_insert_key_def sorted_insort_key)
 
 lemma sorted_insort_insert:
   assumes "sorted xs"
   shows "sorted (insort_insert x xs)"
-  using assms by (simp add: insort_insert_def sorted_insort)
-
-lemma filter_insort_key_triv:
+  using assms sorted_insort_insert_key [of "\<lambda>x. x"] by simp
+
+lemma filter_insort_triv:
   "\<not> P x \<Longrightarrow> filter P (insort_key f x xs) = filter P xs"
   by (induct xs) simp_all
 
-lemma filter_insort_key:
+lemma filter_insort:
   "sorted (map f xs) \<Longrightarrow> P x \<Longrightarrow> filter P (insort_key f x xs) = insort_key f x (filter P xs)"
   using assms by (induct xs)
     (auto simp add: sorted_Cons, subst insort_is_Cons, auto)
 
-lemma filter_sort_key:
+lemma filter_sort:
   "filter P (sort_key f xs) = sort_key f (filter P xs)"
-  by (induct xs) (simp_all add: filter_insort_key_triv filter_insort_key)
+  by (induct xs) (simp_all add: filter_insort_triv filter_insort)
 
 lemma sorted_same [simp]:
   "sorted [x\<leftarrow>xs. x = f xs]"
--- a/src/HOL/Main.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Main.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -1,7 +1,7 @@
 header {* Main HOL *}
 
 theory Main
-imports Plain Record Predicate_Compile Nitpick SMT
+imports Plain Record Predicate_Compile Nitpick
 begin
 
 text {*
--- a/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -357,25 +357,25 @@
     val relevance_fudge = Sledgehammer.relevance_fudge_for_prover prover_name
     val relevance_override = {add = [], del = [], only = false}
     val (_, hyp_ts, concl_t) = Sledgehammer_Util.strip_subgoal goal i
-    val axioms =
+    val facts =
       Sledgehammer_Filter.relevant_facts ctxt full_types relevance_thresholds
           (the_default default_max_relevant max_relevant) irrelevant_consts
           relevance_fudge relevance_override chained_ths hyp_ts concl_t
     val problem =
       {state = st', goal = goal, subgoal = i,
        subgoal_count = Sledgehammer_Util.subgoal_count st,
-       axioms = axioms |> map Sledgehammer.Unprepared, only = true}
+       facts = facts |> map Sledgehammer.Untranslated, only = true}
     val time_limit =
       (case hard_timeout of
         NONE => I
       | SOME secs => TimeLimit.timeLimit (Time.fromSeconds secs))
-    val ({outcome, message, used_axioms, run_time_in_msecs = SOME time_prover, ...}
+    val ({outcome, message, used_facts, run_time_in_msecs = SOME time_prover, ...}
          : Sledgehammer.prover_result,
         time_isa) = time_limit (Mirabelle.cpu_time
                       (prover params (K ""))) problem
   in
     case outcome of
-      NONE => (message, SH_OK (time_isa, time_prover, used_axioms))
+      NONE => (message, SH_OK (time_isa, time_prover, used_facts))
     | SOME _ => (message, SH_FAIL (time_isa, time_prover))
   end
   handle ERROR msg => ("error: " ^ msg, SH_ERROR)
--- a/src/HOL/Multivariate_Analysis/Integration.certs	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Multivariate_Analysis/Integration.certs	Wed Oct 27 11:11:35 2010 -0700
@@ -1,1976 +1,243 @@
-2e3a7f41999e849037c0bc39fd6c0ffa4f5b7eb6 910 0
+19e9a49723add29f06290682aca290aa2ab968ea 238 0
 #2 := false
-#369 := 0::real
-decl f8 :: (-> S5 S2 real)
-decl f11 :: (-> S4 S2)
-decl f7 :: S4
-#13 := f7
-#20 := (f11 f7)
-decl f21 :: (-> S3 S5)
+#48 := 0::real
+decl f19 :: (-> S3 S10 real)
+decl f20 :: S10
+#43 := f20
 decl f4 :: S3
 #8 := f4
-#101 := (f21 f4)
-#1205 := (f8 #101 #20)
-#367 := -1::real
-#1507 := (* -1::real #1205)
-decl f19 :: S3
-#60 := f19
-#88 := (f21 f19)
-#1037 := (f8 #88 #20)
-#1698 := (+ #1037 #1507)
-#1765 := (>= #1698 0::real)
-#1697 := (= #1037 #1205)
-decl f10 :: S5
-#19 := f10
-#21 := (f8 f10 #20)
-#1208 := (= #21 #1205)
-decl f12 :: S5
-#22 := f12
-#1129 := (f8 f12 #20)
-#1207 := (= #1129 #1205)
-decl f6 :: (-> S2 S4)
-#84 := (f6 #20)
-#351 := (= f7 #84)
-#1211 := (ite #351 #1208 #1207)
-decl f9 :: S5
-#16 := f9
-#31 := (f8 f9 #20)
-#1206 := (= #31 #1205)
-#70 := 0::int
-decl f5 :: (-> S4 int)
-#1041 := (f5 #84)
-#155 := -1::int
-#1051 := (* -1::int #1041)
-#14 := (f5 f7)
-#1097 := (+ #14 #1051)
-#1098 := (<= #1097 0::int)
-#1214 := (ite #1098 #1211 #1206)
-#9 := (:var 0 S2)
-#17 := (f8 f9 #9)
-#697 := (pattern #17)
-#23 := (f8 f12 #9)
-#696 := (pattern #23)
-#102 := (f8 #101 #9)
-#695 := (pattern #102)
-#11 := (f6 #9)
-#694 := (pattern #11)
-#580 := (= #17 #102)
-#578 := (= #23 #102)
-#577 := (= #21 #102)
-#18 := (= #11 f7)
-#579 := (ite #18 #577 #578)
-#158 := (* -1::int #14)
-#12 := (f5 #11)
-#159 := (+ #12 #158)
-#157 := (>= #159 0::int)
-#581 := (ite #157 #579 #580)
-#698 := (forall (vars (?v0 S2)) (:pat #694 #695 #696 #697) #581)
-#584 := (forall (vars (?v0 S2)) #581)
-#701 := (iff #584 #698)
-#699 := (iff #581 #581)
-#700 := [refl]: #699
-#702 := [quant-intro #700]: #701
-#24 := (ite #18 #21 #23)
-#165 := (ite #157 #24 #17)
-#486 := (= #102 #165)
-#487 := (forall (vars (?v0 S2)) #486)
-#585 := (iff #487 #584)
-#582 := (iff #486 #581)
-#583 := [rewrite]: #582
-#586 := [quant-intro #583]: #585
-#480 := (~ #487 #487)
-#482 := (~ #486 #486)
-#483 := [refl]: #482
-#481 := [nnf-pos #483]: #480
-decl f3 :: (-> S3 S2 real)
-#10 := (f3 f4 #9)
-#170 := (= #10 #165)
-#173 := (forall (vars (?v0 S2)) #170)
-#488 := (iff #173 #487)
-#130 := (:var 1 S3)
-#133 := (f3 #130 #9)
-#131 := (f21 #130)
-#132 := (f8 #131 #9)
-#134 := (= #132 #133)
-#135 := (forall (vars (?v0 S3) (?v1 S2)) #134)
-#441 := [asserted]: #135
-#489 := [rewrite* #441]: #488
-#15 := (< #12 #14)
-#25 := (ite #15 #17 #24)
-#26 := (= #10 #25)
-#27 := (forall (vars (?v0 S2)) #26)
-#174 := (iff #27 #173)
-#171 := (iff #26 #170)
-#168 := (= #25 #165)
-#156 := (not #157)
-#162 := (ite #156 #17 #24)
+#58 := (f19 f4 f20)
+#109 := -1::real
+#153 := (* -1::real #58)
+decl f5 :: (-> S4 S5 real)
+decl f8 :: S5
+#13 := f8
+decl f22 :: S4
+#54 := f22
+#55 := (f5 f22 f8)
+#154 := (+ #55 #153)
+#137 := (* -1::real #55)
+#144 := (+ #137 #58)
+#188 := (<= #154 0::real)
+#195 := (ite #188 #144 #154)
+#451 := (* -1::real #195)
+#452 := (+ #144 #451)
+#453 := (<= #452 0::real)
+#435 := (= #144 #195)
+decl f21 :: S4
+#45 := f21
+#46 := (f5 f21 f8)
+decl f9 :: S3
+#17 := f9
+#44 := (f19 f9 f20)
+#120 := (* -1::real #44)
+#121 := (+ #120 #46)
+#110 := (* -1::real #46)
+#111 := (+ #44 #110)
+#216 := (>= #111 0::real)
+#223 := (ite #216 #111 #121)
+#447 := (* -1::real #223)
+#450 := (+ #121 #447)
+#454 := (<= #450 0::real)
+#442 := (= #121 #223)
+#217 := (not #216)
+#455 := [hypothesis]: #216
+#184 := (+ #44 #153)
+#185 := (<= #184 0::real)
+#206 := -3::real
+#234 := (* -3::real #223)
+#235 := (+ #137 #234)
+#236 := (+ #46 #235)
+#237 := (<= #236 0::real)
+#238 := (not #237)
+#207 := (* -3::real #195)
+#208 := (+ #137 #207)
+#209 := (+ #46 #208)
+#210 := (<= #209 0::real)
+#211 := (not #210)
+#249 := (and #185 #211 #238)
+#65 := (<= #44 #58)
+#56 := (- #46 #55)
+#52 := 3::real
+#59 := (- #58 #55)
+#61 := (- #59)
+#60 := (< #59 0::real)
+#62 := (ite #60 #61 #59)
+#63 := (* #62 3::real)
+#64 := (< #63 #56)
+#66 := (and #64 #65)
+#47 := (- #44 #46)
+#50 := (- #47)
+#49 := (< #47 0::real)
+#51 := (ite #49 #50 #47)
+#53 := (* #51 3::real)
+#57 := (< #53 #56)
+#67 := (and #57 #66)
+#254 := (iff #67 #249)
+#138 := (+ #46 #137)
+#147 := (< #144 0::real)
+#159 := (ite #147 #154 #144)
+#165 := (* 3::real #159)
+#170 := (< #165 #138)
+#176 := (and #65 #170)
+#114 := (< #111 0::real)
+#126 := (ite #114 #121 #111)
+#132 := (* 3::real #126)
+#141 := (< #132 #138)
+#181 := (and #141 #176)
+#252 := (iff #181 #249)
+#243 := (and #185 #211)
+#246 := (and #238 #243)
+#250 := (iff #246 #249)
+#251 := [rewrite]: #250
+#247 := (iff #181 #246)
+#244 := (iff #176 #243)
+#214 := (iff #170 #211)
+#200 := (* 3::real #195)
+#203 := (< #200 #138)
+#212 := (iff #203 #211)
+#213 := [rewrite]: #212
+#204 := (iff #170 #203)
+#201 := (= #165 #200)
+#198 := (= #159 #195)
+#189 := (not #188)
+#192 := (ite #189 #154 #144)
+#196 := (= #192 #195)
+#197 := [rewrite]: #196
+#193 := (= #159 #192)
+#190 := (iff #147 #189)
+#191 := [rewrite]: #190
+#194 := [monotonicity #191]: #193
+#199 := [trans #194 #197]: #198
+#202 := [monotonicity #199]: #201
+#205 := [monotonicity #202]: #204
+#215 := [trans #205 #213]: #214
+#186 := (iff #65 #185)
+#187 := [rewrite]: #186
+#245 := [monotonicity #187 #215]: #244
+#241 := (iff #141 #238)
+#228 := (* 3::real #223)
+#231 := (< #228 #138)
+#239 := (iff #231 #238)
+#240 := [rewrite]: #239
+#232 := (iff #141 #231)
+#229 := (= #132 #228)
+#226 := (= #126 #223)
+#220 := (ite #217 #121 #111)
+#224 := (= #220 #223)
+#225 := [rewrite]: #224
+#221 := (= #126 #220)
+#218 := (iff #114 #217)
+#219 := [rewrite]: #218
+#222 := [monotonicity #219]: #221
+#227 := [trans #222 #225]: #226
+#230 := [monotonicity #227]: #229
+#233 := [monotonicity #230]: #232
+#242 := [trans #233 #240]: #241
+#248 := [monotonicity #242 #245]: #247
+#253 := [trans #248 #251]: #252
+#182 := (iff #67 #181)
+#179 := (iff #66 #176)
+#173 := (and #170 #65)
+#177 := (iff #173 #176)
+#178 := [rewrite]: #177
+#174 := (iff #66 #173)
+#171 := (iff #64 #170)
+#139 := (= #56 #138)
+#140 := [rewrite]: #139
+#168 := (= #63 #165)
+#162 := (* #159 3::real)
 #166 := (= #162 #165)
 #167 := [rewrite]: #166
-#163 := (= #25 #162)
-#160 := (iff #15 #156)
-#161 := [rewrite]: #160
+#163 := (= #63 #162)
+#160 := (= #62 #159)
+#145 := (= #59 #144)
+#146 := [rewrite]: #145
+#157 := (= #61 #154)
+#150 := (- #144)
+#155 := (= #150 #154)
+#156 := [rewrite]: #155
+#151 := (= #61 #150)
+#152 := [monotonicity #146]: #151
+#158 := [trans #152 #156]: #157
+#148 := (iff #60 #147)
+#149 := [monotonicity #146]: #148
+#161 := [monotonicity #149 #158 #146]: #160
 #164 := [monotonicity #161]: #163
 #169 := [trans #164 #167]: #168
-#172 := [monotonicity #169]: #171
-#175 := [quant-intro #172]: #174
-#152 := [asserted]: #27
-#176 := [mp #152 #175]: #173
-#490 := [mp #176 #489]: #487
-#478 := [mp~ #490 #481]: #487
-#587 := [mp #478 #586]: #584
-#703 := [mp #587 #702]: #698
-#962 := (not #698)
-#1217 := (or #962 #1214)
-#85 := (= #84 f7)
-#1209 := (ite #85 #1208 #1207)
-#1088 := (+ #1041 #158)
-#1089 := (>= #1088 0::int)
-#1210 := (ite #1089 #1209 #1206)
-#1218 := (or #962 #1210)
-#1220 := (iff #1218 #1217)
-#1222 := (iff #1217 #1217)
-#1223 := [rewrite]: #1222
-#1215 := (iff #1210 #1214)
-#1212 := (iff #1209 #1211)
-#353 := (iff #85 #351)
-#354 := [rewrite]: #353
-#1213 := [monotonicity #354]: #1212
-#1101 := (iff #1089 #1098)
-#1091 := (+ #158 #1041)
-#1094 := (>= #1091 0::int)
-#1099 := (iff #1094 #1098)
-#1100 := [rewrite]: #1099
-#1095 := (iff #1089 #1094)
-#1092 := (= #1088 #1091)
-#1093 := [rewrite]: #1092
-#1096 := [monotonicity #1093]: #1095
-#1102 := [trans #1096 #1100]: #1101
-#1216 := [monotonicity #1102 #1213]: #1215
-#1221 := [monotonicity #1216]: #1220
-#1224 := [trans #1221 #1223]: #1220
-#1219 := [quant-inst]: #1218
-#1225 := [mp #1219 #1224]: #1217
-#1767 := [unit-resolution #1225 #703]: #1214
-#1396 := (= #14 #1041)
-#1718 := (= #1041 #14)
-#350 := [asserted]: #85
-#357 := [mp #350 #354]: #351
-#1717 := [symm #357]: #85
-#1719 := [monotonicity #1717]: #1718
-#1749 := [symm #1719]: #1396
-#1750 := (not #1396)
-#1768 := (or #1750 #1098)
-#1769 := [th-lemma]: #1768
-#1770 := [unit-resolution #1769 #1749]: #1098
-#1116 := (not #1098)
-#1238 := (not #1214)
-#1239 := (or #1238 #1116 #1211)
-#1240 := [def-axiom]: #1239
-#1771 := [unit-resolution #1240 #1770 #1767]: #1211
-#1226 := (not #1211)
-#1772 := (or #1226 #1208)
-#1227 := (not #351)
-#1228 := (or #1226 #1227 #1208)
-#1229 := [def-axiom]: #1228
-#1773 := [unit-resolution #1229 #357]: #1772
-#1774 := [unit-resolution #1773 #1771]: #1208
-#1811 := (= #1037 #21)
-#1038 := (= #21 #1037)
-decl f16 :: S4
-#40 := f16
-#41 := (f5 f16)
-#1052 := (+ #41 #1051)
-#1053 := (<= #1052 0::int)
-#1074 := (not #1053)
-#196 := (* -1::int #41)
-#1665 := (+ #14 #196)
-#1666 := (>= #1665 0::int)
-#1754 := (not #1666)
-#1679 := (<= #1665 0::int)
-#44 := (f11 f16)
-#82 := (f6 #44)
-#778 := (f5 #82)
-#788 := (* -1::int #778)
-#835 := (+ #14 #788)
-#836 := (<= #835 0::int)
-decl f18 :: S3
-#55 := f18
-#98 := (f21 f18)
-#823 := (f8 #98 #44)
-#45 := (f8 f10 #44)
-#824 := (= #45 #823)
-#863 := (not #824)
-decl f15 :: S3
-#38 := f15
-#93 := (f21 f15)
-#902 := (f8 #93 #44)
-#1699 := (= #823 #902)
-#1759 := (not #1699)
-#1793 := (iff #1759 #863)
-#1791 := (iff #1699 #824)
-#1786 := (= #823 #45)
-#1789 := (iff #1786 #824)
-#1790 := [commutativity]: #1789
-#1787 := (iff #1699 #1786)
-#1784 := (= #902 #45)
-#905 := (= #45 #902)
-#869 := (f8 f12 #44)
-#904 := (= #869 #902)
-#347 := (= f16 #82)
-#908 := (ite #347 #905 #904)
-#867 := (f8 f9 #44)
-#903 := (= #867 #902)
-#789 := (+ #41 #788)
-#790 := (<= #789 0::int)
-#911 := (ite #790 #908 #903)
-#94 := (f8 #93 #9)
-#713 := (pattern #94)
-#602 := (= #17 #94)
-#600 := (= #23 #94)
-#599 := (= #45 #94)
-#43 := (= #11 f16)
-#601 := (ite #43 #599 #600)
-#197 := (+ #12 #196)
-#195 := (>= #197 0::int)
-#603 := (ite #195 #601 #602)
-#714 := (forall (vars (?v0 S2)) (:pat #694 #713 #696 #697) #603)
-#606 := (forall (vars (?v0 S2)) #603)
-#717 := (iff #606 #714)
-#715 := (iff #603 #603)
-#716 := [refl]: #715
-#718 := [quant-intro #716]: #717
-#46 := (ite #43 #45 #23)
-#203 := (ite #195 #46 #17)
-#497 := (= #94 #203)
-#498 := (forall (vars (?v0 S2)) #497)
-#607 := (iff #498 #606)
-#604 := (iff #497 #603)
-#605 := [rewrite]: #604
-#608 := [quant-intro #605]: #607
-#470 := (~ #498 #498)
-#472 := (~ #497 #497)
-#473 := [refl]: #472
-#471 := [nnf-pos #473]: #470
-#39 := (f3 f15 #9)
-#208 := (= #39 #203)
-#211 := (forall (vars (?v0 S2)) #208)
-#499 := (iff #211 #498)
-#500 := [rewrite* #441]: #499
-#42 := (< #12 #41)
-#47 := (ite #42 #17 #46)
-#48 := (= #39 #47)
-#49 := (forall (vars (?v0 S2)) #48)
-#212 := (iff #49 #211)
-#209 := (iff #48 #208)
-#206 := (= #47 #203)
-#194 := (not #195)
-#200 := (ite #194 #17 #46)
-#204 := (= #200 #203)
-#205 := [rewrite]: #204
-#201 := (= #47 #200)
-#198 := (iff #42 #194)
-#199 := [rewrite]: #198
-#202 := [monotonicity #199]: #201
-#207 := [trans #202 #205]: #206
-#210 := [monotonicity #207]: #209
-#213 := [quant-intro #210]: #212
-#154 := [asserted]: #49
-#214 := [mp #154 #213]: #211
-#501 := [mp #214 #500]: #498
-#468 := [mp~ #501 #471]: #498
-#609 := [mp #468 #608]: #606
-#719 := [mp #609 #718]: #714
-#914 := (not #714)
-#915 := (or #914 #911)
-#83 := (= #82 f16)
-#906 := (ite #83 #905 #904)
-#779 := (+ #778 #196)
-#780 := (>= #779 0::int)
-#907 := (ite #780 #906 #903)
-#916 := (or #914 #907)
-#918 := (iff #916 #915)
-#920 := (iff #915 #915)
-#921 := [rewrite]: #920
-#912 := (iff #907 #911)
-#909 := (iff #906 #908)
-#348 := (iff #83 #347)
-#349 := [rewrite]: #348
-#910 := [monotonicity #349]: #909
-#793 := (iff #780 #790)
-#782 := (+ #196 #778)
-#785 := (>= #782 0::int)
-#791 := (iff #785 #790)
-#792 := [rewrite]: #791
-#786 := (iff #780 #785)
-#783 := (= #779 #782)
-#784 := [rewrite]: #783
-#787 := [monotonicity #784]: #786
-#794 := [trans #787 #792]: #793
-#913 := [monotonicity #794 #910]: #912
-#919 := [monotonicity #913]: #918
-#922 := [trans #919 #921]: #918
-#917 := [quant-inst]: #916
-#923 := [mp #917 #922]: #915
-#1776 := [unit-resolution #923 #719]: #911
-#1394 := (= #41 #778)
-#1674 := (= #778 #41)
-#346 := [asserted]: #83
-#352 := [mp #346 #349]: #347
-#1673 := [symm #352]: #83
-#1675 := [monotonicity #1673]: #1674
-#1676 := [symm #1675]: #1394
-#1739 := (not #1394)
-#1777 := (or #1739 #790)
-#1778 := [th-lemma]: #1777
-#1779 := [unit-resolution #1778 #1676]: #790
-#812 := (not #790)
-#936 := (not #911)
-#937 := (or #936 #812 #908)
-#938 := [def-axiom]: #937
-#1780 := [unit-resolution #938 #1779 #1776]: #908
-#924 := (not #908)
-#1781 := (or #924 #905)
-#925 := (not #347)
-#926 := (or #924 #925 #905)
-#927 := [def-axiom]: #926
-#1782 := [unit-resolution #927 #352]: #1781
-#1783 := [unit-resolution #1782 #1780]: #905
-#1785 := [symm #1783]: #1784
-#1788 := [monotonicity #1785]: #1787
-#1792 := [trans #1788 #1790]: #1791
-#1794 := [monotonicity #1792]: #1793
-#1462 := (* -1::real #902)
-#1709 := (+ #823 #1462)
-#1711 := (>= #1709 0::real)
-#1708 := (not #1711)
-decl f22 :: S5
-#90 := f22
-#1463 := (f8 f22 #44)
-#1466 := (* -1::real #1463)
-#1477 := (+ #902 #1466)
-#1478 := (<= #1477 0::real)
-#1502 := (not #1478)
-#774 := (f8 #88 #44)
-#1467 := (+ #774 #1466)
-#1468 := (>= #1467 0::real)
-#1483 := (or #1468 #1478)
-#1486 := (not #1483)
-#91 := (f8 f22 #9)
-#761 := (pattern #91)
-#89 := (f8 #88 #9)
-#734 := (pattern #89)
-#376 := (* -1::real #94)
-#377 := (+ #91 #376)
-#375 := (>= #377 0::real)
-#371 := (* -1::real #91)
-#372 := (+ #89 #371)
-#370 := (>= #372 0::real)
-#561 := (or #370 #375)
-#562 := (not #561)
-#762 := (forall (vars (?v0 S2)) (:pat #734 #761 #713) #562)
-#565 := (forall (vars (?v0 S2)) #562)
-#765 := (iff #565 #762)
-#763 := (iff #562 #562)
-#764 := [refl]: #763
-#766 := [quant-intro #764]: #765
-#378 := (not #375)
-#368 := (not #370)
-#381 := (and #368 #378)
-#384 := (forall (vars (?v0 S2)) #381)
-#566 := (iff #384 #565)
-#563 := (iff #381 #562)
-#564 := [rewrite]: #563
-#567 := [quant-intro #564]: #566
-#553 := (~ #384 #384)
-#551 := (~ #381 #381)
-#552 := [refl]: #551
-#554 := [nnf-pos #552]: #553
-#394 := (* -1::real #102)
-#395 := (+ #91 #394)
-#393 := (>= #395 0::real)
-#396 := (not #393)
-#99 := (f8 #98 #9)
-#387 := (* -1::real #99)
-#388 := (+ #91 #387)
-#389 := (<= #388 0::real)
-#390 := (not #389)
-#399 := (and #390 #396)
-#402 := (forall (vars (?v0 S2)) #399)
-#405 := (and #384 #402)
-#103 := (< #91 #102)
-#100 := (< #99 #91)
-#104 := (and #100 #103)
-#105 := (forall (vars (?v0 S2)) #104)
-#95 := (< #91 #94)
-#92 := (< #89 #91)
-#96 := (and #92 #95)
-#97 := (forall (vars (?v0 S2)) #96)
-#106 := (and #97 #105)
-#406 := (iff #106 #405)
-#403 := (iff #105 #402)
-#400 := (iff #104 #399)
-#397 := (iff #103 #396)
-#398 := [rewrite]: #397
-#391 := (iff #100 #390)
-#392 := [rewrite]: #391
-#401 := [monotonicity #392 #398]: #400
-#404 := [quant-intro #401]: #403
-#385 := (iff #97 #384)
-#382 := (iff #96 #381)
-#379 := (iff #95 #378)
-#380 := [rewrite]: #379
-#373 := (iff #92 #368)
-#374 := [rewrite]: #373
-#383 := [monotonicity #374 #380]: #382
-#386 := [quant-intro #383]: #385
-#407 := [monotonicity #386 #404]: #406
-#363 := [asserted]: #106
-#408 := [mp #363 #407]: #405
-#409 := [and-elim #408]: #384
-#555 := [mp~ #409 #554]: #384
-#568 := [mp #555 #567]: #565
-#767 := [mp #568 #766]: #762
-#1489 := (not #762)
-#1490 := (or #1489 #1486)
-#1464 := (+ #1463 #1462)
-#1465 := (>= #1464 0::real)
-#1469 := (or #1468 #1465)
-#1470 := (not #1469)
-#1491 := (or #1489 #1470)
-#1493 := (iff #1491 #1490)
-#1495 := (iff #1490 #1490)
-#1496 := [rewrite]: #1495
-#1487 := (iff #1470 #1486)
-#1484 := (iff #1469 #1483)
-#1481 := (iff #1465 #1478)
-#1471 := (+ #1462 #1463)
-#1474 := (>= #1471 0::real)
-#1479 := (iff #1474 #1478)
-#1480 := [rewrite]: #1479
-#1475 := (iff #1465 #1474)
-#1472 := (= #1464 #1471)
-#1473 := [rewrite]: #1472
-#1476 := [monotonicity #1473]: #1475
-#1482 := [trans #1476 #1480]: #1481
-#1485 := [monotonicity #1482]: #1484
-#1488 := [monotonicity #1485]: #1487
-#1494 := [monotonicity #1488]: #1493
-#1497 := [trans #1494 #1496]: #1493
-#1492 := [quant-inst]: #1491
-#1498 := [mp #1492 #1497]: #1490
-#1701 := [unit-resolution #1498 #767]: #1486
-#1503 := (or #1483 #1502)
-#1504 := [def-axiom]: #1503
-#1702 := [unit-resolution #1504 #1701]: #1502
-#1579 := (+ #823 #1466)
-#1580 := (>= #1579 0::real)
-#1612 := (not #1580)
-#946 := (f8 #101 #44)
-#1591 := (+ #946 #1466)
-#1592 := (<= #1591 0::real)
-#1597 := (or #1580 #1592)
-#1600 := (not #1597)
-#727 := (pattern #99)
-#569 := (or #389 #393)
-#570 := (not #569)
-#768 := (forall (vars (?v0 S2)) (:pat #761 #727 #695) #570)
-#573 := (forall (vars (?v0 S2)) #570)
-#771 := (iff #573 #768)
-#769 := (iff #570 #570)
-#770 := [refl]: #769
-#772 := [quant-intro #770]: #771
-#574 := (iff #402 #573)
-#571 := (iff #399 #570)
-#572 := [rewrite]: #571
-#575 := [quant-intro #572]: #574
-#558 := (~ #402 #402)
-#556 := (~ #399 #399)
-#557 := [refl]: #556
-#559 := [nnf-pos #557]: #558
-#410 := [and-elim #408]: #402
-#560 := [mp~ #410 #559]: #402
-#576 := [mp #560 #575]: #573
-#773 := [mp #576 #772]: #768
-#1547 := (not #768)
-#1603 := (or #1547 #1600)
-#1565 := (* -1::real #946)
-#1566 := (+ #1463 #1565)
-#1567 := (>= #1566 0::real)
-#1568 := (* -1::real #823)
-#1569 := (+ #1463 #1568)
-#1570 := (<= #1569 0::real)
-#1571 := (or #1570 #1567)
-#1572 := (not #1571)
-#1604 := (or #1547 #1572)
-#1606 := (iff #1604 #1603)
-#1608 := (iff #1603 #1603)
-#1609 := [rewrite]: #1608
-#1601 := (iff #1572 #1600)
-#1598 := (iff #1571 #1597)
-#1595 := (iff #1567 #1592)
-#1585 := (+ #1565 #1463)
-#1588 := (>= #1585 0::real)
-#1593 := (iff #1588 #1592)
-#1594 := [rewrite]: #1593
-#1589 := (iff #1567 #1588)
-#1586 := (= #1566 #1585)
-#1587 := [rewrite]: #1586
-#1590 := [monotonicity #1587]: #1589
-#1596 := [trans #1590 #1594]: #1595
-#1583 := (iff #1570 #1580)
-#1573 := (+ #1568 #1463)
-#1576 := (<= #1573 0::real)
-#1581 := (iff #1576 #1580)
-#1582 := [rewrite]: #1581
-#1577 := (iff #1570 #1576)
-#1574 := (= #1569 #1573)
-#1575 := [rewrite]: #1574
-#1578 := [monotonicity #1575]: #1577
-#1584 := [trans #1578 #1582]: #1583
-#1599 := [monotonicity #1584 #1596]: #1598
-#1602 := [monotonicity #1599]: #1601
-#1607 := [monotonicity #1602]: #1606
-#1610 := [trans #1607 #1609]: #1606
-#1605 := [quant-inst]: #1604
-#1611 := [mp #1605 #1610]: #1603
-#1703 := [unit-resolution #1611 #773]: #1600
-#1613 := (or #1597 #1612)
-#1614 := [def-axiom]: #1613
-#1704 := [unit-resolution #1614 #1703]: #1612
-#1713 := (or #1708 #1580 #1478)
-#1714 := [th-lemma]: #1713
-#1715 := [unit-resolution #1714 #1704 #1702]: #1708
-#1760 := (or #1759 #1711)
-#1761 := [th-lemma]: #1760
-#1775 := [unit-resolution #1761 #1715]: #1759
-#1795 := [mp #1775 #1794]: #863
-#1797 := (or #836 #824)
-decl f14 :: S5
-#32 := f14
-#776 := (f8 f14 #44)
-#825 := (= #776 #823)
-#841 := (ite #836 #825 #824)
-#30 := (f8 f10 #9)
-#706 := (pattern #30)
-#33 := (f8 f14 #9)
-#705 := (pattern #33)
-#620 := (= #30 #99)
-#619 := (= #33 #99)
-#621 := (ite #157 #619 #620)
-#728 := (forall (vars (?v0 S2)) (:pat #694 #705 #727 #706) #621)
-#624 := (forall (vars (?v0 S2)) #621)
-#731 := (iff #624 #728)
-#729 := (iff #621 #621)
-#730 := [refl]: #729
-#732 := [quant-intro #730]: #731
-#235 := (ite #157 #33 #30)
-#508 := (= #99 #235)
-#509 := (forall (vars (?v0 S2)) #508)
-#625 := (iff #509 #624)
-#622 := (iff #508 #621)
-#623 := [rewrite]: #622
-#626 := [quant-intro #623]: #625
-#460 := (~ #509 #509)
-#462 := (~ #508 #508)
-#463 := [refl]: #462
-#461 := [nnf-pos #463]: #460
-#56 := (f3 f18 #9)
-#240 := (= #56 #235)
-#243 := (forall (vars (?v0 S2)) #240)
-#510 := (iff #243 #509)
-#511 := [rewrite* #441]: #510
-#57 := (ite #15 #30 #33)
-#58 := (= #56 #57)
-#59 := (forall (vars (?v0 S2)) #58)
-#244 := (iff #59 #243)
-#241 := (iff #58 #240)
-#238 := (= #57 #235)
-#232 := (ite #156 #30 #33)
-#236 := (= #232 #235)
-#237 := [rewrite]: #236
-#233 := (= #57 #232)
-#234 := [monotonicity #161]: #233
-#239 := [trans #234 #237]: #238
-#242 := [monotonicity #239]: #241
-#245 := [quant-intro #242]: #244
-#193 := [asserted]: #59
-#246 := [mp #193 #245]: #243
-#512 := [mp #246 #511]: #509
-#530 := [mp~ #512 #461]: #509
-#627 := [mp #530 #626]: #624
-#733 := [mp #627 #732]: #728
-#844 := (not #728)
-#845 := (or #844 #841)
-#826 := (+ #778 #158)
-#827 := (>= #826 0::int)
-#828 := (ite #827 #825 #824)
-#846 := (or #844 #828)
-#848 := (iff #846 #845)
-#850 := (iff #845 #845)
-#851 := [rewrite]: #850
-#842 := (iff #828 #841)
-#839 := (iff #827 #836)
-#829 := (+ #158 #778)
-#832 := (>= #829 0::int)
-#837 := (iff #832 #836)
-#838 := [rewrite]: #837
-#833 := (iff #827 #832)
-#830 := (= #826 #829)
-#831 := [rewrite]: #830
-#834 := [monotonicity #831]: #833
-#840 := [trans #834 #838]: #839
-#843 := [monotonicity #840]: #842
-#849 := [monotonicity #843]: #848
-#852 := [trans #849 #851]: #848
-#847 := [quant-inst]: #846
-#853 := [mp #847 #852]: #845
-#1796 := [unit-resolution #853 #733]: #841
-#854 := (not #841)
-#858 := (or #854 #836 #824)
-#859 := [def-axiom]: #858
-#1798 := [unit-resolution #859 #1796]: #1797
-#1799 := [unit-resolution #1798 #1795]: #836
-#855 := (not #836)
-#1746 := (or #1679 #855)
-#1738 := [hypothesis]: #836
-#1395 := (>= #789 0::int)
-#1740 := (or #1739 #1395)
-#1741 := [th-lemma]: #1740
-#1742 := [unit-resolution #1741 #1676]: #1395
-#1743 := (not #1679)
-#1744 := [hypothesis]: #1743
-#1745 := [th-lemma #1744 #1742 #1738]: false
-#1747 := [lemma #1745]: #1746
-#1800 := [unit-resolution #1747 #1799]: #1679
-#1803 := (or #1743 #1754)
-#1712 := (= #14 #41)
-#1736 := (not #1712)
-#356 := (= f7 f16)
-#953 := (= f7 #82)
-decl f20 :: (-> int S4)
-#1398 := (f20 #778)
-#1727 := (= #1398 #82)
-#1399 := (= #82 #1398)
-#65 := (:var 0 S4)
-#66 := (f5 #65)
-#741 := (pattern #66)
-#67 := (f20 #66)
-#247 := (= #65 #67)
-#742 := (forall (vars (?v0 S4)) (:pat #741) #247)
-#265 := (forall (vars (?v0 S4)) #247)
-#745 := (iff #265 #742)
-#743 := (iff #247 #247)
-#744 := [refl]: #743
-#746 := [quant-intro #744]: #745
-#538 := (~ #265 #265)
-#536 := (~ #247 #247)
-#537 := [refl]: #536
-#539 := [nnf-pos #537]: #538
-#68 := (= #67 #65)
-#69 := (forall (vars (?v0 S4)) #68)
-#266 := (iff #69 #265)
-#263 := (iff #68 #247)
-#264 := [rewrite]: #263
-#267 := [quant-intro #264]: #266
-#231 := [asserted]: #69
-#270 := [mp #231 #267]: #265
-#540 := [mp~ #270 #539]: #265
-#747 := [mp #540 #746]: #742
-#1294 := (not #742)
-#1402 := (or #1294 #1399)
-#1403 := [quant-inst]: #1402
-#1672 := [unit-resolution #1403 #747]: #1399
-#1728 := [symm #1672]: #1727
-#1731 := (= f7 #1398)
-#1400 := (f20 #1041)
-#1725 := (= #1400 #1398)
-#1722 := (= #1041 #778)
-#1720 := (= #1041 #41)
-#1716 := [hypothesis]: #1712
-#1721 := [trans #1719 #1716]: #1720
-#1723 := [trans #1721 #1676]: #1722
-#1726 := [monotonicity #1723]: #1725
-#1729 := (= f7 #1400)
-#1401 := (= #84 #1400)
-#1406 := (or #1294 #1401)
-#1407 := [quant-inst]: #1406
-#1724 := [unit-resolution #1407 #747]: #1401
-#1730 := [trans #357 #1724]: #1729
-#1732 := [trans #1730 #1726]: #1731
-#1733 := [trans #1732 #1728]: #953
-#1734 := [trans #1733 #1673]: #356
-#360 := (not #356)
-#86 := (= f16 f7)
-#87 := (not #86)
-#361 := (iff #87 #360)
-#358 := (iff #86 #356)
-#359 := [rewrite]: #358
-#362 := [monotonicity #359]: #361
-#355 := [asserted]: #87
-#365 := [mp #355 #362]: #360
-#1735 := [unit-resolution #365 #1734]: false
-#1737 := [lemma #1735]: #1736
-#1801 := (or #1712 #1743 #1754)
-#1802 := [th-lemma]: #1801
-#1804 := [unit-resolution #1802 #1737]: #1803
-#1805 := [unit-resolution #1804 #1800]: #1754
-#1757 := (or #1666 #1074)
-#1748 := [hypothesis]: #1053
-#1397 := (>= #1097 0::int)
-#1751 := (or #1750 #1397)
-#1752 := [th-lemma]: #1751
-#1753 := [unit-resolution #1752 #1749]: #1397
-#1755 := [hypothesis]: #1754
-#1756 := [th-lemma #1755 #1753 #1748]: false
-#1758 := [lemma #1756]: #1757
-#1806 := [unit-resolution #1758 #1805]: #1074
-#1808 := (or #1053 #1038)
-#1039 := (f8 f14 #20)
-#1058 := (= #1037 #1039)
-#1061 := (ite #1053 #1058 #1038)
-#629 := (= #30 #89)
-#628 := (= #33 #89)
-#630 := (ite #195 #628 #629)
-#735 := (forall (vars (?v0 S2)) (:pat #694 #705 #734 #706) #630)
-#633 := (forall (vars (?v0 S2)) #630)
-#738 := (iff #633 #735)
-#736 := (iff #630 #630)
-#737 := [refl]: #736
-#739 := [quant-intro #737]: #738
-#251 := (ite #195 #33 #30)
-#513 := (= #89 #251)
-#514 := (forall (vars (?v0 S2)) #513)
-#634 := (iff #514 #633)
-#631 := (iff #513 #630)
-#632 := [rewrite]: #631
-#635 := [quant-intro #632]: #634
-#533 := (~ #514 #514)
-#531 := (~ #513 #513)
-#532 := [refl]: #531
-#534 := [nnf-pos #532]: #533
-#61 := (f3 f19 #9)
-#256 := (= #61 #251)
-#259 := (forall (vars (?v0 S2)) #256)
-#515 := (iff #259 #514)
-#516 := [rewrite* #441]: #515
-#62 := (ite #42 #30 #33)
-#63 := (= #61 #62)
-#64 := (forall (vars (?v0 S2)) #63)
-#260 := (iff #64 #259)
-#257 := (iff #63 #256)
-#254 := (= #62 #251)
-#248 := (ite #194 #30 #33)
-#252 := (= #248 #251)
-#253 := [rewrite]: #252
-#249 := (= #62 #248)
-#250 := [monotonicity #199]: #249
-#255 := [trans #250 #253]: #254
-#258 := [monotonicity #255]: #257
-#261 := [quant-intro #258]: #260
-#215 := [asserted]: #64
-#262 := [mp #215 #261]: #259
-#517 := [mp #262 #516]: #514
-#535 := [mp~ #517 #534]: #514
-#636 := [mp #535 #635]: #633
-#740 := [mp #636 #739]: #735
-#801 := (not #735)
-#1064 := (or #801 #1061)
-#1040 := (= #1039 #1037)
-#1042 := (+ #1041 #196)
-#1043 := (>= #1042 0::int)
-#1044 := (ite #1043 #1040 #1038)
-#1065 := (or #801 #1044)
-#1067 := (iff #1065 #1064)
-#1069 := (iff #1064 #1064)
-#1070 := [rewrite]: #1069
-#1062 := (iff #1044 #1061)
-#1059 := (iff #1040 #1058)
-#1060 := [rewrite]: #1059
-#1056 := (iff #1043 #1053)
-#1045 := (+ #196 #1041)
-#1048 := (>= #1045 0::int)
-#1054 := (iff #1048 #1053)
-#1055 := [rewrite]: #1054
-#1049 := (iff #1043 #1048)
-#1046 := (= #1042 #1045)
-#1047 := [rewrite]: #1046
-#1050 := [monotonicity #1047]: #1049
-#1057 := [trans #1050 #1055]: #1056
-#1063 := [monotonicity #1057 #1060]: #1062
-#1068 := [monotonicity #1063]: #1067
-#1071 := [trans #1068 #1070]: #1067
-#1066 := [quant-inst]: #1065
-#1072 := [mp #1066 #1071]: #1064
-#1807 := [unit-resolution #1072 #740]: #1061
-#1073 := (not #1061)
-#1077 := (or #1073 #1053 #1038)
-#1078 := [def-axiom]: #1077
-#1809 := [unit-resolution #1078 #1807]: #1808
-#1810 := [unit-resolution #1809 #1806]: #1038
-#1812 := [symm #1810]: #1811
-#1813 := [trans #1812 #1774]: #1697
-#1814 := (not #1697)
-#1815 := (or #1814 #1765)
-#1816 := [th-lemma]: #1815
-#1817 := [unit-resolution #1816 #1813]: #1765
-#1508 := (f8 f22 #20)
-#1522 := (* -1::real #1508)
-#1535 := (+ #1205 #1522)
-#1536 := (<= #1535 0::real)
-#1560 := (not #1536)
-#1085 := (f8 #98 #20)
-#1523 := (+ #1085 #1522)
-#1524 := (>= #1523 0::real)
-#1541 := (or #1524 #1536)
-#1544 := (not #1541)
-#1548 := (or #1547 #1544)
-#1509 := (+ #1508 #1507)
-#1510 := (>= #1509 0::real)
-#1511 := (* -1::real #1085)
-#1512 := (+ #1508 #1511)
-#1513 := (<= #1512 0::real)
-#1514 := (or #1513 #1510)
-#1515 := (not #1514)
-#1549 := (or #1547 #1515)
-#1551 := (iff #1549 #1548)
-#1553 := (iff #1548 #1548)
-#1554 := [rewrite]: #1553
-#1545 := (iff #1515 #1544)
-#1542 := (iff #1514 #1541)
-#1539 := (iff #1510 #1536)
-#1529 := (+ #1507 #1508)
-#1532 := (>= #1529 0::real)
-#1537 := (iff #1532 #1536)
-#1538 := [rewrite]: #1537
-#1533 := (iff #1510 #1532)
-#1530 := (= #1509 #1529)
-#1531 := [rewrite]: #1530
-#1534 := [monotonicity #1531]: #1533
-#1540 := [trans #1534 #1538]: #1539
-#1527 := (iff #1513 #1524)
-#1516 := (+ #1511 #1508)
-#1519 := (<= #1516 0::real)
-#1525 := (iff #1519 #1524)
-#1526 := [rewrite]: #1525
-#1520 := (iff #1513 #1519)
-#1517 := (= #1512 #1516)
-#1518 := [rewrite]: #1517
-#1521 := [monotonicity #1518]: #1520
-#1528 := [trans #1521 #1526]: #1527
-#1543 := [monotonicity #1528 #1540]: #1542
-#1546 := [monotonicity #1543]: #1545
-#1552 := [monotonicity #1546]: #1551
-#1555 := [trans #1552 #1554]: #1551
-#1550 := [quant-inst]: #1549
-#1556 := [mp #1550 #1555]: #1548
-#1818 := [unit-resolution #1556 #773]: #1544
-#1561 := (or #1541 #1560)
-#1562 := [def-axiom]: #1561
-#1819 := [unit-resolution #1562 #1818]: #1560
-#1623 := (+ #1037 #1522)
-#1624 := (>= #1623 0::real)
-#1654 := (not #1624)
-#1158 := (f8 #93 #20)
-#1633 := (+ #1158 #1522)
-#1634 := (<= #1633 0::real)
-#1639 := (or #1624 #1634)
-#1642 := (not #1639)
-#1645 := (or #1489 #1642)
-#1620 := (* -1::real #1158)
-#1621 := (+ #1508 #1620)
-#1622 := (>= #1621 0::real)
-#1625 := (or #1624 #1622)
-#1626 := (not #1625)
-#1646 := (or #1489 #1626)
-#1648 := (iff #1646 #1645)
-#1650 := (iff #1645 #1645)
-#1651 := [rewrite]: #1650
-#1643 := (iff #1626 #1642)
-#1640 := (iff #1625 #1639)
-#1637 := (iff #1622 #1634)
-#1627 := (+ #1620 #1508)
-#1630 := (>= #1627 0::real)
-#1635 := (iff #1630 #1634)
-#1636 := [rewrite]: #1635
-#1631 := (iff #1622 #1630)
-#1628 := (= #1621 #1627)
-#1629 := [rewrite]: #1628
-#1632 := [monotonicity #1629]: #1631
-#1638 := [trans #1632 #1636]: #1637
-#1641 := [monotonicity #1638]: #1640
-#1644 := [monotonicity #1641]: #1643
-#1649 := [monotonicity #1644]: #1648
-#1652 := [trans #1649 #1651]: #1648
-#1647 := [quant-inst]: #1646
-#1653 := [mp #1647 #1652]: #1645
-#1820 := [unit-resolution #1653 #767]: #1642
-#1655 := (or #1639 #1654)
-#1656 := [def-axiom]: #1655
-#1821 := [unit-resolution #1656 #1820]: #1654
-[th-lemma #1821 #1819 #1817]: false
+#172 := [monotonicity #169 #140]: #171
+#175 := [monotonicity #172]: #174
+#180 := [trans #175 #178]: #179
+#142 := (iff #57 #141)
+#135 := (= #53 #132)
+#129 := (* #126 3::real)
+#133 := (= #129 #132)
+#134 := [rewrite]: #133
+#130 := (= #53 #129)
+#127 := (= #51 #126)
+#112 := (= #47 #111)
+#113 := [rewrite]: #112
+#124 := (= #50 #121)
+#117 := (- #111)
+#122 := (= #117 #121)
+#123 := [rewrite]: #122
+#118 := (= #50 #117)
+#119 := [monotonicity #113]: #118
+#125 := [trans #119 #123]: #124
+#115 := (iff #49 #114)
+#116 := [monotonicity #113]: #115
+#128 := [monotonicity #116 #125 #113]: #127
+#131 := [monotonicity #128]: #130
+#136 := [trans #131 #134]: #135
+#143 := [monotonicity #136 #140]: #142
+#183 := [monotonicity #143 #180]: #182
+#255 := [trans #183 #253]: #254
+#108 := [asserted]: #67
+#256 := [mp #108 #255]: #249
+#257 := [and-elim #256]: #185
+#259 := [and-elim #256]: #238
+#448 := (+ #111 #447)
+#449 := (<= #448 0::real)
+#441 := (= #111 #223)
+#443 := (or #217 #441)
+#444 := [def-axiom]: #443
+#467 := [unit-resolution #444 #455]: #441
+#468 := (not #441)
+#469 := (or #468 #449)
+#470 := [th-lemma]: #469
+#471 := [unit-resolution #470 #467]: #449
+#463 := (or #189 #217)
+#456 := [hypothesis]: #188
+#258 := [and-elim #256]: #211
+#437 := (or #189 #435)
+#438 := [def-axiom]: #437
+#457 := [unit-resolution #438 #456]: #435
+#458 := (not #435)
+#459 := (or #458 #453)
+#460 := [th-lemma]: #459
+#461 := [unit-resolution #460 #457]: #453
+#462 := [th-lemma #257 #461 #258 #456 #455]: false
+#464 := [lemma #462]: #463
+#472 := [unit-resolution #464 #455]: #189
+#473 := [th-lemma #472 #471 #259 #257 #455]: false
+#474 := [lemma #473]: #217
+#445 := (or #216 #442)
+#446 := [def-axiom]: #445
+#475 := [unit-resolution #446 #474]: #442
+#476 := (not #442)
+#477 := (or #476 #454)
+#478 := [th-lemma]: #477
+#479 := [unit-resolution #478 #475]: #454
+#481 := (not #185)
+#480 := (not #454)
+#482 := (or #188 #480 #237 #481 #216)
+#483 := [th-lemma]: #482
+#484 := [unit-resolution #483 #257 #474 #259 #479]: #188
+#485 := [unit-resolution #438 #484]: #435
+#486 := [unit-resolution #460 #485]: #453
+[th-lemma #479 #259 #257 #474 #258 #486]: false
 unsat
-6c73093b27236ef09bc4a53162dee78b6dc31895 422 0
-#2 := false
-decl f12 :: S2
-#42 := f12
-decl f5 :: S2
-#25 := f5
-#45 := (= f5 f12)
-decl f3 :: (-> int S2)
-decl f4 :: (-> S2 int)
-#43 := (f4 f12)
-#598 := (f3 #43)
-#696 := (= #598 f12)
-#599 := (= f12 #598)
-#8 := (:var 0 S2)
-#9 := (f4 #8)
-#551 := (pattern #9)
-#10 := (f3 #9)
-#98 := (= #8 #10)
-#552 := (forall (vars (?v0 S2)) (:pat #551) #98)
-#101 := (forall (vars (?v0 S2)) #98)
-#555 := (iff #101 #552)
-#553 := (iff #98 #98)
-#554 := [refl]: #553
-#556 := [quant-intro #554]: #555
-#455 := (~ #101 #101)
-#457 := (~ #98 #98)
-#458 := [refl]: #457
-#456 := [nnf-pos #458]: #455
-#11 := (= #10 #8)
-#12 := (forall (vars (?v0 S2)) #11)
-#102 := (iff #12 #101)
-#99 := (iff #11 #98)
-#100 := [rewrite]: #99
-#103 := [quant-intro #100]: #102
-#97 := [asserted]: #12
-#106 := [mp #97 #103]: #101
-#453 := [mp~ #106 #456]: #101
-#557 := [mp #453 #556]: #552
-#600 := (not #552)
-#605 := (or #600 #599)
-#606 := [quant-inst]: #605
-#690 := [unit-resolution #606 #557]: #599
-#697 := [symm #690]: #696
-#698 := (= f5 #598)
-#26 := (f4 f5)
-#596 := (f3 #26)
-#694 := (= #596 #598)
-#692 := (= #598 #596)
-#688 := (= #43 #26)
-#686 := (= #26 #43)
-#13 := 0::int
-#231 := -1::int
-#234 := (* -1::int #43)
-#235 := (+ #26 #234)
-#295 := (<= #235 0::int)
-#74 := (<= #26 #43)
-#393 := (iff #74 #295)
-#394 := [rewrite]: #393
-#346 := [asserted]: #74
-#395 := [mp #346 #394]: #295
-#233 := (>= #235 0::int)
-decl f6 :: (-> S3 S4 real)
-decl f8 :: (-> S2 S4)
-#29 := (f8 f5)
-decl f7 :: S3
-#28 := f7
-#30 := (f6 f7 #29)
-decl f9 :: S3
-#31 := f9
-#32 := (f6 f9 #29)
-#46 := (f8 f12)
-decl f11 :: S3
-#37 := f11
-#47 := (f6 f11 #46)
-#48 := (ite #45 #47 #32)
-#241 := (ite #233 #48 #30)
-#572 := (= #30 #241)
-#658 := (not #572)
-#199 := 0::real
-#197 := -1::real
-#249 := (* -1::real #241)
-#647 := (+ #30 #249)
-#648 := (<= #647 0::real)
-#652 := (not #648)
-#650 := [hypothesis]: #648
-decl f10 :: S3
-#34 := f10
-#35 := (f6 f10 #29)
-#250 := (+ #35 #249)
-#251 := (<= #250 0::real)
-#252 := (not #251)
-#38 := (f6 f11 #29)
-decl f13 :: S3
-#51 := f13
-#52 := (f6 f13 #29)
-#260 := (ite #233 #52 #38)
-#269 := (* -1::real #260)
-#270 := (+ #35 #269)
-#268 := (>= #270 0::real)
-#271 := (not #268)
-#276 := (and #252 #271)
-#44 := (< #26 #43)
-#53 := (ite #44 #38 #52)
-#54 := (< #35 #53)
-#49 := (ite #44 #30 #48)
-#50 := (< #49 #35)
-#55 := (and #50 #54)
-#277 := (iff #55 #276)
-#274 := (iff #54 #271)
-#265 := (< #35 #260)
-#272 := (iff #265 #271)
-#273 := [rewrite]: #272
-#266 := (iff #54 #265)
-#263 := (= #53 #260)
-#232 := (not #233)
-#257 := (ite #232 #38 #52)
-#261 := (= #257 #260)
-#262 := [rewrite]: #261
-#258 := (= #53 #257)
-#236 := (iff #44 #232)
-#237 := [rewrite]: #236
-#259 := [monotonicity #237]: #258
-#264 := [trans #259 #262]: #263
-#267 := [monotonicity #264]: #266
-#275 := [trans #267 #273]: #274
-#255 := (iff #50 #252)
-#246 := (< #241 #35)
-#253 := (iff #246 #252)
-#254 := [rewrite]: #253
-#247 := (iff #50 #246)
-#244 := (= #49 #241)
-#238 := (ite #232 #30 #48)
-#242 := (= #238 #241)
-#243 := [rewrite]: #242
-#239 := (= #49 #238)
-#240 := [monotonicity #237]: #239
-#245 := [trans #240 #243]: #244
-#248 := [monotonicity #245]: #247
-#256 := [trans #248 #254]: #255
-#278 := [monotonicity #256 #275]: #277
-#183 := [asserted]: #55
-#279 := [mp #183 #278]: #276
-#280 := [and-elim #279]: #252
-#201 := (* -1::real #35)
-#217 := (+ #30 #201)
-#218 := (<= #217 0::real)
-#219 := (not #218)
-#202 := (+ #32 #201)
-#200 := (>= #202 0::real)
-#198 := (not #200)
-#224 := (and #198 #219)
-#27 := (< #26 #26)
-#39 := (ite #27 #38 #30)
-#40 := (< #35 #39)
-#33 := (ite #27 #30 #32)
-#36 := (< #33 #35)
-#41 := (and #36 #40)
-#225 := (iff #41 #224)
-#222 := (iff #40 #219)
-#214 := (< #35 #30)
-#220 := (iff #214 #219)
-#221 := [rewrite]: #220
-#215 := (iff #40 #214)
-#212 := (= #39 #30)
-#207 := (ite false #38 #30)
-#210 := (= #207 #30)
-#211 := [rewrite]: #210
-#208 := (= #39 #207)
-#185 := (iff #27 false)
-#186 := [rewrite]: #185
-#209 := [monotonicity #186]: #208
-#213 := [trans #209 #211]: #212
-#216 := [monotonicity #213]: #215
-#223 := [trans #216 #221]: #222
-#205 := (iff #36 #198)
-#194 := (< #32 #35)
-#203 := (iff #194 #198)
-#204 := [rewrite]: #203
-#195 := (iff #36 #194)
-#192 := (= #33 #32)
-#187 := (ite false #30 #32)
-#190 := (= #187 #32)
-#191 := [rewrite]: #190
-#188 := (= #33 #187)
-#189 := [monotonicity #186]: #188
-#193 := [trans #189 #191]: #192
-#196 := [monotonicity #193]: #195
-#206 := [trans #196 #204]: #205
-#226 := [monotonicity #206 #223]: #225
-#182 := [asserted]: #41
-#227 := [mp #182 #226]: #224
-#229 := [and-elim #227]: #219
-#651 := [th-lemma #229 #280 #650]: false
-#653 := [lemma #651]: #652
-#657 := [hypothesis]: #572
-#659 := (or #658 #648)
-#660 := [th-lemma]: #659
-#661 := [unit-resolution #660 #657 #653]: false
-#662 := [lemma #661]: #658
-#582 := (or #233 #572)
-#583 := [def-axiom]: #582
-#685 := [unit-resolution #583 #662]: #233
-#687 := [th-lemma #685 #395]: #686
-#689 := [symm #687]: #688
-#693 := [monotonicity #689]: #692
-#695 := [symm #693]: #694
-#597 := (= f5 #596)
-#601 := (or #600 #597)
-#602 := [quant-inst]: #601
-#691 := [unit-resolution #602 #557]: #597
-#699 := [trans #691 #695]: #698
-#700 := [trans #699 #697]: #45
-#575 := (not #45)
-#63 := (f6 f13 #46)
-#283 := (ite #45 #30 #63)
-#466 := (* -1::real #283)
-#642 := (+ #30 #466)
-#644 := (>= #642 0::real)
-#590 := (= #30 #283)
-#666 := [hypothesis]: #45
-#592 := (or #575 #590)
-#593 := [def-axiom]: #592
-#667 := [unit-resolution #593 #666]: #590
-#668 := (not #590)
-#669 := (or #668 #644)
-#670 := [th-lemma]: #669
-#671 := [unit-resolution #670 #667]: #644
-#60 := (f6 f10 #46)
-#362 := (* -1::real #60)
-#363 := (+ #47 #362)
-#361 := (>= #363 0::real)
-#360 := (not #361)
-#379 := (* -1::real #63)
-#380 := (+ #60 #379)
-#378 := (>= #380 0::real)
-#381 := (not #378)
-#386 := (and #360 #381)
-#68 := (< #43 #43)
-#71 := (ite #68 #47 #63)
-#72 := (< #60 #71)
-#57 := (f6 f7 #46)
-#69 := (ite #68 #57 #47)
-#70 := (< #69 #60)
-#73 := (and #70 #72)
-#387 := (iff #73 #386)
-#384 := (iff #72 #381)
-#375 := (< #60 #63)
-#382 := (iff #375 #381)
-#383 := [rewrite]: #382
-#376 := (iff #72 #375)
-#373 := (= #71 #63)
-#368 := (ite false #47 #63)
-#371 := (= #368 #63)
-#372 := [rewrite]: #371
-#369 := (= #71 #368)
-#348 := (iff #68 false)
-#349 := [rewrite]: #348
-#370 := [monotonicity #349]: #369
-#374 := [trans #370 #372]: #373
-#377 := [monotonicity #374]: #376
-#385 := [trans #377 #383]: #384
-#366 := (iff #70 #360)
-#357 := (< #47 #60)
-#364 := (iff #357 #360)
-#365 := [rewrite]: #364
-#358 := (iff #70 #357)
-#355 := (= #69 #47)
-#350 := (ite false #57 #47)
-#353 := (= #350 #47)
-#354 := [rewrite]: #353
-#351 := (= #69 #350)
-#352 := [monotonicity #349]: #351
-#356 := [trans #352 #354]: #355
-#359 := [monotonicity #356]: #358
-#367 := [trans #359 #365]: #366
-#388 := [monotonicity #367 #385]: #387
-#345 := [asserted]: #73
-#389 := [mp #345 #388]: #386
-#390 := [and-elim #389]: #360
-#399 := (* -1::real #57)
-#400 := (+ #47 #399)
-#398 := (>= #400 0::real)
-#58 := (f6 f9 #46)
-#407 := (* -1::real #58)
-#408 := (+ #57 #407)
-#406 := (>= #408 0::real)
-#402 := (+ #47 #379)
-#403 := (<= #402 0::real)
-#417 := (and #398 #403 #406)
-#77 := (<= #47 #63)
-#76 := (<= #57 #47)
-#78 := (and #76 #77)
-#75 := (<= #58 #57)
-#79 := (and #75 #78)
-#420 := (iff #79 #417)
-#411 := (and #398 #403)
-#414 := (and #406 #411)
-#418 := (iff #414 #417)
-#419 := [rewrite]: #418
-#415 := (iff #79 #414)
-#412 := (iff #78 #411)
-#404 := (iff #77 #403)
-#405 := [rewrite]: #404
-#397 := (iff #76 #398)
-#401 := [rewrite]: #397
-#413 := [monotonicity #401 #405]: #412
-#409 := (iff #75 #406)
-#410 := [rewrite]: #409
-#416 := [monotonicity #410 #413]: #415
-#421 := [trans #416 #419]: #420
-#347 := [asserted]: #79
-#422 := [mp #347 #421]: #417
-#423 := [and-elim #422]: #398
-#655 := (+ #30 #399)
-#656 := (<= #655 0::real)
-#654 := (= #30 #57)
-#676 := (= #57 #30)
-#674 := (= #46 #29)
-#672 := (= #29 #46)
-#673 := [monotonicity #666]: #672
-#675 := [symm #673]: #674
-#677 := [monotonicity #675]: #676
-#678 := [symm #677]: #654
-#679 := (not #654)
-#680 := (or #679 #656)
-#681 := [th-lemma]: #680
-#682 := [unit-resolution #681 #678]: #656
-#469 := (+ #60 #466)
-#472 := (>= #469 0::real)
-#445 := (not #472)
-#321 := (ite #295 #283 #47)
-#331 := (* -1::real #321)
-#332 := (+ #60 #331)
-#330 := (>= #332 0::real)
-#329 := (not #330)
-#446 := (iff #329 #445)
-#473 := (iff #330 #472)
-#470 := (= #332 #469)
-#467 := (= #331 #466)
-#464 := (= #321 #283)
-#1 := true
-#459 := (ite true #283 #47)
-#462 := (= #459 #283)
-#463 := [rewrite]: #462
-#460 := (= #321 #459)
-#451 := (iff #295 true)
-#452 := [iff-true #395]: #451
-#461 := [monotonicity #452]: #460
-#465 := [trans #461 #463]: #464
-#468 := [monotonicity #465]: #467
-#471 := [monotonicity #468]: #470
-#474 := [monotonicity #471]: #473
-#475 := [monotonicity #474]: #446
-#302 := (ite #295 #58 #57)
-#310 := (* -1::real #302)
-#311 := (+ #60 #310)
-#312 := (<= #311 0::real)
-#313 := (not #312)
-#337 := (and #313 #329)
-#62 := (= f12 f5)
-#64 := (ite #62 #30 #63)
-#56 := (< #43 #26)
-#65 := (ite #56 #47 #64)
-#66 := (< #60 #65)
-#59 := (ite #56 #57 #58)
-#61 := (< #59 #60)
-#67 := (and #61 #66)
-#340 := (iff #67 #337)
-#286 := (ite #56 #47 #283)
-#289 := (< #60 #286)
-#292 := (and #61 #289)
-#338 := (iff #292 #337)
-#335 := (iff #289 #329)
-#326 := (< #60 #321)
-#333 := (iff #326 #329)
-#334 := [rewrite]: #333
-#327 := (iff #289 #326)
-#324 := (= #286 #321)
-#296 := (not #295)
-#318 := (ite #296 #47 #283)
-#322 := (= #318 #321)
-#323 := [rewrite]: #322
-#319 := (= #286 #318)
-#297 := (iff #56 #296)
-#298 := [rewrite]: #297
-#320 := [monotonicity #298]: #319
-#325 := [trans #320 #323]: #324
-#328 := [monotonicity #325]: #327
-#336 := [trans #328 #334]: #335
-#316 := (iff #61 #313)
-#307 := (< #302 #60)
-#314 := (iff #307 #313)
-#315 := [rewrite]: #314
-#308 := (iff #61 #307)
-#305 := (= #59 #302)
-#299 := (ite #296 #57 #58)
-#303 := (= #299 #302)
-#304 := [rewrite]: #303
-#300 := (= #59 #299)
-#301 := [monotonicity #298]: #300
-#306 := [trans #301 #304]: #305
-#309 := [monotonicity #306]: #308
-#317 := [trans #309 #315]: #316
-#339 := [monotonicity #317 #336]: #338
-#293 := (iff #67 #292)
-#290 := (iff #66 #289)
-#287 := (= #65 #286)
-#284 := (= #64 #283)
-#230 := (iff #62 #45)
-#282 := [rewrite]: #230
-#285 := [monotonicity #282]: #284
-#288 := [monotonicity #285]: #287
-#291 := [monotonicity #288]: #290
-#294 := [monotonicity #291]: #293
-#341 := [trans #294 #339]: #340
-#184 := [asserted]: #67
-#342 := [mp #184 #341]: #337
-#344 := [and-elim #342]: #329
-#476 := [mp #344 #475]: #445
-#683 := [th-lemma #476 #682 #423 #390 #671]: false
-#684 := [lemma #683]: #575
-[unit-resolution #684 #700]: false
-unsat
-5ee060971856d2def7cc6d40549073dace7efe45 428 0
-#2 := false
-decl f12 :: S2
-#42 := f12
-decl f5 :: S2
-#25 := f5
-#49 := (= f5 f12)
-decl f3 :: (-> int S2)
-decl f4 :: (-> S2 int)
-#43 := (f4 f12)
-#593 := (f3 #43)
-#691 := (= #593 f12)
-#594 := (= f12 #593)
-#8 := (:var 0 S2)
-#9 := (f4 #8)
-#546 := (pattern #9)
-#10 := (f3 #9)
-#98 := (= #8 #10)
-#547 := (forall (vars (?v0 S2)) (:pat #546) #98)
-#101 := (forall (vars (?v0 S2)) #98)
-#550 := (iff #101 #547)
-#548 := (iff #98 #98)
-#549 := [refl]: #548
-#551 := [quant-intro #549]: #550
-#461 := (~ #101 #101)
-#463 := (~ #98 #98)
-#464 := [refl]: #463
-#462 := [nnf-pos #464]: #461
-#11 := (= #10 #8)
-#12 := (forall (vars (?v0 S2)) #11)
-#102 := (iff #12 #101)
-#99 := (iff #11 #98)
-#100 := [rewrite]: #99
-#103 := [quant-intro #100]: #102
-#97 := [asserted]: #12
-#106 := [mp #97 #103]: #101
-#459 := [mp~ #106 #462]: #101
-#552 := [mp #459 #551]: #547
-#595 := (not #547)
-#600 := (or #595 #594)
-#601 := [quant-inst]: #600
-#685 := [unit-resolution #601 #552]: #594
-#692 := [symm #685]: #691
-#693 := (= f5 #593)
-#26 := (f4 f5)
-#591 := (f3 #26)
-#689 := (= #591 #593)
-#687 := (= #593 #591)
-#683 := (= #43 #26)
-#681 := (= #26 #43)
-#13 := 0::int
-#232 := -1::int
-#235 := (* -1::int #43)
-#236 := (+ #26 #235)
-#301 := (<= #236 0::int)
-#74 := (<= #26 #43)
-#398 := (iff #74 #301)
-#399 := [rewrite]: #398
-#352 := [asserted]: #74
-#400 := [mp #352 #399]: #301
-#234 := (>= #236 0::int)
-decl f6 :: (-> S3 S4 real)
-decl f8 :: (-> S2 S4)
-#29 := (f8 f5)
-decl f9 :: S3
-#31 := f9
-#32 := (f6 f9 #29)
-decl f11 :: S3
-#37 := f11
-#38 := (f6 f11 #29)
-#50 := (f8 f12)
-decl f7 :: S3
-#28 := f7
-#51 := (f6 f7 #50)
-#52 := (ite #49 #51 #38)
-#261 := (ite #234 #52 #32)
-#573 := (= #32 #261)
-#653 := (not #573)
-#199 := 0::real
-#197 := -1::real
-#270 := (* -1::real #261)
-#645 := (+ #32 #270)
-#647 := (>= #645 0::real)
-#650 := (not #647)
-#648 := [hypothesis]: #647
-decl f10 :: S3
-#34 := f10
-#35 := (f6 f10 #29)
-#271 := (+ #35 #270)
-#269 := (>= #271 0::real)
-#272 := (not #269)
-#30 := (f6 f7 #29)
-decl f13 :: S3
-#45 := f13
-#46 := (f6 f13 #29)
-#242 := (ite #234 #46 #30)
-#250 := (* -1::real #242)
-#251 := (+ #35 #250)
-#252 := (<= #251 0::real)
-#253 := (not #252)
-#277 := (and #253 #272)
-#44 := (< #26 #43)
-#53 := (ite #44 #32 #52)
-#54 := (< #35 #53)
-#47 := (ite #44 #30 #46)
-#48 := (< #47 #35)
-#55 := (and #48 #54)
-#278 := (iff #55 #277)
-#275 := (iff #54 #272)
-#266 := (< #35 #261)
-#273 := (iff #266 #272)
-#274 := [rewrite]: #273
-#267 := (iff #54 #266)
-#264 := (= #53 #261)
-#233 := (not #234)
-#258 := (ite #233 #32 #52)
-#262 := (= #258 #261)
-#263 := [rewrite]: #262
-#259 := (= #53 #258)
-#237 := (iff #44 #233)
-#238 := [rewrite]: #237
-#260 := [monotonicity #238]: #259
-#265 := [trans #260 #263]: #264
-#268 := [monotonicity #265]: #267
-#276 := [trans #268 #274]: #275
-#256 := (iff #48 #253)
-#247 := (< #242 #35)
-#254 := (iff #247 #253)
-#255 := [rewrite]: #254
-#248 := (iff #48 #247)
-#245 := (= #47 #242)
-#239 := (ite #233 #30 #46)
-#243 := (= #239 #242)
-#244 := [rewrite]: #243
-#240 := (= #47 #239)
-#241 := [monotonicity #238]: #240
-#246 := [trans #241 #244]: #245
-#249 := [monotonicity #246]: #248
-#257 := [trans #249 #255]: #256
-#279 := [monotonicity #257 #276]: #278
-#183 := [asserted]: #55
-#280 := [mp #183 #279]: #277
-#282 := [and-elim #280]: #272
-#201 := (* -1::real #35)
-#202 := (+ #32 #201)
-#200 := (>= #202 0::real)
-#198 := (not #200)
-#218 := (* -1::real #38)
-#219 := (+ #35 #218)
-#217 := (>= #219 0::real)
-#220 := (not #217)
-#225 := (and #198 #220)
-#27 := (< #26 #26)
-#39 := (ite #27 #32 #38)
-#40 := (< #35 #39)
-#33 := (ite #27 #30 #32)
-#36 := (< #33 #35)
-#41 := (and #36 #40)
-#226 := (iff #41 #225)
-#223 := (iff #40 #220)
-#214 := (< #35 #38)
-#221 := (iff #214 #220)
-#222 := [rewrite]: #221
-#215 := (iff #40 #214)
-#212 := (= #39 #38)
-#207 := (ite false #32 #38)
-#210 := (= #207 #38)
-#211 := [rewrite]: #210
-#208 := (= #39 #207)
-#185 := (iff #27 false)
-#186 := [rewrite]: #185
-#209 := [monotonicity #186]: #208
-#213 := [trans #209 #211]: #212
-#216 := [monotonicity #213]: #215
-#224 := [trans #216 #222]: #223
-#205 := (iff #36 #198)
-#194 := (< #32 #35)
-#203 := (iff #194 #198)
-#204 := [rewrite]: #203
-#195 := (iff #36 #194)
-#192 := (= #33 #32)
-#187 := (ite false #30 #32)
-#190 := (= #187 #32)
-#191 := [rewrite]: #190
-#188 := (= #33 #187)
-#189 := [monotonicity #186]: #188
-#193 := [trans #189 #191]: #192
-#196 := [monotonicity #193]: #195
-#206 := [trans #196 #204]: #205
-#227 := [monotonicity #206 #224]: #226
-#182 := [asserted]: #41
-#228 := [mp #182 #227]: #225
-#229 := [and-elim #228]: #198
-#649 := [th-lemma #229 #282 #648]: false
-#651 := [lemma #649]: #650
-#652 := [hypothesis]: #573
-#654 := (or #653 #647)
-#655 := [th-lemma]: #654
-#656 := [unit-resolution #655 #652 #651]: false
-#657 := [lemma #656]: #653
-#583 := (or #234 #573)
-#584 := [def-axiom]: #583
-#680 := [unit-resolution #584 #657]: #234
-#682 := [th-lemma #680 #400]: #681
-#684 := [symm #682]: #683
-#688 := [monotonicity #684]: #687
-#690 := [symm #688]: #689
-#592 := (= f5 #591)
-#596 := (or #595 #592)
-#597 := [quant-inst]: #596
-#686 := [unit-resolution #597 #552]: #592
-#694 := [trans #686 #690]: #693
-#695 := [trans #694 #692]: #49
-#576 := (not #49)
-#58 := (f6 f13 #50)
-#284 := (ite #49 #32 #58)
-#472 := (* -1::real #284)
-#637 := (+ #32 #472)
-#638 := (<= #637 0::real)
-#585 := (= #32 #284)
-#661 := [hypothesis]: #49
-#587 := (or #576 #585)
-#588 := [def-axiom]: #587
-#662 := [unit-resolution #588 #661]: #585
-#663 := (not #585)
-#664 := (or #663 #638)
-#665 := [th-lemma]: #664
-#666 := [unit-resolution #665 #662]: #638
-#61 := (f6 f10 #50)
-#368 := (* -1::real #61)
-#384 := (+ #51 #368)
-#385 := (<= #384 0::real)
-#386 := (not #385)
-#369 := (+ #58 #368)
-#367 := (>= #369 0::real)
-#366 := (not #367)
-#391 := (and #366 #386)
-#63 := (f6 f9 #50)
-#68 := (< #43 #43)
-#71 := (ite #68 #63 #51)
-#72 := (< #61 #71)
-#69 := (ite #68 #51 #58)
-#70 := (< #69 #61)
-#73 := (and #70 #72)
-#392 := (iff #73 #391)
-#389 := (iff #72 #386)
-#381 := (< #61 #51)
-#387 := (iff #381 #386)
-#388 := [rewrite]: #387
-#382 := (iff #72 #381)
-#379 := (= #71 #51)
-#374 := (ite false #63 #51)
-#377 := (= #374 #51)
-#378 := [rewrite]: #377
-#375 := (= #71 #374)
-#354 := (iff #68 false)
-#355 := [rewrite]: #354
-#376 := [monotonicity #355]: #375
-#380 := [trans #376 #378]: #379
-#383 := [monotonicity #380]: #382
-#390 := [trans #383 #388]: #389
-#372 := (iff #70 #366)
-#363 := (< #58 #61)
-#370 := (iff #363 #366)
-#371 := [rewrite]: #370
-#364 := (iff #70 #363)
-#361 := (= #69 #58)
-#356 := (ite false #51 #58)
-#359 := (= #356 #58)
-#360 := [rewrite]: #359
-#357 := (= #69 #356)
-#358 := [monotonicity #355]: #357
-#362 := [trans #358 #360]: #361
-#365 := [monotonicity #362]: #364
-#373 := [trans #365 #371]: #372
-#393 := [monotonicity #373 #390]: #392
-#351 := [asserted]: #73
-#394 := [mp #351 #393]: #391
-#396 := [and-elim #394]: #386
-#402 := (* -1::real #63)
-#403 := (+ #51 #402)
-#404 := (<= #403 0::real)
-#414 := (* -1::real #58)
-#415 := (+ #51 #414)
-#413 := (>= #415 0::real)
-#64 := (f6 f11 #50)
-#407 := (* -1::real #64)
-#408 := (+ #63 #407)
-#409 := (<= #408 0::real)
-#423 := (and #404 #409 #413)
-#77 := (<= #63 #64)
-#76 := (<= #51 #63)
-#78 := (and #76 #77)
-#75 := (<= #58 #51)
-#79 := (and #75 #78)
-#426 := (iff #79 #423)
-#417 := (and #404 #409)
-#420 := (and #413 #417)
-#424 := (iff #420 #423)
-#425 := [rewrite]: #424
-#421 := (iff #79 #420)
-#418 := (iff #78 #417)
-#410 := (iff #77 #409)
-#411 := [rewrite]: #410
-#405 := (iff #76 #404)
-#406 := [rewrite]: #405
-#419 := [monotonicity #406 #411]: #418
-#412 := (iff #75 #413)
-#416 := [rewrite]: #412
-#422 := [monotonicity #416 #419]: #421
-#427 := [trans #422 #425]: #426
-#353 := [asserted]: #79
-#428 := [mp #353 #427]: #423
-#429 := [and-elim #428]: #404
-#642 := (+ #32 #402)
-#644 := (>= #642 0::real)
-#641 := (= #32 #63)
-#671 := (= #63 #32)
-#669 := (= #50 #29)
-#667 := (= #29 #50)
-#668 := [monotonicity #661]: #667
-#670 := [symm #668]: #669
-#672 := [monotonicity #670]: #671
-#673 := [symm #672]: #641
-#674 := (not #641)
-#675 := (or #674 #644)
-#676 := [th-lemma]: #675
-#677 := [unit-resolution #676 #673]: #644
-#475 := (+ #61 #472)
-#478 := (<= #475 0::real)
-#451 := (not #478)
-#327 := (ite #301 #284 #51)
-#335 := (* -1::real #327)
-#336 := (+ #61 #335)
-#337 := (<= #336 0::real)
-#338 := (not #337)
-#452 := (iff #338 #451)
-#479 := (iff #337 #478)
-#476 := (= #336 #475)
-#473 := (= #335 #472)
-#470 := (= #327 #284)
-#1 := true
-#465 := (ite true #284 #51)
-#468 := (= #465 #284)
-#469 := [rewrite]: #468
-#466 := (= #327 #465)
-#457 := (iff #301 true)
-#458 := [iff-true #400]: #457
-#467 := [monotonicity #458]: #466
-#471 := [trans #467 #469]: #470
-#474 := [monotonicity #471]: #473
-#477 := [monotonicity #474]: #476
-#480 := [monotonicity #477]: #479
-#481 := [monotonicity #480]: #452
-#308 := (ite #301 #64 #63)
-#318 := (* -1::real #308)
-#319 := (+ #61 #318)
-#317 := (>= #319 0::real)
-#316 := (not #317)
-#343 := (and #316 #338)
-#56 := (< #43 #26)
-#65 := (ite #56 #63 #64)
-#66 := (< #61 #65)
-#57 := (= f12 f5)
-#59 := (ite #57 #32 #58)
-#60 := (ite #56 #51 #59)
-#62 := (< #60 #61)
-#67 := (and #62 #66)
-#346 := (iff #67 #343)
-#287 := (ite #56 #51 #284)
-#290 := (< #287 #61)
-#296 := (and #66 #290)
-#344 := (iff #296 #343)
-#341 := (iff #290 #338)
-#332 := (< #327 #61)
-#339 := (iff #332 #338)
-#340 := [rewrite]: #339
-#333 := (iff #290 #332)
-#330 := (= #287 #327)
-#302 := (not #301)
-#324 := (ite #302 #51 #284)
-#328 := (= #324 #327)
-#329 := [rewrite]: #328
-#325 := (= #287 #324)
-#303 := (iff #56 #302)
-#304 := [rewrite]: #303
-#326 := [monotonicity #304]: #325
-#331 := [trans #326 #329]: #330
-#334 := [monotonicity #331]: #333
-#342 := [trans #334 #340]: #341
-#322 := (iff #66 #316)
-#313 := (< #61 #308)
-#320 := (iff #313 #316)
-#321 := [rewrite]: #320
-#314 := (iff #66 #313)
-#311 := (= #65 #308)
-#305 := (ite #302 #63 #64)
-#309 := (= #305 #308)
-#310 := [rewrite]: #309
-#306 := (= #65 #305)
-#307 := [monotonicity #304]: #306
-#312 := [trans #307 #310]: #311
-#315 := [monotonicity #312]: #314
-#323 := [trans #315 #321]: #322
-#345 := [monotonicity #323 #342]: #344
-#299 := (iff #67 #296)
-#293 := (and #290 #66)
-#297 := (iff #293 #296)
-#298 := [rewrite]: #297
-#294 := (iff #67 #293)
-#291 := (iff #62 #290)
-#288 := (= #60 #287)
-#285 := (= #59 #284)
-#231 := (iff #57 #49)
-#283 := [rewrite]: #231
-#286 := [monotonicity #283]: #285
-#289 := [monotonicity #286]: #288
-#292 := [monotonicity #289]: #291
-#295 := [monotonicity #292]: #294
-#300 := [trans #295 #298]: #299
-#347 := [trans #300 #345]: #346
-#184 := [asserted]: #67
-#348 := [mp #184 #347]: #343
-#350 := [and-elim #348]: #338
-#482 := [mp #350 #481]: #451
-#678 := [th-lemma #482 #677 #429 #396 #666]: false
-#679 := [lemma #678]: #576
-[unit-resolution #679 #695]: false
-unsat
-ca942f6174c1f53254d5ef1b69b0e75f0d4027d4 208 0
-#2 := false
-#37 := 0::real
-decl f13 :: (-> S6 S7 real)
-decl f17 :: S7
-#32 := f17
-decl f18 :: S6
-#34 := f18
-#35 := (f13 f18 f17)
-decl f14 :: (-> S8 S9 S6)
-decl f16 :: S9
-#30 := f16
-decl f15 :: (-> S2 S8)
-decl f5 :: S2
-#11 := f5
-#29 := (f15 f5)
-#31 := (f14 #29 f16)
-#33 := (f13 #31 f17)
-#96 := -1::real
-#107 := (* -1::real #33)
-#108 := (+ #107 #35)
-#97 := (* -1::real #35)
-#98 := (+ #33 #97)
-#135 := (>= #98 0::real)
-#142 := (ite #135 #98 #108)
-#150 := (* -1::real #142)
-#383 := (+ #108 #150)
-#384 := (<= #383 0::real)
-#369 := (= #108 #142)
-#136 := (not #135)
-decl f3 :: S2
-#8 := f3
-#47 := (f15 f3)
-#48 := (f14 #47 f16)
-#49 := (f13 #48 f17)
-#172 := (* -1::real #49)
-decl f19 :: S6
-#41 := f19
-#42 := (f13 f19 f17)
-#173 := (+ #42 #172)
-#116 := (* -1::real #42)
-#163 := (+ #116 #49)
-#184 := (<= #173 0::real)
-#191 := (ite #184 #163 #173)
-#199 := (* -1::real #191)
-#382 := (+ #173 #199)
-#385 := (<= #382 0::real)
-#375 := (= #173 #191)
-#185 := (not #184)
-#386 := [hypothesis]: #184
-#394 := (or #136 #185)
-#125 := -1/3::real
-#126 := (* -1/3::real #42)
-#200 := (+ #126 #199)
-#123 := 1/3::real
-#124 := (* 1/3::real #35)
-#201 := (+ #124 #200)
-#202 := (<= #201 0::real)
-#203 := (not #202)
-#44 := 3::real
-#43 := (- #35 #42)
-#45 := (/ #43 3::real)
-#50 := (- #49 #42)
-#52 := (- #50)
-#51 := (< #50 0::real)
-#53 := (ite #51 #52 #50)
-#54 := (< #53 #45)
-#208 := (iff #54 #203)
-#127 := (+ #124 #126)
-#166 := (< #163 0::real)
-#178 := (ite #166 #173 #163)
-#181 := (< #178 #127)
-#206 := (iff #181 #203)
-#196 := (< #191 #127)
-#204 := (iff #196 #203)
-#205 := [rewrite]: #204
-#197 := (iff #181 #196)
-#194 := (= #178 #191)
-#188 := (ite #185 #173 #163)
-#192 := (= #188 #191)
-#193 := [rewrite]: #192
-#189 := (= #178 #188)
-#186 := (iff #166 #185)
-#187 := [rewrite]: #186
-#190 := [monotonicity #187]: #189
-#195 := [trans #190 #193]: #194
-#198 := [monotonicity #195]: #197
-#207 := [trans #198 #205]: #206
-#182 := (iff #54 #181)
-#130 := (= #45 #127)
-#117 := (+ #35 #116)
-#120 := (/ #117 3::real)
-#128 := (= #120 #127)
-#129 := [rewrite]: #128
-#121 := (= #45 #120)
-#118 := (= #43 #117)
-#119 := [rewrite]: #118
-#122 := [monotonicity #119]: #121
-#131 := [trans #122 #129]: #130
-#179 := (= #53 #178)
-#164 := (= #50 #163)
-#165 := [rewrite]: #164
-#176 := (= #52 #173)
-#169 := (- #163)
-#174 := (= #169 #173)
-#175 := [rewrite]: #174
-#170 := (= #52 #169)
-#171 := [monotonicity #165]: #170
-#177 := [trans #171 #175]: #176
-#167 := (iff #51 #166)
-#168 := [monotonicity #165]: #167
-#180 := [monotonicity #168 #177 #165]: #179
-#183 := [monotonicity #180 #131]: #182
-#209 := [trans #183 #207]: #208
-#162 := [asserted]: #54
-#210 := [mp #162 #209]: #203
-#380 := (+ #163 #199)
-#381 := (<= #380 0::real)
-#374 := (= #163 #191)
-#376 := (or #185 #374)
-#377 := [def-axiom]: #376
-#387 := [unit-resolution #377 #386]: #374
-#388 := (not #374)
-#389 := (or #388 #381)
-#390 := [th-lemma]: #389
-#391 := [unit-resolution #390 #387]: #381
-#392 := [hypothesis]: #135
-#214 := (+ #33 #172)
-#215 := (<= #214 0::real)
-#55 := (<= #33 #49)
-#216 := (iff #55 #215)
-#217 := [rewrite]: #216
-#211 := [asserted]: #55
-#218 := [mp #211 #217]: #215
-#393 := [th-lemma #218 #392 #391 #210 #386]: false
-#395 := [lemma #393]: #394
-#396 := [unit-resolution #395 #386]: #136
-#151 := (+ #126 #150)
-#152 := (+ #124 #151)
-#153 := (<= #152 0::real)
-#154 := (not #153)
-#36 := (- #33 #35)
-#39 := (- #36)
-#38 := (< #36 0::real)
-#40 := (ite #38 #39 #36)
-#46 := (< #40 #45)
-#159 := (iff #46 #154)
-#101 := (< #98 0::real)
-#113 := (ite #101 #108 #98)
-#132 := (< #113 #127)
-#157 := (iff #132 #154)
-#147 := (< #142 #127)
-#155 := (iff #147 #154)
-#156 := [rewrite]: #155
-#148 := (iff #132 #147)
-#145 := (= #113 #142)
-#139 := (ite #136 #108 #98)
-#143 := (= #139 #142)
-#144 := [rewrite]: #143
-#140 := (= #113 #139)
-#137 := (iff #101 #136)
-#138 := [rewrite]: #137
-#141 := [monotonicity #138]: #140
-#146 := [trans #141 #144]: #145
-#149 := [monotonicity #146]: #148
-#158 := [trans #149 #156]: #157
-#133 := (iff #46 #132)
-#114 := (= #40 #113)
-#99 := (= #36 #98)
-#100 := [rewrite]: #99
-#111 := (= #39 #108)
-#104 := (- #98)
-#109 := (= #104 #108)
-#110 := [rewrite]: #109
-#105 := (= #39 #104)
-#106 := [monotonicity #100]: #105
-#112 := [trans #106 #110]: #111
-#102 := (iff #38 #101)
-#103 := [monotonicity #100]: #102
-#115 := [monotonicity #103 #112 #100]: #114
-#134 := [monotonicity #115 #131]: #133
-#160 := [trans #134 #158]: #159
-#95 := [asserted]: #46
-#161 := [mp #95 #160]: #154
-#372 := (or #135 #369)
-#373 := [def-axiom]: #372
-#397 := [unit-resolution #373 #396]: #369
-#398 := (not #369)
-#399 := (or #398 #384)
-#400 := [th-lemma]: #399
-#401 := [unit-resolution #400 #397]: #384
-#402 := [th-lemma #401 #161 #391 #210 #218 #396]: false
-#403 := [lemma #402]: #185
-#378 := (or #184 #375)
-#379 := [def-axiom]: #378
-#406 := [unit-resolution #379 #403]: #375
-#407 := (not #375)
-#408 := (or #407 #385)
-#409 := [th-lemma]: #408
-#410 := [unit-resolution #409 #406]: #385
-#412 := (not #215)
-#411 := (not #385)
-#413 := (or #136 #411 #412 #202 #184)
-#414 := [th-lemma]: #413
-#415 := [unit-resolution #414 #403 #210 #218 #410]: #136
-#416 := [unit-resolution #373 #415]: #369
-#417 := [unit-resolution #400 #416]: #384
-[th-lemma #410 #218 #210 #403 #161 #417]: false
-unsat
-504ce5f4f6961a0f59840c0aa303f063d46936a5 333 0
+96c3a1d2dea133da23272bb1832adcfffc19c650 333 0
 #2 := false
 #11 := 0::real
 decl ?v2!1 :: real
@@ -2304,173 +571,7 @@
 #408 := [lemma #407]: #230
 [th-lemma #406 #301 #408 #396 #303]: false
 unsat
-024080ea9e6de105c72225d6d55cc8b136a93933 165 0
-#2 := false
-#22 := 0::real
-decl f3 :: (-> S3 S2 real)
-decl ?v0!0 :: S2
-#118 := ?v0!0
-decl f5 :: S3
-#11 := f5
-#119 := (f3 f5 ?v0!0)
-#49 := -1::real
-#117 := (* -1::real #119)
-decl f4 :: S3
-#8 := f4
-#115 := (f3 f4 ?v0!0)
-#120 := (+ #115 #117)
-#121 := (>= #120 0::real)
-decl f6 :: (-> S2 S4 S1)
-decl f7 :: (-> S4 S4 S4)
-decl f9 :: (-> S2 S4 S4)
-decl f11 :: S4
-#17 := f11
-decl f10 :: S2
-#16 := f10
-#18 := (f9 f10 f11)
-decl f8 :: S4
-#15 := f8
-#19 := (f7 f8 #18)
-#123 := (f6 ?v0!0 #19)
-decl f1 :: S1
-#4 := f1
-#124 := (= f1 #123)
-#9 := (:var 0 S2)
-#12 := (f3 f5 #9)
-#81 := (* -1::real #12)
-#10 := (f3 f4 #9)
-#82 := (+ #10 #81)
-#83 := (>= #82 0::real)
-#84 := (not #83)
-#89 := (forall (vars (?v1 S2)) #84)
-#158 := (and #89 #121 #124)
-#122 := (not #121)
-#133 := (not #122)
-#125 := (not #124)
-#130 := (not #125)
-#143 := (and #130 #133 #89)
-#161 := (iff #143 #158)
-#155 := (and #124 #121 #89)
-#159 := (iff #155 #158)
-#160 := [rewrite]: #159
-#156 := (iff #143 #155)
-#153 := (iff #133 #121)
-#154 := [rewrite]: #153
-#151 := (iff #130 #124)
-#152 := [rewrite]: #151
-#157 := [monotonicity #152 #154]: #156
-#162 := [trans #157 #160]: #161
-#92 := (not #89)
-#20 := (f6 #9 #19)
-#46 := (= f1 #20)
-#60 := (not #46)
-#101 := (or #60 #84 #92)
-#106 := (forall (vars (?v0 S2)) #101)
-#109 := (not #106)
-#146 := (~ #109 #143)
-#126 := (or #125 #122 #92)
-#127 := (not #126)
-#144 := (~ #127 #143)
-#140 := (not #92)
-#141 := (~ #140 #89)
-#138 := (~ #89 #89)
-#136 := (~ #84 #84)
-#137 := [refl]: #136
-#139 := [nnf-pos #137]: #138
-#142 := [nnf-neg #139]: #141
-#134 := (~ #133 #133)
-#135 := [refl]: #134
-#131 := (~ #130 #130)
-#132 := [refl]: #131
-#145 := [nnf-neg #132 #135 #142]: #144
-#128 := (~ #109 #127)
-#129 := [sk]: #128
-#147 := [trans #129 #145]: #146
-#23 := (- #12 #10)
-#24 := (< 0::real #23)
-#21 := (= #20 f1)
-#25 := (implies #21 #24)
-#13 := (< #10 #12)
-#14 := (forall (vars (?v1 S2)) #13)
-#26 := (implies #14 #25)
-#27 := (forall (vars (?v0 S2)) #26)
-#28 := (not #27)
-#112 := (iff #28 #109)
-#50 := (* -1::real #10)
-#51 := (+ #50 #12)
-#54 := (< 0::real #51)
-#61 := (or #60 #54)
-#69 := (not #14)
-#70 := (or #69 #61)
-#75 := (forall (vars (?v0 S2)) #70)
-#78 := (not #75)
-#110 := (iff #78 #109)
-#107 := (iff #75 #106)
-#104 := (iff #70 #101)
-#95 := (or #60 #84)
-#98 := (or #92 #95)
-#102 := (iff #98 #101)
-#103 := [rewrite]: #102
-#99 := (iff #70 #98)
-#96 := (iff #61 #95)
-#85 := (iff #54 #84)
-#86 := [rewrite]: #85
-#97 := [monotonicity #86]: #96
-#93 := (iff #69 #92)
-#90 := (iff #14 #89)
-#87 := (iff #13 #84)
-#88 := [rewrite]: #87
-#91 := [quant-intro #88]: #90
-#94 := [monotonicity #91]: #93
-#100 := [monotonicity #94 #97]: #99
-#105 := [trans #100 #103]: #104
-#108 := [quant-intro #105]: #107
-#111 := [monotonicity #108]: #110
-#79 := (iff #28 #78)
-#76 := (iff #27 #75)
-#73 := (iff #26 #70)
-#66 := (implies #14 #61)
-#71 := (iff #66 #70)
-#72 := [rewrite]: #71
-#67 := (iff #26 #66)
-#64 := (iff #25 #61)
-#57 := (implies #46 #54)
-#62 := (iff #57 #61)
-#63 := [rewrite]: #62
-#58 := (iff #25 #57)
-#55 := (iff #24 #54)
-#52 := (= #23 #51)
-#53 := [rewrite]: #52
-#56 := [monotonicity #53]: #55
-#47 := (iff #21 #46)
-#48 := [rewrite]: #47
-#59 := [monotonicity #48 #56]: #58
-#65 := [trans #59 #63]: #64
-#68 := [monotonicity #65]: #67
-#74 := [trans #68 #72]: #73
-#77 := [quant-intro #74]: #76
-#80 := [monotonicity #77]: #79
-#113 := [trans #80 #111]: #112
-#45 := [asserted]: #28
-#114 := [mp #45 #113]: #109
-#148 := [mp~ #114 #147]: #143
-#149 := [mp #148 #162]: #158
-#163 := [and-elim #149]: #121
-#223 := (pattern #12)
-#222 := (pattern #10)
-#224 := (forall (vars (?v1 S2)) (:pat #222 #223) #84)
-#227 := (iff #89 #224)
-#225 := (iff #84 #84)
-#226 := [refl]: #225
-#228 := [quant-intro #226]: #227
-#150 := [and-elim #149]: #89
-#229 := [mp #150 #228]: #224
-#232 := (not #224)
-#233 := (or #232 #122)
-#234 := [quant-inst]: #233
-[unit-resolution #234 #229 #163]: false
-unsat
-116b1dd4c85396a326f34f6c1266b1ad85116049 57 0
+d083434592d4f729dd1c100279aa204b663fa35c 57 0
 #2 := false
 decl f13 :: (-> S4 S4 S5)
 #44 := (:var 0 S4)
@@ -2528,810 +629,827 @@
 #264 := [quant-inst]: #263
 [unit-resolution #264 #269 #258]: false
 unsat
-74073317ccefcdf35878e5154f8155d12c8475cf 91 0
+93a3d21a7f84dc839f14b9a0cfcafdfff6c6b0f1 318 0
 #2 := false
-#43 := 0::real
-decl f3 :: (-> S2 S3 real)
-decl f5 :: S3
-#9 := f5
-decl f6 :: S2
-#11 := f6
-#12 := (f3 f6 f5)
-#40 := -1::real
-#41 := (* -1::real #12)
-decl f4 :: S2
-#8 := f4
-#10 := (f3 f4 f5)
-#42 := (+ #10 #41)
-#135 := (>= #42 0::real)
-#160 := (not #135)
-#48 := (= #10 #12)
-#60 := (not #48)
-#19 := (= #12 #10)
-#20 := (not #19)
-#61 := (iff #20 #60)
-#58 := (iff #19 #48)
-#59 := [rewrite]: #58
-#62 := [monotonicity #59]: #61
-#39 := [asserted]: #20
-#65 := [mp #39 #62]: #60
-#163 := (or #48 #160)
-#44 := (<= #42 0::real)
-#13 := (<= #10 #12)
-#45 := (iff #13 #44)
-#46 := [rewrite]: #45
-#37 := [asserted]: #13
-#47 := [mp #37 #46]: #44
-#159 := (not #44)
-#161 := (or #48 #159 #160)
-#162 := [th-lemma]: #161
-#164 := [unit-resolution #162 #47]: #163
-#165 := [unit-resolution #164 #65]: #160
-#14 := (:var 0 S3)
-#16 := (f3 f4 #14)
-#128 := (pattern #16)
-#15 := (f3 f6 #14)
-#127 := (pattern #15)
-#49 := (* -1::real #16)
-#50 := (+ #15 #49)
-#51 := (<= #50 0::real)
-#129 := (forall (vars (?v0 S3)) (:pat #127 #128) #51)
-#54 := (forall (vars (?v0 S3)) #51)
-#132 := (iff #54 #129)
-#130 := (iff #51 #51)
-#131 := [refl]: #130
-#133 := [quant-intro #131]: #132
-#69 := (~ #54 #54)
-#71 := (~ #51 #51)
-#72 := [refl]: #71
-#70 := [nnf-pos #72]: #69
-#17 := (<= #15 #16)
-#18 := (forall (vars (?v0 S3)) #17)
-#55 := (iff #18 #54)
-#52 := (iff #17 #51)
-#53 := [rewrite]: #52
-#56 := [quant-intro #53]: #55
-#38 := [asserted]: #18
-#57 := [mp #38 #56]: #54
-#67 := [mp~ #57 #70]: #54
-#134 := [mp #67 #133]: #129
-#149 := (not #129)
-#150 := (or #149 #135)
-#136 := (* -1::real #10)
-#137 := (+ #12 #136)
-#138 := (<= #137 0::real)
-#151 := (or #149 #138)
-#153 := (iff #151 #150)
-#155 := (iff #150 #150)
-#156 := [rewrite]: #155
-#147 := (iff #138 #135)
-#139 := (+ #136 #12)
-#142 := (<= #139 0::real)
-#145 := (iff #142 #135)
-#146 := [rewrite]: #145
-#143 := (iff #138 #142)
-#140 := (= #137 #139)
-#141 := [rewrite]: #140
-#144 := [monotonicity #141]: #143
-#148 := [trans #144 #146]: #147
-#154 := [monotonicity #148]: #153
-#157 := [trans #154 #156]: #153
-#152 := [quant-inst]: #151
-#158 := [mp #152 #157]: #150
-[unit-resolution #158 #134 #165]: false
-unsat
-9ecd5f8eb0c8f78bd68a366175093e04632f1f73 149 0
-#2 := false
-#23 := 0::real
-decl f3 :: (-> S2 S3 real)
-decl f5 :: S3
-#9 := f5
-decl f6 :: S2
-#11 := f6
-#12 := (f3 f6 f5)
-#49 := -1::real
-#161 := (* -1::real #12)
-decl f4 :: S2
-#8 := f4
-#10 := (f3 f4 f5)
-#208 := (+ #10 #161)
-#210 := (>= #208 0::real)
-#13 := (= #10 #12)
-#45 := [asserted]: #13
-#213 := (not #13)
-#214 := (or #213 #210)
-#215 := [th-lemma]: #214
-#216 := [unit-resolution #215 #45]: #210
-decl f7 :: S2
-#16 := f7
-#26 := (f3 f7 f5)
-#165 := (* -1::real #26)
-#166 := (+ #10 #165)
-#212 := (>= #166 0::real)
-#227 := (not #212)
-#211 := (= #10 #26)
-#221 := (not #211)
-#67 := (= #12 #26)
-#75 := (not #67)
-#222 := (iff #75 #221)
-#219 := (iff #67 #211)
-#217 := (iff #211 #67)
-#218 := [monotonicity #45]: #217
-#220 := [symm #218]: #219
-#223 := [monotonicity #220]: #222
-#27 := (= #26 #12)
-#28 := (not #27)
-#76 := (iff #28 #75)
-#73 := (iff #27 #67)
-#74 := [rewrite]: #73
-#77 := [monotonicity #74]: #76
-#48 := [asserted]: #28
-#80 := [mp #48 #77]: #75
-#224 := [mp #80 #223]: #221
-#230 := (or #211 #227)
-#167 := (<= #166 0::real)
-#177 := (+ #12 #165)
-#178 := (>= #177 0::real)
-#183 := (not #178)
+#25 := 0::real
+decl f8 :: (-> S4 S2 real)
+decl f11 :: S2
+#40 := f11
+decl f9 :: S4
+#32 := f9
+#47 := (f8 f9 f11)
+decl f12 :: S4
+#44 := f12
+#45 := (f8 f12 f11)
+#192 := -1::real
+#232 := (* -1::real #45)
+#233 := (+ #232 #47)
+decl f5 :: real
+#26 := f5
+#268 := (* -1::real #47)
+#271 := (+ #45 #268)
+#274 := (+ f5 #271)
+#275 := (<= #274 0::real)
+#278 := (ite #275 f5 #233)
+#593 := (* -1::real #278)
+#594 := (+ f5 #593)
+#595 := (<= #594 0::real)
+#603 := (not #595)
+#223 := 1/2::real
+#281 := (* 1/2::real #278)
+#457 := (<= #281 0::real)
+#292 := (= #281 0::real)
+#306 := (<= #271 0::real)
+decl f10 :: S4
+#34 := f10
+#43 := (f8 f10 f11)
+#302 := (+ #43 #232)
+#303 := (<= #302 0::real)
+#309 := (and #303 #306)
+#13 := 0::int
+decl f4 :: (-> S2 int)
+#41 := (f4 f11)
+#178 := -1::int
+#213 := (* -1::int #41)
+decl f6 :: (-> S3 S2)
+decl f7 :: S3
+#28 := f7
+#29 := (f6 f7)
+#30 := (f4 #29)
+#214 := (+ #30 #213)
+#215 := (<= #214 0::int)
+#312 := (or #215 #309)
+#225 := (* 1/2::real #47)
+#272 := (+ #232 #225)
+#224 := (* 1/2::real #43)
+#273 := (+ #224 #272)
+#270 := (>= #273 0::real)
+#321 := (and #270 #292 #312)
+#52 := 2::real
+#55 := (- #47 #45)
+#56 := (<= f5 #55)
+#57 := (ite #56 f5 #55)
+#58 := (/ #57 2::real)
+#59 := (+ #45 #58)
+#60 := (= #59 #45)
+#51 := (+ #43 #47)
+#53 := (/ #51 2::real)
+#54 := (<= #45 #53)
+#61 := (and #54 #60)
+#48 := (<= #45 #47)
+#46 := (<= #43 #45)
+#49 := (and #46 #48)
+#42 := (< #41 #30)
+#50 := (implies #42 #49)
+#62 := (and #50 #61)
+#326 := (iff #62 #321)
+#236 := (<= f5 #233)
+#239 := (ite #236 f5 #233)
+#245 := (* 1/2::real #239)
+#250 := (+ #45 #245)
+#256 := (= #45 #250)
+#226 := (+ #224 #225)
+#229 := (<= #45 #226)
+#261 := (and #229 #256)
+#212 := (not #42)
+#220 := (or #212 #49)
+#264 := (and #220 #261)
+#324 := (iff #264 #321)
+#315 := (and #270 #292)
+#318 := (and #312 #315)
+#322 := (iff #318 #321)
+#323 := [rewrite]: #322
+#319 := (iff #264 #318)
+#316 := (iff #261 #315)
+#293 := (iff #256 #292)
+#284 := (+ #45 #281)
+#287 := (= #45 #284)
+#290 := (iff #287 #292)
+#291 := [rewrite]: #290
+#288 := (iff #256 #287)
+#285 := (= #250 #284)
+#282 := (= #245 #281)
+#279 := (= #239 #278)
+#276 := (iff #236 #275)
+#277 := [rewrite]: #276
+#280 := [monotonicity #277]: #279
+#283 := [monotonicity #280]: #282
+#286 := [monotonicity #283]: #285
+#289 := [monotonicity #286]: #288
+#294 := [trans #289 #291]: #293
+#267 := (iff #229 #270)
+#269 := [rewrite]: #267
+#317 := [monotonicity #269 #294]: #316
+#313 := (iff #220 #312)
+#310 := (iff #49 #309)
+#307 := (iff #48 #306)
+#308 := [rewrite]: #307
+#304 := (iff #46 #303)
+#305 := [rewrite]: #304
+#311 := [monotonicity #305 #308]: #310
+#300 := (iff #212 #215)
+#216 := (not #215)
+#295 := (not #216)
+#298 := (iff #295 #215)
+#299 := [rewrite]: #298
+#296 := (iff #212 #295)
+#217 := (iff #42 #216)
+#218 := [rewrite]: #217
+#297 := [monotonicity #218]: #296
+#301 := [trans #297 #299]: #300
+#314 := [monotonicity #301 #311]: #313
+#320 := [monotonicity #314 #317]: #319
+#325 := [trans #320 #323]: #324
+#265 := (iff #62 #264)
+#262 := (iff #61 #261)
+#259 := (iff #60 #256)
+#253 := (= #250 #45)
+#257 := (iff #253 #256)
+#258 := [rewrite]: #257
+#254 := (iff #60 #253)
+#251 := (= #59 #250)
+#248 := (= #58 #245)
+#242 := (/ #239 2::real)
+#246 := (= #242 #245)
+#247 := [rewrite]: #246
+#243 := (= #58 #242)
+#240 := (= #57 #239)
+#234 := (= #55 #233)
+#235 := [rewrite]: #234
+#237 := (iff #56 #236)
+#238 := [monotonicity #235]: #237
+#241 := [monotonicity #238 #235]: #240
+#244 := [monotonicity #241]: #243
+#249 := [trans #244 #247]: #248
+#252 := [monotonicity #249]: #251
+#255 := [monotonicity #252]: #254
+#260 := [trans #255 #258]: #259
+#230 := (iff #54 #229)
+#227 := (= #53 #226)
+#228 := [rewrite]: #227
+#231 := [monotonicity #228]: #230
+#263 := [monotonicity #231 #260]: #262
+#221 := (iff #50 #220)
+#222 := [rewrite]: #221
+#266 := [monotonicity #222 #263]: #265
+#327 := [trans #266 #325]: #326
+#211 := [asserted]: #62
+#328 := [mp #211 #327]: #321
+#330 := [and-elim #328]: #292
+#597 := (not #292)
+#598 := (or #597 #457)
+#599 := [th-lemma]: #598
+#600 := [unit-resolution #599 #330]: #457
+#601 := [hypothesis]: #595
+#167 := (<= f5 0::real)
 #168 := (not #167)
-#186 := (or #168 #183)
-#189 := (not #186)
-#14 := (:var 0 S3)
-#19 := (f3 f6 #14)
-#154 := (pattern #19)
-#17 := (f3 f7 #14)
-#153 := (pattern #17)
-#15 := (f3 f4 #14)
-#152 := (pattern #15)
-#55 := (* -1::real #19)
-#56 := (+ #17 #55)
-#57 := (<= #56 0::real)
-#82 := (not #57)
-#50 := (* -1::real #17)
-#51 := (+ #15 #50)
-#52 := (<= #51 0::real)
-#85 := (not #52)
-#83 := (or #85 #82)
-#81 := (not #83)
-#155 := (forall (vars (?v0 S3)) (:pat #152 #153 #154) #81)
-#91 := (forall (vars (?v0 S3)) #81)
-#158 := (iff #91 #155)
-#156 := (iff #81 #81)
-#157 := [refl]: #156
-#159 := [quant-intro #157]: #158
-#60 := (and #52 #57)
-#63 := (forall (vars (?v0 S3)) #60)
-#92 := (iff #63 #91)
-#78 := (iff #60 #81)
-#90 := [rewrite]: #78
-#93 := [quant-intro #90]: #92
-#86 := (~ #63 #63)
-#88 := (~ #60 #60)
-#89 := [refl]: #88
-#87 := [nnf-pos #89]: #86
-#20 := (<= #17 #19)
-#18 := (<= #15 #17)
-#21 := (and #18 #20)
-#22 := (forall (vars (?v0 S3)) #21)
-#64 := (iff #22 #63)
-#61 := (iff #21 #60)
-#58 := (iff #20 #57)
-#59 := [rewrite]: #58
-#53 := (iff #18 #52)
-#54 := [rewrite]: #53
-#62 := [monotonicity #54 #59]: #61
-#65 := [quant-intro #62]: #64
-#46 := [asserted]: #22
-#66 := [mp #46 #65]: #63
-#84 := [mp~ #66 #87]: #63
-#94 := [mp #84 #93]: #91
-#160 := [mp #94 #159]: #155
-#192 := (not #155)
-#193 := (or #192 #189)
-#162 := (+ #26 #161)
-#163 := (<= #162 0::real)
-#164 := (not #163)
-#169 := (or #168 #164)
-#170 := (not #169)
-#194 := (or #192 #170)
-#196 := (iff #194 #193)
-#198 := (iff #193 #193)
-#199 := [rewrite]: #198
-#190 := (iff #170 #189)
-#187 := (iff #169 #186)
-#184 := (iff #164 #183)
-#181 := (iff #163 #178)
-#171 := (+ #161 #26)
-#174 := (<= #171 0::real)
-#179 := (iff #174 #178)
-#180 := [rewrite]: #179
-#175 := (iff #163 #174)
-#172 := (= #162 #171)
-#173 := [rewrite]: #172
-#176 := [monotonicity #173]: #175
-#182 := [trans #176 #180]: #181
-#185 := [monotonicity #182]: #184
-#188 := [monotonicity #185]: #187
-#191 := [monotonicity #188]: #190
-#197 := [monotonicity #191]: #196
-#200 := [trans #197 #199]: #196
-#195 := [quant-inst]: #194
-#201 := [mp #195 #200]: #193
-#225 := [unit-resolution #201 #160]: #189
-#202 := (or #186 #167)
-#203 := [def-axiom]: #202
-#226 := [unit-resolution #203 #225]: #167
-#228 := (or #211 #168 #227)
-#229 := [th-lemma]: #228
-#231 := [unit-resolution #229 #226]: #230
-#232 := [unit-resolution #231 #224]: #227
-#204 := (or #186 #178)
-#205 := [def-axiom]: #204
-#233 := [unit-resolution #205 #225]: #178
-[th-lemma #233 #232 #216]: false
-unsat
-ada412db5ba79d588ff49226c319d0dae76f5f87 271 0
-#2 := false
-#8 := 0::real
-decl f4 :: (-> S3 S2 real)
-decl f7 :: S2
-#19 := f7
-decl f5 :: S3
-#11 := f5
-#24 := (f4 f5 f7)
-decl f8 :: S3
-#21 := f8
-#22 := (f4 f8 f7)
-#66 := -1::real
-#87 := (* -1::real #22)
-#88 := (+ #87 #24)
-decl f3 :: real
-#9 := f3
-#148 := (* -1::real #24)
-#149 := (+ #22 #148)
-#150 := (+ f3 #149)
-#151 := (<= #150 0::real)
-#154 := (ite #151 f3 #88)
-#320 := (* -1::real #154)
-#321 := (+ f3 #320)
-#322 := (<= #321 0::real)
-#329 := (not #322)
-#65 := 1/2::real
-#157 := (* 1/2::real #154)
-#289 := (<= #157 0::real)
-#168 := (= #157 0::real)
-#178 := (<= #149 0::real)
-decl f6 :: S3
-#14 := f6
-#20 := (f4 f6 f7)
-#174 := (+ #20 #87)
-#175 := (<= #174 0::real)
-#181 := (and #175 #178)
-#184 := (not #181)
-#171 := (not #168)
-#80 := (* 1/2::real #24)
-#145 := (+ #87 #80)
-#79 := (* 1/2::real #20)
-#146 := (+ #79 #145)
-#143 := (>= #146 0::real)
-#141 := (not #143)
-#193 := (or #141 #171 #184)
-#198 := (not #193)
-#28 := 2::real
-#31 := (- #24 #22)
-#32 := (<= f3 #31)
-#33 := (ite #32 f3 #31)
-#34 := (/ #33 2::real)
-#35 := (+ #22 #34)
-#36 := (= #35 #22)
+#27 := (< 0::real f5)
+#169 := (iff #27 #168)
+#170 := [rewrite]: #169
+#164 := [asserted]: #27
+#171 := [mp #164 #170]: #168
+#602 := [th-lemma #171 #601 #600]: false
+#604 := [lemma #602]: #603
+#450 := (= f5 #278)
+#451 := (= #233 #278)
+#613 := (not #451)
+#596 := (+ #233 #593)
+#605 := (<= #596 0::real)
+#610 := (not #605)
+#532 := (+ #43 #268)
+#533 := (>= #532 0::real)
+#538 := (not #533)
+#210 := [asserted]: #42
+#219 := [mp #210 #218]: #216
+#8 := (:var 0 S2)
+#35 := (f8 f10 #8)
+#443 := (pattern #35)
+#33 := (f8 f9 #8)
+#442 := (pattern #33)
+#9 := (f4 #8)
+#422 := (pattern #9)
+#193 := (* -1::real #35)
+#194 := (+ #33 #193)
+#195 := (<= #194 0::real)
+#198 := (not #195)
+#181 := (* -1::int #30)
+#182 := (+ #9 #181)
+#180 := (>= #182 0::int)
+#201 := (or #180 #198)
+#444 := (forall (vars (?v0 S2)) (:pat #422 #442 #443) #201)
+#204 := (forall (vars (?v0 S2)) #201)
+#447 := (iff #204 #444)
+#445 := (iff #201 #201)
+#446 := [refl]: #445
+#448 := [quant-intro #446]: #447
+#362 := (~ #204 #204)
+#332 := (~ #201 #201)
+#361 := [refl]: #332
+#363 := [nnf-pos #361]: #362
+#36 := (<= #33 #35)
 #37 := (not #36)
-#27 := (+ #20 #24)
-#29 := (/ #27 2::real)
-#30 := (<= #22 #29)
-#38 := (implies #30 #37)
-#25 := (<= #22 #24)
-#23 := (<= #20 #22)
-#26 := (and #23 #25)
-#39 := (implies #26 #38)
-#40 := (not #39)
-#201 := (iff #40 #198)
-#91 := (<= f3 #88)
-#94 := (ite #91 f3 #88)
-#100 := (* 1/2::real #94)
-#105 := (+ #22 #100)
-#111 := (= #22 #105)
-#116 := (not #111)
-#81 := (+ #79 #80)
-#84 := (<= #22 #81)
-#122 := (not #84)
-#123 := (or #122 #116)
-#131 := (not #26)
-#132 := (or #131 #123)
-#137 := (not #132)
-#199 := (iff #137 #198)
-#196 := (iff #132 #193)
-#187 := (or #141 #171)
-#190 := (or #184 #187)
-#194 := (iff #190 #193)
-#195 := [rewrite]: #194
-#191 := (iff #132 #190)
-#188 := (iff #123 #187)
-#172 := (iff #116 #171)
-#169 := (iff #111 #168)
-#160 := (+ #22 #157)
-#163 := (= #22 #160)
-#166 := (iff #163 #168)
-#167 := [rewrite]: #166
-#164 := (iff #111 #163)
-#161 := (= #105 #160)
-#158 := (= #100 #157)
-#155 := (= #94 #154)
-#152 := (iff #91 #151)
-#153 := [rewrite]: #152
-#156 := [monotonicity #153]: #155
-#159 := [monotonicity #156]: #158
-#162 := [monotonicity #159]: #161
-#165 := [monotonicity #162]: #164
-#170 := [trans #165 #167]: #169
-#173 := [monotonicity #170]: #172
-#144 := (iff #122 #141)
-#140 := (iff #84 #143)
-#142 := [rewrite]: #140
-#147 := [monotonicity #142]: #144
-#189 := [monotonicity #147 #173]: #188
-#185 := (iff #131 #184)
-#182 := (iff #26 #181)
-#179 := (iff #25 #178)
-#180 := [rewrite]: #179
-#176 := (iff #23 #175)
-#177 := [rewrite]: #176
-#183 := [monotonicity #177 #180]: #182
-#186 := [monotonicity #183]: #185
-#192 := [monotonicity #186 #189]: #191
-#197 := [trans #192 #195]: #196
+#31 := (< #9 #30)
+#38 := (implies #31 #37)
+#39 := (forall (vars (?v0 S2)) #38)
+#207 := (iff #39 #204)
+#166 := (not #31)
+#172 := (or #166 #37)
+#175 := (forall (vars (?v0 S2)) #172)
+#205 := (iff #175 #204)
+#202 := (iff #172 #201)
+#199 := (iff #37 #198)
+#196 := (iff #36 #195)
+#197 := [rewrite]: #196
 #200 := [monotonicity #197]: #199
-#138 := (iff #40 #137)
-#135 := (iff #39 #132)
-#128 := (implies #26 #123)
-#133 := (iff #128 #132)
-#134 := [rewrite]: #133
-#129 := (iff #39 #128)
-#126 := (iff #38 #123)
-#119 := (implies #84 #116)
-#124 := (iff #119 #123)
-#125 := [rewrite]: #124
-#120 := (iff #38 #119)
-#117 := (iff #37 #116)
-#114 := (iff #36 #111)
-#108 := (= #105 #22)
-#112 := (iff #108 #111)
-#113 := [rewrite]: #112
-#109 := (iff #36 #108)
-#106 := (= #35 #105)
-#103 := (= #34 #100)
-#97 := (/ #94 2::real)
-#101 := (= #97 #100)
-#102 := [rewrite]: #101
-#98 := (= #34 #97)
-#95 := (= #33 #94)
-#89 := (= #31 #88)
-#90 := [rewrite]: #89
-#92 := (iff #32 #91)
-#93 := [monotonicity #90]: #92
-#96 := [monotonicity #93 #90]: #95
-#99 := [monotonicity #96]: #98
-#104 := [trans #99 #102]: #103
-#107 := [monotonicity #104]: #106
-#110 := [monotonicity #107]: #109
-#115 := [trans #110 #113]: #114
-#118 := [monotonicity #115]: #117
-#85 := (iff #30 #84)
-#82 := (= #29 #81)
-#83 := [rewrite]: #82
-#86 := [monotonicity #83]: #85
-#121 := [monotonicity #86 #118]: #120
-#127 := [trans #121 #125]: #126
-#130 := [monotonicity #127]: #129
-#136 := [trans #130 #134]: #135
-#139 := [monotonicity #136]: #138
-#202 := [trans #139 #200]: #201
-#59 := [asserted]: #40
-#203 := [mp #59 #202]: #198
-#205 := [not-or-elim #203]: #168
-#324 := (or #171 #289)
-#325 := [th-lemma]: #324
-#326 := [unit-resolution #325 #205]: #289
-#327 := [hypothesis]: #322
-#60 := (<= f3 0::real)
-#61 := (not #60)
-#10 := (< 0::real f3)
-#62 := (iff #10 #61)
-#63 := [rewrite]: #62
-#57 := [asserted]: #10
-#64 := [mp #57 #63]: #61
-#328 := [th-lemma #64 #327 #326]: false
-#330 := [lemma #328]: #329
-#282 := (= f3 #154)
-#283 := (= #88 #154)
-#339 := (not #283)
-#323 := (+ #88 #320)
-#331 := (<= #323 0::real)
-#336 := (not #331)
-#301 := (+ #20 #148)
-#302 := (>= #301 0::real)
-#307 := (not #302)
-#12 := (:var 0 S2)
-#15 := (f4 f6 #12)
-#275 := (pattern #15)
-#13 := (f4 f5 #12)
-#274 := (pattern #13)
-#67 := (* -1::real #15)
-#68 := (+ #13 #67)
-#69 := (<= #68 0::real)
-#218 := (not #69)
-#276 := (forall (vars (?v0 S2)) (:pat #274 #275) #218)
-#223 := (forall (vars (?v0 S2)) #218)
-#279 := (iff #223 #276)
-#277 := (iff #218 #218)
-#278 := [refl]: #277
-#280 := [quant-intro #278]: #279
-#72 := (exists (vars (?v0 S2)) #69)
-#75 := (not #72)
-#220 := (~ #75 #223)
-#219 := (~ #218 #218)
-#222 := [refl]: #219
-#221 := [nnf-neg #222]: #220
-#16 := (<= #13 #15)
-#17 := (exists (vars (?v0 S2)) #16)
-#18 := (not #17)
-#76 := (iff #18 #75)
-#73 := (iff #17 #72)
-#70 := (iff #16 #69)
-#71 := [rewrite]: #70
-#74 := [quant-intro #71]: #73
-#77 := [monotonicity #74]: #76
-#58 := [asserted]: #18
-#78 := [mp #58 #77]: #75
-#216 := [mp~ #78 #221]: #223
-#281 := [mp #216 #280]: #276
-#310 := (not #276)
-#311 := (or #310 #307)
-#291 := (* -1::real #20)
-#292 := (+ #24 #291)
-#293 := (<= #292 0::real)
-#294 := (not #293)
-#312 := (or #310 #294)
-#314 := (iff #312 #311)
-#316 := (iff #311 #311)
-#317 := [rewrite]: #316
-#308 := (iff #294 #307)
-#305 := (iff #293 #302)
-#295 := (+ #291 #24)
-#298 := (<= #295 0::real)
-#303 := (iff #298 #302)
-#304 := [rewrite]: #303
-#299 := (iff #293 #298)
-#296 := (= #292 #295)
-#297 := [rewrite]: #296
-#300 := [monotonicity #297]: #299
-#306 := [trans #300 #304]: #305
-#309 := [monotonicity #306]: #308
-#315 := [monotonicity #309]: #314
-#318 := [trans #315 #317]: #314
-#313 := [quant-inst]: #312
-#319 := [mp #313 #318]: #311
-#333 := [unit-resolution #319 #281]: #307
-#204 := [not-or-elim #203]: #143
-#334 := [hypothesis]: #331
-#335 := [th-lemma #334 #204 #333 #326]: false
-#337 := [lemma #335]: #336
-#338 := [hypothesis]: #283
-#340 := (or #339 #331)
-#341 := [th-lemma]: #340
-#342 := [unit-resolution #341 #338 #337]: false
-#343 := [lemma #342]: #339
-#287 := (or #151 #283)
-#288 := [def-axiom]: #287
-#344 := [unit-resolution #288 #343]: #151
-#284 := (not #151)
-#285 := (or #284 #282)
-#286 := [def-axiom]: #285
-#345 := [unit-resolution #286 #344]: #282
-#346 := (not #282)
-#347 := (or #346 #322)
-#348 := [th-lemma]: #347
-[unit-resolution #348 #345 #330]: false
+#190 := (iff #166 #180)
+#179 := (not #180)
+#185 := (not #179)
+#188 := (iff #185 #180)
+#189 := [rewrite]: #188
+#186 := (iff #166 #185)
+#183 := (iff #31 #179)
+#184 := [rewrite]: #183
+#187 := [monotonicity #184]: #186
+#191 := [trans #187 #189]: #190
+#203 := [monotonicity #191 #200]: #202
+#206 := [quant-intro #203]: #205
+#176 := (iff #39 #175)
+#173 := (iff #38 #172)
+#174 := [rewrite]: #173
+#177 := [quant-intro #174]: #176
+#208 := [trans #177 #206]: #207
+#165 := [asserted]: #39
+#209 := [mp #165 #208]: #204
+#364 := [mp~ #209 #363]: #204
+#449 := [mp #364 #448]: #444
+#544 := (not #444)
+#545 := (or #544 #215 #538)
+#507 := (* -1::real #43)
+#508 := (+ #47 #507)
+#511 := (<= #508 0::real)
+#512 := (not #511)
+#513 := (+ #41 #181)
+#514 := (>= #513 0::int)
+#515 := (or #514 #512)
+#546 := (or #544 #515)
+#553 := (iff #546 #545)
+#541 := (or #215 #538)
+#548 := (or #544 #541)
+#551 := (iff #548 #545)
+#552 := [rewrite]: #551
+#549 := (iff #546 #548)
+#542 := (iff #515 #541)
+#539 := (iff #512 #538)
+#536 := (iff #511 #533)
+#526 := (+ #507 #47)
+#529 := (<= #526 0::real)
+#534 := (iff #529 #533)
+#535 := [rewrite]: #534
+#530 := (iff #511 #529)
+#527 := (= #508 #526)
+#528 := [rewrite]: #527
+#531 := [monotonicity #528]: #530
+#537 := [trans #531 #535]: #536
+#540 := [monotonicity #537]: #539
+#524 := (iff #514 #215)
+#516 := (+ #181 #41)
+#519 := (>= #516 0::int)
+#522 := (iff #519 #215)
+#523 := [rewrite]: #522
+#520 := (iff #514 #519)
+#517 := (= #513 #516)
+#518 := [rewrite]: #517
+#521 := [monotonicity #518]: #520
+#525 := [trans #521 #523]: #524
+#543 := [monotonicity #525 #540]: #542
+#550 := [monotonicity #543]: #549
+#554 := [trans #550 #552]: #553
+#547 := [quant-inst]: #546
+#555 := [mp #547 #554]: #545
+#607 := [unit-resolution #555 #449 #219]: #538
+#329 := [and-elim #328]: #270
+#608 := [hypothesis]: #605
+#609 := [th-lemma #608 #329 #607 #600]: false
+#611 := [lemma #609]: #610
+#612 := [hypothesis]: #451
+#614 := (or #613 #605)
+#615 := [th-lemma]: #614
+#616 := [unit-resolution #615 #612 #611]: false
+#617 := [lemma #616]: #613
+#455 := (or #275 #451)
+#456 := [def-axiom]: #455
+#618 := [unit-resolution #456 #617]: #275
+#452 := (not #275)
+#453 := (or #452 #450)
+#454 := [def-axiom]: #453
+#619 := [unit-resolution #454 #618]: #450
+#620 := (not #450)
+#621 := (or #620 #595)
+#622 := [th-lemma]: #621
+[unit-resolution #622 #619 #604]: false
 unsat
-3f6125a99a8cb462db3a2586a1eae0021b892091 288 0
+26bfbe5c6decad8b4a2114b85a11ff3552d6b060 295 0
 #2 := false
-#8 := 0::real
-decl f4 :: (-> S3 S2 real)
+#25 := 0::real
+decl f8 :: (-> S4 S2 real)
+decl f11 :: S2
+#40 := f11
+decl f12 :: S4
+#44 := f12
+#45 := (f8 f12 f11)
+decl f10 :: S4
+#34 := f10
+#43 := (f8 f10 f11)
+#199 := -1::real
+#263 := (* -1::real #43)
+#264 := (+ #263 #45)
+decl f5 :: real
+#26 := f5
+#242 := (* -1::real #45)
+#331 := (+ #43 #242)
+#332 := (+ f5 #331)
+#333 := (<= #332 0::real)
+#336 := (ite #333 f5 #264)
+#666 := (* -1::real #336)
+#667 := (+ f5 #666)
+#668 := (<= #667 0::real)
+#676 := (not #668)
+#230 := 1/2::real
+#414 := (* 1/2::real #336)
+#531 := (<= #414 0::real)
+#415 := (= #414 0::real)
+#284 := -1/2::real
+#339 := (* -1/2::real #336)
+#342 := (+ #45 #339)
+decl f9 :: S4
+#32 := f9
+#47 := (f8 f9 f11)
+#243 := (+ #242 #47)
+#316 := (* -1::real #47)
+#317 := (+ #45 #316)
+#318 := (+ f5 #317)
+#319 := (<= #318 0::real)
+#322 := (ite #319 f5 #243)
+#325 := (* 1/2::real #322)
+#328 := (+ #45 #325)
+#232 := (* 1/2::real #47)
+#312 := (+ #242 #232)
+#231 := (* 1/2::real #43)
+#313 := (+ #231 #312)
+#310 := (>= #313 0::real)
+#345 := (ite #310 #328 #342)
+#348 := (= #45 #345)
+#418 := (iff #348 #415)
+#411 := (= #45 #342)
+#416 := (iff #411 #415)
+#417 := [rewrite]: #416
+#412 := (iff #348 #411)
+#409 := (= #345 #342)
+#404 := (ite false #328 #342)
+#407 := (= #404 #342)
+#408 := [rewrite]: #407
+#405 := (= #345 #404)
+#402 := (iff #310 false)
+#309 := (not #310)
+#361 := (<= #317 0::real)
+#358 := (<= #331 0::real)
+#364 := (and #358 #361)
+#13 := 0::int
+decl f4 :: (-> S2 int)
+#41 := (f4 f11)
+#185 := -1::int
+#220 := (* -1::int #41)
+decl f6 :: (-> S3 S2)
+decl f7 :: S3
+#28 := f7
+#29 := (f6 f7)
+#30 := (f4 #29)
+#221 := (+ #30 #220)
+#222 := (<= #221 0::int)
+#367 := (or #222 #364)
+#376 := (and #309 #348 #367)
+#52 := 2::real
+#61 := (- #45 #43)
+#62 := (<= f5 #61)
+#63 := (ite #62 f5 #61)
+#64 := (/ #63 2::real)
+#65 := (- #45 #64)
+#56 := (- #47 #45)
+#57 := (<= f5 #56)
+#58 := (ite #57 f5 #56)
+#59 := (/ #58 2::real)
+#60 := (+ #45 #59)
+#51 := (+ #43 #47)
+#53 := (/ #51 2::real)
+#55 := (<= #45 #53)
+#66 := (ite #55 #60 #65)
+#67 := (= #66 #45)
+#54 := (< #53 #45)
+#68 := (and #54 #67)
+#48 := (<= #45 #47)
+#46 := (<= #43 #45)
+#49 := (and #46 #48)
+#42 := (< #41 #30)
+#50 := (implies #42 #49)
+#69 := (and #50 #68)
+#381 := (iff #69 #376)
+#267 := (<= f5 #264)
+#270 := (ite #267 f5 #264)
+#285 := (* -1/2::real #270)
+#286 := (+ #45 #285)
+#246 := (<= f5 #243)
+#249 := (ite #246 f5 #243)
+#255 := (* 1/2::real #249)
+#260 := (+ #45 #255)
+#233 := (+ #231 #232)
+#239 := (<= #45 #233)
+#291 := (ite #239 #260 #286)
+#297 := (= #45 #291)
+#236 := (< #233 #45)
+#302 := (and #236 #297)
+#219 := (not #42)
+#227 := (or #219 #49)
+#305 := (and #227 #302)
+#379 := (iff #305 #376)
+#370 := (and #309 #348)
+#373 := (and #367 #370)
+#377 := (iff #373 #376)
+#378 := [rewrite]: #377
+#374 := (iff #305 #373)
+#371 := (iff #302 #370)
+#349 := (iff #297 #348)
+#346 := (= #291 #345)
+#343 := (= #286 #342)
+#340 := (= #285 #339)
+#337 := (= #270 #336)
+#334 := (iff #267 #333)
+#335 := [rewrite]: #334
+#338 := [monotonicity #335]: #337
+#341 := [monotonicity #338]: #340
+#344 := [monotonicity #341]: #343
+#329 := (= #260 #328)
+#326 := (= #255 #325)
+#323 := (= #249 #322)
+#320 := (iff #246 #319)
+#321 := [rewrite]: #320
+#324 := [monotonicity #321]: #323
+#327 := [monotonicity #324]: #326
+#330 := [monotonicity #327]: #329
+#315 := (iff #239 #310)
+#314 := [rewrite]: #315
+#347 := [monotonicity #314 #330 #344]: #346
+#350 := [monotonicity #347]: #349
+#308 := (iff #236 #309)
+#311 := [rewrite]: #308
+#372 := [monotonicity #311 #350]: #371
+#368 := (iff #227 #367)
+#365 := (iff #49 #364)
+#362 := (iff #48 #361)
+#363 := [rewrite]: #362
+#359 := (iff #46 #358)
+#360 := [rewrite]: #359
+#366 := [monotonicity #360 #363]: #365
+#356 := (iff #219 #222)
+#223 := (not #222)
+#351 := (not #223)
+#354 := (iff #351 #222)
+#355 := [rewrite]: #354
+#352 := (iff #219 #351)
+#224 := (iff #42 #223)
+#225 := [rewrite]: #224
+#353 := [monotonicity #225]: #352
+#357 := [trans #353 #355]: #356
+#369 := [monotonicity #357 #366]: #368
+#375 := [monotonicity #369 #372]: #374
+#380 := [trans #375 #378]: #379
+#306 := (iff #69 #305)
+#303 := (iff #68 #302)
+#300 := (iff #67 #297)
+#294 := (= #291 #45)
+#298 := (iff #294 #297)
+#299 := [rewrite]: #298
+#295 := (iff #67 #294)
+#292 := (= #66 #291)
+#289 := (= #65 #286)
+#276 := (* 1/2::real #270)
+#281 := (- #45 #276)
+#287 := (= #281 #286)
+#288 := [rewrite]: #287
+#282 := (= #65 #281)
+#279 := (= #64 #276)
+#273 := (/ #270 2::real)
+#277 := (= #273 #276)
+#278 := [rewrite]: #277
+#274 := (= #64 #273)
+#271 := (= #63 #270)
+#265 := (= #61 #264)
+#266 := [rewrite]: #265
+#268 := (iff #62 #267)
+#269 := [monotonicity #266]: #268
+#272 := [monotonicity #269 #266]: #271
+#275 := [monotonicity #272]: #274
+#280 := [trans #275 #278]: #279
+#283 := [monotonicity #280]: #282
+#290 := [trans #283 #288]: #289
+#261 := (= #60 #260)
+#258 := (= #59 #255)
+#252 := (/ #249 2::real)
+#256 := (= #252 #255)
+#257 := [rewrite]: #256
+#253 := (= #59 #252)
+#250 := (= #58 #249)
+#244 := (= #56 #243)
+#245 := [rewrite]: #244
+#247 := (iff #57 #246)
+#248 := [monotonicity #245]: #247
+#251 := [monotonicity #248 #245]: #250
+#254 := [monotonicity #251]: #253
+#259 := [trans #254 #257]: #258
+#262 := [monotonicity #259]: #261
+#240 := (iff #55 #239)
+#234 := (= #53 #233)
+#235 := [rewrite]: #234
+#241 := [monotonicity #235]: #240
+#293 := [monotonicity #241 #262 #290]: #292
+#296 := [monotonicity #293]: #295
+#301 := [trans #296 #299]: #300
+#237 := (iff #54 #236)
+#238 := [monotonicity #235]: #237
+#304 := [monotonicity #238 #301]: #303
+#228 := (iff #50 #227)
+#229 := [rewrite]: #228
+#307 := [monotonicity #229 #304]: #306
+#382 := [trans #307 #380]: #381
+#218 := [asserted]: #69
+#383 := [mp #218 #382]: #376
+#384 := [and-elim #383]: #309
+#403 := [iff-false #384]: #402
+#406 := [monotonicity #403]: #405
+#410 := [trans #406 #408]: #409
+#413 := [monotonicity #410]: #412
+#419 := [trans #413 #417]: #418
+#385 := [and-elim #383]: #348
+#420 := [mp #385 #419]: #415
+#670 := (not #415)
+#671 := (or #670 #531)
+#672 := [th-lemma]: #671
+#673 := [unit-resolution #672 #420]: #531
+#674 := [hypothesis]: #668
+#174 := (<= f5 0::real)
+#175 := (not #174)
+#27 := (< 0::real f5)
+#176 := (iff #27 #175)
+#177 := [rewrite]: #176
+#171 := [asserted]: #27
+#178 := [mp #171 #177]: #175
+#675 := [th-lemma #178 #674 #673]: false
+#677 := [lemma #675]: #676
+#524 := (= f5 #336)
+#525 := (= #264 #336)
+#685 := (not #525)
+#669 := (+ #264 #666)
+#678 := (<= #669 0::real)
+#682 := (not #678)
+#428 := (iff #367 #364)
+#423 := (or false #364)
+#426 := (iff #423 #364)
+#427 := [rewrite]: #426
+#424 := (iff #367 #423)
+#400 := (iff #222 false)
+#217 := [asserted]: #42
+#226 := [mp #217 #225]: #223
+#401 := [iff-false #226]: #400
+#425 := [monotonicity #401]: #424
+#429 := [trans #425 #427]: #428
+#386 := [and-elim #383]: #367
+#430 := [mp #386 #429]: #364
+#422 := [and-elim #430]: #361
+#680 := [hypothesis]: #678
+#681 := [th-lemma #680 #422 #384 #673]: false
+#683 := [lemma #681]: #682
+#684 := [hypothesis]: #525
+#686 := (or #685 #678)
+#687 := [th-lemma]: #686
+#688 := [unit-resolution #687 #684 #683]: false
+#689 := [lemma #688]: #685
+#529 := (or #333 #525)
+#530 := [def-axiom]: #529
+#690 := [unit-resolution #530 #689]: #333
+#526 := (not #333)
+#527 := (or #526 #524)
+#528 := [def-axiom]: #527
+#691 := [unit-resolution #528 #690]: #524
+#692 := (not #524)
+#693 := (or #692 #668)
+#694 := [th-lemma]: #693
+[unit-resolution #694 #691 #677]: false
+unsat
+5d75e298a98eb55cd39dacd22665bd3062cf5345 204 0
+#2 := false
+#46 := 0::real
+decl f5 :: (-> S3 S2 real)
 decl f7 :: S2
-#19 := f7
+#26 := f7
+decl f11 :: S3
+#38 := f11
+#49 := (f5 f11 f7)
+#183 := -1::real
+#349 := (* -1::real #49)
 decl f8 :: S3
-#21 := f8
-#22 := (f4 f8 f7)
+#28 := f8
+#29 := (f5 f8 f7)
+#362 := (+ #29 #349)
+#363 := (>= #362 0::real)
+#368 := (not #363)
 decl f6 :: S3
-#14 := f6
-#20 := (f4 f6 f7)
-#73 := -1::real
-#118 := (* -1::real #20)
-#119 := (+ #118 #22)
-decl f3 :: real
-#9 := f3
-#97 := (* -1::real #22)
-#211 := (+ #20 #97)
-#212 := (+ f3 #211)
-#213 := (<= #212 0::real)
-#216 := (ite #213 f3 #119)
-#397 := (* -1::real #216)
-#398 := (+ f3 #397)
-#399 := (<= #398 0::real)
-#407 := (not #399)
-#72 := 1/2::real
-#287 := (* 1/2::real #216)
-#367 := (<= #287 0::real)
-#288 := (= #287 0::real)
-#139 := -1/2::real
-#219 := (* -1/2::real #216)
-#222 := (+ #22 #219)
-decl f5 :: S3
-#11 := f5
-#24 := (f4 f5 f7)
-#98 := (+ #97 #24)
-#196 := (* -1::real #24)
-#197 := (+ #22 #196)
-#198 := (+ f3 #197)
-#199 := (<= #198 0::real)
-#202 := (ite #199 f3 #98)
-#205 := (* 1/2::real #202)
-#208 := (+ #22 #205)
-#87 := (* 1/2::real #24)
-#185 := (+ #97 #87)
-#86 := (* 1/2::real #20)
-#186 := (+ #86 #185)
-#183 := (>= #186 0::real)
-#225 := (ite #183 #208 #222)
-#228 := (= #22 #225)
-#291 := (iff #228 #288)
-#284 := (= #22 #222)
-#289 := (iff #284 #288)
-#290 := [rewrite]: #289
-#285 := (iff #228 #284)
-#282 := (= #225 #222)
-#277 := (ite false #208 #222)
-#280 := (= #277 #222)
-#281 := [rewrite]: #280
-#278 := (= #225 #277)
-#275 := (iff #183 false)
-#182 := (not #183)
-#237 := (<= #197 0::real)
-#234 := (<= #211 0::real)
-#240 := (and #234 #237)
-#243 := (not #240)
-#231 := (not #228)
-#252 := (or #183 #231 #243)
-#257 := (not #252)
-#28 := 2::real
-#37 := (- #22 #20)
-#38 := (<= f3 #37)
-#39 := (ite #38 f3 #37)
-#40 := (/ #39 2::real)
-#41 := (- #22 #40)
-#32 := (- #24 #22)
-#33 := (<= f3 #32)
-#34 := (ite #33 f3 #32)
-#35 := (/ #34 2::real)
-#36 := (+ #22 #35)
-#27 := (+ #20 #24)
-#29 := (/ #27 2::real)
-#31 := (<= #22 #29)
-#42 := (ite #31 #36 #41)
-#43 := (= #42 #22)
-#44 := (not #43)
-#30 := (< #29 #22)
-#45 := (implies #30 #44)
-#25 := (<= #22 #24)
-#23 := (<= #20 #22)
-#26 := (and #23 #25)
-#46 := (implies #26 #45)
-#47 := (not #46)
-#260 := (iff #47 #257)
-#122 := (<= f3 #119)
-#125 := (ite #122 f3 #119)
-#140 := (* -1/2::real #125)
-#141 := (+ #22 #140)
-#101 := (<= f3 #98)
-#104 := (ite #101 f3 #98)
-#110 := (* 1/2::real #104)
-#115 := (+ #22 #110)
-#88 := (+ #86 #87)
-#94 := (<= #22 #88)
-#146 := (ite #94 #115 #141)
-#152 := (= #22 #146)
-#157 := (not #152)
-#91 := (< #88 #22)
-#163 := (not #91)
-#164 := (or #163 #157)
-#172 := (not #26)
-#173 := (or #172 #164)
-#178 := (not #173)
-#258 := (iff #178 #257)
-#255 := (iff #173 #252)
-#246 := (or #183 #231)
-#249 := (or #243 #246)
-#253 := (iff #249 #252)
-#254 := [rewrite]: #253
-#250 := (iff #173 #249)
-#247 := (iff #164 #246)
-#232 := (iff #157 #231)
-#229 := (iff #152 #228)
-#226 := (= #146 #225)
-#223 := (= #141 #222)
-#220 := (= #140 #219)
-#217 := (= #125 #216)
-#214 := (iff #122 #213)
+#25 := f6
+#27 := (f5 f6 f7)
+#350 := (+ #27 #349)
+#351 := (<= #350 0::real)
+#352 := (not #351)
+#371 := (or #352 #368)
+#374 := (not #371)
+#8 := (:var 0 S2)
+#41 := (f5 f8 #8)
+#333 := (pattern #41)
+#39 := (f5 f11 #8)
+#332 := (pattern #39)
+#37 := (f5 f6 #8)
+#331 := (pattern #37)
+decl f4 :: (-> S2 int)
+#9 := (f4 #8)
+#311 := (pattern #9)
+#189 := (* -1::real #41)
+#190 := (+ #39 #189)
+#191 := (<= #190 0::real)
+#242 := (not #191)
+#184 := (* -1::real #39)
+#185 := (+ #37 #184)
+#186 := (<= #185 0::real)
+#241 := (not #186)
+#243 := (or #241 #242)
+#244 := (not #243)
+#13 := 0::int
+decl f9 :: (-> S4 S2)
+decl f10 :: S4
+#32 := f10
+#33 := (f9 f10)
+#34 := (f4 #33)
+#157 := -1::int
+#160 := (* -1::int #34)
+#173 := (+ #9 #160)
+#172 := (>= #173 0::int)
+#247 := (or #172 #244)
+#334 := (forall (vars (?v0 S2)) (:pat #311 #331 #332 #333) #247)
+#250 := (forall (vars (?v0 S2)) #247)
+#337 := (iff #250 #334)
+#335 := (iff #247 #247)
+#336 := [refl]: #335
+#338 := [quant-intro #336]: #337
+#194 := (and #186 #191)
+#197 := (or #172 #194)
+#200 := (forall (vars (?v0 S2)) #197)
+#251 := (iff #200 #250)
+#248 := (iff #197 #247)
+#245 := (iff #194 #244)
+#246 := [rewrite]: #245
+#249 := [monotonicity #246]: #248
+#252 := [quant-intro #249]: #251
+#219 := (~ #200 #200)
+#224 := (~ #197 #197)
+#222 := [refl]: #224
+#239 := [nnf-pos #222]: #219
+#42 := (<= #39 #41)
+#40 := (<= #37 #39)
+#43 := (and #40 #42)
+#36 := (< #9 #34)
+#44 := (implies #36 #43)
+#45 := (forall (vars (?v0 S2)) #44)
+#203 := (iff #45 #200)
+#156 := (not #36)
+#165 := (or #156 #43)
+#168 := (forall (vars (?v0 S2)) #165)
+#201 := (iff #168 #200)
+#198 := (iff #165 #197)
+#195 := (iff #43 #194)
+#192 := (iff #42 #191)
+#193 := [rewrite]: #192
+#187 := (iff #40 #186)
+#188 := [rewrite]: #187
+#196 := [monotonicity #188 #193]: #195
+#181 := (iff #156 #172)
+#171 := (not #172)
+#176 := (not #171)
+#179 := (iff #176 #172)
+#180 := [rewrite]: #179
+#177 := (iff #156 #176)
+#174 := (iff #36 #171)
+#175 := [rewrite]: #174
+#178 := [monotonicity #175]: #177
+#182 := [trans #178 #180]: #181
+#199 := [monotonicity #182 #196]: #198
+#202 := [quant-intro #199]: #201
+#169 := (iff #45 #168)
+#166 := (iff #44 #165)
+#167 := [rewrite]: #166
+#170 := [quant-intro #167]: #169
+#204 := [trans #170 #202]: #203
+#155 := [asserted]: #45
+#205 := [mp #155 #204]: #200
+#240 := [mp~ #205 #239]: #200
+#253 := [mp #240 #252]: #250
+#339 := [mp #253 #338]: #334
+#31 := (f4 f7)
+#161 := (+ #31 #160)
+#159 := (>= #161 0::int)
+#158 := (not #159)
+#35 := (< #31 #34)
+#162 := (iff #35 #158)
+#163 := [rewrite]: #162
+#154 := [asserted]: #35
+#164 := [mp #154 #163]: #158
+#380 := (not #334)
+#381 := (or #380 #159 #374)
+#342 := (* -1::real #29)
+#343 := (+ #49 #342)
+#347 := (<= #343 0::real)
+#348 := (not #347)
+#353 := (or #352 #348)
+#354 := (not #353)
+#355 := (or #159 #354)
+#382 := (or #380 #355)
+#389 := (iff #382 #381)
+#377 := (or #159 #374)
+#384 := (or #380 #377)
+#387 := (iff #384 #381)
+#388 := [rewrite]: #387
+#385 := (iff #382 #384)
+#378 := (iff #355 #377)
+#375 := (iff #354 #374)
+#372 := (iff #353 #371)
+#369 := (iff #348 #368)
+#366 := (iff #347 #363)
+#356 := (+ #342 #49)
+#359 := (<= #356 0::real)
+#364 := (iff #359 #363)
+#365 := [rewrite]: #364
+#360 := (iff #347 #359)
+#357 := (= #343 #356)
+#358 := [rewrite]: #357
+#361 := [monotonicity #358]: #360
+#367 := [trans #361 #365]: #366
+#370 := [monotonicity #367]: #369
+#373 := [monotonicity #370]: #372
+#376 := [monotonicity #373]: #375
+#379 := [monotonicity #376]: #378
+#386 := [monotonicity #379]: #385
+#390 := [trans #386 #388]: #389
+#383 := [quant-inst]: #382
+#391 := [mp #383 #390]: #381
+#513 := [unit-resolution #391 #164 #339]: #374
+#394 := (or #371 #363)
+#395 := [def-axiom]: #394
+#514 := [unit-resolution #395 #513]: #363
+#475 := (>= #350 0::real)
+#524 := (not #475)
+#474 := (= #27 #49)
+#519 := (not #474)
+#208 := (= #29 #49)
+#216 := (not #208)
+#520 := (iff #216 #519)
+#517 := (iff #208 #474)
+#515 := (iff #474 #208)
+#30 := (= #27 #29)
+#153 := [asserted]: #30
+#516 := [monotonicity #153]: #515
+#518 := [symm #516]: #517
+#521 := [monotonicity #518]: #520
+#50 := (= #49 #29)
+#51 := (not #50)
+#217 := (iff #51 #216)
+#214 := (iff #50 #208)
 #215 := [rewrite]: #214
 #218 := [monotonicity #215]: #217
-#221 := [monotonicity #218]: #220
-#224 := [monotonicity #221]: #223
-#209 := (= #115 #208)
-#206 := (= #110 #205)
-#203 := (= #104 #202)
-#200 := (iff #101 #199)
-#201 := [rewrite]: #200
-#204 := [monotonicity #201]: #203
-#207 := [monotonicity #204]: #206
-#210 := [monotonicity #207]: #209
-#195 := (iff #94 #183)
-#194 := [rewrite]: #195
-#227 := [monotonicity #194 #210 #224]: #226
-#230 := [monotonicity #227]: #229
-#233 := [monotonicity #230]: #232
-#192 := (iff #163 #183)
-#187 := (not #182)
-#190 := (iff #187 #183)
-#191 := [rewrite]: #190
-#188 := (iff #163 #187)
-#181 := (iff #91 #182)
-#184 := [rewrite]: #181
-#189 := [monotonicity #184]: #188
-#193 := [trans #189 #191]: #192
-#248 := [monotonicity #193 #233]: #247
-#244 := (iff #172 #243)
-#241 := (iff #26 #240)
-#238 := (iff #25 #237)
-#239 := [rewrite]: #238
-#235 := (iff #23 #234)
-#236 := [rewrite]: #235
-#242 := [monotonicity #236 #239]: #241
-#245 := [monotonicity #242]: #244
-#251 := [monotonicity #245 #248]: #250
-#256 := [trans #251 #254]: #255
-#259 := [monotonicity #256]: #258
-#179 := (iff #47 #178)
-#176 := (iff #46 #173)
-#169 := (implies #26 #164)
-#174 := (iff #169 #173)
-#175 := [rewrite]: #174
-#170 := (iff #46 #169)
-#167 := (iff #45 #164)
-#160 := (implies #91 #157)
-#165 := (iff #160 #164)
-#166 := [rewrite]: #165
-#161 := (iff #45 #160)
-#158 := (iff #44 #157)
-#155 := (iff #43 #152)
-#149 := (= #146 #22)
-#153 := (iff #149 #152)
-#154 := [rewrite]: #153
-#150 := (iff #43 #149)
-#147 := (= #42 #146)
-#144 := (= #41 #141)
-#131 := (* 1/2::real #125)
-#136 := (- #22 #131)
-#142 := (= #136 #141)
-#143 := [rewrite]: #142
-#137 := (= #41 #136)
-#134 := (= #40 #131)
-#128 := (/ #125 2::real)
-#132 := (= #128 #131)
-#133 := [rewrite]: #132
-#129 := (= #40 #128)
-#126 := (= #39 #125)
-#120 := (= #37 #119)
-#121 := [rewrite]: #120
-#123 := (iff #38 #122)
-#124 := [monotonicity #121]: #123
-#127 := [monotonicity #124 #121]: #126
-#130 := [monotonicity #127]: #129
-#135 := [trans #130 #133]: #134
-#138 := [monotonicity #135]: #137
-#145 := [trans #138 #143]: #144
-#116 := (= #36 #115)
-#113 := (= #35 #110)
-#107 := (/ #104 2::real)
-#111 := (= #107 #110)
-#112 := [rewrite]: #111
-#108 := (= #35 #107)
-#105 := (= #34 #104)
-#99 := (= #32 #98)
-#100 := [rewrite]: #99
-#102 := (iff #33 #101)
-#103 := [monotonicity #100]: #102
-#106 := [monotonicity #103 #100]: #105
-#109 := [monotonicity #106]: #108
-#114 := [trans #109 #112]: #113
-#117 := [monotonicity #114]: #116
-#95 := (iff #31 #94)
-#89 := (= #29 #88)
-#90 := [rewrite]: #89
-#96 := [monotonicity #90]: #95
-#148 := [monotonicity #96 #117 #145]: #147
-#151 := [monotonicity #148]: #150
-#156 := [trans #151 #154]: #155
-#159 := [monotonicity #156]: #158
-#92 := (iff #30 #91)
-#93 := [monotonicity #90]: #92
-#162 := [monotonicity #93 #159]: #161
-#168 := [trans #162 #166]: #167
-#171 := [monotonicity #168]: #170
-#177 := [trans #171 #175]: #176
-#180 := [monotonicity #177]: #179
-#261 := [trans #180 #259]: #260
-#66 := [asserted]: #47
-#262 := [mp #66 #261]: #257
-#263 := [not-or-elim #262]: #182
-#276 := [iff-false #263]: #275
-#279 := [monotonicity #276]: #278
-#283 := [trans #279 #281]: #282
-#286 := [monotonicity #283]: #285
-#292 := [trans #286 #290]: #291
-#264 := [not-or-elim #262]: #228
-#293 := [mp #264 #292]: #288
-#401 := (not #288)
-#402 := (or #401 #367)
-#403 := [th-lemma]: #402
-#404 := [unit-resolution #403 #293]: #367
-#405 := [hypothesis]: #399
-#67 := (<= f3 0::real)
-#68 := (not #67)
-#10 := (< 0::real f3)
-#69 := (iff #10 #68)
-#70 := [rewrite]: #69
-#64 := [asserted]: #10
-#71 := [mp #64 #70]: #68
-#406 := [th-lemma #71 #405 #404]: false
-#408 := [lemma #406]: #407
-#360 := (= f3 #216)
-#361 := (= #119 #216)
-#416 := (not #361)
-#400 := (+ #119 #397)
-#409 := (<= #400 0::real)
-#413 := (not #409)
-#265 := [not-or-elim #262]: #240
-#267 := [and-elim #265]: #237
-#411 := [hypothesis]: #409
-#412 := [th-lemma #411 #267 #263 #404]: false
-#414 := [lemma #412]: #413
-#415 := [hypothesis]: #361
-#417 := (or #416 #409)
-#418 := [th-lemma]: #417
-#419 := [unit-resolution #418 #415 #414]: false
-#420 := [lemma #419]: #416
-#365 := (or #213 #361)
-#366 := [def-axiom]: #365
-#421 := [unit-resolution #366 #420]: #213
-#362 := (not #213)
-#363 := (or #362 #360)
-#364 := [def-axiom]: #363
-#422 := [unit-resolution #364 #421]: #360
-#423 := (not #360)
-#424 := (or #423 #399)
-#425 := [th-lemma]: #424
-[unit-resolution #425 #422 #408]: false
+#207 := [asserted]: #51
+#221 := [mp #207 #218]: #216
+#522 := [mp #221 #521]: #519
+#527 := (or #474 #524)
+#392 := (or #371 #351)
+#393 := [def-axiom]: #392
+#523 := [unit-resolution #393 #513]: #351
+#525 := (or #474 #352 #524)
+#526 := [th-lemma]: #525
+#528 := [unit-resolution #526 #523]: #527
+#529 := [unit-resolution #528 #522]: #524
+#471 := (+ #27 #342)
+#473 := (>= #471 0::real)
+#530 := (not #30)
+#531 := (or #530 #473)
+#532 := [th-lemma]: #531
+#533 := [unit-resolution #532 #153]: #473
+[th-lemma #533 #529 #514]: false
 unsat
-2dea73fd0603d00ddaec5e14116c465addb0b89e 870 0
+7b235442f6856a575f67cb04d6eaa03042f5336e 870 0
 #2 := false
 #11 := 0::real
 decl f5 :: real
@@ -4202,5165 +2320,3 @@
 #949 := [unit-resolution #948 #945]: #819
 [th-lemma #926 #949 #646 #648 #899 #651 #938 #944 #644 #642]: false
 unsat
-fda4738b9d427b4c846961908e5b41bb384b40b6 89 0
-f1 -> val!2
-f2 -> val!0
-f23 -> val!1
-f24 -> val!12
-f7 -> val!20
-f25 -> val!18
-f21 -> val!3
-f10 -> val!5
-f9 -> val!6
-f14 -> val!7
-f16 -> val!8
-f4 -> val!10
-f13 -> val!13
-f15 -> val!14
-f31 -> val!17
-f22 -> {
-  val!1 val!12 -> val!2
-  else -> val!2
-}
-f6 -> {
-  val!18 -> val!21
-  else -> val!21
-}
-f5 -> {
-  val!20 -> 7720
-  val!21 -> 7719
-  val!17 -> 8365
-  val!18 -> 1796
-  val!19 -> 1797
-  val!15 -> 1
-  val!22 -> 8366
-  else -> 8366
-}
-f20 -> {
-  val!3 -> val!19
-  else -> val!19
-}
-f26 -> {
-  val!18 val!19 -> val!4
-  else -> val!4
-}
-f8 -> {
-  val!5 val!18 -> 0
-  val!6 val!18 -> 0
-  val!7 val!18 -> 0
-  val!1 val!18 -> -1
-  else -> -1
-}
-f28 -> {
-  val!8 -> val!9
-  val!10 -> val!11
-  else -> val!11
-}
-f27 -> {
-  val!9 val!11 -> val!12
-  else -> val!12
-}
-f17 -> {
-  1 -> val!15
-  8366 -> val!22
-  8365 -> val!17
-  1796 -> val!18
-  1797 -> val!19
-  7720 -> val!20
-  7719 -> val!21
-  else -> val!21
-}
-f30 -> {
-  val!15 val!22 -> val!16
-  else -> val!16
-}
-f29 -> {
-  val!20 val!16 -> val!2
-  else -> val!2
-}
-f3 -> {
-  val!13 val!18 -> 0
-  val!8 val!18 -> 0
-  val!14 val!18 -> 0
-  val!10 val!18 -> 0
-  else -> 0
-}
-f18 -> (ite (forall (?v2 S2)
-              (or (not (= f1 (f19 ?v2 (f20 f21))))
-                  (<= (+ (f8 #1 ?v2) (* -1 (f8 #2 ?v2))) 0))
-              :qid {k!45})
-            f1
-            (f18!0 #0 #1))
-unknown
-7f975502da925ceb9dd4add3271a5a407e743846 262 0
-f1 -> val!7
-f2 -> val!0
-f16 -> val!2
-f7 -> val!4
-f27 -> val!25
-f28 -> val!12
-f18 -> val!8
-f4 -> val!10
-f13 -> val!13
-f17 -> val!14
-f12 -> val!16
-f9 -> val!17
-f14 -> val!18
-f10 -> val!19
-f15 -> val!20
-f19 -> val!21
-f24 -> val!23
-f11 -> {
-  val!2 -> val!1
-  val!4 -> val!3
-  else -> val!3
-}
-f6 -> {
-  val!1 -> val!2
-  val!3 -> val!4
-  val!26 -> val!29
-  val!4 -> val!28
-  val!25 -> val!29
-  val!2 -> val!27
-  val!5 -> val!29
-  val!27 -> val!30
-  val!28 -> val!4
-  val!29 -> val!31
-  val!30 -> val!4
-  val!31 -> val!2
-  else -> val!2
-}
-f20 -> {
-  1 -> val!5
-  7720 -> val!26
-  7719 -> val!25
-  14272 -> val!2
-  14270 -> val!4
-  8365 -> val!27
-  14269 -> val!28
-  14271 -> val!29
-  8457 -> val!30
-  0 -> val!31
-  else -> val!31
-}
-f26 -> {
-  val!5 val!25 -> val!6
-  val!5 val!26 -> val!15
-  else -> val!15
-}
-f25 -> {
-  val!2 val!6 -> val!7
-  val!4 val!6 -> val!7
-  val!4 val!15 -> val!7
-  else -> val!7
-}
-f23 -> {
-  val!8 -> val!9
-  val!10 -> val!11
-  val!21 -> val!22
-  val!20 -> val!24
-  else -> val!24
-}
-f29 -> {
-  val!9 val!11 -> val!12
-  else -> val!12
-}
-f5 -> {
-  val!25 -> 7719
-  val!2 -> 14272
-  val!4 -> 14270
-  val!5 -> 1
-  val!26 -> 7720
-  val!29 -> 14271
-  val!28 -> 14269
-  val!27 -> 8365
-  val!30 -> 8457
-  val!31 -> 0
-  else -> 0
-}
-f3 -> {
-  val!21 val!1 -> 2
-  val!8 val!1 -> 2
-  val!14 val!1 -> 3
-  val!20 val!1 -> 4
-  val!10 val!1 -> 3
-  val!13 val!1 -> 2
-  val!21 val!3 -> 7
-  val!8 val!3 -> 5
-  val!14 val!3 -> 6
-  val!20 val!3 -> 8
-  val!10 val!3 -> 7
-  val!13 val!3 -> 8
-  val!13 val!4 -> 29
-  val!10 val!2 -> 30
-  val!10 val!5 -> 36
-  val!10 val!4 -> 31
-  val!10 val!25 -> 37
-  val!10 val!26 -> 38
-  val!13 val!2 -> 32
-  val!13 val!5 -> 39
-  val!13 val!25 -> 40
-  val!13 val!26 -> 41
-  val!20 val!2 -> 30
-  val!20 val!5 -> 33
-  val!20 val!4 -> 31
-  val!20 val!25 -> 34
-  val!20 val!26 -> 35
-  val!14 val!2 -> 30
-  val!14 val!5 -> 36
-  val!14 val!4 -> 31
-  val!14 val!25 -> 37
-  val!14 val!26 -> 38
-  val!8 val!2 -> 32
-  val!8 val!5 -> 39
-  val!8 val!4 -> 29
-  val!8 val!25 -> 40
-  val!8 val!26 -> 41
-  val!21 val!2 -> 32
-  val!21 val!5 -> 42
-  val!21 val!4 -> 29
-  val!21 val!25 -> 43
-  val!21 val!26 -> 44
-  val!10 val!29 -> 73
-  val!10 val!28 -> 7
-  val!10 val!27 -> 75
-  val!13 val!29 -> 78
-  val!13 val!28 -> 8
-  val!13 val!27 -> 80
-  val!20 val!29 -> 73
-  val!20 val!28 -> 74
-  val!20 val!27 -> 75
-  val!14 val!29 -> 73
-  val!14 val!28 -> 76
-  val!14 val!27 -> 75
-  val!8 val!29 -> 78
-  val!8 val!28 -> 77
-  val!8 val!27 -> 80
-  val!21 val!29 -> 78
-  val!21 val!28 -> 79
-  val!21 val!27 -> 80
-  val!10 val!30 -> 7
-  val!13 val!30 -> 8
-  val!20 val!30 -> 93
-  val!14 val!30 -> 94
-  val!8 val!30 -> 95
-  val!21 val!30 -> 96
-  val!10 val!31 -> 101
-  val!13 val!31 -> 102
-  val!20 val!31 -> 4
-  val!14 val!31 -> 101
-  val!8 val!31 -> 102
-  val!21 val!31 -> 102
-  else -> 102
-}
-f8 -> {
-  val!18 val!1 -> 2
-  val!19 val!1 -> 4
-  val!16 val!1 -> 3
-  val!19 val!3 -> 7
-  val!17 val!3 -> 8
-  val!18 val!3 -> 5
-  val!16 val!3 -> 6
-  val!23 val!25 -> 0
-  val!9 val!25 -> -1
-  val!11 val!25 -> 1
-  val!24 val!25 -> 1
-  val!22 val!25 -> -1
-  val!23 val!2 -> 0
-  val!9 val!2 -> -1
-  val!11 val!2 -> 1
-  val!24 val!2 -> 1
-  val!22 val!2 -> -1
-  val!23 val!4 -> 0
-  val!9 val!4 -> -1
-  val!11 val!4 -> 1
-  val!24 val!4 -> 1
-  val!22 val!4 -> -1
-  val!23 val!5 -> 0
-  val!9 val!5 -> -1
-  val!11 val!5 -> 1
-  val!24 val!5 -> 1
-  val!22 val!5 -> -1
-  val!23 val!26 -> 0
-  val!9 val!26 -> -1
-  val!11 val!26 -> 1
-  val!24 val!26 -> 1
-  val!22 val!26 -> -1
-  val!17 val!2 -> 30
-  val!16 val!5 -> 36
-  val!17 val!5 -> 33
-  val!17 val!4 -> 31
-  val!16 val!25 -> 37
-  val!17 val!25 -> 34
-  val!16 val!26 -> 38
-  val!17 val!26 -> 35
-  val!19 val!2 -> 32
-  val!18 val!5 -> 39
-  val!19 val!5 -> 42
-  val!19 val!4 -> 29
-  val!18 val!25 -> 40
-  val!19 val!25 -> 43
-  val!18 val!26 -> 41
-  val!19 val!26 -> 44
-  val!23 val!3 -> 0
-  val!22 val!3 -> -1
-  val!24 val!3 -> 1
-  val!23 val!1 -> 0
-  val!22 val!1 -> -1
-  val!24 val!1 -> 1
-  val!11 val!3 -> 1
-  val!9 val!3 -> -1
-  val!11 val!1 -> 1
-  val!9 val!1 -> -1
-  val!23 val!27 -> 0
-  val!9 val!27 -> -1
-  val!11 val!27 -> 1
-  val!24 val!27 -> 1
-  val!22 val!27 -> -1
-  val!23 val!28 -> 0
-  val!9 val!28 -> -1
-  val!11 val!28 -> 1
-  val!24 val!28 -> 1
-  val!22 val!28 -> -1
-  val!23 val!29 -> 0
-  val!9 val!29 -> -1
-  val!11 val!29 -> 1
-  val!24 val!29 -> 1
-  val!22 val!29 -> -1
-  val!17 val!29 -> 73
-  val!16 val!28 -> 76
-  val!17 val!28 -> 74
-  val!17 val!27 -> 75
-  val!19 val!29 -> 78
-  val!18 val!28 -> 77
-  val!19 val!28 -> 79
-  val!19 val!27 -> 80
-  val!23 val!31 -> 0
-  val!9 val!31 -> -1
-  val!11 val!31 -> 1
-  val!24 val!31 -> 1
-  val!22 val!31 -> -1
-  val!23 val!30 -> 0
-  val!9 val!30 -> -1
-  val!11 val!30 -> 1
-  val!24 val!30 -> 1
-  val!22 val!30 -> -1
-  val!16 val!30 -> 94
-  val!17 val!30 -> 93
-  val!18 val!30 -> 95
-  val!19 val!30 -> 96
-  val!16 val!31 -> 101
-  val!18 val!31 -> 102
-  else -> 102
-}
-f32 -> (f30 (f31 #0) #1)
-unknown
-bd6da22de14f35502495633a6d03f6d719a5ebda 538 0
-f1 -> val!7
-f2 -> val!0
-f16 -> val!2
-f7 -> val!4
-f27 -> val!27
-f28 -> val!12
-f18 -> val!8
-f4 -> val!10
-f13 -> val!13
-f17 -> val!14
-f12 -> val!16
-f9 -> val!17
-f14 -> val!18
-f10 -> val!19
-f15 -> val!20
-f19 -> val!21
-f24 -> val!23
-f22 -> val!25
-f11 -> {
-  val!2 -> val!1
-  val!4 -> val!3
-  else -> val!3
-}
-f6 -> {
-  val!1 -> val!2
-  val!3 -> val!4
-  val!30 -> val!44
-  val!2 -> val!31
-  val!27 -> val!32
-  val!5 -> val!33
-  val!4 -> val!34
-  val!31 -> val!37
-  val!32 -> val!35
-  val!34 -> val!36
-  val!33 -> val!38
-  val!36 -> val!44
-  val!37 -> val!41
-  val!38 -> val!39
-  val!35 -> val!40
-  val!39 -> val!46
-  val!40 -> val!43
-  val!41 -> val!42
-  val!42 -> val!50
-  val!43 -> val!44
-  val!44 -> val!45
-  val!45 -> val!48
-  val!46 -> val!47
-  val!47 -> val!49
-  val!48 -> val!48
-  val!49 -> val!44
-  val!50 -> val!44
-  else -> val!44
-}
-f20 -> {
-  1 -> val!5
-  7720 -> val!30
-  7719 -> val!27
-  13211 -> val!2
-  13214 -> val!4
-  2944 -> val!31
-  8366 -> val!32
-  5724 -> val!33
-  1597 -> val!34
-  1176 -> val!35
-  1143 -> val!36
-  8753 -> val!37
-  10625 -> val!38
-  4877 -> val!39
-  2854 -> val!40
-  13213 -> val!41
-  5166 -> val!42
-  13216 -> val!43
-  6737 -> val!46
-  2853 -> val!45
-  1276 -> val!47
-  13210 -> val!48
-  2140 -> val!49
-  13212 -> val!44
-  10932 -> val!50
-  else -> val!50
-}
-f26 -> {
-  val!5 val!27 -> val!6
-  val!5 val!30 -> val!15
-  else -> val!15
-}
-f25 -> {
-  val!2 val!6 -> val!7
-  val!4 val!6 -> val!7
-  val!4 val!15 -> val!7
-  else -> val!7
-}
-f23 -> {
-  val!8 -> val!9
-  val!10 -> val!11
-  val!21 -> val!22
-  val!20 -> val!24
-  else -> val!24
-}
-f29 -> {
-  val!9 val!11 -> val!12
-  else -> val!12
-}
-f5 -> {
-  val!27 -> 7719
-  val!2 -> 13211
-  val!4 -> 13214
-  val!5 -> 1
-  val!30 -> 7720
-  val!31 -> 2944
-  val!32 -> 8366
-  val!33 -> 5724
-  val!34 -> 1597
-  val!37 -> 8753
-  val!35 -> 1176
-  val!36 -> 1143
-  val!38 -> 10625
-  val!41 -> 13213
-  val!39 -> 4877
-  val!40 -> 2854
-  val!46 -> 6737
-  val!43 -> 13216
-  val!48 -> 13210
-  val!42 -> 5166
-  val!50 -> 10932
-  val!44 -> 13212
-  val!45 -> 2853
-  val!47 -> 1276
-  val!49 -> 2140
-  else -> 2140
-}
-f21 -> {
-  val!25 -> val!26
-  else -> val!26
-}
-f3 -> {
-  val!21 val!1 -> 2
-  val!8 val!1 -> 4
-  val!14 val!1 -> 3
-  val!20 val!1 -> 4
-  val!10 val!1 -> 3
-  val!13 val!1 -> 4
-  val!21 val!3 -> 5
-  val!8 val!3 -> 5
-  val!14 val!3 -> 6
-  val!20 val!3 -> 6
-  val!10 val!3 -> 7
-  val!13 val!3 -> 8
-  val!10 val!30 -> 37
-  val!13 val!36 -> 30
-  val!10 val!36 -> 96
-  val!10 val!40 -> 32
-  val!21 val!27 -> 40
-  val!20 val!27 -> 35
-  val!21 val!2 -> 38
-  val!20 val!2 -> 33
-  val!10 val!27 -> 35
-  val!8 val!27 -> 40
-  val!10 val!4 -> 36
-  val!8 val!4 -> 41
-  val!10 val!2 -> 33
-  val!8 val!2 -> 38
-  val!21 val!4 -> 41
-  val!20 val!4 -> 36
-  val!10 val!5 -> 34
-  val!8 val!5 -> 39
-  val!8 val!30 -> 42
-  val!21 val!30 -> 43
-  val!20 val!30 -> 29
-  val!21 val!5 -> 39
-  val!20 val!5 -> 34
-  val!13 val!2 -> 38
-  val!13 val!5 -> 39
-  val!13 val!27 -> 40
-  val!13 val!4 -> 41
-  val!13 val!30 -> 42
-  val!14 val!2 -> 33
-  val!14 val!5 -> 34
-  val!14 val!27 -> 35
-  val!14 val!4 -> 36
-  val!14 val!30 -> 37
-  val!10 val!34 -> 70
-  val!8 val!34 -> 74
-  val!10 val!33 -> 68
-  val!8 val!33 -> 72
-  val!21 val!32 -> 73
-  val!20 val!32 -> 69
-  val!21 val!34 -> 74
-  val!20 val!34 -> 70
-  val!21 val!33 -> 72
-  val!20 val!33 -> 68
-  val!10 val!31 -> 67
-  val!8 val!31 -> 71
-  val!21 val!31 -> 71
-  val!20 val!31 -> 67
-  val!10 val!32 -> 69
-  val!8 val!32 -> 73
-  val!13 val!31 -> 71
-  val!13 val!33 -> 72
-  val!13 val!32 -> 73
-  val!13 val!34 -> 74
-  val!14 val!31 -> 67
-  val!14 val!33 -> 68
-  val!14 val!32 -> 69
-  val!14 val!34 -> 70
-  val!10 val!37 -> 98
-  val!8 val!37 -> 102
-  val!21 val!37 -> 103
-  val!20 val!37 -> 97
-  val!21 val!35 -> 104
-  val!20 val!35 -> 99
-  val!10 val!35 -> 99
-  val!8 val!35 -> 104
-  val!21 val!38 -> 100
-  val!20 val!38 -> 95
-  val!21 val!36 -> 101
-  val!20 val!36 -> 31
-  val!8 val!36 -> 30
-  val!10 val!38 -> 95
-  val!8 val!38 -> 100
-  val!13 val!38 -> 100
-  val!13 val!37 -> 102
-  val!13 val!35 -> 104
-  val!14 val!38 -> 95
-  val!14 val!36 -> 96
-  val!14 val!37 -> 98
-  val!14 val!35 -> 99
-  val!10 val!39 -> 125
-  val!8 val!39 -> 128
-  val!10 val!41 -> 127
-  val!8 val!41 -> 130
-  val!21 val!39 -> 128
-  val!20 val!39 -> 125
-  val!21 val!41 -> 130
-  val!20 val!41 -> 127
-  val!21 val!40 -> 131
-  val!20 val!40 -> 32
-  val!8 val!40 -> 131
-  val!13 val!39 -> 128
-  val!13 val!41 -> 130
-  val!13 val!40 -> 131
-  val!14 val!39 -> 125
-  val!14 val!41 -> 127
-  val!14 val!40 -> 32
-  val!10 val!48 -> 153
-  val!8 val!48 -> 159
-  val!8 val!42 -> 158
-  val!10 val!42 -> 152
-  val!21 val!43 -> 161
-  val!20 val!43 -> 154
-  val!21 val!48 -> 159
-  val!20 val!48 -> 153
-  val!10 val!43 -> 155
-  val!8 val!43 -> 160
-  val!10 val!46 -> 157
-  val!8 val!46 -> 163
-  val!21 val!46 -> 163
-  val!20 val!46 -> 157
-  val!20 val!42 -> 152
-  val!21 val!42 -> 158
-  val!13 val!42 -> 158
-  val!13 val!48 -> 159
-  val!13 val!43 -> 160
-  val!13 val!46 -> 163
-  val!14 val!42 -> 152
-  val!14 val!48 -> 153
-  val!14 val!43 -> 155
-  val!14 val!46 -> 157
-  val!21 val!44 -> 129
-  val!20 val!44 -> 126
-  val!10 val!44 -> 126
-  val!8 val!44 -> 129
-  val!20 val!47 -> 186
-  val!21 val!47 -> 189
-  val!10 val!47 -> 186
-  val!8 val!47 -> 189
-  val!21 val!50 -> 188
-  val!20 val!50 -> 184
-  val!10 val!50 -> 185
-  val!8 val!50 -> 187
-  val!13 val!44 -> 129
-  val!13 val!50 -> 187
-  val!13 val!47 -> 189
-  val!14 val!44 -> 126
-  val!14 val!50 -> 185
-  val!14 val!47 -> 186
-  val!10 val!45 -> 156
-  val!8 val!45 -> 162
-  val!10 val!49 -> 207
-  val!8 val!49 -> 208
-  val!21 val!45 -> 162
-  val!20 val!45 -> 156
-  val!21 val!49 -> 209
-  val!20 val!49 -> 206
-  val!13 val!45 -> 162
-  val!13 val!49 -> 208
-  val!14 val!45 -> 156
-  val!14 val!49 -> 207
-  else -> 207
-}
-f8 -> {
-  val!18 val!1 -> 2
-  val!19 val!1 -> 4
-  val!17 val!1 -> 3
-  val!19 val!3 -> 7
-  val!17 val!3 -> 8
-  val!18 val!3 -> 5
-  val!16 val!3 -> 6
-  val!23 val!27 -> 0
-  val!9 val!27 -> -1
-  val!11 val!27 -> 1
-  val!24 val!27 -> 1
-  val!22 val!27 -> -1
-  val!23 val!2 -> 0
-  val!9 val!2 -> -1
-  val!11 val!2 -> 1
-  val!24 val!2 -> 1
-  val!22 val!2 -> -1
-  val!23 val!4 -> 0
-  val!9 val!4 -> -1
-  val!11 val!4 -> 1
-  val!24 val!4 -> 1
-  val!22 val!4 -> -1
-  val!23 val!5 -> 0
-  val!9 val!5 -> -1
-  val!11 val!5 -> 1
-  val!24 val!5 -> 1
-  val!22 val!5 -> -1
-  val!22 val!1 -> -1
-  val!9 val!3 -> -1
-  val!24 val!1 -> 1
-  val!11 val!3 -> 1
-  val!23 val!30 -> 0
-  val!9 val!30 -> -1
-  val!11 val!30 -> 1
-  val!24 val!30 -> 1
-  val!22 val!30 -> -1
-  val!16 val!30 -> 29
-  val!16 val!36 -> 31
-  val!16 val!40 -> 32
-  val!9 val!1 -> -1
-  val!11 val!1 -> 1
-  val!22 val!3 -> -1
-  val!24 val!3 -> 1
-  val!17 val!2 -> 33
-  val!17 val!5 -> 34
-  val!17 val!27 -> 35
-  val!17 val!4 -> 36
-  val!17 val!30 -> 37
-  val!19 val!2 -> 38
-  val!19 val!5 -> 39
-  val!19 val!27 -> 40
-  val!19 val!4 -> 41
-  val!18 val!30 -> 43
-  val!19 val!30 -> 42
-  val!23 val!3 -> 0
-  val!23 val!1 -> 0
-  val!23 val!31 -> 0
-  val!9 val!31 -> -1
-  val!11 val!31 -> 1
-  val!24 val!31 -> 1
-  val!22 val!31 -> -1
-  val!23 val!33 -> 0
-  val!9 val!33 -> -1
-  val!11 val!33 -> 1
-  val!24 val!33 -> 1
-  val!22 val!33 -> -1
-  val!23 val!32 -> 0
-  val!9 val!32 -> -1
-  val!11 val!32 -> 1
-  val!24 val!32 -> 1
-  val!22 val!32 -> -1
-  val!23 val!34 -> 0
-  val!9 val!34 -> -1
-  val!11 val!34 -> 1
-  val!24 val!34 -> 1
-  val!22 val!34 -> -1
-  val!17 val!31 -> 67
-  val!17 val!33 -> 68
-  val!17 val!32 -> 69
-  val!17 val!34 -> 70
-  val!19 val!31 -> 71
-  val!19 val!33 -> 72
-  val!19 val!32 -> 73
-  val!19 val!34 -> 74
-  val!23 val!37 -> 0
-  val!9 val!37 -> -1
-  val!11 val!37 -> 1
-  val!24 val!37 -> 1
-  val!22 val!37 -> -1
-  val!23 val!38 -> 0
-  val!9 val!38 -> -1
-  val!11 val!38 -> 1
-  val!24 val!38 -> 1
-  val!22 val!38 -> -1
-  val!23 val!35 -> 0
-  val!9 val!35 -> -1
-  val!11 val!35 -> 1
-  val!24 val!35 -> 1
-  val!22 val!35 -> -1
-  val!23 val!36 -> 0
-  val!9 val!36 -> -1
-  val!11 val!36 -> 1
-  val!24 val!36 -> 1
-  val!22 val!36 -> -1
-  val!17 val!38 -> 95
-  val!17 val!36 -> 96
-  val!16 val!37 -> 97
-  val!17 val!37 -> 98
-  val!17 val!35 -> 99
-  val!19 val!38 -> 100
-  val!18 val!36 -> 101
-  val!19 val!36 -> 30
-  val!18 val!37 -> 103
-  val!19 val!37 -> 102
-  val!19 val!35 -> 104
-  val!23 val!39 -> 0
-  val!9 val!39 -> -1
-  val!11 val!39 -> 1
-  val!24 val!39 -> 1
-  val!22 val!39 -> -1
-  val!23 val!41 -> 0
-  val!9 val!41 -> -1
-  val!11 val!41 -> 1
-  val!24 val!41 -> 1
-  val!22 val!41 -> -1
-  val!23 val!40 -> 0
-  val!9 val!40 -> -1
-  val!11 val!40 -> 1
-  val!24 val!40 -> 1
-  val!22 val!40 -> -1
-  val!17 val!39 -> 125
-  val!17 val!41 -> 127
-  val!19 val!39 -> 128
-  val!19 val!41 -> 130
-  val!18 val!40 -> 131
-  val!23 val!48 -> 0
-  val!9 val!48 -> -1
-  val!11 val!48 -> 1
-  val!24 val!48 -> 1
-  val!22 val!48 -> -1
-  val!23 val!46 -> 0
-  val!9 val!46 -> -1
-  val!11 val!46 -> 1
-  val!24 val!46 -> 1
-  val!22 val!46 -> -1
-  val!23 val!42 -> 0
-  val!11 val!42 -> 1
-  val!9 val!42 -> -1
-  val!22 val!42 -> -1
-  val!24 val!42 -> 1
-  val!23 val!43 -> 0
-  val!9 val!43 -> -1
-  val!11 val!43 -> 1
-  val!24 val!43 -> 1
-  val!22 val!43 -> -1
-  val!17 val!42 -> 152
-  val!17 val!48 -> 153
-  val!16 val!43 -> 154
-  val!17 val!43 -> 155
-  val!17 val!46 -> 157
-  val!19 val!42 -> 158
-  val!19 val!48 -> 159
-  val!18 val!43 -> 161
-  val!19 val!43 -> 160
-  val!19 val!46 -> 163
-  val!23 val!50 -> 0
-  val!9 val!50 -> -1
-  val!11 val!50 -> 1
-  val!24 val!50 -> 1
-  val!22 val!50 -> -1
-  val!23 val!44 -> 0
-  val!9 val!44 -> -1
-  val!11 val!44 -> 1
-  val!24 val!44 -> 1
-  val!22 val!44 -> -1
-  val!23 val!47 -> 0
-  val!9 val!47 -> -1
-  val!11 val!47 -> 1
-  val!22 val!47 -> -1
-  val!24 val!47 -> 1
-  val!17 val!44 -> 126
-  val!16 val!50 -> 184
-  val!17 val!50 -> 185
-  val!17 val!47 -> 186
-  val!19 val!44 -> 129
-  val!18 val!50 -> 188
-  val!19 val!50 -> 187
-  val!19 val!47 -> 189
-  val!23 val!45 -> 0
-  val!9 val!45 -> -1
-  val!11 val!45 -> 1
-  val!24 val!45 -> 1
-  val!22 val!45 -> -1
-  val!23 val!49 -> 0
-  val!9 val!49 -> -1
-  val!11 val!49 -> 1
-  val!24 val!49 -> 1
-  val!22 val!49 -> -1
-  val!17 val!45 -> 156
-  val!16 val!49 -> 206
-  val!17 val!49 -> 207
-  val!19 val!45 -> 162
-  val!18 val!49 -> 209
-  val!19 val!49 -> 208
-  else -> 208
-}
-f30 -> {
-  val!1 val!26 -> val!28
-  val!3 val!26 -> val!29
-  val!27 val!26 -> val!51
-  val!2 val!26 -> val!52
-  val!4 val!26 -> val!53
-  val!5 val!26 -> val!54
-  val!30 val!26 -> val!55
-  val!34 val!26 -> val!56
-  val!33 val!26 -> val!57
-  val!32 val!26 -> val!58
-  val!31 val!26 -> val!59
-  val!37 val!26 -> val!60
-  val!35 val!26 -> val!61
-  val!38 val!26 -> val!62
-  val!36 val!26 -> val!63
-  val!39 val!26 -> val!64
-  val!41 val!26 -> val!65
-  val!40 val!26 -> val!67
-  val!48 val!26 -> val!68
-  val!42 val!26 -> val!69
-  val!43 val!26 -> val!71
-  val!46 val!26 -> val!72
-  val!44 val!26 -> val!66
-  val!47 val!26 -> val!73
-  val!50 val!26 -> val!74
-  val!45 val!26 -> val!70
-  val!49 val!26 -> val!75
-  else -> val!75
-}
-unknown
-68b48c983404c9f923b5c2c801b81a1ebfa54f09 60 0
-f1 -> val!0
-f2 -> val!1
-f7 -> val!3
-f8 -> val!5
-f12 -> val!13
-f11 -> val!6
-f17 -> val!8
-f18 -> val!9
-f16 -> val!10
-f20 -> val!11
-f14 -> val!12
-f6 -> {
-  val!3 -> val!2
-  val!5 -> val!4
-  else -> val!4
-}
-f5 -> {
-  val!2 -> val!3
-  val!4 -> val!5
-  else -> val!5
-}
-f4 -> {
-  val!5 -> 2
-  val!13 -> 40
-  val!14 -> 1
-  val!3 -> 3
-  val!17 -> 0
-  val!4 -> 1236
-  else -> 1236
-}
-f3 -> {
-  1 -> val!14
-  2 -> val!5
-  40 -> val!13
-  3 -> val!3
-  0 -> val!17
-  1236 -> val!4
-  else -> val!4
-}
-f10 -> {
-  val!6 val!13 -> val!7
-  else -> val!7
-}
-f13 -> {
-  val!12 -> val!17
-  else -> val!17
-}
-f9 -> {
-  val!3 val!7 -> val!15
-  val!5 val!7 -> val!16
-  else -> val!16
-}
-f15 -> {
-  val!9 val!2 -> 0
-  val!8 val!2 -> -1
-  val!10 val!2 -> 1
-  val!11 val!2 -> 1
-  else -> 1
-}
-unknown
-2cf0827cf5826be278dbca8a4dccba886e2b01ef 611 0
-f1 -> val!7
-f2 -> val!0
-f16 -> val!2
-f7 -> val!4
-f24 -> val!31
-f29 -> val!16
-f18 -> val!8
-f4 -> val!10
-f13 -> val!12
-f17 -> val!14
-f12 -> val!18
-f9 -> val!19
-f14 -> val!20
-f10 -> val!21
-f15 -> val!22
-f19 -> val!23
-f23 -> val!24
-f28 -> val!27
-f26 -> val!29
-f11 -> {
-  val!2 -> val!1
-  val!4 -> val!3
-  else -> val!3
-}
-f6 -> {
-  val!1 -> val!2
-  val!3 -> val!4
-  val!36 -> val!40
-  val!31 -> val!41
-  val!5 -> val!37
-  val!2 -> val!38
-  val!4 -> val!39
-  val!37 -> val!43
-  val!38 -> val!54
-  val!39 -> val!44
-  val!40 -> val!45
-  val!41 -> val!42
-  val!43 -> val!47
-  val!44 -> val!46
-  val!45 -> val!56
-  val!42 -> val!51
-  val!46 -> val!53
-  val!47 -> val!48
-  val!48 -> val!49
-  val!49 -> val!50
-  val!50 -> val!49
-  val!51 -> val!52
-  val!54 -> val!59
-  val!53 -> val!55
-  val!57 -> val!54
-  val!55 -> val!2
-  val!56 -> val!58
-  val!52 -> val!57
-  val!58 -> val!54
-  val!59 -> val!60
-  val!60 -> val!86
-  val!86 -> val!2
-  else -> val!2
-}
-f20 -> {
-  1 -> val!5
-  7720 -> val!36
-  7719 -> val!31
-  31267 -> val!2
-  36564 -> val!4
-  8366 -> val!37
-  9246 -> val!38
-  8852 -> val!39
-  0 -> val!40
-  2944 -> val!41
-  26221 -> val!42
-  29498 -> val!43
-  15353 -> val!44
-  6996 -> val!45
-  2997 -> val!46
-  9245 -> val!47
-  19679 -> val!49
-  45859 -> val!50
-  9531 -> val!51
-  24738 -> val!52
-  30204 -> val!57
-  9247 -> val!53
-  23570 -> val!55
-  26809 -> val!56
-  23086 -> val!58
-  5529 -> val!59
-  36563 -> val!60
-  36565 -> val!48
-  27246 -> val!86
-  31266 -> val!54
-  else -> val!54
-}
-f22 -> {
-  val!5 val!31 -> val!6
-  val!5 val!36 -> val!17
-  val!24 val!31 -> val!25
-  else -> val!25
-}
-f21 -> {
-  val!2 val!6 -> val!7
-  val!4 val!6 -> val!7
-  val!4 val!17 -> val!7
-  val!2 val!25 -> val!32
-  val!4 val!25 -> val!33
-  else -> val!33
-}
-f27 -> {
-  val!8 -> val!9
-  val!10 -> val!11
-  val!12 -> val!13
-  val!14 -> val!15
-  val!23 -> val!26
-  val!22 -> val!28
-  else -> val!28
-}
-f30 -> {
-  val!13 val!15 -> val!16
-  else -> val!16
-}
-f5 -> {
-  val!31 -> 7719
-  val!2 -> 31267
-  val!4 -> 36564
-  val!5 -> 1
-  val!36 -> 7720
-  val!40 -> 0
-  val!41 -> 2944
-  val!37 -> 8366
-  val!38 -> 9246
-  val!39 -> 8852
-  val!44 -> 15353
-  val!43 -> 29498
-  val!42 -> 26221
-  val!45 -> 6996
-  val!46 -> 2997
-  val!56 -> 26809
-  val!51 -> 9531
-  val!53 -> 9247
-  val!47 -> 9245
-  val!48 -> 36565
-  val!49 -> 19679
-  val!50 -> 45859
-  val!52 -> 24738
-  val!59 -> 5529
-  val!55 -> 23570
-  val!54 -> 31266
-  val!58 -> 23086
-  val!57 -> 30204
-  val!60 -> 36563
-  val!86 -> 27246
-  else -> 27246
-}
-f25 -> {
-  val!29 -> val!30
-  else -> val!30
-}
-f3 -> {
-  val!23 val!1 -> 2
-  val!8 val!1 -> 4
-  val!14 val!1 -> 3
-  val!22 val!1 -> 4
-  val!10 val!1 -> 3
-  val!12 val!1 -> 4
-  val!23 val!3 -> 5
-  val!8 val!3 -> 5
-  val!14 val!3 -> 6
-  val!22 val!3 -> 6
-  val!10 val!3 -> 7
-  val!12 val!3 -> 8
-  val!22 val!31 -> 32
-  val!22 val!37 -> 62
-  val!22 val!40 -> 60
-  val!22 val!43 -> 88
-  val!10 val!5 -> 30
-  val!8 val!5 -> 35
-  val!22 val!60 -> 150
-  val!23 val!36 -> 37
-  val!22 val!36 -> 31
-  val!23 val!31 -> 36
-  val!10 val!2 -> 28
-  val!8 val!2 -> 33
-  val!23 val!5 -> 35
-  val!22 val!5 -> 30
-  val!10 val!31 -> 32
-  val!8 val!31 -> 36
-  val!23 val!2 -> 33
-  val!22 val!2 -> 28
-  val!10 val!4 -> 29
-  val!8 val!4 -> 34
-  val!23 val!4 -> 34
-  val!22 val!4 -> 29
-  val!10 val!36 -> 31
-  val!8 val!36 -> 37
-  val!12 val!2 -> 33
-  val!12 val!4 -> 34
-  val!12 val!5 -> 35
-  val!12 val!31 -> 36
-  val!12 val!36 -> 37
-  val!14 val!2 -> 28
-  val!14 val!4 -> 29
-  val!14 val!5 -> 30
-  val!14 val!31 -> 32
-  val!14 val!36 -> 31
-  val!10 val!38 -> 59
-  val!8 val!38 -> 64
-  val!23 val!38 -> 64
-  val!22 val!38 -> 59
-  val!10 val!39 -> 58
-  val!8 val!39 -> 63
-  val!10 val!37 -> 62
-  val!8 val!37 -> 67
-  val!23 val!39 -> 63
-  val!22 val!39 -> 58
-  val!10 val!40 -> 60
-  val!8 val!40 -> 65
-  val!23 val!40 -> 65
-  val!23 val!41 -> 66
-  val!22 val!41 -> 61
-  val!23 val!37 -> 67
-  val!10 val!41 -> 61
-  val!8 val!41 -> 66
-  val!12 val!39 -> 63
-  val!12 val!38 -> 64
-  val!12 val!40 -> 65
-  val!12 val!41 -> 66
-  val!12 val!37 -> 67
-  val!14 val!39 -> 58
-  val!14 val!38 -> 59
-  val!14 val!40 -> 60
-  val!14 val!41 -> 61
-  val!14 val!37 -> 62
-  val!10 val!42 -> 89
-  val!8 val!42 -> 93
-  val!23 val!54 -> 94
-  val!22 val!54 -> 90
-  val!10 val!43 -> 88
-  val!8 val!43 -> 92
-  val!23 val!43 -> 92
-  val!23 val!42 -> 93
-  val!22 val!42 -> 89
-  val!23 val!44 -> 95
-  val!22 val!44 -> 91
-  val!10 val!44 -> 91
-  val!8 val!44 -> 95
-  val!10 val!54 -> 90
-  val!8 val!54 -> 94
-  val!23 val!45 -> 107
-  val!22 val!45 -> 106
-  val!10 val!45 -> 106
-  val!8 val!45 -> 107
-  val!12 val!43 -> 92
-  val!12 val!42 -> 93
-  val!12 val!54 -> 94
-  val!12 val!44 -> 95
-  val!14 val!43 -> 88
-  val!14 val!42 -> 89
-  val!14 val!54 -> 90
-  val!14 val!44 -> 91
-  val!12 val!45 -> 107
-  val!14 val!45 -> 106
-  val!10 val!46 -> 118
-  val!8 val!46 -> 122
-  val!23 val!46 -> 122
-  val!22 val!46 -> 118
-  val!23 val!47 -> 123
-  val!22 val!47 -> 127
-  val!8 val!59 -> 126
-  val!10 val!59 -> 120
-  val!23 val!59 -> 125
-  val!22 val!59 -> 119
-  val!10 val!47 -> 127
-  val!8 val!47 -> 123
-  val!10 val!51 -> 116
-  val!8 val!51 -> 124
-  val!23 val!51 -> 124
-  val!22 val!51 -> 116
-  val!23 val!56 -> 121
-  val!22 val!56 -> 117
-  val!10 val!56 -> 117
-  val!8 val!56 -> 121
-  val!12 val!56 -> 121
-  val!12 val!46 -> 122
-  val!12 val!47 -> 123
-  val!12 val!51 -> 124
-  val!12 val!59 -> 126
-  val!14 val!56 -> 117
-  val!14 val!46 -> 118
-  val!14 val!47 -> 127
-  val!14 val!51 -> 116
-  val!14 val!59 -> 120
-  val!23 val!58 -> 153
-  val!22 val!58 -> 148
-  val!10 val!53 -> 151
-  val!8 val!53 -> 156
-  val!23 val!53 -> 156
-  val!22 val!53 -> 151
-  val!10 val!48 -> 149
-  val!8 val!48 -> 154
-  val!23 val!48 -> 154
-  val!22 val!48 -> 149
-  val!23 val!52 -> 157
-  val!22 val!52 -> 152
-  val!10 val!58 -> 148
-  val!8 val!58 -> 153
-  val!10 val!60 -> 150
-  val!8 val!60 -> 155
-  val!10 val!52 -> 152
-  val!8 val!52 -> 157
-  val!23 val!60 -> 155
-  val!12 val!58 -> 153
-  val!12 val!48 -> 154
-  val!12 val!60 -> 155
-  val!12 val!53 -> 156
-  val!12 val!52 -> 157
-  val!14 val!58 -> 148
-  val!14 val!48 -> 149
-  val!14 val!60 -> 150
-  val!14 val!53 -> 151
-  val!14 val!52 -> 152
-  val!23 val!49 -> 186
-  val!22 val!49 -> 187
-  val!10 val!55 -> 179
-  val!8 val!55 -> 182
-  val!23 val!55 -> 178
-  val!22 val!55 -> 4
-  val!10 val!86 -> 181
-  val!8 val!86 -> 184
-  val!23 val!86 -> 185
-  val!22 val!86 -> 4
-  val!23 val!57 -> 183
-  val!22 val!57 -> 180
-  val!10 val!49 -> 187
-  val!8 val!49 -> 186
-  val!8 val!57 -> 183
-  val!10 val!57 -> 180
-  val!12 val!55 -> 182
-  val!12 val!57 -> 183
-  val!12 val!86 -> 184
-  val!12 val!49 -> 186
-  val!14 val!55 -> 179
-  val!14 val!57 -> 180
-  val!14 val!86 -> 181
-  val!14 val!49 -> 187
-  val!23 val!50 -> 205
-  val!22 val!50 -> 204
-  val!10 val!50 -> 204
-  val!8 val!50 -> 205
-  val!12 val!50 -> 205
-  val!14 val!50 -> 204
-  else -> 204
-}
-f8 -> {
-  val!20 val!1 -> 2
-  val!21 val!1 -> 4
-  val!19 val!1 -> 3
-  val!21 val!3 -> 7
-  val!19 val!3 -> 8
-  val!20 val!3 -> 5
-  val!18 val!3 -> 6
-  val!27 val!31 -> 0
-  val!9 val!31 -> -1
-  val!11 val!31 -> 1
-  val!28 val!31 -> 1
-  val!26 val!31 -> -1
-  val!27 val!2 -> 0
-  val!9 val!2 -> -1
-  val!11 val!2 -> 1
-  val!28 val!2 -> 1
-  val!26 val!2 -> -1
-  val!27 val!4 -> 0
-  val!9 val!4 -> -1
-  val!11 val!4 -> 1
-  val!28 val!4 -> 1
-  val!26 val!4 -> -1
-  val!27 val!5 -> 0
-  val!9 val!5 -> -1
-  val!11 val!5 -> 1
-  val!28 val!5 -> 1
-  val!26 val!5 -> -1
-  val!9 val!36 -> -1
-  val!27 val!36 -> 0
-  val!11 val!36 -> 1
-  val!28 val!36 -> 1
-  val!26 val!36 -> -1
-  val!19 val!2 -> 28
-  val!19 val!4 -> 29
-  val!19 val!5 -> 30
-  val!19 val!36 -> 31
-  val!19 val!31 -> 32
-  val!21 val!2 -> 33
-  val!21 val!4 -> 34
-  val!21 val!5 -> 35
-  val!21 val!31 -> 36
-  val!21 val!36 -> 37
-  val!27 val!38 -> 0
-  val!9 val!38 -> -1
-  val!11 val!38 -> 1
-  val!28 val!38 -> 1
-  val!26 val!38 -> -1
-  val!27 val!39 -> 0
-  val!9 val!39 -> -1
-  val!11 val!39 -> 1
-  val!28 val!39 -> 1
-  val!26 val!39 -> -1
-  val!27 val!37 -> 0
-  val!9 val!37 -> -1
-  val!11 val!37 -> 1
-  val!28 val!37 -> 1
-  val!26 val!37 -> -1
-  val!27 val!40 -> 0
-  val!9 val!40 -> -1
-  val!11 val!40 -> 1
-  val!28 val!40 -> 1
-  val!26 val!40 -> -1
-  val!27 val!41 -> 0
-  val!9 val!41 -> -1
-  val!11 val!41 -> 1
-  val!28 val!41 -> 1
-  val!26 val!41 -> -1
-  val!19 val!39 -> 58
-  val!19 val!38 -> 59
-  val!19 val!40 -> 60
-  val!19 val!41 -> 61
-  val!19 val!37 -> 62
-  val!21 val!39 -> 63
-  val!21 val!38 -> 64
-  val!21 val!40 -> 65
-  val!21 val!41 -> 66
-  val!21 val!37 -> 67
-  val!27 val!44 -> 0
-  val!9 val!44 -> -1
-  val!11 val!44 -> 1
-  val!28 val!44 -> 1
-  val!26 val!44 -> -1
-  val!27 val!45 -> 0
-  val!9 val!45 -> -1
-  val!11 val!45 -> 1
-  val!28 val!45 -> 1
-  val!26 val!45 -> -1
-  val!27 val!42 -> 0
-  val!9 val!42 -> -1
-  val!11 val!42 -> 1
-  val!28 val!42 -> 1
-  val!26 val!42 -> -1
-  val!27 val!43 -> 0
-  val!9 val!43 -> -1
-  val!11 val!43 -> 1
-  val!28 val!43 -> 1
-  val!26 val!43 -> -1
-  val!19 val!43 -> 88
-  val!19 val!42 -> 89
-  val!19 val!54 -> 90
-  val!19 val!44 -> 91
-  val!21 val!43 -> 92
-  val!21 val!42 -> 93
-  val!21 val!54 -> 94
-  val!21 val!44 -> 95
-  val!27 val!47 -> 0
-  val!9 val!47 -> -1
-  val!11 val!47 -> 1
-  val!28 val!47 -> 1
-  val!26 val!47 -> -1
-  val!27 val!59 -> 0
-  val!11 val!59 -> 1
-  val!9 val!59 -> -1
-  val!28 val!59 -> 1
-  val!26 val!59 -> -1
-  val!27 val!46 -> 0
-  val!9 val!46 -> -1
-  val!11 val!46 -> 1
-  val!28 val!46 -> 1
-  val!26 val!46 -> -1
-  val!19 val!45 -> 106
-  val!21 val!45 -> 107
-  val!27 val!51 -> 0
-  val!9 val!51 -> -1
-  val!11 val!51 -> 1
-  val!28 val!51 -> 1
-  val!26 val!51 -> -1
-  val!27 val!56 -> 0
-  val!9 val!56 -> -1
-  val!11 val!56 -> 1
-  val!28 val!56 -> 1
-  val!26 val!56 -> -1
-  val!19 val!56 -> 117
-  val!19 val!46 -> 118
-  val!18 val!47 -> 127
-  val!19 val!51 -> 116
-  val!18 val!59 -> 119
-  val!19 val!59 -> 120
-  val!21 val!56 -> 121
-  val!21 val!46 -> 122
-  val!20 val!47 -> 123
-  val!21 val!51 -> 124
-  val!20 val!59 -> 125
-  val!21 val!59 -> 126
-  val!27 val!58 -> 0
-  val!9 val!58 -> -1
-  val!11 val!58 -> 1
-  val!28 val!58 -> 1
-  val!26 val!58 -> -1
-  val!27 val!53 -> 0
-  val!9 val!53 -> -1
-  val!11 val!53 -> 1
-  val!28 val!53 -> 1
-  val!26 val!53 -> -1
-  val!27 val!48 -> 0
-  val!9 val!48 -> -1
-  val!11 val!48 -> 1
-  val!28 val!48 -> 1
-  val!26 val!48 -> -1
-  val!27 val!52 -> 0
-  val!9 val!52 -> -1
-  val!11 val!52 -> 1
-  val!28 val!52 -> 1
-  val!26 val!52 -> -1
-  val!27 val!60 -> 0
-  val!9 val!60 -> -1
-  val!11 val!60 -> 1
-  val!26 val!60 -> -1
-  val!28 val!60 -> 1
-  val!19 val!58 -> 148
-  val!19 val!48 -> 149
-  val!19 val!60 -> 150
-  val!19 val!53 -> 151
-  val!19 val!52 -> 152
-  val!21 val!58 -> 153
-  val!21 val!48 -> 154
-  val!21 val!60 -> 155
-  val!21 val!53 -> 156
-  val!21 val!52 -> 157
-  val!27 val!49 -> 0
-  val!9 val!49 -> -1
-  val!11 val!49 -> 1
-  val!28 val!49 -> 1
-  val!26 val!49 -> -1
-  val!27 val!86 -> 0
-  val!9 val!86 -> -1
-  val!11 val!86 -> 1
-  val!28 val!86 -> 1
-  val!26 val!86 -> -1
-  val!27 val!55 -> 0
-  val!9 val!55 -> -1
-  val!11 val!55 -> 1
-  val!28 val!55 -> 1
-  val!26 val!55 -> -1
-  val!27 val!57 -> 0
-  val!11 val!57 -> 1
-  val!9 val!57 -> -1
-  val!28 val!57 -> 1
-  val!26 val!57 -> -1
-  val!19 val!55 -> 179
-  val!19 val!57 -> 180
-  val!19 val!86 -> 181
-  val!18 val!49 -> 187
-  val!20 val!55 -> 178
-  val!21 val!55 -> 182
-  val!21 val!57 -> 183
-  val!20 val!86 -> 185
-  val!21 val!86 -> 184
-  val!20 val!49 -> 186
-  val!27 val!50 -> 0
-  val!9 val!50 -> -1
-  val!11 val!50 -> 1
-  val!28 val!50 -> 1
-  val!26 val!50 -> -1
-  val!27 val!54 -> 0
-  val!9 val!54 -> -1
-  val!11 val!54 -> 1
-  val!28 val!54 -> 1
-  val!26 val!54 -> -1
-  val!19 val!50 -> 204
-  val!21 val!50 -> 205
-  else -> 205
-}
-f31 -> {
-  val!1 val!30 -> val!34
-  val!3 val!30 -> val!35
-  val!5 val!30 -> val!61
-  val!36 val!30 -> val!62
-  val!31 val!30 -> val!63
-  val!2 val!30 -> val!64
-  val!4 val!30 -> val!65
-  val!38 val!30 -> val!66
-  val!39 val!30 -> val!67
-  val!37 val!30 -> val!68
-  val!40 val!30 -> val!69
-  val!41 val!30 -> val!70
-  val!42 val!30 -> val!71
-  val!54 val!30 -> val!72
-  val!43 val!30 -> val!73
-  val!44 val!30 -> val!74
-  val!45 val!30 -> val!75
-  val!46 val!30 -> val!76
-  val!47 val!30 -> val!77
-  val!59 val!30 -> val!78
-  val!51 val!30 -> val!79
-  val!56 val!30 -> val!80
-  val!58 val!30 -> val!81
-  val!53 val!30 -> val!82
-  val!48 val!30 -> val!83
-  val!52 val!30 -> val!84
-  val!60 val!30 -> val!85
-  val!49 val!30 -> val!87
-  val!55 val!30 -> val!88
-  val!86 val!30 -> val!89
-  val!57 val!30 -> val!90
-  val!50 val!30 -> val!91
-  else -> val!91
-}
-unknown
-a1ff1dee861d393a5412b6a4273eb86deab7593f 77 0
-f1 -> val!0
-f2 -> val!1
-f7 -> val!3
-f8 -> val!5
-f14 -> val!6
-f11 -> val!7
-f12 -> val!8
-f17 -> val!10
-f18 -> val!11
-f16 -> val!12
-f20 -> val!13
-f19 -> val!14
-f6 -> {
-  val!3 -> val!2
-  val!5 -> val!4
-  else -> val!4
-}
-f5 -> {
-  val!2 -> val!3
-  val!4 -> val!5
-  val!16 -> val!19
-  val!5 -> val!16
-  val!19 -> val!16
-  val!3 -> val!16
-  else -> val!16
-}
-f13 -> {
-  val!6 -> val!15
-  else -> val!15
-}
-f4 -> {
-  val!15 -> 40
-  val!5 -> 2
-  val!16 -> 1
-  val!3 -> 3
-  val!4 -> 1276
-  val!19 -> 0
-  else -> 0
-}
-f3 -> {
-  1 -> val!16
-  40 -> val!15
-  2 -> val!5
-  3 -> val!3
-  0 -> val!19
-  1276 -> val!4
-  else -> val!4
-}
-f10 -> {
-  val!7 val!8 -> val!9
-  else -> val!9
-}
-f9 -> {
-  val!3 val!9 -> val!17
-  val!5 val!9 -> val!18
-  else -> val!18
-}
-f15 -> {
-  val!11 val!2 -> 0
-  val!10 val!2 -> -1
-  val!12 val!2 -> 1
-  val!13 val!2 -> 1
-  val!11 val!16 -> 0
-  val!12 val!16 -> -1
-  val!14 val!16 -> 1
-  val!11 val!3 -> 0
-  val!12 val!3 -> -1
-  val!14 val!3 -> 1
-  val!11 val!5 -> 0
-  val!12 val!5 -> -1
-  val!14 val!5 -> 1
-  val!11 val!19 -> 0
-  val!12 val!19 -> -1
-  val!14 val!19 -> 1
-  else -> 1
-}
-unknown
-26a56d9f61c23c26da64882e31e65a36923764cb 72 0
-f1 -> val!0
-f2 -> val!1
-f7 -> val!3
-f8 -> val!5
-f12 -> val!13
-f11 -> val!6
-f17 -> val!8
-f18 -> val!9
-f16 -> val!10
-f20 -> val!11
-f19 -> val!12
-f6 -> {
-  val!3 -> val!2
-  val!5 -> val!4
-  else -> val!4
-}
-f5 -> {
-  val!2 -> val!3
-  val!4 -> val!5
-  val!3 -> val!17
-  val!14 -> val!14
-  val!5 -> val!14
-  val!17 -> val!14
-  else -> val!14
-}
-f4 -> {
-  val!13 -> 40
-  val!5 -> 2
-  val!14 -> 1
-  val!3 -> 3
-  val!4 -> 1276
-  val!17 -> 0
-  else -> 0
-}
-f3 -> {
-  1 -> val!14
-  40 -> val!13
-  2 -> val!5
-  3 -> val!3
-  1276 -> val!4
-  0 -> val!17
-  else -> val!17
-}
-f10 -> {
-  val!6 val!13 -> val!7
-  else -> val!7
-}
-f9 -> {
-  val!3 val!7 -> val!15
-  val!5 val!7 -> val!16
-  else -> val!16
-}
-f15 -> {
-  val!9 val!2 -> 0
-  val!8 val!2 -> -1
-  val!10 val!2 -> 1
-  val!11 val!2 -> 1
-  val!9 val!3 -> 0
-  val!10 val!3 -> -1
-  val!12 val!3 -> 1
-  val!9 val!14 -> 0
-  val!10 val!14 -> -1
-  val!12 val!14 -> 1
-  val!9 val!5 -> 0
-  val!10 val!5 -> -1
-  val!12 val!5 -> 1
-  val!9 val!17 -> 0
-  val!10 val!17 -> -1
-  val!12 val!17 -> 1
-  else -> 1
-}
-unknown
-a4a52206bfedbc1fa95069df86dc718ab11f5b3a 59 0
-f1 -> val!0
-f2 -> val!1
-f7 -> val!3
-f8 -> val!5
-f17 -> val!6
-f20 -> val!7
-f19 -> val!8
-f18 -> val!9
-f16 -> val!10
-f14 -> val!11
-f11 -> val!12
-f12 -> val!13
-f6 -> {
-  val!3 -> val!2
-  val!5 -> val!4
-  else -> val!4
-}
-f5 -> {
-  val!2 -> val!3
-  val!4 -> val!5
-  else -> val!5
-}
-f15 -> {
-  val!6 val!2 -> 0
-  val!7 val!2 -> 1
-  val!8 val!2 -> -1
-  val!6 val!4 -> 0
-  val!9 val!4 -> 1
-  val!10 val!4 -> -1
-  else -> -1
-}
-f4 -> {
-  val!3 -> 1237
-  val!5 -> 1238
-  val!15 -> 8957
-  val!16 -> 1
-  else -> 1
-}
-f13 -> {
-  val!11 -> val!15
-  else -> val!15
-}
-f3 -> {
-  1 -> val!16
-  1237 -> val!3
-  1238 -> val!5
-  8957 -> val!15
-  else -> val!15
-}
-f10 -> {
-  val!12 val!13 -> val!14
-  else -> val!14
-}
-f9 -> {
-  val!3 val!14 -> val!17
-  val!5 val!14 -> val!18
-  else -> val!18
-}
-unknown
-14333af854a5b65017dd8bd92066cf80fe83d74b 43 0
-f1 -> val!0
-f2 -> val!1
-f15 -> val!2
-f6 -> val!9
-f12 -> val!3
-f16 -> val!4
-f5 -> val!9
-f13 -> val!5
-f10 -> val!7
-f8 -> val!8
-f11 -> {
-  val!9 -> val!6
-  else -> val!6
-}
-f9 -> {
-  val!5 val!6 -> 1
-  val!3 val!6 -> 0
-  val!4 val!6 -> 1
-  val!2 val!6 -> -1
-  val!7 val!6 -> -1
-  else -> -1
-}
-f4 -> {
-  val!9 -> 7720
-  val!10 -> 7758
-  val!11 -> 1
-  else -> 1
-}
-f14 -> {
-  val!6 -> val!9
-  else -> val!9
-}
-f7 -> {
-  val!8 -> val!10
-  else -> val!10
-}
-f3 -> {
-  1 -> val!11
-  7720 -> val!9
-  7758 -> val!10
-  else -> val!10
-}
-unknown
-5f2c36ac6c49043bec7b255aa0d7a0c690b930a4 21 0
-f1 -> val!0
-f2 -> val!1
-f17 -> val!2
-f4 -> val!4
-f14 -> val!7
-f16 -> val!9
-f7 -> val!12
-f10 -> val!12
-f20 -> {
-  val!2 -> val!3
-  val!4 -> val!5
-  val!7 -> val!8
-  val!9 -> val!10
-  else -> val!10
-}
-f19 -> {
-  val!3 val!5 -> val!6
-  val!8 val!10 -> val!11
-  else -> val!11
-}
-unknown
-4bd0e90cc9ab46b702c811a3ea671a168ec22aba 65 0
-f1 -> val!0
-f2 -> val!1
-f14 -> val!19
-f21 -> val!3
-f13 -> val!6
-f23 -> val!8
-f17 -> val!9
-f16 -> val!11
-f4 -> val!13
-f7 -> val!19
-f10 -> val!15
-f9 -> val!16
-f12 -> val!17
-f15 -> val!18
-f11 -> {
-  val!19 -> val!2
-  else -> val!2
-}
-f20 -> {
-  val!3 -> val!4
-  else -> val!4
-}
-f19 -> {
-  val!2 val!4 -> val!5
-  else -> val!5
-}
-f22 -> {
-  val!6 -> val!7
-  val!9 -> val!10
-  val!11 -> val!12
-  val!13 -> val!14
-  else -> val!14
-}
-f8 -> {
-  val!7 val!2 -> -1
-  val!8 val!2 -> 0
-  val!10 val!2 -> 1
-  val!12 val!2 -> -1
-  val!14 val!2 -> 1
-  val!15 val!2 -> 0
-  val!16 val!2 -> -1
-  val!17 val!2 -> 1
-  val!18 val!2 -> 0
-  else -> 0
-}
-f5 -> {
-  val!19 -> 0
-  else -> 0
-}
-f18 -> {
-  0 -> val!19
-  else -> val!19
-}
-f6 -> {
-  val!2 -> val!19
-  else -> val!19
-}
-f3 -> {
-  val!13 val!2 -> 0
-  val!6 val!2 -> -1
-  val!11 val!2 -> 0
-  val!9 val!2 -> 1
-  else -> 1
-}
-unknown
-cdc3b5c1b7947a0fe95f5d2f65f941e2774116a3 63 0
-f1 -> val!0
-f2 -> val!1
-f14 -> val!17
-f21 -> val!3
-f13 -> val!6
-f23 -> val!8
-f17 -> val!9
-f16 -> val!11
-f4 -> val!13
-f7 -> val!17
-f10 -> val!15
-f9 -> val!16
-f11 -> {
-  val!17 -> val!2
-  else -> val!2
-}
-f20 -> {
-  val!3 -> val!4
-  else -> val!4
-}
-f19 -> {
-  val!2 val!4 -> val!5
-  else -> val!5
-}
-f22 -> {
-  val!6 -> val!7
-  val!9 -> val!10
-  val!11 -> val!12
-  val!13 -> val!14
-  else -> val!14
-}
-f8 -> {
-  val!7 val!2 -> -1
-  val!8 val!2 -> 0
-  val!10 val!2 -> 1
-  val!12 val!2 -> -1
-  val!14 val!2 -> 1
-  val!15 val!2 -> 0
-  val!16 val!2 -> 2
-  else -> 2
-}
-f5 -> {
-  val!17 -> 1
-  val!18 -> 0
-  else -> 0
-}
-f18 -> {
-  1 -> val!17
-  0 -> val!18
-  else -> val!18
-}
-f3 -> {
-  val!13 val!2 -> 2
-  val!11 val!2 -> 0
-  val!6 val!2 -> 0
-  val!9 val!2 -> 2
-  else -> 2
-}
-f6 -> {
-  val!2 -> val!18
-  else -> val!18
-}
-unknown
-e7b427c80202d9c4c29f2a9278c2bcdecdabe460 42 0
-f1 -> val!0
-f2 -> val!1
-f3 -> val!2
-f4 -> val!2
-f5 -> val!3
-f6 -> val!3
-f18 -> val!4
-f19 -> val!5
-f20 -> val!6
-f17 -> val!8
-f23 -> val!14
-f13 -> {
-  val!4 val!5 -> 0
-  val!6 val!5 -> -3
-  val!10 val!5 -> 0
-  val!13 val!5 -> 0
-  else -> 0
-}
-f16 -> {
-  val!3 -> val!7
-  val!2 -> val!11
-  else -> val!11
-}
-f15 -> {
-  val!7 val!8 -> val!9
-  val!11 val!8 -> val!12
-  else -> val!12
-}
-f14 -> {
-  val!9 val!4 -> val!10
-  val!12 val!6 -> val!13
-  else -> val!13
-}
-f22 -> {
-  val!14 -> val!15
-  else -> val!15
-}
-f21 -> {
-  val!5 val!15 -> val!16
-  else -> val!16
-}
-unknown
-fc367e02b7e503644f5c70cc9716512639619a51 52 0
-f1 -> val!0
-f2 -> val!1
-f3 -> val!2
-f4 -> val!2
-f5 -> val!3
-f6 -> val!3
-f20 -> val!4
-f21 -> val!15
-f22 -> val!5
-f19 -> val!7
-f25 -> val!13
-f15 -> {
-  val!4 val!15 -> 6
-  val!5 val!15 -> 0
-  val!9 val!15 -> 1
-  val!12 val!15 -> -1
-  else -> -1
-}
-f18 -> {
-  val!3 -> val!6
-  val!2 -> val!10
-  else -> val!10
-}
-f17 -> {
-  val!6 val!7 -> val!8
-  val!10 val!7 -> val!11
-  else -> val!11
-}
-f16 -> {
-  val!8 val!4 -> val!9
-  val!11 val!5 -> val!12
-  else -> val!12
-}
-f24 -> {
-  val!13 -> val!16
-  else -> val!16
-}
-f23 -> {
-  val!15 val!16 -> val!14
-  else -> val!14
-}
-f14 -> {
-  val!15 -> 0
-  val!16 -> 1
-  else -> 1
-}
-f13 -> {
-  0 -> val!15
-  1 -> val!16
-  else -> val!16
-}
-unknown
-2b6a5d632b112638788d36706b106b4e4d10b262 35 0
-f1 -> val!0
-f2 -> val!1
-f3 -> val!2
-f4 -> val!2
-f5 -> val!3
-f6 -> val!3
-f18 -> val!4
-f19 -> val!5
-f20 -> val!6
-f17 -> val!8
-f13 -> {
-  val!4 val!5 -> 0
-  val!6 val!5 -> -3
-  val!10 val!5 -> 0
-  val!13 val!5 -> 0
-  val!9 val!5 -> 0
-  val!12 val!5 -> 0
-  else -> 0
-}
-f16 -> {
-  val!3 -> val!7
-  val!2 -> val!11
-  else -> val!11
-}
-f15 -> {
-  val!7 val!8 -> val!9
-  val!11 val!8 -> val!12
-  else -> val!12
-}
-f14 -> {
-  val!9 val!4 -> val!10
-  val!12 val!6 -> val!13
-  else -> val!13
-}
-unknown
-6220e161feeb3d35604e0055563cdf27e5108197 35 0
-f1 -> val!0
-f2 -> val!1
-f8 -> val!2
-f12 -> val!2
-f11 -> val!3
-f13 -> val!3
-f22 -> val!4
-f9 -> val!5
-f23 -> val!6
-f10 -> val!7
-f21 -> val!8
-f4 -> val!9
-f5 -> {
-  val!4 val!5 -> 3
-  val!6 val!5 -> 0
-  val!11 val!5 -> 0
-  val!13 val!5 -> 0
-  else -> 0
-}
-f20 -> {
-  val!7 val!8 -> 3
-  val!9 val!8 -> 0
-  else -> 0
-}
-f7 -> {
-  val!3 -> val!10
-  val!2 -> val!12
-  else -> val!12
-}
-f24 -> {
-  val!10 val!8 -> val!11
-  val!12 val!8 -> val!13
-  else -> val!13
-}
-unknown
-c29b6c64de39317156c532c91451be225c15adb2 238 0
-#2 := false
-#48 := 0::real
-decl f19 :: (-> S3 S10 real)
-decl f20 :: S10
-#43 := f20
-decl f4 :: S3
-#8 := f4
-#58 := (f19 f4 f20)
-#109 := -1::real
-#153 := (* -1::real #58)
-decl f5 :: (-> S4 S5 real)
-decl f8 :: S5
-#13 := f8
-decl f22 :: S4
-#54 := f22
-#55 := (f5 f22 f8)
-#154 := (+ #55 #153)
-#137 := (* -1::real #55)
-#144 := (+ #137 #58)
-#188 := (<= #154 0::real)
-#195 := (ite #188 #144 #154)
-#451 := (* -1::real #195)
-#452 := (+ #144 #451)
-#453 := (<= #452 0::real)
-#435 := (= #144 #195)
-decl f21 :: S4
-#45 := f21
-#46 := (f5 f21 f8)
-decl f9 :: S3
-#17 := f9
-#44 := (f19 f9 f20)
-#120 := (* -1::real #44)
-#121 := (+ #120 #46)
-#110 := (* -1::real #46)
-#111 := (+ #44 #110)
-#216 := (>= #111 0::real)
-#223 := (ite #216 #111 #121)
-#447 := (* -1::real #223)
-#450 := (+ #121 #447)
-#454 := (<= #450 0::real)
-#442 := (= #121 #223)
-#217 := (not #216)
-#455 := [hypothesis]: #216
-#184 := (+ #44 #153)
-#185 := (<= #184 0::real)
-#206 := -3::real
-#234 := (* -3::real #223)
-#235 := (+ #137 #234)
-#236 := (+ #46 #235)
-#237 := (<= #236 0::real)
-#238 := (not #237)
-#207 := (* -3::real #195)
-#208 := (+ #137 #207)
-#209 := (+ #46 #208)
-#210 := (<= #209 0::real)
-#211 := (not #210)
-#249 := (and #185 #211 #238)
-#65 := (<= #44 #58)
-#56 := (- #46 #55)
-#52 := 3::real
-#59 := (- #58 #55)
-#61 := (- #59)
-#60 := (< #59 0::real)
-#62 := (ite #60 #61 #59)
-#63 := (* #62 3::real)
-#64 := (< #63 #56)
-#66 := (and #64 #65)
-#47 := (- #44 #46)
-#50 := (- #47)
-#49 := (< #47 0::real)
-#51 := (ite #49 #50 #47)
-#53 := (* #51 3::real)
-#57 := (< #53 #56)
-#67 := (and #57 #66)
-#254 := (iff #67 #249)
-#138 := (+ #46 #137)
-#147 := (< #144 0::real)
-#159 := (ite #147 #154 #144)
-#165 := (* 3::real #159)
-#170 := (< #165 #138)
-#176 := (and #65 #170)
-#114 := (< #111 0::real)
-#126 := (ite #114 #121 #111)
-#132 := (* 3::real #126)
-#141 := (< #132 #138)
-#181 := (and #141 #176)
-#252 := (iff #181 #249)
-#243 := (and #185 #211)
-#246 := (and #238 #243)
-#250 := (iff #246 #249)
-#251 := [rewrite]: #250
-#247 := (iff #181 #246)
-#244 := (iff #176 #243)
-#214 := (iff #170 #211)
-#200 := (* 3::real #195)
-#203 := (< #200 #138)
-#212 := (iff #203 #211)
-#213 := [rewrite]: #212
-#204 := (iff #170 #203)
-#201 := (= #165 #200)
-#198 := (= #159 #195)
-#189 := (not #188)
-#192 := (ite #189 #154 #144)
-#196 := (= #192 #195)
-#197 := [rewrite]: #196
-#193 := (= #159 #192)
-#190 := (iff #147 #189)
-#191 := [rewrite]: #190
-#194 := [monotonicity #191]: #193
-#199 := [trans #194 #197]: #198
-#202 := [monotonicity #199]: #201
-#205 := [monotonicity #202]: #204
-#215 := [trans #205 #213]: #214
-#186 := (iff #65 #185)
-#187 := [rewrite]: #186
-#245 := [monotonicity #187 #215]: #244
-#241 := (iff #141 #238)
-#228 := (* 3::real #223)
-#231 := (< #228 #138)
-#239 := (iff #231 #238)
-#240 := [rewrite]: #239
-#232 := (iff #141 #231)
-#229 := (= #132 #228)
-#226 := (= #126 #223)
-#220 := (ite #217 #121 #111)
-#224 := (= #220 #223)
-#225 := [rewrite]: #224
-#221 := (= #126 #220)
-#218 := (iff #114 #217)
-#219 := [rewrite]: #218
-#222 := [monotonicity #219]: #221
-#227 := [trans #222 #225]: #226
-#230 := [monotonicity #227]: #229
-#233 := [monotonicity #230]: #232
-#242 := [trans #233 #240]: #241
-#248 := [monotonicity #242 #245]: #247
-#253 := [trans #248 #251]: #252
-#182 := (iff #67 #181)
-#179 := (iff #66 #176)
-#173 := (and #170 #65)
-#177 := (iff #173 #176)
-#178 := [rewrite]: #177
-#174 := (iff #66 #173)
-#171 := (iff #64 #170)
-#139 := (= #56 #138)
-#140 := [rewrite]: #139
-#168 := (= #63 #165)
-#162 := (* #159 3::real)
-#166 := (= #162 #165)
-#167 := [rewrite]: #166
-#163 := (= #63 #162)
-#160 := (= #62 #159)
-#145 := (= #59 #144)
-#146 := [rewrite]: #145
-#157 := (= #61 #154)
-#150 := (- #144)
-#155 := (= #150 #154)
-#156 := [rewrite]: #155
-#151 := (= #61 #150)
-#152 := [monotonicity #146]: #151
-#158 := [trans #152 #156]: #157
-#148 := (iff #60 #147)
-#149 := [monotonicity #146]: #148
-#161 := [monotonicity #149 #158 #146]: #160
-#164 := [monotonicity #161]: #163
-#169 := [trans #164 #167]: #168
-#172 := [monotonicity #169 #140]: #171
-#175 := [monotonicity #172]: #174
-#180 := [trans #175 #178]: #179
-#142 := (iff #57 #141)
-#135 := (= #53 #132)
-#129 := (* #126 3::real)
-#133 := (= #129 #132)
-#134 := [rewrite]: #133
-#130 := (= #53 #129)
-#127 := (= #51 #126)
-#112 := (= #47 #111)
-#113 := [rewrite]: #112
-#124 := (= #50 #121)
-#117 := (- #111)
-#122 := (= #117 #121)
-#123 := [rewrite]: #122
-#118 := (= #50 #117)
-#119 := [monotonicity #113]: #118
-#125 := [trans #119 #123]: #124
-#115 := (iff #49 #114)
-#116 := [monotonicity #113]: #115
-#128 := [monotonicity #116 #125 #113]: #127
-#131 := [monotonicity #128]: #130
-#136 := [trans #131 #134]: #135
-#143 := [monotonicity #136 #140]: #142
-#183 := [monotonicity #143 #180]: #182
-#255 := [trans #183 #253]: #254
-#108 := [asserted]: #67
-#256 := [mp #108 #255]: #249
-#257 := [and-elim #256]: #185
-#259 := [and-elim #256]: #238
-#448 := (+ #111 #447)
-#449 := (<= #448 0::real)
-#441 := (= #111 #223)
-#443 := (or #217 #441)
-#444 := [def-axiom]: #443
-#467 := [unit-resolution #444 #455]: #441
-#468 := (not #441)
-#469 := (or #468 #449)
-#470 := [th-lemma]: #469
-#471 := [unit-resolution #470 #467]: #449
-#463 := (or #189 #217)
-#456 := [hypothesis]: #188
-#258 := [and-elim #256]: #211
-#437 := (or #189 #435)
-#438 := [def-axiom]: #437
-#457 := [unit-resolution #438 #456]: #435
-#458 := (not #435)
-#459 := (or #458 #453)
-#460 := [th-lemma]: #459
-#461 := [unit-resolution #460 #457]: #453
-#462 := [th-lemma #257 #461 #258 #456 #455]: false
-#464 := [lemma #462]: #463
-#472 := [unit-resolution #464 #455]: #189
-#473 := [th-lemma #472 #471 #259 #257 #455]: false
-#474 := [lemma #473]: #217
-#445 := (or #216 #442)
-#446 := [def-axiom]: #445
-#475 := [unit-resolution #446 #474]: #442
-#476 := (not #442)
-#477 := (or #476 #454)
-#478 := [th-lemma]: #477
-#479 := [unit-resolution #478 #475]: #454
-#481 := (not #185)
-#480 := (not #454)
-#482 := (or #188 #480 #237 #481 #216)
-#483 := [th-lemma]: #482
-#484 := [unit-resolution #483 #257 #474 #259 #479]: #188
-#485 := [unit-resolution #438 #484]: #435
-#486 := [unit-resolution #460 #485]: #453
-[th-lemma #479 #259 #257 #474 #258 #486]: false
-unsat
-e14093e1f049d6c695e00b050a39f30b5ee75c84 41 0
-f1 -> val!6
-f2 -> val!0
-?v0!0 -> val!11
-f12 -> val!1
-f14 -> val!2
-f15 -> val!3
-f9 -> val!7
-f8 -> val!8
-f6 -> val!9
-f13 -> {
-  val!2 val!3 -> val!4
-  else -> val!4
-}
-f11 -> {
-  val!1 val!4 -> val!5
-  else -> val!5
-}
-f10 -> {
-  val!11 val!5 -> val!6
-  else -> val!6
-}
-f7 -> {
-  val!7 val!11 -> 0
-  val!8 val!11 -> 0
-  else -> 0
-}
-f5 -> {
-  val!9 -> val!10
-  else -> val!10
-}
-f4 -> {
-  val!10 -> 0
-  val!11 -> 38
-  else -> 38
-}
-f3 -> {
-  0 -> val!10
-  38 -> val!11
-  else -> val!11
-}
-unknown
-b39ccddf1d4f138215d1b195d383905c9937a82f 44 0
-f1 -> val!7
-f2 -> val!0
-?v0!0 -> val!11
-f6 -> val!1
-f14 -> val!3
-f15 -> val!4
-f9 -> val!8
-f8 -> val!9
-f5 -> {
-  val!1 -> val!10
-  else -> val!10
-}
-f12 -> {
-  val!10 -> val!2
-  else -> val!2
-}
-f13 -> {
-  val!3 val!4 -> val!5
-  else -> val!5
-}
-f11 -> {
-  val!2 val!5 -> val!6
-  else -> val!6
-}
-f10 -> {
-  val!11 val!6 -> val!7
-  else -> val!7
-}
-f7 -> {
-  val!8 val!11 -> 0
-  val!9 val!11 -> 0
-  else -> 0
-}
-f4 -> {
-  val!10 -> 0
-  val!11 -> 38
-  else -> 38
-}
-f3 -> {
-  0 -> val!10
-  38 -> val!11
-  else -> val!11
-}
-unknown
-e03a7d5d1030a490d5669c1c764d61d605302d39 48 0
-f1 -> val!6
-f2 -> val!0
-f5 -> val!9
-f7 -> val!1
-?v0!0 -> val!11
-f15 -> val!3
-f10 -> val!7
-f9 -> val!8
-f4 -> {
-  val!9 -> 0
-  val!10 -> 1
-  val!11 -> 1237
-  else -> 1237
-}
-f6 -> {
-  val!1 -> val!10
-  else -> val!10
-}
-f13 -> {
-  val!10 -> val!2
-  else -> val!2
-}
-f14 -> {
-  val!9 val!3 -> val!4
-  else -> val!4
-}
-f12 -> {
-  val!2 val!4 -> val!5
-  else -> val!5
-}
-f11 -> {
-  val!11 val!5 -> val!6
-  else -> val!6
-}
-f8 -> {
-  val!7 val!11 -> 0
-  val!8 val!11 -> 0
-  val!7 val!9 -> -1
-  val!8 val!9 -> 0
-  else -> 0
-}
-f3 -> {
-  0 -> val!9
-  1 -> val!10
-  1237 -> val!11
-  else -> val!11
-}
-unknown
-e7836464956771492f5454e164bcd43446470f5a 48 0
-f1 -> val!6
-f2 -> val!0
-f5 -> val!9
-f7 -> val!1
-?v0!0 -> val!11
-f15 -> val!3
-f10 -> val!7
-f9 -> val!8
-f4 -> {
-  val!9 -> 0
-  val!10 -> 1
-  val!11 -> 1237
-  else -> 1237
-}
-f6 -> {
-  val!1 -> val!10
-  else -> val!10
-}
-f13 -> {
-  val!10 -> val!2
-  else -> val!2
-}
-f14 -> {
-  val!9 val!3 -> val!4
-  else -> val!4
-}
-f12 -> {
-  val!2 val!4 -> val!5
-  else -> val!5
-}
-f11 -> {
-  val!11 val!5 -> val!6
-  else -> val!6
-}
-f8 -> {
-  val!7 val!11 -> 0
-  val!8 val!11 -> 0
-  val!7 val!9 -> 1
-  val!8 val!9 -> 0
-  else -> 0
-}
-f3 -> {
-  0 -> val!9
-  1 -> val!10
-  1237 -> val!11
-  else -> val!11
-}
-unknown
-94aef9051873492c2e350c3ee8f6b777964eebe4 147 0
-#2 := false
-#159 := 0::real
-decl f8 :: (-> S4 S2 real)
-decl f5 :: S2
-#25 := f5
-decl f10 :: S4
-#33 := f10
-#34 := (f8 f10 f5)
-#156 := -1::real
-#157 := (* -1::real #34)
-decl f9 :: S4
-#31 := f9
-#32 := (f8 f9 f5)
-#158 := (+ #32 #157)
-#315 := (>= #158 0::real)
-#392 := (not #315)
-#202 := (= #32 #34)
-#205 := (not #202)
-#43 := (= #34 #32)
-#44 := (not #43)
-#206 := (iff #44 #205)
-#203 := (iff #43 #202)
-#204 := [rewrite]: #203
-#207 := [monotonicity #204]: #206
-#201 := [asserted]: #44
-#210 := [mp #201 #207]: #205
-#395 := (or #202 #392)
-#160 := (<= #158 0::real)
-#13 := 0::int
-decl f4 :: (-> S2 int)
-decl f6 :: (-> S3 S2)
-decl f7 :: S3
-#27 := f7
-#28 := (f6 f7)
-#29 := (f4 #28)
-#149 := -1::int
-#152 := (* -1::int #29)
-#26 := (f4 f5)
-#153 := (+ #26 #152)
-#151 := (>= #153 0::int)
-#150 := (not #151)
-#163 := (and #150 #160)
-#35 := (<= #32 #34)
-#30 := (< #26 #29)
-#36 := (and #30 #35)
-#164 := (iff #36 #163)
-#161 := (iff #35 #160)
-#162 := [rewrite]: #161
-#154 := (iff #30 #150)
-#155 := [rewrite]: #154
-#165 := [monotonicity #155 #162]: #164
-#146 := [asserted]: #36
-#166 := [mp #146 #165]: #163
-#168 := [and-elim #166]: #160
-#391 := (not #160)
-#393 := (or #202 #391 #392)
-#394 := [th-lemma]: #393
-#396 := [unit-resolution #394 #168]: #395
-#397 := [unit-resolution #396 #210]: #392
-#8 := (:var 0 S2)
-#39 := (f8 f9 #8)
-#308 := (pattern #39)
-#38 := (f8 f10 #8)
-#307 := (pattern #38)
-#9 := (f4 #8)
-#287 := (pattern #9)
-#187 := (* -1::real #39)
-#188 := (+ #38 #187)
-#189 := (<= #188 0::real)
-#177 := (+ #9 #152)
-#176 := (>= #177 0::int)
-#192 := (or #176 #189)
-#309 := (forall (vars (?v0 S2)) (:pat #287 #307 #308) #192)
-#195 := (forall (vars (?v0 S2)) #192)
-#312 := (iff #195 #309)
-#310 := (iff #192 #192)
-#311 := [refl]: #310
-#313 := [quant-intro #311]: #312
-#227 := (~ #195 #195)
-#208 := (~ #192 #192)
-#226 := [refl]: #208
-#228 := [nnf-pos #226]: #227
-#40 := (<= #38 #39)
-#37 := (< #9 #29)
-#41 := (implies #37 #40)
-#42 := (forall (vars (?v0 S2)) #41)
-#198 := (iff #42 #195)
-#148 := (not #37)
-#169 := (or #148 #40)
-#172 := (forall (vars (?v0 S2)) #169)
-#196 := (iff #172 #195)
-#193 := (iff #169 #192)
-#190 := (iff #40 #189)
-#191 := [rewrite]: #190
-#185 := (iff #148 #176)
-#175 := (not #176)
-#180 := (not #175)
-#183 := (iff #180 #176)
-#184 := [rewrite]: #183
-#181 := (iff #148 #180)
-#178 := (iff #37 #175)
-#179 := [rewrite]: #178
-#182 := [monotonicity #179]: #181
-#186 := [trans #182 #184]: #185
-#194 := [monotonicity #186 #191]: #193
-#197 := [quant-intro #194]: #196
-#173 := (iff #42 #172)
-#170 := (iff #41 #169)
-#171 := [rewrite]: #170
-#174 := [quant-intro #171]: #173
-#199 := [trans #174 #197]: #198
-#147 := [asserted]: #42
-#200 := [mp #147 #199]: #195
-#229 := [mp~ #200 #228]: #195
-#314 := [mp #229 #313]: #309
-#167 := [and-elim #166]: #150
-#338 := (not #309)
-#339 := (or #338 #151 #315)
-#318 := (* -1::real #32)
-#319 := (+ #34 #318)
-#323 := (<= #319 0::real)
-#324 := (or #151 #323)
-#340 := (or #338 #324)
-#347 := (iff #340 #339)
-#335 := (or #151 #315)
-#342 := (or #338 #335)
-#345 := (iff #342 #339)
-#346 := [rewrite]: #345
-#343 := (iff #340 #342)
-#336 := (iff #324 #335)
-#333 := (iff #323 #315)
-#325 := (+ #318 #34)
-#328 := (<= #325 0::real)
-#331 := (iff #328 #315)
-#332 := [rewrite]: #331
-#329 := (iff #323 #328)
-#326 := (= #319 #325)
-#327 := [rewrite]: #326
-#330 := [monotonicity #327]: #329
-#334 := [trans #330 #332]: #333
-#337 := [monotonicity #334]: #336
-#344 := [monotonicity #337]: #343
-#348 := [trans #344 #346]: #347
-#341 := [quant-inst]: #340
-#349 := [mp #341 #348]: #339
-[unit-resolution #349 #167 #314 #397]: false
-unsat
-a776c8d80c92e36c87766b3cd2a3d1707a6a4f17 29 0
-f1 -> val!0
-f2 -> val!1
-f6 -> val!2
-f7 -> val!7
-f8 -> val!3
-f12 -> 1
-f11 -> val!4
-f10 -> val!5
-f5 -> {
-  val!2 val!7 -> 0
-  val!3 val!7 -> 0
-  val!4 val!7 -> 1
-  else -> 1
-}
-f9 -> {
-  val!5 -> val!6
-  else -> val!6
-}
-f4 -> {
-  val!6 -> 0
-  val!7 -> 7719
-  else -> 7719
-}
-f3 -> {
-  0 -> val!6
-  7719 -> val!7
-  else -> val!7
-}
-unknown
-2b1630352ec036f9060c61ea0f08be80276b84d2 29 0
-f1 -> val!0
-f2 -> val!1
-f6 -> val!2
-f7 -> val!7
-f8 -> val!3
-f12 -> 1
-f11 -> val!4
-f10 -> val!5
-f5 -> {
-  val!2 val!7 -> 0
-  val!3 val!7 -> 0
-  val!4 val!7 -> 1
-  else -> 1
-}
-f9 -> {
-  val!5 -> val!6
-  else -> val!6
-}
-f4 -> {
-  val!6 -> 0
-  val!7 -> 7719
-  else -> 7719
-}
-f3 -> {
-  0 -> val!6
-  7719 -> val!7
-  else -> val!7
-}
-unknown
-478564f7f355dca34baaecde71bc3afeb25b293b 204 0
-#2 := false
-#46 := 0::real
-decl f5 :: (-> S3 S2 real)
-decl f7 :: S2
-#26 := f7
-decl f11 :: S3
-#38 := f11
-#49 := (f5 f11 f7)
-#183 := -1::real
-#349 := (* -1::real #49)
-decl f8 :: S3
-#28 := f8
-#29 := (f5 f8 f7)
-#362 := (+ #29 #349)
-#363 := (>= #362 0::real)
-#368 := (not #363)
-decl f6 :: S3
-#25 := f6
-#27 := (f5 f6 f7)
-#350 := (+ #27 #349)
-#351 := (<= #350 0::real)
-#352 := (not #351)
-#371 := (or #352 #368)
-#374 := (not #371)
-#8 := (:var 0 S2)
-#41 := (f5 f8 #8)
-#333 := (pattern #41)
-#39 := (f5 f11 #8)
-#332 := (pattern #39)
-#37 := (f5 f6 #8)
-#331 := (pattern #37)
-decl f4 :: (-> S2 int)
-#9 := (f4 #8)
-#311 := (pattern #9)
-#189 := (* -1::real #41)
-#190 := (+ #39 #189)
-#191 := (<= #190 0::real)
-#242 := (not #191)
-#184 := (* -1::real #39)
-#185 := (+ #37 #184)
-#186 := (<= #185 0::real)
-#241 := (not #186)
-#243 := (or #241 #242)
-#244 := (not #243)
-#13 := 0::int
-decl f9 :: (-> S4 S2)
-decl f10 :: S4
-#32 := f10
-#33 := (f9 f10)
-#34 := (f4 #33)
-#157 := -1::int
-#160 := (* -1::int #34)
-#173 := (+ #9 #160)
-#172 := (>= #173 0::int)
-#247 := (or #172 #244)
-#334 := (forall (vars (?v0 S2)) (:pat #311 #331 #332 #333) #247)
-#250 := (forall (vars (?v0 S2)) #247)
-#337 := (iff #250 #334)
-#335 := (iff #247 #247)
-#336 := [refl]: #335
-#338 := [quant-intro #336]: #337
-#194 := (and #186 #191)
-#197 := (or #172 #194)
-#200 := (forall (vars (?v0 S2)) #197)
-#251 := (iff #200 #250)
-#248 := (iff #197 #247)
-#245 := (iff #194 #244)
-#246 := [rewrite]: #245
-#249 := [monotonicity #246]: #248
-#252 := [quant-intro #249]: #251
-#219 := (~ #200 #200)
-#224 := (~ #197 #197)
-#222 := [refl]: #224
-#239 := [nnf-pos #222]: #219
-#42 := (<= #39 #41)
-#40 := (<= #37 #39)
-#43 := (and #40 #42)
-#36 := (< #9 #34)
-#44 := (implies #36 #43)
-#45 := (forall (vars (?v0 S2)) #44)
-#203 := (iff #45 #200)
-#156 := (not #36)
-#165 := (or #156 #43)
-#168 := (forall (vars (?v0 S2)) #165)
-#201 := (iff #168 #200)
-#198 := (iff #165 #197)
-#195 := (iff #43 #194)
-#192 := (iff #42 #191)
-#193 := [rewrite]: #192
-#187 := (iff #40 #186)
-#188 := [rewrite]: #187
-#196 := [monotonicity #188 #193]: #195
-#181 := (iff #156 #172)
-#171 := (not #172)
-#176 := (not #171)
-#179 := (iff #176 #172)
-#180 := [rewrite]: #179
-#177 := (iff #156 #176)
-#174 := (iff #36 #171)
-#175 := [rewrite]: #174
-#178 := [monotonicity #175]: #177
-#182 := [trans #178 #180]: #181
-#199 := [monotonicity #182 #196]: #198
-#202 := [quant-intro #199]: #201
-#169 := (iff #45 #168)
-#166 := (iff #44 #165)
-#167 := [rewrite]: #166
-#170 := [quant-intro #167]: #169
-#204 := [trans #170 #202]: #203
-#155 := [asserted]: #45
-#205 := [mp #155 #204]: #200
-#240 := [mp~ #205 #239]: #200
-#253 := [mp #240 #252]: #250
-#339 := [mp #253 #338]: #334
-#31 := (f4 f7)
-#161 := (+ #31 #160)
-#159 := (>= #161 0::int)
-#158 := (not #159)
-#35 := (< #31 #34)
-#162 := (iff #35 #158)
-#163 := [rewrite]: #162
-#154 := [asserted]: #35
-#164 := [mp #154 #163]: #158
-#380 := (not #334)
-#381 := (or #380 #159 #374)
-#342 := (* -1::real #29)
-#343 := (+ #49 #342)
-#347 := (<= #343 0::real)
-#348 := (not #347)
-#353 := (or #352 #348)
-#354 := (not #353)
-#355 := (or #159 #354)
-#382 := (or #380 #355)
-#389 := (iff #382 #381)
-#377 := (or #159 #374)
-#384 := (or #380 #377)
-#387 := (iff #384 #381)
-#388 := [rewrite]: #387
-#385 := (iff #382 #384)
-#378 := (iff #355 #377)
-#375 := (iff #354 #374)
-#372 := (iff #353 #371)
-#369 := (iff #348 #368)
-#366 := (iff #347 #363)
-#356 := (+ #342 #49)
-#359 := (<= #356 0::real)
-#364 := (iff #359 #363)
-#365 := [rewrite]: #364
-#360 := (iff #347 #359)
-#357 := (= #343 #356)
-#358 := [rewrite]: #357
-#361 := [monotonicity #358]: #360
-#367 := [trans #361 #365]: #366
-#370 := [monotonicity #367]: #369
-#373 := [monotonicity #370]: #372
-#376 := [monotonicity #373]: #375
-#379 := [monotonicity #376]: #378
-#386 := [monotonicity #379]: #385
-#390 := [trans #386 #388]: #389
-#383 := [quant-inst]: #382
-#391 := [mp #383 #390]: #381
-#513 := [unit-resolution #391 #164 #339]: #374
-#394 := (or #371 #363)
-#395 := [def-axiom]: #394
-#514 := [unit-resolution #395 #513]: #363
-#475 := (>= #350 0::real)
-#524 := (not #475)
-#474 := (= #27 #49)
-#519 := (not #474)
-#208 := (= #29 #49)
-#216 := (not #208)
-#520 := (iff #216 #519)
-#517 := (iff #208 #474)
-#515 := (iff #474 #208)
-#30 := (= #27 #29)
-#153 := [asserted]: #30
-#516 := [monotonicity #153]: #515
-#518 := [symm #516]: #517
-#521 := [monotonicity #518]: #520
-#50 := (= #49 #29)
-#51 := (not #50)
-#217 := (iff #51 #216)
-#214 := (iff #50 #208)
-#215 := [rewrite]: #214
-#218 := [monotonicity #215]: #217
-#207 := [asserted]: #51
-#221 := [mp #207 #218]: #216
-#522 := [mp #221 #521]: #519
-#527 := (or #474 #524)
-#392 := (or #371 #351)
-#393 := [def-axiom]: #392
-#523 := [unit-resolution #393 #513]: #351
-#525 := (or #474 #352 #524)
-#526 := [th-lemma]: #525
-#528 := [unit-resolution #526 #523]: #527
-#529 := [unit-resolution #528 #522]: #524
-#471 := (+ #27 #342)
-#473 := (>= #471 0::real)
-#530 := (not #30)
-#531 := (or #530 #473)
-#532 := [th-lemma]: #531
-#533 := [unit-resolution #532 #153]: #473
-[th-lemma #533 #529 #514]: false
-unsat
-07499e3240b5b9ff9168a39bf4f0848339ecb137 29 0
-f5 -> 1
-f12 -> val!0
-f11 -> val!7
-f9 -> val!1
-f1 -> val!2
-f2 -> val!3
-f10 -> val!4
-f7 -> val!5
-f8 -> {
-  val!0 val!7 -> 0
-  val!1 val!7 -> 0
-  val!4 val!7 -> 0
-  else -> 0
-}
-f6 -> {
-  val!5 -> val!6
-  else -> val!6
-}
-f4 -> {
-  val!6 -> 0
-  val!7 -> 38
-  else -> 38
-}
-f3 -> {
-  0 -> val!6
-  38 -> val!7
-  else -> val!7
-}
-unknown
-3521e5c8a799f779cef25ded86c422c58ebfc7fd 318 0
-#2 := false
-#25 := 0::real
-decl f8 :: (-> S4 S2 real)
-decl f11 :: S2
-#40 := f11
-decl f9 :: S4
-#32 := f9
-#47 := (f8 f9 f11)
-decl f12 :: S4
-#44 := f12
-#45 := (f8 f12 f11)
-#192 := -1::real
-#232 := (* -1::real #45)
-#233 := (+ #232 #47)
-decl f5 :: real
-#26 := f5
-#268 := (* -1::real #47)
-#271 := (+ #45 #268)
-#274 := (+ f5 #271)
-#275 := (<= #274 0::real)
-#278 := (ite #275 f5 #233)
-#593 := (* -1::real #278)
-#594 := (+ f5 #593)
-#595 := (<= #594 0::real)
-#603 := (not #595)
-#223 := 1/2::real
-#281 := (* 1/2::real #278)
-#457 := (<= #281 0::real)
-#292 := (= #281 0::real)
-#306 := (<= #271 0::real)
-decl f10 :: S4
-#34 := f10
-#43 := (f8 f10 f11)
-#302 := (+ #43 #232)
-#303 := (<= #302 0::real)
-#309 := (and #303 #306)
-#13 := 0::int
-decl f4 :: (-> S2 int)
-#41 := (f4 f11)
-#178 := -1::int
-#213 := (* -1::int #41)
-decl f6 :: (-> S3 S2)
-decl f7 :: S3
-#28 := f7
-#29 := (f6 f7)
-#30 := (f4 #29)
-#214 := (+ #30 #213)
-#215 := (<= #214 0::int)
-#312 := (or #215 #309)
-#225 := (* 1/2::real #47)
-#272 := (+ #232 #225)
-#224 := (* 1/2::real #43)
-#273 := (+ #224 #272)
-#270 := (>= #273 0::real)
-#321 := (and #270 #292 #312)
-#52 := 2::real
-#55 := (- #47 #45)
-#56 := (<= f5 #55)
-#57 := (ite #56 f5 #55)
-#58 := (/ #57 2::real)
-#59 := (+ #45 #58)
-#60 := (= #59 #45)
-#51 := (+ #43 #47)
-#53 := (/ #51 2::real)
-#54 := (<= #45 #53)
-#61 := (and #54 #60)
-#48 := (<= #45 #47)
-#46 := (<= #43 #45)
-#49 := (and #46 #48)
-#42 := (< #41 #30)
-#50 := (implies #42 #49)
-#62 := (and #50 #61)
-#326 := (iff #62 #321)
-#236 := (<= f5 #233)
-#239 := (ite #236 f5 #233)
-#245 := (* 1/2::real #239)
-#250 := (+ #45 #245)
-#256 := (= #45 #250)
-#226 := (+ #224 #225)
-#229 := (<= #45 #226)
-#261 := (and #229 #256)
-#212 := (not #42)
-#220 := (or #212 #49)
-#264 := (and #220 #261)
-#324 := (iff #264 #321)
-#315 := (and #270 #292)
-#318 := (and #312 #315)
-#322 := (iff #318 #321)
-#323 := [rewrite]: #322
-#319 := (iff #264 #318)
-#316 := (iff #261 #315)
-#293 := (iff #256 #292)
-#284 := (+ #45 #281)
-#287 := (= #45 #284)
-#290 := (iff #287 #292)
-#291 := [rewrite]: #290
-#288 := (iff #256 #287)
-#285 := (= #250 #284)
-#282 := (= #245 #281)
-#279 := (= #239 #278)
-#276 := (iff #236 #275)
-#277 := [rewrite]: #276
-#280 := [monotonicity #277]: #279
-#283 := [monotonicity #280]: #282
-#286 := [monotonicity #283]: #285
-#289 := [monotonicity #286]: #288
-#294 := [trans #289 #291]: #293
-#267 := (iff #229 #270)
-#269 := [rewrite]: #267
-#317 := [monotonicity #269 #294]: #316
-#313 := (iff #220 #312)
-#310 := (iff #49 #309)
-#307 := (iff #48 #306)
-#308 := [rewrite]: #307
-#304 := (iff #46 #303)
-#305 := [rewrite]: #304
-#311 := [monotonicity #305 #308]: #310
-#300 := (iff #212 #215)
-#216 := (not #215)
-#295 := (not #216)
-#298 := (iff #295 #215)
-#299 := [rewrite]: #298
-#296 := (iff #212 #295)
-#217 := (iff #42 #216)
-#218 := [rewrite]: #217
-#297 := [monotonicity #218]: #296
-#301 := [trans #297 #299]: #300
-#314 := [monotonicity #301 #311]: #313
-#320 := [monotonicity #314 #317]: #319
-#325 := [trans #320 #323]: #324
-#265 := (iff #62 #264)
-#262 := (iff #61 #261)
-#259 := (iff #60 #256)
-#253 := (= #250 #45)
-#257 := (iff #253 #256)
-#258 := [rewrite]: #257
-#254 := (iff #60 #253)
-#251 := (= #59 #250)
-#248 := (= #58 #245)
-#242 := (/ #239 2::real)
-#246 := (= #242 #245)
-#247 := [rewrite]: #246
-#243 := (= #58 #242)
-#240 := (= #57 #239)
-#234 := (= #55 #233)
-#235 := [rewrite]: #234
-#237 := (iff #56 #236)
-#238 := [monotonicity #235]: #237
-#241 := [monotonicity #238 #235]: #240
-#244 := [monotonicity #241]: #243
-#249 := [trans #244 #247]: #248
-#252 := [monotonicity #249]: #251
-#255 := [monotonicity #252]: #254
-#260 := [trans #255 #258]: #259
-#230 := (iff #54 #229)
-#227 := (= #53 #226)
-#228 := [rewrite]: #227
-#231 := [monotonicity #228]: #230
-#263 := [monotonicity #231 #260]: #262
-#221 := (iff #50 #220)
-#222 := [rewrite]: #221
-#266 := [monotonicity #222 #263]: #265
-#327 := [trans #266 #325]: #326
-#211 := [asserted]: #62
-#328 := [mp #211 #327]: #321
-#330 := [and-elim #328]: #292
-#597 := (not #292)
-#598 := (or #597 #457)
-#599 := [th-lemma]: #598
-#600 := [unit-resolution #599 #330]: #457
-#601 := [hypothesis]: #595
-#167 := (<= f5 0::real)
-#168 := (not #167)
-#27 := (< 0::real f5)
-#169 := (iff #27 #168)
-#170 := [rewrite]: #169
-#164 := [asserted]: #27
-#171 := [mp #164 #170]: #168
-#602 := [th-lemma #171 #601 #600]: false
-#604 := [lemma #602]: #603
-#450 := (= f5 #278)
-#451 := (= #233 #278)
-#613 := (not #451)
-#596 := (+ #233 #593)
-#605 := (<= #596 0::real)
-#610 := (not #605)
-#532 := (+ #43 #268)
-#533 := (>= #532 0::real)
-#538 := (not #533)
-#210 := [asserted]: #42
-#219 := [mp #210 #218]: #216
-#8 := (:var 0 S2)
-#35 := (f8 f10 #8)
-#443 := (pattern #35)
-#33 := (f8 f9 #8)
-#442 := (pattern #33)
-#9 := (f4 #8)
-#422 := (pattern #9)
-#193 := (* -1::real #35)
-#194 := (+ #33 #193)
-#195 := (<= #194 0::real)
-#198 := (not #195)
-#181 := (* -1::int #30)
-#182 := (+ #9 #181)
-#180 := (>= #182 0::int)
-#201 := (or #180 #198)
-#444 := (forall (vars (?v0 S2)) (:pat #422 #442 #443) #201)
-#204 := (forall (vars (?v0 S2)) #201)
-#447 := (iff #204 #444)
-#445 := (iff #201 #201)
-#446 := [refl]: #445
-#448 := [quant-intro #446]: #447
-#362 := (~ #204 #204)
-#332 := (~ #201 #201)
-#361 := [refl]: #332
-#363 := [nnf-pos #361]: #362
-#36 := (<= #33 #35)
-#37 := (not #36)
-#31 := (< #9 #30)
-#38 := (implies #31 #37)
-#39 := (forall (vars (?v0 S2)) #38)
-#207 := (iff #39 #204)
-#166 := (not #31)
-#172 := (or #166 #37)
-#175 := (forall (vars (?v0 S2)) #172)
-#205 := (iff #175 #204)
-#202 := (iff #172 #201)
-#199 := (iff #37 #198)
-#196 := (iff #36 #195)
-#197 := [rewrite]: #196
-#200 := [monotonicity #197]: #199
-#190 := (iff #166 #180)
-#179 := (not #180)
-#185 := (not #179)
-#188 := (iff #185 #180)
-#189 := [rewrite]: #188
-#186 := (iff #166 #185)
-#183 := (iff #31 #179)
-#184 := [rewrite]: #183
-#187 := [monotonicity #184]: #186
-#191 := [trans #187 #189]: #190
-#203 := [monotonicity #191 #200]: #202
-#206 := [quant-intro #203]: #205
-#176 := (iff #39 #175)
-#173 := (iff #38 #172)
-#174 := [rewrite]: #173
-#177 := [quant-intro #174]: #176
-#208 := [trans #177 #206]: #207
-#165 := [asserted]: #39
-#209 := [mp #165 #208]: #204
-#364 := [mp~ #209 #363]: #204
-#449 := [mp #364 #448]: #444
-#544 := (not #444)
-#545 := (or #544 #215 #538)
-#507 := (* -1::real #43)
-#508 := (+ #47 #507)
-#511 := (<= #508 0::real)
-#512 := (not #511)
-#513 := (+ #41 #181)
-#514 := (>= #513 0::int)
-#515 := (or #514 #512)
-#546 := (or #544 #515)
-#553 := (iff #546 #545)
-#541 := (or #215 #538)
-#548 := (or #544 #541)
-#551 := (iff #548 #545)
-#552 := [rewrite]: #551
-#549 := (iff #546 #548)
-#542 := (iff #515 #541)
-#539 := (iff #512 #538)
-#536 := (iff #511 #533)
-#526 := (+ #507 #47)
-#529 := (<= #526 0::real)
-#534 := (iff #529 #533)
-#535 := [rewrite]: #534
-#530 := (iff #511 #529)
-#527 := (= #508 #526)
-#528 := [rewrite]: #527
-#531 := [monotonicity #528]: #530
-#537 := [trans #531 #535]: #536
-#540 := [monotonicity #537]: #539
-#524 := (iff #514 #215)
-#516 := (+ #181 #41)
-#519 := (>= #516 0::int)
-#522 := (iff #519 #215)
-#523 := [rewrite]: #522
-#520 := (iff #514 #519)
-#517 := (= #513 #516)
-#518 := [rewrite]: #517
-#521 := [monotonicity #518]: #520
-#525 := [trans #521 #523]: #524
-#543 := [monotonicity #525 #540]: #542
-#550 := [monotonicity #543]: #549
-#554 := [trans #550 #552]: #553
-#547 := [quant-inst]: #546
-#555 := [mp #547 #554]: #545
-#607 := [unit-resolution #555 #449 #219]: #538
-#329 := [and-elim #328]: #270
-#608 := [hypothesis]: #605
-#609 := [th-lemma #608 #329 #607 #600]: false
-#611 := [lemma #609]: #610
-#612 := [hypothesis]: #451
-#614 := (or #613 #605)
-#615 := [th-lemma]: #614
-#616 := [unit-resolution #615 #612 #611]: false
-#617 := [lemma #616]: #613
-#455 := (or #275 #451)
-#456 := [def-axiom]: #455
-#618 := [unit-resolution #456 #617]: #275
-#452 := (not #275)
-#453 := (or #452 #450)
-#454 := [def-axiom]: #453
-#619 := [unit-resolution #454 #618]: #450
-#620 := (not #450)
-#621 := (or #620 #595)
-#622 := [th-lemma]: #621
-[unit-resolution #622 #619 #604]: false
-unsat
-f7a332c4ab50576b47f6154d9c204a030b7f3346 295 0
-#2 := false
-#25 := 0::real
-decl f8 :: (-> S4 S2 real)
-decl f11 :: S2
-#40 := f11
-decl f12 :: S4
-#44 := f12
-#45 := (f8 f12 f11)
-decl f10 :: S4
-#34 := f10
-#43 := (f8 f10 f11)
-#199 := -1::real
-#263 := (* -1::real #43)
-#264 := (+ #263 #45)
-decl f5 :: real
-#26 := f5
-#242 := (* -1::real #45)
-#331 := (+ #43 #242)
-#332 := (+ f5 #331)
-#333 := (<= #332 0::real)
-#336 := (ite #333 f5 #264)
-#666 := (* -1::real #336)
-#667 := (+ f5 #666)
-#668 := (<= #667 0::real)
-#676 := (not #668)
-#230 := 1/2::real
-#414 := (* 1/2::real #336)
-#531 := (<= #414 0::real)
-#415 := (= #414 0::real)
-#284 := -1/2::real
-#339 := (* -1/2::real #336)
-#342 := (+ #45 #339)
-decl f9 :: S4
-#32 := f9
-#47 := (f8 f9 f11)
-#243 := (+ #242 #47)
-#316 := (* -1::real #47)
-#317 := (+ #45 #316)
-#318 := (+ f5 #317)
-#319 := (<= #318 0::real)
-#322 := (ite #319 f5 #243)
-#325 := (* 1/2::real #322)
-#328 := (+ #45 #325)
-#232 := (* 1/2::real #47)
-#312 := (+ #242 #232)
-#231 := (* 1/2::real #43)
-#313 := (+ #231 #312)
-#310 := (>= #313 0::real)
-#345 := (ite #310 #328 #342)
-#348 := (= #45 #345)
-#418 := (iff #348 #415)
-#411 := (= #45 #342)
-#416 := (iff #411 #415)
-#417 := [rewrite]: #416
-#412 := (iff #348 #411)
-#409 := (= #345 #342)
-#404 := (ite false #328 #342)
-#407 := (= #404 #342)
-#408 := [rewrite]: #407
-#405 := (= #345 #404)
-#402 := (iff #310 false)
-#309 := (not #310)
-#361 := (<= #317 0::real)
-#358 := (<= #331 0::real)
-#364 := (and #358 #361)
-#13 := 0::int
-decl f4 :: (-> S2 int)
-#41 := (f4 f11)
-#185 := -1::int
-#220 := (* -1::int #41)
-decl f6 :: (-> S3 S2)
-decl f7 :: S3
-#28 := f7
-#29 := (f6 f7)
-#30 := (f4 #29)
-#221 := (+ #30 #220)
-#222 := (<= #221 0::int)
-#367 := (or #222 #364)
-#376 := (and #309 #348 #367)
-#52 := 2::real
-#61 := (- #45 #43)
-#62 := (<= f5 #61)
-#63 := (ite #62 f5 #61)
-#64 := (/ #63 2::real)
-#65 := (- #45 #64)
-#56 := (- #47 #45)
-#57 := (<= f5 #56)
-#58 := (ite #57 f5 #56)
-#59 := (/ #58 2::real)
-#60 := (+ #45 #59)
-#51 := (+ #43 #47)
-#53 := (/ #51 2::real)
-#55 := (<= #45 #53)
-#66 := (ite #55 #60 #65)
-#67 := (= #66 #45)
-#54 := (< #53 #45)
-#68 := (and #54 #67)
-#48 := (<= #45 #47)
-#46 := (<= #43 #45)
-#49 := (and #46 #48)
-#42 := (< #41 #30)
-#50 := (implies #42 #49)
-#69 := (and #50 #68)
-#381 := (iff #69 #376)
-#267 := (<= f5 #264)
-#270 := (ite #267 f5 #264)
-#285 := (* -1/2::real #270)
-#286 := (+ #45 #285)
-#246 := (<= f5 #243)
-#249 := (ite #246 f5 #243)
-#255 := (* 1/2::real #249)
-#260 := (+ #45 #255)
-#233 := (+ #231 #232)
-#239 := (<= #45 #233)
-#291 := (ite #239 #260 #286)
-#297 := (= #45 #291)
-#236 := (< #233 #45)
-#302 := (and #236 #297)
-#219 := (not #42)
-#227 := (or #219 #49)
-#305 := (and #227 #302)
-#379 := (iff #305 #376)
-#370 := (and #309 #348)
-#373 := (and #367 #370)
-#377 := (iff #373 #376)
-#378 := [rewrite]: #377
-#374 := (iff #305 #373)
-#371 := (iff #302 #370)
-#349 := (iff #297 #348)
-#346 := (= #291 #345)
-#343 := (= #286 #342)
-#340 := (= #285 #339)
-#337 := (= #270 #336)
-#334 := (iff #267 #333)
-#335 := [rewrite]: #334
-#338 := [monotonicity #335]: #337
-#341 := [monotonicity #338]: #340
-#344 := [monotonicity #341]: #343
-#329 := (= #260 #328)
-#326 := (= #255 #325)
-#323 := (= #249 #322)
-#320 := (iff #246 #319)
-#321 := [rewrite]: #320
-#324 := [monotonicity #321]: #323
-#327 := [monotonicity #324]: #326
-#330 := [monotonicity #327]: #329
-#315 := (iff #239 #310)
-#314 := [rewrite]: #315
-#347 := [monotonicity #314 #330 #344]: #346
-#350 := [monotonicity #347]: #349
-#308 := (iff #236 #309)
-#311 := [rewrite]: #308
-#372 := [monotonicity #311 #350]: #371
-#368 := (iff #227 #367)
-#365 := (iff #49 #364)
-#362 := (iff #48 #361)
-#363 := [rewrite]: #362
-#359 := (iff #46 #358)
-#360 := [rewrite]: #359
-#366 := [monotonicity #360 #363]: #365
-#356 := (iff #219 #222)
-#223 := (not #222)
-#351 := (not #223)
-#354 := (iff #351 #222)
-#355 := [rewrite]: #354
-#352 := (iff #219 #351)
-#224 := (iff #42 #223)
-#225 := [rewrite]: #224
-#353 := [monotonicity #225]: #352
-#357 := [trans #353 #355]: #356
-#369 := [monotonicity #357 #366]: #368
-#375 := [monotonicity #369 #372]: #374
-#380 := [trans #375 #378]: #379
-#306 := (iff #69 #305)
-#303 := (iff #68 #302)
-#300 := (iff #67 #297)
-#294 := (= #291 #45)
-#298 := (iff #294 #297)
-#299 := [rewrite]: #298
-#295 := (iff #67 #294)
-#292 := (= #66 #291)
-#289 := (= #65 #286)
-#276 := (* 1/2::real #270)
-#281 := (- #45 #276)
-#287 := (= #281 #286)
-#288 := [rewrite]: #287
-#282 := (= #65 #281)
-#279 := (= #64 #276)
-#273 := (/ #270 2::real)
-#277 := (= #273 #276)
-#278 := [rewrite]: #277
-#274 := (= #64 #273)
-#271 := (= #63 #270)
-#265 := (= #61 #264)
-#266 := [rewrite]: #265
-#268 := (iff #62 #267)
-#269 := [monotonicity #266]: #268
-#272 := [monotonicity #269 #266]: #271
-#275 := [monotonicity #272]: #274
-#280 := [trans #275 #278]: #279
-#283 := [monotonicity #280]: #282
-#290 := [trans #283 #288]: #289
-#261 := (= #60 #260)
-#258 := (= #59 #255)
-#252 := (/ #249 2::real)
-#256 := (= #252 #255)
-#257 := [rewrite]: #256
-#253 := (= #59 #252)
-#250 := (= #58 #249)
-#244 := (= #56 #243)
-#245 := [rewrite]: #244
-#247 := (iff #57 #246)
-#248 := [monotonicity #245]: #247
-#251 := [monotonicity #248 #245]: #250
-#254 := [monotonicity #251]: #253
-#259 := [trans #254 #257]: #258
-#262 := [monotonicity #259]: #261
-#240 := (iff #55 #239)
-#234 := (= #53 #233)
-#235 := [rewrite]: #234
-#241 := [monotonicity #235]: #240
-#293 := [monotonicity #241 #262 #290]: #292
-#296 := [monotonicity #293]: #295
-#301 := [trans #296 #299]: #300
-#237 := (iff #54 #236)
-#238 := [monotonicity #235]: #237
-#304 := [monotonicity #238 #301]: #303
-#228 := (iff #50 #227)
-#229 := [rewrite]: #228
-#307 := [monotonicity #229 #304]: #306
-#382 := [trans #307 #380]: #381
-#218 := [asserted]: #69
-#383 := [mp #218 #382]: #376
-#384 := [and-elim #383]: #309
-#403 := [iff-false #384]: #402
-#406 := [monotonicity #403]: #405
-#410 := [trans #406 #408]: #409
-#413 := [monotonicity #410]: #412
-#419 := [trans #413 #417]: #418
-#385 := [and-elim #383]: #348
-#420 := [mp #385 #419]: #415
-#670 := (not #415)
-#671 := (or #670 #531)
-#672 := [th-lemma]: #671
-#673 := [unit-resolution #672 #420]: #531
-#674 := [hypothesis]: #668
-#174 := (<= f5 0::real)
-#175 := (not #174)
-#27 := (< 0::real f5)
-#176 := (iff #27 #175)
-#177 := [rewrite]: #176
-#171 := [asserted]: #27
-#178 := [mp #171 #177]: #175
-#675 := [th-lemma #178 #674 #673]: false
-#677 := [lemma #675]: #676
-#524 := (= f5 #336)
-#525 := (= #264 #336)
-#685 := (not #525)
-#669 := (+ #264 #666)
-#678 := (<= #669 0::real)
-#682 := (not #678)
-#428 := (iff #367 #364)
-#423 := (or false #364)
-#426 := (iff #423 #364)
-#427 := [rewrite]: #426
-#424 := (iff #367 #423)
-#400 := (iff #222 false)
-#217 := [asserted]: #42
-#226 := [mp #217 #225]: #223
-#401 := [iff-false #226]: #400
-#425 := [monotonicity #401]: #424
-#429 := [trans #425 #427]: #428
-#386 := [and-elim #383]: #367
-#430 := [mp #386 #429]: #364
-#422 := [and-elim #430]: #361
-#680 := [hypothesis]: #678
-#681 := [th-lemma #680 #422 #384 #673]: false
-#683 := [lemma #681]: #682
-#684 := [hypothesis]: #525
-#686 := (or #685 #678)
-#687 := [th-lemma]: #686
-#688 := [unit-resolution #687 #684 #683]: false
-#689 := [lemma #688]: #685
-#529 := (or #333 #525)
-#530 := [def-axiom]: #529
-#690 := [unit-resolution #530 #689]: #333
-#526 := (not #333)
-#527 := (or #526 #524)
-#528 := [def-axiom]: #527
-#691 := [unit-resolution #528 #690]: #524
-#692 := (not #524)
-#693 := (or #692 #668)
-#694 := [th-lemma]: #693
-[unit-resolution #694 #691 #677]: false
-unsat
-5c9bfd1097eaec646745ccbc19c98d1ba847f0f5 308 0
-#2 := false
-#10 := 0::real
-decl f6 :: real
-#15 := f6
-#143 := (<= f6 0::real)
-#144 := (not #143)
-decl f5 :: real
-#13 := f5
-#53 := -1::real
-#203 := (* -1::real f5)
-decl f4 :: real
-#8 := f4
-#204 := (+ f4 #203)
-#202 := (>= #204 0::real)
-decl f3 :: (-> real real)
-#9 := (f3 f4)
-#54 := (* -1::real #9)
-#154 := (>= #9 0::real)
-#161 := (ite #154 #9 #54)
-#166 := (* f6 #161)
-#187 := (* f5 #166)
-#193 := (* -1::real #187)
-#66 := (* #9 f6)
-#78 := (* f4 #66)
-#72 := (* f5 f6)
-#73 := (* #9 #72)
-#96 := (* -1::real #73)
-#97 := (+ #96 #78)
-#84 := (* -1::real #78)
-#85 := (+ #73 #84)
-#172 := (>= #85 0::real)
-#179 := (ite #172 #85 #97)
-#194 := (+ #179 #193)
-#169 := (* f4 #166)
-#195 := (+ #169 #194)
-#196 := (<= #195 0::real)
-#11 := (= #9 0::real)
-#223 := (or #11 #143 #196 #202)
-#228 := (not #223)
-#18 := (- #9)
-#17 := (< #9 0::real)
-#19 := (ite #17 #18 #9)
-#20 := (* f6 #19)
-#30 := (* f5 #20)
-#22 := (* f6 #9)
-#24 := (* f4 #22)
-#23 := (* f5 #22)
-#25 := (- #23 #24)
-#27 := (- #25)
-#26 := (< #25 0::real)
-#28 := (ite #26 #27 #25)
-#21 := (* f4 #20)
-#29 := (+ #21 #28)
-#31 := (<= #29 #30)
-#16 := (< 0::real f6)
-#32 := (implies #16 #31)
-#14 := (< f4 f5)
-#33 := (implies #14 #32)
-#12 := (not #11)
-#34 := (implies #12 #33)
-#35 := (not #34)
-#231 := (iff #35 #228)
-#57 := (ite #17 #54 #9)
-#60 := (* f6 #57)
-#108 := (* f5 #60)
-#90 := (< #85 0::real)
-#102 := (ite #90 #97 #85)
-#63 := (* f4 #60)
-#105 := (+ #63 #102)
-#111 := (<= #105 #108)
-#117 := (not #16)
-#118 := (or #117 #111)
-#126 := (not #14)
-#127 := (or #126 #118)
-#135 := (or #11 #127)
-#140 := (not #135)
-#229 := (iff #140 #228)
-#226 := (iff #135 #223)
-#214 := (or #143 #196)
-#217 := (or #202 #214)
-#220 := (or #11 #217)
-#224 := (iff #220 #223)
-#225 := [rewrite]: #224
-#221 := (iff #135 #220)
-#218 := (iff #127 #217)
-#215 := (iff #118 #214)
-#199 := (iff #111 #196)
-#184 := (+ #169 #179)
-#190 := (<= #184 #187)
-#197 := (iff #190 #196)
-#198 := [rewrite]: #197
-#191 := (iff #111 #190)
-#188 := (= #108 #187)
-#167 := (= #60 #166)
-#164 := (= #57 #161)
-#155 := (not #154)
-#158 := (ite #155 #54 #9)
-#162 := (= #158 #161)
-#163 := [rewrite]: #162
-#159 := (= #57 #158)
-#156 := (iff #17 #155)
-#157 := [rewrite]: #156
-#160 := [monotonicity #157]: #159
-#165 := [trans #160 #163]: #164
-#168 := [monotonicity #165]: #167
-#189 := [monotonicity #168]: #188
-#185 := (= #105 #184)
-#182 := (= #102 #179)
-#173 := (not #172)
-#176 := (ite #173 #97 #85)
-#180 := (= #176 #179)
-#181 := [rewrite]: #180
-#177 := (= #102 #176)
-#174 := (iff #90 #173)
-#175 := [rewrite]: #174
-#178 := [monotonicity #175]: #177
-#183 := [trans #178 #181]: #182
-#170 := (= #63 #169)
-#171 := [monotonicity #168]: #170
-#186 := [monotonicity #171 #183]: #185
-#192 := [monotonicity #186 #189]: #191
-#200 := [trans #192 #198]: #199
-#152 := (iff #117 #143)
-#147 := (not #144)
-#150 := (iff #147 #143)
-#151 := [rewrite]: #150
-#148 := (iff #117 #147)
-#145 := (iff #16 #144)
-#146 := [rewrite]: #145
-#149 := [monotonicity #146]: #148
-#153 := [trans #149 #151]: #152
-#216 := [monotonicity #153 #200]: #215
-#212 := (iff #126 #202)
-#201 := (not #202)
-#207 := (not #201)
-#210 := (iff #207 #202)
-#211 := [rewrite]: #210
-#208 := (iff #126 #207)
-#205 := (iff #14 #201)
-#206 := [rewrite]: #205
-#209 := [monotonicity #206]: #208
-#213 := [trans #209 #211]: #212
-#219 := [monotonicity #213 #216]: #218
-#222 := [monotonicity #219]: #221
-#227 := [trans #222 #225]: #226
-#230 := [monotonicity #227]: #229
-#141 := (iff #35 #140)
-#138 := (iff #34 #135)
-#132 := (implies #12 #127)
-#136 := (iff #132 #135)
-#137 := [rewrite]: #136
-#133 := (iff #34 #132)
-#130 := (iff #33 #127)
-#123 := (implies #14 #118)
-#128 := (iff #123 #127)
-#129 := [rewrite]: #128
-#124 := (iff #33 #123)
-#121 := (iff #32 #118)
-#114 := (implies #16 #111)
-#119 := (iff #114 #118)
-#120 := [rewrite]: #119
-#115 := (iff #32 #114)
-#112 := (iff #31 #111)
-#109 := (= #30 #108)
-#61 := (= #20 #60)
-#58 := (= #19 #57)
-#55 := (= #18 #54)
-#56 := [rewrite]: #55
-#59 := [monotonicity #56]: #58
-#62 := [monotonicity #59]: #61
-#110 := [monotonicity #62]: #109
-#106 := (= #29 #105)
-#103 := (= #28 #102)
-#88 := (= #25 #85)
-#81 := (- #73 #78)
-#86 := (= #81 #85)
-#87 := [rewrite]: #86
-#82 := (= #25 #81)
-#79 := (= #24 #78)
-#67 := (= #22 #66)
-#68 := [rewrite]: #67
-#80 := [monotonicity #68]: #79
-#76 := (= #23 #73)
-#69 := (* f5 #66)
-#74 := (= #69 #73)
-#75 := [rewrite]: #74
-#70 := (= #23 #69)
-#71 := [monotonicity #68]: #70
-#77 := [trans #71 #75]: #76
-#83 := [monotonicity #77 #80]: #82
-#89 := [trans #83 #87]: #88
-#100 := (= #27 #97)
-#93 := (- #85)
-#98 := (= #93 #97)
-#99 := [rewrite]: #98
-#94 := (= #27 #93)
-#95 := [monotonicity #89]: #94
-#101 := [trans #95 #99]: #100
-#91 := (iff #26 #90)
-#92 := [monotonicity #89]: #91
-#104 := [monotonicity #92 #101 #89]: #103
-#64 := (= #21 #63)
-#65 := [monotonicity #62]: #64
-#107 := [monotonicity #65 #104]: #106
-#113 := [monotonicity #107 #110]: #112
-#116 := [monotonicity #113]: #115
-#122 := [trans #116 #120]: #121
-#125 := [monotonicity #122]: #124
-#131 := [trans #125 #129]: #130
-#134 := [monotonicity #131]: #133
-#139 := [trans #134 #137]: #138
-#142 := [monotonicity #139]: #141
-#232 := [trans #142 #230]: #231
-#52 := [asserted]: #35
-#233 := [mp #52 #232]: #228
-#235 := [not-or-elim #233]: #144
-#241 := (<= #9 0::real)
-#317 := (not #241)
-#251 := (= #54 #161)
-#290 := (not #251)
-#236 := (not #196)
-#237 := [not-or-elim #233]: #236
-#239 := (* -1::real #179)
-#299 := (+ #97 #239)
-#301 := (>= #299 0::real)
-#247 := (= #97 #179)
-#246 := (= #85 #179)
-#280 := (not #246)
-#250 := (= #9 #161)
-#264 := (not #250)
-#254 := (+ #85 #239)
-#256 := (>= #254 0::real)
-#279 := [hypothesis]: #246
-#281 := (or #280 #256)
-#282 := [th-lemma]: #281
-#283 := [unit-resolution #282 #279]: #256
-#255 := (<= #254 0::real)
-#284 := (or #280 #255)
-#285 := [th-lemma]: #284
-#286 := [unit-resolution #285 #279]: #255
-#273 := (not #256)
-#272 := (not #255)
-#274 := (or #264 #272 #273)
-#261 := [hypothesis]: #256
-#262 := [hypothesis]: #255
-#257 := (* -1::real #161)
-#258 := (+ #9 #257)
-#260 := (>= #258 0::real)
-#263 := [hypothesis]: #250
-#265 := (or #264 #260)
-#266 := [th-lemma]: #265
-#267 := [unit-resolution #266 #263]: #260
-#259 := (<= #258 0::real)
-#268 := (or #264 #259)
-#269 := [th-lemma]: #268
-#270 := [unit-resolution #269 #263]: #259
-#271 := [th-lemma #270 #267 #262 #261 #237]: false
-#275 := [lemma #271]: #274
-#287 := [unit-resolution #275 #286 #283]: #264
-#252 := (or #155 #250)
-#253 := [def-axiom]: #252
-#288 := [unit-resolution #253 #287]: #155
-#238 := [not-or-elim #233]: #201
-#276 := (+ #54 #257)
-#278 := (>= #276 0::real)
-#248 := (or #154 #251)
-#249 := [def-axiom]: #248
-#289 := [unit-resolution #249 #288]: #251
-#291 := (or #290 #278)
-#292 := [th-lemma]: #291
-#293 := [unit-resolution #292 #289]: #278
-#277 := (<= #276 0::real)
-#294 := (or #290 #277)
-#295 := [th-lemma]: #294
-#296 := [unit-resolution #295 #289]: #277
-#297 := [th-lemma #296 #293 #286 #283 #237 #238 #288 #235]: false
-#298 := [lemma #297]: #280
-#244 := (or #173 #246)
-#245 := [def-axiom]: #244
-#302 := [unit-resolution #245 #298]: #173
-#242 := (or #172 #247)
-#243 := [def-axiom]: #242
-#303 := [unit-resolution #243 #302]: #247
-#304 := (not #247)
-#305 := (or #304 #301)
-#306 := [th-lemma]: #305
-#307 := [unit-resolution #306 #303]: #301
-#300 := (<= #299 0::real)
-#308 := (or #304 #300)
-#309 := [th-lemma]: #308
-#310 := [unit-resolution #309 #303]: #300
-#311 := [hypothesis]: #251
-#312 := [unit-resolution #292 #311]: #278
-#313 := [unit-resolution #295 #311]: #277
-#314 := [th-lemma #313 #312 #310 #307 #237]: false
-#315 := [lemma #314]: #290
-#316 := [unit-resolution #249 #315]: #154
-#320 := (or #317 #155)
-#234 := [not-or-elim #233]: #12
-#318 := (or #11 #317 #155)
-#319 := [th-lemma]: #318
-#321 := [unit-resolution #319 #234]: #320
-#322 := [unit-resolution #321 #316]: #317
-#323 := [unit-resolution #253 #316]: #250
-#324 := [unit-resolution #266 #323]: #260
-#325 := [unit-resolution #269 #323]: #259
-[th-lemma #325 #324 #310 #307 #237 #238 #322 #235]: false
-unsat
-9171919cd704be8e7c103a2053e7cd7353c1a910 19 0
-f1 -> val!0
-f2 -> val!1
-f3 -> 0
-f4 -> 1
-f5 -> 1
-f6 -> {
-  0 -> 3
-  else -> 3
-}
-f7 -> {
-  3 -> 1
-  8 -> 1
-  else -> 1
-}
-f8 -> {
-  1236 3 -> 8
-  else -> 8
-}
-unknown
-2890cc0ea13f6a8b550bc60d1dd56915e31e1ba5 386 0
-#2 := false
-#11 := 0::real
-decl f5 :: real
-#12 := f5
-#62 := (<= f5 0::real)
-#63 := (not #62)
-#13 := (< 0::real f5)
-#64 := (iff #13 #63)
-#65 := [rewrite]: #64
-#51 := [asserted]: #13
-#66 := [mp #51 #65]: #63
-decl f4 :: real
-#9 := f4
-#53 := -1::real
-#56 := (* -1::real f4)
-decl f3 :: real
-#8 := f3
-#57 := (+ f3 #56)
-#55 := (>= #57 0::real)
-#54 := (not #55)
-#10 := (< f3 f4)
-#58 := (iff #10 #54)
-#59 := [rewrite]: #58
-#50 := [asserted]: #10
-#60 := [mp #50 #59]: #54
-decl f6 :: (-> real real)
-#14 := (f6 f3)
-#99 := (* -1::real #14)
-#152 := (>= #14 0::real)
-#159 := (ite #152 #14 #99)
-#271 := (<= #159 0::real)
-#336 := (not #271)
-#252 := (<= #14 0::real)
-#397 := (not #252)
-#153 := (not #152)
-#384 := [hypothesis]: #153
-#267 := (* -1::real #159)
-#305 := (+ #99 #267)
-#306 := (<= #305 0::real)
-#241 := (= #99 #159)
-#244 := (or #152 #241)
-#245 := [def-axiom]: #244
-#385 := [unit-resolution #245 #384]: #241
-#314 := (not #241)
-#318 := (or #314 #306)
-#319 := [th-lemma]: #318
-#386 := [unit-resolution #319 #385]: #306
-#337 := (not #306)
-#338 := (or #336 #152 #337)
-#339 := [th-lemma]: #338
-#387 := [unit-resolution #339 #386 #384]: #336
-#81 := (* f4 f5)
-#90 := 1/8::real
-#93 := (* 1/8::real #81)
-#79 := (* f3 f5)
-#91 := -1/8::real
-#92 := (* -1/8::real #79)
-#94 := (+ #92 #93)
-#164 := (/ #94 #159)
-#170 := (<= #164 0::real)
-#176 := (* #14 #164)
-#188 := (* -1::real #176)
-#183 := (>= #176 0::real)
-#194 := (ite #183 #176 #188)
-#203 := (* -1/8::real #81)
-#204 := (+ #203 #194)
-#202 := (* 1/8::real #79)
-#205 := (+ #202 #204)
-#206 := (<= #205 0::real)
-#382 := (or #206 #314)
-#237 := (not #206)
-#273 := [hypothesis]: #237
-#263 := (* -1::real #194)
-#348 := (+ #188 #263)
-#350 := (>= #348 0::real)
-#247 := (= #188 #194)
-#182 := (not #183)
-#246 := (= #176 #194)
-#325 := (not #246)
-#346 := (or #325 #206)
-#226 := (= #14 #159)
-#293 := (not #226)
-#264 := (+ #176 #263)
-#266 := (>= #264 0::real)
-#324 := [hypothesis]: #246
-#326 := (or #325 #266)
-#327 := [th-lemma]: #326
-#328 := [unit-resolution #327 #324]: #266
-#265 := (<= #264 0::real)
-#329 := (or #325 #265)
-#330 := [th-lemma]: #329
-#331 := [unit-resolution #330 #324]: #265
-#302 := (not #266)
-#301 := (not #265)
-#303 := (or #293 #301 #302 #206)
-#254 := (* #159 #164)
-#258 := (+ #203 #254)
-#259 := (+ #202 #258)
-#257 := (>= #259 0::real)
-#260 := (= #259 0::real)
-#253 := (= #159 0::real)
-#277 := (not #253)
-#15 := (= #14 0::real)
-#16 := (not #15)
-#278 := (iff #16 #277)
-#275 := (iff #15 #253)
-#274 := [hypothesis]: #226
-#276 := [monotonicity #274]: #275
-#279 := [monotonicity #276]: #278
-#171 := (not #170)
-#211 := (and #171 #206)
-#214 := (or #15 #211)
-#217 := (not #214)
-#19 := 8::real
-#17 := (- f4 f3)
-#18 := (* f5 #17)
-#20 := (/ #18 8::real)
-#22 := (- #14)
-#21 := (< #14 0::real)
-#23 := (ite #21 #22 #14)
-#24 := (/ #20 #23)
-#26 := (* #24 #14)
-#28 := (- #26)
-#27 := (< #26 0::real)
-#29 := (ite #27 #28 #26)
-#30 := (<= #29 #20)
-#25 := (< 0::real #24)
-#31 := (and #25 #30)
-#32 := (implies #16 #31)
-#33 := (not #32)
-#220 := (iff #33 #217)
-#102 := (ite #21 #99 #14)
-#105 := (/ #94 #102)
-#114 := (* #14 #105)
-#125 := (* -1::real #114)
-#119 := (< #114 0::real)
-#130 := (ite #119 #125 #114)
-#133 := (<= #130 #94)
-#108 := (< 0::real #105)
-#136 := (and #108 #133)
-#61 := (= 0::real #14)
-#142 := (or #61 #136)
-#147 := (not #142)
-#218 := (iff #147 #217)
-#215 := (iff #142 #214)
-#212 := (iff #136 #211)
-#209 := (iff #133 #206)
-#199 := (<= #194 #94)
-#207 := (iff #199 #206)
-#208 := [rewrite]: #207
-#200 := (iff #133 #199)
-#197 := (= #130 #194)
-#191 := (ite #182 #188 #176)
-#195 := (= #191 #194)
-#196 := [rewrite]: #195
-#192 := (= #130 #191)
-#177 := (= #114 #176)
-#165 := (= #105 #164)
-#162 := (= #102 #159)
-#156 := (ite #153 #99 #14)
-#160 := (= #156 #159)
-#161 := [rewrite]: #160
-#157 := (= #102 #156)
-#154 := (iff #21 #153)
-#155 := [rewrite]: #154
-#158 := [monotonicity #155]: #157
-#163 := [trans #158 #161]: #162
-#166 := [monotonicity #163]: #165
-#178 := [monotonicity #166]: #177
-#189 := (= #125 #188)
-#190 := [monotonicity #178]: #189
-#186 := (iff #119 #182)
-#179 := (< #176 0::real)
-#184 := (iff #179 #182)
-#185 := [rewrite]: #184
-#180 := (iff #119 #179)
-#181 := [monotonicity #178]: #180
-#187 := [trans #181 #185]: #186
-#193 := [monotonicity #187 #190 #178]: #192
-#198 := [trans #193 #196]: #197
-#201 := [monotonicity #198]: #200
-#210 := [trans #201 #208]: #209
-#174 := (iff #108 #171)
-#167 := (< 0::real #164)
-#172 := (iff #167 #171)
-#173 := [rewrite]: #172
-#168 := (iff #108 #167)
-#169 := [monotonicity #166]: #168
-#175 := [trans #169 #173]: #174
-#213 := [monotonicity #175 #210]: #212
-#150 := (iff #61 #15)
-#151 := [rewrite]: #150
-#216 := [monotonicity #151 #213]: #215
-#219 := [monotonicity #216]: #218
-#148 := (iff #33 #147)
-#145 := (iff #32 #142)
-#69 := (not #61)
-#139 := (implies #69 #136)
-#143 := (iff #139 #142)
-#144 := [rewrite]: #143
-#140 := (iff #32 #139)
-#137 := (iff #31 #136)
-#134 := (iff #30 #133)
-#97 := (= #20 #94)
-#80 := (* -1::real #79)
-#82 := (+ #80 #81)
-#87 := (/ #82 8::real)
-#95 := (= #87 #94)
-#96 := [rewrite]: #95
-#88 := (= #20 #87)
-#85 := (= #18 #82)
-#72 := (* -1::real f3)
-#73 := (+ #72 f4)
-#76 := (* f5 #73)
-#83 := (= #76 #82)
-#84 := [rewrite]: #83
-#77 := (= #18 #76)
-#74 := (= #17 #73)
-#75 := [rewrite]: #74
-#78 := [monotonicity #75]: #77
-#86 := [trans #78 #84]: #85
-#89 := [monotonicity #86]: #88
-#98 := [trans #89 #96]: #97
-#131 := (= #29 #130)
-#117 := (= #26 #114)
-#111 := (* #105 #14)
-#115 := (= #111 #114)
-#116 := [rewrite]: #115
-#112 := (= #26 #111)
-#106 := (= #24 #105)
-#103 := (= #23 #102)
-#100 := (= #22 #99)
-#101 := [rewrite]: #100
-#104 := [monotonicity #101]: #103
-#107 := [monotonicity #98 #104]: #106
-#113 := [monotonicity #107]: #112
-#118 := [trans #113 #116]: #117
-#128 := (= #28 #125)
-#122 := (- #114)
-#126 := (= #122 #125)
-#127 := [rewrite]: #126
-#123 := (= #28 #122)
-#124 := [monotonicity #118]: #123
-#129 := [trans #124 #127]: #128
-#120 := (iff #27 #119)
-#121 := [monotonicity #118]: #120
-#132 := [monotonicity #121 #129 #118]: #131
-#135 := [monotonicity #132 #98]: #134
-#109 := (iff #25 #108)
-#110 := [monotonicity #107]: #109
-#138 := [monotonicity #110 #135]: #137
-#70 := (iff #16 #69)
-#67 := (iff #15 #61)
-#68 := [rewrite]: #67
-#71 := [monotonicity #68]: #70
-#141 := [monotonicity #71 #138]: #140
-#146 := [trans #141 #144]: #145
-#149 := [monotonicity #146]: #148
-#221 := [trans #149 #219]: #220
-#52 := [asserted]: #33
-#222 := [mp #52 #221]: #217
-#223 := [not-or-elim #222]: #16
-#280 := [mp #223 #279]: #277
-#281 := (or #253 #260)
-#282 := [th-lemma]: #281
-#283 := [unit-resolution #282 #280]: #260
-#284 := (not #260)
-#285 := (or #284 #257)
-#286 := [th-lemma]: #285
-#287 := [unit-resolution #286 #283]: #257
-#256 := (<= #259 0::real)
-#288 := (or #284 #256)
-#289 := [th-lemma]: #288
-#290 := [unit-resolution #289 #283]: #256
-#291 := [hypothesis]: #266
-#292 := [hypothesis]: #265
-#268 := (+ #14 #267)
-#270 := (>= #268 0::real)
-#294 := (or #293 #270)
-#295 := [th-lemma]: #294
-#296 := [unit-resolution #295 #274]: #270
-#269 := (<= #268 0::real)
-#297 := (or #293 #269)
-#298 := [th-lemma]: #297
-#299 := [unit-resolution #298 #274]: #269
-#300 := [th-lemma #299 #296 #292 #291 #290 #287 #273]: false
-#304 := [lemma #300]: #303
-#332 := [unit-resolution #304 #331 #328 #273]: #293
-#242 := (or #153 #226)
-#243 := [def-axiom]: #242
-#333 := [unit-resolution #243 #332]: #153
-#334 := [unit-resolution #245 #333]: #241
-#335 := [unit-resolution #319 #334]: #306
-#340 := [unit-resolution #339 #333 #335]: #336
-#322 := (or #284 #301 #302 #206)
-#308 := [hypothesis]: #260
-#309 := [unit-resolution #286 #308]: #257
-#310 := [unit-resolution #289 #308]: #256
-#307 := (>= #305 0::real)
-#311 := [unit-resolution #304 #292 #291 #273]: #293
-#312 := [unit-resolution #243 #311]: #153
-#313 := [unit-resolution #245 #312]: #241
-#315 := (or #314 #307)
-#316 := [th-lemma]: #315
-#317 := [unit-resolution #316 #313]: #307
-#320 := [unit-resolution #319 #313]: #306
-#321 := [th-lemma #320 #317 #292 #291 #310 #309 #273 #60 #66]: false
-#323 := [lemma #321]: #322
-#341 := [unit-resolution #323 #331 #328 #273]: #284
-#342 := [unit-resolution #282 #341]: #253
-#343 := (or #277 #271)
-#344 := [th-lemma]: #343
-#345 := [unit-resolution #344 #342 #340]: false
-#347 := [lemma #345]: #346
-#365 := [unit-resolution #347 #273]: #325
-#248 := (or #182 #246)
-#249 := [def-axiom]: #248
-#366 := [unit-resolution #249 #365]: #182
-#250 := (or #183 #247)
-#251 := [def-axiom]: #250
-#367 := [unit-resolution #251 #366]: #247
-#368 := (not #247)
-#369 := (or #368 #350)
-#370 := [th-lemma]: #369
-#371 := [unit-resolution #370 #367]: #350
-#349 := (<= #348 0::real)
-#372 := (or #368 #349)
-#373 := [th-lemma]: #372
-#374 := [unit-resolution #373 #367]: #349
-#351 := [hypothesis]: #253
-#352 := [unit-resolution #344 #351]: #271
-#357 := (iff #16 #293)
-#355 := (iff #15 #226)
-#353 := (iff #226 #15)
-#354 := [monotonicity #351]: #353
-#356 := [symm #354]: #355
-#358 := [monotonicity #356]: #357
-#359 := [mp #223 #358]: #293
-#360 := [unit-resolution #243 #359]: #153
-#361 := [unit-resolution #339 #360 #352]: #337
-#362 := [unit-resolution #245 #360]: #241
-#363 := [unit-resolution #319 #362 #361]: false
-#364 := [lemma #363]: #277
-#375 := [unit-resolution #282 #364]: #260
-#376 := [unit-resolution #286 #375]: #257
-#377 := [unit-resolution #289 #375]: #256
-#378 := [hypothesis]: #241
-#379 := [unit-resolution #316 #378]: #307
-#380 := [unit-resolution #319 #378]: #306
-#381 := [th-lemma #380 #379 #377 #376 #374 #371 #273]: false
-#383 := [lemma #381]: #382
-#388 := [unit-resolution #383 #385]: #206
-#238 := (or #170 #237)
-#224 := (not #211)
-#229 := (iff #224 #238)
-#239 := (not #238)
-#236 := (not #239)
-#231 := (iff #236 #238)
-#232 := [rewrite]: #231
-#233 := (iff #224 #236)
-#240 := (iff #211 #239)
-#235 := [rewrite]: #240
-#234 := [monotonicity #235]: #233
-#230 := [trans #234 #232]: #229
-#225 := [not-or-elim #222]: #224
-#228 := [mp #225 #230]: #238
-#389 := [unit-resolution #228 #388]: #170
-#390 := [th-lemma #377 #376 #389 #387 #60 #66]: false
-#391 := [lemma #390]: #152
-#400 := (or #397 #153)
-#398 := (or #15 #397 #153)
-#399 := [th-lemma]: #398
-#401 := [unit-resolution #399 #223]: #400
-#402 := [unit-resolution #401 #391]: #397
-#392 := [unit-resolution #243 #391]: #226
-#394 := [unit-resolution #298 #392]: #269
-#403 := (not #269)
-#404 := (or #336 #252 #403)
-#405 := [th-lemma]: #404
-#406 := [unit-resolution #405 #394 #402]: #336
-#393 := [unit-resolution #295 #392]: #270
-#395 := [th-lemma #394 #393 #377 #376 #374 #371 #273 #60 #66]: false
-#396 := [lemma #395]: #206
-#407 := [unit-resolution #228 #396]: #170
-[th-lemma #377 #376 #407 #406 #60 #66]: false
-unsat
-726e0e73e0888d06cfe039319b4774bfabee92ef 421 0
-#2 := false
-#10 := 0::real
-decl f6 :: real
-#15 := f6
-#150 := (<= f6 0::real)
-#151 := (not #150)
-decl f5 :: real
-#13 := f5
-#53 := -1::real
-#223 := (* -1::real f5)
-decl f4 :: real
-#8 := f4
-#224 := (+ f4 #223)
-#225 := (>= #224 0::real)
-decl f3 :: (-> real real)
-#9 := (f3 f4)
-#81 := (* -1::real #9)
-#161 := (>= #9 0::real)
-#168 := (ite #161 #9 #81)
-#63 := (* f5 f6)
-#72 := 1/8::real
-#75 := (* 1/8::real #63)
-#61 := (* f4 f6)
-#73 := -1/8::real
-#74 := (* -1/8::real #61)
-#76 := (+ #74 #75)
-#173 := (/ #76 #168)
-#185 := (* #9 #173)
-#197 := (* -1::real #185)
-#192 := (>= #185 0::real)
-#203 := (ite #192 #185 #197)
-#212 := (* -1/8::real #63)
-#213 := (+ #212 #203)
-#211 := (* 1/8::real #61)
-#214 := (+ #211 #213)
-#215 := (<= #214 0::real)
-#179 := (<= #173 0::real)
-#180 := (not #179)
-#220 := (and #180 #215)
-#11 := (= #9 0::real)
-#245 := (or #11 #150 #220 #225)
-#250 := (not #245)
-#19 := 8::real
-#17 := (- f5 f4)
-#18 := (* f6 #17)
-#20 := (/ #18 8::real)
-#22 := (- #9)
-#21 := (< #9 0::real)
-#23 := (ite #21 #22 #9)
-#24 := (/ #20 #23)
-#26 := (* #24 #9)
-#28 := (- #26)
-#27 := (< #26 0::real)
-#29 := (ite #27 #28 #26)
-#30 := (<= #29 #20)
-#25 := (< 0::real #24)
-#31 := (and #25 #30)
-#16 := (< 0::real f6)
-#32 := (implies #16 #31)
-#14 := (< f4 f5)
-#33 := (implies #14 #32)
-#12 := (not #11)
-#34 := (implies #12 #33)
-#35 := (not #34)
-#253 := (iff #35 #250)
-#84 := (ite #21 #81 #9)
-#87 := (/ #76 #84)
-#96 := (* #9 #87)
-#107 := (* -1::real #96)
-#101 := (< #96 0::real)
-#112 := (ite #101 #107 #96)
-#115 := (<= #112 #76)
-#90 := (< 0::real #87)
-#118 := (and #90 #115)
-#124 := (not #16)
-#125 := (or #124 #118)
-#133 := (not #14)
-#134 := (or #133 #125)
-#142 := (or #11 #134)
-#147 := (not #142)
-#251 := (iff #147 #250)
-#248 := (iff #142 #245)
-#236 := (or #150 #220)
-#239 := (or #225 #236)
-#242 := (or #11 #239)
-#246 := (iff #242 #245)
-#247 := [rewrite]: #246
-#243 := (iff #142 #242)
-#240 := (iff #134 #239)
-#237 := (iff #125 #236)
-#221 := (iff #118 #220)
-#218 := (iff #115 #215)
-#208 := (<= #203 #76)
-#216 := (iff #208 #215)
-#217 := [rewrite]: #216
-#209 := (iff #115 #208)
-#206 := (= #112 #203)
-#191 := (not #192)
-#200 := (ite #191 #197 #185)
-#204 := (= #200 #203)
-#205 := [rewrite]: #204
-#201 := (= #112 #200)
-#186 := (= #96 #185)
-#174 := (= #87 #173)
-#171 := (= #84 #168)
-#162 := (not #161)
-#165 := (ite #162 #81 #9)
-#169 := (= #165 #168)
-#170 := [rewrite]: #169
-#166 := (= #84 #165)
-#163 := (iff #21 #162)
-#164 := [rewrite]: #163
-#167 := [monotonicity #164]: #166
-#172 := [trans #167 #170]: #171
-#175 := [monotonicity #172]: #174
-#187 := [monotonicity #175]: #186
-#198 := (= #107 #197)
-#199 := [monotonicity #187]: #198
-#195 := (iff #101 #191)
-#188 := (< #185 0::real)
-#193 := (iff #188 #191)
-#194 := [rewrite]: #193
-#189 := (iff #101 #188)
-#190 := [monotonicity #187]: #189
-#196 := [trans #190 #194]: #195
-#202 := [monotonicity #196 #199 #187]: #201
-#207 := [trans #202 #205]: #206
-#210 := [monotonicity #207]: #209
-#219 := [trans #210 #217]: #218
-#183 := (iff #90 #180)
-#176 := (< 0::real #173)
-#181 := (iff #176 #180)
-#182 := [rewrite]: #181
-#177 := (iff #90 #176)
-#178 := [monotonicity #175]: #177
-#184 := [trans #178 #182]: #183
-#222 := [monotonicity #184 #219]: #221
-#159 := (iff #124 #150)
-#154 := (not #151)
-#157 := (iff #154 #150)
-#158 := [rewrite]: #157
-#155 := (iff #124 #154)
-#152 := (iff #16 #151)
-#153 := [rewrite]: #152
-#156 := [monotonicity #153]: #155
-#160 := [trans #156 #158]: #159
-#238 := [monotonicity #160 #222]: #237
-#234 := (iff #133 #225)
-#226 := (not #225)
-#229 := (not #226)
-#232 := (iff #229 #225)
-#233 := [rewrite]: #232
-#230 := (iff #133 #229)
-#227 := (iff #14 #226)
-#228 := [rewrite]: #227
-#231 := [monotonicity #228]: #230
-#235 := [trans #231 #233]: #234
-#241 := [monotonicity #235 #238]: #240
-#244 := [monotonicity #241]: #243
-#249 := [trans #244 #247]: #248
-#252 := [monotonicity #249]: #251
-#148 := (iff #35 #147)
-#145 := (iff #34 #142)
-#139 := (implies #12 #134)
-#143 := (iff #139 #142)
-#144 := [rewrite]: #143
-#140 := (iff #34 #139)
-#137 := (iff #33 #134)
-#130 := (implies #14 #125)
-#135 := (iff #130 #134)
-#136 := [rewrite]: #135
-#131 := (iff #33 #130)
-#128 := (iff #32 #125)
-#121 := (implies #16 #118)
-#126 := (iff #121 #125)
-#127 := [rewrite]: #126
-#122 := (iff #32 #121)
-#119 := (iff #31 #118)
-#116 := (iff #30 #115)
-#79 := (= #20 #76)
-#62 := (* -1::real #61)
-#64 := (+ #62 #63)
-#69 := (/ #64 8::real)
-#77 := (= #69 #76)
-#78 := [rewrite]: #77
-#70 := (= #20 #69)
-#67 := (= #18 #64)
-#54 := (* -1::real f4)
-#55 := (+ #54 f5)
-#58 := (* f6 #55)
-#65 := (= #58 #64)
-#66 := [rewrite]: #65
-#59 := (= #18 #58)
-#56 := (= #17 #55)
-#57 := [rewrite]: #56
-#60 := [monotonicity #57]: #59
-#68 := [trans #60 #66]: #67
-#71 := [monotonicity #68]: #70
-#80 := [trans #71 #78]: #79
-#113 := (= #29 #112)
-#99 := (= #26 #96)
-#93 := (* #87 #9)
-#97 := (= #93 #96)
-#98 := [rewrite]: #97
-#94 := (= #26 #93)
-#88 := (= #24 #87)
-#85 := (= #23 #84)
-#82 := (= #22 #81)
-#83 := [rewrite]: #82
-#86 := [monotonicity #83]: #85
-#89 := [monotonicity #80 #86]: #88
-#95 := [monotonicity #89]: #94
-#100 := [trans #95 #98]: #99
-#110 := (= #28 #107)
-#104 := (- #96)
-#108 := (= #104 #107)
-#109 := [rewrite]: #108
-#105 := (= #28 #104)
-#106 := [monotonicity #100]: #105
-#111 := [trans #106 #109]: #110
-#102 := (iff #27 #101)
-#103 := [monotonicity #100]: #102
-#114 := [monotonicity #103 #111 #100]: #113
-#117 := [monotonicity #114 #80]: #116
-#91 := (iff #25 #90)
-#92 := [monotonicity #89]: #91
-#120 := [monotonicity #92 #117]: #119
-#123 := [monotonicity #120]: #122
-#129 := [trans #123 #127]: #128
-#132 := [monotonicity #129]: #131
-#138 := [trans #132 #136]: #137
-#141 := [monotonicity #138]: #140
-#146 := [trans #141 #144]: #145
-#149 := [monotonicity #146]: #148
-#254 := [trans #149 #252]: #253
-#52 := [asserted]: #35
-#255 := [mp #52 #254]: #250
-#257 := [not-or-elim #255]: #151
-#260 := [not-or-elim #255]: #226
-#306 := (<= #168 0::real)
-#371 := (not #306)
-#287 := (<= #9 0::real)
-#432 := (not #287)
-#419 := [hypothesis]: #162
-#302 := (* -1::real #168)
-#340 := (+ #81 #302)
-#341 := (<= #340 0::real)
-#276 := (= #81 #168)
-#279 := (or #161 #276)
-#280 := [def-axiom]: #279
-#420 := [unit-resolution #280 #419]: #276
-#349 := (not #276)
-#353 := (or #349 #341)
-#354 := [th-lemma]: #353
-#421 := [unit-resolution #354 #420]: #341
-#372 := (not #341)
-#373 := (or #371 #161 #372)
-#374 := [th-lemma]: #373
-#422 := [unit-resolution #374 #421 #419]: #371
-#417 := (or #215 #349)
-#272 := (not #215)
-#308 := [hypothesis]: #272
-#298 := (* -1::real #203)
-#383 := (+ #197 #298)
-#385 := (>= #383 0::real)
-#282 := (= #197 #203)
-#281 := (= #185 #203)
-#360 := (not #281)
-#381 := (or #360 #215)
-#261 := (= #9 #168)
-#328 := (not #261)
-#299 := (+ #185 #298)
-#301 := (>= #299 0::real)
-#359 := [hypothesis]: #281
-#361 := (or #360 #301)
-#362 := [th-lemma]: #361
-#363 := [unit-resolution #362 #359]: #301
-#300 := (<= #299 0::real)
-#364 := (or #360 #300)
-#365 := [th-lemma]: #364
-#366 := [unit-resolution #365 #359]: #300
-#337 := (not #301)
-#336 := (not #300)
-#338 := (or #328 #336 #337 #215)
-#289 := (* #168 #173)
-#293 := (+ #212 #289)
-#294 := (+ #211 #293)
-#292 := (>= #294 0::real)
-#295 := (= #294 0::real)
-#288 := (= #168 0::real)
-#312 := (not #288)
-#313 := (iff #12 #312)
-#310 := (iff #11 #288)
-#309 := [hypothesis]: #261
-#311 := [monotonicity #309]: #310
-#314 := [monotonicity #311]: #313
-#256 := [not-or-elim #255]: #12
-#315 := [mp #256 #314]: #312
-#316 := (or #288 #295)
-#317 := [th-lemma]: #316
-#318 := [unit-resolution #317 #315]: #295
-#319 := (not #295)
-#320 := (or #319 #292)
-#321 := [th-lemma]: #320
-#322 := [unit-resolution #321 #318]: #292
-#291 := (<= #294 0::real)
-#323 := (or #319 #291)
-#324 := [th-lemma]: #323
-#325 := [unit-resolution #324 #318]: #291
-#326 := [hypothesis]: #301
-#327 := [hypothesis]: #300
-#303 := (+ #9 #302)
-#305 := (>= #303 0::real)
-#329 := (or #328 #305)
-#330 := [th-lemma]: #329
-#331 := [unit-resolution #330 #309]: #305
-#304 := (<= #303 0::real)
-#332 := (or #328 #304)
-#333 := [th-lemma]: #332
-#334 := [unit-resolution #333 #309]: #304
-#335 := [th-lemma #334 #331 #327 #326 #325 #322 #308]: false
-#339 := [lemma #335]: #338
-#367 := [unit-resolution #339 #366 #363 #308]: #328
-#277 := (or #162 #261)
-#278 := [def-axiom]: #277
-#368 := [unit-resolution #278 #367]: #162
-#369 := [unit-resolution #280 #368]: #276
-#370 := [unit-resolution #354 #369]: #341
-#375 := [unit-resolution #374 #368 #370]: #371
-#357 := (or #319 #336 #337 #215)
-#343 := [hypothesis]: #295
-#344 := [unit-resolution #321 #343]: #292
-#345 := [unit-resolution #324 #343]: #291
-#342 := (>= #340 0::real)
-#346 := [unit-resolution #339 #327 #326 #308]: #328
-#347 := [unit-resolution #278 #346]: #162
-#348 := [unit-resolution #280 #347]: #276
-#350 := (or #349 #342)
-#351 := [th-lemma]: #350
-#352 := [unit-resolution #351 #348]: #342
-#355 := [unit-resolution #354 #348]: #341
-#356 := [th-lemma #355 #352 #327 #326 #345 #344 #308 #260 #257]: false
-#358 := [lemma #356]: #357
-#376 := [unit-resolution #358 #366 #363 #308]: #319
-#377 := [unit-resolution #317 #376]: #288
-#378 := (or #312 #306)
-#379 := [th-lemma]: #378
-#380 := [unit-resolution #379 #377 #375]: false
-#382 := [lemma #380]: #381
-#400 := [unit-resolution #382 #308]: #360
-#283 := (or #191 #281)
-#284 := [def-axiom]: #283
-#401 := [unit-resolution #284 #400]: #191
-#285 := (or #192 #282)
-#286 := [def-axiom]: #285
-#402 := [unit-resolution #286 #401]: #282
-#403 := (not #282)
-#404 := (or #403 #385)
-#405 := [th-lemma]: #404
-#406 := [unit-resolution #405 #402]: #385
-#384 := (<= #383 0::real)
-#407 := (or #403 #384)
-#408 := [th-lemma]: #407
-#409 := [unit-resolution #408 #402]: #384
-#386 := [hypothesis]: #288
-#387 := [unit-resolution #379 #386]: #306
-#392 := (iff #12 #328)
-#390 := (iff #11 #261)
-#388 := (iff #261 #11)
-#389 := [monotonicity #386]: #388
-#391 := [symm #389]: #390
-#393 := [monotonicity #391]: #392
-#394 := [mp #256 #393]: #328
-#395 := [unit-resolution #278 #394]: #162
-#396 := [unit-resolution #374 #395 #387]: #372
-#397 := [unit-resolution #280 #395]: #276
-#398 := [unit-resolution #354 #397 #396]: false
-#399 := [lemma #398]: #312
-#410 := [unit-resolution #317 #399]: #295
-#411 := [unit-resolution #321 #410]: #292
-#412 := [unit-resolution #324 #410]: #291
-#413 := [hypothesis]: #276
-#414 := [unit-resolution #351 #413]: #342
-#415 := [unit-resolution #354 #413]: #341
-#416 := [th-lemma #415 #414 #412 #411 #409 #406 #308]: false
-#418 := [lemma #416]: #417
-#423 := [unit-resolution #418 #420]: #215
-#273 := (or #179 #272)
-#258 := (not #220)
-#264 := (iff #258 #273)
-#274 := (not #273)
-#271 := (not #274)
-#266 := (iff #271 #273)
-#267 := [rewrite]: #266
-#268 := (iff #258 #271)
-#275 := (iff #220 #274)
-#270 := [rewrite]: #275
-#269 := [monotonicity #270]: #268
-#265 := [trans #269 #267]: #264
-#259 := [not-or-elim #255]: #258
-#263 := [mp #259 #265]: #273
-#424 := [unit-resolution #263 #423]: #179
-#425 := [th-lemma #412 #411 #424 #422 #260 #257]: false
-#426 := [lemma #425]: #161
-#435 := (or #432 #162)
-#433 := (or #11 #432 #162)
-#434 := [th-lemma]: #433
-#436 := [unit-resolution #434 #256]: #435
-#437 := [unit-resolution #436 #426]: #432
-#427 := [unit-resolution #278 #426]: #261
-#429 := [unit-resolution #333 #427]: #304
-#438 := (not #304)
-#439 := (or #371 #287 #438)
-#440 := [th-lemma]: #439
-#441 := [unit-resolution #440 #429 #437]: #371
-#428 := [unit-resolution #330 #427]: #305
-#430 := [th-lemma #429 #428 #412 #411 #409 #406 #308 #260 #257]: false
-#431 := [lemma #430]: #215
-#442 := [unit-resolution #263 #431]: #179
-[th-lemma #412 #411 #442 #441 #260 #257]: false
-unsat
-bae4db9bec568394a074211dd1bdf1addbc19bd3 401 0
-#2 := false
-#16 := 0::real
-decl f3 :: real
-#8 := f3
-#165 := (<= f3 0::real)
-#166 := (not #165)
-decl f5 :: real
-#10 := f5
-#54 := -1::real
-#55 := (* -1::real f5)
-decl f4 :: real
-#9 := f4
-#56 := (+ f4 #55)
-#237 := (<= #56 0::real)
-decl f6 :: (-> real real)
-#15 := (f6 f5)
-#82 := (* -1::real #15)
-#176 := (>= #15 0::real)
-#183 := (ite #176 #15 #82)
-#63 := (* f3 f5)
-#75 := -1/8::real
-#76 := (* -1/8::real #63)
-#62 := (* f3 f4)
-#73 := 1/8::real
-#74 := (* 1/8::real #62)
-#77 := (+ #74 #76)
-#188 := (/ #77 #183)
-#228 := (<= #188 0::real)
-#229 := (not #228)
-#191 := (* #15 #188)
-#203 := (* -1::real #191)
-#198 := (>= #191 0::real)
-#209 := (ite #198 #191 #203)
-#221 := (* -1::real #209)
-#222 := (+ #76 #221)
-#223 := (+ #74 #222)
-#219 := (>= #223 0::real)
-#234 := (and #219 #229)
-#248 := (not #219)
-#26 := (= #15 0::real)
-#263 := (or #26 #165 #248 #234 #237)
-#268 := (not #263)
-#13 := 8::real
-#11 := (- f4 f5)
-#12 := (* f3 #11)
-#14 := (/ #12 8::real)
-#18 := (- #15)
-#17 := (< #15 0::real)
-#19 := (ite #17 #18 #15)
-#20 := (/ #14 #19)
-#21 := (* #20 #15)
-#23 := (- #21)
-#22 := (< #21 0::real)
-#24 := (ite #22 #23 #21)
-#25 := (<= #24 #14)
-#30 := (< 0::real #20)
-#31 := (and #30 #25)
-#29 := (< 0::real f3)
-#32 := (implies #29 #31)
-#28 := (< f5 f4)
-#33 := (implies #28 #32)
-#27 := (not #26)
-#34 := (implies #27 #33)
-#35 := (implies #25 #34)
-#36 := (not #35)
-#271 := (iff #36 #268)
-#85 := (ite #17 #82 #15)
-#88 := (/ #77 #85)
-#116 := (< 0::real #88)
-#94 := (* #15 #88)
-#105 := (* -1::real #94)
-#99 := (< #94 0::real)
-#110 := (ite #99 #105 #94)
-#113 := (<= #110 #77)
-#122 := (and #113 #116)
-#130 := (not #29)
-#131 := (or #130 #122)
-#139 := (not #28)
-#140 := (or #139 #131)
-#148 := (or #26 #140)
-#156 := (not #113)
-#157 := (or #156 #148)
-#162 := (not #157)
-#269 := (iff #162 #268)
-#266 := (iff #157 #263)
-#251 := (or #165 #234)
-#254 := (or #237 #251)
-#257 := (or #26 #254)
-#260 := (or #248 #257)
-#264 := (iff #260 #263)
-#265 := [rewrite]: #264
-#261 := (iff #157 #260)
-#258 := (iff #148 #257)
-#255 := (iff #140 #254)
-#252 := (iff #131 #251)
-#235 := (iff #122 #234)
-#232 := (iff #116 #229)
-#225 := (< 0::real #188)
-#230 := (iff #225 #229)
-#231 := [rewrite]: #230
-#226 := (iff #116 #225)
-#189 := (= #88 #188)
-#186 := (= #85 #183)
-#177 := (not #176)
-#180 := (ite #177 #82 #15)
-#184 := (= #180 #183)
-#185 := [rewrite]: #184
-#181 := (= #85 #180)
-#178 := (iff #17 #177)
-#179 := [rewrite]: #178
-#182 := [monotonicity #179]: #181
-#187 := [trans #182 #185]: #186
-#190 := [monotonicity #187]: #189
-#227 := [monotonicity #190]: #226
-#233 := [trans #227 #231]: #232
-#220 := (iff #113 #219)
-#214 := (<= #209 #77)
-#218 := (iff #214 #219)
-#217 := [rewrite]: #218
-#215 := (iff #113 #214)
-#212 := (= #110 #209)
-#197 := (not #198)
-#206 := (ite #197 #203 #191)
-#210 := (= #206 #209)
-#211 := [rewrite]: #210
-#207 := (= #110 #206)
-#192 := (= #94 #191)
-#193 := [monotonicity #190]: #192
-#204 := (= #105 #203)
-#205 := [monotonicity #193]: #204
-#201 := (iff #99 #197)
-#194 := (< #191 0::real)
-#199 := (iff #194 #197)
-#200 := [rewrite]: #199
-#195 := (iff #99 #194)
-#196 := [monotonicity #193]: #195
-#202 := [trans #196 #200]: #201
-#208 := [monotonicity #202 #205 #193]: #207
-#213 := [trans #208 #211]: #212
-#216 := [monotonicity #213]: #215
-#224 := [trans #216 #217]: #220
-#236 := [monotonicity #224 #233]: #235
-#174 := (iff #130 #165)
-#169 := (not #166)
-#172 := (iff #169 #165)
-#173 := [rewrite]: #172
-#170 := (iff #130 #169)
-#167 := (iff #29 #166)
-#168 := [rewrite]: #167
-#171 := [monotonicity #168]: #170
-#175 := [trans #171 #173]: #174
-#253 := [monotonicity #175 #236]: #252
-#246 := (iff #139 #237)
-#238 := (not #237)
-#241 := (not #238)
-#244 := (iff #241 #237)
-#245 := [rewrite]: #244
-#242 := (iff #139 #241)
-#239 := (iff #28 #238)
-#240 := [rewrite]: #239
-#243 := [monotonicity #240]: #242
-#247 := [trans #243 #245]: #246
-#256 := [monotonicity #247 #253]: #255
-#259 := [monotonicity #256]: #258
-#249 := (iff #156 #248)
-#250 := [monotonicity #224]: #249
-#262 := [monotonicity #250 #259]: #261
-#267 := [trans #262 #265]: #266
-#270 := [monotonicity #267]: #269
-#163 := (iff #36 #162)
-#160 := (iff #35 #157)
-#153 := (implies #113 #148)
-#158 := (iff #153 #157)
-#159 := [rewrite]: #158
-#154 := (iff #35 #153)
-#151 := (iff #34 #148)
-#145 := (implies #27 #140)
-#149 := (iff #145 #148)
-#150 := [rewrite]: #149
-#146 := (iff #34 #145)
-#143 := (iff #33 #140)
-#136 := (implies #28 #131)
-#141 := (iff #136 #140)
-#142 := [rewrite]: #141
-#137 := (iff #33 #136)
-#134 := (iff #32 #131)
-#127 := (implies #29 #122)
-#132 := (iff #127 #131)
-#133 := [rewrite]: #132
-#128 := (iff #32 #127)
-#125 := (iff #31 #122)
-#119 := (and #116 #113)
-#123 := (iff #119 #122)
-#124 := [rewrite]: #123
-#120 := (iff #31 #119)
-#114 := (iff #25 #113)
-#80 := (= #14 #77)
-#64 := (* -1::real #63)
-#65 := (+ #62 #64)
-#70 := (/ #65 8::real)
-#78 := (= #70 #77)
-#79 := [rewrite]: #78
-#71 := (= #14 #70)
-#68 := (= #12 #65)
-#59 := (* f3 #56)
-#66 := (= #59 #65)
-#67 := [rewrite]: #66
-#60 := (= #12 #59)
-#57 := (= #11 #56)
-#58 := [rewrite]: #57
-#61 := [monotonicity #58]: #60
-#69 := [trans #61 #67]: #68
-#72 := [monotonicity #69]: #71
-#81 := [trans #72 #79]: #80
-#111 := (= #24 #110)
-#97 := (= #21 #94)
-#91 := (* #88 #15)
-#95 := (= #91 #94)
-#96 := [rewrite]: #95
-#92 := (= #21 #91)
-#89 := (= #20 #88)
-#86 := (= #19 #85)
-#83 := (= #18 #82)
-#84 := [rewrite]: #83
-#87 := [monotonicity #84]: #86
-#90 := [monotonicity #81 #87]: #89
-#93 := [monotonicity #90]: #92
-#98 := [trans #93 #96]: #97
-#108 := (= #23 #105)
-#102 := (- #94)
-#106 := (= #102 #105)
-#107 := [rewrite]: #106
-#103 := (= #23 #102)
-#104 := [monotonicity #98]: #103
-#109 := [trans #104 #107]: #108
-#100 := (iff #22 #99)
-#101 := [monotonicity #98]: #100
-#112 := [monotonicity #101 #109 #98]: #111
-#115 := [monotonicity #112 #81]: #114
-#117 := (iff #30 #116)
-#118 := [monotonicity #90]: #117
-#121 := [monotonicity #118 #115]: #120
-#126 := [trans #121 #124]: #125
-#129 := [monotonicity #126]: #128
-#135 := [trans #129 #133]: #134
-#138 := [monotonicity #135]: #137
-#144 := [trans #138 #142]: #143
-#147 := [monotonicity #144]: #146
-#152 := [trans #147 #150]: #151
-#155 := [monotonicity #115 #152]: #154
-#161 := [trans #155 #159]: #160
-#164 := [monotonicity #161]: #163
-#272 := [trans #164 #270]: #271
-#53 := [asserted]: #36
-#273 := [mp #53 #272]: #268
-#275 := [not-or-elim #273]: #166
-#279 := [not-or-elim #273]: #238
-#322 := (<= #183 0::real)
-#370 := (not #322)
-#318 := (<= #15 0::real)
-#404 := (not #318)
-#276 := [not-or-elim #273]: #219
-#352 := (+ #191 #221)
-#353 := (<= #352 0::real)
-#306 := (= #191 #209)
-#386 := [hypothesis]: #177
-#389 := (or #198 #176)
-#277 := (not #234)
-#301 := (iff #277 #228)
-#296 := (not #229)
-#299 := (iff #296 #228)
-#300 := [rewrite]: #299
-#297 := (iff #277 #296)
-#294 := (iff #234 #229)
-#1 := true
-#289 := (and true #229)
-#292 := (iff #289 #229)
-#293 := [rewrite]: #292
-#290 := (iff #234 #289)
-#287 := (iff #219 true)
-#288 := [iff-true #276]: #287
-#291 := [monotonicity #288]: #290
-#295 := [trans #291 #293]: #294
-#298 := [monotonicity #295]: #297
-#302 := [trans #298 #300]: #301
-#278 := [not-or-elim #273]: #277
-#303 := [mp #278 #302]: #228
-#387 := [hypothesis]: #197
-#388 := [th-lemma #387 #303 #386]: false
-#390 := [lemma #388]: #389
-#376 := [unit-resolution #390 #386]: #198
-#282 := (or #197 #306)
-#280 := [def-axiom]: #282
-#377 := [unit-resolution #280 #376]: #306
-#366 := (not #306)
-#367 := (or #366 #353)
-#368 := [th-lemma]: #367
-#391 := [unit-resolution #368 #377]: #353
-#324 := (* -1::real #183)
-#325 := (+ #82 #324)
-#326 := (<= #325 0::real)
-#305 := (= #82 #183)
-#283 := (or #176 #305)
-#284 := [def-axiom]: #283
-#392 := [unit-resolution #284 #386]: #305
-#342 := (not #305)
-#343 := (or #342 #326)
-#344 := [th-lemma]: #343
-#393 := [unit-resolution #344 #392]: #326
-#394 := (not #326)
-#395 := (or #370 #176 #394)
-#396 := [th-lemma]: #395
-#397 := [unit-resolution #396 #386 #393]: #370
-#311 := (* #183 #188)
-#319 := (* -1::real #311)
-#320 := (+ #76 #319)
-#321 := (+ #74 #320)
-#314 := (>= #321 0::real)
-#317 := (= #321 0::real)
-#310 := (= #183 0::real)
-#346 := (not #310)
-#304 := (= #15 #183)
-#336 := (not #304)
-#337 := (iff #27 #336)
-#334 := (iff #26 #304)
-#332 := (iff #304 #26)
-#331 := [hypothesis]: #310
-#333 := [monotonicity #331]: #332
-#335 := [symm #333]: #334
-#338 := [monotonicity #335]: #337
-#274 := [not-or-elim #273]: #27
-#339 := [mp #274 #338]: #336
-#285 := (or #177 #304)
-#286 := [def-axiom]: #285
-#340 := [unit-resolution #286 #339]: #177
-#341 := [unit-resolution #284 #340]: #305
-#345 := [unit-resolution #344 #341]: #326
-#347 := (or #346 #322)
-#348 := [th-lemma]: #347
-#349 := [unit-resolution #348 #331]: #322
-#350 := [th-lemma #340 #349 #345]: false
-#351 := [lemma #350]: #346
-#357 := (or #310 #317)
-#358 := [th-lemma]: #357
-#359 := [unit-resolution #358 #351]: #317
-#360 := (not #317)
-#398 := (or #360 #314)
-#399 := [th-lemma]: #398
-#400 := [unit-resolution #399 #359]: #314
-#313 := (<= #321 0::real)
-#361 := (or #360 #313)
-#362 := [th-lemma]: #361
-#363 := [unit-resolution #362 #359]: #313
-#401 := [th-lemma #363 #400 #303 #397 #376 #391 #276 #279 #275]: false
-#402 := [lemma #401]: #176
-#407 := (or #404 #177)
-#405 := (or #26 #404 #177)
-#406 := [th-lemma]: #405
-#408 := [unit-resolution #406 #274]: #407
-#409 := [unit-resolution #408 #402]: #404
-#328 := (+ #15 #324)
-#329 := (<= #328 0::real)
-#410 := [unit-resolution #286 #402]: #304
-#411 := (or #336 #329)
-#412 := [th-lemma]: #411
-#413 := [unit-resolution #412 #410]: #329
-#414 := (not #329)
-#415 := (or #370 #318 #414)
-#416 := [th-lemma]: #415
-#417 := [unit-resolution #416 #413 #409]: #370
-#327 := (+ #203 #221)
-#354 := (<= #327 0::real)
-#307 := (= #203 #209)
-#384 := (or #307 #322)
-#355 := (not #307)
-#356 := [hypothesis]: #355
-#308 := (or #198 #307)
-#309 := [def-axiom]: #308
-#364 := [unit-resolution #309 #356]: #198
-#365 := [unit-resolution #280 #364]: #306
-#380 := (= #203 #191)
-#378 := (= 0::real #191)
-#372 := (= #191 0::real)
-#369 := [unit-resolution #368 #365]: #353
-#371 := [hypothesis]: #370
-#373 := [th-lemma #364 #371 #364 #369 #276 #363 #303]: #372
-#379 := [symm #373]: #378
-#374 := (= #203 0::real)
-#375 := [th-lemma #364 #371 #364 #369 #276 #363 #303]: #374
-#381 := [trans #375 #379]: #380
-#382 := [trans #381 #365]: #307
-#383 := [unit-resolution #356 #382]: false
-#385 := [lemma #383]: #384
-#418 := [unit-resolution #385 #417]: #307
-#419 := (or #355 #354)
-#420 := [th-lemma]: #419
-#421 := [unit-resolution #420 #418]: #354
-#422 := [th-lemma #387 #421 #276 #363 #303 #417]: false
-#423 := [lemma #422]: #198
-[th-lemma #363 #400 #303 #409 #423 #279 #275]: false
-unsat
--- a/src/HOL/Multivariate_Analysis/Integration.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Multivariate_Analysis/Integration.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -8,8 +8,8 @@
 begin
 
 declare [[smt_certificates="~~/src/HOL/Multivariate_Analysis/Integration.certs"]]
-declare [[smt_fixed=false]]
-declare [[z3_proofs=true]]
+declare [[smt_fixed=true]]
+declare [[smt_oracle=false]]
 
 setup {* Arith_Data.add_tactic "Ferrante-Rackoff" (K FerranteRackoff.dlo_tac) *}
 
--- a/src/HOL/Mutabelle/MutabelleExtra.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Mutabelle/MutabelleExtra.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -19,10 +19,11 @@
      "mutabelle_extra.ML"
 begin
 
-ML {* val old_tr = !Output.tracing_fn *}
-ML {* val old_wr = !Output.writeln_fn *}
-ML {* val old_pr = !Output.priority_fn *}
-ML {* val old_wa = !Output.warning_fn *}
+(* FIXME !?!?! *)
+ML {* val old_tr = !Output.Private_Hooks.tracing_fn *}
+ML {* val old_wr = !Output.Private_Hooks.writeln_fn *}
+ML {* val old_ur = !Output.Private_Hooks.urgent_message_fn *}
+ML {* val old_wa = !Output.Private_Hooks.warning_fn *}
 
 quickcheck_params [size = 5, iterations = 1000]
 (*
@@ -48,9 +49,10 @@
 *)
  *}
 
-ML {* Output.tracing_fn := old_tr *}
-ML {* Output.writeln_fn := old_wr *}
-ML {* Output.priority_fn := old_pr *}
-ML {* Output.warning_fn := old_wa *}
+(* FIXME !?!?! *)
+ML {* Output.Private_Hooks.tracing_fn := old_tr *}
+ML {* Output.Private_Hooks.writeln_fn := old_wr *}
+ML {* Output.Private_Hooks.urgent_message_fn := old_ur *}
+ML {* Output.Private_Hooks.warning_fn := old_wa *}
 
 end
--- a/src/HOL/Mutabelle/mutabelle.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Mutabelle/mutabelle.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -499,14 +499,14 @@
 fun is_executable thy insts th =
  (Quickcheck.test_term
     (ProofContext.init_global thy) false (SOME (!testgen_name)) 1 1 (preprocess thy insts (prop_of th));
-  priority "executable"; true) handle ERROR s =>
-    (priority ("not executable: " ^ s); false);
+  Output.urgent_message "executable"; true) handle ERROR s =>
+    (Output.urgent_message ("not executable: " ^ s); false);
 
 fun qc_recursive usedthy [] insts sz qciter acc = rev acc
  | qc_recursive usedthy (x::xs) insts sz qciter acc = qc_recursive usedthy xs insts sz qciter 
-     (priority ("qc_recursive: " ^ string_of_int (length xs)); ((x, pretty (the_default [] (Quickcheck.test_term
+     (Output.urgent_message ("qc_recursive: " ^ string_of_int (length xs)); ((x, pretty (the_default [] (Quickcheck.test_term
        (ProofContext.init_global usedthy) false (SOME (!testgen_name)) sz qciter (preprocess usedthy insts x)))) :: acc))
-          handle ERROR msg => (priority msg; qc_recursive usedthy xs insts sz qciter acc);
+          handle ERROR msg => (Output.urgent_message msg; qc_recursive usedthy xs insts sz qciter acc);
 
 
 (*quickcheck-test the mutants created by function mutate with type-instantiation insts, 
@@ -721,11 +721,11 @@
 fun mutqc_thystat option p mutthy usedthy commutatives forbidden_funs iter insts sz qciter filename =
  let
    val thmlist = filter
-     (fn (s, th) => not (p s th) andalso (priority s; is_executable mutthy insts th)) (thms_of mutthy)
+     (fn (s, th) => not (p s th) andalso (Output.urgent_message s; is_executable mutthy insts th)) (thms_of mutthy)
    fun mutthmrec [] = ()
    |   mutthmrec ((name,thm)::xs) =
      let          
-       val _ = priority ("mutthmrec: " ^ name);
+       val _ = Output.urgent_message ("mutthmrec: " ^ name);
        val mutated = mutate option (prop_of thm) usedthy
            commutatives forbidden_funs iter
        val (passednum,_,cenum,_) = qc_test mutated insts usedthy sz qciter
--- a/src/HOL/Mutabelle/mutabelle_extra.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Mutabelle/mutabelle_extra.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -109,7 +109,7 @@
 fun invoke_refute thy t =
   let
     val res = MyRefute.refute_term thy [] t
-    val _ = priority ("Refute: " ^ res)
+    val _ = Output.urgent_message ("Refute: " ^ res)
   in
     case res of
       "genuine" => GenuineCex
@@ -137,7 +137,7 @@
   in
     let
       val (res, _) = Nitpick.pick_nits_in_term state (Nitpick_Isar.default_params thy []) false [] t
-      val _ = priority ("Nitpick: " ^ res)
+      val _ = Output.urgent_message ("Nitpick: " ^ res)
     in
       case res of
         "genuine" => GenuineCex
@@ -182,7 +182,7 @@
            (error ("Unhandled Refute error (" ^ quote loc ^ "): " ^ details ^
                    "."))
          | Exn.Interrupt => raise Exn.Interrupt
-         | _ => (priority ("Unknown error in Nitpick"); Error)
+         | _ => (Output.urgent_message ("Unknown error in Nitpick"); Error)
   end
 val nitpick_mtd = ("nitpick", invoke_nitpick)
 *)
@@ -281,13 +281,13 @@
 
 fun safe_invoke_mtd thy (mtd_name, invoke_mtd) t =
   let
-    val _ = priority ("Invoking " ^ mtd_name)
+    val _ = Output.urgent_message ("Invoking " ^ mtd_name)
     val ((res, (timing, reports)), time) = cpu_time "total time"
       (fn () => case try (invoke_mtd thy) t of
           SOME (res, (timing, reports)) => (res, (timing, reports))
-        | NONE => (priority ("**** PROBLEMS WITH " ^ Syntax.string_of_term_global thy t);
+        | NONE => (Output.urgent_message ("**** PROBLEMS WITH " ^ Syntax.string_of_term_global thy t);
            (Error , ([], NONE))))
-    val _ = priority (" Done")
+    val _ = Output.urgent_message (" Done")
   in (res, (time :: timing, reports)) end
 
 (* theory -> term list -> mtd -> subentry *)
@@ -316,7 +316,7 @@
 fun mutate_theorem create_entry thy mtds thm =
   let
     val exec = is_executable_thm thy thm
-    val _ = priority (if exec then "EXEC" else "NOEXEC")
+    val _ = Output.urgent_message (if exec then "EXEC" else "NOEXEC")
     val mutants =
           (if num_mutations = 0 then
              [Thm.prop_of thm]
@@ -327,7 +327,7 @@
     val mutants =
       if exec then
         let
-          val _ = priority ("BEFORE PARTITION OF " ^
+          val _ = Output.urgent_message ("BEFORE PARTITION OF " ^
                             Int.toString (length mutants) ^ " MUTANTS")
           val (execs, noexecs) = List.partition (is_executable_term thy) (take_random (20 * max_mutants) mutants)
           val _ = tracing ("AFTER PARTITION (" ^ Int.toString (length execs) ^
@@ -342,7 +342,7 @@
           |> map Mutabelle.freeze |> map freezeT
 (*          |> filter (not o is_forbidden_mutant) *)
           |> List.mapPartial (try (Sign.cert_term thy))
-    val _ = map (fn t => priority ("MUTANT: " ^ Syntax.string_of_term_global thy t)) mutants
+    val _ = map (fn t => Output.urgent_message ("MUTANT: " ^ Syntax.string_of_term_global thy t)) mutants
   in
     create_entry thy thm exec mutants mtds
   end
@@ -421,7 +421,7 @@
 (* theory -> mtd list -> thm list -> string -> unit *)
 fun mutate_theorems_and_write_report thy mtds thms file_name =
   let
-    val _ = priority "Starting Mutabelle..."
+    val _ = Output.urgent_message "Starting Mutabelle..."
     val path = Path.explode file_name
     (* for normal report: *)
     (*val (gen_create_entry, gen_string_for_entry) = (create_entry, string_for_entry)*)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Partial_Function.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -0,0 +1,247 @@
+(* Title:    HOL/Partial_Function.thy
+   Author:   Alexander Krauss, TU Muenchen
+*)
+
+header {* Partial Function Definitions *}
+
+theory Partial_Function
+imports Complete_Partial_Order Option
+uses 
+  "Tools/Function/function_lib.ML" 
+  "Tools/Function/partial_function.ML" 
+begin
+
+setup Partial_Function.setup
+
+subsection {* Axiomatic setup *}
+
+text {* This techical locale constains the requirements for function
+  definitions with ccpo fixed points.  *}
+
+definition "fun_ord ord f g \<longleftrightarrow> (\<forall>x. ord (f x) (g x))"
+definition "fun_lub L A = (\<lambda>x. L {y. \<exists>f\<in>A. y = f x})"
+definition "img_ord f ord = (\<lambda>x y. ord (f x) (f y))"
+definition "img_lub f g Lub = (\<lambda>A. g (Lub (f ` A)))"
+
+lemma call_mono[partial_function_mono]: "monotone (fun_ord ord) ord (\<lambda>f. f t)"
+by (rule monotoneI) (auto simp: fun_ord_def)
+
+lemma if_mono[partial_function_mono]: "monotone orda ordb F 
+\<Longrightarrow> monotone orda ordb G
+\<Longrightarrow> monotone orda ordb (\<lambda>f. if c then F f else G f)"
+unfolding monotone_def by simp
+
+definition "mk_less R = (\<lambda>x y. R x y \<and> \<not> R y x)"
+
+locale partial_function_definitions = 
+  fixes leq :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
+  fixes lub :: "'a set \<Rightarrow> 'a"
+  assumes leq_refl: "leq x x"
+  assumes leq_trans: "leq x y \<Longrightarrow> leq y z \<Longrightarrow> leq x z"
+  assumes leq_antisym: "leq x y \<Longrightarrow> leq y x \<Longrightarrow> x = y"
+  assumes lub_upper: "chain leq A \<Longrightarrow> x \<in> A \<Longrightarrow> leq x (lub A)"
+  assumes lub_least: "chain leq A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> leq x z) \<Longrightarrow> leq (lub A) z"
+
+lemma partial_function_lift:
+  assumes "partial_function_definitions ord lb"
+  shows "partial_function_definitions (fun_ord ord) (fun_lub lb)" (is "partial_function_definitions ?ordf ?lubf")
+proof -
+  interpret partial_function_definitions ord lb by fact
+
+  { fix A a assume A: "chain ?ordf A"
+    have "chain ord {y. \<exists>f\<in>A. y = f a}" (is "chain ord ?C")
+    proof (rule chainI)
+      fix x y assume "x \<in> ?C" "y \<in> ?C"
+      then obtain f g where fg: "f \<in> A" "g \<in> A" 
+        and [simp]: "x = f a" "y = g a" by blast
+      from chainD[OF A fg]
+      show "ord x y \<or> ord y x" unfolding fun_ord_def by auto
+    qed }
+  note chain_fun = this
+
+  show ?thesis
+  proof
+    fix x show "?ordf x x"
+      unfolding fun_ord_def by (auto simp: leq_refl)
+  next
+    fix x y z assume "?ordf x y" "?ordf y z"
+    thus "?ordf x z" unfolding fun_ord_def 
+      by (force dest: leq_trans)
+  next
+    fix x y assume "?ordf x y" "?ordf y x"
+    thus "x = y" unfolding fun_ord_def
+      by (force intro!: ext dest: leq_antisym)
+  next
+    fix A f assume f: "f \<in> A" and A: "chain ?ordf A"
+    thus "?ordf f (?lubf A)"
+      unfolding fun_lub_def fun_ord_def
+      by (blast intro: lub_upper chain_fun[OF A] f)
+  next
+    fix A :: "('b \<Rightarrow> 'a) set" and g :: "'b \<Rightarrow> 'a"
+    assume A: "chain ?ordf A" and g: "\<And>f. f \<in> A \<Longrightarrow> ?ordf f g"
+    show "?ordf (?lubf A) g" unfolding fun_lub_def fun_ord_def
+      by (blast intro: lub_least chain_fun[OF A] dest: g[unfolded fun_ord_def])
+   qed
+qed
+
+lemma ccpo: assumes "partial_function_definitions ord lb"
+  shows "class.ccpo ord (mk_less ord) lb"
+using assms unfolding partial_function_definitions_def mk_less_def
+by unfold_locales blast+
+
+lemma partial_function_image:
+  assumes "partial_function_definitions ord Lub"
+  assumes inj: "\<And>x y. f x = f y \<Longrightarrow> x = y"
+  assumes inv: "\<And>x. f (g x) = x"
+  shows "partial_function_definitions (img_ord f ord) (img_lub f g Lub)"
+proof -
+  let ?iord = "img_ord f ord"
+  let ?ilub = "img_lub f g Lub"
+
+  interpret partial_function_definitions ord Lub by fact
+  show ?thesis
+  proof
+    fix A x assume "chain ?iord A" "x \<in> A"
+    then have "chain ord (f ` A)" "f x \<in> f ` A"
+      by (auto simp: img_ord_def intro: chainI dest: chainD)
+    thus "?iord x (?ilub A)"
+      unfolding inv img_lub_def img_ord_def by (rule lub_upper)
+  next
+    fix A x assume "chain ?iord A"
+      and 1: "\<And>z. z \<in> A \<Longrightarrow> ?iord z x"
+    then have "chain ord (f ` A)"
+      by (auto simp: img_ord_def intro: chainI dest: chainD)
+    thus "?iord (?ilub A) x"
+      unfolding inv img_lub_def img_ord_def
+      by (rule lub_least) (auto dest: 1[unfolded img_ord_def])
+  qed (auto simp: img_ord_def intro: leq_refl dest: leq_trans leq_antisym inj)
+qed
+
+context partial_function_definitions
+begin
+
+abbreviation "le_fun \<equiv> fun_ord leq"
+abbreviation "lub_fun \<equiv> fun_lub lub"
+abbreviation "fixp_fun == ccpo.fixp le_fun lub_fun"
+abbreviation "mono_body \<equiv> monotone le_fun leq"
+
+text {* Interpret manually, to avoid flooding everything with facts about
+  orders *}
+
+lemma ccpo: "class.ccpo le_fun (mk_less le_fun) lub_fun"
+apply (rule ccpo)
+apply (rule partial_function_lift)
+apply (rule partial_function_definitions_axioms)
+done
+
+text {* The crucial fixed-point theorem *}
+
+lemma mono_body_fixp: 
+  "(\<And>x. mono_body (\<lambda>f. F f x)) \<Longrightarrow> fixp_fun F = F (fixp_fun F)"
+by (rule ccpo.fixp_unfold[OF ccpo]) (auto simp: monotone_def fun_ord_def)
+
+text {* Version with curry/uncurry combinators, to be used by package *}
+
+lemma fixp_rule_uc:
+  fixes F :: "'c \<Rightarrow> 'c" and
+    U :: "'c \<Rightarrow> 'b \<Rightarrow> 'a" and
+    C :: "('b \<Rightarrow> 'a) \<Rightarrow> 'c"
+  assumes mono: "\<And>x. mono_body (\<lambda>f. U (F (C f)) x)"
+  assumes eq: "f \<equiv> C (fixp_fun (\<lambda>f. U (F (C f))))"
+  assumes inverse: "\<And>f. C (U f) = f"
+  shows "f = F f"
+proof -
+  have "f = C (fixp_fun (\<lambda>f. U (F (C f))))" by (simp add: eq)
+  also have "... = C (U (F (C (fixp_fun (\<lambda>f. U (F (C f)))))))"
+    by (subst mono_body_fixp[of "%f. U (F (C f))", OF mono]) (rule refl)
+  also have "... = F (C (fixp_fun (\<lambda>f. U (F (C f)))))" by (rule inverse)
+  also have "... = F f" by (simp add: eq)
+  finally show "f = F f" .
+qed
+
+text {* Rules for @{term mono_body}: *}
+
+lemma const_mono[partial_function_mono]: "monotone ord leq (\<lambda>f. c)"
+by (rule monotoneI) (rule leq_refl)
+
+declaration {* Partial_Function.init @{term fixp_fun}
+  @{term mono_body} @{thm fixp_rule_uc} *}
+
+end
+
+
+subsection {* Flat interpretation: tailrec and option *}
+
+definition 
+  "flat_ord b x y \<longleftrightarrow> x = b \<or> x = y"
+
+definition 
+  "flat_lub b A = (if A \<subseteq> {b} then b else (THE x. x \<in> A - {b}))"
+
+lemma flat_interpretation:
+  "partial_function_definitions (flat_ord b) (flat_lub b)"
+proof
+  fix A x assume 1: "chain (flat_ord b) A" "x \<in> A"
+  show "flat_ord b x (flat_lub b A)"
+  proof cases
+    assume "x = b"
+    thus ?thesis by (simp add: flat_ord_def)
+  next
+    assume "x \<noteq> b"
+    with 1 have "A - {b} = {x}"
+      by (auto elim: chainE simp: flat_ord_def)
+    then have "flat_lub b A = x"
+      by (auto simp: flat_lub_def)
+    thus ?thesis by (auto simp: flat_ord_def)
+  qed
+next
+  fix A z assume A: "chain (flat_ord b) A"
+    and z: "\<And>x. x \<in> A \<Longrightarrow> flat_ord b x z"
+  show "flat_ord b (flat_lub b A) z"
+  proof cases
+    assume "A \<subseteq> {b}"
+    thus ?thesis
+      by (auto simp: flat_lub_def flat_ord_def)
+  next
+    assume nb: "\<not> A \<subseteq> {b}"
+    then obtain y where y: "y \<in> A" "y \<noteq> b" by auto
+    with A have "A - {b} = {y}"
+      by (auto elim: chainE simp: flat_ord_def)
+    with nb have "flat_lub b A = y"
+      by (auto simp: flat_lub_def)
+    with z y show ?thesis by auto    
+  qed
+qed (auto simp: flat_ord_def)
+
+interpretation tailrec!:
+  partial_function_definitions "flat_ord undefined" "flat_lub undefined"
+by (rule flat_interpretation)
+
+interpretation option!:
+  partial_function_definitions "flat_ord None" "flat_lub None"
+by (rule flat_interpretation)
+
+abbreviation "option_ord \<equiv> flat_ord None"
+abbreviation "mono_option \<equiv> monotone (fun_ord option_ord) option_ord"
+
+lemma bind_mono[partial_function_mono]:
+assumes mf: "mono_option B" and mg: "\<And>y. mono_option (\<lambda>f. C y f)"
+shows "mono_option (\<lambda>f. Option.bind (B f) (\<lambda>y. C y f))"
+proof (rule monotoneI)
+  fix f g :: "'a \<Rightarrow> 'b option" assume fg: "fun_ord option_ord f g"
+  with mf
+  have "option_ord (B f) (B g)" by (rule monotoneD[of _ _ _ f g])
+  then have "option_ord (Option.bind (B f) (\<lambda>y. C y f)) (Option.bind (B g) (\<lambda>y. C y f))"
+    unfolding flat_ord_def by auto    
+  also from mg
+  have "\<And>y'. option_ord (C y' f) (C y' g)"
+    by (rule monotoneD) (rule fg)
+  then have "option_ord (Option.bind (B g) (\<lambda>y'. C y' f)) (Option.bind (B g) (\<lambda>y'. C y' g))"
+    unfolding flat_ord_def by (cases "B g") auto
+  finally (option.leq_trans)
+  show "option_ord (Option.bind (B f) (\<lambda>y. C y f)) (Option.bind (B g) (\<lambda>y'. C y' g))" .
+qed
+
+
+end
+
--- a/src/HOL/Predicate_Compile_Examples/Hotel_Example.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Predicate_Compile_Examples/Hotel_Example.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -1,5 +1,5 @@
 theory Hotel_Example
-imports Predicate_Compile_Alternative_Defs Code_Prolog
+imports Main
 begin
 
 datatype guest = Guest0 | Guest1
@@ -71,145 +71,21 @@
   Enter g r (k,k') \<Rightarrow> (k,k') : cards s g & (roomk s r : {k, k'}) |
   Exit g r \<Rightarrow> g : isin s r))"
 
-lemma issued_nil: "issued [] = {Key0}"
-by (auto simp add: initk_def)
-
-lemmas issued_simps[code, code_pred_def] = issued_nil issued.simps(2)
-
-declare Let_def[code_pred_inline]
-
-lemma [code_pred_inline]: "insert == (%y A x. y = x | A x)"
-by (auto simp add: insert_iff[unfolded mem_def] fun_eq_iff intro!: eq_reflection)
-
-lemma [code_pred_inline]: "(op -) == (%A B x. A x \<and> \<not> B x)"
-by (auto simp add: Diff_iff[unfolded mem_def] fun_eq_iff intro!: eq_reflection)
-
-setup {* Code_Prolog.map_code_options (K
-  {ensure_groundness = true,
-  limit_globally = NONE,
-  limited_types = [],
-  limited_predicates = [],
-  replacing = [],
-  manual_reorder = []}) *}
-
-values 40 "{s. hotel s}"
-
-
-setup {* Context.theory_map (Quickcheck.add_generator ("prolog", Code_Prolog.quickcheck)) *}
-
-lemma "\<lbrakk> hotel s; g \<in> isin s r \<rbrakk> \<Longrightarrow> owns s r = Some g"
-(*quickcheck[generator = code, iterations = 100000, report]*)
-quickcheck[generator = prolog, iterations = 1, expect = counterexample]
-oops
-
-
 definition no_Check_in :: "event list \<Rightarrow> room \<Rightarrow> bool" where(*>*)
 [code del]: "no_Check_in s r \<equiv> \<not>(\<exists>g c. Check_in g r c \<in> set s)"
 
-
 definition feels_safe :: "event list \<Rightarrow> room \<Rightarrow> bool"
 where
   "feels_safe s r = (\<exists>s\<^isub>1 s\<^isub>2 s\<^isub>3 g c c'.
    s = s\<^isub>3 @ [Enter g r c] @ s\<^isub>2 @ [Check_in g r c'] @ s\<^isub>1 \<and>
    no_Check_in (s\<^isub>3 @ s\<^isub>2) r \<and> isin (s\<^isub>2 @ [Check_in g r c] @ s\<^isub>1) r = {})"
 
-section {* Manual setup to find the counterexample *}
 
-setup {* Code_Prolog.map_code_options (K 
-  {ensure_groundness = true,
-   limit_globally = NONE,
-   limited_types = [],
-   limited_predicates = [(["hotel"], 4)],
-   replacing = [(("hotel", "limited_hotel"), "quickcheck")],
-   manual_reorder = []}) *}
-
-lemma
-  "hotel s ==> feels_safe s r ==> g \<in> isin s r ==> owns s r = Some g"
-quickcheck[generator = prolog, iterations = 1, expect = no_counterexample]
-oops
-
-setup {* Code_Prolog.map_code_options (K 
-  {ensure_groundness = true,
-   limit_globally = NONE,
-   limited_types = [],
-   limited_predicates = [(["hotel"], 5)],
-   replacing = [(("hotel", "limited_hotel"), "quickcheck")],
-   manual_reorder = []}) *}
-
-lemma
-  "hotel s ==> feels_safe s r ==> g \<in> isin s r ==> owns s r = Some g"
-quickcheck[generator = prolog, iterations = 1, expect = counterexample]
-oops
-
-section {* Simulating a global depth limit manually by limiting all predicates *}
-
-setup {*
-  Code_Prolog.map_code_options (K
-  {ensure_groundness = true,
-  limit_globally = NONE,
-  limited_types = [],
-  limited_predicates = [(["feelssafe", "feelssafeaux", "noCheckin", "noCheckinaux", "appendP", "ownsP", "hotel", "hotelaux", "hotelauxaux", "roomkP", "issued", "currkP", "initkP",
-    "cards", "cardsauxauxaux", "cardsauxaux", "cardsaux", "isin", "isinauxauxa", "isinauxauxaux", "isinauxaux", "isinaux", "set"], 12)],
-  replacing = [(("hotel", "limited_hotel"), "quickcheck"), (("feelssafe", "limited_feelssafe"), "quickcheck"), (("isin", "limited_isin"), "quickcheck")],
-  manual_reorder = []})
-*}
-
-lemma
-  "hotel s ==> feels_safe s r ==> g \<in> isin s r ==> owns s r = Some g"
-quickcheck[generator = prolog, iterations = 1, expect = no_counterexample]
-oops
-
+section {* Some setup *}
 
-setup {*
-  Code_Prolog.map_code_options (K
-  {ensure_groundness = true,
-  limit_globally = NONE,
-  limited_types = [],
-  limited_predicates = [(["feelssafe", "feelssafeaux", "noCheckin", "noCheckinaux", "appendP", "ownsP", "hotel", "hotelaux", "hotelauxaux", "roomkP", "issued", "currkP", "initkP",
-    "cards", "cardsauxauxaux", "cardsauxaux", "cardsaux", "isin", "isinauxauxa", "isinauxauxaux", "isinauxaux", "isinaux", "set"], 13)],
-  replacing = [(("hotel", "limited_hotel"), "quickcheck"), (("feelssafe", "limited_feelssafe"), "quickcheck"), (("isin", "limited_isin"), "quickcheck")],
-  manual_reorder = []})
-*}
-
-lemma
-  "hotel s ==> feels_safe s r ==> g \<in> isin s r ==> owns s r = Some g"
-quickcheck[generator = prolog, iterations = 1, expect = counterexample]
-oops
-
-section {* Using a global limit for limiting the execution *} 
-
-text {* A global depth limit of 13 does not suffice to find the counterexample. *}
+lemma issued_nil: "issued [] = {Key0}"
+by (auto simp add: initk_def)
 
-setup {*
-  Code_Prolog.map_code_options (K
-  {ensure_groundness = true,
-  limit_globally = SOME 13,
-  limited_types = [],
-  limited_predicates = [],
-  replacing = [],
-  manual_reorder = []})
-*}
-
-lemma
-  "hotel s ==> feels_safe s r ==> g \<in> isin s r ==> owns s r = Some g"
-quickcheck[generator = prolog, iterations = 1, expect = no_counterexample]
-oops
-
-text {* But a global depth limit of 14 does. *}
-
-setup {*
-  Code_Prolog.map_code_options (K
-  {ensure_groundness = true,
-  limit_globally = SOME 14,
-  limited_types = [],
-  limited_predicates = [],
-  replacing = [],
-  manual_reorder = []})
-*}
-
-lemma
-  "hotel s ==> feels_safe s r ==> g \<in> isin s r ==> owns s r = Some g"
-quickcheck[generator = prolog, iterations = 1, expect = counterexample]
-oops
+lemmas issued_simps[code, code_pred_def] = issued_nil issued.simps(2)
 
 end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Predicate_Compile_Examples/Hotel_Example_Prolog.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -0,0 +1,128 @@
+theory Hotel_Example_Prolog
+imports Hotel_Example Predicate_Compile_Alternative_Defs Code_Prolog
+begin
+
+declare Let_def[code_pred_inline]
+
+lemma [code_pred_inline]: "insert == (%y A x. y = x | A x)"
+by (auto simp add: insert_iff[unfolded mem_def] fun_eq_iff intro!: eq_reflection)
+
+lemma [code_pred_inline]: "(op -) == (%A B x. A x \<and> \<not> B x)"
+by (auto simp add: Diff_iff[unfolded mem_def] fun_eq_iff intro!: eq_reflection)
+
+setup {* Code_Prolog.map_code_options (K
+  {ensure_groundness = true,
+  limit_globally = NONE,
+  limited_types = [],
+  limited_predicates = [],
+  replacing = [],
+  manual_reorder = []}) *}
+
+values 40 "{s. hotel s}"
+
+setup {* Context.theory_map (Quickcheck.add_generator ("prolog", Code_Prolog.quickcheck)) *}
+
+lemma "\<lbrakk> hotel s; g \<in> isin s r \<rbrakk> \<Longrightarrow> owns s r = Some g"
+quickcheck[generator = code, iterations = 10000, report]
+quickcheck[generator = prolog, iterations = 1, expect = counterexample]
+oops
+
+section {* Manual setup to find the counterexample *}
+
+setup {* Code_Prolog.map_code_options (K 
+  {ensure_groundness = true,
+   limit_globally = NONE,
+   limited_types = [],
+   limited_predicates = [(["hotel"], 4)],
+   replacing = [(("hotel", "limited_hotel"), "quickcheck")],
+   manual_reorder = []}) *}
+
+lemma
+  "hotel s ==> feels_safe s r ==> g \<in> isin s r ==> owns s r = Some g"
+quickcheck[generator = prolog, iterations = 1, expect = no_counterexample]
+oops
+
+setup {* Code_Prolog.map_code_options (K 
+  {ensure_groundness = true,
+   limit_globally = NONE,
+   limited_types = [],
+   limited_predicates = [(["hotel"], 5)],
+   replacing = [(("hotel", "limited_hotel"), "quickcheck")],
+   manual_reorder = []}) *}
+
+lemma
+  "hotel s ==> feels_safe s r ==> g \<in> isin s r ==> owns s r = Some g"
+quickcheck[generator = prolog, iterations = 1, expect = counterexample]
+oops
+
+section {* Simulating a global depth limit manually by limiting all predicates *}
+
+setup {*
+  Code_Prolog.map_code_options (K
+  {ensure_groundness = true,
+  limit_globally = NONE,
+  limited_types = [],
+  limited_predicates = [(["feelssafe", "feelssafeaux", "noCheckin", "noCheckinaux", "appendP", "ownsP", "hotel", "hotelaux", "hotelauxaux", "roomkP", "issued", "currkP", "initkP",
+    "cards", "cardsauxauxaux", "cardsauxaux", "cardsaux", "isin", "isinauxauxa", "isinauxauxaux", "isinauxaux", "isinaux", "set"], 12)],
+  replacing = [(("hotel", "limited_hotel"), "quickcheck"), (("feelssafe", "limited_feelssafe"), "quickcheck"), (("isin", "limited_isin"), "quickcheck")],
+  manual_reorder = []})
+*}
+
+lemma
+  "hotel s ==> feels_safe s r ==> g \<in> isin s r ==> owns s r = Some g"
+quickcheck[generator = prolog, iterations = 1, expect = no_counterexample]
+oops
+
+setup {*
+  Code_Prolog.map_code_options (K
+  {ensure_groundness = true,
+  limit_globally = NONE,
+  limited_types = [],
+  limited_predicates = [(["feelssafe", "feelssafeaux", "noCheckin", "noCheckinaux", "appendP", "ownsP", "hotel", "hotelaux", "hotelauxaux", "roomkP", "issued", "currkP", "initkP",
+    "cards", "cardsauxauxaux", "cardsauxaux", "cardsaux", "isin", "isinauxauxa", "isinauxauxaux", "isinauxaux", "isinaux", "set"], 13)],
+  replacing = [(("hotel", "limited_hotel"), "quickcheck"), (("feelssafe", "limited_feelssafe"), "quickcheck"), (("isin", "limited_isin"), "quickcheck")],
+  manual_reorder = []})
+*}
+
+lemma
+  "hotel s ==> feels_safe s r ==> g \<in> isin s r ==> owns s r = Some g"
+quickcheck[generator = prolog, iterations = 1, expect = counterexample]
+oops
+
+section {* Using a global limit for limiting the execution *} 
+
+text {* A global depth limit of 13 does not suffice to find the counterexample. *}
+
+setup {*
+  Code_Prolog.map_code_options (K
+  {ensure_groundness = true,
+  limit_globally = SOME 13,
+  limited_types = [],
+  limited_predicates = [],
+  replacing = [],
+  manual_reorder = []})
+*}
+
+lemma
+  "hotel s ==> feels_safe s r ==> g \<in> isin s r ==> owns s r = Some g"
+quickcheck[generator = prolog, iterations = 1, expect = no_counterexample]
+oops
+
+text {* But a global depth limit of 14 does. *}
+
+setup {*
+  Code_Prolog.map_code_options (K
+  {ensure_groundness = true,
+  limit_globally = SOME 14,
+  limited_types = [],
+  limited_predicates = [],
+  replacing = [],
+  manual_reorder = []})
+*}
+
+lemma
+  "hotel s ==> feels_safe s r ==> g \<in> isin s r ==> owns s r = Some g"
+quickcheck[generator = prolog, iterations = 1, expect = counterexample]
+oops
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Predicate_Compile_Examples/Hotel_Example_Small_Generator.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -0,0 +1,58 @@
+theory Hotel_Example_Small_Generator
+imports Hotel_Example Predicate_Compile_Alternative_Defs
+uses "~~/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML"
+begin
+
+declare Let_def[code_pred_inline]
+
+lemma [code_pred_inline]: "insert == (%y A x. y = x | A x)"
+by (auto simp add: insert_iff[unfolded mem_def] fun_eq_iff intro!: eq_reflection)
+
+lemma [code_pred_inline]: "(op -) == (%A B x. A x \<and> \<not> B x)"
+by (auto simp add: Diff_iff[unfolded mem_def] fun_eq_iff intro!: eq_reflection)
+
+instantiation room :: small_lazy
+begin
+
+definition
+  "small_lazy i = Lazy_Sequence.single Room0"
+
+instance ..
+
+end
+
+instantiation key :: small_lazy
+begin
+
+definition
+  "small_lazy i = Lazy_Sequence.append (Lazy_Sequence.single Key0) (Lazy_Sequence.append (Lazy_Sequence.single Key1) (Lazy_Sequence.append (Lazy_Sequence.single Key2) (Lazy_Sequence.single Key3)))"
+
+instance ..
+
+end
+
+instantiation guest :: small_lazy
+begin
+
+definition
+  "small_lazy i = Lazy_Sequence.append (Lazy_Sequence.single Guest0) (Lazy_Sequence.single Guest1)"
+
+instance ..
+
+end
+
+setup {* Context.theory_map (Quickcheck.add_generator ("small_generators_depth_14", Predicate_Compile_Quickcheck.quickcheck_compile_term
+  Predicate_Compile_Aux.Pos_Generator_DSeq true true 14)) *}
+
+
+setup {* Context.theory_map (Quickcheck.add_generator ("small_generators_depth_15", Predicate_Compile_Quickcheck.quickcheck_compile_term
+  Predicate_Compile_Aux.Pos_Generator_DSeq true true 15)) *}
+
+lemma
+  "hotel s ==> feels_safe s r ==> g \<in> isin s r ==> owns s r = Some g"
+quickcheck[generator = small_generators_depth_14, iterations = 1, size = 1, expect = no_counterexample]
+quickcheck[generator = small_generators_depth_15, iterations = 1, size = 1, expect = counterexample]
+oops
+
+
+end
\ No newline at end of file
--- a/src/HOL/Predicate_Compile_Examples/List_Examples.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Predicate_Compile_Examples/List_Examples.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -16,7 +16,7 @@
    manual_reorder = []}) *}
 
 lemma "(xs :: nat list) = ys @ ys --> rev xs = xs"
-quickcheck[generator = code, iterations = 200000, expect = counterexample]
+quickcheck[generator = code, iterations = 10000]
 quickcheck[generator = predicate_compile_wo_ff, iterations = 1, expect = counterexample]
 quickcheck[generator = prolog, expect = counterexample]
 oops
--- a/src/HOL/Predicate_Compile_Examples/Predicate_Compile_Tests.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Predicate_Compile_Examples/Predicate_Compile_Tests.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -70,14 +70,13 @@
 where
   "False \<Longrightarrow> False''"
 
-code_pred (expected_modes: []) False'' .
+code_pred (expected_modes: bool) False'' .
 
 inductive EmptySet'' :: "'a \<Rightarrow> bool"
 where
   "False \<Longrightarrow> EmptySet'' x"
 
-code_pred (expected_modes: [1]) EmptySet'' .
-code_pred (expected_modes: [], [1]) [inductify] EmptySet'' .
+code_pred (expected_modes: i => bool, o => bool) [inductify] EmptySet'' .
 *)
 
 consts a' :: 'a
@@ -975,9 +974,10 @@
 
 code_pred [new_random_dseq inductify] avl .
 thm avl.new_random_dseq_equation
+(* TODO: has highly non-deterministic execution time!
 
 values [new_random_dseq 2, 1, 7] 5 "{t:: int tree. avl t}"
-
+*)
 fun set_of
 where
 "set_of ET = {}"
--- a/src/HOL/Predicate_Compile_Examples/ROOT.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Predicate_Compile_Examples/ROOT.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -3,10 +3,11 @@
   "Predicate_Compile_Tests",
 (*  "Predicate_Compile_Quickcheck_Examples", -- should be added again soon *)
   "Specialisation_Examples",
+  "Hotel_Example_Small_Generator",
   "IMP_1",
-  "IMP_2",
-  "IMP_3",
-  "IMP_4"];
+  "IMP_2"
+(*  "IMP_3",
+  "IMP_4"*)];
 
 if getenv "EXEC_SWIPL" = "" andalso getenv "EXEC_YAP" = "" then
   (warning "No prolog system found - skipping some example theories"; ())
@@ -14,7 +15,7 @@
   (use_thys [
     "Code_Prolog_Examples",
     "Context_Free_Grammar_Example",
-    "Hotel_Example",
+    "Hotel_Example_Prolog",
     "Lambda_Example",
     "List_Examples"];
    Unsynchronized.setmp quick_and_dirty true use_thys ["Reg_Exp_Example"]);
--- a/src/HOL/Probability/Euclidean_Lebesgue.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Probability/Euclidean_Lebesgue.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -122,7 +122,6 @@
   let ?u = "\<lambda>i x. real (u i x)" and ?f = "\<lambda>x. real (f x)"
   have u_simple:"\<And>k. lebesgue.simple_integral (u k) = lebesgue.positive_integral (u k)"
     apply(subst lebesgue.positive_integral_eq_simple_integral[THEN sym,OF u(1)]) ..
-    (*unfolding image_iff defer apply(rule) using u(2) by smt*)
   have int_u_le:"\<And>k. lebesgue.simple_integral (u k) \<le> lebesgue.positive_integral f"
     unfolding u_simple apply(rule lebesgue.positive_integral_mono)
     using isoton_Sup[OF u(3)] unfolding le_fun_def by auto
--- a/src/HOL/SMT.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/SMT.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -19,9 +19,7 @@
   ("Tools/SMT/z3_proof_reconstruction.ML")
   ("Tools/SMT/z3_model.ML")
   ("Tools/SMT/z3_interface.ML")
-  ("Tools/SMT/z3_solver.ML")
-  ("Tools/SMT/cvc3_solver.ML")
-  ("Tools/SMT/yices_solver.ML")
+  ("Tools/SMT/smt_setup_solvers.ML")
 begin
 
 
@@ -124,16 +122,12 @@
 use "Tools/SMT/z3_proof_literals.ML"
 use "Tools/SMT/z3_proof_reconstruction.ML"
 use "Tools/SMT/z3_model.ML"
-use "Tools/SMT/z3_solver.ML"
-use "Tools/SMT/cvc3_solver.ML"
-use "Tools/SMT/yices_solver.ML"
+use "Tools/SMT/smt_setup_solvers.ML"
 
 setup {*
   SMT_Solver.setup #>
   Z3_Proof_Reconstruction.setup #>
-  Z3_Solver.setup #>
-  CVC3_Solver.setup #>
-  Yices_Solver.setup
+  SMT_Setup_Solvers.setup
 *}
 
 
@@ -171,6 +165,31 @@
 
 declare [[ smt_timeout = 20 ]]
 
+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 [[ smt_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 = "", yices_options = "", z3_options = "" ]]
+
+text {*
+Enable the following option to use built-in support for datatypes and
+records.  Currently, this is only implemented for Z3 running in oracle
+mode.
+*}
+
+declare [[ smt_datatypes = false ]]
+
 
 
 subsection {* Certificates *}
@@ -213,41 +232,14 @@
 
 declare [[ smt_trace = false ]]
 
-
-
-subsection {* Z3-specific options *}
-
 text {*
-Z3 is the only SMT solver whose proofs are checked (or reconstructed)
-in Isabelle (all other solvers are implemented as oracles).  Enabling
-or disabling proof reconstruction for Z3 is controlled by the option
-@{text z3_proofs}. 
+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 smt_solver} and @{text smt_oracle} above).
 *}
 
-declare [[ z3_proofs = true ]]
-
-text {*
-From the set of assumptions given to Z3, those assumptions used in
-the proof are traced when the option @{text z3_trace_assms} is set to
-@{term true}.
-*}
-
-declare [[ z3_trace_assms = false ]]
-
-text {*
-Z3 provides several commandline options to tweak its behaviour.  They
-can be configured by writing them literally as value for the option
-@{text z3_options}.
-*}
-
-declare [[ z3_options = "" ]]
-
-text {*
-The following configuration option may be used to enable mapping of
-HOL datatypes and records to native datatypes provided by Z3.
-*}
-
-declare [[ z3_datatypes = false ]]
+declare [[ smt_trace_used_facts = false ]]
 
 
 
--- a/src/HOL/SMT_Examples/SMT_Examples.certs	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/SMT_Examples/SMT_Examples.certs	Wed Oct 27 11:11:35 2010 -0700
@@ -1,4 +1,4 @@
-ef3118b3c538f5cd123d726ccf13a3c2d3725bb0 8 0
+d4b1481d70f20d4d6280999633c905e21d180d13 8 0
 #2 := false
 #1 := true
 #8 := (not true)
@@ -7,7 +7,7 @@
 #25 := [asserted]: #8
 [mp #25 #27]: false
 unsat
-b59b2f429ffd47c407e4ef1af006540275b1a26c 33 0
+a38979487be0aafcf2d387180e92dfd3bd43e483 33 0
 #2 := false
 decl f1 :: S1
 #4 := f1
@@ -41,7 +41,7 @@
 #29 := [asserted]: #12
 [mp #29 #49]: false
 unsat
-2e7e9d422d65be29e15f02c69c60971cab2c6ff5 37 0
+fd75276ee28b486a6ea1b6b677d04008959d13f8 37 0
 #2 := false
 decl f1 :: S1
 #4 := f1
@@ -79,7 +79,7 @@
 #29 := [asserted]: #12
 [mp #29 #53]: false
 unsat
-c09f42a4df50b3429cb0f5ddbf56d93bf0b06b7b 66 0
+58aaefc3b25d8181c01a2139430686aeca861bf3 66 0
 #2 := false
 decl f4 :: S1
 #10 := f4
@@ -148,7 +148,7 @@
 unsat
 a3969d96ad7c1fdd20f7d69abeb13b18bac4be8d 1 0
 unsat
-8f1394900829ee257ef5682b19118d98b3c0c820 94 0
+694292afe609b4087c0435b347aed1c0c0b6b45b 94 0
 #2 := false
 decl f1 :: S1
 #4 := f1
@@ -243,7 +243,7 @@
 #39 := [asserted]: #22
 [mp #39 #108]: false
 unsat
-41b693f504b47ef3e9345eeec00c566a4d93af04 72 0
+fcb0f1f1909bdd097de1e485a000708289f9f1ed 72 0
 #2 := false
 decl f1 :: S1
 #4 := f1
@@ -316,7 +316,7 @@
 #36 := [asserted]: #19
 [mp #36 #88]: false
 unsat
-1bb9b155bb9b619999029c3bea6bb00c7c21cd3b 234 0
+36c95d1da5779fd86d36cb33e0f80e4f14e6c4f0 234 0
 #2 := false
 decl f6 :: S1
 #14 := f6
@@ -551,7 +551,7 @@
 #96 := [mp #69 #95]: #91
 [mp #96 #299]: false
 unsat
-c02efa254926732c7b107199403ddd23251d413c 52 0
+44ca45de077963316e54f5cdfef5cb8cd66ac906 52 0
 #2 := false
 decl f4 :: (-> S2 S2 S2)
 #16 := (:var 1 S2)
@@ -604,7 +604,7 @@
 #118 := [quant-inst]: #204
 [unit-resolution #118 #53 #541]: false
 unsat
-950f494f909b64c53d64f243647b66f8a05dd7f6 2578 0
+f79d03683394ca6e785cbd8e87f4249c3d42307a 2578 0
 #2 := false
 decl f35 :: S1
 #112 := f35
@@ -3183,7 +3183,7 @@
 #3638 := [unit-resolution #705 #3637 #3636 #3635]: #683
 [unit-resolution #1968 #3638 #3634]: false
 unsat
-d9089f92efb7cf3cd5f913b2fd8f4e6cc7fca01b 95 0
+3221f6bae508cb0c6bc508c2b8b4600cdb00669d 95 0
 #2 := false
 decl f3 :: (-> int S1)
 decl ?v0!0 :: int
@@ -3279,7 +3279,7 @@
 #102 := [and-elim #101]: #76
 [unit-resolution #102 #115]: false
 unsat
-c79d3e134c40dc60b6991210fa65efc043a31981 146 0
+517d7705593795e4a28f7687bc92edbaaad5e6a6 146 0
 #2 := false
 decl f3 :: (-> S1 S2 S1)
 #9 := (:var 0 S2)
@@ -3426,7 +3426,7 @@
 #261 := [quant-inst]: #256
 [unit-resolution #261 #584 #284]: false
 unsat
-90a02b978126d0cf254666f4e32434ce827e91f3 146 2
+f7796908cca3a39a93067ec7456dca4edfe09069 146 2
 #2 := false
 decl f3 :: (-> S1 S2 S1)
 #9 := (:var 0 S2)
@@ -3575,7 +3575,7 @@
 unsat
 WARNING: failed to find a pattern for quantifier (quantifier id: k!13)
 
-f01333d944f2f80d82a424a44d896dcad3d2ab6b 91 0
+3d560cfeafa07e8fb4e5aedbf244ed81193a1bdd 91 0
 #2 := false
 decl f3 :: (-> S2 S1)
 decl f4 :: S2
@@ -3667,7 +3667,7 @@
 #162 := [quant-inst]: #248
 [unit-resolution #162 #582 #79]: false
 unsat
-63e9e9c72f5ef2d46bc28af44941977dc976d3b5 17 0
+6f4135b5f443c9489ab6f839e6371eb6c05021a1 17 0
 #2 := false
 #8 := 3::int
 #9 := (= 3::int 3::int)
@@ -3685,7 +3685,7 @@
 #27 := [asserted]: #10
 [mp #27 #36]: false
 unsat
-711fabeeccec1b44c93a2cd4391df71f5fdafd74 17 0
+1fb88686c084cbd29d5379055cf31c165231f0d2 17 0
 #2 := false
 #8 := 3::real
 #9 := (= 3::real 3::real)
@@ -3703,7 +3703,7 @@
 #27 := [asserted]: #10
 [mp #27 #36]: false
 unsat
-eb7a82422e5eb210ae6e04c9f7170b09a67998d0 26 0
+4439cb9b553ba6a7d3d7adaf0bccb7b403389495 26 0
 #2 := false
 #11 := 4::int
 #9 := 1::int
@@ -3730,7 +3730,7 @@
 #30 := [asserted]: #13
 [mp #30 #45]: false
 unsat
-5764777c5ab04100aa6409067059f42abecddd25 41 0
+f4b4baf086c1953b46b6c8516c9a7e45f6c086ee 41 0
 #2 := false
 decl f3 :: int
 #8 := f3
@@ -3772,7 +3772,7 @@
 #33 := [asserted]: #16
 [mp #33 #57]: false
 unsat
-f10e15cb480fc7147b50d697554b03e8d45998ff 35 0
+5926d4083f7d2a89c57531d8e87a5d554a39421a 35 0
 #2 := false
 #9 := 3::int
 #10 := 8::int
@@ -3808,7 +3808,7 @@
 #31 := [asserted]: #14
 [mp #31 #56]: false
 unsat
-85e154b43f9eaa0aee17c7df60f7b297367e646f 209 0
+53ebe62bcea7dc90158b3082dda2089fed123703 209 0
 #2 := false
 #11 := 0::real
 decl f4 :: real
@@ -4018,7 +4018,7 @@
 #231 := [unit-resolution #159 #230]: #150
 [th-lemma #125 #220 #227 #231]: false
 unsat
-c7afd26a7e4c735b3f9da1bcf41b42fb4f611502 42 0
+cbc9c5a04b9570a386fb97f8ff4491630835be54 42 0
 #2 := false
 decl f3 :: (-> S1 S2)
 decl f1 :: S1
@@ -4061,7 +4061,7 @@
 #32 := [asserted]: #15
 [mp #32 #58]: false
 unsat
-b1921561d90efa4396e2195f6befedbbc08c3f93 54 0
+732b399d3ae1d48d68db7a2b4fb14062c7ea7650 54 0
 #2 := false
 #13 := 1::int
 decl f3 :: int
@@ -4116,7 +4116,7 @@
 #33 := [asserted]: #16
 [mp #33 #72]: false
 unsat
-186b71f451f1ced8ef7a2af2f7333861c89cddbd 63 0
+a511ac5eff75bd85fb28de979615bdeb27073e2f 63 0
 #2 := false
 #15 := 0::int
 decl f4 :: int
@@ -4180,7 +4180,7 @@
 #81 := [mp #57 #80]: #68
 [mp #81 #93]: false
 unsat
-ed9f5b4ddb34caba3e930028ab3abbea1141767c 35 0
+5e99525435673f50d9e2cedb056a3eef3ca48d0f 35 0
 #2 := false
 #10 := 5::int
 #8 := 2::int
@@ -4216,7 +4216,7 @@
 #30 := [asserted]: #13
 [mp #30 #54]: false
 unsat
-e6971c89ab6428863738be19afcd7e6157397835 45 0
+cb1ff2dbd74d2844a905d03c2bb5dfdcef7b3703 45 0
 #2 := false
 #15 := 4::real
 decl f4 :: real
@@ -4262,7 +4262,7 @@
 #65 := [mp #41 #64]: #56
 [th-lemma #65 #52 #43]: false
 unsat
-4bafaa4fbebd6a7c5c47a5e807b07fb360f1fe97 59 0
+7357ea99a84507698d605d047a8912506786b005 59 0
 #2 := false
 #20 := (not false)
 decl f4 :: int
@@ -4322,7 +4322,7 @@
 #39 := [asserted]: #22
 [mp #39 #76]: false
 unsat
-c2be8bcf98f782ce6e39f901eead42a359d07f16 43 0
+a2e7fe792ce7e828740707e47e962b9013bfaa8c 43 0
 #2 := false
 decl f3 :: int
 #8 := f3
@@ -4366,7 +4366,7 @@
 #31 := [asserted]: #14
 [mp #31 #62]: false
 unsat
-f13f6152ed2127a9ebc1d2b9377e88e7270d2645 86 0
+f746fa8e8060fa7d6624d49afb6c22d0c2467190 86 0
 #2 := false
 decl f3 :: int
 #9 := f3
@@ -4453,7 +4453,7 @@
 #104 := [unit-resolution #82 #103 #96 #87]: #65
 [unit-resolution #104 #117]: false
 unsat
-faa44c5ec66ed5260e47611608af7290f07ca77d 551 0
+96de0a1de3d9897dfb99656d5dd83bfdb80fa75e 551 0
 #2 := false
 #174 := 0::int
 decl f4 :: int
@@ -5005,7 +5005,7 @@
 #617 := [unit-resolution #411 #614]: #421
 [unit-resolution #601 #617 #593 #616]: false
 unsat
-4dd2cecd4113f0eaf15e733b006ad0e01202872f 2109 0
+be4350e057381018abf4a5df1b96e44aa5440970 2109 0
 #2 := false
 #10 := 0::int
 decl f3 :: int
@@ -7115,7 +7115,7 @@
 #2119 := [unit-resolution #2118 #2109 #2056 #2114 #2113 #2112 #2096 #2108]: #945
 [unit-resolution #1969 #2119 #2109 #2096 #2042 #2105 #2108]: false
 unsat
-91a72ccddaee78c8eef27b68b8a29a516e5fa812 52 0
+9371d51a82e0a1300620d46b4e7e3f2e34923066 52 0
 #2 := false
 #12 := 1::real
 decl f3 :: real
@@ -7168,7 +7168,7 @@
 #34 := [asserted]: #17
 [mp #34 #70]: false
 unsat
-fcbcf3a00f1d2be93e0056636348f79e12052409 348 0
+7d808a299ec3aa6543955e0b52da2db4f14e2ebc 348 0
 #2 := false
 #11 := 0::int
 decl f5 :: int
@@ -7517,7 +7517,7 @@
 #638 := [unit-resolution #642 #647]: #643
 [unit-resolution #638 #640 #637]: false
 unsat
-3ff5c7aca258ce7606adbaf5b380622b5f987775 354 0
+a58c3bad13ae6ef8848c31ae2ecb4d141ea9bdc6 354 0
 #2 := false
 #11 := 0::int
 decl f5 :: int
@@ -7872,7 +7872,7 @@
 #626 := [unit-resolution #621 #634]: #753
 [unit-resolution #626 #619 #508]: false
 unsat
-37b2210a0af8b22facfab20f06af5ae3c558113d 101 0
+cfd6420dc9709ceb6bf8be602fb4605c1ddb5222 101 0
 #2 := false
 #9 := 0::real
 decl f3 :: real
@@ -7974,7 +7974,7 @@
 #120 := [th-lemma]: #119
 [unit-resolution #120 #116 #109 #43]: false
 unsat
-70f2c00c05f2849d6805e35842428b4464a986bf 943 0
+957b603c67e0eb51df6576d5eeddcc6aa6ccd469 943 0
 #2 := false
 #49 := 1::int
 decl f4 :: (-> int int int)
@@ -8918,7 +8918,7 @@
 #1294 := [unit-resolution #1107 #1293 #1239]: #405
 [unit-resolution #1252 #1294 #1292]: false
 unsat
-659511229cba86eafd88b265af83c150a34ef866 24 0
+e0bd8674b0a8871e3484c1aaf5e9760d6cdde8fa 24 0
 #2 := false
 #8 := (exists (vars (?v0 int)) false)
 #9 := (not #8)
@@ -8943,7 +8943,7 @@
 #27 := [asserted]: #10
 [mp #27 #43]: false
 unsat
-9315f166813c097237f939d5bfec03336d48d6a9 24 0
+20e7b35733689b5c0fe416fc2b22bfe44b252442 24 0
 #2 := false
 #8 := (exists (vars (?v0 real)) false)
 #9 := (not #8)
@@ -8974,7 +8974,7 @@
 unsat
 0e0dcf19cd5b120a9f262cfa9c01b662a315ee94 1 0
 unsat
-2336c05cc0db64fbd723579d41aafc2ab55fbe93 67 0
+1781c1a480d6415ddb235a621259ce2092732ba5 67 0
 #2 := false
 #9 := 0::int
 #12 := 1::int
@@ -9042,7 +9042,7 @@
 #143 := [trans #141 #80]: #142
 [mp #143 #145]: false
 unsat
-5efd84778495b1b150aa0b9f4a0b4263b1687cbc 82 0
+cb54bc1c8221ce33a204ef4756c350ab3a9165fd 82 0
 #2 := false
 #9 := (:var 0 int)
 #11 := 0::int
@@ -9125,7 +9125,7 @@
 #35 := [asserted]: #18
 [mp #35 #101]: false
 unsat
-7f5cf2a606c1f6fc9684f5e78e16538004b7482b 78 0
+05e26467c63b2926bb1206f4d9498749ac8572cd 78 0
 #2 := false
 #9 := (:var 0 int)
 #11 := 2::int
@@ -9204,7 +9204,7 @@
 #36 := [asserted]: #19
 [mp #36 #97]: false
 unsat
-2a13654aa992c8bed90af2572886d27708a51b8d 61 0
+9a25b8cd095131859f14e6e15166f75f654f8b36 61 0
 #2 := false
 #13 := (:var 0 int)
 #8 := 2::int
@@ -9266,7 +9266,7 @@
 #35 := [asserted]: #18
 [mp #35 #80]: false
 unsat
-00a64958e6476c1107006d163ef1b97bd464c3ab 111 0
+b845be4ad91437fd7de4e1e43e09de89dff87060 111 0
 #2 := false
 #8 := 2::int
 decl ?v0!1 :: int
@@ -9378,7 +9378,7 @@
 #189 := [th-lemma]: #188
 [unit-resolution #189 #132 #130 #131]: false
 unsat
-3f82af9f48300cb861109298f1c808dfeedc27df 89 0
+d516d2245b3e1198319c778f5d87f60364d9480b 89 0
 #2 := false
 #8 := 0::int
 decl ?v0!0 :: int
@@ -9468,7 +9468,7 @@
 #172 := [unit-resolution #157 #95]: #171
 [unit-resolution #172 #170 #167]: false
 unsat
-acea26b9fcd9353bdce76a59fb81f2d230c19002 83 2
+16973ba3459b683ba763cfd9d1b131b50880d52b 83 2
 #2 := false
 #9 := 0::int
 #8 := (:var 0 int)
@@ -9554,7 +9554,7 @@
 unsat
 WARNING: failed to find a pattern for quantifier (quantifier id: k!9)
 
-baa910cd7ead809257c9ce1201462fe467c60aef 180 2
+6a2477ce7e396f5a6b8350494c34556e3187382b 180 2
 #2 := false
 #16 := 3::int
 #37 := -1::int
@@ -9737,7 +9737,7 @@
 unsat
 WARNING: failed to find a pattern for quantifier (quantifier id: k!9)
 
-7a877ca9edb4f2c73b1ef0625cc9c1f31c342b91 68 0
+2c2ba5876a1ae5b4aacba3ec088a92e8517d1db5 68 0
 #2 := false
 #16 := 1::int
 #13 := (:var 1 int)
@@ -9806,7 +9806,7 @@
 #37 := [asserted]: #20
 [mp #37 #88]: false
 unsat
-c8d9f4511ee6959de9f02b091245c0b85d6b9a13 107 0
+3bc68a4890101c78401460850d978a5b784ef891 107 0
 #2 := false
 #8 := 0::int
 decl ?v1!1 :: int
@@ -9914,7 +9914,7 @@
 #128 := [and-elim #106]: #93
 [th-lemma #128 #129 #130]: false
 unsat
-64c53f605d405ecd0170c0241c2194ef25f17257 117 0
+69398382e4236c1f3e9ea09ef96d5056d4f3fa6e 117 0
 #2 := false
 #8 := 0::int
 decl ?v1!1 :: int
@@ -10032,7 +10032,7 @@
 #193 := [unit-resolution #192 #115]: #103
 [unit-resolution #193 #135]: false
 unsat
-7bc9510644b978b75cac18b8fb186f89e0d2c0fd 148 0
+4e8319ba61c4c37d1677e9a262b47203f2bf5468 148 0
 #2 := false
 #149 := (not false)
 #11 := 0::int
@@ -10181,7 +10181,7 @@
 #163 := [mp #131 #162]: #158
 [mp #163 #186]: false
 unsat
-23471f6b1ed12f9477f5b746b57d0271e3fe26ad 67 0
+a21b19aa84de923d8a0d2dc0af8016cd64248558 67 0
 #2 := false
 #8 := (:var 0 int)
 #9 := (pattern #8)
@@ -10251,7 +10251,7 @@
 unsat
 53ac944328a76df2270bb4117390a226d92d4e8d 1 0
 unsat
-5645c40190078bc259023463ba288394a1ee2d31 75 0
+80d97536e18b652f6d95d83165f7e113fa8206dd 75 0
 #2 := false
 #10 := 1::int
 decl f5 :: int
@@ -10327,7 +10327,7 @@
 #37 := [asserted]: #20
 [mp #37 #91]: false
 unsat
-dfe5809b750cf5dfc5ba8d902ad5002b8123814d 62 0
+1290567b7974f50913459a9370404d1fb27c5c04 62 0
 #2 := false
 decl f4 :: real
 #10 := f4
@@ -10390,7 +10390,7 @@
 #37 := [asserted]: #20
 [mp #37 #79]: false
 unsat
-3845bb98c4dd099de3474412f61c35c8b79e798c 141 0
+539e99c284f7a3ed7bb350da08f51bfa73f9e832 141 0
 #2 := false
 decl f6 :: int
 #13 := f6
@@ -10532,7 +10532,7 @@
 #50 := [asserted]: #33
 [mp #50 #155]: false
 unsat
-522e7a274d393467135232ebd14a729963459539 252 0
+95f708529404688f5c32519f5db83b1972672eb0 252 0
 #2 := false
 #13 := 0::int
 decl f4 :: (-> S2 int)
@@ -10785,7 +10785,7 @@
 #544 := [unit-resolution #537 #456]: #562
 [th-lemma #460 #544 #542 #551]: false
 unsat
-4d15f891322a8916dc12e3586edb994e590a9767 227 0
+069a886d4e32ad09f893e803be1b967f3568fe31 227 0
 #2 := false
 #27 := 3::int
 decl f4 :: (-> S2 int)
@@ -11013,7 +11013,7 @@
 #607 := [unit-resolution #602 #605]: #632
 [th-lemma #161 #607 #160]: false
 unsat
-ee8d523b01e5668527c65aa2ae3489e22daf3c31 367 0
+dc0cfb9f55d790de5cad944289080b012df32109 367 0
 #2 := false
 #13 := 0::int
 decl f4 :: (-> S2 int)
@@ -11381,7 +11381,7 @@
 #461 := [th-lemma]: #460
 [unit-resolution #461 #469 #457]: false
 unsat
-fe0d866a3da682e02aab50789c1002badb2ec0c9 299 0
+1119baa29d78de9d8299759a1e2778fd64cb45ac 299 0
 #2 := false
 #13 := 0::int
 decl f4 :: (-> S2 int)
@@ -11681,7 +11681,7 @@
 #603 := [unit-resolution #617 #615]: #685
 [th-lemma #625 #190 #603 #630]: false
 unsat
-0be1401474e291c6d6933a302a6418a34a233b36 458 0
+bc26a0abfecc650600742f8182ea35bd8e9c0657 458 0
 #2 := false
 #13 := 0::int
 decl f4 :: (-> S2 int)
@@ -12140,7 +12140,7 @@
 #355 := [unit-resolution #374 #372]: #373
 [unit-resolution #355 #328]: false
 unsat
-75dd8d1e40b17d4d494dfbbca0b51228568d2c11 161 0
+d28cb53d0b19f6ffbd9f6906e69a9e4028ab5365 161 0
 #2 := false
 #13 := 0::int
 decl f5 :: int
@@ -12302,7 +12302,7 @@
 #366 := [unit-resolution #644 #660]: #652
 [th-lemma #661 #366 #266]: false
 unsat
-c37debf59e8d6c8b243d5aaa6c9491c65a29d492 604 0
+3bdd721710ce740508c48327f8e9174f1a8cdb9c 604 0
 #2 := false
 #13 := 0::int
 decl f4 :: (-> S2 int)
@@ -12907,7 +12907,7 @@
 #1005 := [unit-resolution #599 #859]: #455
 [unit-resolution #1005 #1004 #994]: false
 unsat
-8add4983d79a2879b8718ea3b1ded6de7ef7d64c 61 0
+9702672d0e4ef4f07fbabb8ea1413e6a7036c6b3 61 0
 #2 := false
 decl f7 :: S3
 #12 := f7
@@ -12969,7 +12969,7 @@
 #59 := [not-or-elim #58]: #57
 [unit-resolution #59 #251]: false
 unsat
-6a6cf93b65ddcaf1b125eb57f4a626298f29e47d 116 0
+eb96c266a8e3d434df367635dd7c28e36177b61a 116 0
 #2 := false
 decl f10 :: (-> S5 S3)
 decl f7 :: S5
@@ -13086,7 +13086,7 @@
 #81 := [not-or-elim #78]: #80
 [unit-resolution #81 #612]: false
 unsat
-c4ff195222fb458c30d04a5dbddcd6f66862822e 192 0
+9aa5c4fb11199431a576a38868b15130e8351add 192 0
 #2 := false
 decl f6 :: (-> S3 S2 S4)
 decl f3 :: S2
@@ -13279,7 +13279,7 @@
 #73 := [not-or-elim #70]: #72
 [unit-resolution #73 #505]: false
 unsat
-2dc4b9dba54c10e431adb8ebd743c805e8784c37 80 0
+25babfdee044def80821e354fcd8abca6147a0b0 80 0
 #2 := false
 decl f6 :: (-> S2 S3 S1)
 decl f5 :: S3
@@ -13360,7 +13360,7 @@
 #75 := [not-or-elim #69]: #74
 [mp #75 #96]: false
 unsat
-39345d35925f930054a812342385e19283e673d4 128 0
+608cd0071a67598a13af0858c9e1341886caf8cc 128 0
 #2 := false
 decl f3 :: (-> S2 S2)
 #15 := (:var 0 S2)
@@ -13489,7 +13489,7 @@
 #224 := [quant-inst]: #566
 [unit-resolution #224 #575 #583]: false
 unsat
-c66a2580a73c3d496a195fadc270ed658f997968 469 0
+737f3f173e29756f346a0a821d7647d108111546 469 0
 #2 := false
 decl f8 :: (-> S2 S4 S4)
 decl f9 :: S4
@@ -13959,7 +13959,7 @@
 #173 := [asserted]: #43
 [unit-resolution #173 #814]: false
 unsat
-4cc658eed58c39b67c9bb50c6dd977ba3d6d31a0 38 0
+5cfea173c08e120333c1dea606b9463cfcc6450a 38 0
 #2 := false
 decl f1 :: S1
 #4 := f1
@@ -13998,7 +13998,7 @@
 #31 := [asserted]: #14
 [mp #31 #54]: false
 unsat
-8134ded6ddb01de87deb72c99e279ab6f9bebab8 366 0
+9cf33fdca403c6c514f1cd0e17b6b0a250477095 366 0
 #2 := false
 decl f3 :: (-> int S2)
 #41 := 6::int
@@ -14365,7 +14365,7 @@
 #187 := [asserted]: #44
 [unit-resolution #187 #404]: false
 unsat
-2a228f77dfa0e1f9c335db511955697c625e5b3e 1271 0
+f72e33edaef326a0ed844a14ae54b706945ff4d5 1271 0
 #2 := false
 decl f6 :: (-> int int int)
 #12 := 2::int
@@ -15637,7 +15637,7 @@
 #615 := [unit-resolution #638 #1762]: #367
 [unit-resolution #615 #1764]: false
 unsat
-5234c113c47a34efd714baca9777f6d9bb4fe31e 76 0
+59b518b007738eb00dbca158af8db9e067af1dee 76 0
 #2 := false
 decl f3 :: (-> int S1)
 #12 := (:var 0 int)
@@ -15714,7 +15714,7 @@
 #689 := [quant-inst]: #688
 [unit-resolution #689 #158 #1022]: false
 unsat
-b9f47e2d664572f569f3105a25d0e77e894b5cdd 478 0
+997a0c6c7b0815f9412f3484edb90c328a2f5296 478 0
 #2 := false
 decl f7 :: (-> S4 S2)
 decl f8 :: (-> S1 S4)
--- a/src/HOL/SMT_Examples/SMT_Examples.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/SMT_Examples/SMT_Examples.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -8,7 +8,7 @@
 imports Complex_Main
 begin
 
-declare [[smt_solver=z3, z3_proofs=true]]
+declare [[smt_solver=z3, smt_oracle=false]]
 declare [[smt_certificates="~~/src/HOL/SMT_Examples/SMT_Examples.certs"]]
 declare [[smt_fixed=true]]
 
@@ -25,7 +25,7 @@
 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)"
-  using [[z3_proofs=false]] (* no Z3 proof *)
+  using [[smt_oracle=true]] (* no Z3 proof *)
   by smt
 
 lemma "(p1 \<and> p2) \<or> p3 \<longrightarrow> (p1 \<longrightarrow> (p3 \<and> p2) \<or> (p1 \<and> p3)) \<or> p1" by smt
@@ -358,15 +358,15 @@
 lemma "~ (\<exists>x::real. False)" by smt
 
 lemma "\<exists>x::int. 0 < x"
-  using [[z3_proofs=false]] (* no Z3 proof *)
+  using [[smt_oracle=true]] (* no Z3 proof *)
   by smt
 
 lemma "\<exists>x::real. 0 < x"
-  using [[z3_proofs=false]] (* no Z3 proof *)
+  using [[smt_oracle=true]] (* no Z3 proof *)
   by smt
 
 lemma "\<forall>x::int. \<exists>y. y > x"
-  using [[z3_proofs=false]] (* no Z3 proof *)
+  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
@@ -400,7 +400,7 @@
 subsection {* Non-linear arithmetic over integers and reals *}
 
 lemma "a > (0::int) \<Longrightarrow> a*b > 0 \<Longrightarrow> b > 0"
-  using [[z3_proofs=false]]  -- {* Isabelle's arithmetic decision procedures
+  using [[smt_oracle=true]]  -- {* Isabelle's arithmetic decision procedures
     are too weak to automatically prove @{thm zero_less_mult_pos}. *}
   by smt (* FIXME: use z3_rule *)
 
--- a/src/HOL/SMT_Examples/SMT_Tests.certs	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/SMT_Examples/SMT_Tests.certs	Wed Oct 27 11:11:35 2010 -0700
@@ -1,4 +1,4 @@
-ef3118b3c538f5cd123d726ccf13a3c2d3725bb0 8 0
+d4b1481d70f20d4d6280999633c905e21d180d13 8 0
 #2 := false
 #1 := true
 #8 := (not true)
@@ -7,7 +7,7 @@
 #25 := [asserted]: #8
 [mp #25 #27]: false
 unsat
-300ac981df663a0c38577b3d315e53c71dd0988e 16 0
+ac5b1925e47f9c59b180fdccd321c68a4a01a710 16 0
 #2 := false
 #8 := (not false)
 #9 := (not #8)
@@ -24,7 +24,7 @@
 #26 := [asserted]: #9
 [mp #26 #35]: false
 unsat
-0b184a84ce871fe108c2b5dce0e6482dd3c0450a 21 0
+b13645ce0b47b432b924b38e0bebde6cbcb236f9 21 0
 #2 := false
 #1 := true
 #8 := (not true)
@@ -46,7 +46,7 @@
 #27 := [asserted]: #10
 [mp #27 #40]: false
 unsat
-9acfa8385beeafd8546dddb4d92b4b4b9a1a0bd8 16 0
+8ec60d0baade8bddd39b6e00ec153f6bd8949184 16 0
 #2 := false
 #1 := true
 #8 := (and true true)
@@ -63,7 +63,7 @@
 #26 := [asserted]: #9
 [mp #26 #35]: false
 unsat
-d35ed39a7bb1024906bec1f4afa8168bb568efe7 16 0
+b448617db1567a507fc596f5b21ade645ca6e0f9 16 0
 #2 := false
 #1 := true
 #8 := (or true false)
@@ -80,7 +80,7 @@
 #26 := [asserted]: #9
 [mp #26 #35]: false
 unsat
-a3266a8f006b3750ff4748b6a1a76e7bde92780d 16 0
+bf545ecaa58049cbf454eac70ee3077f64564cca 16 0
 #2 := false
 #1 := true
 #8 := (implies false true)
@@ -97,7 +97,7 @@
 #26 := [asserted]: #9
 [mp #26 #35]: false
 unsat
-b59b2f429ffd47c407e4ef1af006540275b1a26c 33 0
+a38979487be0aafcf2d387180e92dfd3bd43e483 33 0
 #2 := false
 decl f1 :: S1
 #4 := f1
@@ -131,7 +131,7 @@
 #29 := [asserted]: #12
 [mp #29 #49]: false
 unsat
-3115e8e7b694b406a07c11043a6921a0dba8d5ea 41 0
+b2faf72eaa234b8ac1acedfef5768f99ca3da86b 41 0
 #2 := false
 decl f1 :: S1
 #4 := f1
@@ -173,157 +173,7 @@
 #30 := [asserted]: #13
 [mp #30 #57]: false
 unsat
-e3d163795277d672f4f54561a9fed7fa5b707b37 149 0
-#2 := false
-decl f5 :: S1
-#12 := f5
-decl f1 :: S1
-#4 := f1
-#44 := (= f1 f5)
-decl f4 :: S1
-#10 := f4
-#41 := (= f1 f4)
-decl f3 :: S1
-#8 := f3
-#38 := (= f1 f3)
-#47 := (ite #38 #41 #44)
-#53 := (not #38)
-#54 := (or #53 #41)
-#64 := (or #38 #44)
-#91 := (not #64)
-#90 := (not #54)
-#92 := (or #90 #91)
-#143 := [hypothesis]: #90
-#128 := (or #92 #54)
-#129 := [def-axiom]: #128
-#144 := [unit-resolution #129 #143]: #92
-#78 := (not #47)
-#116 := (or #54 #38)
-#117 := [def-axiom]: #116
-#145 := [unit-resolution #117 #143]: #38
-#110 := (not #41)
-#118 := (or #54 #110)
-#119 := [def-axiom]: #118
-#146 := [unit-resolution #119 #143]: #110
-#106 := (or #78 #53 #41)
-#107 := [def-axiom]: #106
-#147 := [unit-resolution #107 #146 #145]: #78
-#93 := (not #92)
-#137 := (or #47 #93)
-#100 := (iff #47 #92)
-#69 := (and #54 #64)
-#79 := (iff #69 #78)
-#103 := (iff #79 #100)
-#95 := (iff #92 #47)
-#101 := (iff #95 #100)
-#102 := [rewrite]: #101
-#98 := (iff #79 #95)
-#87 := (iff #93 #78)
-#96 := (iff #87 #95)
-#97 := [rewrite]: #96
-#84 := (iff #79 #87)
-#88 := (iff #69 #93)
-#89 := [rewrite]: #88
-#94 := [monotonicity #89]: #84
-#99 := [trans #94 #97]: #98
-#104 := [trans #99 #102]: #103
-#13 := (= f5 f1)
-#9 := (= f3 f1)
-#16 := (not #9)
-#17 := (implies #16 #13)
-#11 := (= f4 f1)
-#15 := (implies #9 #11)
-#18 := (and #15 #17)
-#14 := (ite #9 #11 #13)
-#19 := (iff #14 #18)
-#20 := (not #19)
-#82 := (iff #20 #79)
-#72 := (iff #47 #69)
-#75 := (not #72)
-#80 := (iff #75 #79)
-#81 := [rewrite]: #80
-#76 := (iff #20 #75)
-#73 := (iff #19 #72)
-#70 := (iff #18 #69)
-#67 := (iff #17 #64)
-#61 := (implies #53 #44)
-#65 := (iff #61 #64)
-#66 := [rewrite]: #65
-#62 := (iff #17 #61)
-#45 := (iff #13 #44)
-#46 := [rewrite]: #45
-#59 := (iff #16 #53)
-#39 := (iff #9 #38)
-#40 := [rewrite]: #39
-#60 := [monotonicity #40]: #59
-#63 := [monotonicity #60 #46]: #62
-#68 := [trans #63 #66]: #67
-#57 := (iff #15 #54)
-#50 := (implies #38 #41)
-#55 := (iff #50 #54)
-#56 := [rewrite]: #55
-#51 := (iff #15 #50)
-#42 := (iff #11 #41)
-#43 := [rewrite]: #42
-#52 := [monotonicity #40 #43]: #51
-#58 := [trans #52 #56]: #57
-#71 := [monotonicity #58 #68]: #70
-#48 := (iff #14 #47)
-#49 := [monotonicity #40 #43 #46]: #48
-#74 := [monotonicity #49 #71]: #73
-#77 := [monotonicity #74]: #76
-#83 := [trans #77 #81]: #82
-#37 := [asserted]: #20
-#86 := [mp #37 #83]: #79
-#105 := [mp #86 #104]: #100
-#134 := (not #100)
-#135 := (or #47 #93 #134)
-#136 := [def-axiom]: #135
-#138 := [unit-resolution #136 #105]: #137
-#148 := [unit-resolution #138 #147 #144]: false
-#149 := [lemma #148]: #54
-#150 := [hypothesis]: #78
-#156 := (or #38 #47)
-#151 := [unit-resolution #138 #150]: #93
-#130 := (or #92 #64)
-#131 := [def-axiom]: #130
-#152 := [unit-resolution #131 #151]: #64
-#153 := [hypothesis]: #53
-#113 := (not #44)
-#114 := (or #47 #38 #113)
-#115 := [def-axiom]: #114
-#154 := [unit-resolution #115 #153 #150]: #113
-#126 := (or #91 #38 #44)
-#127 := [def-axiom]: #126
-#155 := [unit-resolution #127 #154 #153 #152]: false
-#157 := [lemma #155]: #156
-#158 := [unit-resolution #157 #150]: #38
-#111 := (or #47 #53 #110)
-#112 := [def-axiom]: #111
-#159 := [unit-resolution #112 #158 #150]: #110
-#120 := (or #90 #53 #41)
-#121 := [def-axiom]: #120
-#160 := [unit-resolution #121 #159 #158 #149]: false
-#161 := [lemma #160]: #47
-#141 := (or #78 #92)
-#139 := (or #78 #92 #134)
-#140 := [def-axiom]: #139
-#142 := [unit-resolution #140 #105]: #141
-#162 := [unit-resolution #142 #161]: #92
-#132 := (or #93 #90 #91)
-#133 := [def-axiom]: #132
-#163 := [unit-resolution #133 #162 #149]: #91
-#122 := (or #64 #53)
-#123 := [def-axiom]: #122
-#164 := [unit-resolution #123 #163]: #53
-#124 := (or #64 #113)
-#125 := [def-axiom]: #124
-#165 := [unit-resolution #125 #163]: #113
-#108 := (or #78 #38 #44)
-#109 := [def-axiom]: #108
-[unit-resolution #109 #165 #164 #161]: false
-unsat
-293682654769f9a30e4727176ece5ebbfa2edbf9 65 0
+fa3a2c67c2833da196c4009cb283e94a9f22829b 65 0
 #2 := false
 decl f1 :: S1
 #4 := f1
@@ -389,74 +239,7 @@
 #33 := [asserted]: #16
 [mp #33 #81]: false
 unsat
-b7f1180181cd52aa741865e3ac17b4d97cf7408a 66 0
-#2 := false
-decl f3 :: S1
-#8 := f3
-decl f1 :: S1
-#4 := f1
-#33 := (= f1 f3)
-#51 := (not #33)
-#87 := [hypothesis]: #33
-decl f4 :: S1
-#10 := f4
-#36 := (= f1 f4)
-#42 := (not #36)
-#43 := (or #33 #42)
-#69 := (or #43 #51)
-#70 := [def-axiom]: #69
-#88 := [unit-resolution #70 #87]: #43
-#67 := (not #43)
-#89 := (or #51 #67)
-#52 := (or #51 #36)
-#57 := (ite #33 #43 #52)
-#60 := (not #57)
-#11 := (= f4 f1)
-#9 := (= f3 f1)
-#13 := (implies #9 #11)
-#12 := (implies #11 #9)
-#14 := (ite #9 #12 #13)
-#15 := (not #14)
-#61 := (iff #15 #60)
-#58 := (iff #14 #57)
-#55 := (iff #13 #52)
-#48 := (implies #33 #36)
-#53 := (iff #48 #52)
-#54 := [rewrite]: #53
-#49 := (iff #13 #48)
-#37 := (iff #11 #36)
-#38 := [rewrite]: #37
-#34 := (iff #9 #33)
-#35 := [rewrite]: #34
-#50 := [monotonicity #35 #38]: #49
-#56 := [trans #50 #54]: #55
-#46 := (iff #12 #43)
-#39 := (implies #36 #33)
-#44 := (iff #39 #43)
-#45 := [rewrite]: #44
-#40 := (iff #12 #39)
-#41 := [monotonicity #38 #35]: #40
-#47 := [trans #41 #45]: #46
-#59 := [monotonicity #35 #47 #56]: #58
-#62 := [monotonicity #59]: #61
-#32 := [asserted]: #15
-#65 := [mp #32 #62]: #60
-#83 := (or #57 #51 #67)
-#84 := [def-axiom]: #83
-#90 := [unit-resolution #84 #65]: #89
-#91 := [unit-resolution #90 #88 #87]: false
-#92 := [lemma #91]: #51
-#63 := (or #52 #33)
-#73 := [def-axiom]: #63
-#93 := [unit-resolution #73 #92]: #52
-#76 := (not #52)
-#94 := (or #33 #76)
-#85 := (or #57 #33 #76)
-#86 := [def-axiom]: #85
-#95 := [unit-resolution #86 #65]: #94
-[unit-resolution #95 #93 #92]: false
-unsat
-001f5ceb25bbdf443b66d76c2d62a655a1155fa7 29 0
+85a9b542a9c6d64e778f2449f892d6d55e153689 29 0
 #2 := false
 decl f1 :: S1
 #4 := f1
@@ -486,7 +269,7 @@
 #28 := [asserted]: #11
 [mp #28 #45]: false
 unsat
-140af760eaf1680b3107622d1bad9733191529ce 41 0
+a21434fdffa6243cc7410eb0aff669b84c5be474 41 0
 #2 := false
 decl f1 :: S1
 #4 := f1
@@ -528,75 +311,7 @@
 #30 := [asserted]: #13
 [mp #30 #57]: false
 unsat
-4ef549c8b4c07c939a8ccb4f819e8b3925600dc3 67 0
-#2 := false
-decl f1 :: S1
-#4 := f1
-decl f6 :: S1
-#15 := f6
-#16 := (= f6 f1)
-decl f5 :: S1
-#13 := f5
-#14 := (= f5 f1)
-decl f4 :: S1
-#10 := f4
-#11 := (= f4 f1)
-#18 := (ite #11 #14 #16)
-decl f3 :: S1
-#8 := f3
-#9 := (= f3 f1)
-#19 := (ite #9 #14 #18)
-#12 := (or #9 #11)
-#17 := (ite #12 #14 #16)
-#20 := (iff #17 #19)
-#21 := (not #20)
-#79 := (iff #21 false)
-#1 := true
-#74 := (not true)
-#77 := (iff #74 false)
-#78 := [rewrite]: #77
-#75 := (iff #21 #74)
-#72 := (iff #20 true)
-#51 := (= f1 f6)
-#48 := (= f1 f5)
-#42 := (= f1 f4)
-#39 := (= f1 f3)
-#45 := (or #39 #42)
-#54 := (ite #45 #48 #51)
-#67 := (iff #54 #54)
-#70 := (iff #67 true)
-#71 := [rewrite]: #70
-#68 := (iff #20 #67)
-#65 := (iff #19 #54)
-#57 := (ite #42 #48 #51)
-#60 := (ite #39 #48 #57)
-#63 := (iff #60 #54)
-#64 := [rewrite]: #63
-#61 := (iff #19 #60)
-#58 := (iff #18 #57)
-#52 := (iff #16 #51)
-#53 := [rewrite]: #52
-#49 := (iff #14 #48)
-#50 := [rewrite]: #49
-#43 := (iff #11 #42)
-#44 := [rewrite]: #43
-#59 := [monotonicity #44 #50 #53]: #58
-#40 := (iff #9 #39)
-#41 := [rewrite]: #40
-#62 := [monotonicity #41 #50 #59]: #61
-#66 := [trans #62 #64]: #65
-#55 := (iff #17 #54)
-#46 := (iff #12 #45)
-#47 := [monotonicity #41 #44]: #46
-#56 := [monotonicity #47 #50 #53]: #55
-#69 := [monotonicity #56 #66]: #68
-#73 := [trans #69 #71]: #72
-#76 := [monotonicity #73]: #75
-#80 := [trans #76 #78]: #79
-#38 := [asserted]: #21
-[mp #38 #80]: false
-unsat
-6e4da81b8f15e2e0630be83dbec8d4ceab075fda 47 0
+a6a664cc4c3ff45fb3211b72c0f520208b902c9b 47 0
 #2 := false
 decl f1 :: S1
 #4 := f1
@@ -644,75 +359,7 @@
 #32 := [asserted]: #15
 [mp #32 #62]: false
 unsat
-00916919fe29a5b81ca835c8a92b8b60250f42f6 67 0
-#2 := false
-decl f1 :: S1
-#4 := f1
-decl f6 :: S1
-#15 := f6
-#16 := (= f6 f1)
-decl f5 :: S1
-#13 := f5
-#14 := (= f5 f1)
-decl f4 :: S1
-#10 := f4
-#11 := (= f4 f1)
-#18 := (ite #11 #14 #16)
-decl f3 :: S1
-#8 := f3
-#9 := (= f3 f1)
-#19 := (ite #9 #18 #16)
-#12 := (and #9 #11)
-#17 := (ite #12 #14 #16)
-#20 := (iff #17 #19)
-#21 := (not #20)
-#79 := (iff #21 false)
-#1 := true
-#74 := (not true)
-#77 := (iff #74 false)
-#78 := [rewrite]: #77
-#75 := (iff #21 #74)
-#72 := (iff #20 true)
-#51 := (= f1 f6)
-#48 := (= f1 f5)
-#42 := (= f1 f4)
-#39 := (= f1 f3)
-#45 := (and #39 #42)
-#54 := (ite #45 #48 #51)
-#67 := (iff #54 #54)
-#70 := (iff #67 true)
-#71 := [rewrite]: #70
-#68 := (iff #20 #67)
-#65 := (iff #19 #54)
-#57 := (ite #42 #48 #51)
-#60 := (ite #39 #57 #51)
-#63 := (iff #60 #54)
-#64 := [rewrite]: #63
-#61 := (iff #19 #60)
-#52 := (iff #16 #51)
-#53 := [rewrite]: #52
-#58 := (iff #18 #57)
-#49 := (iff #14 #48)
-#50 := [rewrite]: #49
-#43 := (iff #11 #42)
-#44 := [rewrite]: #43
-#59 := [monotonicity #44 #50 #53]: #58
-#40 := (iff #9 #39)
-#41 := [rewrite]: #40
-#62 := [monotonicity #41 #59 #53]: #61
-#66 := [trans #62 #64]: #65
-#55 := (iff #17 #54)
-#46 := (iff #12 #45)
-#47 := [monotonicity #41 #44]: #46
-#56 := [monotonicity #47 #50 #53]: #55
-#69 := [monotonicity #56 #66]: #68
-#73 := [trans #69 #71]: #72
-#76 := [monotonicity #73]: #75
-#80 := [trans #76 #78]: #79
-#38 := [asserted]: #21
-[mp #38 #80]: false
-unsat
-d5bc9e39a4bb8116809502fbfd19f3ca0119f9c0 47 0
+d99e41f56c960ba88eb8ca95c5a49bd2181a7626 47 0
 #2 := false
 decl f1 :: S1
 #4 := f1
@@ -760,194 +407,7 @@
 #32 := [asserted]: #15
 [mp #32 #62]: false
 unsat
-7d731ec0802cbb5844eb34ccffb568c08f773a66 186 0
-#2 := false
-decl f5 :: S1
-#12 := f5
-decl f1 :: S1
-#4 := f1
-#47 := (= f1 f5)
-decl f3 :: S1
-#8 := f3
-#41 := (= f1 f3)
-#59 := (not #41)
-#76 := (or #59 #47)
-#119 := (not #76)
-decl f4 :: S1
-#10 := f4
-#44 := (= f1 f4)
-#68 := (or #59 #44)
-decl f6 :: S1
-#14 := f6
-#50 := (= f1 f6)
-#84 := (or #59 #50)
-#89 := (ite #68 #76 #84)
-#130 := (not #89)
-#53 := (ite #44 #47 #50)
-#60 := (or #59 #53)
-#112 := (not #44)
-#165 := [hypothesis]: #112
-#172 := (or #60 #44)
-#98 := (not #60)
-#163 := [hypothesis]: #98
-#148 := (or #60 #41)
-#149 := [def-axiom]: #148
-#164 := [unit-resolution #149 #163]: #41
-#124 := (not #50)
-#139 := (not #53)
-#150 := (or #60 #139)
-#151 := [def-axiom]: #150
-#166 := [unit-resolution #151 #163]: #139
-#146 := (or #53 #44 #124)
-#147 := [def-axiom]: #146
-#167 := [unit-resolution #147 #166 #165]: #124
-#157 := (or #89 #60)
-#99 := (iff #89 #98)
-#15 := (= f6 f1)
-#9 := (= f3 f1)
-#20 := (implies #9 #15)
-#13 := (= f5 f1)
-#19 := (implies #9 #13)
-#11 := (= f4 f1)
-#18 := (implies #9 #11)
-#21 := (ite #18 #19 #20)
-#16 := (ite #11 #13 #15)
-#17 := (implies #9 #16)
-#22 := (iff #17 #21)
-#23 := (not #22)
-#102 := (iff #23 #99)
-#92 := (iff #60 #89)
-#95 := (not #92)
-#100 := (iff #95 #99)
-#101 := [rewrite]: #100
-#96 := (iff #23 #95)
-#93 := (iff #22 #92)
-#90 := (iff #21 #89)
-#87 := (iff #20 #84)
-#81 := (implies #41 #50)
-#85 := (iff #81 #84)
-#86 := [rewrite]: #85
-#82 := (iff #20 #81)
-#51 := (iff #15 #50)
-#52 := [rewrite]: #51
-#42 := (iff #9 #41)
-#43 := [rewrite]: #42
-#83 := [monotonicity #43 #52]: #82
-#88 := [trans #83 #86]: #87
-#79 := (iff #19 #76)
-#73 := (implies #41 #47)
-#77 := (iff #73 #76)
-#78 := [rewrite]: #77
-#74 := (iff #19 #73)
-#48 := (iff #13 #47)
-#49 := [rewrite]: #48
-#75 := [monotonicity #43 #49]: #74
-#80 := [trans #75 #78]: #79
-#71 := (iff #18 #68)
-#65 := (implies #41 #44)
-#69 := (iff #65 #68)
-#70 := [rewrite]: #69
-#66 := (iff #18 #65)
-#45 := (iff #11 #44)
-#46 := [rewrite]: #45
-#67 := [monotonicity #43 #46]: #66
-#72 := [trans #67 #70]: #71
-#91 := [monotonicity #72 #80 #88]: #90
-#63 := (iff #17 #60)
-#56 := (implies #41 #53)
-#61 := (iff #56 #60)
-#62 := [rewrite]: #61
-#57 := (iff #17 #56)
-#54 := (iff #16 #53)
-#55 := [monotonicity #46 #49 #52]: #54
-#58 := [monotonicity #43 #55]: #57
-#64 := [trans #58 #62]: #63
-#94 := [monotonicity #64 #91]: #93
-#97 := [monotonicity #94]: #96
-#103 := [trans #97 #101]: #102
-#40 := [asserted]: #23
-#106 := [mp #40 #103]: #99
-#154 := (not #99)
-#155 := (or #89 #60 #154)
-#156 := [def-axiom]: #155
-#158 := [unit-resolution #156 #106]: #157
-#168 := [unit-resolution #158 #163]: #89
-#109 := (not #68)
-#107 := (or #109 #59 #44)
-#104 := [def-axiom]: #107
-#169 := [unit-resolution #104 #164 #165]: #109
-#133 := (or #130 #68 #84)
-#134 := [def-axiom]: #133
-#170 := [unit-resolution #134 #169 #168]: #84
-#127 := (not #84)
-#128 := (or #127 #59 #50)
-#129 := [def-axiom]: #128
-#171 := [unit-resolution #129 #170 #167 #164]: false
-#173 := [lemma #171]: #172
-#176 := [unit-resolution #173 #165]: #60
-#161 := (or #130 #98)
-#159 := (or #130 #98 #154)
-#160 := [def-axiom]: #159
-#162 := [unit-resolution #160 #106]: #161
-#182 := [unit-resolution #162 #176]: #130
-#180 := (or #84 #44)
-#174 := [hypothesis]: #127
-#125 := (or #84 #124)
-#126 := [def-axiom]: #125
-#175 := [unit-resolution #126 #174]: #124
-#122 := (or #84 #41)
-#123 := [def-axiom]: #122
-#177 := [unit-resolution #123 #174]: #41
-#152 := (or #98 #59 #53)
-#153 := [def-axiom]: #152
-#178 := [unit-resolution #153 #177 #176]: #53
-#142 := (or #139 #44 #50)
-#143 := [def-axiom]: #142
-#179 := [unit-resolution #143 #178 #175 #165]: false
-#181 := [lemma #179]: #180
-#183 := [unit-resolution #181 #165]: #84
-#137 := (or #89 #68 #127)
-#138 := [def-axiom]: #137
-#184 := [unit-resolution #138 #183 #182]: #68
-#135 := (or #89 #109 #119)
-#136 := [def-axiom]: #135
-#185 := [unit-resolution #136 #184 #182]: #119
-#186 := [unit-resolution #104 #184 #165]: #59
-#114 := (or #76 #41)
-#115 := [def-axiom]: #114
-#187 := [unit-resolution #115 #186 #185]: false
-#188 := [lemma #187]: #44
-#113 := (or #68 #112)
-#108 := [def-axiom]: #113
-#191 := [unit-resolution #108 #188]: #68
-#189 := [hypothesis]: #59
-#190 := [unit-resolution #149 #189]: #60
-#192 := [unit-resolution #115 #189]: #76
-#193 := [unit-resolution #136 #192 #191]: #89
-#194 := [unit-resolution #162 #193 #190]: false
-#195 := [lemma #194]: #41
-#116 := (not #47)
-#144 := (or #53 #112 #116)
-#145 := [def-axiom]: #144
-#196 := [unit-resolution #145 #166 #188]: #116
-#131 := (or #130 #109 #76)
-#132 := [def-axiom]: #131
-#197 := [unit-resolution #132 #168 #191]: #76
-#120 := (or #119 #59 #47)
-#121 := [def-axiom]: #120
-#198 := [unit-resolution #121 #197 #196 #195]: false
-#199 := [lemma #198]: #60
-#200 := [unit-resolution #162 #199]: #130
-#201 := [unit-resolution #136 #200 #191]: #119
-#202 := [unit-resolution #153 #199 #195]: #53
-#140 := (or #139 #112 #47)
-#141 := [def-axiom]: #140
-#203 := [unit-resolution #141 #202 #188]: #47
-#117 := (or #76 #116)
-#118 := [def-axiom]: #117
-[unit-resolution #118 #203 #201]: false
-unsat
-8f2ba8a450b38e7bdf2f188b8cd61ae5399f3de8 55 0
+a8e5086b6e21c7bf961a22426cb7ae14847089d7 55 0
 #2 := false
 decl f3 :: S1
 #8 := f3
@@ -1003,41 +463,7 @@
 #71 := [and-elim #70]: #33
 [mp #71 #77]: false
 unsat
-4c5797ba6bf6e1b3bb6b8306b16e65642fb4c91c 33 0
-#2 := false
-decl f1 :: S1
-#4 := f1
-decl f3 :: S1
-#8 := f3
-#9 := (= f3 f1)
-#10 := (not #9)
-#11 := (ite #9 #9 #10)
-#12 := (not #11)
-#48 := (iff #12 false)
-#1 := true
-#43 := (not true)
-#46 := (iff #43 false)
-#47 := [rewrite]: #46
-#44 := (iff #12 #43)
-#41 := (iff #11 true)
-#30 := (= f1 f3)
-#33 := (not #30)
-#36 := (ite #30 #30 #33)
-#39 := (iff #36 true)
-#40 := [rewrite]: #39
-#37 := (iff #11 #36)
-#34 := (iff #10 #33)
-#31 := (iff #9 #30)
-#32 := [rewrite]: #31
-#35 := [monotonicity #32]: #34
-#38 := [monotonicity #32 #32 #35]: #37
-#42 := [trans #38 #40]: #41
-#45 := [monotonicity #42]: #44
-#49 := [trans #45 #47]: #48
-#29 := [asserted]: #12
-[mp #29 #49]: false
-unsat
-97918cca48c9094bd5b3664f47e49327114322de 55 0
+e2d0843914ae918fb9bd06a2d6949232327217fa 55 0
 #2 := false
 decl f1 :: S1
 #4 := f1
@@ -1093,41 +519,7 @@
 #33 := [asserted]: #16
 [mp #33 #70]: false
 unsat
-9bd9eaea49f2e1ee09a5710de988a261a59b3161 33 0
-#2 := false
-decl f1 :: S1
-#4 := f1
-decl f3 :: S1
-#8 := f3
-#9 := (= f3 f1)
-#10 := (not #9)
-#11 := (ite #10 #10 #9)
-#12 := (not #11)
-#48 := (iff #12 false)
-#1 := true
-#43 := (not true)
-#46 := (iff #43 false)
-#47 := [rewrite]: #46
-#44 := (iff #12 #43)
-#41 := (iff #11 true)
-#30 := (= f1 f3)
-#33 := (not #30)
-#36 := (ite #33 #33 #30)
-#39 := (iff #36 true)
-#40 := [rewrite]: #39
-#37 := (iff #11 #36)
-#31 := (iff #9 #30)
-#32 := [rewrite]: #31
-#34 := (iff #10 #33)
-#35 := [monotonicity #32]: #34
-#38 := [monotonicity #35 #35 #32]: #37
-#42 := [trans #38 #40]: #41
-#45 := [monotonicity #42]: #44
-#49 := [trans #45 #47]: #48
-#29 := [asserted]: #12
-[mp #29 #49]: false
-unsat
-54241528b45f8f4faa0c0db0443b55a0791bdebf 55 0
+00d9c16458cf2825e776c32505a513a412ede70f 55 0
 #2 := false
 decl f1 :: S1
 #4 := f1
@@ -1183,26 +575,7 @@
 #33 := [asserted]: #16
 [mp #33 #70]: false
 unsat
-7e9279650e1782ee3eacd8b5ddc9abff7cf3939f 18 0
-#2 := false
-decl f3 :: S2
-#8 := f3
-#9 := (= f3 f3)
-#10 := (not #9)
-#35 := (iff #10 false)
-#1 := true
-#30 := (not true)
-#33 := (iff #30 false)
-#34 := [rewrite]: #33
-#31 := (iff #10 #30)
-#28 := (iff #9 true)
-#29 := [rewrite]: #28
-#32 := [monotonicity #29]: #31
-#36 := [trans #32 #34]: #35
-#27 := [asserted]: #10
-[mp #27 #36]: false
-unsat
-fad369d322dc0c4cbd802fd43ac8e39269b90c24 49 0
+bca31cbd8bb6114896c32a11dfae7025f2fc07b1 49 0
 #2 := false
 decl f3 :: S1
 #8 := f3
@@ -1252,37 +625,7 @@
 #63 := [and-elim #61]: #34
 [mp #63 #69]: false
 unsat
-4f15136164606c4b4de495e81f7a15d214b1a16d 29 0
-#2 := false
-decl f3 :: S2
-#8 := f3
-decl f4 :: S2
-#9 := f4
-#11 := (= f4 f3)
-#10 := (= f3 f4)
-#12 := (implies #10 #11)
-#13 := (not #12)
-#45 := (iff #13 false)
-#1 := true
-#40 := (not true)
-#43 := (iff #40 false)
-#44 := [rewrite]: #43
-#41 := (iff #13 #40)
-#38 := (iff #12 true)
-#33 := (implies #10 #10)
-#36 := (iff #33 true)
-#37 := [rewrite]: #36
-#34 := (iff #12 #33)
-#31 := (iff #11 #10)
-#32 := [rewrite]: #31
-#35 := [monotonicity #32]: #34
-#39 := [trans #35 #37]: #38
-#42 := [monotonicity #39]: #41
-#46 := [trans #42 #44]: #45
-#30 := [asserted]: #13
-[mp #30 #46]: false
-unsat
-b24d83f0e79204c7c5bc93564e0299d7be355898 54 0
+f5031e01f448cd128a3cce562b470d23d1a8bc59 54 0
 #2 := false
 decl f3 :: S1
 #8 := f3
@@ -1337,42 +680,7 @@
 #68 := [and-elim #66]: #34
 [mp #68 #74]: false
 unsat
-dc9c622e9c5a73b4906108c844d17870f8f467f3 34 0
-#2 := false
-decl f5 :: S2
-#11 := f5
-decl f3 :: S2
-#8 := f3
-#14 := (= f3 f5)
-decl f4 :: S2
-#9 := f4
-#12 := (= f4 f5)
-#58 := (iff #12 #14)
-#56 := (iff #14 #12)
-#10 := (= f3 f4)
-#13 := (and #10 #12)
-#34 := (not #13)
-#35 := (or #34 #14)
-#38 := (not #35)
-#15 := (implies #13 #14)
-#16 := (not #15)
-#39 := (iff #16 #38)
-#36 := (iff #15 #35)
-#37 := [rewrite]: #36
-#40 := [monotonicity #37]: #39
-#33 := [asserted]: #16
-#43 := [mp #33 #40]: #38
-#41 := [not-or-elim #43]: #13
-#42 := [and-elim #41]: #10
-#57 := [monotonicity #42]: #56
-#59 := [symm #57]: #58
-#44 := [and-elim #41]: #12
-#54 := [mp #44 #59]: #14
-#45 := (not #14)
-#46 := [not-or-elim #43]: #45
-[unit-resolution #46 #54]: false
-unsat
-0f759d4f64321dbbf3a7fb1265c4748ce0ceab54 60 0
+38d7d5f12a8c933e456bba4a04c7645f4c87ea49 60 0
 #2 := false
 decl f4 :: S1
 #10 := f4
@@ -1433,35 +741,7 @@
 #70 := [mp #35 #67]: #64
 [mp #70 #68]: false
 unsat
-a8195f22c8d1b577a8f5ae2243196c5ebe015804 27 0
-#2 := false
-decl f5 :: (-> S2 S2)
-decl f4 :: S2
-#9 := f4
-#12 := (f5 f4)
-decl f3 :: S2
-#8 := f3
-#11 := (f5 f3)
-#13 := (= #11 #12)
-#10 := (= f3 f4)
-#33 := (not #10)
-#34 := (or #33 #13)
-#37 := (not #34)
-#14 := (implies #10 #13)
-#15 := (not #14)
-#38 := (iff #15 #37)
-#35 := (iff #14 #34)
-#36 := [rewrite]: #35
-#39 := [monotonicity #36]: #38
-#32 := [asserted]: #15
-#42 := [mp #32 #39]: #37
-#40 := [not-or-elim #42]: #10
-#51 := [monotonicity #40]: #13
-#41 := (not #13)
-#43 := [not-or-elim #42]: #41
-[unit-resolution #43 #51]: false
-unsat
-56fbd97995f4b7bc4d7561f89b146084e2ece4a4 79 0
+27121de8ceb5985c14cca2c4563e9cad4d2ebd3c 79 0
 #2 := false
 decl f1 :: S1
 #4 := f1
@@ -1541,39 +821,7 @@
 #36 := [asserted]: #19
 [mp #36 #93]: false
 unsat
-f0071ffac5f0293e3340ba2c318e33d4b7b540a9 31 0
-#2 := false
-decl f5 :: (-> S2 S2 S3)
-decl f3 :: S2
-#8 := f3
-decl f4 :: S2
-#9 := f4
-#12 := (f5 f4 f3)
-#11 := (f5 f3 f4)
-#13 := (= #11 #12)
-#53 := (= #12 #11)
-#10 := (= f3 f4)
-#33 := (not #10)
-#34 := (or #33 #13)
-#37 := (not #34)
-#14 := (implies #10 #13)
-#15 := (not #14)
-#38 := (iff #15 #37)
-#35 := (iff #14 #34)
-#36 := [rewrite]: #35
-#39 := [monotonicity #36]: #38
-#32 := [asserted]: #15
-#42 := [mp #32 #39]: #37
-#40 := [not-or-elim #42]: #10
-#51 := (= f4 f3)
-#52 := [symm #40]: #51
-#54 := [monotonicity #52 #40]: #53
-#49 := [symm #54]: #13
-#41 := (not #13)
-#43 := [not-or-elim #42]: #41
-[unit-resolution #43 #49]: false
-unsat
-06f5c6a6bea9a2fa1b9f87412f4567c4e365964a 94 0
+dc14c19019afcaa63507e4b503184496ed6bfd2f 94 0
 #2 := false
 decl f5 :: S1
 #13 := f5
@@ -1668,71 +916,7 @@
 #96 := [not-or-elim #89]: #70
 [mp #96 #112]: false
 unsat
-16b0a3ece5062e22059e25c49e83e51c0b937a88 63 0
-#2 := false
-decl f3 :: (-> S2 S2)
-decl f4 :: S2
-#8 := f4
-#9 := (f3 f4)
-#46 := (= f4 #9)
-#10 := (f3 #9)
-#12 := (f3 #10)
-#78 := (= #12 #9)
-#76 := (= #9 #12)
-#37 := (= f4 #10)
-#13 := (f3 #12)
-#14 := (f3 #13)
-#40 := (= f4 #14)
-#43 := (and #37 #40)
-#52 := (not #43)
-#53 := (or #52 #46)
-#58 := (not #53)
-#17 := (= #9 f4)
-#15 := (= #14 f4)
-#11 := (= #10 f4)
-#16 := (and #11 #15)
-#18 := (implies #16 #17)
-#19 := (not #18)
-#59 := (iff #19 #58)
-#56 := (iff #18 #53)
-#49 := (implies #43 #46)
-#54 := (iff #49 #53)
-#55 := [rewrite]: #54
-#50 := (iff #18 #49)
-#47 := (iff #17 #46)
-#48 := [rewrite]: #47
-#44 := (iff #16 #43)
-#41 := (iff #15 #40)
-#42 := [rewrite]: #41
-#38 := (iff #11 #37)
-#39 := [rewrite]: #38
-#45 := [monotonicity #39 #42]: #44
-#51 := [monotonicity #45 #48]: #50
-#57 := [trans #51 #55]: #56
-#60 := [monotonicity #57]: #59
-#36 := [asserted]: #19
-#63 := [mp #36 #60]: #58
-#61 := [not-or-elim #63]: #43
-#62 := [and-elim #61]: #37
-#77 := [monotonicity #62]: #76
-#79 := [symm #77]: #78
-#81 := (= f4 #12)
-#67 := (= #14 #12)
-#70 := (= #12 #14)
-#72 := (= #10 #13)
-#74 := (= #13 #10)
-#75 := [monotonicity #79]: #74
-#73 := [symm #75]: #72
-#71 := [monotonicity #73]: #70
-#80 := [symm #71]: #67
-#64 := [and-elim #61]: #40
-#82 := [trans #64 #80]: #81
-#83 := [trans #82 #79]: #46
-#65 := (not #46)
-#66 := [not-or-elim #63]: #65
-[unit-resolution #66 #83]: false
-unsat
-f45903baa95c438d411754363ad5219754af16ca 114 0
+682b764e75beeea86518e9d265c22d014c6c1f1d 114 0
 #2 := false
 decl f5 :: S1
 #13 := f5
@@ -1847,203 +1031,7 @@
 #136 := [unit-resolution #109 #135]: #87
 [unit-resolution #93 #136 #134]: false
 unsat
-4b1d0031a7c5da8fe83e5d22bb3a19a4727fa881 195 0
-#2 := false
-decl f6 :: S2
-#13 := f6
-decl f4 :: S2
-#10 := f4
-#15 := (= f4 f6)
-decl f5 :: S2
-#11 := f5
-decl f3 :: S1
-#8 := f3
-decl f1 :: S1
-#4 := f1
-#40 := (= f1 f3)
-#43 := (ite #40 f4 f5)
-#49 := (= f6 #43)
-#200 := (iff #49 #15)
-#198 := (iff #15 #49)
-#46 := (= #43 f6)
-#50 := (iff #46 #49)
-#197 := [commutativity]: #50
-#195 := (iff #15 #46)
-#110 := (= f4 #43)
-#111 := (= f5 #43)
-#57 := (not #40)
-#180 := [hypothesis]: #57
-#114 := (or #40 #111)
-#115 := [def-axiom]: #114
-#184 := [unit-resolution #115 #180]: #111
-#185 := (= f6 f5)
-#18 := (= f5 f6)
-#174 := (iff #110 #15)
-#172 := (iff #15 #110)
-#68 := (or #18 #40)
-#95 := (not #68)
-#58 := (or #15 #57)
-#94 := (not #58)
-#96 := (or #94 #95)
-#123 := (not #18)
-#147 := [hypothesis]: #123
-#157 := (or #96 #18)
-#97 := (not #96)
-#145 := [hypothesis]: #97
-#132 := (or #96 #68)
-#133 := [def-axiom]: #132
-#148 := [unit-resolution #133 #145]: #68
-#128 := (or #95 #18 #40)
-#129 := [def-axiom]: #128
-#149 := [unit-resolution #129 #148 #147]: #40
-#112 := (or #57 #110)
-#113 := [def-axiom]: #112
-#150 := [unit-resolution #113 #149]: #110
-#153 := (= f6 f4)
-#130 := (or #96 #58)
-#131 := [def-axiom]: #130
-#151 := [unit-resolution #131 #145]: #58
-#121 := (or #94 #15 #57)
-#122 := [def-axiom]: #121
-#152 := [unit-resolution #122 #149 #151]: #15
-#154 := [symm #152]: #153
-#155 := [trans #154 #150]: #49
-#82 := (not #49)
-#143 := (or #82 #96)
-#104 := (iff #49 #96)
-#73 := (and #58 #68)
-#83 := (iff #73 #82)
-#107 := (iff #83 #104)
-#99 := (iff #96 #49)
-#105 := (iff #99 #104)
-#106 := [rewrite]: #105
-#102 := (iff #83 #99)
-#91 := (iff #97 #82)
-#100 := (iff #91 #99)
-#101 := [rewrite]: #100
-#88 := (iff #83 #91)
-#92 := (iff #73 #97)
-#93 := [rewrite]: #92
-#98 := [monotonicity #93]: #88
-#103 := [trans #98 #101]: #102
-#108 := [trans #103 #106]: #107
-#9 := (= f3 f1)
-#17 := (not #9)
-#19 := (implies #17 #18)
-#16 := (implies #9 #15)
-#20 := (and #16 #19)
-#12 := (ite #9 f4 f5)
-#14 := (= #12 f6)
-#21 := (iff #14 #20)
-#22 := (not #21)
-#86 := (iff #22 #83)
-#76 := (iff #49 #73)
-#79 := (not #76)
-#84 := (iff #79 #83)
-#85 := [rewrite]: #84
-#80 := (iff #22 #79)
-#77 := (iff #21 #76)
-#74 := (iff #20 #73)
-#71 := (iff #19 #68)
-#65 := (implies #57 #18)
-#69 := (iff #65 #68)
-#70 := [rewrite]: #69
-#66 := (iff #19 #65)
-#63 := (iff #17 #57)
-#41 := (iff #9 #40)
-#42 := [rewrite]: #41
-#64 := [monotonicity #42]: #63
-#67 := [monotonicity #64]: #66
-#72 := [trans #67 #70]: #71
-#61 := (iff #16 #58)
-#54 := (implies #40 #15)
-#59 := (iff #54 #58)
-#60 := [rewrite]: #59
-#55 := (iff #16 #54)
-#56 := [monotonicity #42]: #55
-#62 := [trans #56 #60]: #61
-#75 := [monotonicity #62 #72]: #74
-#52 := (iff #14 #49)
-#51 := [rewrite]: #50
-#47 := (iff #14 #46)
-#44 := (= #12 #43)
-#45 := [monotonicity #42]: #44
-#48 := [monotonicity #45]: #47
-#53 := [trans #48 #51]: #52
-#78 := [monotonicity #53 #75]: #77
-#81 := [monotonicity #78]: #80
-#87 := [trans #81 #85]: #86
-#39 := [asserted]: #22
-#90 := [mp #39 #87]: #83
-#109 := [mp #90 #108]: #104
-#136 := (not #104)
-#141 := (or #82 #96 #136)
-#142 := [def-axiom]: #141
-#144 := [unit-resolution #142 #109]: #143
-#146 := [unit-resolution #144 #145]: #82
-#156 := [unit-resolution #146 #155]: false
-#158 := [lemma #156]: #157
-#159 := [unit-resolution #158 #147]: #96
-#139 := (or #49 #97)
-#137 := (or #49 #97 #136)
-#138 := [def-axiom]: #137
-#140 := [unit-resolution #138 #109]: #139
-#160 := [unit-resolution #140 #159]: #49
-#173 := [monotonicity #160]: #172
-#175 := [symm #173]: #174
-#163 := (not #111)
-#164 := (iff #123 #163)
-#161 := (iff #18 #111)
-#162 := [monotonicity #160]: #161
-#165 := [monotonicity #162]: #164
-#166 := [mp #147 #165]: #163
-#167 := [unit-resolution #115 #166]: #40
-#171 := [unit-resolution #113 #167]: #110
-#176 := [mp #171 #175]: #15
-#116 := (not #15)
-#126 := (or #68 #57)
-#127 := [def-axiom]: #126
-#168 := [unit-resolution #127 #167]: #68
-#134 := (or #97 #94 #95)
-#135 := [def-axiom]: #134
-#169 := [unit-resolution #135 #168 #159]: #94
-#117 := (or #58 #116)
-#118 := [def-axiom]: #117
-#170 := [unit-resolution #118 #169]: #116
-#177 := [unit-resolution #170 #176]: false
-#178 := [lemma #177]: #18
-#186 := [symm #178]: #185
-#187 := [trans #186 #184]: #49
-#124 := (or #68 #123)
-#125 := [def-axiom]: #124
-#179 := [unit-resolution #125 #178]: #68
-#119 := (or #58 #40)
-#120 := [def-axiom]: #119
-#181 := [unit-resolution #120 #180]: #58
-#182 := [unit-resolution #135 #181 #179]: #97
-#183 := [unit-resolution #144 #182]: #82
-#188 := [unit-resolution #183 #187]: false
-#189 := [lemma #188]: #40
-#194 := [unit-resolution #113 #189]: #110
-#196 := [monotonicity #194]: #195
-#199 := [trans #196 #197]: #198
-#201 := [symm #199]: #200
-#202 := (iff #82 #116)
-#203 := [monotonicity #201]: #202
-#190 := [hypothesis]: #82
-#204 := [mp #190 #203]: #116
-#191 := [unit-resolution #140 #190]: #97
-#192 := [unit-resolution #131 #191]: #58
-#193 := [unit-resolution #122 #192 #189]: #15
-#205 := [unit-resolution #193 #204]: false
-#206 := [lemma #205]: #49
-#210 := [mp #206 #201]: #15
-#207 := [unit-resolution #144 #206]: #96
-#208 := [unit-resolution #135 #207 #179]: #94
-#209 := [unit-resolution #118 #208]: #116
-[unit-resolution #209 #210]: false
-unsat
-d88d2fa5d625a6669d08e24799b598bc3243212b 121 0
+c134dad603b5004748b4c36796dfbdf41bc232cc 121 0
 #2 := false
 decl f5 :: S1
 #13 := f5
@@ -2165,44 +1153,7 @@
 #99 := [not-or-elim #92]: #98
 [mp #99 #138]: false
 unsat
-08da4c2954fa7e715a65caa44d6cfda212fd8151 36 0
-#2 := false
-decl f5 :: S2
-#10 := f5
-decl f3 :: S2
-#8 := f3
-#12 := (= f3 f5)
-#13 := (not #12)
-decl f4 :: S2
-#9 := f4
-#11 := (distinct f3 f4 f5)
-#33 := (not #11)
-#34 := (or #33 #13)
-#37 := (not #34)
-#14 := (implies #11 #13)
-#15 := (not #14)
-#38 := (iff #15 #37)
-#35 := (iff #14 #34)
-#36 := [rewrite]: #35
-#39 := [monotonicity #36]: #38
-#32 := [asserted]: #15
-#42 := [mp #32 #39]: #37
-#41 := [not-or-elim #42]: #12
-#52 := (= f4 f5)
-#53 := (not #52)
-#50 := (= f3 f4)
-#51 := (not #50)
-#48 := (and #51 #13 #53)
-#40 := [not-or-elim #42]: #11
-#58 := (or #33 #48)
-#59 := [def-axiom]: #58
-#62 := [unit-resolution #59 #40]: #48
-#49 := (not #48)
-#45 := (or #49 #13)
-#43 := [def-axiom]: #45
-[unit-resolution #43 #62 #41]: false
-unsat
-bd50ed75e1a24fdaf488b7bfb75ceea321293722 110 0
+7afff0259006c1c4b6a802d510f1b96695b3e9ad 110 0
 #2 := false
 decl f4 :: S1
 #10 := f4
@@ -2313,85 +1264,7 @@
 #84 := [and-elim #83]: #44
 [mp #84 #129]: false
 unsat
-8dc240ee402eb47862312ada03d54f1453a80352 77 0
-#2 := false
-decl f4 :: S2
-#9 := f4
-decl f3 :: S2
-#8 := f3
-#81 := (= f3 f4)
-decl f6 :: S2
-#12 := f6
-#36 := (= f4 f6)
-#100 := (iff #36 #81)
-#98 := (iff #81 #36)
-#13 := (= f6 f4)
-#37 := (iff #13 #36)
-#97 := [commutativity]: #37
-#95 := (iff #81 #13)
-#14 := (= f3 f6)
-#42 := (not #36)
-#15 := (not #14)
-decl f5 :: S2
-#10 := f5
-#11 := (distinct f3 f4 f5)
-#51 := (not #11)
-#60 := (or #51 #15 #42)
-#63 := (not #60)
-#16 := (implies #13 #15)
-#17 := (implies #11 #16)
-#18 := (not #17)
-#66 := (iff #18 #63)
-#43 := (or #15 #42)
-#52 := (or #51 #43)
-#57 := (not #52)
-#64 := (iff #57 #63)
-#61 := (iff #52 #60)
-#62 := [rewrite]: #61
-#65 := [monotonicity #62]: #64
-#58 := (iff #18 #57)
-#55 := (iff #17 #52)
-#48 := (implies #11 #43)
-#53 := (iff #48 #52)
-#54 := [rewrite]: #53
-#49 := (iff #17 #48)
-#46 := (iff #16 #43)
-#39 := (implies #36 #15)
-#44 := (iff #39 #43)
-#45 := [rewrite]: #44
-#40 := (iff #16 #39)
-#38 := [rewrite]: #37
-#41 := [monotonicity #38]: #40
-#47 := [trans #41 #45]: #46
-#50 := [monotonicity #47]: #49
-#56 := [trans #50 #54]: #55
-#59 := [monotonicity #56]: #58
-#67 := [trans #59 #65]: #66
-#35 := [asserted]: #18
-#68 := [mp #35 #67]: #63
-#70 := [not-or-elim #68]: #14
-#96 := [monotonicity #70]: #95
-#99 := [trans #96 #97]: #98
-#101 := [symm #99]: #100
-#71 := [not-or-elim #68]: #36
-#102 := [mp #71 #101]: #81
-#82 := (not #81)
-#79 := (= f4 f5)
-#80 := (not #79)
-#83 := (= f3 f5)
-#84 := (not #83)
-#77 := (and #82 #84 #80)
-#69 := [not-or-elim #68]: #11
-#89 := (or #51 #77)
-#90 := [def-axiom]: #89
-#93 := [unit-resolution #90 #69]: #77
-#78 := (not #77)
-#75 := (or #78 #82)
-#76 := [def-axiom]: #75
-#94 := [unit-resolution #76 #93]: #82
-[unit-resolution #94 #102]: false
-unsat
-82e8476a2b40c7b4ddf6aad1b870323c7e2adbfc 127 0
+83b90287eef3783de4d8fa4aad70d03dd0abddbf 127 0
 #2 := false
 decl f5 :: S1
 #13 := f5
@@ -2519,36 +1392,7 @@
 #149 := [unit-resolution #131 #147]: #97
 [unit-resolution #108 #149 #148]: false
 unsat
-3ec381e2388d16595dee423cb23410b8ee65f5be 28 0
-#2 := false
-decl f4 :: S2
-#9 := f4
-decl f3 :: S2
-#8 := f3
-#10 := (distinct f3 f4 f3 f4)
-#11 := (not #10)
-#12 := (not #11)
-#44 := (iff #12 false)
-#1 := true
-#39 := (not true)
-#42 := (iff #39 false)
-#43 := [rewrite]: #42
-#40 := (iff #12 #39)
-#37 := (iff #11 true)
-#32 := (not false)
-#35 := (iff #32 true)
-#36 := [rewrite]: #35
-#33 := (iff #11 #32)
-#30 := (iff #10 false)
-#31 := [rewrite]: #30
-#34 := [monotonicity #31]: #33
-#38 := [trans #34 #36]: #37
-#41 := [monotonicity #38]: #40
-#45 := [trans #41 #43]: #44
-#29 := [asserted]: #12
-[mp #29 #45]: false
-unsat
-d88d10125d202994145262227e4bbc27de00325d 66 0
+a7f2c916e2192b153aaac29ea6f75d0842fd84d2 66 0
 #2 := false
 decl f3 :: S1
 #8 := f3
@@ -2615,38 +1459,7 @@
 #74 := [not-or-elim #84]: #33
 [unit-resolution #69 #74]: false
 unsat
-7978400bf5569e9c21731e843185e21f42a67667 30 0
-#2 := false
-decl f4 :: S2
-#9 := f4
-decl f3 :: S2
-#8 := f3
-#10 := (= f3 f4)
-#11 := (not #10)
-#12 := (not #11)
-#13 := (implies #10 #12)
-#14 := (not #13)
-#46 := (iff #14 false)
-#1 := true
-#41 := (not true)
-#44 := (iff #41 false)
-#45 := [rewrite]: #44
-#42 := (iff #14 #41)
-#39 := (iff #13 true)
-#34 := (implies #10 #10)
-#37 := (iff #34 true)
-#38 := [rewrite]: #37
-#35 := (iff #13 #34)
-#32 := (iff #12 #10)
-#33 := [rewrite]: #32
-#36 := [monotonicity #33]: #35
-#40 := [trans #36 #38]: #39
-#43 := [monotonicity #40]: #42
-#47 := [trans #43 #45]: #46
-#31 := [asserted]: #14
-[mp #31 #47]: false
-unsat
-123f558f9f3ffbb84aa29cf95bfa60a7fd8f9c53 156 0
+d398c106ccd0941d46bab8e5feda32347dd02740 156 0
 #2 := false
 decl f5 :: S1
 #13 := f5
@@ -2803,48 +1616,7 @@
 #114 := [def-axiom]: #113
 [unit-resolution #114 #172 #170]: false
 unsat
-10f35a88989b72a52fb3f98d7f7428c1e1b3e122 40 0
-#2 := false
-decl f5 :: S2
-#11 := f5
-decl f4 :: S2
-#9 := f4
-decl f3 :: S2
-#8 := f3
-#14 := (distinct f3 f4 f5)
-#15 := (not #14)
-#12 := (= f3 f5)
-#10 := (= f3 f4)
-#13 := (and #10 #12)
-#35 := (not #13)
-#36 := (or #35 #15)
-#39 := (not #36)
-#16 := (implies #13 #15)
-#17 := (not #16)
-#40 := (iff #17 #39)
-#37 := (iff #16 #36)
-#38 := [rewrite]: #37
-#41 := [monotonicity #38]: #40
-#34 := [asserted]: #17
-#44 := [mp #34 #41]: #39
-#46 := [not-or-elim #44]: #14
-#58 := (= f4 f5)
-#59 := (not #58)
-#57 := (not #12)
-#56 := (not #10)
-#54 := (and #56 #57 #59)
-#55 := (not #54)
-#42 := [not-or-elim #44]: #13
-#43 := [and-elim #42]: #10
-#52 := (or #55 #56)
-#53 := [def-axiom]: #52
-#66 := [unit-resolution #53 #43]: #55
-#62 := (or #15 #54)
-#63 := [def-axiom]: #62
-#67 := [unit-resolution #63 #66]: #15
-[unit-resolution #67 #46]: false
-unsat
-b3152f8558491ad9fde8dfd8e10a714b3e98cb57 114 0
+78ae843b278e5f0cc65999ee60776dec7c9f683f 114 0
 #2 := false
 decl f5 :: S1
 #13 := f5
@@ -2959,41 +1731,7 @@
 #101 := [mp #36 #100]: #94
 [mp #101 #130]: false
 unsat
-ed5b4a9da6fcf2483f2081f8e182d1c39bc7b1d0 33 0
-#2 := false
-decl f3 :: S2
-#8 := f3
-decl f5 :: S2
-#10 := f5
-decl f4 :: S2
-#9 := f4
-decl f6 :: S2
-#11 := f6
-#13 := (distinct f6 f4 f5 f3)
-#12 := (distinct f3 f4 f5 f6)
-#14 := (implies #12 #13)
-#15 := (not #14)
-#47 := (iff #15 false)
-#1 := true
-#42 := (not true)
-#45 := (iff #42 false)
-#46 := [rewrite]: #45
-#43 := (iff #15 #42)
-#40 := (iff #14 true)
-#35 := (implies #12 #12)
-#38 := (iff #35 true)
-#39 := [rewrite]: #38
-#36 := (iff #14 #35)
-#33 := (iff #13 #12)
-#34 := [rewrite]: #33
-#37 := [monotonicity #34]: #36
-#41 := [trans #37 #39]: #40
-#44 := [monotonicity #41]: #43
-#48 := [trans #44 #46]: #47
-#32 := [asserted]: #15
-[mp #32 #48]: false
-unsat
-710a09ff1423bdde95e9fb0541e56c839786cb31 198 0
+016fee2f7f490b7727a4579bedbeaa0a03271f5c 198 0
 #2 := false
 decl f5 :: S1
 #14 := f5
@@ -3192,99 +1930,7 @@
 #141 := [not-or-elim #136]: #93
 [mp #141 #215]: false
 unsat
-65f8a76a6f6a38a4bead0ee4b39d32d58aff533c 91 0
-#2 := false
-decl f5 :: S2
-#10 := f5
-decl f4 :: S2
-#9 := f4
-decl f3 :: S2
-#8 := f3
-#13 := (distinct f3 f4 f5)
-#67 := (= f4 f5)
-#68 := (not #67)
-#63 := (= f3 f5)
-#64 := (not #63)
-#61 := (= f3 f4)
-#62 := (not #61)
-#93 := (and #62 #64 #68)
-decl f6 :: S2
-#11 := f6
-#71 := (= f5 f6)
-#72 := (not #71)
-#69 := (= f4 f6)
-#70 := (not #69)
-#65 := (= f3 f6)
-#66 := (not #65)
-#73 := (and #62 #64 #66 #68 #70 #72)
-#12 := (distinct f3 f4 f5 f6)
-#14 := (distinct f4 f5 f6)
-#15 := (and #13 #14)
-#35 := (not #12)
-#36 := (or #35 #15)
-#39 := (not #36)
-#16 := (implies #12 #15)
-#17 := (not #16)
-#40 := (iff #17 #39)
-#37 := (iff #16 #36)
-#38 := [rewrite]: #37
-#41 := [monotonicity #38]: #40
-#34 := [asserted]: #17
-#44 := [mp #34 #41]: #39
-#42 := [not-or-elim #44]: #12
-#89 := (or #35 #73)
-#90 := [def-axiom]: #89
-#121 := [unit-resolution #90 #42]: #73
-#74 := (not #73)
-#75 := (or #74 #62)
-#76 := [def-axiom]: #75
-#122 := [unit-resolution #76 #121]: #62
-#81 := (or #74 #68)
-#82 := [def-axiom]: #81
-#123 := [unit-resolution #82 #121]: #68
-#77 := (or #74 #64)
-#78 := [def-axiom]: #77
-#124 := [unit-resolution #78 #121]: #64
-#101 := (or #93 #61 #63 #67)
-#102 := [def-axiom]: #101
-#125 := [unit-resolution #102 #124 #123 #122]: #93
-#94 := (not #93)
-#105 := (or #13 #94)
-#106 := [def-axiom]: #105
-#126 := [unit-resolution #106 #125]: #13
-#107 := (and #68 #70 #72)
-#85 := (or #74 #72)
-#86 := [def-axiom]: #85
-#127 := [unit-resolution #86 #121]: #72
-#83 := (or #74 #70)
-#84 := [def-axiom]: #83
-#128 := [unit-resolution #84 #121]: #70
-#115 := (or #107 #67 #69 #71)
-#116 := [def-axiom]: #115
-#129 := [unit-resolution #116 #128 #127 #123]: #107
-#108 := (not #107)
-#119 := (or #14 #108)
-#120 := [def-axiom]: #119
-#130 := [unit-resolution #120 #129]: #14
-#54 := (not #14)
-#53 := (not #13)
-#55 := (or #53 #54)
-#43 := (not #15)
-#58 := (iff #43 #55)
-#56 := (not #55)
-#49 := (not #56)
-#46 := (iff #49 #55)
-#57 := [rewrite]: #46
-#50 := (iff #43 #49)
-#51 := (iff #15 #56)
-#52 := [rewrite]: #51
-#48 := [monotonicity #52]: #50
-#59 := [trans #48 #57]: #58
-#45 := [not-or-elim #44]: #43
-#60 := [mp #45 #59]: #55
-[unit-resolution #60 #130 #126]: false
-unsat
-c2ceb7f67b73bb3bdaed86d97f31bbabffade4ad 134 0
+3e0305075f83d65498b4a11b8ca72526b8e2d46d 134 0
 #2 := false
 decl f5 :: S1
 #14 := f5
@@ -3419,7 +2065,7 @@
 #121 := [not-or-elim #116]: #81
 [mp #121 #151]: false
 unsat
-77f105a2335d8f399aa84eb3f5b2c6fc18cb5cb0 162 0
+78ecef8a6265ee63380698aee66f7437c6eb9791 162 0
 #2 := false
 decl f5 :: S1
 #12 := f5
@@ -3582,7 +2228,7 @@
 #178 := [unit-resolution #128 #177]: #70
 [unit-resolution #136 #178 #176 #168]: false
 unsat
-f7a1358537b42b76857f25d2bb6dab6a65b957d3 54 0
+cbabdc011b8a302b61906c6191f0d7af27e25599 54 0
 #2 := false
 decl f1 :: S1
 #4 := f1
@@ -3637,7 +2283,7 @@
 #31 := [asserted]: #14
 [mp #31 #69]: false
 unsat
-cb5396a335ba3f235c4f126793b6452d8c22b8a0 144 0
+ae8651ca265765f3bb029935ca54bf350a439257 144 0
 #2 := false
 decl f5 :: S1
 #12 := f5
@@ -3782,33 +2428,7 @@
 #126 := [not-or-elim #121]: #102
 [mp #126 #161]: false
 unsat
-828c6cc5274797cbab0d66f9715797229aa86661 25 0
-#2 := false
-#8 := (:var 0 S2)
-#9 := (= #8 #8)
-#10 := (forall (vars (?v0 S2)) #9)
-#11 := (not #10)
-#43 := (iff #11 false)
-#1 := true
-#38 := (not true)
-#41 := (iff #38 false)
-#42 := [rewrite]: #41
-#39 := (iff #11 #38)
-#36 := (iff #10 true)
-#31 := (forall (vars (?v0 S2)) true)
-#34 := (iff #31 true)
-#35 := [elim-unused]: #34
-#32 := (iff #10 #31)
-#29 := (iff #9 true)
-#30 := [rewrite]: #29
-#33 := [quant-intro #30]: #32
-#37 := [trans #33 #35]: #36
-#40 := [monotonicity #37]: #39
-#44 := [trans #40 #42]: #43
-#28 := [asserted]: #11
-[mp #28 #44]: false
-unsat
-614eb9a555c5a60772dafc0939c267db1405600c 121 0
+f037e85f433fe4b6db0edebe38d7599d0d5f9a56 121 0
 #2 := false
 decl f5 :: S1
 #13 := f5
@@ -3930,42 +2550,7 @@
 #115 := [not-or-elim #109]: #90
 [mp #115 #138]: false
 unsat
-95dd35edf8ef21f4e7cf317ce58b9a2bec541f8f 34 0
-#2 := false
-decl f1 :: S1
-#4 := f1
-decl f3 :: (-> S2 S1)
-#8 := (:var 0 S2)
-#9 := (f3 #8)
-#10 := (= #9 f1)
-#11 := (forall (vars (?v0 S2)) #10)
-#12 := (iff #11 #11)
-#13 := (not #12)
-#49 := (iff #13 false)
-#1 := true
-#44 := (not true)
-#47 := (iff #44 false)
-#48 := [rewrite]: #47
-#45 := (iff #13 #44)
-#42 := (iff #12 true)
-#31 := (= f1 #9)
-#34 := (forall (vars (?v0 S2)) #31)
-#37 := (iff #34 #34)
-#40 := (iff #37 true)
-#41 := [rewrite]: #40
-#38 := (iff #12 #37)
-#35 := (iff #11 #34)
-#32 := (iff #10 #31)
-#33 := [rewrite]: #32
-#36 := [quant-intro #33]: #35
-#39 := [monotonicity #36 #36]: #38
-#43 := [trans #39 #41]: #42
-#46 := [monotonicity #43]: #45
-#50 := [trans #46 #48]: #49
-#30 := [asserted]: #13
-[mp #30 #50]: false
-unsat
-93c5fb9808d8360640a347075369425791ec133b 123 0
+a79fa106ccc873005f00005d34ba1798389c2876 123 0
 #2 := false
 decl f4 :: S1
 #10 := f4
@@ -4089,7 +2674,1914 @@
 #92 := [not-or-elim #91]: #72
 [mp #92 #141]: false
 unsat
-b4400db85e456f9b3856673990716e60114a7552 95 0
+7e097d3ef4e0c4927519d96187a725503177b49b 60 0
+#2 := false
+decl f1 :: S1
+#4 := f1
+decl f3 :: S1
+#8 := f3
+#9 := (= f3 f1)
+#14 := (not #9)
+decl f4 :: S1
+#10 := f4
+#11 := (= f4 f1)
+#13 := (not #11)
+#15 := (implies #13 #14)
+#12 := (implies #9 #11)
+#16 := (implies #12 #15)
+#17 := (not #16)
+#74 := (iff #17 false)
+#1 := true
+#69 := (not true)
+#72 := (iff #69 false)
+#73 := [rewrite]: #72
+#70 := (iff #17 #69)
+#67 := (iff #16 true)
+#38 := (= f1 f4)
+#35 := (= f1 f3)
+#44 := (not #35)
+#45 := (or #44 #38)
+#62 := (implies #45 #45)
+#65 := (iff #62 true)
+#66 := [rewrite]: #65
+#63 := (iff #16 #62)
+#60 := (iff #15 #45)
+#50 := (not #38)
+#55 := (implies #50 #44)
+#58 := (iff #55 #45)
+#59 := [rewrite]: #58
+#56 := (iff #15 #55)
+#53 := (iff #14 #44)
+#36 := (iff #9 #35)
+#37 := [rewrite]: #36
+#54 := [monotonicity #37]: #53
+#51 := (iff #13 #50)
+#39 := (iff #11 #38)
+#40 := [rewrite]: #39
+#52 := [monotonicity #40]: #51
+#57 := [monotonicity #52 #54]: #56
+#61 := [trans #57 #59]: #60
+#48 := (iff #12 #45)
+#41 := (implies #35 #38)
+#46 := (iff #41 #45)
+#47 := [rewrite]: #46
+#42 := (iff #12 #41)
+#43 := [monotonicity #37 #40]: #42
+#49 := [trans #43 #47]: #48
+#64 := [monotonicity #49 #61]: #63
+#68 := [trans #64 #66]: #67
+#71 := [monotonicity #68]: #70
+#75 := [trans #71 #73]: #74
+#34 := [asserted]: #17
+[mp #34 #75]: false
+unsat
+cb74096d43b27b6df4bed4b142ce429518455d83 121 0
+#2 := false
+decl f5 :: S1
+#12 := f5
+decl f1 :: S1
+#4 := f1
+#44 := (= f1 f5)
+decl f4 :: S1
+#10 := f4
+#41 := (= f1 f4)
+decl f3 :: S1
+#8 := f3
+#38 := (= f1 f3)
+#53 := (not #38)
+#90 := (or #53 #41 #44)
+#137 := (iff #90 false)
+#132 := (or false false false)
+#135 := (iff #132 false)
+#136 := [rewrite]: #135
+#133 := (iff #90 #132)
+#123 := (iff #44 false)
+#113 := (not #44)
+#93 := (not #90)
+#99 := (or #53 #41 #44 #93)
+#104 := (not #99)
+#13 := (= f5 f1)
+#9 := (= f3 f1)
+#17 := (implies #9 #13)
+#11 := (= f4 f1)
+#16 := (implies #9 #11)
+#18 := (or #16 #17)
+#14 := (or #11 #13)
+#15 := (implies #9 #14)
+#19 := (implies #15 #18)
+#20 := (not #19)
+#107 := (iff #20 #104)
+#70 := (or #53 #44)
+#62 := (or #53 #41)
+#75 := (or #62 #70)
+#47 := (or #41 #44)
+#54 := (or #53 #47)
+#81 := (not #54)
+#82 := (or #81 #75)
+#87 := (not #82)
+#105 := (iff #87 #104)
+#102 := (iff #82 #99)
+#96 := (or #93 #75)
+#100 := (iff #96 #99)
+#101 := [rewrite]: #100
+#97 := (iff #82 #96)
+#94 := (iff #81 #93)
+#91 := (iff #54 #90)
+#92 := [rewrite]: #91
+#95 := [monotonicity #92]: #94
+#98 := [monotonicity #95]: #97
+#103 := [trans #98 #101]: #102
+#106 := [monotonicity #103]: #105
+#88 := (iff #20 #87)
+#85 := (iff #19 #82)
+#78 := (implies #54 #75)
+#83 := (iff #78 #82)
+#84 := [rewrite]: #83
+#79 := (iff #19 #78)
+#76 := (iff #18 #75)
+#73 := (iff #17 #70)
+#67 := (implies #38 #44)
+#71 := (iff #67 #70)
+#72 := [rewrite]: #71
+#68 := (iff #17 #67)
+#45 := (iff #13 #44)
+#46 := [rewrite]: #45
+#39 := (iff #9 #38)
+#40 := [rewrite]: #39
+#69 := [monotonicity #40 #46]: #68
+#74 := [trans #69 #72]: #73
+#65 := (iff #16 #62)
+#59 := (implies #38 #41)
+#63 := (iff #59 #62)
+#64 := [rewrite]: #63
+#60 := (iff #16 #59)
+#42 := (iff #11 #41)
+#43 := [rewrite]: #42
+#61 := [monotonicity #40 #43]: #60
+#66 := [trans #61 #64]: #65
+#77 := [monotonicity #66 #74]: #76
+#57 := (iff #15 #54)
+#50 := (implies #38 #47)
+#55 := (iff #50 #54)
+#56 := [rewrite]: #55
+#51 := (iff #15 #50)
+#48 := (iff #14 #47)
+#49 := [monotonicity #43 #46]: #48
+#52 := [monotonicity #40 #49]: #51
+#58 := [trans #52 #56]: #57
+#80 := [monotonicity #58 #77]: #79
+#86 := [trans #80 #84]: #85
+#89 := [monotonicity #86]: #88
+#108 := [trans #89 #106]: #107
+#37 := [asserted]: #20
+#109 := [mp #37 #108]: #104
+#114 := [not-or-elim #109]: #113
+#124 := [iff-false #114]: #123
+#121 := (iff #41 false)
+#111 := (not #41)
+#112 := [not-or-elim #109]: #111
+#122 := [iff-false #112]: #121
+#130 := (iff #53 false)
+#1 := true
+#125 := (not true)
+#128 := (iff #125 false)
+#129 := [rewrite]: #128
+#126 := (iff #53 #125)
+#119 := (iff #38 true)
+#110 := [not-or-elim #109]: #38
+#120 := [iff-true #110]: #119
+#127 := [monotonicity #120]: #126
+#131 := [trans #127 #129]: #130
+#134 := [monotonicity #131 #122 #124]: #133
+#138 := [trans #134 #136]: #137
+#115 := [not-or-elim #109]: #90
+[mp #115 #138]: false
+unsat
+f67715569c1aa586d3c8e9a94e103b17f1f34614 84 0
+#2 := false
+decl f4 :: S1
+#10 := f4
+decl f1 :: S1
+#4 := f1
+#38 := (= f1 f4)
+#53 := (not #38)
+#97 := [hypothesis]: #53
+decl f3 :: S1
+#8 := f3
+#35 := (= f1 f3)
+#44 := (not #35)
+#45 := (or #44 #38)
+#54 := (or #35 #53)
+#59 := (and #45 #54)
+#62 := (iff #35 #38)
+#68 := (not #59)
+#69 := (or #68 #62)
+#74 := (not #69)
+#11 := (= f4 f1)
+#9 := (= f3 f1)
+#15 := (iff #9 #11)
+#13 := (implies #11 #9)
+#12 := (implies #9 #11)
+#14 := (and #12 #13)
+#16 := (implies #14 #15)
+#17 := (not #16)
+#75 := (iff #17 #74)
+#72 := (iff #16 #69)
+#65 := (implies #59 #62)
+#70 := (iff #65 #69)
+#71 := [rewrite]: #70
+#66 := (iff #16 #65)
+#63 := (iff #15 #62)
+#39 := (iff #11 #38)
+#40 := [rewrite]: #39
+#36 := (iff #9 #35)
+#37 := [rewrite]: #36
+#64 := [monotonicity #37 #40]: #63
+#60 := (iff #14 #59)
+#57 := (iff #13 #54)
+#50 := (implies #38 #35)
+#55 := (iff #50 #54)
+#56 := [rewrite]: #55
+#51 := (iff #13 #50)
+#52 := [monotonicity #40 #37]: #51
+#58 := [trans #52 #56]: #57
+#48 := (iff #12 #45)
+#41 := (implies #35 #38)
+#46 := (iff #41 #45)
+#47 := [rewrite]: #46
+#42 := (iff #12 #41)
+#43 := [monotonicity #37 #40]: #42
+#49 := [trans #43 #47]: #48
+#61 := [monotonicity #49 #58]: #60
+#67 := [monotonicity #61 #64]: #66
+#73 := [trans #67 #71]: #72
+#76 := [monotonicity #73]: #75
+#34 := [asserted]: #17
+#79 := [mp #34 #76]: #74
+#77 := [not-or-elim #79]: #59
+#78 := [and-elim #77]: #45
+#98 := [unit-resolution #78 #97]: #44
+#89 := (or #38 #35)
+#90 := (iff #38 #44)
+#81 := (not #62)
+#91 := (iff #81 #90)
+#92 := [rewrite]: #91
+#82 := [not-or-elim #79]: #81
+#93 := [mp #82 #92]: #90
+#94 := (not #90)
+#95 := (or #38 #35 #94)
+#88 := [def-axiom]: #95
+#86 := [unit-resolution #88 #93]: #89
+#99 := [unit-resolution #86 #98 #97]: false
+#100 := [lemma #99]: #38
+#80 := [and-elim #77]: #54
+#101 := [unit-resolution #80 #100]: #35
+#83 := (or #53 #44)
+#87 := (or #53 #44 #94)
+#85 := [def-axiom]: #87
+#96 := [unit-resolution #85 #93]: #83
+[unit-resolution #96 #101 #100]: false
+unsat
+8abc924327822787c2f5c4e6e7680303fbae57f8 47 0
+#2 := false
+decl f1 :: S1
+#4 := f1
+decl f3 :: S1
+#8 := f3
+#9 := (= f3 f1)
+decl f4 :: S1
+#10 := f4
+#11 := (= f4 f1)
+#13 := (iff #11 #9)
+#12 := (iff #9 #11)
+#14 := (iff #12 #13)
+#15 := (not #14)
+#61 := (iff #15 false)
+#1 := true
+#56 := (not true)
+#59 := (iff #56 false)
+#60 := [rewrite]: #59
+#57 := (iff #15 #56)
+#54 := (iff #14 true)
+#36 := (= f1 f4)
+#33 := (= f1 f3)
+#39 := (iff #33 #36)
+#49 := (iff #39 #39)
+#52 := (iff #49 true)
+#53 := [rewrite]: #52
+#50 := (iff #14 #49)
+#47 := (iff #13 #39)
+#42 := (iff #36 #33)
+#45 := (iff #42 #39)
+#46 := [rewrite]: #45
+#43 := (iff #13 #42)
+#34 := (iff #9 #33)
+#35 := [rewrite]: #34
+#37 := (iff #11 #36)
+#38 := [rewrite]: #37
+#44 := [monotonicity #38 #35]: #43
+#48 := [trans #44 #46]: #47
+#40 := (iff #12 #39)
+#41 := [monotonicity #35 #38]: #40
+#51 := [monotonicity #41 #48]: #50
+#55 := [trans #51 #53]: #54
+#58 := [monotonicity #55]: #57
+#62 := [trans #58 #60]: #61
+#32 := [asserted]: #15
+[mp #32 #62]: false
+unsat
+3731fdf00230f94750072246058d6aec5d30ad57 41 0
+#2 := false
+decl f1 :: S1
+#4 := f1
+decl f3 :: S1
+#8 := f3
+#9 := (= f3 f1)
+#10 := (not #9)
+#11 := (iff #9 #10)
+#12 := (not #11)
+#13 := (not #12)
+#56 := (iff #13 false)
+#1 := true
+#51 := (not true)
+#54 := (iff #51 false)
+#55 := [rewrite]: #54
+#52 := (iff #13 #51)
+#49 := (iff #12 true)
+#44 := (not false)
+#47 := (iff #44 true)
+#48 := [rewrite]: #47
+#45 := (iff #12 #44)
+#42 := (iff #11 false)
+#31 := (= f1 f3)
+#34 := (not #31)
+#37 := (iff #31 #34)
+#40 := (iff #37 false)
+#41 := [rewrite]: #40
+#38 := (iff #11 #37)
+#35 := (iff #10 #34)
+#32 := (iff #9 #31)
+#33 := [rewrite]: #32
+#36 := [monotonicity #33]: #35
+#39 := [monotonicity #33 #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
+#30 := [asserted]: #13
+[mp #30 #57]: false
+unsat
+c4682bd721e64af9e51bd86d6d03d59662baa782 60 0
+#2 := false
+decl f1 :: S1
+#4 := f1
+decl f3 :: S1
+#8 := f3
+#9 := (= f3 f1)
+#14 := (not #9)
+decl f4 :: S1
+#10 := f4
+#11 := (= f4 f1)
+#13 := (not #11)
+#15 := (implies #13 #14)
+#12 := (implies #9 #11)
+#16 := (iff #12 #15)
+#17 := (not #16)
+#74 := (iff #17 false)
+#1 := true
+#69 := (not true)
+#72 := (iff #69 false)
+#73 := [rewrite]: #72
+#70 := (iff #17 #69)
+#67 := (iff #16 true)
+#38 := (= f1 f4)
+#35 := (= f1 f3)
+#44 := (not #35)
+#45 := (or #44 #38)
+#62 := (iff #45 #45)
+#65 := (iff #62 true)
+#66 := [rewrite]: #65
+#63 := (iff #16 #62)
+#60 := (iff #15 #45)
+#50 := (not #38)
+#55 := (implies #50 #44)
+#58 := (iff #55 #45)
+#59 := [rewrite]: #58
+#56 := (iff #15 #55)
+#53 := (iff #14 #44)
+#36 := (iff #9 #35)
+#37 := [rewrite]: #36
+#54 := [monotonicity #37]: #53
+#51 := (iff #13 #50)
+#39 := (iff #11 #38)
+#40 := [rewrite]: #39
+#52 := [monotonicity #40]: #51
+#57 := [monotonicity #52 #54]: #56
+#61 := [trans #57 #59]: #60
+#48 := (iff #12 #45)
+#41 := (implies #35 #38)
+#46 := (iff #41 #45)
+#47 := [rewrite]: #46
+#42 := (iff #12 #41)
+#43 := [monotonicity #37 #40]: #42
+#49 := [trans #43 #47]: #48
+#64 := [monotonicity #49 #61]: #63
+#68 := [trans #64 #66]: #67
+#71 := [monotonicity #68]: #70
+#75 := [trans #71 #73]: #74
+#34 := [asserted]: #17
+[mp #34 #75]: false
+unsat
+36d1352b395221a431006c238d124165e2b745a3 72 0
+#2 := false
+decl f1 :: S1
+#4 := f1
+decl f3 :: S1
+#8 := f3
+#9 := (= f3 f1)
+#10 := (iff #9 #9)
+#11 := (iff #9 #10)
+#12 := (iff #9 #11)
+#13 := (iff #9 #12)
+#14 := (iff #9 #13)
+#15 := (iff #9 #14)
+#16 := (iff #9 #15)
+#17 := (iff #9 #16)
+#18 := (iff #9 #17)
+#19 := (not #18)
+#87 := (iff #19 false)
+#1 := true
+#82 := (not true)
+#85 := (iff #82 false)
+#86 := [rewrite]: #85
+#83 := (iff #19 #82)
+#80 := (iff #18 true)
+#37 := (= f1 f3)
+#40 := (iff #37 #37)
+#43 := (iff #40 true)
+#44 := [rewrite]: #43
+#78 := (iff #18 #40)
+#76 := (iff #17 #37)
+#47 := (iff #37 true)
+#50 := (iff #47 #37)
+#51 := [rewrite]: #50
+#74 := (iff #17 #47)
+#72 := (iff #16 true)
+#70 := (iff #16 #40)
+#68 := (iff #15 #37)
+#66 := (iff #15 #47)
+#64 := (iff #14 true)
+#62 := (iff #14 #40)
+#60 := (iff #13 #37)
+#58 := (iff #13 #47)
+#56 := (iff #12 true)
+#54 := (iff #12 #40)
+#52 := (iff #11 #37)
+#48 := (iff #11 #47)
+#45 := (iff #10 true)
+#41 := (iff #10 #40)
+#38 := (iff #9 #37)
+#39 := [rewrite]: #38
+#42 := [monotonicity #39 #39]: #41
+#46 := [trans #42 #44]: #45
+#49 := [monotonicity #39 #46]: #48
+#53 := [trans #49 #51]: #52
+#55 := [monotonicity #39 #53]: #54
+#57 := [trans #55 #44]: #56
+#59 := [monotonicity #39 #57]: #58
+#61 := [trans #59 #51]: #60
+#63 := [monotonicity #39 #61]: #62
+#65 := [trans #63 #44]: #64
+#67 := [monotonicity #39 #65]: #66
+#69 := [trans #67 #51]: #68
+#71 := [monotonicity #39 #69]: #70
+#73 := [trans #71 #44]: #72
+#75 := [monotonicity #39 #73]: #74
+#77 := [trans #75 #51]: #76
+#79 := [monotonicity #39 #77]: #78
+#81 := [trans #79 #44]: #80
+#84 := [monotonicity #81]: #83
+#88 := [trans #84 #86]: #87
+#36 := [asserted]: #19
+[mp #36 #88]: false
+unsat
+9a169ba4d23a9080465edad5616230c8b4e1ef98 149 0
+#2 := false
+decl f5 :: S1
+#12 := f5
+decl f1 :: S1
+#4 := f1
+#44 := (= f1 f5)
+decl f4 :: S1
+#10 := f4
+#41 := (= f1 f4)
+decl f3 :: S1
+#8 := f3
+#38 := (= f1 f3)
+#47 := (ite #38 #41 #44)
+#53 := (not #38)
+#54 := (or #53 #41)
+#64 := (or #38 #44)
+#91 := (not #64)
+#90 := (not #54)
+#92 := (or #90 #91)
+#143 := [hypothesis]: #90
+#128 := (or #92 #54)
+#129 := [def-axiom]: #128
+#144 := [unit-resolution #129 #143]: #92
+#78 := (not #47)
+#116 := (or #54 #38)
+#117 := [def-axiom]: #116
+#145 := [unit-resolution #117 #143]: #38
+#110 := (not #41)
+#118 := (or #54 #110)
+#119 := [def-axiom]: #118
+#146 := [unit-resolution #119 #143]: #110
+#106 := (or #78 #53 #41)
+#107 := [def-axiom]: #106
+#147 := [unit-resolution #107 #146 #145]: #78
+#93 := (not #92)
+#137 := (or #47 #93)
+#100 := (iff #47 #92)
+#69 := (and #54 #64)
+#79 := (iff #69 #78)
+#103 := (iff #79 #100)
+#95 := (iff #92 #47)
+#101 := (iff #95 #100)
+#102 := [rewrite]: #101
+#98 := (iff #79 #95)
+#87 := (iff #93 #78)
+#96 := (iff #87 #95)
+#97 := [rewrite]: #96
+#84 := (iff #79 #87)
+#88 := (iff #69 #93)
+#89 := [rewrite]: #88
+#94 := [monotonicity #89]: #84
+#99 := [trans #94 #97]: #98
+#104 := [trans #99 #102]: #103
+#13 := (= f5 f1)
+#9 := (= f3 f1)
+#16 := (not #9)
+#17 := (implies #16 #13)
+#11 := (= f4 f1)
+#15 := (implies #9 #11)
+#18 := (and #15 #17)
+#14 := (ite #9 #11 #13)
+#19 := (iff #14 #18)
+#20 := (not #19)
+#82 := (iff #20 #79)
+#72 := (iff #47 #69)
+#75 := (not #72)
+#80 := (iff #75 #79)
+#81 := [rewrite]: #80
+#76 := (iff #20 #75)
+#73 := (iff #19 #72)
+#70 := (iff #18 #69)
+#67 := (iff #17 #64)
+#61 := (implies #53 #44)
+#65 := (iff #61 #64)
+#66 := [rewrite]: #65
+#62 := (iff #17 #61)
+#45 := (iff #13 #44)
+#46 := [rewrite]: #45
+#59 := (iff #16 #53)
+#39 := (iff #9 #38)
+#40 := [rewrite]: #39
+#60 := [monotonicity #40]: #59
+#63 := [monotonicity #60 #46]: #62
+#68 := [trans #63 #66]: #67
+#57 := (iff #15 #54)
+#50 := (implies #38 #41)
+#55 := (iff #50 #54)
+#56 := [rewrite]: #55
+#51 := (iff #15 #50)
+#42 := (iff #11 #41)
+#43 := [rewrite]: #42
+#52 := [monotonicity #40 #43]: #51
+#58 := [trans #52 #56]: #57
+#71 := [monotonicity #58 #68]: #70
+#48 := (iff #14 #47)
+#49 := [monotonicity #40 #43 #46]: #48
+#74 := [monotonicity #49 #71]: #73
+#77 := [monotonicity #74]: #76
+#83 := [trans #77 #81]: #82
+#37 := [asserted]: #20
+#86 := [mp #37 #83]: #79
+#105 := [mp #86 #104]: #100
+#134 := (not #100)
+#135 := (or #47 #93 #134)
+#136 := [def-axiom]: #135
+#138 := [unit-resolution #136 #105]: #137
+#148 := [unit-resolution #138 #147 #144]: false
+#149 := [lemma #148]: #54
+#150 := [hypothesis]: #78
+#156 := (or #38 #47)
+#151 := [unit-resolution #138 #150]: #93
+#130 := (or #92 #64)
+#131 := [def-axiom]: #130
+#152 := [unit-resolution #131 #151]: #64
+#153 := [hypothesis]: #53
+#113 := (not #44)
+#114 := (or #47 #38 #113)
+#115 := [def-axiom]: #114
+#154 := [unit-resolution #115 #153 #150]: #113
+#126 := (or #91 #38 #44)
+#127 := [def-axiom]: #126
+#155 := [unit-resolution #127 #154 #153 #152]: false
+#157 := [lemma #155]: #156
+#158 := [unit-resolution #157 #150]: #38
+#111 := (or #47 #53 #110)
+#112 := [def-axiom]: #111
+#159 := [unit-resolution #112 #158 #150]: #110
+#120 := (or #90 #53 #41)
+#121 := [def-axiom]: #120
+#160 := [unit-resolution #121 #159 #158 #149]: false
+#161 := [lemma #160]: #47
+#141 := (or #78 #92)
+#139 := (or #78 #92 #134)
+#140 := [def-axiom]: #139
+#142 := [unit-resolution #140 #105]: #141
+#162 := [unit-resolution #142 #161]: #92
+#132 := (or #93 #90 #91)
+#133 := [def-axiom]: #132
+#163 := [unit-resolution #133 #162 #149]: #91
+#122 := (or #64 #53)
+#123 := [def-axiom]: #122
+#164 := [unit-resolution #123 #163]: #53
+#124 := (or #64 #113)
+#125 := [def-axiom]: #124
+#165 := [unit-resolution #125 #163]: #113
+#108 := (or #78 #38 #44)
+#109 := [def-axiom]: #108
+[unit-resolution #109 #165 #164 #161]: false
+unsat
+8d07b6d671ef5c9eedcaa35c22d5f1c7124504cb 66 0
+#2 := false
+decl f3 :: S1
+#8 := f3
+decl f1 :: S1
+#4 := f1
+#33 := (= f1 f3)
+#51 := (not #33)
+#87 := [hypothesis]: #33
+decl f4 :: S1
+#10 := f4
+#36 := (= f1 f4)
+#42 := (not #36)
+#43 := (or #33 #42)
+#69 := (or #43 #51)
+#70 := [def-axiom]: #69
+#88 := [unit-resolution #70 #87]: #43
+#67 := (not #43)
+#89 := (or #51 #67)
+#52 := (or #51 #36)
+#57 := (ite #33 #43 #52)
+#60 := (not #57)
+#11 := (= f4 f1)
+#9 := (= f3 f1)
+#13 := (implies #9 #11)
+#12 := (implies #11 #9)
+#14 := (ite #9 #12 #13)
+#15 := (not #14)
+#61 := (iff #15 #60)
+#58 := (iff #14 #57)
+#55 := (iff #13 #52)
+#48 := (implies #33 #36)
+#53 := (iff #48 #52)
+#54 := [rewrite]: #53
+#49 := (iff #13 #48)
+#37 := (iff #11 #36)
+#38 := [rewrite]: #37
+#34 := (iff #9 #33)
+#35 := [rewrite]: #34
+#50 := [monotonicity #35 #38]: #49
+#56 := [trans #50 #54]: #55
+#46 := (iff #12 #43)
+#39 := (implies #36 #33)
+#44 := (iff #39 #43)
+#45 := [rewrite]: #44
+#40 := (iff #12 #39)
+#41 := [monotonicity #38 #35]: #40
+#47 := [trans #41 #45]: #46
+#59 := [monotonicity #35 #47 #56]: #58
+#62 := [monotonicity #59]: #61
+#32 := [asserted]: #15
+#65 := [mp #32 #62]: #60
+#83 := (or #57 #51 #67)
+#84 := [def-axiom]: #83
+#90 := [unit-resolution #84 #65]: #89
+#91 := [unit-resolution #90 #88 #87]: false
+#92 := [lemma #91]: #51
+#63 := (or #52 #33)
+#73 := [def-axiom]: #63
+#93 := [unit-resolution #73 #92]: #52
+#76 := (not #52)
+#94 := (or #33 #76)
+#85 := (or #57 #33 #76)
+#86 := [def-axiom]: #85
+#95 := [unit-resolution #86 #65]: #94
+[unit-resolution #95 #93 #92]: false
+unsat
+415ad6a91bb75e131bb05d34f1965ab914825b25 67 0
+#2 := false
+decl f1 :: S1
+#4 := f1
+decl f6 :: S1
+#15 := f6
+#16 := (= f6 f1)
+decl f5 :: S1
+#13 := f5
+#14 := (= f5 f1)
+decl f4 :: S1
+#10 := f4
+#11 := (= f4 f1)
+#18 := (ite #11 #14 #16)
+decl f3 :: S1
+#8 := f3
+#9 := (= f3 f1)
+#19 := (ite #9 #14 #18)
+#12 := (or #9 #11)
+#17 := (ite #12 #14 #16)
+#20 := (iff #17 #19)
+#21 := (not #20)
+#79 := (iff #21 false)
+#1 := true
+#74 := (not true)
+#77 := (iff #74 false)
+#78 := [rewrite]: #77
+#75 := (iff #21 #74)
+#72 := (iff #20 true)
+#51 := (= f1 f6)
+#48 := (= f1 f5)
+#42 := (= f1 f4)
+#39 := (= f1 f3)
+#45 := (or #39 #42)
+#54 := (ite #45 #48 #51)
+#67 := (iff #54 #54)
+#70 := (iff #67 true)
+#71 := [rewrite]: #70
+#68 := (iff #20 #67)
+#65 := (iff #19 #54)
+#57 := (ite #42 #48 #51)
+#60 := (ite #39 #48 #57)
+#63 := (iff #60 #54)
+#64 := [rewrite]: #63
+#61 := (iff #19 #60)
+#58 := (iff #18 #57)
+#52 := (iff #16 #51)
+#53 := [rewrite]: #52
+#49 := (iff #14 #48)
+#50 := [rewrite]: #49
+#43 := (iff #11 #42)
+#44 := [rewrite]: #43
+#59 := [monotonicity #44 #50 #53]: #58
+#40 := (iff #9 #39)
+#41 := [rewrite]: #40
+#62 := [monotonicity #41 #50 #59]: #61
+#66 := [trans #62 #64]: #65
+#55 := (iff #17 #54)
+#46 := (iff #12 #45)
+#47 := [monotonicity #41 #44]: #46
+#56 := [monotonicity #47 #50 #53]: #55
+#69 := [monotonicity #56 #66]: #68
+#73 := [trans #69 #71]: #72
+#76 := [monotonicity #73]: #75
+#80 := [trans #76 #78]: #79
+#38 := [asserted]: #21
+[mp #38 #80]: false
+unsat
+8b8741030ee2ece2c8d6b3906a1951bc56d9b885 67 0
+#2 := false
+decl f1 :: S1
+#4 := f1
+decl f6 :: S1
+#15 := f6
+#16 := (= f6 f1)
+decl f5 :: S1
+#13 := f5
+#14 := (= f5 f1)
+decl f4 :: S1
+#10 := f4
+#11 := (= f4 f1)
+#18 := (ite #11 #14 #16)
+decl f3 :: S1
+#8 := f3
+#9 := (= f3 f1)
+#19 := (ite #9 #18 #16)
+#12 := (and #9 #11)
+#17 := (ite #12 #14 #16)
+#20 := (iff #17 #19)
+#21 := (not #20)
+#79 := (iff #21 false)
+#1 := true
+#74 := (not true)
+#77 := (iff #74 false)
+#78 := [rewrite]: #77
+#75 := (iff #21 #74)
+#72 := (iff #20 true)
+#51 := (= f1 f6)
+#48 := (= f1 f5)
+#42 := (= f1 f4)
+#39 := (= f1 f3)
+#45 := (and #39 #42)
+#54 := (ite #45 #48 #51)
+#67 := (iff #54 #54)
+#70 := (iff #67 true)
+#71 := [rewrite]: #70
+#68 := (iff #20 #67)
+#65 := (iff #19 #54)
+#57 := (ite #42 #48 #51)
+#60 := (ite #39 #57 #51)
+#63 := (iff #60 #54)
+#64 := [rewrite]: #63
+#61 := (iff #19 #60)
+#52 := (iff #16 #51)
+#53 := [rewrite]: #52
+#58 := (iff #18 #57)
+#49 := (iff #14 #48)
+#50 := [rewrite]: #49
+#43 := (iff #11 #42)
+#44 := [rewrite]: #43
+#59 := [monotonicity #44 #50 #53]: #58
+#40 := (iff #9 #39)
+#41 := [rewrite]: #40
+#62 := [monotonicity #41 #59 #53]: #61
+#66 := [trans #62 #64]: #65
+#55 := (iff #17 #54)
+#46 := (iff #12 #45)
+#47 := [monotonicity #41 #44]: #46
+#56 := [monotonicity #47 #50 #53]: #55
+#69 := [monotonicity #56 #66]: #68
+#73 := [trans #69 #71]: #72
+#76 := [monotonicity #73]: #75
+#80 := [trans #76 #78]: #79
+#38 := [asserted]: #21
+[mp #38 #80]: false
+unsat
+a0ca672c65c25c024851f833f39bfb2c705ec387 186 0
+#2 := false
+decl f5 :: S1
+#12 := f5
+decl f1 :: S1
+#4 := f1
+#47 := (= f1 f5)
+decl f3 :: S1
+#8 := f3
+#41 := (= f1 f3)
+#59 := (not #41)
+#76 := (or #59 #47)
+#119 := (not #76)
+decl f4 :: S1
+#10 := f4
+#44 := (= f1 f4)
+#68 := (or #59 #44)
+decl f6 :: S1
+#14 := f6
+#50 := (= f1 f6)
+#84 := (or #59 #50)
+#89 := (ite #68 #76 #84)
+#130 := (not #89)
+#53 := (ite #44 #47 #50)
+#60 := (or #59 #53)
+#112 := (not #44)
+#165 := [hypothesis]: #112
+#172 := (or #60 #44)
+#98 := (not #60)
+#163 := [hypothesis]: #98
+#148 := (or #60 #41)
+#149 := [def-axiom]: #148
+#164 := [unit-resolution #149 #163]: #41
+#124 := (not #50)
+#139 := (not #53)
+#150 := (or #60 #139)
+#151 := [def-axiom]: #150
+#166 := [unit-resolution #151 #163]: #139
+#146 := (or #53 #44 #124)
+#147 := [def-axiom]: #146
+#167 := [unit-resolution #147 #166 #165]: #124
+#157 := (or #89 #60)
+#99 := (iff #89 #98)
+#15 := (= f6 f1)
+#9 := (= f3 f1)
+#20 := (implies #9 #15)
+#13 := (= f5 f1)
+#19 := (implies #9 #13)
+#11 := (= f4 f1)
+#18 := (implies #9 #11)
+#21 := (ite #18 #19 #20)
+#16 := (ite #11 #13 #15)
+#17 := (implies #9 #16)
+#22 := (iff #17 #21)
+#23 := (not #22)
+#102 := (iff #23 #99)
+#92 := (iff #60 #89)
+#95 := (not #92)
+#100 := (iff #95 #99)
+#101 := [rewrite]: #100
+#96 := (iff #23 #95)
+#93 := (iff #22 #92)
+#90 := (iff #21 #89)
+#87 := (iff #20 #84)
+#81 := (implies #41 #50)
+#85 := (iff #81 #84)
+#86 := [rewrite]: #85
+#82 := (iff #20 #81)
+#51 := (iff #15 #50)
+#52 := [rewrite]: #51
+#42 := (iff #9 #41)
+#43 := [rewrite]: #42
+#83 := [monotonicity #43 #52]: #82
+#88 := [trans #83 #86]: #87
+#79 := (iff #19 #76)
+#73 := (implies #41 #47)
+#77 := (iff #73 #76)
+#78 := [rewrite]: #77
+#74 := (iff #19 #73)
+#48 := (iff #13 #47)
+#49 := [rewrite]: #48
+#75 := [monotonicity #43 #49]: #74
+#80 := [trans #75 #78]: #79
+#71 := (iff #18 #68)
+#65 := (implies #41 #44)
+#69 := (iff #65 #68)
+#70 := [rewrite]: #69
+#66 := (iff #18 #65)
+#45 := (iff #11 #44)
+#46 := [rewrite]: #45
+#67 := [monotonicity #43 #46]: #66
+#72 := [trans #67 #70]: #71
+#91 := [monotonicity #72 #80 #88]: #90
+#63 := (iff #17 #60)
+#56 := (implies #41 #53)
+#61 := (iff #56 #60)
+#62 := [rewrite]: #61
+#57 := (iff #17 #56)
+#54 := (iff #16 #53)
+#55 := [monotonicity #46 #49 #52]: #54
+#58 := [monotonicity #43 #55]: #57
+#64 := [trans #58 #62]: #63
+#94 := [monotonicity #64 #91]: #93
+#97 := [monotonicity #94]: #96
+#103 := [trans #97 #101]: #102
+#40 := [asserted]: #23
+#106 := [mp #40 #103]: #99
+#154 := (not #99)
+#155 := (or #89 #60 #154)
+#156 := [def-axiom]: #155
+#158 := [unit-resolution #156 #106]: #157
+#168 := [unit-resolution #158 #163]: #89
+#109 := (not #68)
+#107 := (or #109 #59 #44)
+#104 := [def-axiom]: #107
+#169 := [unit-resolution #104 #164 #165]: #109
+#133 := (or #130 #68 #84)
+#134 := [def-axiom]: #133
+#170 := [unit-resolution #134 #169 #168]: #84
+#127 := (not #84)
+#128 := (or #127 #59 #50)
+#129 := [def-axiom]: #128
+#171 := [unit-resolution #129 #170 #167 #164]: false
+#173 := [lemma #171]: #172
+#176 := [unit-resolution #173 #165]: #60
+#161 := (or #130 #98)
+#159 := (or #130 #98 #154)
+#160 := [def-axiom]: #159
+#162 := [unit-resolution #160 #106]: #161
+#182 := [unit-resolution #162 #176]: #130
+#180 := (or #84 #44)
+#174 := [hypothesis]: #127
+#125 := (or #84 #124)
+#126 := [def-axiom]: #125
+#175 := [unit-resolution #126 #174]: #124
+#122 := (or #84 #41)
+#123 := [def-axiom]: #122
+#177 := [unit-resolution #123 #174]: #41
+#152 := (or #98 #59 #53)
+#153 := [def-axiom]: #152
+#178 := [unit-resolution #153 #177 #176]: #53
+#142 := (or #139 #44 #50)
+#143 := [def-axiom]: #142
+#179 := [unit-resolution #143 #178 #175 #165]: false
+#181 := [lemma #179]: #180
+#183 := [unit-resolution #181 #165]: #84
+#137 := (or #89 #68 #127)
+#138 := [def-axiom]: #137
+#184 := [unit-resolution #138 #183 #182]: #68
+#135 := (or #89 #109 #119)
+#136 := [def-axiom]: #135
+#185 := [unit-resolution #136 #184 #182]: #119
+#186 := [unit-resolution #104 #184 #165]: #59
+#114 := (or #76 #41)
+#115 := [def-axiom]: #114
+#187 := [unit-resolution #115 #186 #185]: false
+#188 := [lemma #187]: #44
+#113 := (or #68 #112)
+#108 := [def-axiom]: #113
+#191 := [unit-resolution #108 #188]: #68
+#189 := [hypothesis]: #59
+#190 := [unit-resolution #149 #189]: #60
+#192 := [unit-resolution #115 #189]: #76
+#193 := [unit-resolution #136 #192 #191]: #89
+#194 := [unit-resolution #162 #193 #190]: false
+#195 := [lemma #194]: #41
+#116 := (not #47)
+#144 := (or #53 #112 #116)
+#145 := [def-axiom]: #144
+#196 := [unit-resolution #145 #166 #188]: #116
+#131 := (or #130 #109 #76)
+#132 := [def-axiom]: #131
+#197 := [unit-resolution #132 #168 #191]: #76
+#120 := (or #119 #59 #47)
+#121 := [def-axiom]: #120
+#198 := [unit-resolution #121 #197 #196 #195]: false
+#199 := [lemma #198]: #60
+#200 := [unit-resolution #162 #199]: #130
+#201 := [unit-resolution #136 #200 #191]: #119
+#202 := [unit-resolution #153 #199 #195]: #53
+#140 := (or #139 #112 #47)
+#141 := [def-axiom]: #140
+#203 := [unit-resolution #141 #202 #188]: #47
+#117 := (or #76 #116)
+#118 := [def-axiom]: #117
+[unit-resolution #118 #203 #201]: false
+unsat
+6bcac7ed9b4a6f3af0933f64ee0cb316b51a44ed 33 0
+#2 := false
+decl f1 :: S1
+#4 := f1
+decl f3 :: S1
+#8 := f3
+#9 := (= f3 f1)
+#10 := (not #9)
+#11 := (ite #9 #9 #10)
+#12 := (not #11)
+#48 := (iff #12 false)
+#1 := true
+#43 := (not true)
+#46 := (iff #43 false)
+#47 := [rewrite]: #46
+#44 := (iff #12 #43)
+#41 := (iff #11 true)
+#30 := (= f1 f3)
+#33 := (not #30)
+#36 := (ite #30 #30 #33)
+#39 := (iff #36 true)
+#40 := [rewrite]: #39
+#37 := (iff #11 #36)
+#34 := (iff #10 #33)
+#31 := (iff #9 #30)
+#32 := [rewrite]: #31
+#35 := [monotonicity #32]: #34
+#38 := [monotonicity #32 #32 #35]: #37
+#42 := [trans #38 #40]: #41
+#45 := [monotonicity #42]: #44
+#49 := [trans #45 #47]: #48
+#29 := [asserted]: #12
+[mp #29 #49]: false
+unsat
+be3ab8b5e09f8f942728ba1af0512891f56e0c5b 33 0
+#2 := false
+decl f1 :: S1
+#4 := f1
+decl f3 :: S1
+#8 := f3
+#9 := (= f3 f1)
+#10 := (not #9)
+#11 := (ite #10 #10 #9)
+#12 := (not #11)
+#48 := (iff #12 false)
+#1 := true
+#43 := (not true)
+#46 := (iff #43 false)
+#47 := [rewrite]: #46
+#44 := (iff #12 #43)
+#41 := (iff #11 true)
+#30 := (= f1 f3)
+#33 := (not #30)
+#36 := (ite #33 #33 #30)
+#39 := (iff #36 true)
+#40 := [rewrite]: #39
+#37 := (iff #11 #36)
+#31 := (iff #9 #30)
+#32 := [rewrite]: #31
+#34 := (iff #10 #33)
+#35 := [monotonicity #32]: #34
+#38 := [monotonicity #35 #35 #32]: #37
+#42 := [trans #38 #40]: #41
+#45 := [monotonicity #42]: #44
+#49 := [trans #45 #47]: #48
+#29 := [asserted]: #12
+[mp #29 #49]: false
+unsat
+aa78c56be2c8312419ab25323cf112e3e8809e9f 18 0
+#2 := false
+decl f3 :: S2
+#8 := f3
+#9 := (= f3 f3)
+#10 := (not #9)
+#35 := (iff #10 false)
+#1 := true
+#30 := (not true)
+#33 := (iff #30 false)
+#34 := [rewrite]: #33
+#31 := (iff #10 #30)
+#28 := (iff #9 true)
+#29 := [rewrite]: #28
+#32 := [monotonicity #29]: #31
+#36 := [trans #32 #34]: #35
+#27 := [asserted]: #10
+[mp #27 #36]: false
+unsat
+8d3a435e96c2d773a90e94525d09aa8a352d27d6 29 0
+#2 := false
+decl f3 :: S2
+#8 := f3
+decl f4 :: S2
+#9 := f4
+#11 := (= f4 f3)
+#10 := (= f3 f4)
+#12 := (implies #10 #11)
+#13 := (not #12)
+#45 := (iff #13 false)
+#1 := true
+#40 := (not true)
+#43 := (iff #40 false)
+#44 := [rewrite]: #43
+#41 := (iff #13 #40)
+#38 := (iff #12 true)
+#33 := (implies #10 #10)
+#36 := (iff #33 true)
+#37 := [rewrite]: #36
+#34 := (iff #12 #33)
+#31 := (iff #11 #10)
+#32 := [rewrite]: #31
+#35 := [monotonicity #32]: #34
+#39 := [trans #35 #37]: #38
+#42 := [monotonicity #39]: #41
+#46 := [trans #42 #44]: #45
+#30 := [asserted]: #13
+[mp #30 #46]: false
+unsat
+7d6e447cf6e1145f4f9a866bcc3cde4710e17c79 34 0
+#2 := false
+decl f5 :: S2
+#11 := f5
+decl f3 :: S2
+#8 := f3
+#14 := (= f3 f5)
+decl f4 :: S2
+#9 := f4
+#12 := (= f4 f5)
+#58 := (iff #12 #14)
+#56 := (iff #14 #12)
+#10 := (= f3 f4)
+#13 := (and #10 #12)
+#34 := (not #13)
+#35 := (or #34 #14)
+#38 := (not #35)
+#15 := (implies #13 #14)
+#16 := (not #15)
+#39 := (iff #16 #38)
+#36 := (iff #15 #35)
+#37 := [rewrite]: #36
+#40 := [monotonicity #37]: #39
+#33 := [asserted]: #16
+#43 := [mp #33 #40]: #38
+#41 := [not-or-elim #43]: #13
+#42 := [and-elim #41]: #10
+#57 := [monotonicity #42]: #56
+#59 := [symm #57]: #58
+#44 := [and-elim #41]: #12
+#54 := [mp #44 #59]: #14
+#45 := (not #14)
+#46 := [not-or-elim #43]: #45
+[unit-resolution #46 #54]: false
+unsat
+810792c4ea80c0bc3287d99f969f1a65dc039da5 27 0
+#2 := false
+decl f5 :: (-> S2 S2)
+decl f4 :: S2
+#9 := f4
+#12 := (f5 f4)
+decl f3 :: S2
+#8 := f3
+#11 := (f5 f3)
+#13 := (= #11 #12)
+#10 := (= f3 f4)
+#33 := (not #10)
+#34 := (or #33 #13)
+#37 := (not #34)
+#14 := (implies #10 #13)
+#15 := (not #14)
+#38 := (iff #15 #37)
+#35 := (iff #14 #34)
+#36 := [rewrite]: #35
+#39 := [monotonicity #36]: #38
+#32 := [asserted]: #15
+#42 := [mp #32 #39]: #37
+#40 := [not-or-elim #42]: #10
+#51 := [monotonicity #40]: #13
+#41 := (not #13)
+#43 := [not-or-elim #42]: #41
+[unit-resolution #43 #51]: false
+unsat
+09e564f98c53b63bab667092b7072fa2b674fa08 31 0
+#2 := false
+decl f5 :: (-> S2 S2 S3)
+decl f3 :: S2
+#8 := f3
+decl f4 :: S2
+#9 := f4
+#12 := (f5 f4 f3)
+#11 := (f5 f3 f4)
+#13 := (= #11 #12)
+#53 := (= #12 #11)
+#10 := (= f3 f4)
+#33 := (not #10)
+#34 := (or #33 #13)
+#37 := (not #34)
+#14 := (implies #10 #13)
+#15 := (not #14)
+#38 := (iff #15 #37)
+#35 := (iff #14 #34)
+#36 := [rewrite]: #35
+#39 := [monotonicity #36]: #38
+#32 := [asserted]: #15
+#42 := [mp #32 #39]: #37
+#40 := [not-or-elim #42]: #10
+#51 := (= f4 f3)
+#52 := [symm #40]: #51
+#54 := [monotonicity #52 #40]: #53
+#49 := [symm #54]: #13
+#41 := (not #13)
+#43 := [not-or-elim #42]: #41
+[unit-resolution #43 #49]: false
+unsat
+f97c055cc3b87ad54efe59b6c22e3e3633409f7e 63 0
+#2 := false
+decl f3 :: (-> S2 S2)
+decl f4 :: S2
+#8 := f4
+#9 := (f3 f4)
+#46 := (= f4 #9)
+#10 := (f3 #9)
+#12 := (f3 #10)
+#78 := (= #12 #9)
+#76 := (= #9 #12)
+#37 := (= f4 #10)
+#13 := (f3 #12)
+#14 := (f3 #13)
+#40 := (= f4 #14)
+#43 := (and #37 #40)
+#52 := (not #43)
+#53 := (or #52 #46)
+#58 := (not #53)
+#17 := (= #9 f4)
+#15 := (= #14 f4)
+#11 := (= #10 f4)
+#16 := (and #11 #15)
+#18 := (implies #16 #17)
+#19 := (not #18)
+#59 := (iff #19 #58)
+#56 := (iff #18 #53)
+#49 := (implies #43 #46)
+#54 := (iff #49 #53)
+#55 := [rewrite]: #54
+#50 := (iff #18 #49)
+#47 := (iff #17 #46)
+#48 := [rewrite]: #47
+#44 := (iff #16 #43)
+#41 := (iff #15 #40)
+#42 := [rewrite]: #41
+#38 := (iff #11 #37)
+#39 := [rewrite]: #38
+#45 := [monotonicity #39 #42]: #44
+#51 := [monotonicity #45 #48]: #50
+#57 := [trans #51 #55]: #56
+#60 := [monotonicity #57]: #59
+#36 := [asserted]: #19
+#63 := [mp #36 #60]: #58
+#61 := [not-or-elim #63]: #43
+#62 := [and-elim #61]: #37
+#77 := [monotonicity #62]: #76
+#79 := [symm #77]: #78
+#81 := (= f4 #12)
+#67 := (= #14 #12)
+#70 := (= #12 #14)
+#72 := (= #10 #13)
+#74 := (= #13 #10)
+#75 := [monotonicity #79]: #74
+#73 := [symm #75]: #72
+#71 := [monotonicity #73]: #70
+#80 := [symm #71]: #67
+#64 := [and-elim #61]: #40
+#82 := [trans #64 #80]: #81
+#83 := [trans #82 #79]: #46
+#65 := (not #46)
+#66 := [not-or-elim #63]: #65
+[unit-resolution #66 #83]: false
+unsat
+0b20a1cbaefe018fa6502d17013ec89dc657717b 195 0
+#2 := false
+decl f6 :: S2
+#13 := f6
+decl f4 :: S2
+#10 := f4
+#15 := (= f4 f6)
+decl f5 :: S2
+#11 := f5
+decl f3 :: S1
+#8 := f3
+decl f1 :: S1
+#4 := f1
+#40 := (= f1 f3)
+#43 := (ite #40 f4 f5)
+#49 := (= f6 #43)
+#200 := (iff #49 #15)
+#198 := (iff #15 #49)
+#46 := (= #43 f6)
+#50 := (iff #46 #49)
+#197 := [commutativity]: #50
+#195 := (iff #15 #46)
+#110 := (= f4 #43)
+#111 := (= f5 #43)
+#57 := (not #40)
+#180 := [hypothesis]: #57
+#114 := (or #40 #111)
+#115 := [def-axiom]: #114
+#184 := [unit-resolution #115 #180]: #111
+#185 := (= f6 f5)
+#18 := (= f5 f6)
+#174 := (iff #110 #15)
+#172 := (iff #15 #110)
+#68 := (or #18 #40)
+#95 := (not #68)
+#58 := (or #15 #57)
+#94 := (not #58)
+#96 := (or #94 #95)
+#123 := (not #18)
+#147 := [hypothesis]: #123
+#157 := (or #96 #18)
+#97 := (not #96)
+#145 := [hypothesis]: #97
+#132 := (or #96 #68)
+#133 := [def-axiom]: #132
+#148 := [unit-resolution #133 #145]: #68
+#128 := (or #95 #18 #40)
+#129 := [def-axiom]: #128
+#149 := [unit-resolution #129 #148 #147]: #40
+#112 := (or #57 #110)
+#113 := [def-axiom]: #112
+#150 := [unit-resolution #113 #149]: #110
+#153 := (= f6 f4)
+#130 := (or #96 #58)
+#131 := [def-axiom]: #130
+#151 := [unit-resolution #131 #145]: #58
+#121 := (or #94 #15 #57)
+#122 := [def-axiom]: #121
+#152 := [unit-resolution #122 #149 #151]: #15
+#154 := [symm #152]: #153
+#155 := [trans #154 #150]: #49
+#82 := (not #49)
+#143 := (or #82 #96)
+#104 := (iff #49 #96)
+#73 := (and #58 #68)
+#83 := (iff #73 #82)
+#107 := (iff #83 #104)
+#99 := (iff #96 #49)
+#105 := (iff #99 #104)
+#106 := [rewrite]: #105
+#102 := (iff #83 #99)
+#91 := (iff #97 #82)
+#100 := (iff #91 #99)
+#101 := [rewrite]: #100
+#88 := (iff #83 #91)
+#92 := (iff #73 #97)
+#93 := [rewrite]: #92
+#98 := [monotonicity #93]: #88
+#103 := [trans #98 #101]: #102
+#108 := [trans #103 #106]: #107
+#9 := (= f3 f1)
+#17 := (not #9)
+#19 := (implies #17 #18)
+#16 := (implies #9 #15)
+#20 := (and #16 #19)
+#12 := (ite #9 f4 f5)
+#14 := (= #12 f6)
+#21 := (iff #14 #20)
+#22 := (not #21)
+#86 := (iff #22 #83)
+#76 := (iff #49 #73)
+#79 := (not #76)
+#84 := (iff #79 #83)
+#85 := [rewrite]: #84
+#80 := (iff #22 #79)
+#77 := (iff #21 #76)
+#74 := (iff #20 #73)
+#71 := (iff #19 #68)
+#65 := (implies #57 #18)
+#69 := (iff #65 #68)
+#70 := [rewrite]: #69
+#66 := (iff #19 #65)
+#63 := (iff #17 #57)
+#41 := (iff #9 #40)
+#42 := [rewrite]: #41
+#64 := [monotonicity #42]: #63
+#67 := [monotonicity #64]: #66
+#72 := [trans #67 #70]: #71
+#61 := (iff #16 #58)
+#54 := (implies #40 #15)
+#59 := (iff #54 #58)
+#60 := [rewrite]: #59
+#55 := (iff #16 #54)
+#56 := [monotonicity #42]: #55
+#62 := [trans #56 #60]: #61
+#75 := [monotonicity #62 #72]: #74
+#52 := (iff #14 #49)
+#51 := [rewrite]: #50
+#47 := (iff #14 #46)
+#44 := (= #12 #43)
+#45 := [monotonicity #42]: #44
+#48 := [monotonicity #45]: #47
+#53 := [trans #48 #51]: #52
+#78 := [monotonicity #53 #75]: #77
+#81 := [monotonicity #78]: #80
+#87 := [trans #81 #85]: #86
+#39 := [asserted]: #22
+#90 := [mp #39 #87]: #83
+#109 := [mp #90 #108]: #104
+#136 := (not #104)
+#141 := (or #82 #96 #136)
+#142 := [def-axiom]: #141
+#144 := [unit-resolution #142 #109]: #143
+#146 := [unit-resolution #144 #145]: #82
+#156 := [unit-resolution #146 #155]: false
+#158 := [lemma #156]: #157
+#159 := [unit-resolution #158 #147]: #96
+#139 := (or #49 #97)
+#137 := (or #49 #97 #136)
+#138 := [def-axiom]: #137
+#140 := [unit-resolution #138 #109]: #139
+#160 := [unit-resolution #140 #159]: #49
+#173 := [monotonicity #160]: #172
+#175 := [symm #173]: #174
+#163 := (not #111)
+#164 := (iff #123 #163)
+#161 := (iff #18 #111)
+#162 := [monotonicity #160]: #161
+#165 := [monotonicity #162]: #164
+#166 := [mp #147 #165]: #163
+#167 := [unit-resolution #115 #166]: #40
+#171 := [unit-resolution #113 #167]: #110
+#176 := [mp #171 #175]: #15
+#116 := (not #15)
+#126 := (or #68 #57)
+#127 := [def-axiom]: #126
+#168 := [unit-resolution #127 #167]: #68
+#134 := (or #97 #94 #95)
+#135 := [def-axiom]: #134
+#169 := [unit-resolution #135 #168 #159]: #94
+#117 := (or #58 #116)
+#118 := [def-axiom]: #117
+#170 := [unit-resolution #118 #169]: #116
+#177 := [unit-resolution #170 #176]: false
+#178 := [lemma #177]: #18
+#186 := [symm #178]: #185
+#187 := [trans #186 #184]: #49
+#124 := (or #68 #123)
+#125 := [def-axiom]: #124
+#179 := [unit-resolution #125 #178]: #68
+#119 := (or #58 #40)
+#120 := [def-axiom]: #119
+#181 := [unit-resolution #120 #180]: #58
+#182 := [unit-resolution #135 #181 #179]: #97
+#183 := [unit-resolution #144 #182]: #82
+#188 := [unit-resolution #183 #187]: false
+#189 := [lemma #188]: #40
+#194 := [unit-resolution #113 #189]: #110
+#196 := [monotonicity #194]: #195
+#199 := [trans #196 #197]: #198
+#201 := [symm #199]: #200
+#202 := (iff #82 #116)
+#203 := [monotonicity #201]: #202
+#190 := [hypothesis]: #82
+#204 := [mp #190 #203]: #116
+#191 := [unit-resolution #140 #190]: #97
+#192 := [unit-resolution #131 #191]: #58
+#193 := [unit-resolution #122 #192 #189]: #15
+#205 := [unit-resolution #193 #204]: false
+#206 := [lemma #205]: #49
+#210 := [mp #206 #201]: #15
+#207 := [unit-resolution #144 #206]: #96
+#208 := [unit-resolution #135 #207 #179]: #94
+#209 := [unit-resolution #118 #208]: #116
+[unit-resolution #209 #210]: false
+unsat
+ba53dbe39a10406601e0a359b3a1de14129ef1a1 36 0
+#2 := false
+decl f5 :: S2
+#10 := f5
+decl f3 :: S2
+#8 := f3
+#12 := (= f3 f5)
+#13 := (not #12)
+decl f4 :: S2
+#9 := f4
+#11 := (distinct f3 f4 f5)
+#33 := (not #11)
+#34 := (or #33 #13)
+#37 := (not #34)
+#14 := (implies #11 #13)
+#15 := (not #14)
+#38 := (iff #15 #37)
+#35 := (iff #14 #34)
+#36 := [rewrite]: #35
+#39 := [monotonicity #36]: #38
+#32 := [asserted]: #15
+#42 := [mp #32 #39]: #37
+#41 := [not-or-elim #42]: #12
+#52 := (= f4 f5)
+#53 := (not #52)
+#50 := (= f3 f4)
+#51 := (not #50)
+#48 := (and #51 #13 #53)
+#40 := [not-or-elim #42]: #11
+#58 := (or #33 #48)
+#59 := [def-axiom]: #58
+#62 := [unit-resolution #59 #40]: #48
+#49 := (not #48)
+#45 := (or #49 #13)
+#43 := [def-axiom]: #45
+[unit-resolution #43 #62 #41]: false
+unsat
+08e0778705271e8f33e9bdf3c26b749a4c39eda2 77 0
+#2 := false
+decl f4 :: S2
+#9 := f4
+decl f3 :: S2
+#8 := f3
+#81 := (= f3 f4)
+decl f6 :: S2
+#12 := f6
+#36 := (= f4 f6)
+#100 := (iff #36 #81)
+#98 := (iff #81 #36)
+#13 := (= f6 f4)
+#37 := (iff #13 #36)
+#97 := [commutativity]: #37
+#95 := (iff #81 #13)
+#14 := (= f3 f6)
+#42 := (not #36)
+#15 := (not #14)
+decl f5 :: S2
+#10 := f5
+#11 := (distinct f3 f4 f5)
+#51 := (not #11)
+#60 := (or #51 #15 #42)
+#63 := (not #60)
+#16 := (implies #13 #15)
+#17 := (implies #11 #16)
+#18 := (not #17)
+#66 := (iff #18 #63)
+#43 := (or #15 #42)
+#52 := (or #51 #43)
+#57 := (not #52)
+#64 := (iff #57 #63)
+#61 := (iff #52 #60)
+#62 := [rewrite]: #61
+#65 := [monotonicity #62]: #64
+#58 := (iff #18 #57)
+#55 := (iff #17 #52)
+#48 := (implies #11 #43)
+#53 := (iff #48 #52)
+#54 := [rewrite]: #53
+#49 := (iff #17 #48)
+#46 := (iff #16 #43)
+#39 := (implies #36 #15)
+#44 := (iff #39 #43)
+#45 := [rewrite]: #44
+#40 := (iff #16 #39)
+#38 := [rewrite]: #37
+#41 := [monotonicity #38]: #40
+#47 := [trans #41 #45]: #46
+#50 := [monotonicity #47]: #49
+#56 := [trans #50 #54]: #55
+#59 := [monotonicity #56]: #58
+#67 := [trans #59 #65]: #66
+#35 := [asserted]: #18
+#68 := [mp #35 #67]: #63
+#70 := [not-or-elim #68]: #14
+#96 := [monotonicity #70]: #95
+#99 := [trans #96 #97]: #98
+#101 := [symm #99]: #100
+#71 := [not-or-elim #68]: #36
+#102 := [mp #71 #101]: #81
+#82 := (not #81)
+#79 := (= f4 f5)
+#80 := (not #79)
+#83 := (= f3 f5)
+#84 := (not #83)
+#77 := (and #82 #84 #80)
+#69 := [not-or-elim #68]: #11
+#89 := (or #51 #77)
+#90 := [def-axiom]: #89
+#93 := [unit-resolution #90 #69]: #77
+#78 := (not #77)
+#75 := (or #78 #82)
+#76 := [def-axiom]: #75
+#94 := [unit-resolution #76 #93]: #82
+[unit-resolution #94 #102]: false
+unsat
+32fbe7827173348f5f83cc50b7e41574fd04dba5 28 0
+#2 := false
+decl f4 :: S2
+#9 := f4
+decl f3 :: S2
+#8 := f3
+#10 := (distinct f3 f4 f3 f4)
+#11 := (not #10)
+#12 := (not #11)
+#44 := (iff #12 false)
+#1 := true
+#39 := (not true)
+#42 := (iff #39 false)
+#43 := [rewrite]: #42
+#40 := (iff #12 #39)
+#37 := (iff #11 true)
+#32 := (not false)
+#35 := (iff #32 true)
+#36 := [rewrite]: #35
+#33 := (iff #11 #32)
+#30 := (iff #10 false)
+#31 := [rewrite]: #30
+#34 := [monotonicity #31]: #33
+#38 := [trans #34 #36]: #37
+#41 := [monotonicity #38]: #40
+#45 := [trans #41 #43]: #44
+#29 := [asserted]: #12
+[mp #29 #45]: false
+unsat
+00d94d0af0e9abf6798a28ecb39a441c7c299833 30 0
+#2 := false
+decl f4 :: S2
+#9 := f4
+decl f3 :: S2
+#8 := f3
+#10 := (= f3 f4)
+#11 := (not #10)
+#12 := (not #11)
+#13 := (implies #10 #12)
+#14 := (not #13)
+#46 := (iff #14 false)
+#1 := true
+#41 := (not true)
+#44 := (iff #41 false)
+#45 := [rewrite]: #44
+#42 := (iff #14 #41)
+#39 := (iff #13 true)
+#34 := (implies #10 #10)
+#37 := (iff #34 true)
+#38 := [rewrite]: #37
+#35 := (iff #13 #34)
+#32 := (iff #12 #10)
+#33 := [rewrite]: #32
+#36 := [monotonicity #33]: #35
+#40 := [trans #36 #38]: #39
+#43 := [monotonicity #40]: #42
+#47 := [trans #43 #45]: #46
+#31 := [asserted]: #14
+[mp #31 #47]: false
+unsat
+11d1529c7a7d7cf3014e0031300d5be1524ea6df 40 0
+#2 := false
+decl f5 :: S2
+#11 := f5
+decl f4 :: S2
+#9 := f4
+decl f3 :: S2
+#8 := f3
+#14 := (distinct f3 f4 f5)
+#15 := (not #14)
+#12 := (= f3 f5)
+#10 := (= f3 f4)
+#13 := (and #10 #12)
+#35 := (not #13)
+#36 := (or #35 #15)
+#39 := (not #36)
+#16 := (implies #13 #15)
+#17 := (not #16)
+#40 := (iff #17 #39)
+#37 := (iff #16 #36)
+#38 := [rewrite]: #37
+#41 := [monotonicity #38]: #40
+#34 := [asserted]: #17
+#44 := [mp #34 #41]: #39
+#46 := [not-or-elim #44]: #14
+#58 := (= f4 f5)
+#59 := (not #58)
+#57 := (not #12)
+#56 := (not #10)
+#54 := (and #56 #57 #59)
+#55 := (not #54)
+#42 := [not-or-elim #44]: #13
+#43 := [and-elim #42]: #10
+#52 := (or #55 #56)
+#53 := [def-axiom]: #52
+#66 := [unit-resolution #53 #43]: #55
+#62 := (or #15 #54)
+#63 := [def-axiom]: #62
+#67 := [unit-resolution #63 #66]: #15
+[unit-resolution #67 #46]: false
+unsat
+8585364a9de57f2c76854513f54ec42d9a599f64 33 0
+#2 := false
+decl f3 :: S2
+#8 := f3
+decl f5 :: S2
+#10 := f5
+decl f4 :: S2
+#9 := f4
+decl f6 :: S2
+#11 := f6
+#13 := (distinct f6 f4 f5 f3)
+#12 := (distinct f3 f4 f5 f6)
+#14 := (implies #12 #13)
+#15 := (not #14)
+#47 := (iff #15 false)
+#1 := true
+#42 := (not true)
+#45 := (iff #42 false)
+#46 := [rewrite]: #45
+#43 := (iff #15 #42)
+#40 := (iff #14 true)
+#35 := (implies #12 #12)
+#38 := (iff #35 true)
+#39 := [rewrite]: #38
+#36 := (iff #14 #35)
+#33 := (iff #13 #12)
+#34 := [rewrite]: #33
+#37 := [monotonicity #34]: #36
+#41 := [trans #37 #39]: #40
+#44 := [monotonicity #41]: #43
+#48 := [trans #44 #46]: #47
+#32 := [asserted]: #15
+[mp #32 #48]: false
+unsat
+a785d8d462362aedcc845e3ec87cbb8d13b87ae7 91 0
+#2 := false
+decl f5 :: S2
+#10 := f5
+decl f4 :: S2
+#9 := f4
+decl f3 :: S2
+#8 := f3
+#13 := (distinct f3 f4 f5)
+#67 := (= f4 f5)
+#68 := (not #67)
+#63 := (= f3 f5)
+#64 := (not #63)
+#61 := (= f3 f4)
+#62 := (not #61)
+#93 := (and #62 #64 #68)
+decl f6 :: S2
+#11 := f6
+#71 := (= f5 f6)
+#72 := (not #71)
+#69 := (= f4 f6)
+#70 := (not #69)
+#65 := (= f3 f6)
+#66 := (not #65)
+#73 := (and #62 #64 #66 #68 #70 #72)
+#12 := (distinct f3 f4 f5 f6)
+#14 := (distinct f4 f5 f6)
+#15 := (and #13 #14)
+#35 := (not #12)
+#36 := (or #35 #15)
+#39 := (not #36)
+#16 := (implies #12 #15)
+#17 := (not #16)
+#40 := (iff #17 #39)
+#37 := (iff #16 #36)
+#38 := [rewrite]: #37
+#41 := [monotonicity #38]: #40
+#34 := [asserted]: #17
+#44 := [mp #34 #41]: #39
+#42 := [not-or-elim #44]: #12
+#89 := (or #35 #73)
+#90 := [def-axiom]: #89
+#121 := [unit-resolution #90 #42]: #73
+#74 := (not #73)
+#75 := (or #74 #62)
+#76 := [def-axiom]: #75
+#122 := [unit-resolution #76 #121]: #62
+#81 := (or #74 #68)
+#82 := [def-axiom]: #81
+#123 := [unit-resolution #82 #121]: #68
+#77 := (or #74 #64)
+#78 := [def-axiom]: #77
+#124 := [unit-resolution #78 #121]: #64
+#101 := (or #93 #61 #63 #67)
+#102 := [def-axiom]: #101
+#125 := [unit-resolution #102 #124 #123 #122]: #93
+#94 := (not #93)
+#105 := (or #13 #94)
+#106 := [def-axiom]: #105
+#126 := [unit-resolution #106 #125]: #13
+#107 := (and #68 #70 #72)
+#85 := (or #74 #72)
+#86 := [def-axiom]: #85
+#127 := [unit-resolution #86 #121]: #72
+#83 := (or #74 #70)
+#84 := [def-axiom]: #83
+#128 := [unit-resolution #84 #121]: #70
+#115 := (or #107 #67 #69 #71)
+#116 := [def-axiom]: #115
+#129 := [unit-resolution #116 #128 #127 #123]: #107
+#108 := (not #107)
+#119 := (or #14 #108)
+#120 := [def-axiom]: #119
+#130 := [unit-resolution #120 #129]: #14
+#54 := (not #14)
+#53 := (not #13)
+#55 := (or #53 #54)
+#43 := (not #15)
+#58 := (iff #43 #55)
+#56 := (not #55)
+#49 := (not #56)
+#46 := (iff #49 #55)
+#57 := [rewrite]: #46
+#50 := (iff #43 #49)
+#51 := (iff #15 #56)
+#52 := [rewrite]: #51
+#48 := [monotonicity #52]: #50
+#59 := [trans #48 #57]: #58
+#45 := [not-or-elim #44]: #43
+#60 := [mp #45 #59]: #55
+[unit-resolution #60 #130 #126]: false
+unsat
+273a6dc6f10616be4347838d6cbb54a0b0ad5636 25 0
+#2 := false
+#8 := (:var 0 S2)
+#9 := (= #8 #8)
+#10 := (forall (vars (?v0 S2)) #9)
+#11 := (not #10)
+#43 := (iff #11 false)
+#1 := true
+#38 := (not true)
+#41 := (iff #38 false)
+#42 := [rewrite]: #41
+#39 := (iff #11 #38)
+#36 := (iff #10 true)
+#31 := (forall (vars (?v0 S2)) true)
+#34 := (iff #31 true)
+#35 := [elim-unused]: #34
+#32 := (iff #10 #31)
+#29 := (iff #9 true)
+#30 := [rewrite]: #29
+#33 := [quant-intro #30]: #32
+#37 := [trans #33 #35]: #36
+#40 := [monotonicity #37]: #39
+#44 := [trans #40 #42]: #43
+#28 := [asserted]: #11
+[mp #28 #44]: false
+unsat
+2b9ce0054ea55ee81c55715bd6d24aa6896b491c 34 0
+#2 := false
+decl f1 :: S1
+#4 := f1
+decl f3 :: (-> S2 S1)
+#8 := (:var 0 S2)
+#9 := (f3 #8)
+#10 := (= #9 f1)
+#11 := (forall (vars (?v0 S2)) #10)
+#12 := (iff #11 #11)
+#13 := (not #12)
+#49 := (iff #13 false)
+#1 := true
+#44 := (not true)
+#47 := (iff #44 false)
+#48 := [rewrite]: #47
+#45 := (iff #13 #44)
+#42 := (iff #12 true)
+#31 := (= f1 #9)
+#34 := (forall (vars (?v0 S2)) #31)
+#37 := (iff #34 #34)
+#40 := (iff #37 true)
+#41 := [rewrite]: #40
+#38 := (iff #12 #37)
+#35 := (iff #11 #34)
+#32 := (iff #10 #31)
+#33 := [rewrite]: #32
+#36 := [quant-intro #33]: #35
+#39 := [monotonicity #36 #36]: #38
+#43 := [trans #39 #41]: #42
+#46 := [monotonicity #43]: #45
+#50 := [trans #46 #48]: #49
+#30 := [asserted]: #13
+[mp #30 #50]: false
+unsat
+4905f6794c524726db3be755d752d3ea53fc4f9a 95 0
 #2 := false
 decl f3 :: (-> S2 S1)
 decl ?v0!0 :: S2
@@ -4185,68 +4677,7 @@
 #102 := [and-elim #101]: #76
 [unit-resolution #102 #115]: false
 unsat
-f23e61ff98637733b3412f875d5b578e0e8387fe 60 0
-#2 := false
-decl f1 :: S1
-#4 := f1
-decl f3 :: S1
-#8 := f3
-#9 := (= f3 f1)
-#14 := (not #9)
-decl f4 :: S1
-#10 := f4
-#11 := (= f4 f1)
-#13 := (not #11)
-#15 := (implies #13 #14)
-#12 := (implies #9 #11)
-#16 := (implies #12 #15)
-#17 := (not #16)
-#74 := (iff #17 false)
-#1 := true
-#69 := (not true)
-#72 := (iff #69 false)
-#73 := [rewrite]: #72
-#70 := (iff #17 #69)
-#67 := (iff #16 true)
-#38 := (= f1 f4)
-#35 := (= f1 f3)
-#44 := (not #35)
-#45 := (or #44 #38)
-#62 := (implies #45 #45)
-#65 := (iff #62 true)
-#66 := [rewrite]: #65
-#63 := (iff #16 #62)
-#60 := (iff #15 #45)
-#50 := (not #38)
-#55 := (implies #50 #44)
-#58 := (iff #55 #45)
-#59 := [rewrite]: #58
-#56 := (iff #15 #55)
-#53 := (iff #14 #44)
-#36 := (iff #9 #35)
-#37 := [rewrite]: #36
-#54 := [monotonicity #37]: #53
-#51 := (iff #13 #50)
-#39 := (iff #11 #38)
-#40 := [rewrite]: #39
-#52 := [monotonicity #40]: #51
-#57 := [monotonicity #52 #54]: #56
-#61 := [trans #57 #59]: #60
-#48 := (iff #12 #45)
-#41 := (implies #35 #38)
-#46 := (iff #41 #45)
-#47 := [rewrite]: #46
-#42 := (iff #12 #41)
-#43 := [monotonicity #37 #40]: #42
-#49 := [trans #43 #47]: #48
-#64 := [monotonicity #49 #61]: #63
-#68 := [trans #64 #66]: #67
-#71 := [monotonicity #68]: #70
-#75 := [trans #71 #73]: #74
-#34 := [asserted]: #17
-[mp #34 #75]: false
-unsat
-99ac0be7380191db43f2261910b85b072853ffc6 300 0
+48a51f989350a0e584db03a2950a272906be29b0 300 0
 #2 := false
 decl f4 :: (-> S2 S1)
 decl ?v0!1 :: S2
@@ -4547,129 +4978,7 @@
 #609 := [quant-inst]: #615
 [unit-resolution #609 #257 #603]: false
 unsat
-a8bd8ccd5c190cb3f1fa752363a309006c96634d 121 0
-#2 := false
-decl f5 :: S1
-#12 := f5
-decl f1 :: S1
-#4 := f1
-#44 := (= f1 f5)
-decl f4 :: S1
-#10 := f4
-#41 := (= f1 f4)
-decl f3 :: S1
-#8 := f3
-#38 := (= f1 f3)
-#53 := (not #38)
-#90 := (or #53 #41 #44)
-#137 := (iff #90 false)
-#132 := (or false false false)
-#135 := (iff #132 false)
-#136 := [rewrite]: #135
-#133 := (iff #90 #132)
-#123 := (iff #44 false)
-#113 := (not #44)
-#93 := (not #90)
-#99 := (or #53 #41 #44 #93)
-#104 := (not #99)
-#13 := (= f5 f1)
-#9 := (= f3 f1)
-#17 := (implies #9 #13)
-#11 := (= f4 f1)
-#16 := (implies #9 #11)
-#18 := (or #16 #17)
-#14 := (or #11 #13)
-#15 := (implies #9 #14)
-#19 := (implies #15 #18)
-#20 := (not #19)
-#107 := (iff #20 #104)
-#70 := (or #53 #44)
-#62 := (or #53 #41)
-#75 := (or #62 #70)
-#47 := (or #41 #44)
-#54 := (or #53 #47)
-#81 := (not #54)
-#82 := (or #81 #75)
-#87 := (not #82)
-#105 := (iff #87 #104)
-#102 := (iff #82 #99)
-#96 := (or #93 #75)
-#100 := (iff #96 #99)
-#101 := [rewrite]: #100
-#97 := (iff #82 #96)
-#94 := (iff #81 #93)
-#91 := (iff #54 #90)
-#92 := [rewrite]: #91
-#95 := [monotonicity #92]: #94
-#98 := [monotonicity #95]: #97
-#103 := [trans #98 #101]: #102
-#106 := [monotonicity #103]: #105
-#88 := (iff #20 #87)
-#85 := (iff #19 #82)
-#78 := (implies #54 #75)
-#83 := (iff #78 #82)
-#84 := [rewrite]: #83
-#79 := (iff #19 #78)
-#76 := (iff #18 #75)
-#73 := (iff #17 #70)
-#67 := (implies #38 #44)
-#71 := (iff #67 #70)
-#72 := [rewrite]: #71
-#68 := (iff #17 #67)
-#45 := (iff #13 #44)
-#46 := [rewrite]: #45
-#39 := (iff #9 #38)
-#40 := [rewrite]: #39
-#69 := [monotonicity #40 #46]: #68
-#74 := [trans #69 #72]: #73
-#65 := (iff #16 #62)
-#59 := (implies #38 #41)
-#63 := (iff #59 #62)
-#64 := [rewrite]: #63
-#60 := (iff #16 #59)
-#42 := (iff #11 #41)
-#43 := [rewrite]: #42
-#61 := [monotonicity #40 #43]: #60
-#66 := [trans #61 #64]: #65
-#77 := [monotonicity #66 #74]: #76
-#57 := (iff #15 #54)
-#50 := (implies #38 #47)
-#55 := (iff #50 #54)
-#56 := [rewrite]: #55
-#51 := (iff #15 #50)
-#48 := (iff #14 #47)
-#49 := [monotonicity #43 #46]: #48
-#52 := [monotonicity #40 #49]: #51
-#58 := [trans #52 #56]: #57
-#80 := [monotonicity #58 #77]: #79
-#86 := [trans #80 #84]: #85
-#89 := [monotonicity #86]: #88
-#108 := [trans #89 #106]: #107
-#37 := [asserted]: #20
-#109 := [mp #37 #108]: #104
-#114 := [not-or-elim #109]: #113
-#124 := [iff-false #114]: #123
-#121 := (iff #41 false)
-#111 := (not #41)
-#112 := [not-or-elim #109]: #111
-#122 := [iff-false #112]: #121
-#130 := (iff #53 false)
-#1 := true
-#125 := (not true)
-#128 := (iff #125 false)
-#129 := [rewrite]: #128
-#126 := (iff #53 #125)
-#119 := (iff #38 true)
-#110 := [not-or-elim #109]: #38
-#120 := [iff-true #110]: #119
-#127 := [monotonicity #120]: #126
-#131 := [trans #127 #129]: #130
-#134 := [monotonicity #131 #122 #124]: #133
-#138 := [trans #134 #136]: #137
-#115 := [not-or-elim #109]: #90
-[mp #115 #138]: false
-unsat
-cadcb33e2fa210fd60afd1f576e2a8b43f4a3cdf 201 0
+8d1c135b088bf498c079d2958e8764a8092679cf 201 0
 #2 := false
 decl f3 :: (-> S2 S1)
 #8 := (:var 0 S2)
@@ -4871,92 +5180,7 @@
 #211 := [quant-inst]: #320
 [unit-resolution #211 #595 #319]: false
 unsat
-49f39d63404aef857c81d69ce3ed08db1cf19a62 84 0
-#2 := false
-decl f4 :: S1
-#10 := f4
-decl f1 :: S1
-#4 := f1
-#38 := (= f1 f4)
-#53 := (not #38)
-#97 := [hypothesis]: #53
-decl f3 :: S1
-#8 := f3
-#35 := (= f1 f3)
-#44 := (not #35)
-#45 := (or #44 #38)
-#54 := (or #35 #53)
-#59 := (and #45 #54)
-#62 := (iff #35 #38)
-#68 := (not #59)
-#69 := (or #68 #62)
-#74 := (not #69)
-#11 := (= f4 f1)
-#9 := (= f3 f1)
-#15 := (iff #9 #11)
-#13 := (implies #11 #9)
-#12 := (implies #9 #11)
-#14 := (and #12 #13)
-#16 := (implies #14 #15)
-#17 := (not #16)
-#75 := (iff #17 #74)
-#72 := (iff #16 #69)
-#65 := (implies #59 #62)
-#70 := (iff #65 #69)
-#71 := [rewrite]: #70
-#66 := (iff #16 #65)
-#63 := (iff #15 #62)
-#39 := (iff #11 #38)
-#40 := [rewrite]: #39
-#36 := (iff #9 #35)
-#37 := [rewrite]: #36
-#64 := [monotonicity #37 #40]: #63
-#60 := (iff #14 #59)
-#57 := (iff #13 #54)
-#50 := (implies #38 #35)
-#55 := (iff #50 #54)
-#56 := [rewrite]: #55
-#51 := (iff #13 #50)
-#52 := [monotonicity #40 #37]: #51
-#58 := [trans #52 #56]: #57
-#48 := (iff #12 #45)
-#41 := (implies #35 #38)
-#46 := (iff #41 #45)
-#47 := [rewrite]: #46
-#42 := (iff #12 #41)
-#43 := [monotonicity #37 #40]: #42
-#49 := [trans #43 #47]: #48
-#61 := [monotonicity #49 #58]: #60
-#67 := [monotonicity #61 #64]: #66
-#73 := [trans #67 #71]: #72
-#76 := [monotonicity #73]: #75
-#34 := [asserted]: #17
-#79 := [mp #34 #76]: #74
-#77 := [not-or-elim #79]: #59
-#78 := [and-elim #77]: #45
-#98 := [unit-resolution #78 #97]: #44
-#89 := (or #38 #35)
-#90 := (iff #38 #44)
-#81 := (not #62)
-#91 := (iff #81 #90)
-#92 := [rewrite]: #91
-#82 := [not-or-elim #79]: #81
-#93 := [mp #82 #92]: #90
-#94 := (not #90)
-#95 := (or #38 #35 #94)
-#88 := [def-axiom]: #95
-#86 := [unit-resolution #88 #93]: #89
-#99 := [unit-resolution #86 #98 #97]: false
-#100 := [lemma #99]: #38
-#80 := [and-elim #77]: #54
-#101 := [unit-resolution #80 #100]: #35
-#83 := (or #53 #44)
-#87 := (or #53 #44 #94)
-#85 := [def-axiom]: #87
-#96 := [unit-resolution #85 #93]: #83
-[unit-resolution #96 #101 #100]: false
-unsat
-275457b2bda52ffcdd4253789835ed95d90c1f17 244 0
+c2e19ee7461e4fbfe65c34d98111e3914f852baf 244 0
 #2 := false
 decl f3 :: (-> S2 S1)
 decl ?v0!0 :: S2
@@ -5201,55 +5425,7 @@
 #616 := [mp #343 #237]: #621
 [unit-resolution #616 #629 #330]: false
 unsat
-bbcd2845c0842c2c58bd699481f0fecfa298d8a7 47 0
-#2 := false
-decl f1 :: S1
-#4 := f1
-decl f3 :: S1
-#8 := f3
-#9 := (= f3 f1)
-decl f4 :: S1
-#10 := f4
-#11 := (= f4 f1)
-#13 := (iff #11 #9)
-#12 := (iff #9 #11)
-#14 := (iff #12 #13)
-#15 := (not #14)
-#61 := (iff #15 false)
-#1 := true
-#56 := (not true)
-#59 := (iff #56 false)
-#60 := [rewrite]: #59
-#57 := (iff #15 #56)
-#54 := (iff #14 true)
-#36 := (= f1 f4)
-#33 := (= f1 f3)
-#39 := (iff #33 #36)
-#49 := (iff #39 #39)
-#52 := (iff #49 true)
-#53 := [rewrite]: #52
-#50 := (iff #14 #49)
-#47 := (iff #13 #39)
-#42 := (iff #36 #33)
-#45 := (iff #42 #39)
-#46 := [rewrite]: #45
-#43 := (iff #13 #42)
-#34 := (iff #9 #33)
-#35 := [rewrite]: #34
-#37 := (iff #11 #36)
-#38 := [rewrite]: #37
-#44 := [monotonicity #38 #35]: #43
-#48 := [trans #44 #46]: #47
-#40 := (iff #12 #39)
-#41 := [monotonicity #35 #38]: #40
-#51 := [monotonicity #41 #48]: #50
-#55 := [trans #51 #53]: #54
-#58 := [monotonicity #55]: #57
-#62 := [trans #58 #60]: #61
-#32 := [asserted]: #15
-[mp #32 #62]: false
-unsat
-b11d7db349f0b800b6d28726c33bbc51dfd5ec41 49 0
+35b6826e84c97e2db1e3cd91ad69abb16f131b93 49 0
 #2 := false
 decl f1 :: S1
 #4 := f1
@@ -5299,49 +5475,7 @@
 #35 := [asserted]: #18
 [mp #35 #65]: false
 unsat
-d6ad6bdc191dec543b99c99d75040cdba1ff04e8 41 0
-#2 := false
-decl f1 :: S1
-#4 := f1
-decl f3 :: S1
-#8 := f3
-#9 := (= f3 f1)
-#10 := (not #9)
-#11 := (iff #9 #10)
-#12 := (not #11)
-#13 := (not #12)
-#56 := (iff #13 false)
-#1 := true
-#51 := (not true)
-#54 := (iff #51 false)
-#55 := [rewrite]: #54
-#52 := (iff #13 #51)
-#49 := (iff #12 true)
-#44 := (not false)
-#47 := (iff #44 true)
-#48 := [rewrite]: #47
-#45 := (iff #12 #44)
-#42 := (iff #11 false)
-#31 := (= f1 f3)
-#34 := (not #31)
-#37 := (iff #31 #34)
-#40 := (iff #37 false)
-#41 := [rewrite]: #40
-#38 := (iff #11 #37)
-#35 := (iff #10 #34)
-#32 := (iff #9 #31)
-#33 := [rewrite]: #32
-#36 := [monotonicity #33]: #35
-#39 := [monotonicity #33 #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
-#30 := [asserted]: #13
-[mp #30 #57]: false
-unsat
-b5e5edf486660b04e5c2cc8e67825f081c9c49a9 136 0
+2e6f238fe51369f58bc6454a136532e4c6f3bd82 136 0
 #2 := false
 decl f3 :: (-> S2 S2 S1)
 decl f4 :: S2
@@ -5478,68 +5612,7 @@
 #593 := [quant-inst]: #592
 [unit-resolution #593 #610 #258]: false
 unsat
-81bf38fb0456bcb900ea52fab12378a5588effa7 60 0
-#2 := false
-decl f1 :: S1
-#4 := f1
-decl f3 :: S1
-#8 := f3
-#9 := (= f3 f1)
-#14 := (not #9)
-decl f4 :: S1
-#10 := f4
-#11 := (= f4 f1)
-#13 := (not #11)
-#15 := (implies #13 #14)
-#12 := (implies #9 #11)
-#16 := (iff #12 #15)
-#17 := (not #16)
-#74 := (iff #17 false)
-#1 := true
-#69 := (not true)
-#72 := (iff #69 false)
-#73 := [rewrite]: #72
-#70 := (iff #17 #69)
-#67 := (iff #16 true)
-#38 := (= f1 f4)
-#35 := (= f1 f3)
-#44 := (not #35)
-#45 := (or #44 #38)
-#62 := (iff #45 #45)
-#65 := (iff #62 true)
-#66 := [rewrite]: #65
-#63 := (iff #16 #62)
-#60 := (iff #15 #45)
-#50 := (not #38)
-#55 := (implies #50 #44)
-#58 := (iff #55 #45)
-#59 := [rewrite]: #58
-#56 := (iff #15 #55)
-#53 := (iff #14 #44)
-#36 := (iff #9 #35)
-#37 := [rewrite]: #36
-#54 := [monotonicity #37]: #53
-#51 := (iff #13 #50)
-#39 := (iff #11 #38)
-#40 := [rewrite]: #39
-#52 := [monotonicity #40]: #51
-#57 := [monotonicity #52 #54]: #56
-#61 := [trans #57 #59]: #60
-#48 := (iff #12 #45)
-#41 := (implies #35 #38)
-#46 := (iff #41 #45)
-#47 := [rewrite]: #46
-#42 := (iff #12 #41)
-#43 := [monotonicity #37 #40]: #42
-#49 := [trans #43 #47]: #48
-#64 := [monotonicity #49 #61]: #63
-#68 := [trans #64 #66]: #67
-#71 := [monotonicity #68]: #70
-#75 := [trans #71 #73]: #74
-#34 := [asserted]: #17
-[mp #34 #75]: false
-unsat
-9d9671876c5da4fc5d4c634fb934bc459b3fb5e8 124 0
+a15d00efdb475e30bdd25ef3ed2227a8ea9c519f 124 0
 #2 := false
 decl f3 :: (-> S2 S1)
 decl f4 :: (-> S2 S2)
@@ -5664,80 +5737,7 @@
 #207 := [mp #217 #206]: #569
 [unit-resolution #207 #576 #88 #209]: false
 unsat
-7d3ca4c016c80378051605a2446d32cfa542aaf2 72 0
-#2 := false
-decl f1 :: S1
-#4 := f1
-decl f3 :: S1
-#8 := f3
-#9 := (= f3 f1)
-#10 := (iff #9 #9)
-#11 := (iff #9 #10)
-#12 := (iff #9 #11)
-#13 := (iff #9 #12)
-#14 := (iff #9 #13)
-#15 := (iff #9 #14)
-#16 := (iff #9 #15)
-#17 := (iff #9 #16)
-#18 := (iff #9 #17)
-#19 := (not #18)
-#87 := (iff #19 false)
-#1 := true
-#82 := (not true)
-#85 := (iff #82 false)
-#86 := [rewrite]: #85
-#83 := (iff #19 #82)
-#80 := (iff #18 true)
-#37 := (= f1 f3)
-#40 := (iff #37 #37)
-#43 := (iff #40 true)
-#44 := [rewrite]: #43
-#78 := (iff #18 #40)
-#76 := (iff #17 #37)
-#47 := (iff #37 true)
-#50 := (iff #47 #37)
-#51 := [rewrite]: #50
-#74 := (iff #17 #47)
-#72 := (iff #16 true)
-#70 := (iff #16 #40)
-#68 := (iff #15 #37)
-#66 := (iff #15 #47)
-#64 := (iff #14 true)
-#62 := (iff #14 #40)
-#60 := (iff #13 #37)
-#58 := (iff #13 #47)
-#56 := (iff #12 true)
-#54 := (iff #12 #40)
-#52 := (iff #11 #37)
-#48 := (iff #11 #47)
-#45 := (iff #10 true)
-#41 := (iff #10 #40)
-#38 := (iff #9 #37)
-#39 := [rewrite]: #38
-#42 := [monotonicity #39 #39]: #41
-#46 := [trans #42 #44]: #45
-#49 := [monotonicity #39 #46]: #48
-#53 := [trans #49 #51]: #52
-#55 := [monotonicity #39 #53]: #54
-#57 := [trans #55 #44]: #56
-#59 := [monotonicity #39 #57]: #58
-#61 := [trans #59 #51]: #60
-#63 := [monotonicity #39 #61]: #62
-#65 := [trans #63 #44]: #64
-#67 := [monotonicity #39 #65]: #66
-#69 := [trans #67 #51]: #68
-#71 := [monotonicity #39 #69]: #70
-#73 := [trans #71 #44]: #72
-#75 := [monotonicity #39 #73]: #74
-#77 := [trans #75 #51]: #76
-#79 := [monotonicity #39 #77]: #78
-#81 := [trans #79 #44]: #80
-#84 := [monotonicity #81]: #83
-#88 := [trans #84 #86]: #87
-#36 := [asserted]: #19
-[mp #36 #88]: false
-unsat
-6fba77c7ee170a1620030b502ad2db7a8a1aed9e 64 0
+0e7fa24f774b5d4840baee92f17adcdc6e36469d 64 0
 #2 := false
 decl f3 :: (-> S2 S2 S3)
 decl f4 :: S2
@@ -5802,33 +5802,7 @@
 #130 := [quant-inst]: #216
 [unit-resolution #130 #553 #64]: false
 unsat
-35917dcc4acd32693fbb42d1a381fa58c2e775dd 25 0
-#2 := false
-#8 := (:var 0 S2)
-#9 := (= #8 #8)
-#10 := (exists (vars (?v0 S2)) #9)
-#11 := (not #10)
-#43 := (iff #11 false)
-#1 := true
-#38 := (not true)
-#41 := (iff #38 false)
-#42 := [rewrite]: #41
-#39 := (iff #11 #38)
-#36 := (iff #10 true)
-#31 := (exists (vars (?v0 S2)) true)
-#34 := (iff #31 true)
-#35 := [elim-unused]: #34
-#32 := (iff #10 #31)
-#29 := (iff #9 true)
-#30 := [rewrite]: #29
-#33 := [quant-intro #30]: #32
-#37 := [trans #33 #35]: #36
-#40 := [monotonicity #37]: #39
-#44 := [trans #40 #42]: #43
-#28 := [asserted]: #11
-[mp #28 #44]: false
-unsat
-223d86a753fa4111a2efb19d2bcec167505d6a7e 250 0
+e97ea15e571d05d4542f0baceec935a45a20adf5 250 0
 #2 := false
 decl f3 :: (-> S2 S1)
 decl f7 :: S2
@@ -6079,7 +6053,33 @@
 #699 := [mp #350 #698]: #693
 [unit-resolution #699 #173 #714 #398]: false
 unsat
-48e86f3a23bbb7c517dedfe6660826ca25059fca 34 0
+150a137e969806c6900cabc66bb27970472c7ba1 25 0
+#2 := false
+#8 := (:var 0 S2)
+#9 := (= #8 #8)
+#10 := (exists (vars (?v0 S2)) #9)
+#11 := (not #10)
+#43 := (iff #11 false)
+#1 := true
+#38 := (not true)
+#41 := (iff #38 false)
+#42 := [rewrite]: #41
+#39 := (iff #11 #38)
+#36 := (iff #10 true)
+#31 := (exists (vars (?v0 S2)) true)
+#34 := (iff #31 true)
+#35 := [elim-unused]: #34
+#32 := (iff #10 #31)
+#29 := (iff #9 true)
+#30 := [rewrite]: #29
+#33 := [quant-intro #30]: #32
+#37 := [trans #33 #35]: #36
+#40 := [monotonicity #37]: #39
+#44 := [trans #40 #42]: #43
+#28 := [asserted]: #11
+[mp #28 #44]: false
+unsat
+b28b1c7eb24edf6acf5466ddd1304a6b881a60f9 34 0
 #2 := false
 decl f1 :: S1
 #4 := f1
@@ -6114,11 +6114,7 @@
 #30 := [asserted]: #13
 [mp #30 #50]: false
 unsat
-532713577ffc51767de4833038556b4e3293fda4 1 2
-unsat
-WARNING: failed to find a pattern for quantifier (quantifier id: k!9)
-
-fc7f88181fbb95ae1b43bd1ec021651eb1a6e29c 258 0
+043713d1f8d615cef2b57b8a8ae1162f9c51de61 258 0
 #2 := false
 decl f4 :: (-> S2 S1)
 decl ?v0!2 :: S2
@@ -6377,9 +6373,7 @@
 #579 := [quant-inst]: #578
 [unit-resolution #579 #432 #573]: false
 unsat
-b9a1dbc05a0fcdc4ae2c7e03b2b8a9ad6df98756 1 0
-unsat
-9db388838cd05c32098d6cee4786c7311602e951 222 0
+7fcf7ce1521f6f1599b2fe6f93ba5dee1cbd7b76 222 0
 #2 := false
 decl f3 :: (-> S2 S1)
 decl ?v0!0 :: S2
@@ -6602,9 +6596,7 @@
 #610 := [quant-inst]: #609
 [unit-resolution #610 #339 #320]: false
 unsat
-cbac2f7e87b549893b94e3792e09a58dadcc84e9 1 0
-unsat
-5815749abaa2e9d579422ebe3d14186ba8803728 49 0
+8031eb3bcd3fb97da7b874e36b0d4d2ae215ec7b 49 0
 #2 := false
 decl f1 :: S1
 #4 := f1
@@ -6654,9 +6646,7 @@
 #35 := [asserted]: #18
 [mp #35 #65]: false
 unsat
-1493ff264c0f89a95d188930a1476aec2a1aa1bb 1 0
-unsat
-e053e47f320825e5b9b67b9497406e00070abe43 122 0
+301a246cbbd414b8b75c6b801fee2f734cf2af47 122 0
 #2 := false
 decl f3 :: (-> S2 S1)
 decl ?v0!1 :: S2
@@ -6779,9 +6769,21 @@
 #287 := [quant-inst]: #200
 [unit-resolution #287 #620 #138]: false
 unsat
-1860fe1f45d9df9b729809ed662950c61f79f829 1 0
-unsat
-12d1461d1fa5477cb514ef80c18c85a012607364 105 0
+2317912b27971033dc831a6ce5ef74d0f20fbca2 1 2
+unsat
+WARNING: failed to find a pattern for quantifier (quantifier id: k!9)
+
+ee898f52c96d73444993788b39566855271caaec 1 0
+unsat
+f7daa562701d4999ed8a3f5ec133f08651bf6190 1 0
+unsat
+fdbe39cdd485517d8ec1964f2c4fb7b6646fe608 1 0
+unsat
+779c5fc38fd08b8484ec1da52f6a17e6b849eac4 1 0
+unsat
+b720bb30cc87882c3f660a5bf9c8737dde4c14c4 1 0
+unsat
+2acdf9e468d060988547fd7ad4c93238ed78a5bb 105 0
 #2 := false
 decl f3 :: (-> S1 S1)
 decl ?v0!1 :: S1
@@ -6887,9 +6889,7 @@
 #261 := [quant-inst]: #188
 [unit-resolution #261 #590 #259]: false
 unsat
-58aeee112e6b81dcd3353f63fc46ef0b0f3eb24c 1 0
-unsat
-ed0af7544f19477dc25f65a160752e70bd0bbf21 230 0
+ff2e1b7168f418022909b6d3ec628836562df57d 230 0
 #2 := false
 decl f3 :: (-> S1 S1)
 decl ?v0!0 :: S1
@@ -7120,7 +7120,7 @@
 #309 := [quant-inst]: #308
 [unit-resolution #309 #590 #312]: false
 unsat
-446f4f1e86496dc31d9127dadba175aab73fcdf4 156 0
+258fe7f9c05eff8e107778930f55bcfa407e2fde 156 0
 #2 := false
 decl f3 :: (-> S1 S2 S1)
 #9 := (:var 0 S2)
@@ -7277,7 +7277,7 @@
 #269 := [quant-inst]: #264
 [unit-resolution #269 #592 #292]: false
 unsat
-73417ccd5b8c2fb3c0c21a8c23809282ab12f8b8 74 0
+6c0f54222b59d5e628ff948749ac67e96364c6a1 74 0
 #2 := false
 decl f3 :: (-> S2 S3 S4)
 #9 := (:var 0 S3)
@@ -7352,7 +7352,7 @@
 #548 := [mp #55 #547]: #543
 [unit-resolution #548 #537]: false
 unsat
-e3e78e7b3355a0e12a6553b06865dbb99fa1ace5 108 0
+cba176f179e597e7b48393532028be8634e0c4f7 108 0
 #2 := false
 decl f3 :: (-> S1 S1)
 #11 := (:var 0 S1)
@@ -7461,13 +7461,7 @@
 #179 := [quant-inst]: #265
 [unit-resolution #179 #96 #599]: false
 unsat
-da85096b57c08af8febf5871c1333121a34c6d28 1 0
-unsat
-6c41d71d8ec4d73ffa088d938320c03ca09b9201 1 2
-unsat
-WARNING: failed to find a pattern for quantifier (quantifier id: k!12)
-
-12e3d2e5e3ee6fdfd1904d6ec3b7cee27863319e 183 2
+1f79e2aaf530f1cad27b8439c2a93a00022ad034 183 2
 #2 := false
 decl f3 :: (-> S1 S2 S1)
 #9 := (:var 0 S2)
@@ -7653,6 +7647,12 @@
 unsat
 WARNING: failed to find a pattern for quantifier (quantifier id: k!12)
 
+da85096b57c08af8febf5871c1333121a34c6d28 1 0
+unsat
+6c41d71d8ec4d73ffa088d938320c03ca09b9201 1 2
+unsat
+WARNING: failed to find a pattern for quantifier (quantifier id: k!12)
+
 5d15e0a34749300f4ce5f4c920d3ca259f7027cd 1 0
 unsat
 fa1639cb0c965ae73d4377750883cc09dbe87562 1 0
@@ -7663,7 +7663,7 @@
 unsat
 22c7d1d1136fafc0a09e243a12ede1ee4fee416f 1 0
 unsat
-92de3e43df292b3d7f698eeecd3084ad63385f52 107 0
+cfd9d2f647046e282a1080aae51b4b88be291157 107 0
 #2 := false
 decl f3 :: (-> S2 S1)
 #8 := (:var 0 S2)
@@ -7771,7 +7771,7 @@
 #279 := [quant-inst]: #193
 [unit-resolution #279 #102 #618]: false
 unsat
-ab595871a8f439823819ae39abbb09300ded5d4d 75 0
+9ea0d22c6270219f7cef6ee3d8f429562349b2a3 75 0
 #2 := false
 decl f1 :: S1
 #4 := f1
@@ -7847,7 +7847,7 @@
 #41 := [asserted]: #24
 [mp #41 #91]: false
 unsat
-66a02ca33e858454a913431cd41d1e5a58d7ca83 207 0
+51de51a0d85b126e271283f7eca5a0bb3a5a5bca 207 0
 #2 := false
 decl f3 :: (-> S2 S1)
 decl ?v1!0 :: (-> S2 S2)
@@ -8055,7 +8055,7 @@
 #622 := [mp #616 #257]: #277
 [unit-resolution #622 #645 #591 #596]: false
 unsat
-070caa247cea32871c723a813303329e9c48e89d 250 0
+628763241ae198f2ea4a54c641ca6e8df6d851ca 250 0
 #2 := false
 decl f3 :: (-> S2 S1)
 decl ?v1!1 :: (-> S2 S2)
@@ -8306,7 +8306,7 @@
 #616 := [mp #628 #615]: #622
 [unit-resolution #616 #128 #676 #607 #606]: false
 unsat
-ffccf343597d7c1ee1fd555056e55c0f14472255 238 0
+64bd43a9d49de65761e5220e25dbbe3dc269bf0e 238 0
 #2 := false
 decl f3 :: (-> S2 S1)
 decl ?v1!1 :: (-> S2 S2)
@@ -8545,7 +8545,215 @@
 #338 := [mp #643 #632]: #276
 [unit-resolution #338 #664 #617 #616]: false
 unsat
-8ce5d808d97d9f937477d79f321db3630294ec1d 43 0
+b8d80ad0d8c340491041b0c69d71b3944eaf90e6 93 0
+#2 := false
+decl f5 :: (-> S2 S1)
+decl f6 :: S2
+#16 := f6
+#20 := (f5 f6)
+decl f1 :: S1
+#4 := f1
+#65 := (= f1 #20)
+#84 := (not #65)
+decl f3 :: (-> S2 S3 S1)
+decl f4 :: S3
+#9 := f4
+#17 := (f3 f6 f4)
+#59 := (= f1 #17)
+#8 := (:var 0 S2)
+#12 := (f5 #8)
+#44 := (= f1 #12)
+#10 := (f3 #8 f4)
+#41 := (= f1 #10)
+#50 := (not #41)
+#51 := (or #50 #44)
+#56 := (forall (vars (?v0 S2)) #51)
+#62 := (and #56 #59)
+#71 := (not #62)
+#72 := (or #71 #65)
+#77 := (not #72)
+#21 := (= #20 f1)
+#18 := (= #17 f1)
+#13 := (= #12 f1)
+#11 := (= #10 f1)
+#14 := (implies #11 #13)
+#15 := (forall (vars (?v0 S2)) #14)
+#19 := (and #15 #18)
+#22 := (implies #19 #21)
+#23 := (not #22)
+#78 := (iff #23 #77)
+#75 := (iff #22 #72)
+#68 := (implies #62 #65)
+#73 := (iff #68 #72)
+#74 := [rewrite]: #73
+#69 := (iff #22 #68)
+#66 := (iff #21 #65)
+#67 := [rewrite]: #66
+#63 := (iff #19 #62)
+#60 := (iff #18 #59)
+#61 := [rewrite]: #60
+#57 := (iff #15 #56)
+#54 := (iff #14 #51)
+#47 := (implies #41 #44)
+#52 := (iff #47 #51)
+#53 := [rewrite]: #52
+#48 := (iff #14 #47)
+#45 := (iff #13 #44)
+#46 := [rewrite]: #45
+#42 := (iff #11 #41)
+#43 := [rewrite]: #42
+#49 := [monotonicity #43 #46]: #48
+#55 := [trans #49 #53]: #54
+#58 := [quant-intro #55]: #57
+#64 := [monotonicity #58 #61]: #63
+#70 := [monotonicity #64 #67]: #69
+#76 := [trans #70 #74]: #75
+#79 := [monotonicity #76]: #78
+#40 := [asserted]: #23
+#82 := [mp #40 #79]: #77
+#85 := [not-or-elim #82]: #84
+#80 := [not-or-elim #82]: #62
+#83 := [and-elim #80]: #59
+#572 := (pattern #12)
+#571 := (pattern #10)
+#573 := (forall (vars (?v0 S2)) (:pat #571 #572) #51)
+#576 := (iff #56 #573)
+#574 := (iff #51 #51)
+#575 := [refl]: #574
+#577 := [quant-intro #575]: #576
+#97 := (~ #56 #56)
+#95 := (~ #51 #51)
+#96 := [refl]: #95
+#98 := [nnf-pos #96]: #97
+#81 := [and-elim #80]: #56
+#88 := [mp~ #81 #98]: #56
+#578 := [mp #88 #577]: #573
+#155 := (not #59)
+#157 := (not #573)
+#244 := (or #157 #155 #65)
+#242 := (or #155 #65)
+#235 := (or #157 #242)
+#247 := (iff #235 #244)
+#175 := [rewrite]: #247
+#246 := [quant-inst]: #235
+#248 := [mp #246 #175]: #244
+[unit-resolution #248 #578 #83 #85]: false
+unsat
+0f8bf9d230b6b97c3b9a64ed4f1bb8757963fbc7 113 0
+#2 := false
+decl f3 :: (-> S2 S3 S1)
+decl f4 :: S3
+#9 := f4
+decl f6 :: S2
+#16 := f6
+#19 := (f3 f6 f4)
+decl f1 :: S1
+#4 := f1
+#57 := (= f1 #19)
+decl f5 :: (-> S2 S1)
+#17 := (f5 f6)
+#54 := (= f1 #17)
+#60 := (and #54 #57)
+#63 := (not #60)
+#8 := (:var 0 S2)
+#12 := (f5 #8)
+#45 := (= f1 #12)
+#10 := (f3 #8 f4)
+#42 := (= f1 #10)
+#48 := (and #42 #45)
+#51 := (exists (vars (?v0 S2)) #48)
+#66 := (or #51 #63)
+#69 := (not #66)
+#20 := (= #19 f1)
+#18 := (= #17 f1)
+#21 := (and #18 #20)
+#22 := (not #21)
+#13 := (= #12 f1)
+#11 := (= #10 f1)
+#14 := (and #11 #13)
+#15 := (exists (vars (?v0 S2)) #14)
+#23 := (or #15 #22)
+#24 := (not #23)
+#70 := (iff #24 #69)
+#67 := (iff #23 #66)
+#64 := (iff #22 #63)
+#61 := (iff #21 #60)
+#58 := (iff #20 #57)
+#59 := [rewrite]: #58
+#55 := (iff #18 #54)
+#56 := [rewrite]: #55
+#62 := [monotonicity #56 #59]: #61
+#65 := [monotonicity #62]: #64
+#52 := (iff #15 #51)
+#49 := (iff #14 #48)
+#46 := (iff #13 #45)
+#47 := [rewrite]: #46
+#43 := (iff #11 #42)
+#44 := [rewrite]: #43
+#50 := [monotonicity #44 #47]: #49
+#53 := [quant-intro #50]: #52
+#68 := [monotonicity #53 #65]: #67
+#71 := [monotonicity #68]: #70
+#41 := [asserted]: #24
+#74 := [mp #41 #71]: #69
+#75 := [not-or-elim #74]: #60
+#77 := [and-elim #75]: #57
+#76 := [and-elim #75]: #54
+#585 := (pattern #12)
+#584 := (pattern #10)
+#93 := (not #45)
+#92 := (not #42)
+#94 := (or #92 #93)
+#586 := (forall (vars (?v0 S2)) (:pat #584 #585) #94)
+#101 := (forall (vars (?v0 S2)) #94)
+#589 := (iff #101 #586)
+#587 := (iff #94 #94)
+#588 := [refl]: #587
+#590 := [quant-intro #588]: #589
+#87 := (not #48)
+#90 := (forall (vars (?v0 S2)) #87)
+#102 := (iff #90 #101)
+#99 := (iff #87 #94)
+#95 := (not #94)
+#83 := (not #95)
+#97 := (iff #83 #94)
+#98 := [rewrite]: #97
+#84 := (iff #87 #83)
+#85 := (iff #48 #95)
+#86 := [rewrite]: #85
+#96 := [monotonicity #86]: #84
+#100 := [trans #96 #98]: #99
+#103 := [quant-intro #100]: #102
+#72 := (not #51)
+#80 := (~ #72 #90)
+#88 := (~ #87 #87)
+#89 := [refl]: #88
+#78 := [nnf-neg #89]: #80
+#73 := [not-or-elim #74]: #72
+#91 := [mp~ #73 #78]: #90
+#104 := [mp #91 #103]: #101
+#591 := [mp #104 #590]: #586
+#255 := (not #57)
+#168 := (not #54)
+#248 := (not #586)
+#259 := (or #248 #168 #255)
+#169 := (or #255 #168)
+#260 := (or #248 #169)
+#578 := (iff #260 #259)
+#256 := (or #168 #255)
+#261 := (or #248 #256)
+#241 := (iff #261 #259)
+#576 := [rewrite]: #241
+#258 := (iff #260 #261)
+#170 := (iff #169 #256)
+#257 := [rewrite]: #170
+#262 := [monotonicity #257]: #258
+#235 := [trans #262 #576]: #578
+#188 := [quant-inst]: #260
+#365 := [mp #188 #235]: #259
+[unit-resolution #365 #591 #76 #77]: false
+unsat
+42f3c8af4c50150e11519c085a0ad69e107e1705 43 0
 #2 := false
 decl f1 :: S1
 #4 := f1
@@ -8589,7 +8797,7 @@
 #32 := [asserted]: #15
 [mp #32 #58]: false
 unsat
-ee1e5475ea11ff0f473b33b8c572488c1f039a0e 27 0
+e44a18e086890afdae37ccfa17aec44f2a634103 27 0
 #2 := false
 #1 := true
 #9 := (or false true)
@@ -8617,7 +8825,7 @@
 #28 := [asserted]: #11
 [mp #28 #46]: false
 unsat
-82c66294f6bc190a20d250a4062b09707f16aa4c 31 0
+48f0da3339d9c9f7707291c89458bdaf3e50c114 31 0
 #2 := false
 decl f1 :: S1
 #4 := f1
@@ -8649,7 +8857,7 @@
 #29 := [asserted]: #12
 [mp #29 #46]: false
 unsat
-f32b862e1d7121899074795ddf1c0725bc5cc874 33 0
+e7059fb5d2b001f75897ae16a2ace8ce919e02a1 33 0
 #2 := false
 decl f1 :: S1
 #4 := f1
@@ -8683,7 +8891,43 @@
 #30 := [asserted]: #13
 [mp #30 #47]: false
 unsat
-e070982f250a1ca780eab96adc6095281038c59d 38 0
+f5fef851c10f40408540848a3507970291adb3e5 35 0
+#2 := false
+decl f1 :: S1
+#4 := f1
+decl f3 :: (-> S2 S1)
+decl f4 :: S2
+#8 := f4
+#9 := (f3 f4)
+#10 := (= #9 f1)
+#11 := (not #10)
+#12 := (ite #10 #10 #11)
+#13 := (not #12)
+#49 := (iff #13 false)
+#1 := true
+#44 := (not true)
+#47 := (iff #44 false)
+#48 := [rewrite]: #47
+#45 := (iff #13 #44)
+#42 := (iff #12 true)
+#31 := (= f1 #9)
+#34 := (not #31)
+#37 := (ite #31 #31 #34)
+#40 := (iff #37 true)
+#41 := [rewrite]: #40
+#38 := (iff #12 #37)
+#35 := (iff #11 #34)
+#32 := (iff #10 #31)
+#33 := [rewrite]: #32
+#36 := [monotonicity #33]: #35
+#39 := [monotonicity #33 #33 #36]: #38
+#43 := [trans #39 #41]: #42
+#46 := [monotonicity #43]: #45
+#50 := [trans #46 #48]: #49
+#30 := [asserted]: #13
+[mp #30 #50]: false
+unsat
+9facdecb47b87e40c25944deaa804ef85fc985e3 38 0
 #2 := false
 decl f3 :: (-> S2 S2)
 decl f4 :: S2
@@ -8722,43 +8966,7 @@
 #114 := [quant-inst]: #200
 [unit-resolution #114 #51 #50]: false
 unsat
-5566817723d9b0dab514cd1865ced56d8e166fe8 35 0
-#2 := false
-decl f1 :: S1
-#4 := f1
-decl f3 :: (-> S2 S1)
-decl f4 :: S2
-#8 := f4
-#9 := (f3 f4)
-#10 := (= #9 f1)
-#11 := (not #10)
-#12 := (ite #10 #10 #11)
-#13 := (not #12)
-#49 := (iff #13 false)
-#1 := true
-#44 := (not true)
-#47 := (iff #44 false)
-#48 := [rewrite]: #47
-#45 := (iff #13 #44)
-#42 := (iff #12 true)
-#31 := (= f1 #9)
-#34 := (not #31)
-#37 := (ite #31 #31 #34)
-#40 := (iff #37 true)
-#41 := [rewrite]: #40
-#38 := (iff #12 #37)
-#35 := (iff #11 #34)
-#32 := (iff #10 #31)
-#33 := [rewrite]: #32
-#36 := [monotonicity #33]: #35
-#39 := [monotonicity #33 #33 #36]: #38
-#43 := [trans #39 #41]: #42
-#46 := [monotonicity #43]: #45
-#50 := [trans #46 #48]: #49
-#30 := [asserted]: #13
-[mp #30 #50]: false
-unsat
-bc840300be3a9e2dc587394b4317175f8f11e65e 30 0
+ae9dfcc290bc6bcccab3cd3af533d10c7c03b3b8 30 0
 #2 := false
 decl f4 :: (-> S3 S4)
 decl f6 :: S3
@@ -8789,7 +8997,7 @@
 #104 := [quant-inst]: #190
 [unit-resolution #104 #41 #38]: false
 unsat
-4e0dadc89c97a124d33549e6d90d8c6fb94ffbbc 16 0
+90ce83a254a12d0f1bfd0c5eb5a99c36ab0fe295 16 0
 #2 := false
 #1 := true
 #8 := (implies true true)
@@ -8806,7 +9014,7 @@
 #26 := [asserted]: #9
 [mp #26 #35]: false
 unsat
-300f22ba0865fbc4bc611b5c5320b9fb475b953c 16 0
+9a0ddca6dee00a602bc428366efd931bcfcee94b 16 0
 #2 := false
 #8 := (implies false false)
 #9 := (not #8)
@@ -8823,7 +9031,7 @@
 #26 := [asserted]: #9
 [mp #26 #35]: false
 unsat
-5cdd14f812034153746401ccd7149e6a8e608edc 31 0
+cabaa8d0303be40a64e7064642dabce9e7bfcad7 31 0
 #2 := false
 decl f1 :: S1
 #4 := f1
@@ -8855,7 +9063,7 @@
 #29 := [asserted]: #12
 [mp #29 #46]: false
 unsat
-6c182dd23633c7853c1a91f152d7ab135330412d 48 0
+ac00789473bf074e49a8a6419c1f2a1042912d6a 48 0
 #2 := false
 decl f1 :: S1
 #4 := f1
@@ -8904,7 +9112,7 @@
 #31 := [asserted]: #14
 [mp #31 #63]: false
 unsat
-85c24c97e452e2742450de7a829a24bef0b7745e 53 0
+75052914158f379cd343f693f354895fa579a89c 53 0
 #2 := false
 decl f1 :: S1
 #4 := f1
@@ -8958,7 +9166,7 @@
 #31 := [asserted]: #14
 [mp #31 #68]: false
 unsat
-9c92a59e1b2801a69940e4a222a4ab6da3c64d26 56 0
+4bcd26d70bcaf24bb295031b732101a3dec2c643 56 0
 #2 := false
 decl f1 :: S1
 #4 := f1
@@ -9015,7 +9223,7 @@
 #32 := [asserted]: #15
 [mp #32 #71]: false
 unsat
-e45e8738077ad6b1be47d22a573712217695edc5 89 0
+7eeb56537bd13a6f6bc35096a4ddfb33dceb295f 89 0
 #2 := false
 decl f4 :: S1
 #10 := f4
@@ -9105,7 +9313,7 @@
 #86 := [not-or-elim #83]: #48
 [mp #86 #107]: false
 unsat
-54fa403d565ef2e106d2d33e4ac25062fa36ce4d 57 0
+38ce6b99e91d92d7feb75b7b5ff7006aee1a887f 57 0
 #2 := false
 decl f3 :: (-> S2 S1)
 decl f4 :: S2
@@ -9163,7 +9371,7 @@
 #59 := [not-or-elim #57]: #58
 [mp #59 #74]: false
 unsat
-65635b02b15ecbf13faf998b6e281ec4741d131a 50 0
+8e8baf35daf48781c204954ba8d7df5254aac6a8 50 0
 #2 := false
 decl f5 :: S2
 #11 := f5
@@ -9214,7 +9422,7 @@
 #62 := [not-or-elim #58]: #61
 [unit-resolution #62 #70]: false
 unsat
-317fa0f363a89944322628da895574f970f0290e 27 0
+7bbe1bf9e8878e3709a71589e39f12e31278f1b7 27 0
 #2 := false
 decl f5 :: (-> S2 S3)
 decl f4 :: S2
@@ -9242,7 +9450,7 @@
 #43 := [not-or-elim #42]: #41
 [unit-resolution #43 #51]: false
 unsat
-1d26941a5bb2830adb1f9668b97240fa88d0de2c 73 0
+a8014c2ba4ebe428717a3516d1900ca3b7591555 73 0
 #2 := false
 decl f3 :: (-> S1 S1)
 decl f4 :: S1
@@ -9316,7 +9524,7 @@
 #146 := [quant-inst]: #232
 [unit-resolution #146 #566 #74]: false
 unsat
-43323a2e96998bec6557b77d2200c04f63428dbf 106 0
+c0906c23e4b0650863a2ef0d68bafa4e06a9dd19 106 0
 #2 := false
 decl f3 :: (-> S2 S2 S1)
 decl ?v0!0 :: S2
@@ -9423,7 +9631,7 @@
 #239 := [mp #568 #238]: #254
 [unit-resolution #239 #583 #88]: false
 unsat
-c09f42a4df50b3429cb0f5ddbf56d93bf0b06b7b 66 0
+58aaefc3b25d8181c01a2139430686aeca861bf3 66 0
 #2 := false
 decl f4 :: S1
 #10 := f4
@@ -9490,7 +9698,7 @@
 #71 := [and-elim #69]: #40
 [mp #71 #85]: false
 unsat
-0e357e9971053d2577fa677046eba1d5718e4aa2 59 0
+d80d7bb97197c962c0b81d44c7a8cfd617b6ca33 59 0
 #2 := false
 decl f1 :: S1
 #4 := f1
@@ -9550,11 +9758,11 @@
 #37 := [asserted]: #20
 [mp #37 #72]: false
 unsat
-d654180cf0c8fb52012ec18288c7b9cba4237e00 3 0
+46acb4cd4b5353c32c529ffe509fc6f26ac2f2a1 3 0
 #2 := false
 [asserted]: false
 unsat
-499cbfe2406983d4607365f26173b22f6a49074a 18 0
+040123bf68e09d0b72cf990436bb29a623c102b8 18 0
 #2 := false
 #9 := 1::int
 #8 := 0::int
@@ -9573,7 +9781,7 @@
 #28 := [asserted]: #11
 [mp #28 #39]: false
 unsat
-7669ecc3fd6958d7a6535d8186e984760faa6558 18 0
+be1600749b493592f603ae3a7f9dcbc68d688187 18 0
 #2 := false
 #9 := 1::int
 #8 := 0::int
@@ -9592,7 +9800,7 @@
 #28 := [asserted]: #11
 [mp #28 #39]: false
 unsat
-399d7032965ae0b6a791d2007093e9c881b1b6df 18 0
+d6985ded37fcab252cac3fbb09a30e977c68973d 18 0
 #2 := false
 #9 := 2345678901::int
 #8 := 123456789::int
@@ -9611,7 +9819,7 @@
 #28 := [asserted]: #11
 [mp #28 #39]: false
 unsat
-27d5c260fdf33e6295946c76babe053e1d5624d6 30 0
+12f34f8245ed278a095c475d68eb5a5711d23981 30 0
 #2 := false
 decl f3 :: (-> int S2)
 #25 := 1::int
@@ -9642,7 +9850,7 @@
 #132 := [asserted]: #30
 [mp #132 #149]: false
 unsat
-8b231c32569080edc95ccdce9f961a9b49869ef0 225 0
+fb0035d341cb2e8ac6449dbe61b4ac5d9e12a151 225 0
 #2 := false
 #13 := 0::int
 decl f4 :: (-> S2 int)
@@ -9868,7 +10076,7 @@
 #616 := [unit-resolution #619 #618]: #347
 [th-lemma #617 #616 #166]: false
 unsat
-9c6198ccc440358afe5fbefeb129ea4a00cbbb0c 428 0
+f572500dcd49c30edfbceabc272b009fb11814d2 428 0
 #2 := false
 decl f4 :: (-> S2 int)
 decl f6 :: S2
@@ -10297,7 +10505,7 @@
 #393 := [unit-resolution #403 #411 #414]: #443
 [unit-resolution #393 #424]: false
 unsat
-206e819433d78c632645e8c99645a12ddf3b51c0 409 0
+1b539d34eedd00cd7cacdb9d650b014baeb31cd1 409 0
 #2 := false
 #181 := -1::int
 decl f4 :: (-> S2 int)
@@ -10707,7 +10915,7 @@
 #420 := [th-lemma]: #419
 [unit-resolution #420 #417 #412]: false
 unsat
-42b37d88064ab90bea2088cc1497480fb8953990 60 0
+b265c055659e0ea1b48e0304606789db36e05540 60 0
 #2 := false
 decl f3 :: (-> int S2)
 decl f4 :: (-> S2 int)
@@ -10768,7 +10976,7 @@
 #311 := [quant-inst]: #310
 [unit-resolution #311 #646 #151]: false
 unsat
-9798b09402a5018871b7e69145c530bf687b5055 60 0
+75c3bae150e305a0fc97df1a1e35cb97fdd88bd7 60 0
 #2 := false
 decl f3 :: (-> int S2)
 decl f4 :: (-> S2 int)
@@ -10829,7 +11037,7 @@
 #311 := [quant-inst]: #310
 [unit-resolution #311 #646 #151]: false
 unsat
-a870002129082499022465a324347d99a190812c 37 0
+c52f3166d15e1102a7ca505cd72ba484966ee318 37 0
 #2 := false
 decl f3 :: (-> int S2)
 decl f4 :: (-> S2 int)
@@ -10867,7 +11075,7 @@
 #136 := [asserted]: #34
 [mp #136 #154]: false
 unsat
-54e17ed584134deb7dd3985de3ff45b3769746ed 42 0
+b5b340366dfb8289e173cd594b95953e97e705ac 42 0
 #2 := false
 decl f3 :: (-> int S2)
 decl f4 :: (-> S2 int)
@@ -10910,68 +11118,7 @@
 #140 := [asserted]: #38
 [mp #140 #158]: false
 unsat
-6d74fb5fdf45d4b9f42d13b10663aa3e27c3c330 60 0
-#2 := false
-decl f3 :: (-> int S2)
-decl f4 :: (-> S2 int)
-decl f5 :: S2
-#25 := f5
-#26 := (f4 f5)
-#135 := (f3 #26)
-#141 := (= f5 #135)
-#146 := (not #141)
-#13 := 0::int
-#27 := (- #26 0::int)
-#28 := (f3 #27)
-#29 := (= #28 f5)
-#30 := (not #29)
-#147 := (iff #30 #146)
-#144 := (iff #29 #141)
-#138 := (= #135 f5)
-#142 := (iff #138 #141)
-#143 := [rewrite]: #142
-#139 := (iff #29 #138)
-#136 := (= #28 #135)
-#133 := (= #27 #26)
-#134 := [rewrite]: #133
-#137 := [monotonicity #134]: #136
-#140 := [monotonicity #137]: #139
-#145 := [trans #140 #143]: #144
-#148 := [monotonicity #145]: #147
-#132 := [asserted]: #30
-#151 := [mp #132 #148]: #146
-#8 := (:var 0 S2)
-#9 := (f4 #8)
-#639 := (pattern #9)
-#10 := (f3 #9)
-#48 := (= #8 #10)
-#640 := (forall (vars (?v0 S2)) (:pat #639) #48)
-#51 := (forall (vars (?v0 S2)) #48)
-#641 := (iff #51 #640)
-#643 := (iff #640 #640)
-#644 := [rewrite]: #643
-#642 := [rewrite]: #641
-#645 := [trans #642 #644]: #641
-#163 := (~ #51 #51)
-#161 := (~ #48 #48)
-#162 := [refl]: #161
-#164 := [nnf-pos #162]: #163
-#11 := (= #10 #8)
-#12 := (forall (vars (?v0 S2)) #11)
-#52 := (iff #12 #51)
-#49 := (iff #11 #48)
-#50 := [rewrite]: #49
-#53 := [quant-intro #50]: #52
-#47 := [asserted]: #12
-#56 := [mp #47 #53]: #51
-#152 := [mp~ #56 #164]: #51
-#646 := [mp #152 #645]: #640
-#224 := (not #640)
-#310 := (or #224 #141)
-#311 := [quant-inst]: #310
-[unit-resolution #311 #646 #151]: false
-unsat
-b9b9aa4c3f51dafeb0a0d1867b6ec055cde87e79 438 0
+3c1f7d30ea379c9ce2207a011fc3d11958eb42f1 438 0
 #2 := false
 #13 := 0::int
 decl f4 :: (-> S2 int)
@@ -11410,7 +11557,68 @@
 #404 := [unit-resolution #379 #434]: #485
 [unit-resolution #404 #396]: false
 unsat
-487dc506324bbdc4f045ceae0d9a8d2729196eb1 179 0
+b04995911f57617d06c9591edf349edadf325b74 60 0
+#2 := false
+decl f3 :: (-> int S2)
+decl f4 :: (-> S2 int)
+decl f5 :: S2
+#25 := f5
+#26 := (f4 f5)
+#135 := (f3 #26)
+#141 := (= f5 #135)
+#146 := (not #141)
+#13 := 0::int
+#27 := (- #26 0::int)
+#28 := (f3 #27)
+#29 := (= #28 f5)
+#30 := (not #29)
+#147 := (iff #30 #146)
+#144 := (iff #29 #141)
+#138 := (= #135 f5)
+#142 := (iff #138 #141)
+#143 := [rewrite]: #142
+#139 := (iff #29 #138)
+#136 := (= #28 #135)
+#133 := (= #27 #26)
+#134 := [rewrite]: #133
+#137 := [monotonicity #134]: #136
+#140 := [monotonicity #137]: #139
+#145 := [trans #140 #143]: #144
+#148 := [monotonicity #145]: #147
+#132 := [asserted]: #30
+#151 := [mp #132 #148]: #146
+#8 := (:var 0 S2)
+#9 := (f4 #8)
+#639 := (pattern #9)
+#10 := (f3 #9)
+#48 := (= #8 #10)
+#640 := (forall (vars (?v0 S2)) (:pat #639) #48)
+#51 := (forall (vars (?v0 S2)) #48)
+#641 := (iff #51 #640)
+#643 := (iff #640 #640)
+#644 := [rewrite]: #643
+#642 := [rewrite]: #641
+#645 := [trans #642 #644]: #641
+#163 := (~ #51 #51)
+#161 := (~ #48 #48)
+#162 := [refl]: #161
+#164 := [nnf-pos #162]: #163
+#11 := (= #10 #8)
+#12 := (forall (vars (?v0 S2)) #11)
+#52 := (iff #12 #51)
+#49 := (iff #11 #48)
+#50 := [rewrite]: #49
+#53 := [quant-intro #50]: #52
+#47 := [asserted]: #12
+#56 := [mp #47 #53]: #51
+#152 := [mp~ #56 #164]: #51
+#646 := [mp #152 #645]: #640
+#224 := (not #640)
+#310 := (or #224 #141)
+#311 := [quant-inst]: #310
+[unit-resolution #311 #646 #151]: false
+unsat
+2a686ca258d75e408aba2310eab7bfde718856c5 179 0
 #2 := false
 decl f3 :: (-> int S2)
 decl f4 :: (-> S2 int)
@@ -11590,7 +11798,7 @@
 #175 := [and-elim #174]: #154
 [unit-resolution #175 #587]: false
 unsat
-2f9c070703deaa5a8000b1eca828e079373a654c 321 0
+d2811f187a26b77cca78876fe4e7b2248df3c75e 321 0
 #2 := false
 decl f3 :: (-> int S2)
 decl f4 :: (-> S2 int)
@@ -11912,7 +12120,7 @@
 #181 := [and-elim #183]: #157
 [unit-resolution #181 #478]: false
 unsat
-d31a177ad3abcc23163a30e1d5b84f367c2764b8 329 0
+6750007dac6f87e186d233d8af85a2f566a05649 329 0
 #2 := false
 decl f6 :: S2
 #27 := f6
@@ -12242,41 +12450,7 @@
 #500 := [trans #499 #542]: #188
 [unit-resolution #198 #500]: false
 unsat
-2d9d46280589be884838dc154131e7cc5b93a904 33 0
-#2 := false
-decl f3 :: (-> int S2)
-#13 := 0::int
-#29 := (f3 0::int)
-decl f4 :: (-> S2 int)
-decl f5 :: S2
-#25 := f5
-#26 := (f4 f5)
-#27 := (* #26 0::int)
-#28 := (f3 #27)
-#30 := (= #28 #29)
-#31 := (not #30)
-#149 := (iff #31 false)
-#1 := true
-#144 := (not true)
-#147 := (iff #144 false)
-#148 := [rewrite]: #147
-#145 := (iff #31 #144)
-#142 := (iff #30 true)
-#137 := (= #29 #29)
-#140 := (iff #137 true)
-#141 := [rewrite]: #140
-#138 := (iff #30 #137)
-#134 := (= #27 0::int)
-#135 := [rewrite]: #134
-#136 := [monotonicity #135]: #30
-#139 := [monotonicity #136]: #138
-#143 := [trans #139 #141]: #142
-#146 := [monotonicity #143]: #145
-#150 := [trans #146 #148]: #149
-#133 := [asserted]: #31
-[mp #133 #150]: false
-unsat
-13fdb649a31033e319de9267921fa923a1c6b965 517 0
+38a43c830e9f2f628495aa4330122f2c2da981cd 517 0
 #2 := false
 #13 := 0::int
 decl f4 :: (-> S2 int)
@@ -12794,7 +12968,41 @@
 #282 := [unit-resolution #419 #286]: #319
 [unit-resolution #424 #282 #300]: false
 unsat
-24450ce0a41cb0a0348eb41438999f733a6e381f 33 0
+73fe8e8f8efc05a9592e80e471f1f73a8ea6b41e 33 0
+#2 := false
+decl f3 :: (-> int S2)
+#13 := 0::int
+#29 := (f3 0::int)
+decl f4 :: (-> S2 int)
+decl f5 :: S2
+#25 := f5
+#26 := (f4 f5)
+#27 := (* #26 0::int)
+#28 := (f3 #27)
+#30 := (= #28 #29)
+#31 := (not #30)
+#149 := (iff #31 false)
+#1 := true
+#144 := (not true)
+#147 := (iff #144 false)
+#148 := [rewrite]: #147
+#145 := (iff #31 #144)
+#142 := (iff #30 true)
+#137 := (= #29 #29)
+#140 := (iff #137 true)
+#141 := [rewrite]: #140
+#138 := (iff #30 #137)
+#134 := (= #27 0::int)
+#135 := [rewrite]: #134
+#136 := [monotonicity #135]: #30
+#139 := [monotonicity #136]: #138
+#143 := [trans #139 #141]: #142
+#146 := [monotonicity #143]: #145
+#150 := [trans #146 #148]: #149
+#133 := [asserted]: #31
+[mp #133 #150]: false
+unsat
+b8b27e6325657430dd1af6ebdf5787c6efe3284f 33 0
 #2 := false
 decl f3 :: (-> int S2)
 #13 := 0::int
@@ -12828,7 +13036,7 @@
 #133 := [asserted]: #31
 [mp #133 #150]: false
 unsat
-1e8972e2630fa2285c0c67ff5ba0a1263d55a492 60 0
+21ee0df98aa191c83f2b436328c15786df798f7a 60 0
 #2 := false
 decl f3 :: (-> int S2)
 decl f4 :: (-> S2 int)
@@ -12889,7 +13097,7 @@
 #311 := [quant-inst]: #310
 [unit-resolution #311 #646 #152]: false
 unsat
-e21662dbc702e0a68a796cfee66c613e7c9aa24d 60 0
+f2bd5534623dcef31ae5e3e9a7322f435fd34452 60 0
 #2 := false
 decl f3 :: (-> int S2)
 decl f4 :: (-> S2 int)
@@ -12950,7 +13158,7 @@
 #311 := [quant-inst]: #310
 [unit-resolution #311 #646 #152]: false
 unsat
-839b4d54676d59b13d32133c52ac13a6709c5d35 35 0
+563a522b2ffcab4b8b92a3674b4a7e71c6c68e79 35 0
 #2 := false
 decl f3 :: (-> int S2)
 #25 := 3::int
@@ -12986,7 +13194,7 @@
 #135 := [asserted]: #33
 [mp #135 #153]: false
 unsat
-68ef8375d310e3f53bec3a99bed61517fecbfd6b 250 0
+2f498bc61394e79ef8b8e82a20d5e6a57d031207 250 0
 #2 := false
 decl f5 :: (-> int S2)
 #11 := 0::int
@@ -13237,7 +13445,7 @@
 #269 := [asserted]: #57
 [unit-resolution #269 #746]: false
 unsat
-a3c5f7082796ba5aeb6ea604bf07f9672a32dbef 265 0
+344dd4c43f015bcb0632176ba256e3bea4a444a5 265 0
 #2 := false
 decl f5 :: (-> int S2)
 #11 := 0::int
@@ -13503,282 +13711,7 @@
 #271 := [asserted]: #59
 [unit-resolution #271 #738]: false
 unsat
-557fcb6227e5e97d879f09a9b899d2b5ad291311 274 0
-#2 := false
-decl f5 :: (-> int S2)
-#11 := 0::int
-#55 := (f5 0::int)
-decl f4 :: (-> int int int)
-#53 := (f4 0::int 0::int)
-#54 := (f5 #53)
-#56 := (= #54 #55)
-#831 := (= #53 0::int)
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#29 := (f4 #8 #9)
-#850 := (pattern #29)
-#81 := -1::int
-#85 := (* -1::int #9)
-#82 := (* -1::int #8)
-#140 := (mod #82 #85)
-#361 := (+ #29 #140)
-#362 := (= #361 0::int)
-#30 := (mod #8 #9)
-#358 := (* -1::int #30)
-#359 := (+ #29 #358)
-#360 := (= #359 0::int)
-#107 := (<= #9 0::int)
-#103 := (<= #8 0::int)
-#300 := (or #103 #107)
-#301 := (not #300)
-#114 := (>= #8 0::int)
-#283 := (or #107 #114)
-#284 := (not #283)
-#307 := (or #284 #301)
-#363 := (ite #307 #360 #362)
-#357 := (= #29 0::int)
-#12 := (= #8 0::int)
-#364 := (ite #12 #357 #363)
-#356 := (= #8 #29)
-#13 := (= #9 0::int)
-#365 := (ite #13 #356 #364)
-#851 := (forall (vars (?v0 int) (?v1 int)) (:pat #850) #365)
-#368 := (forall (vars (?v0 int) (?v1 int)) #365)
-#854 := (iff #368 #851)
-#852 := (iff #365 #365)
-#853 := [refl]: #852
-#855 := [quant-intro #853]: #854
-#146 := (* -1::int #140)
-#325 := (ite #307 #30 #146)
-#328 := (ite #12 0::int #325)
-#331 := (ite #13 #8 #328)
-#334 := (= #29 #331)
-#337 := (forall (vars (?v0 int) (?v1 int)) #334)
-#369 := (iff #337 #368)
-#366 := (iff #334 #365)
-#367 := [rewrite]: #366
-#370 := [quant-intro #367]: #369
-#115 := (not #114)
-#108 := (not #107)
-#118 := (and #108 #115)
-#104 := (not #103)
-#111 := (and #104 #108)
-#121 := (or #111 #118)
-#166 := (ite #121 #30 #146)
-#169 := (ite #12 0::int #166)
-#172 := (ite #13 #8 #169)
-#175 := (= #29 #172)
-#178 := (forall (vars (?v0 int) (?v1 int)) #175)
-#338 := (iff #178 #337)
-#335 := (iff #175 #334)
-#332 := (= #172 #331)
-#329 := (= #169 #328)
-#326 := (= #166 #325)
-#310 := (iff #121 #307)
-#304 := (or #301 #284)
-#308 := (iff #304 #307)
-#309 := [rewrite]: #308
-#305 := (iff #121 #304)
-#302 := (iff #118 #284)
-#303 := [rewrite]: #302
-#281 := (iff #111 #301)
-#282 := [rewrite]: #281
-#306 := [monotonicity #282 #303]: #305
-#311 := [trans #306 #309]: #310
-#327 := [monotonicity #311]: #326
-#330 := [monotonicity #327]: #329
-#333 := [monotonicity #330]: #332
-#336 := [monotonicity #333]: #335
-#339 := [quant-intro #336]: #338
-#273 := (~ #178 #178)
-#270 := (~ #175 #175)
-#289 := [refl]: #270
-#274 := [nnf-pos #289]: #273
-#23 := (- #9)
-#22 := (- #8)
-#31 := (mod #22 #23)
-#32 := (- #31)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#33 := (ite #20 #30 #32)
-#34 := (ite #12 0::int #33)
-#35 := (ite #13 #8 #34)
-#36 := (= #29 #35)
-#37 := (forall (vars (?v0 int) (?v1 int)) #36)
-#181 := (iff #37 #178)
-#75 := (and #16 #18)
-#78 := (or #17 #75)
-#151 := (ite #78 #30 #146)
-#154 := (ite #12 0::int #151)
-#157 := (ite #13 #8 #154)
-#160 := (= #29 #157)
-#163 := (forall (vars (?v0 int) (?v1 int)) #160)
-#179 := (iff #163 #178)
-#176 := (iff #160 #175)
-#173 := (= #157 #172)
-#170 := (= #154 #169)
-#167 := (= #151 #166)
-#122 := (iff #78 #121)
-#119 := (iff #75 #118)
-#116 := (iff #18 #115)
-#117 := [rewrite]: #116
-#109 := (iff #16 #108)
-#110 := [rewrite]: #109
-#120 := [monotonicity #110 #117]: #119
-#112 := (iff #17 #111)
-#105 := (iff #15 #104)
-#106 := [rewrite]: #105
-#113 := [monotonicity #106 #110]: #112
-#123 := [monotonicity #113 #120]: #122
-#168 := [monotonicity #123]: #167
-#171 := [monotonicity #168]: #170
-#174 := [monotonicity #171]: #173
-#177 := [monotonicity #174]: #176
-#180 := [quant-intro #177]: #179
-#164 := (iff #37 #163)
-#161 := (iff #36 #160)
-#158 := (= #35 #157)
-#155 := (= #34 #154)
-#152 := (= #33 #151)
-#149 := (= #32 #146)
-#143 := (- #140)
-#147 := (= #143 #146)
-#148 := [rewrite]: #147
-#144 := (= #32 #143)
-#141 := (= #31 #140)
-#86 := (= #23 #85)
-#87 := [rewrite]: #86
-#83 := (= #22 #82)
-#84 := [rewrite]: #83
-#142 := [monotonicity #84 #87]: #141
-#145 := [monotonicity #142]: #144
-#150 := [trans #145 #148]: #149
-#79 := (iff #20 #78)
-#76 := (iff #19 #75)
-#77 := [rewrite]: #76
-#80 := [monotonicity #77]: #79
-#153 := [monotonicity #80 #150]: #152
-#156 := [monotonicity #153]: #155
-#159 := [monotonicity #156]: #158
-#162 := [monotonicity #159]: #161
-#165 := [quant-intro #162]: #164
-#182 := [trans #165 #180]: #181
-#139 := [asserted]: #37
-#183 := [mp #139 #182]: #178
-#290 := [mp~ #183 #274]: #178
-#340 := [mp #290 #339]: #337
-#371 := [mp #340 #370]: #368
-#856 := [mp #371 #855]: #851
-#646 := (not #851)
-#788 := (or #646 #831)
-#429 := (* -1::int 0::int)
-#514 := (mod #429 #429)
-#515 := (+ #53 #514)
-#516 := (= #515 0::int)
-#507 := (mod 0::int 0::int)
-#518 := (* -1::int #507)
-#519 := (+ #53 #518)
-#447 := (= #519 0::int)
-#520 := (<= 0::int 0::int)
-#517 := (or #520 #520)
-#521 := (not #517)
-#500 := (>= 0::int 0::int)
-#835 := (or #520 #500)
-#837 := (not #835)
-#494 := (or #837 #521)
-#624 := (ite #494 #447 #516)
-#505 := (= 0::int 0::int)
-#506 := (ite #505 #831 #624)
-#838 := (= 0::int #53)
-#839 := (ite #505 #838 #506)
-#789 := (or #646 #839)
-#791 := (iff #789 #788)
-#786 := (iff #788 #788)
-#792 := [rewrite]: #786
-#644 := (iff #839 #831)
-#1 := true
-#796 := (ite true #831 #831)
-#797 := (iff #796 #831)
-#803 := [rewrite]: #797
-#801 := (iff #839 #796)
-#800 := (iff #506 #831)
-#536 := (+ #53 #507)
-#811 := (= #536 0::int)
-#808 := (ite true #831 #811)
-#798 := (iff #808 #831)
-#799 := [rewrite]: #798
-#805 := (iff #506 #808)
-#522 := (iff #624 #811)
-#526 := (ite false #447 #811)
-#806 := (iff #526 #811)
-#807 := [rewrite]: #806
-#527 := (iff #624 #526)
-#815 := (iff #516 #811)
-#810 := (= #515 #536)
-#813 := (= #514 #507)
-#435 := (= #429 0::int)
-#812 := [rewrite]: #435
-#535 := [monotonicity #812 #812]: #813
-#814 := [monotonicity #535]: #810
-#525 := [monotonicity #814]: #815
-#541 := (iff #494 false)
-#830 := (or false false)
-#539 := (iff #830 false)
-#540 := [rewrite]: #539
-#816 := (iff #494 #830)
-#829 := (iff #521 false)
-#484 := (not true)
-#822 := (iff #484 false)
-#823 := [rewrite]: #822
-#468 := (iff #521 #484)
-#826 := (iff #517 true)
-#493 := (or true true)
-#818 := (iff #493 true)
-#481 := [rewrite]: #818
-#825 := (iff #517 #493)
-#832 := (iff #520 true)
-#492 := [rewrite]: #832
-#463 := [monotonicity #492 #492]: #825
-#828 := [trans #463 #481]: #826
-#469 := [monotonicity #828]: #468
-#827 := [trans #469 #823]: #829
-#824 := (iff #837 false)
-#820 := (iff #837 #484)
-#482 := (iff #835 true)
-#834 := (iff #835 #493)
-#497 := (iff #500 true)
-#833 := [rewrite]: #497
-#477 := [monotonicity #492 #833]: #834
-#483 := [trans #477 #481]: #482
-#821 := [monotonicity #483]: #820
-#819 := [trans #821 #823]: #824
-#817 := [monotonicity #819 #827]: #816
-#542 := [trans #817 #540]: #541
-#528 := [monotonicity #542 #525]: #527
-#804 := [trans #528 #807]: #522
-#840 := (iff #505 true)
-#841 := [rewrite]: #840
-#809 := [monotonicity #841 #804]: #805
-#795 := [trans #809 #799]: #800
-#836 := (iff #838 #831)
-#842 := [rewrite]: #836
-#802 := [monotonicity #841 #842 #795]: #801
-#645 := [trans #802 #803]: #644
-#785 := [monotonicity #645]: #791
-#793 := [trans #785 #792]: #791
-#790 := [quant-inst]: #789
-#787 := [mp #790 #793]: #788
-#741 := [unit-resolution #787 #856]: #831
-#742 := [monotonicity #741]: #56
-#57 := (not #56)
-#269 := [asserted]: #57
-[unit-resolution #269 #742]: false
-unsat
-32d7b5a9ad9a19e13efcf9a11e88fc7fa8b6d32d 287 0
+426e56a2fa6cae689a233b52c380ff3e1f4c5d0f 287 0
 #2 := false
 decl f5 :: (-> int S2)
 #11 := 0::int
@@ -14066,329 +13999,7 @@
 #270 := [asserted]: #58
 [unit-resolution #270 #691]: false
 unsat
-c5c377ca97ceb2064a54831562497528f636353d 321 0
-#2 := false
-decl f5 :: (-> int S2)
-decl f4 :: (-> int int int)
-#11 := 0::int
-decl f6 :: (-> S2 int)
-decl f7 :: S2
-#53 := f7
-#54 := (f6 f7)
-#55 := (f4 #54 0::int)
-#56 := (f5 #55)
-#271 := (= f7 #56)
-#795 := (f5 #54)
-#772 := (= #795 #56)
-#771 := (= #56 #795)
-#768 := (= #55 #54)
-#848 := (= #54 #55)
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#29 := (f4 #8 #9)
-#858 := (pattern #29)
-#82 := -1::int
-#86 := (* -1::int #9)
-#83 := (* -1::int #8)
-#141 := (mod #83 #86)
-#369 := (+ #29 #141)
-#370 := (= #369 0::int)
-#30 := (mod #8 #9)
-#366 := (* -1::int #30)
-#367 := (+ #29 #366)
-#368 := (= #367 0::int)
-#108 := (<= #9 0::int)
-#104 := (<= #8 0::int)
-#308 := (or #104 #108)
-#309 := (not #308)
-#115 := (>= #8 0::int)
-#291 := (or #108 #115)
-#292 := (not #291)
-#315 := (or #292 #309)
-#371 := (ite #315 #368 #370)
-#365 := (= #29 0::int)
-#12 := (= #8 0::int)
-#372 := (ite #12 #365 #371)
-#364 := (= #8 #29)
-#13 := (= #9 0::int)
-#373 := (ite #13 #364 #372)
-#859 := (forall (vars (?v0 int) (?v1 int)) (:pat #858) #373)
-#376 := (forall (vars (?v0 int) (?v1 int)) #373)
-#862 := (iff #376 #859)
-#860 := (iff #373 #373)
-#861 := [refl]: #860
-#863 := [quant-intro #861]: #862
-#147 := (* -1::int #141)
-#333 := (ite #315 #30 #147)
-#336 := (ite #12 0::int #333)
-#339 := (ite #13 #8 #336)
-#342 := (= #29 #339)
-#345 := (forall (vars (?v0 int) (?v1 int)) #342)
-#377 := (iff #345 #376)
-#374 := (iff #342 #373)
-#375 := [rewrite]: #374
-#378 := [quant-intro #375]: #377
-#116 := (not #115)
-#109 := (not #108)
-#119 := (and #109 #116)
-#105 := (not #104)
-#112 := (and #105 #109)
-#122 := (or #112 #119)
-#167 := (ite #122 #30 #147)
-#170 := (ite #12 0::int #167)
-#173 := (ite #13 #8 #170)
-#176 := (= #29 #173)
-#179 := (forall (vars (?v0 int) (?v1 int)) #176)
-#346 := (iff #179 #345)
-#343 := (iff #176 #342)
-#340 := (= #173 #339)
-#337 := (= #170 #336)
-#334 := (= #167 #333)
-#318 := (iff #122 #315)
-#312 := (or #309 #292)
-#316 := (iff #312 #315)
-#317 := [rewrite]: #316
-#313 := (iff #122 #312)
-#310 := (iff #119 #292)
-#311 := [rewrite]: #310
-#289 := (iff #112 #309)
-#290 := [rewrite]: #289
-#314 := [monotonicity #290 #311]: #313
-#319 := [trans #314 #317]: #318
-#335 := [monotonicity #319]: #334
-#338 := [monotonicity #335]: #337
-#341 := [monotonicity #338]: #340
-#344 := [monotonicity #341]: #343
-#347 := [quant-intro #344]: #346
-#281 := (~ #179 #179)
-#277 := (~ #176 #176)
-#297 := [refl]: #277
-#282 := [nnf-pos #297]: #281
-#23 := (- #9)
-#22 := (- #8)
-#31 := (mod #22 #23)
-#32 := (- #31)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#33 := (ite #20 #30 #32)
-#34 := (ite #12 0::int #33)
-#35 := (ite #13 #8 #34)
-#36 := (= #29 #35)
-#37 := (forall (vars (?v0 int) (?v1 int)) #36)
-#182 := (iff #37 #179)
-#76 := (and #16 #18)
-#79 := (or #17 #76)
-#152 := (ite #79 #30 #147)
-#155 := (ite #12 0::int #152)
-#158 := (ite #13 #8 #155)
-#161 := (= #29 #158)
-#164 := (forall (vars (?v0 int) (?v1 int)) #161)
-#180 := (iff #164 #179)
-#177 := (iff #161 #176)
-#174 := (= #158 #173)
-#171 := (= #155 #170)
-#168 := (= #152 #167)
-#123 := (iff #79 #122)
-#120 := (iff #76 #119)
-#117 := (iff #18 #116)
-#118 := [rewrite]: #117
-#110 := (iff #16 #109)
-#111 := [rewrite]: #110
-#121 := [monotonicity #111 #118]: #120
-#113 := (iff #17 #112)
-#106 := (iff #15 #105)
-#107 := [rewrite]: #106
-#114 := [monotonicity #107 #111]: #113
-#124 := [monotonicity #114 #121]: #123
-#169 := [monotonicity #124]: #168
-#172 := [monotonicity #169]: #171
-#175 := [monotonicity #172]: #174
-#178 := [monotonicity #175]: #177
-#181 := [quant-intro #178]: #180
-#165 := (iff #37 #164)
-#162 := (iff #36 #161)
-#159 := (= #35 #158)
-#156 := (= #34 #155)
-#153 := (= #33 #152)
-#150 := (= #32 #147)
-#144 := (- #141)
-#148 := (= #144 #147)
-#149 := [rewrite]: #148
-#145 := (= #32 #144)
-#142 := (= #31 #141)
-#87 := (= #23 #86)
-#88 := [rewrite]: #87
-#84 := (= #22 #83)
-#85 := [rewrite]: #84
-#143 := [monotonicity #85 #88]: #142
-#146 := [monotonicity #143]: #145
-#151 := [trans #146 #149]: #150
-#80 := (iff #20 #79)
-#77 := (iff #19 #76)
-#78 := [rewrite]: #77
-#81 := [monotonicity #78]: #80
-#154 := [monotonicity #81 #151]: #153
-#157 := [monotonicity #154]: #156
-#160 := [monotonicity #157]: #159
-#163 := [monotonicity #160]: #162
-#166 := [quant-intro #163]: #165
-#183 := [trans #166 #181]: #182
-#140 := [asserted]: #37
-#184 := [mp #140 #183]: #179
-#298 := [mp~ #184 #282]: #179
-#348 := [mp #298 #347]: #345
-#379 := [mp #348 #378]: #376
-#864 := [mp #379 #863]: #859
-#653 := (not #859)
-#654 := (or #653 #848)
-#437 := (* -1::int 0::int)
-#522 := (* -1::int #54)
-#523 := (mod #522 #437)
-#524 := (+ #55 #523)
-#515 := (= #524 0::int)
-#526 := (mod #54 0::int)
-#527 := (* -1::int #526)
-#455 := (+ #55 #527)
-#528 := (= #455 0::int)
-#525 := (<= 0::int 0::int)
-#529 := (<= #54 0::int)
-#508 := (or #529 #525)
-#843 := (not #508)
-#845 := (>= #54 0::int)
-#502 := (or #525 #845)
-#632 := (not #502)
-#839 := (or #632 #843)
-#513 := (ite #839 #528 #515)
-#514 := (= #55 0::int)
-#846 := (= #54 0::int)
-#847 := (ite #846 #514 #513)
-#849 := (= 0::int 0::int)
-#844 := (ite #849 #848 #847)
-#796 := (or #653 #844)
-#798 := (iff #796 #654)
-#793 := (iff #654 #654)
-#794 := [rewrite]: #793
-#811 := (iff #844 #848)
-#544 := (mod #522 0::int)
-#819 := (+ #55 #544)
-#534 := (= #819 0::int)
-#806 := (ite #846 #514 #534)
-#1 := true
-#803 := (ite true #848 #806)
-#810 := (iff #803 #848)
-#805 := [rewrite]: #810
-#804 := (iff #844 #803)
-#807 := (iff #847 #806)
-#813 := (iff #513 #534)
-#814 := (ite false #528 #534)
-#812 := (iff #814 #534)
-#816 := [rewrite]: #812
-#815 := (iff #513 #814)
-#535 := (iff #515 #534)
-#823 := (= #524 #819)
-#818 := (= #523 #544)
-#821 := (= #437 0::int)
-#543 := [rewrite]: #821
-#822 := [monotonicity #543]: #818
-#533 := [monotonicity #822]: #823
-#536 := [monotonicity #533]: #535
-#443 := (iff #839 false)
-#825 := (or false false)
-#549 := (iff #825 false)
-#550 := [rewrite]: #549
-#547 := (iff #839 #825)
-#838 := (iff #843 false)
-#491 := (not true)
-#829 := (iff #491 false)
-#830 := [rewrite]: #829
-#837 := (iff #843 #491)
-#476 := (iff #508 true)
-#827 := (or #529 true)
-#834 := (iff #827 true)
-#836 := [rewrite]: #834
-#833 := (iff #508 #827)
-#500 := (iff #525 true)
-#505 := [rewrite]: #500
-#471 := [monotonicity #505]: #833
-#477 := [trans #471 #836]: #476
-#835 := [monotonicity #477]: #837
-#824 := [trans #835 #830]: #838
-#831 := (iff #632 false)
-#492 := (iff #632 #491)
-#489 := (iff #502 true)
-#841 := (or true #845)
-#485 := (iff #841 true)
-#826 := [rewrite]: #485
-#501 := (iff #502 #841)
-#842 := [monotonicity #505]: #501
-#490 := [trans #842 #826]: #489
-#828 := [monotonicity #490]: #492
-#832 := [trans #828 #830]: #831
-#548 := [monotonicity #832 #824]: #547
-#820 := [trans #548 #550]: #443
-#530 := [monotonicity #820 #536]: #815
-#817 := [trans #530 #816]: #813
-#808 := [monotonicity #817]: #807
-#850 := (iff #849 true)
-#840 := [rewrite]: #850
-#809 := [monotonicity #840 #808]: #804
-#652 := [trans #809 #805]: #811
-#799 := [monotonicity #652]: #798
-#800 := [trans #799 #794]: #798
-#797 := [quant-inst]: #796
-#801 := [mp #797 #800]: #654
-#779 := [unit-resolution #801 #864]: #848
-#769 := [symm #779]: #768
-#765 := [monotonicity #769]: #771
-#756 := [symm #765]: #772
-#802 := (= f7 #795)
-#38 := (:var 0 S2)
-#39 := (f6 #38)
-#865 := (pattern #39)
-#40 := (f5 #39)
-#186 := (= #38 #40)
-#866 := (forall (vars (?v0 S2)) (:pat #865) #186)
-#189 := (forall (vars (?v0 S2)) #186)
-#867 := (iff #189 #866)
-#869 := (iff #866 #866)
-#870 := [rewrite]: #869
-#868 := [rewrite]: #867
-#871 := [trans #868 #870]: #867
-#283 := (~ #189 #189)
-#299 := (~ #186 #186)
-#300 := [refl]: #299
-#284 := [nnf-pos #300]: #283
-#41 := (= #40 #38)
-#42 := (forall (vars (?v0 S2)) #41)
-#190 := (iff #42 #189)
-#187 := (iff #41 #186)
-#188 := [rewrite]: #187
-#191 := [quant-intro #188]: #190
-#185 := [asserted]: #42
-#194 := [mp #185 #191]: #189
-#301 := [mp~ #194 #284]: #189
-#872 := [mp #301 #871]: #866
-#634 := (not #866)
-#787 := (or #634 #802)
-#788 := [quant-inst]: #787
-#770 := [unit-resolution #788 #872]: #802
-#757 := [trans #770 #756]: #271
-#274 := (not #271)
-#57 := (= #56 f7)
-#58 := (not #57)
-#275 := (iff #58 #274)
-#272 := (iff #57 #271)
-#273 := [rewrite]: #272
-#276 := [monotonicity #273]: #275
-#270 := [asserted]: #58
-#279 := [mp #270 #276]: #274
-[unit-resolution #279 #757]: false
-unsat
-b75d0d31bfc8b1838bed57a3059c21b7ee56aaf6 287 0
+f595f2976198bb59087cfc3afdc24567f743886b 287 0
 #2 := false
 decl f5 :: (-> int S2)
 #53 := 1::int
@@ -14676,7 +14287,7 @@
 #270 := [asserted]: #58
 [unit-resolution #270 #692]: false
 unsat
-192147dde9bbf2d1457462417b1b13f9cf65afdf 298 0
+69a73700cfd90edae90cc37e225595659ec32fc1 298 0
 #2 := false
 decl f5 :: (-> int S2)
 #53 := 3::int
@@ -14975,312 +14586,7 @@
 #271 := [asserted]: #59
 [unit-resolution #271 #706]: false
 unsat
-64092e43f0b1d98a4db3438cb496ffa4c38dbe0d 304 0
-#2 := false
-decl f5 :: (-> int S2)
-#11 := 0::int
-#56 := (f5 0::int)
-decl f4 :: (-> int int int)
-#53 := 1::int
-#54 := (f4 0::int 1::int)
-#55 := (f5 #54)
-#57 := (= #55 #56)
-#838 := (= #54 0::int)
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#29 := (f4 #8 #9)
-#850 := (pattern #29)
-#82 := -1::int
-#86 := (* -1::int #9)
-#83 := (* -1::int #8)
-#141 := (mod #83 #86)
-#362 := (+ #29 #141)
-#363 := (= #362 0::int)
-#30 := (mod #8 #9)
-#359 := (* -1::int #30)
-#360 := (+ #29 #359)
-#361 := (= #360 0::int)
-#108 := (<= #9 0::int)
-#104 := (<= #8 0::int)
-#301 := (or #104 #108)
-#302 := (not #301)
-#115 := (>= #8 0::int)
-#284 := (or #108 #115)
-#285 := (not #284)
-#308 := (or #285 #302)
-#364 := (ite #308 #361 #363)
-#358 := (= #29 0::int)
-#12 := (= #8 0::int)
-#365 := (ite #12 #358 #364)
-#357 := (= #8 #29)
-#13 := (= #9 0::int)
-#366 := (ite #13 #357 #365)
-#851 := (forall (vars (?v0 int) (?v1 int)) (:pat #850) #366)
-#369 := (forall (vars (?v0 int) (?v1 int)) #366)
-#854 := (iff #369 #851)
-#852 := (iff #366 #366)
-#853 := [refl]: #852
-#855 := [quant-intro #853]: #854
-#147 := (* -1::int #141)
-#326 := (ite #308 #30 #147)
-#329 := (ite #12 0::int #326)
-#332 := (ite #13 #8 #329)
-#335 := (= #29 #332)
-#338 := (forall (vars (?v0 int) (?v1 int)) #335)
-#370 := (iff #338 #369)
-#367 := (iff #335 #366)
-#368 := [rewrite]: #367
-#371 := [quant-intro #368]: #370
-#116 := (not #115)
-#109 := (not #108)
-#119 := (and #109 #116)
-#105 := (not #104)
-#112 := (and #105 #109)
-#122 := (or #112 #119)
-#167 := (ite #122 #30 #147)
-#170 := (ite #12 0::int #167)
-#173 := (ite #13 #8 #170)
-#176 := (= #29 #173)
-#179 := (forall (vars (?v0 int) (?v1 int)) #176)
-#339 := (iff #179 #338)
-#336 := (iff #176 #335)
-#333 := (= #173 #332)
-#330 := (= #170 #329)
-#327 := (= #167 #326)
-#311 := (iff #122 #308)
-#305 := (or #302 #285)
-#309 := (iff #305 #308)
-#310 := [rewrite]: #309
-#306 := (iff #122 #305)
-#303 := (iff #119 #285)
-#304 := [rewrite]: #303
-#282 := (iff #112 #302)
-#283 := [rewrite]: #282
-#307 := [monotonicity #283 #304]: #306
-#312 := [trans #307 #310]: #311
-#328 := [monotonicity #312]: #327
-#331 := [monotonicity #328]: #330
-#334 := [monotonicity #331]: #333
-#337 := [monotonicity #334]: #336
-#340 := [quant-intro #337]: #339
-#274 := (~ #179 #179)
-#271 := (~ #176 #176)
-#290 := [refl]: #271
-#275 := [nnf-pos #290]: #274
-#23 := (- #9)
-#22 := (- #8)
-#31 := (mod #22 #23)
-#32 := (- #31)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#33 := (ite #20 #30 #32)
-#34 := (ite #12 0::int #33)
-#35 := (ite #13 #8 #34)
-#36 := (= #29 #35)
-#37 := (forall (vars (?v0 int) (?v1 int)) #36)
-#182 := (iff #37 #179)
-#76 := (and #16 #18)
-#79 := (or #17 #76)
-#152 := (ite #79 #30 #147)
-#155 := (ite #12 0::int #152)
-#158 := (ite #13 #8 #155)
-#161 := (= #29 #158)
-#164 := (forall (vars (?v0 int) (?v1 int)) #161)
-#180 := (iff #164 #179)
-#177 := (iff #161 #176)
-#174 := (= #158 #173)
-#171 := (= #155 #170)
-#168 := (= #152 #167)
-#123 := (iff #79 #122)
-#120 := (iff #76 #119)
-#117 := (iff #18 #116)
-#118 := [rewrite]: #117
-#110 := (iff #16 #109)
-#111 := [rewrite]: #110
-#121 := [monotonicity #111 #118]: #120
-#113 := (iff #17 #112)
-#106 := (iff #15 #105)
-#107 := [rewrite]: #106
-#114 := [monotonicity #107 #111]: #113
-#124 := [monotonicity #114 #121]: #123
-#169 := [monotonicity #124]: #168
-#172 := [monotonicity #169]: #171
-#175 := [monotonicity #172]: #174
-#178 := [monotonicity #175]: #177
-#181 := [quant-intro #178]: #180
-#165 := (iff #37 #164)
-#162 := (iff #36 #161)
-#159 := (= #35 #158)
-#156 := (= #34 #155)
-#153 := (= #33 #152)
-#150 := (= #32 #147)
-#144 := (- #141)
-#148 := (= #144 #147)
-#149 := [rewrite]: #148
-#145 := (= #32 #144)
-#142 := (= #31 #141)
-#87 := (= #23 #86)
-#88 := [rewrite]: #87
-#84 := (= #22 #83)
-#85 := [rewrite]: #84
-#143 := [monotonicity #85 #88]: #142
-#146 := [monotonicity #143]: #145
-#151 := [trans #146 #149]: #150
-#80 := (iff #20 #79)
-#77 := (iff #19 #76)
-#78 := [rewrite]: #77
-#81 := [monotonicity #78]: #80
-#154 := [monotonicity #81 #151]: #153
-#157 := [monotonicity #154]: #156
-#160 := [monotonicity #157]: #159
-#163 := [monotonicity #160]: #162
-#166 := [quant-intro #163]: #165
-#183 := [trans #166 #181]: #182
-#140 := [asserted]: #37
-#184 := [mp #140 #183]: #179
-#291 := [mp~ #184 #275]: #179
-#341 := [mp #291 #340]: #338
-#372 := [mp #341 #371]: #369
-#856 := [mp #372 #855]: #851
-#676 := (not #851)
-#678 := (or #676 #838)
-#430 := (* -1::int 1::int)
-#514 := (* -1::int 0::int)
-#515 := (mod #514 #430)
-#516 := (+ #54 #515)
-#507 := (= #516 0::int)
-#518 := (mod 0::int 1::int)
-#519 := (* -1::int #518)
-#520 := (+ #54 #519)
-#517 := (= #520 0::int)
-#521 := (<= 1::int 0::int)
-#500 := (<= 0::int 0::int)
-#835 := (or #500 #521)
-#837 := (not #835)
-#494 := (>= 0::int 0::int)
-#624 := (or #521 #494)
-#831 := (not #624)
-#505 := (or #831 #837)
-#506 := (ite #505 #517 #507)
-#839 := (= 0::int 0::int)
-#840 := (ite #839 #838 #506)
-#841 := (= 0::int #54)
-#836 := (= 1::int 0::int)
-#842 := (ite #836 #841 #840)
-#679 := (or #676 #842)
-#680 := (iff #679 #678)
-#682 := (iff #678 #678)
-#683 := [rewrite]: #682
-#776 := (iff #842 #838)
-#625 := (ite false #838 #838)
-#780 := (iff #625 #838)
-#782 := [rewrite]: #780
-#772 := (iff #842 #625)
-#775 := (iff #840 #838)
-#1 := true
-#784 := (ite true #838 #838)
-#668 := (iff #784 #838)
-#627 := [rewrite]: #668
-#666 := (iff #840 #784)
-#783 := (iff #506 #838)
-#626 := (iff #506 #625)
-#794 := (iff #507 #838)
-#793 := (= #516 #54)
-#809 := (+ #54 0::int)
-#800 := (= #809 #54)
-#795 := [rewrite]: #800
-#786 := (= #516 #809)
-#791 := (= #515 0::int)
-#645 := (mod 0::int -1::int)
-#789 := (= #645 0::int)
-#790 := [rewrite]: #789
-#646 := (= #515 #645)
-#803 := (= #430 -1::int)
-#644 := [rewrite]: #803
-#522 := (= #514 0::int)
-#804 := [rewrite]: #522
-#788 := [monotonicity #804 #644]: #646
-#785 := [trans #788 #790]: #791
-#792 := [monotonicity #785]: #786
-#787 := [trans #792 #795]: #793
-#623 := [monotonicity #787]: #794
-#802 := (iff #517 #838)
-#796 := (= #520 #54)
-#798 := (= #520 #809)
-#808 := (= #519 0::int)
-#806 := (= #519 #514)
-#527 := (= #518 0::int)
-#528 := [rewrite]: #527
-#807 := [monotonicity #528]: #806
-#805 := [trans #807 #804]: #808
-#799 := [monotonicity #805]: #798
-#801 := [trans #799 #795]: #796
-#797 := [monotonicity #801]: #802
-#525 := (iff #505 false)
-#536 := (or false false)
-#811 := (iff #536 false)
-#815 := [rewrite]: #811
-#810 := (iff #505 #536)
-#813 := (iff #837 false)
-#819 := (not true)
-#826 := (iff #819 false)
-#828 := [rewrite]: #826
-#436 := (iff #837 #819)
-#541 := (iff #835 true)
-#830 := (or true false)
-#539 := (iff #830 true)
-#540 := [rewrite]: #539
-#816 := (iff #835 #830)
-#477 := (iff #521 false)
-#818 := [rewrite]: #477
-#829 := (iff #500 true)
-#827 := [rewrite]: #829
-#817 := [monotonicity #827 #818]: #816
-#542 := [trans #817 #540]: #541
-#812 := [monotonicity #542]: #436
-#535 := [trans #812 #828]: #813
-#468 := (iff #831 false)
-#825 := (iff #831 #819)
-#823 := (iff #624 true)
-#483 := (or false true)
-#821 := (iff #483 true)
-#822 := [rewrite]: #821
-#484 := (iff #624 #483)
-#481 := (iff #494 true)
-#482 := [rewrite]: #481
-#820 := [monotonicity #818 #482]: #484
-#824 := [trans #820 #822]: #823
-#463 := [monotonicity #824]: #825
-#469 := [trans #463 #828]: #468
-#814 := [monotonicity #469 #535]: #810
-#526 := [trans #814 #815]: #525
-#779 := [monotonicity #526 #797 #623]: #626
-#781 := [trans #779 #782]: #783
-#493 := (iff #839 true)
-#834 := [rewrite]: #493
-#667 := [monotonicity #834 #781]: #666
-#677 := [trans #667 #627]: #775
-#497 := (iff #841 #838)
-#833 := [rewrite]: #497
-#832 := (iff #836 false)
-#492 := [rewrite]: #832
-#773 := [monotonicity #492 #833 #677]: #772
-#661 := [trans #773 #782]: #776
-#681 := [monotonicity #661]: #680
-#684 := [trans #681 #683]: #680
-#672 := [quant-inst]: #679
-#777 := [mp #672 #684]: #678
-#712 := [unit-resolution #777 #856]: #838
-#708 := [monotonicity #712]: #57
-#58 := (not #57)
-#270 := [asserted]: #58
-[unit-resolution #270 #708]: false
-unsat
-9ffddad65860c13b15f4fc36ad26084e810a67ad 360 0
+4183c3ff4a9f3923626099be345775dc19db2d14 360 0
 #2 := false
 decl f5 :: (-> int S2)
 decl f3 :: (-> int int int)
@@ -15641,310 +14947,7 @@
 #557 := [trans #620 #572]: #272
 [unit-resolution #280 #557]: false
 unsat
-d5a593ec8919de39ee6c16ac3310bcb1eb62e0ae 302 0
-#2 := false
-decl f5 :: (-> int S2)
-#11 := 0::int
-#56 := (f5 0::int)
-decl f4 :: (-> int int int)
-#53 := 1::int
-#54 := (f4 1::int 1::int)
-#55 := (f5 #54)
-#57 := (= #55 #56)
-#505 := (= #54 0::int)
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#29 := (f4 #8 #9)
-#850 := (pattern #29)
-#82 := -1::int
-#86 := (* -1::int #9)
-#83 := (* -1::int #8)
-#141 := (mod #83 #86)
-#362 := (+ #29 #141)
-#363 := (= #362 0::int)
-#30 := (mod #8 #9)
-#359 := (* -1::int #30)
-#360 := (+ #29 #359)
-#361 := (= #360 0::int)
-#108 := (<= #9 0::int)
-#104 := (<= #8 0::int)
-#301 := (or #104 #108)
-#302 := (not #301)
-#115 := (>= #8 0::int)
-#284 := (or #108 #115)
-#285 := (not #284)
-#308 := (or #285 #302)
-#364 := (ite #308 #361 #363)
-#358 := (= #29 0::int)
-#12 := (= #8 0::int)
-#365 := (ite #12 #358 #364)
-#357 := (= #8 #29)
-#13 := (= #9 0::int)
-#366 := (ite #13 #357 #365)
-#851 := (forall (vars (?v0 int) (?v1 int)) (:pat #850) #366)
-#369 := (forall (vars (?v0 int) (?v1 int)) #366)
-#854 := (iff #369 #851)
-#852 := (iff #366 #366)
-#853 := [refl]: #852
-#855 := [quant-intro #853]: #854
-#147 := (* -1::int #141)
-#326 := (ite #308 #30 #147)
-#329 := (ite #12 0::int #326)
-#332 := (ite #13 #8 #329)
-#335 := (= #29 #332)
-#338 := (forall (vars (?v0 int) (?v1 int)) #335)
-#370 := (iff #338 #369)
-#367 := (iff #335 #366)
-#368 := [rewrite]: #367
-#371 := [quant-intro #368]: #370
-#116 := (not #115)
-#109 := (not #108)
-#119 := (and #109 #116)
-#105 := (not #104)
-#112 := (and #105 #109)
-#122 := (or #112 #119)
-#167 := (ite #122 #30 #147)
-#170 := (ite #12 0::int #167)
-#173 := (ite #13 #8 #170)
-#176 := (= #29 #173)
-#179 := (forall (vars (?v0 int) (?v1 int)) #176)
-#339 := (iff #179 #338)
-#336 := (iff #176 #335)
-#333 := (= #173 #332)
-#330 := (= #170 #329)
-#327 := (= #167 #326)
-#311 := (iff #122 #308)
-#305 := (or #302 #285)
-#309 := (iff #305 #308)
-#310 := [rewrite]: #309
-#306 := (iff #122 #305)
-#303 := (iff #119 #285)
-#304 := [rewrite]: #303
-#282 := (iff #112 #302)
-#283 := [rewrite]: #282
-#307 := [monotonicity #283 #304]: #306
-#312 := [trans #307 #310]: #311
-#328 := [monotonicity #312]: #327
-#331 := [monotonicity #328]: #330
-#334 := [monotonicity #331]: #333
-#337 := [monotonicity #334]: #336
-#340 := [quant-intro #337]: #339
-#274 := (~ #179 #179)
-#271 := (~ #176 #176)
-#290 := [refl]: #271
-#275 := [nnf-pos #290]: #274
-#23 := (- #9)
-#22 := (- #8)
-#31 := (mod #22 #23)
-#32 := (- #31)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#33 := (ite #20 #30 #32)
-#34 := (ite #12 0::int #33)
-#35 := (ite #13 #8 #34)
-#36 := (= #29 #35)
-#37 := (forall (vars (?v0 int) (?v1 int)) #36)
-#182 := (iff #37 #179)
-#76 := (and #16 #18)
-#79 := (or #17 #76)
-#152 := (ite #79 #30 #147)
-#155 := (ite #12 0::int #152)
-#158 := (ite #13 #8 #155)
-#161 := (= #29 #158)
-#164 := (forall (vars (?v0 int) (?v1 int)) #161)
-#180 := (iff #164 #179)
-#177 := (iff #161 #176)
-#174 := (= #158 #173)
-#171 := (= #155 #170)
-#168 := (= #152 #167)
-#123 := (iff #79 #122)
-#120 := (iff #76 #119)
-#117 := (iff #18 #116)
-#118 := [rewrite]: #117
-#110 := (iff #16 #109)
-#111 := [rewrite]: #110
-#121 := [monotonicity #111 #118]: #120
-#113 := (iff #17 #112)
-#106 := (iff #15 #105)
-#107 := [rewrite]: #106
-#114 := [monotonicity #107 #111]: #113
-#124 := [monotonicity #114 #121]: #123
-#169 := [monotonicity #124]: #168
-#172 := [monotonicity #169]: #171
-#175 := [monotonicity #172]: #174
-#178 := [monotonicity #175]: #177
-#181 := [quant-intro #178]: #180
-#165 := (iff #37 #164)
-#162 := (iff #36 #161)
-#159 := (= #35 #158)
-#156 := (= #34 #155)
-#153 := (= #33 #152)
-#150 := (= #32 #147)
-#144 := (- #141)
-#148 := (= #144 #147)
-#149 := [rewrite]: #148
-#145 := (= #32 #144)
-#142 := (= #31 #141)
-#87 := (= #23 #86)
-#88 := [rewrite]: #87
-#84 := (= #22 #83)
-#85 := [rewrite]: #84
-#143 := [monotonicity #85 #88]: #142
-#146 := [monotonicity #143]: #145
-#151 := [trans #146 #149]: #150
-#80 := (iff #20 #79)
-#77 := (iff #19 #76)
-#78 := [rewrite]: #77
-#81 := [monotonicity #78]: #80
-#154 := [monotonicity #81 #151]: #153
-#157 := [monotonicity #154]: #156
-#160 := [monotonicity #157]: #159
-#163 := [monotonicity #160]: #162
-#166 := [quant-intro #163]: #165
-#183 := [trans #166 #181]: #182
-#140 := [asserted]: #37
-#184 := [mp #140 #183]: #179
-#291 := [mp~ #184 #275]: #179
-#341 := [mp #291 #340]: #338
-#372 := [mp #341 #371]: #369
-#856 := [mp #372 #855]: #851
-#776 := (not #851)
-#661 := (or #776 #505)
-#430 := (* -1::int 1::int)
-#514 := (mod #430 #430)
-#515 := (+ #54 #514)
-#516 := (= #515 0::int)
-#507 := (mod 1::int 1::int)
-#518 := (* -1::int #507)
-#519 := (+ #54 #518)
-#520 := (= #519 0::int)
-#517 := (<= 1::int 0::int)
-#521 := (or #517 #517)
-#500 := (not #521)
-#835 := (>= 1::int 0::int)
-#837 := (or #517 #835)
-#494 := (not #837)
-#624 := (or #494 #500)
-#831 := (ite #624 #520 #516)
-#506 := (= 1::int 0::int)
-#838 := (ite #506 #505 #831)
-#839 := (= 1::int #54)
-#840 := (ite #506 #839 #838)
-#676 := (or #776 #840)
-#679 := (iff #676 #661)
-#680 := (iff #661 #661)
-#681 := [rewrite]: #680
-#772 := (iff #840 #505)
-#832 := (= #54 1::int)
-#667 := (ite false #832 #505)
-#775 := (iff #667 #505)
-#677 := [rewrite]: #775
-#668 := (iff #840 #667)
-#784 := (iff #838 #505)
-#779 := (ite false #505 #505)
-#783 := (iff #779 #505)
-#781 := [rewrite]: #783
-#780 := (iff #838 #779)
-#625 := (iff #831 #505)
-#1 := true
-#792 := (ite true #505 #505)
-#794 := (iff #792 #505)
-#623 := [rewrite]: #794
-#793 := (iff #831 #792)
-#785 := (iff #516 #505)
-#790 := (= #515 #54)
-#807 := (+ #54 0::int)
-#808 := (= #807 #54)
-#805 := [rewrite]: #808
-#788 := (= #515 #807)
-#645 := (= #514 0::int)
-#801 := (mod -1::int -1::int)
-#803 := (= #801 0::int)
-#644 := [rewrite]: #803
-#802 := (= #514 #801)
-#795 := (= #430 -1::int)
-#796 := [rewrite]: #795
-#797 := [monotonicity #796 #796]: #802
-#646 := [trans #797 #644]: #645
-#789 := [monotonicity #646]: #788
-#791 := [trans #789 #805]: #790
-#786 := [monotonicity #791]: #785
-#799 := (iff #520 #505)
-#809 := (= #519 #54)
-#522 := (= #519 #807)
-#528 := (= #518 0::int)
-#811 := (* -1::int 0::int)
-#526 := (= #811 0::int)
-#527 := [rewrite]: #526
-#815 := (= #518 #811)
-#810 := (= #507 0::int)
-#814 := [rewrite]: #810
-#525 := [monotonicity #814]: #815
-#806 := [trans #525 #527]: #528
-#804 := [monotonicity #806]: #522
-#798 := [trans #804 #805]: #809
-#800 := [monotonicity #798]: #799
-#535 := (iff #624 true)
-#477 := (or false true)
-#482 := (iff #477 true)
-#483 := [rewrite]: #482
-#812 := (iff #624 #477)
-#542 := (iff #500 true)
-#816 := (not false)
-#540 := (iff #816 true)
-#541 := [rewrite]: #540
-#817 := (iff #500 #816)
-#827 := (iff #521 false)
-#826 := (or false false)
-#469 := (iff #826 false)
-#829 := [rewrite]: #469
-#828 := (iff #521 #826)
-#497 := (iff #517 false)
-#833 := [rewrite]: #497
-#468 := [monotonicity #833 #833]: #828
-#830 := [trans #468 #829]: #827
-#539 := [monotonicity #830]: #817
-#436 := [trans #539 #541]: #542
-#825 := (iff #494 false)
-#821 := (not true)
-#824 := (iff #821 false)
-#819 := [rewrite]: #824
-#822 := (iff #494 #821)
-#484 := (iff #837 true)
-#818 := (iff #837 #477)
-#493 := (iff #835 true)
-#834 := [rewrite]: #493
-#481 := [monotonicity #833 #834]: #818
-#820 := [trans #481 #483]: #484
-#823 := [monotonicity #820]: #822
-#463 := [trans #823 #819]: #825
-#813 := [monotonicity #463 #436]: #812
-#536 := [trans #813 #483]: #535
-#787 := [monotonicity #536 #800 #786]: #793
-#626 := [trans #787 #623]: #625
-#841 := (iff #506 false)
-#836 := [rewrite]: #841
-#782 := [monotonicity #836 #626]: #780
-#666 := [trans #782 #781]: #784
-#842 := (iff #839 #832)
-#492 := [rewrite]: #842
-#627 := [monotonicity #836 #492 #666]: #668
-#773 := [trans #627 #677]: #772
-#672 := [monotonicity #773]: #679
-#682 := [trans #672 #681]: #679
-#678 := [quant-inst]: #676
-#683 := [mp #678 #682]: #661
-#703 := [unit-resolution #683 #856]: #505
-#699 := [monotonicity #703]: #57
-#58 := (not #57)
-#270 := [asserted]: #58
-[unit-resolution #270 #699]: false
-unsat
-bc318995e409f6e5340f5c18fb5eab5001ed9343 288 0
+f2f039ccfb9f8bb6046ed67b2d5f9e4965b93c1c 288 0
 #2 := false
 decl f5 :: (-> int S2)
 #11 := 0::int
@@ -16233,321 +15236,7 @@
 #270 := [asserted]: #58
 [unit-resolution #270 #692]: false
 unsat
-6860ddff5f419fe559575a6c298a4b8f50cb8ee7 313 0
-#2 := false
-decl f5 :: (-> int S2)
-#11 := 0::int
-#57 := (f5 0::int)
-decl f4 :: (-> int int int)
-#54 := 1::int
-#53 := 3::int
-#55 := (f4 3::int 1::int)
-#56 := (f5 #55)
-#58 := (= #56 #57)
-#839 := (= #55 0::int)
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#29 := (f4 #8 #9)
-#851 := (pattern #29)
-#83 := -1::int
-#87 := (* -1::int #9)
-#84 := (* -1::int #8)
-#142 := (mod #84 #87)
-#363 := (+ #29 #142)
-#364 := (= #363 0::int)
-#30 := (mod #8 #9)
-#360 := (* -1::int #30)
-#361 := (+ #29 #360)
-#362 := (= #361 0::int)
-#109 := (<= #9 0::int)
-#105 := (<= #8 0::int)
-#302 := (or #105 #109)
-#303 := (not #302)
-#116 := (>= #8 0::int)
-#285 := (or #109 #116)
-#286 := (not #285)
-#309 := (or #286 #303)
-#365 := (ite #309 #362 #364)
-#359 := (= #29 0::int)
-#12 := (= #8 0::int)
-#366 := (ite #12 #359 #365)
-#358 := (= #8 #29)
-#13 := (= #9 0::int)
-#367 := (ite #13 #358 #366)
-#852 := (forall (vars (?v0 int) (?v1 int)) (:pat #851) #367)
-#370 := (forall (vars (?v0 int) (?v1 int)) #367)
-#855 := (iff #370 #852)
-#853 := (iff #367 #367)
-#854 := [refl]: #853
-#856 := [quant-intro #854]: #855
-#148 := (* -1::int #142)
-#327 := (ite #309 #30 #148)
-#330 := (ite #12 0::int #327)
-#333 := (ite #13 #8 #330)
-#336 := (= #29 #333)
-#339 := (forall (vars (?v0 int) (?v1 int)) #336)
-#371 := (iff #339 #370)
-#368 := (iff #336 #367)
-#369 := [rewrite]: #368
-#372 := [quant-intro #369]: #371
-#117 := (not #116)
-#110 := (not #109)
-#120 := (and #110 #117)
-#106 := (not #105)
-#113 := (and #106 #110)
-#123 := (or #113 #120)
-#168 := (ite #123 #30 #148)
-#171 := (ite #12 0::int #168)
-#174 := (ite #13 #8 #171)
-#177 := (= #29 #174)
-#180 := (forall (vars (?v0 int) (?v1 int)) #177)
-#340 := (iff #180 #339)
-#337 := (iff #177 #336)
-#334 := (= #174 #333)
-#331 := (= #171 #330)
-#328 := (= #168 #327)
-#312 := (iff #123 #309)
-#306 := (or #303 #286)
-#310 := (iff #306 #309)
-#311 := [rewrite]: #310
-#307 := (iff #123 #306)
-#304 := (iff #120 #286)
-#305 := [rewrite]: #304
-#283 := (iff #113 #303)
-#284 := [rewrite]: #283
-#308 := [monotonicity #284 #305]: #307
-#313 := [trans #308 #311]: #312
-#329 := [monotonicity #313]: #328
-#332 := [monotonicity #329]: #331
-#335 := [monotonicity #332]: #334
-#338 := [monotonicity #335]: #337
-#341 := [quant-intro #338]: #340
-#275 := (~ #180 #180)
-#272 := (~ #177 #177)
-#291 := [refl]: #272
-#276 := [nnf-pos #291]: #275
-#23 := (- #9)
-#22 := (- #8)
-#31 := (mod #22 #23)
-#32 := (- #31)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#33 := (ite #20 #30 #32)
-#34 := (ite #12 0::int #33)
-#35 := (ite #13 #8 #34)
-#36 := (= #29 #35)
-#37 := (forall (vars (?v0 int) (?v1 int)) #36)
-#183 := (iff #37 #180)
-#77 := (and #16 #18)
-#80 := (or #17 #77)
-#153 := (ite #80 #30 #148)
-#156 := (ite #12 0::int #153)
-#159 := (ite #13 #8 #156)
-#162 := (= #29 #159)
-#165 := (forall (vars (?v0 int) (?v1 int)) #162)
-#181 := (iff #165 #180)
-#178 := (iff #162 #177)
-#175 := (= #159 #174)
-#172 := (= #156 #171)
-#169 := (= #153 #168)
-#124 := (iff #80 #123)
-#121 := (iff #77 #120)
-#118 := (iff #18 #117)
-#119 := [rewrite]: #118
-#111 := (iff #16 #110)
-#112 := [rewrite]: #111
-#122 := [monotonicity #112 #119]: #121
-#114 := (iff #17 #113)
-#107 := (iff #15 #106)
-#108 := [rewrite]: #107
-#115 := [monotonicity #108 #112]: #114
-#125 := [monotonicity #115 #122]: #124
-#170 := [monotonicity #125]: #169
-#173 := [monotonicity #170]: #172
-#176 := [monotonicity #173]: #175
-#179 := [monotonicity #176]: #178
-#182 := [quant-intro #179]: #181
-#166 := (iff #37 #165)
-#163 := (iff #36 #162)
-#160 := (= #35 #159)
-#157 := (= #34 #156)
-#154 := (= #33 #153)
-#151 := (= #32 #148)
-#145 := (- #142)
-#149 := (= #145 #148)
-#150 := [rewrite]: #149
-#146 := (= #32 #145)
-#143 := (= #31 #142)
-#88 := (= #23 #87)
-#89 := [rewrite]: #88
-#85 := (= #22 #84)
-#86 := [rewrite]: #85
-#144 := [monotonicity #86 #89]: #143
-#147 := [monotonicity #144]: #146
-#152 := [trans #147 #150]: #151
-#81 := (iff #20 #80)
-#78 := (iff #19 #77)
-#79 := [rewrite]: #78
-#82 := [monotonicity #79]: #81
-#155 := [monotonicity #82 #152]: #154
-#158 := [monotonicity #155]: #157
-#161 := [monotonicity #158]: #160
-#164 := [monotonicity #161]: #163
-#167 := [quant-intro #164]: #166
-#184 := [trans #167 #182]: #183
-#141 := [asserted]: #37
-#185 := [mp #141 #184]: #180
-#292 := [mp~ #185 #276]: #180
-#342 := [mp #292 #341]: #339
-#373 := [mp #342 #372]: #370
-#857 := [mp #373 #856]: #852
-#685 := (not #852)
-#778 := (or #685 #839)
-#431 := (* -1::int 1::int)
-#515 := (* -1::int 3::int)
-#516 := (mod #515 #431)
-#517 := (+ #55 #516)
-#508 := (= #517 0::int)
-#519 := (mod 3::int 1::int)
-#520 := (* -1::int #519)
-#521 := (+ #55 #520)
-#518 := (= #521 0::int)
-#522 := (<= 1::int 0::int)
-#501 := (<= 3::int 0::int)
-#836 := (or #501 #522)
-#838 := (not #836)
-#495 := (>= 3::int 0::int)
-#625 := (or #522 #495)
-#832 := (not #625)
-#506 := (or #832 #838)
-#507 := (ite #506 #518 #508)
-#840 := (= 3::int 0::int)
-#841 := (ite #840 #839 #507)
-#842 := (= 3::int #55)
-#837 := (= 1::int 0::int)
-#843 := (ite #837 #842 #841)
-#775 := (or #685 #843)
-#766 := (iff #775 #778)
-#760 := (iff #778 #778)
-#757 := [rewrite]: #760
-#683 := (iff #843 #839)
-#834 := (= #55 3::int)
-#679 := (ite false #834 #839)
-#681 := (iff #679 #839)
-#682 := [rewrite]: #681
-#680 := (iff #843 #679)
-#662 := (iff #841 #839)
-#776 := (ite false #839 #839)
-#774 := (iff #776 #839)
-#777 := [rewrite]: #774
-#678 := (iff #841 #776)
-#669 := (iff #507 #839)
-#1 := true
-#784 := (ite true #839 #839)
-#667 := (iff #784 #839)
-#668 := [rewrite]: #667
-#782 := (iff #507 #784)
-#781 := (iff #508 #839)
-#627 := (= #517 #55)
-#800 := (+ #55 0::int)
-#797 := (= #800 #55)
-#802 := [rewrite]: #797
-#624 := (= #517 #800)
-#788 := (= #516 0::int)
-#646 := -3::int
-#792 := (mod -3::int -1::int)
-#793 := (= #792 0::int)
-#794 := [rewrite]: #793
-#786 := (= #516 #792)
-#790 := (= #431 -1::int)
-#791 := [rewrite]: #790
-#647 := (= #515 -3::int)
-#789 := [rewrite]: #647
-#787 := [monotonicity #789 #791]: #786
-#795 := [trans #787 #794]: #788
-#626 := [monotonicity #795]: #624
-#780 := [trans #626 #802]: #627
-#783 := [monotonicity #780]: #781
-#804 := (iff #518 #839)
-#803 := (= #521 #55)
-#801 := (= #521 #800)
-#810 := (= #520 0::int)
-#808 := (* -1::int 0::int)
-#809 := (= #808 0::int)
-#806 := [rewrite]: #809
-#523 := (= #520 #808)
-#529 := (= #519 0::int)
-#807 := [rewrite]: #529
-#805 := [monotonicity #807]: #523
-#799 := [trans #805 #806]: #810
-#796 := [monotonicity #799]: #801
-#798 := [trans #796 #802]: #803
-#645 := [monotonicity #798]: #804
-#527 := (iff #506 true)
-#485 := (or false true)
-#823 := (iff #485 true)
-#824 := [rewrite]: #823
-#816 := (iff #506 #485)
-#815 := (iff #838 true)
-#813 := (not false)
-#537 := (iff #813 true)
-#811 := [rewrite]: #537
-#814 := (iff #838 #813)
-#543 := (iff #836 false)
-#817 := (or false false)
-#541 := (iff #817 false)
-#542 := [rewrite]: #541
-#818 := (iff #836 #817)
-#819 := (iff #522 false)
-#482 := [rewrite]: #819
-#828 := (iff #501 false)
-#831 := [rewrite]: #828
-#540 := [monotonicity #831 #482]: #818
-#437 := [trans #540 #542]: #543
-#536 := [monotonicity #437]: #814
-#812 := [trans #536 #811]: #815
-#470 := (iff #832 false)
-#826 := (not true)
-#829 := (iff #826 false)
-#469 := [rewrite]: #829
-#464 := (iff #832 #826)
-#825 := (iff #625 true)
-#821 := (iff #625 #485)
-#483 := (iff #495 true)
-#484 := [rewrite]: #483
-#822 := [monotonicity #482 #484]: #821
-#820 := [trans #822 #824]: #825
-#827 := [monotonicity #820]: #464
-#830 := [trans #827 #469]: #470
-#526 := [monotonicity #830 #812]: #816
-#528 := [trans #526 #824]: #527
-#785 := [monotonicity #528 #645 #783]: #782
-#628 := [trans #785 #668]: #669
-#835 := (iff #840 false)
-#478 := [rewrite]: #835
-#773 := [monotonicity #478 #628]: #678
-#677 := [trans #773 #777]: #662
-#498 := (iff #842 #834)
-#494 := [rewrite]: #498
-#833 := (iff #837 false)
-#493 := [rewrite]: #833
-#673 := [monotonicity #493 #494 #677]: #680
-#684 := [trans #673 #682]: #683
-#768 := [monotonicity #684]: #766
-#759 := [trans #768 #757]: #766
-#779 := [quant-inst]: #775
-#769 := [mp #779 #759]: #778
-#701 := [unit-resolution #769 #857]: #839
-#702 := [monotonicity #701]: #58
-#59 := (not #58)
-#271 := [asserted]: #59
-[unit-resolution #271 #702]: false
-unsat
-55de26cad67b2fd694b0a728f5c662e01cc56f2a 302 0
+ed1dcbda8b48a8192bb34e0311b0082775d33ed7 302 0
 #2 := false
 decl f5 :: (-> int S2)
 #11 := 0::int
@@ -16850,455 +15539,7 @@
 #271 := [asserted]: #59
 [unit-resolution #271 #702]: false
 unsat
-921cbc00d40a6ed7847a6fd4ceb1178d690d0926 447 0
-#2 := false
-#11 := 0::int
-decl f6 :: (-> S2 int)
-decl f7 :: S2
-#53 := f7
-#54 := (f6 f7)
-#496 := (>= #54 0::int)
-decl f5 :: (-> int S2)
-#762 := (f5 #54)
-#708 := (f6 #762)
-#704 := (= #708 0::int)
-#655 := (not #704)
-#841 := (= #54 0::int)
-#674 := (not #841)
-#656 := (iff #674 #655)
-#653 := (iff #841 #704)
-#651 := (iff #704 #841)
-#649 := (= #708 #54)
-#643 := (= #762 f7)
-#763 := (= f7 #762)
-#38 := (:var 0 S2)
-#39 := (f6 #38)
-#859 := (pattern #39)
-#40 := (f5 #39)
-#188 := (= #38 #40)
-#860 := (forall (vars (?v0 S2)) (:pat #859) #188)
-#191 := (forall (vars (?v0 S2)) #188)
-#861 := (iff #191 #860)
-#863 := (iff #860 #860)
-#864 := [rewrite]: #863
-#862 := [rewrite]: #861
-#865 := [trans #862 #864]: #861
-#278 := (~ #191 #191)
-#294 := (~ #188 #188)
-#295 := [refl]: #294
-#279 := [nnf-pos #295]: #278
-#41 := (= #40 #38)
-#42 := (forall (vars (?v0 S2)) #41)
-#192 := (iff #42 #191)
-#189 := (iff #41 #188)
-#190 := [rewrite]: #189
-#193 := [quant-intro #190]: #192
-#187 := [asserted]: #42
-#196 := [mp #187 #193]: #191
-#296 := [mp~ #196 #279]: #191
-#866 := [mp #296 #865]: #860
-#759 := (not #860)
-#766 := (or #759 #763)
-#750 := [quant-inst]: #766
-#688 := [unit-resolution #750 #866]: #763
-#644 := [symm #688]: #643
-#650 := [monotonicity #644]: #649
-#652 := [monotonicity #650]: #651
-#654 := [symm #652]: #653
-#657 := [monotonicity #654]: #656
-#486 := (not #496)
-#673 := [hypothesis]: #486
-#677 := (or #674 #496)
-#687 := [th-lemma]: #677
-#667 := [unit-resolution #687 #673]: #674
-#658 := [mp #667 #657]: #655
-#689 := (or #496 #704)
-#9 := (:var 0 int)
-#44 := (f5 #9)
-#867 := (pattern #44)
-#211 := (>= #9 0::int)
-#45 := (f6 #44)
-#50 := (= #45 0::int)
-#261 := (or #50 #211)
-#874 := (forall (vars (?v0 int)) (:pat #867) #261)
-#266 := (forall (vars (?v0 int)) #261)
-#877 := (iff #266 #874)
-#875 := (iff #261 #261)
-#876 := [refl]: #875
-#878 := [quant-intro #876]: #877
-#282 := (~ #266 #266)
-#300 := (~ #261 #261)
-#301 := [refl]: #300
-#283 := [nnf-pos #301]: #282
-#49 := (< #9 0::int)
-#51 := (implies #49 #50)
-#52 := (forall (vars (?v0 int)) #51)
-#269 := (iff #52 #266)
-#232 := (= 0::int #45)
-#238 := (not #49)
-#239 := (or #238 #232)
-#244 := (forall (vars (?v0 int)) #239)
-#267 := (iff #244 #266)
-#264 := (iff #239 #261)
-#258 := (or #211 #50)
-#262 := (iff #258 #261)
-#263 := [rewrite]: #262
-#259 := (iff #239 #258)
-#256 := (iff #232 #50)
-#257 := [rewrite]: #256
-#254 := (iff #238 #211)
-#214 := (not #211)
-#249 := (not #214)
-#252 := (iff #249 #211)
-#253 := [rewrite]: #252
-#250 := (iff #238 #249)
-#247 := (iff #49 #214)
-#248 := [rewrite]: #247
-#251 := [monotonicity #248]: #250
-#255 := [trans #251 #253]: #254
-#260 := [monotonicity #255 #257]: #259
-#265 := [trans #260 #263]: #264
-#268 := [quant-intro #265]: #267
-#245 := (iff #52 #244)
-#242 := (iff #51 #239)
-#235 := (implies #49 #232)
-#240 := (iff #235 #239)
-#241 := [rewrite]: #240
-#236 := (iff #51 #235)
-#233 := (iff #50 #232)
-#234 := [rewrite]: #233
-#237 := [monotonicity #234]: #236
-#243 := [trans #237 #241]: #242
-#246 := [quant-intro #243]: #245
-#270 := [trans #246 #268]: #269
-#231 := [asserted]: #52
-#271 := [mp #231 #270]: #266
-#302 := [mp~ #271 #283]: #266
-#879 := [mp #302 #878]: #874
-#729 := (not #874)
-#671 := (or #729 #496 #704)
-#709 := (or #704 #496)
-#695 := (or #729 #709)
-#662 := (iff #695 #671)
-#691 := (or #729 #689)
-#672 := (iff #691 #671)
-#631 := [rewrite]: #672
-#697 := (iff #695 #691)
-#635 := (iff #709 #689)
-#690 := [rewrite]: #635
-#665 := [monotonicity #690]: #697
-#664 := [trans #665 #631]: #662
-#696 := [quant-inst]: #695
-#666 := [mp #696 #664]: #671
-#675 := [unit-resolution #666 #879]: #689
-#676 := [unit-resolution #675 #673]: #704
-#659 := [unit-resolution #676 #658]: false
-#660 := [lemma #659]: #496
-#502 := (<= #54 0::int)
-#830 := (not #502)
-#625 := (or #486 #830 #841)
-#684 := (not #625)
-decl f4 :: (-> int int int)
-#55 := 1::int
-#56 := (f4 #54 1::int)
-#840 := (= #56 0::int)
-#760 := (not #840)
-#58 := (f5 0::int)
-#57 := (f5 #56)
-#59 := (= #57 #58)
-#645 := [hypothesis]: #840
-#661 := [monotonicity #645]: #59
-#60 := (not #59)
-#272 := [asserted]: #60
-#622 := [unit-resolution #272 #661]: false
-#623 := [lemma #622]: #760
-#632 := (or #684 #840)
-#84 := -1::int
-#516 := (* -1::int #54)
-#809 := (mod #516 -1::int)
-#810 := (+ #56 #809)
-#800 := (= #810 0::int)
-#781 := (ite #625 #840 #800)
-#8 := (:var 1 int)
-#29 := (f4 #8 #9)
-#852 := (pattern #29)
-#88 := (* -1::int #9)
-#85 := (* -1::int #8)
-#143 := (mod #85 #88)
-#364 := (+ #29 #143)
-#365 := (= #364 0::int)
-#30 := (mod #8 #9)
-#361 := (* -1::int #30)
-#362 := (+ #29 #361)
-#363 := (= #362 0::int)
-#110 := (<= #9 0::int)
-#106 := (<= #8 0::int)
-#303 := (or #106 #110)
-#304 := (not #303)
-#117 := (>= #8 0::int)
-#286 := (or #110 #117)
-#287 := (not #286)
-#310 := (or #287 #304)
-#366 := (ite #310 #363 #365)
-#360 := (= #29 0::int)
-#12 := (= #8 0::int)
-#367 := (ite #12 #360 #366)
-#359 := (= #8 #29)
-#13 := (= #9 0::int)
-#368 := (ite #13 #359 #367)
-#853 := (forall (vars (?v0 int) (?v1 int)) (:pat #852) #368)
-#371 := (forall (vars (?v0 int) (?v1 int)) #368)
-#856 := (iff #371 #853)
-#854 := (iff #368 #368)
-#855 := [refl]: #854
-#857 := [quant-intro #855]: #856
-#149 := (* -1::int #143)
-#328 := (ite #310 #30 #149)
-#331 := (ite #12 0::int #328)
-#334 := (ite #13 #8 #331)
-#337 := (= #29 #334)
-#340 := (forall (vars (?v0 int) (?v1 int)) #337)
-#372 := (iff #340 #371)
-#369 := (iff #337 #368)
-#370 := [rewrite]: #369
-#373 := [quant-intro #370]: #372
-#118 := (not #117)
-#111 := (not #110)
-#121 := (and #111 #118)
-#107 := (not #106)
-#114 := (and #107 #111)
-#124 := (or #114 #121)
-#169 := (ite #124 #30 #149)
-#172 := (ite #12 0::int #169)
-#175 := (ite #13 #8 #172)
-#178 := (= #29 #175)
-#181 := (forall (vars (?v0 int) (?v1 int)) #178)
-#341 := (iff #181 #340)
-#338 := (iff #178 #337)
-#335 := (= #175 #334)
-#332 := (= #172 #331)
-#329 := (= #169 #328)
-#313 := (iff #124 #310)
-#307 := (or #304 #287)
-#311 := (iff #307 #310)
-#312 := [rewrite]: #311
-#308 := (iff #124 #307)
-#305 := (iff #121 #287)
-#306 := [rewrite]: #305
-#284 := (iff #114 #304)
-#285 := [rewrite]: #284
-#309 := [monotonicity #285 #306]: #308
-#314 := [trans #309 #312]: #313
-#330 := [monotonicity #314]: #329
-#333 := [monotonicity #330]: #332
-#336 := [monotonicity #333]: #335
-#339 := [monotonicity #336]: #338
-#342 := [quant-intro #339]: #341
-#276 := (~ #181 #181)
-#273 := (~ #178 #178)
-#292 := [refl]: #273
-#277 := [nnf-pos #292]: #276
-#23 := (- #9)
-#22 := (- #8)
-#31 := (mod #22 #23)
-#32 := (- #31)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#33 := (ite #20 #30 #32)
-#34 := (ite #12 0::int #33)
-#35 := (ite #13 #8 #34)
-#36 := (= #29 #35)
-#37 := (forall (vars (?v0 int) (?v1 int)) #36)
-#184 := (iff #37 #181)
-#78 := (and #16 #18)
-#81 := (or #17 #78)
-#154 := (ite #81 #30 #149)
-#157 := (ite #12 0::int #154)
-#160 := (ite #13 #8 #157)
-#163 := (= #29 #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 #81 #124)
-#122 := (iff #78 #121)
-#119 := (iff #18 #118)
-#120 := [rewrite]: #119
-#112 := (iff #16 #111)
-#113 := [rewrite]: #112
-#123 := [monotonicity #113 #120]: #122
-#115 := (iff #17 #114)
-#108 := (iff #15 #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 #37 #166)
-#164 := (iff #36 #163)
-#161 := (= #35 #160)
-#158 := (= #34 #157)
-#155 := (= #33 #154)
-#152 := (= #32 #149)
-#146 := (- #143)
-#150 := (= #146 #149)
-#151 := [rewrite]: #150
-#147 := (= #32 #146)
-#144 := (= #31 #143)
-#89 := (= #23 #88)
-#90 := [rewrite]: #89
-#86 := (= #22 #85)
-#87 := [rewrite]: #86
-#145 := [monotonicity #87 #90]: #144
-#148 := [monotonicity #145]: #147
-#153 := [trans #148 #151]: #152
-#82 := (iff #20 #81)
-#79 := (iff #19 #78)
-#80 := [rewrite]: #79
-#83 := [monotonicity #80]: #82
-#156 := [monotonicity #83 #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]: #37
-#186 := [mp #142 #185]: #181
-#293 := [mp~ #186 #277]: #181
-#343 := [mp #293 #342]: #340
-#374 := [mp #343 #373]: #371
-#858 := [mp #374 #857]: #853
-#786 := (not #853)
-#668 := (or #786 #781)
-#432 := (* -1::int 1::int)
-#517 := (mod #516 #432)
-#518 := (+ #56 #517)
-#509 := (= #518 0::int)
-#520 := (mod #54 1::int)
-#521 := (* -1::int #520)
-#522 := (+ #56 #521)
-#519 := (= #522 0::int)
-#523 := (<= 1::int 0::int)
-#837 := (or #502 #523)
-#839 := (not #837)
-#626 := (or #523 #496)
-#833 := (not #626)
-#507 := (or #833 #839)
-#508 := (ite #507 #519 #509)
-#842 := (ite #841 #840 #508)
-#843 := (= #54 #56)
-#838 := (= 1::int 0::int)
-#844 := (ite #838 #843 #842)
-#669 := (or #786 #844)
-#629 := (iff #669 #668)
-#679 := (iff #668 #668)
-#774 := [rewrite]: #679
-#785 := (iff #844 #781)
-#831 := (or #486 #830)
-#646 := (or #831 #841)
-#647 := (ite #646 #840 #800)
-#782 := (iff #647 #781)
-#627 := (iff #646 #625)
-#628 := [rewrite]: #627
-#784 := [monotonicity #628]: #782
-#789 := (iff #844 #647)
-#793 := (ite false #843 #647)
-#794 := (iff #793 #647)
-#795 := [rewrite]: #794
-#787 := (iff #844 #793)
-#791 := (iff #842 #647)
-#797 := (ite #831 #840 #800)
-#804 := (ite #841 #840 #797)
-#648 := (iff #804 #647)
-#790 := [rewrite]: #648
-#799 := (iff #842 #804)
-#798 := (iff #508 #797)
-#801 := (iff #509 #800)
-#807 := (= #518 #810)
-#524 := (= #517 #809)
-#530 := (= #432 -1::int)
-#808 := [rewrite]: #530
-#806 := [monotonicity #808]: #524
-#811 := [monotonicity #806]: #807
-#802 := [monotonicity #811]: #801
-#528 := (iff #519 #840)
-#817 := (= #522 #56)
-#537 := (+ #56 0::int)
-#816 := (= #537 #56)
-#813 := [rewrite]: #816
-#538 := (= #522 #537)
-#814 := (= #521 0::int)
-#541 := (* -1::int 0::int)
-#544 := (= #541 0::int)
-#438 := [rewrite]: #544
-#542 := (= #521 #541)
-#818 := (= #520 0::int)
-#819 := [rewrite]: #818
-#543 := [monotonicity #819]: #542
-#815 := [trans #543 #438]: #814
-#812 := [monotonicity #815]: #538
-#527 := [trans #812 #813]: #817
-#529 := [monotonicity #527]: #528
-#829 := (iff #507 #831)
-#470 := (iff #839 #830)
-#465 := (iff #837 #502)
-#824 := (or #502 false)
-#821 := (iff #824 #502)
-#827 := [rewrite]: #821
-#825 := (iff #837 #824)
-#499 := (iff #523 false)
-#835 := [rewrite]: #499
-#826 := [monotonicity #835]: #825
-#828 := [trans #826 #827]: #465
-#471 := [monotonicity #828]: #470
-#822 := (iff #833 #486)
-#484 := (iff #626 #496)
-#495 := (or false #496)
-#820 := (iff #495 #496)
-#483 := [rewrite]: #820
-#836 := (iff #626 #495)
-#479 := [monotonicity #835]: #836
-#485 := [trans #479 #483]: #484
-#823 := [monotonicity #485]: #822
-#832 := [monotonicity #823 #471]: #829
-#803 := [monotonicity #832 #529 #802]: #798
-#805 := [monotonicity #803]: #799
-#792 := [trans #805 #790]: #791
-#834 := (iff #838 false)
-#494 := [rewrite]: #834
-#788 := [monotonicity #494 #792]: #787
-#796 := [trans #788 #795]: #789
-#783 := [trans #796 #784]: #785
-#777 := [monotonicity #783]: #629
-#775 := [trans #777 #774]: #629
-#670 := [quant-inst]: #669
-#778 := [mp #670 #775]: #668
-#630 := [unit-resolution #778 #858]: #781
-#780 := (not #781)
-#767 := (or #780 #684 #840)
-#769 := [def-axiom]: #767
-#633 := [unit-resolution #769 #630]: #632
-#634 := [unit-resolution #633 #623]: #684
-#680 := (or #625 #502)
-#681 := [def-axiom]: #680
-#636 := [unit-resolution #681 #634]: #502
-#682 := (or #625 #674)
-#683 := [def-axiom]: #682
-#637 := [unit-resolution #683 #634]: #674
-#638 := (or #841 #830 #486)
-#639 := [th-lemma]: #638
-[unit-resolution #639 #637 #636 #660]: false
-unsat
-b39f51f6b3dc27b329777d3e0b54f2354e076c4d 292 0
+8f1427496ff61632aac173e792df93e6d2f8fec2 292 0
 #2 := false
 decl f5 :: (-> int S2)
 #56 := 1::int
@@ -17591,7 +15832,7 @@
 #271 := [asserted]: #59
 [unit-resolution #271 #701]: false
 unsat
-8abb85a64c00016f8d51c89dca4f66482ca1b3ed 538 0
+d3447cf6898d788767aea97d325e514adc5f9e24 538 0
 #2 := false
 #11 := 0::int
 decl f6 :: (-> S2 int)
@@ -18130,313 +16371,7 @@
 #593 := [trans #591 #608]: #781
 [unit-resolution #661 #593]: false
 unsat
-ab86a54cb616e2b66ccbd5c4b16ac320b8b6bebe 305 0
-#2 := false
-decl f5 :: (-> int S2)
-#11 := 0::int
-#56 := (f5 0::int)
-decl f4 :: (-> int int int)
-#53 := 3::int
-#54 := (f4 0::int 3::int)
-#55 := (f5 #54)
-#57 := (= #55 #56)
-#507 := (= #54 0::int)
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#29 := (f4 #8 #9)
-#851 := (pattern #29)
-#82 := -1::int
-#86 := (* -1::int #9)
-#83 := (* -1::int #8)
-#141 := (mod #83 #86)
-#362 := (+ #29 #141)
-#363 := (= #362 0::int)
-#30 := (mod #8 #9)
-#359 := (* -1::int #30)
-#360 := (+ #29 #359)
-#361 := (= #360 0::int)
-#108 := (<= #9 0::int)
-#104 := (<= #8 0::int)
-#301 := (or #104 #108)
-#302 := (not #301)
-#115 := (>= #8 0::int)
-#284 := (or #108 #115)
-#285 := (not #284)
-#308 := (or #285 #302)
-#364 := (ite #308 #361 #363)
-#358 := (= #29 0::int)
-#12 := (= #8 0::int)
-#365 := (ite #12 #358 #364)
-#357 := (= #8 #29)
-#13 := (= #9 0::int)
-#366 := (ite #13 #357 #365)
-#852 := (forall (vars (?v0 int) (?v1 int)) (:pat #851) #366)
-#369 := (forall (vars (?v0 int) (?v1 int)) #366)
-#855 := (iff #369 #852)
-#853 := (iff #366 #366)
-#854 := [refl]: #853
-#856 := [quant-intro #854]: #855
-#147 := (* -1::int #141)
-#326 := (ite #308 #30 #147)
-#329 := (ite #12 0::int #326)
-#332 := (ite #13 #8 #329)
-#335 := (= #29 #332)
-#338 := (forall (vars (?v0 int) (?v1 int)) #335)
-#370 := (iff #338 #369)
-#367 := (iff #335 #366)
-#368 := [rewrite]: #367
-#371 := [quant-intro #368]: #370
-#116 := (not #115)
-#109 := (not #108)
-#119 := (and #109 #116)
-#105 := (not #104)
-#112 := (and #105 #109)
-#122 := (or #112 #119)
-#167 := (ite #122 #30 #147)
-#170 := (ite #12 0::int #167)
-#173 := (ite #13 #8 #170)
-#176 := (= #29 #173)
-#179 := (forall (vars (?v0 int) (?v1 int)) #176)
-#339 := (iff #179 #338)
-#336 := (iff #176 #335)
-#333 := (= #173 #332)
-#330 := (= #170 #329)
-#327 := (= #167 #326)
-#311 := (iff #122 #308)
-#305 := (or #302 #285)
-#309 := (iff #305 #308)
-#310 := [rewrite]: #309
-#306 := (iff #122 #305)
-#303 := (iff #119 #285)
-#304 := [rewrite]: #303
-#282 := (iff #112 #302)
-#283 := [rewrite]: #282
-#307 := [monotonicity #283 #304]: #306
-#312 := [trans #307 #310]: #311
-#328 := [monotonicity #312]: #327
-#331 := [monotonicity #328]: #330
-#334 := [monotonicity #331]: #333
-#337 := [monotonicity #334]: #336
-#340 := [quant-intro #337]: #339
-#274 := (~ #179 #179)
-#271 := (~ #176 #176)
-#290 := [refl]: #271
-#275 := [nnf-pos #290]: #274
-#23 := (- #9)
-#22 := (- #8)
-#31 := (mod #22 #23)
-#32 := (- #31)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#33 := (ite #20 #30 #32)
-#34 := (ite #12 0::int #33)
-#35 := (ite #13 #8 #34)
-#36 := (= #29 #35)
-#37 := (forall (vars (?v0 int) (?v1 int)) #36)
-#182 := (iff #37 #179)
-#76 := (and #16 #18)
-#79 := (or #17 #76)
-#152 := (ite #79 #30 #147)
-#155 := (ite #12 0::int #152)
-#158 := (ite #13 #8 #155)
-#161 := (= #29 #158)
-#164 := (forall (vars (?v0 int) (?v1 int)) #161)
-#180 := (iff #164 #179)
-#177 := (iff #161 #176)
-#174 := (= #158 #173)
-#171 := (= #155 #170)
-#168 := (= #152 #167)
-#123 := (iff #79 #122)
-#120 := (iff #76 #119)
-#117 := (iff #18 #116)
-#118 := [rewrite]: #117
-#110 := (iff #16 #109)
-#111 := [rewrite]: #110
-#121 := [monotonicity #111 #118]: #120
-#113 := (iff #17 #112)
-#106 := (iff #15 #105)
-#107 := [rewrite]: #106
-#114 := [monotonicity #107 #111]: #113
-#124 := [monotonicity #114 #121]: #123
-#169 := [monotonicity #124]: #168
-#172 := [monotonicity #169]: #171
-#175 := [monotonicity #172]: #174
-#178 := [monotonicity #175]: #177
-#181 := [quant-intro #178]: #180
-#165 := (iff #37 #164)
-#162 := (iff #36 #161)
-#159 := (= #35 #158)
-#156 := (= #34 #155)
-#153 := (= #33 #152)
-#150 := (= #32 #147)
-#144 := (- #141)
-#148 := (= #144 #147)
-#149 := [rewrite]: #148
-#145 := (= #32 #144)
-#142 := (= #31 #141)
-#87 := (= #23 #86)
-#88 := [rewrite]: #87
-#84 := (= #22 #83)
-#85 := [rewrite]: #84
-#143 := [monotonicity #85 #88]: #142
-#146 := [monotonicity #143]: #145
-#151 := [trans #146 #149]: #150
-#80 := (iff #20 #79)
-#77 := (iff #19 #76)
-#78 := [rewrite]: #77
-#81 := [monotonicity #78]: #80
-#154 := [monotonicity #81 #151]: #153
-#157 := [monotonicity #154]: #156
-#160 := [monotonicity #157]: #159
-#163 := [monotonicity #160]: #162
-#166 := [quant-intro #163]: #165
-#183 := [trans #166 #181]: #182
-#140 := [asserted]: #37
-#184 := [mp #140 #183]: #179
-#291 := [mp~ #184 #275]: #179
-#341 := [mp #291 #340]: #338
-#372 := [mp #341 #371]: #369
-#857 := [mp #372 #856]: #852
-#677 := (not #852)
-#679 := (or #677 #507)
-#430 := (* -1::int 3::int)
-#515 := (* -1::int 0::int)
-#516 := (mod #515 #430)
-#517 := (+ #54 #516)
-#508 := (= #517 0::int)
-#519 := (mod 0::int 3::int)
-#520 := (* -1::int #519)
-#448 := (+ #54 #520)
-#521 := (= #448 0::int)
-#518 := (<= 3::int 0::int)
-#522 := (<= 0::int 0::int)
-#501 := (or #522 #518)
-#836 := (not #501)
-#838 := (>= 0::int 0::int)
-#495 := (or #518 #838)
-#625 := (not #495)
-#832 := (or #625 #836)
-#506 := (ite #832 #521 #508)
-#839 := (= 0::int 0::int)
-#840 := (ite #839 #507 #506)
-#841 := (= 0::int #54)
-#842 := (= 3::int 0::int)
-#837 := (ite #842 #841 #840)
-#680 := (or #677 #837)
-#681 := (iff #680 #679)
-#683 := (iff #679 #679)
-#684 := [rewrite]: #683
-#777 := (iff #837 #507)
-#626 := (ite false #507 #507)
-#781 := (iff #626 #507)
-#783 := [rewrite]: #781
-#773 := (iff #837 #626)
-#776 := (iff #840 #507)
-#1 := true
-#785 := (ite true #507 #507)
-#669 := (iff #785 #507)
-#628 := [rewrite]: #669
-#667 := (iff #840 #785)
-#784 := (iff #506 #507)
-#627 := (iff #506 #626)
-#795 := (iff #508 #507)
-#794 := (= #517 #54)
-#806 := (+ #54 0::int)
-#800 := (= #806 #54)
-#801 := [rewrite]: #800
-#787 := (= #517 #806)
-#792 := (= #516 0::int)
-#798 := -3::int
-#646 := (mod 0::int -3::int)
-#790 := (= #646 0::int)
-#791 := [rewrite]: #790
-#647 := (= #516 #646)
-#804 := (= #430 -3::int)
-#645 := [rewrite]: #804
-#808 := (= #515 0::int)
-#523 := [rewrite]: #808
-#789 := [monotonicity #523 #645]: #647
-#786 := [trans #789 #791]: #792
-#793 := [monotonicity #786]: #787
-#788 := [trans #793 #801]: #794
-#624 := [monotonicity #788]: #795
-#802 := (iff #521 #507)
-#796 := (= #448 #54)
-#810 := (= #448 #806)
-#805 := (= #520 0::int)
-#529 := (= #520 #515)
-#527 := (= #519 0::int)
-#528 := [rewrite]: #527
-#807 := [monotonicity #528]: #529
-#809 := [trans #807 #523]: #805
-#799 := [monotonicity #809]: #810
-#797 := [trans #799 #801]: #796
-#803 := [monotonicity #797]: #802
-#816 := (iff #832 false)
-#536 := (or false false)
-#815 := (iff #536 false)
-#812 := [rewrite]: #815
-#537 := (iff #832 #536)
-#813 := (iff #836 false)
-#825 := (not true)
-#464 := (iff #825 false)
-#827 := [rewrite]: #464
-#543 := (iff #836 #825)
-#541 := (iff #501 true)
-#828 := (or true false)
-#818 := (iff #828 true)
-#540 := [rewrite]: #818
-#831 := (iff #501 #828)
-#835 := (iff #518 false)
-#478 := [rewrite]: #835
-#470 := (iff #522 true)
-#830 := [rewrite]: #470
-#817 := [monotonicity #830 #478]: #831
-#542 := [trans #817 #540]: #541
-#436 := [monotonicity #542]: #543
-#814 := [trans #436 #827]: #813
-#829 := (iff #625 false)
-#820 := (iff #625 #825)
-#823 := (iff #495 true)
-#483 := (or false true)
-#821 := (iff #483 true)
-#822 := [rewrite]: #821
-#484 := (iff #495 #483)
-#819 := (iff #838 true)
-#482 := [rewrite]: #819
-#485 := [monotonicity #478 #482]: #484
-#824 := [trans #485 #822]: #823
-#826 := [monotonicity #824]: #820
-#469 := [trans #826 #827]: #829
-#811 := [monotonicity #469 #814]: #537
-#526 := [trans #811 #812]: #816
-#780 := [monotonicity #526 #803 #624]: #627
-#782 := [trans #780 #783]: #784
-#834 := (iff #839 true)
-#494 := [rewrite]: #834
-#668 := [monotonicity #494 #782]: #667
-#678 := [trans #668 #628]: #776
-#493 := (iff #841 #507)
-#498 := [rewrite]: #493
-#843 := (iff #842 false)
-#833 := [rewrite]: #843
-#774 := [monotonicity #833 #498 #678]: #773
-#662 := [trans #774 #783]: #777
-#682 := [monotonicity #662]: #681
-#685 := [trans #682 #684]: #681
-#673 := [quant-inst]: #680
-#778 := [mp #673 #685]: #679
-#713 := [unit-resolution #778 #857]: #507
-#709 := [monotonicity #713]: #57
-#58 := (not #57)
-#270 := [asserted]: #58
-[unit-resolution #270 #709]: false
-unsat
-b9b522deb0a9246b7d18bd07eb8550de7aa969f7 605 0
+b68469389563a7c2adfa58d6c3a299893ac9030d 605 0
 #2 := false
 decl f5 :: (-> int S2)
 #11 := 0::int
@@ -19042,7 +16977,2280 @@
 #603 := [unit-resolution #535 #671]: #61
 [unit-resolution #603 #596]: false
 unsat
-a249032926dbefc0badb62f9f4f23768776070bc 328 0
+548d39bc73378ffff45c17ea9af834838e4f2aed 274 0
+#2 := false
+decl f5 :: (-> int S2)
+#11 := 0::int
+#55 := (f5 0::int)
+decl f4 :: (-> int int int)
+#53 := (f4 0::int 0::int)
+#54 := (f5 #53)
+#56 := (= #54 #55)
+#831 := (= #53 0::int)
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#29 := (f4 #8 #9)
+#850 := (pattern #29)
+#81 := -1::int
+#85 := (* -1::int #9)
+#82 := (* -1::int #8)
+#140 := (mod #82 #85)
+#361 := (+ #29 #140)
+#362 := (= #361 0::int)
+#30 := (mod #8 #9)
+#358 := (* -1::int #30)
+#359 := (+ #29 #358)
+#360 := (= #359 0::int)
+#107 := (<= #9 0::int)
+#103 := (<= #8 0::int)
+#300 := (or #103 #107)
+#301 := (not #300)
+#114 := (>= #8 0::int)
+#283 := (or #107 #114)
+#284 := (not #283)
+#307 := (or #284 #301)
+#363 := (ite #307 #360 #362)
+#357 := (= #29 0::int)
+#12 := (= #8 0::int)
+#364 := (ite #12 #357 #363)
+#356 := (= #8 #29)
+#13 := (= #9 0::int)
+#365 := (ite #13 #356 #364)
+#851 := (forall (vars (?v0 int) (?v1 int)) (:pat #850) #365)
+#368 := (forall (vars (?v0 int) (?v1 int)) #365)
+#854 := (iff #368 #851)
+#852 := (iff #365 #365)
+#853 := [refl]: #852
+#855 := [quant-intro #853]: #854
+#146 := (* -1::int #140)
+#325 := (ite #307 #30 #146)
+#328 := (ite #12 0::int #325)
+#331 := (ite #13 #8 #328)
+#334 := (= #29 #331)
+#337 := (forall (vars (?v0 int) (?v1 int)) #334)
+#369 := (iff #337 #368)
+#366 := (iff #334 #365)
+#367 := [rewrite]: #366
+#370 := [quant-intro #367]: #369
+#115 := (not #114)
+#108 := (not #107)
+#118 := (and #108 #115)
+#104 := (not #103)
+#111 := (and #104 #108)
+#121 := (or #111 #118)
+#166 := (ite #121 #30 #146)
+#169 := (ite #12 0::int #166)
+#172 := (ite #13 #8 #169)
+#175 := (= #29 #172)
+#178 := (forall (vars (?v0 int) (?v1 int)) #175)
+#338 := (iff #178 #337)
+#335 := (iff #175 #334)
+#332 := (= #172 #331)
+#329 := (= #169 #328)
+#326 := (= #166 #325)
+#310 := (iff #121 #307)
+#304 := (or #301 #284)
+#308 := (iff #304 #307)
+#309 := [rewrite]: #308
+#305 := (iff #121 #304)
+#302 := (iff #118 #284)
+#303 := [rewrite]: #302
+#281 := (iff #111 #301)
+#282 := [rewrite]: #281
+#306 := [monotonicity #282 #303]: #305
+#311 := [trans #306 #309]: #310
+#327 := [monotonicity #311]: #326
+#330 := [monotonicity #327]: #329
+#333 := [monotonicity #330]: #332
+#336 := [monotonicity #333]: #335
+#339 := [quant-intro #336]: #338
+#273 := (~ #178 #178)
+#270 := (~ #175 #175)
+#289 := [refl]: #270
+#274 := [nnf-pos #289]: #273
+#23 := (- #9)
+#22 := (- #8)
+#31 := (mod #22 #23)
+#32 := (- #31)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#33 := (ite #20 #30 #32)
+#34 := (ite #12 0::int #33)
+#35 := (ite #13 #8 #34)
+#36 := (= #29 #35)
+#37 := (forall (vars (?v0 int) (?v1 int)) #36)
+#181 := (iff #37 #178)
+#75 := (and #16 #18)
+#78 := (or #17 #75)
+#151 := (ite #78 #30 #146)
+#154 := (ite #12 0::int #151)
+#157 := (ite #13 #8 #154)
+#160 := (= #29 #157)
+#163 := (forall (vars (?v0 int) (?v1 int)) #160)
+#179 := (iff #163 #178)
+#176 := (iff #160 #175)
+#173 := (= #157 #172)
+#170 := (= #154 #169)
+#167 := (= #151 #166)
+#122 := (iff #78 #121)
+#119 := (iff #75 #118)
+#116 := (iff #18 #115)
+#117 := [rewrite]: #116
+#109 := (iff #16 #108)
+#110 := [rewrite]: #109
+#120 := [monotonicity #110 #117]: #119
+#112 := (iff #17 #111)
+#105 := (iff #15 #104)
+#106 := [rewrite]: #105
+#113 := [monotonicity #106 #110]: #112
+#123 := [monotonicity #113 #120]: #122
+#168 := [monotonicity #123]: #167
+#171 := [monotonicity #168]: #170
+#174 := [monotonicity #171]: #173
+#177 := [monotonicity #174]: #176
+#180 := [quant-intro #177]: #179
+#164 := (iff #37 #163)
+#161 := (iff #36 #160)
+#158 := (= #35 #157)
+#155 := (= #34 #154)
+#152 := (= #33 #151)
+#149 := (= #32 #146)
+#143 := (- #140)
+#147 := (= #143 #146)
+#148 := [rewrite]: #147
+#144 := (= #32 #143)
+#141 := (= #31 #140)
+#86 := (= #23 #85)
+#87 := [rewrite]: #86
+#83 := (= #22 #82)
+#84 := [rewrite]: #83
+#142 := [monotonicity #84 #87]: #141
+#145 := [monotonicity #142]: #144
+#150 := [trans #145 #148]: #149
+#79 := (iff #20 #78)
+#76 := (iff #19 #75)
+#77 := [rewrite]: #76
+#80 := [monotonicity #77]: #79
+#153 := [monotonicity #80 #150]: #152
+#156 := [monotonicity #153]: #155
+#159 := [monotonicity #156]: #158
+#162 := [monotonicity #159]: #161
+#165 := [quant-intro #162]: #164
+#182 := [trans #165 #180]: #181
+#139 := [asserted]: #37
+#183 := [mp #139 #182]: #178
+#290 := [mp~ #183 #274]: #178
+#340 := [mp #290 #339]: #337
+#371 := [mp #340 #370]: #368
+#856 := [mp #371 #855]: #851
+#646 := (not #851)
+#788 := (or #646 #831)
+#429 := (* -1::int 0::int)
+#514 := (mod #429 #429)
+#515 := (+ #53 #514)
+#516 := (= #515 0::int)
+#507 := (mod 0::int 0::int)
+#518 := (* -1::int #507)
+#519 := (+ #53 #518)
+#447 := (= #519 0::int)
+#520 := (<= 0::int 0::int)
+#517 := (or #520 #520)
+#521 := (not #517)
+#500 := (>= 0::int 0::int)
+#835 := (or #520 #500)
+#837 := (not #835)
+#494 := (or #837 #521)
+#624 := (ite #494 #447 #516)
+#505 := (= 0::int 0::int)
+#506 := (ite #505 #831 #624)
+#838 := (= 0::int #53)
+#839 := (ite #505 #838 #506)
+#789 := (or #646 #839)
+#791 := (iff #789 #788)
+#786 := (iff #788 #788)
+#792 := [rewrite]: #786
+#644 := (iff #839 #831)
+#1 := true
+#796 := (ite true #831 #831)
+#797 := (iff #796 #831)
+#803 := [rewrite]: #797
+#801 := (iff #839 #796)
+#800 := (iff #506 #831)
+#536 := (+ #53 #507)
+#811 := (= #536 0::int)
+#808 := (ite true #831 #811)
+#798 := (iff #808 #831)
+#799 := [rewrite]: #798
+#805 := (iff #506 #808)
+#522 := (iff #624 #811)
+#526 := (ite false #447 #811)
+#806 := (iff #526 #811)
+#807 := [rewrite]: #806
+#527 := (iff #624 #526)
+#815 := (iff #516 #811)
+#810 := (= #515 #536)
+#813 := (= #514 #507)
+#435 := (= #429 0::int)
+#812 := [rewrite]: #435
+#535 := [monotonicity #812 #812]: #813
+#814 := [monotonicity #535]: #810
+#525 := [monotonicity #814]: #815
+#541 := (iff #494 false)
+#830 := (or false false)
+#539 := (iff #830 false)
+#540 := [rewrite]: #539
+#816 := (iff #494 #830)
+#829 := (iff #521 false)
+#484 := (not true)
+#822 := (iff #484 false)
+#823 := [rewrite]: #822
+#468 := (iff #521 #484)
+#826 := (iff #517 true)
+#493 := (or true true)
+#818 := (iff #493 true)
+#481 := [rewrite]: #818
+#825 := (iff #517 #493)
+#832 := (iff #520 true)
+#492 := [rewrite]: #832
+#463 := [monotonicity #492 #492]: #825
+#828 := [trans #463 #481]: #826
+#469 := [monotonicity #828]: #468
+#827 := [trans #469 #823]: #829
+#824 := (iff #837 false)
+#820 := (iff #837 #484)
+#482 := (iff #835 true)
+#834 := (iff #835 #493)
+#497 := (iff #500 true)
+#833 := [rewrite]: #497
+#477 := [monotonicity #492 #833]: #834
+#483 := [trans #477 #481]: #482
+#821 := [monotonicity #483]: #820
+#819 := [trans #821 #823]: #824
+#817 := [monotonicity #819 #827]: #816
+#542 := [trans #817 #540]: #541
+#528 := [monotonicity #542 #525]: #527
+#804 := [trans #528 #807]: #522
+#840 := (iff #505 true)
+#841 := [rewrite]: #840
+#809 := [monotonicity #841 #804]: #805
+#795 := [trans #809 #799]: #800
+#836 := (iff #838 #831)
+#842 := [rewrite]: #836
+#802 := [monotonicity #841 #842 #795]: #801
+#645 := [trans #802 #803]: #644
+#785 := [monotonicity #645]: #791
+#793 := [trans #785 #792]: #791
+#790 := [quant-inst]: #789
+#787 := [mp #790 #793]: #788
+#741 := [unit-resolution #787 #856]: #831
+#742 := [monotonicity #741]: #56
+#57 := (not #56)
+#269 := [asserted]: #57
+[unit-resolution #269 #742]: false
+unsat
+2a856348a2d0d12d43659d3f8fbd60f2f395959a 321 0
+#2 := false
+decl f5 :: (-> int S2)
+decl f4 :: (-> int int int)
+#11 := 0::int
+decl f6 :: (-> S2 int)
+decl f7 :: S2
+#53 := f7
+#54 := (f6 f7)
+#55 := (f4 #54 0::int)
+#56 := (f5 #55)
+#271 := (= f7 #56)
+#795 := (f5 #54)
+#772 := (= #795 #56)
+#771 := (= #56 #795)
+#768 := (= #55 #54)
+#848 := (= #54 #55)
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#29 := (f4 #8 #9)
+#858 := (pattern #29)
+#82 := -1::int
+#86 := (* -1::int #9)
+#83 := (* -1::int #8)
+#141 := (mod #83 #86)
+#369 := (+ #29 #141)
+#370 := (= #369 0::int)
+#30 := (mod #8 #9)
+#366 := (* -1::int #30)
+#367 := (+ #29 #366)
+#368 := (= #367 0::int)
+#108 := (<= #9 0::int)
+#104 := (<= #8 0::int)
+#308 := (or #104 #108)
+#309 := (not #308)
+#115 := (>= #8 0::int)
+#291 := (or #108 #115)
+#292 := (not #291)
+#315 := (or #292 #309)
+#371 := (ite #315 #368 #370)
+#365 := (= #29 0::int)
+#12 := (= #8 0::int)
+#372 := (ite #12 #365 #371)
+#364 := (= #8 #29)
+#13 := (= #9 0::int)
+#373 := (ite #13 #364 #372)
+#859 := (forall (vars (?v0 int) (?v1 int)) (:pat #858) #373)
+#376 := (forall (vars (?v0 int) (?v1 int)) #373)
+#862 := (iff #376 #859)
+#860 := (iff #373 #373)
+#861 := [refl]: #860
+#863 := [quant-intro #861]: #862
+#147 := (* -1::int #141)
+#333 := (ite #315 #30 #147)
+#336 := (ite #12 0::int #333)
+#339 := (ite #13 #8 #336)
+#342 := (= #29 #339)
+#345 := (forall (vars (?v0 int) (?v1 int)) #342)
+#377 := (iff #345 #376)
+#374 := (iff #342 #373)
+#375 := [rewrite]: #374
+#378 := [quant-intro #375]: #377
+#116 := (not #115)
+#109 := (not #108)
+#119 := (and #109 #116)
+#105 := (not #104)
+#112 := (and #105 #109)
+#122 := (or #112 #119)
+#167 := (ite #122 #30 #147)
+#170 := (ite #12 0::int #167)
+#173 := (ite #13 #8 #170)
+#176 := (= #29 #173)
+#179 := (forall (vars (?v0 int) (?v1 int)) #176)
+#346 := (iff #179 #345)
+#343 := (iff #176 #342)
+#340 := (= #173 #339)
+#337 := (= #170 #336)
+#334 := (= #167 #333)
+#318 := (iff #122 #315)
+#312 := (or #309 #292)
+#316 := (iff #312 #315)
+#317 := [rewrite]: #316
+#313 := (iff #122 #312)
+#310 := (iff #119 #292)
+#311 := [rewrite]: #310
+#289 := (iff #112 #309)
+#290 := [rewrite]: #289
+#314 := [monotonicity #290 #311]: #313
+#319 := [trans #314 #317]: #318
+#335 := [monotonicity #319]: #334
+#338 := [monotonicity #335]: #337
+#341 := [monotonicity #338]: #340
+#344 := [monotonicity #341]: #343
+#347 := [quant-intro #344]: #346
+#281 := (~ #179 #179)
+#277 := (~ #176 #176)
+#297 := [refl]: #277
+#282 := [nnf-pos #297]: #281
+#23 := (- #9)
+#22 := (- #8)
+#31 := (mod #22 #23)
+#32 := (- #31)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#33 := (ite #20 #30 #32)
+#34 := (ite #12 0::int #33)
+#35 := (ite #13 #8 #34)
+#36 := (= #29 #35)
+#37 := (forall (vars (?v0 int) (?v1 int)) #36)
+#182 := (iff #37 #179)
+#76 := (and #16 #18)
+#79 := (or #17 #76)
+#152 := (ite #79 #30 #147)
+#155 := (ite #12 0::int #152)
+#158 := (ite #13 #8 #155)
+#161 := (= #29 #158)
+#164 := (forall (vars (?v0 int) (?v1 int)) #161)
+#180 := (iff #164 #179)
+#177 := (iff #161 #176)
+#174 := (= #158 #173)
+#171 := (= #155 #170)
+#168 := (= #152 #167)
+#123 := (iff #79 #122)
+#120 := (iff #76 #119)
+#117 := (iff #18 #116)
+#118 := [rewrite]: #117
+#110 := (iff #16 #109)
+#111 := [rewrite]: #110
+#121 := [monotonicity #111 #118]: #120
+#113 := (iff #17 #112)
+#106 := (iff #15 #105)
+#107 := [rewrite]: #106
+#114 := [monotonicity #107 #111]: #113
+#124 := [monotonicity #114 #121]: #123
+#169 := [monotonicity #124]: #168
+#172 := [monotonicity #169]: #171
+#175 := [monotonicity #172]: #174
+#178 := [monotonicity #175]: #177
+#181 := [quant-intro #178]: #180
+#165 := (iff #37 #164)
+#162 := (iff #36 #161)
+#159 := (= #35 #158)
+#156 := (= #34 #155)
+#153 := (= #33 #152)
+#150 := (= #32 #147)
+#144 := (- #141)
+#148 := (= #144 #147)
+#149 := [rewrite]: #148
+#145 := (= #32 #144)
+#142 := (= #31 #141)
+#87 := (= #23 #86)
+#88 := [rewrite]: #87
+#84 := (= #22 #83)
+#85 := [rewrite]: #84
+#143 := [monotonicity #85 #88]: #142
+#146 := [monotonicity #143]: #145
+#151 := [trans #146 #149]: #150
+#80 := (iff #20 #79)
+#77 := (iff #19 #76)
+#78 := [rewrite]: #77
+#81 := [monotonicity #78]: #80
+#154 := [monotonicity #81 #151]: #153
+#157 := [monotonicity #154]: #156
+#160 := [monotonicity #157]: #159
+#163 := [monotonicity #160]: #162
+#166 := [quant-intro #163]: #165
+#183 := [trans #166 #181]: #182
+#140 := [asserted]: #37
+#184 := [mp #140 #183]: #179
+#298 := [mp~ #184 #282]: #179
+#348 := [mp #298 #347]: #345
+#379 := [mp #348 #378]: #376
+#864 := [mp #379 #863]: #859
+#653 := (not #859)
+#654 := (or #653 #848)
+#437 := (* -1::int 0::int)
+#522 := (* -1::int #54)
+#523 := (mod #522 #437)
+#524 := (+ #55 #523)
+#515 := (= #524 0::int)
+#526 := (mod #54 0::int)
+#527 := (* -1::int #526)
+#455 := (+ #55 #527)
+#528 := (= #455 0::int)
+#525 := (<= 0::int 0::int)
+#529 := (<= #54 0::int)
+#508 := (or #529 #525)
+#843 := (not #508)
+#845 := (>= #54 0::int)
+#502 := (or #525 #845)
+#632 := (not #502)
+#839 := (or #632 #843)
+#513 := (ite #839 #528 #515)
+#514 := (= #55 0::int)
+#846 := (= #54 0::int)
+#847 := (ite #846 #514 #513)
+#849 := (= 0::int 0::int)
+#844 := (ite #849 #848 #847)
+#796 := (or #653 #844)
+#798 := (iff #796 #654)
+#793 := (iff #654 #654)
+#794 := [rewrite]: #793
+#811 := (iff #844 #848)
+#544 := (mod #522 0::int)
+#819 := (+ #55 #544)
+#534 := (= #819 0::int)
+#806 := (ite #846 #514 #534)
+#1 := true
+#803 := (ite true #848 #806)
+#810 := (iff #803 #848)
+#805 := [rewrite]: #810
+#804 := (iff #844 #803)
+#807 := (iff #847 #806)
+#813 := (iff #513 #534)
+#814 := (ite false #528 #534)
+#812 := (iff #814 #534)
+#816 := [rewrite]: #812
+#815 := (iff #513 #814)
+#535 := (iff #515 #534)
+#823 := (= #524 #819)
+#818 := (= #523 #544)
+#821 := (= #437 0::int)
+#543 := [rewrite]: #821
+#822 := [monotonicity #543]: #818
+#533 := [monotonicity #822]: #823
+#536 := [monotonicity #533]: #535
+#443 := (iff #839 false)
+#825 := (or false false)
+#549 := (iff #825 false)
+#550 := [rewrite]: #549
+#547 := (iff #839 #825)
+#838 := (iff #843 false)
+#491 := (not true)
+#829 := (iff #491 false)
+#830 := [rewrite]: #829
+#837 := (iff #843 #491)
+#476 := (iff #508 true)
+#827 := (or #529 true)
+#834 := (iff #827 true)
+#836 := [rewrite]: #834
+#833 := (iff #508 #827)
+#500 := (iff #525 true)
+#505 := [rewrite]: #500
+#471 := [monotonicity #505]: #833
+#477 := [trans #471 #836]: #476
+#835 := [monotonicity #477]: #837
+#824 := [trans #835 #830]: #838
+#831 := (iff #632 false)
+#492 := (iff #632 #491)
+#489 := (iff #502 true)
+#841 := (or true #845)
+#485 := (iff #841 true)
+#826 := [rewrite]: #485
+#501 := (iff #502 #841)
+#842 := [monotonicity #505]: #501
+#490 := [trans #842 #826]: #489
+#828 := [monotonicity #490]: #492
+#832 := [trans #828 #830]: #831
+#548 := [monotonicity #832 #824]: #547
+#820 := [trans #548 #550]: #443
+#530 := [monotonicity #820 #536]: #815
+#817 := [trans #530 #816]: #813
+#808 := [monotonicity #817]: #807
+#850 := (iff #849 true)
+#840 := [rewrite]: #850
+#809 := [monotonicity #840 #808]: #804
+#652 := [trans #809 #805]: #811
+#799 := [monotonicity #652]: #798
+#800 := [trans #799 #794]: #798
+#797 := [quant-inst]: #796
+#801 := [mp #797 #800]: #654
+#779 := [unit-resolution #801 #864]: #848
+#769 := [symm #779]: #768
+#765 := [monotonicity #769]: #771
+#756 := [symm #765]: #772
+#802 := (= f7 #795)
+#38 := (:var 0 S2)
+#39 := (f6 #38)
+#865 := (pattern #39)
+#40 := (f5 #39)
+#186 := (= #38 #40)
+#866 := (forall (vars (?v0 S2)) (:pat #865) #186)
+#189 := (forall (vars (?v0 S2)) #186)
+#867 := (iff #189 #866)
+#869 := (iff #866 #866)
+#870 := [rewrite]: #869
+#868 := [rewrite]: #867
+#871 := [trans #868 #870]: #867
+#283 := (~ #189 #189)
+#299 := (~ #186 #186)
+#300 := [refl]: #299
+#284 := [nnf-pos #300]: #283
+#41 := (= #40 #38)
+#42 := (forall (vars (?v0 S2)) #41)
+#190 := (iff #42 #189)
+#187 := (iff #41 #186)
+#188 := [rewrite]: #187
+#191 := [quant-intro #188]: #190
+#185 := [asserted]: #42
+#194 := [mp #185 #191]: #189
+#301 := [mp~ #194 #284]: #189
+#872 := [mp #301 #871]: #866
+#634 := (not #866)
+#787 := (or #634 #802)
+#788 := [quant-inst]: #787
+#770 := [unit-resolution #788 #872]: #802
+#757 := [trans #770 #756]: #271
+#274 := (not #271)
+#57 := (= #56 f7)
+#58 := (not #57)
+#275 := (iff #58 #274)
+#272 := (iff #57 #271)
+#273 := [rewrite]: #272
+#276 := [monotonicity #273]: #275
+#270 := [asserted]: #58
+#279 := [mp #270 #276]: #274
+[unit-resolution #279 #757]: false
+unsat
+42cd145e516d85b450559c9cc1b54c1ed7265de3 304 0
+#2 := false
+decl f5 :: (-> int S2)
+#11 := 0::int
+#56 := (f5 0::int)
+decl f4 :: (-> int int int)
+#53 := 1::int
+#54 := (f4 0::int 1::int)
+#55 := (f5 #54)
+#57 := (= #55 #56)
+#838 := (= #54 0::int)
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#29 := (f4 #8 #9)
+#850 := (pattern #29)
+#82 := -1::int
+#86 := (* -1::int #9)
+#83 := (* -1::int #8)
+#141 := (mod #83 #86)
+#362 := (+ #29 #141)
+#363 := (= #362 0::int)
+#30 := (mod #8 #9)
+#359 := (* -1::int #30)
+#360 := (+ #29 #359)
+#361 := (= #360 0::int)
+#108 := (<= #9 0::int)
+#104 := (<= #8 0::int)
+#301 := (or #104 #108)
+#302 := (not #301)
+#115 := (>= #8 0::int)
+#284 := (or #108 #115)
+#285 := (not #284)
+#308 := (or #285 #302)
+#364 := (ite #308 #361 #363)
+#358 := (= #29 0::int)
+#12 := (= #8 0::int)
+#365 := (ite #12 #358 #364)
+#357 := (= #8 #29)
+#13 := (= #9 0::int)
+#366 := (ite #13 #357 #365)
+#851 := (forall (vars (?v0 int) (?v1 int)) (:pat #850) #366)
+#369 := (forall (vars (?v0 int) (?v1 int)) #366)
+#854 := (iff #369 #851)
+#852 := (iff #366 #366)
+#853 := [refl]: #852
+#855 := [quant-intro #853]: #854
+#147 := (* -1::int #141)
+#326 := (ite #308 #30 #147)
+#329 := (ite #12 0::int #326)
+#332 := (ite #13 #8 #329)
+#335 := (= #29 #332)
+#338 := (forall (vars (?v0 int) (?v1 int)) #335)
+#370 := (iff #338 #369)
+#367 := (iff #335 #366)
+#368 := [rewrite]: #367
+#371 := [quant-intro #368]: #370
+#116 := (not #115)
+#109 := (not #108)
+#119 := (and #109 #116)
+#105 := (not #104)
+#112 := (and #105 #109)
+#122 := (or #112 #119)
+#167 := (ite #122 #30 #147)
+#170 := (ite #12 0::int #167)
+#173 := (ite #13 #8 #170)
+#176 := (= #29 #173)
+#179 := (forall (vars (?v0 int) (?v1 int)) #176)
+#339 := (iff #179 #338)
+#336 := (iff #176 #335)
+#333 := (= #173 #332)
+#330 := (= #170 #329)
+#327 := (= #167 #326)
+#311 := (iff #122 #308)
+#305 := (or #302 #285)
+#309 := (iff #305 #308)
+#310 := [rewrite]: #309
+#306 := (iff #122 #305)
+#303 := (iff #119 #285)
+#304 := [rewrite]: #303
+#282 := (iff #112 #302)
+#283 := [rewrite]: #282
+#307 := [monotonicity #283 #304]: #306
+#312 := [trans #307 #310]: #311
+#328 := [monotonicity #312]: #327
+#331 := [monotonicity #328]: #330
+#334 := [monotonicity #331]: #333
+#337 := [monotonicity #334]: #336
+#340 := [quant-intro #337]: #339
+#274 := (~ #179 #179)
+#271 := (~ #176 #176)
+#290 := [refl]: #271
+#275 := [nnf-pos #290]: #274
+#23 := (- #9)
+#22 := (- #8)
+#31 := (mod #22 #23)
+#32 := (- #31)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#33 := (ite #20 #30 #32)
+#34 := (ite #12 0::int #33)
+#35 := (ite #13 #8 #34)
+#36 := (= #29 #35)
+#37 := (forall (vars (?v0 int) (?v1 int)) #36)
+#182 := (iff #37 #179)
+#76 := (and #16 #18)
+#79 := (or #17 #76)
+#152 := (ite #79 #30 #147)
+#155 := (ite #12 0::int #152)
+#158 := (ite #13 #8 #155)
+#161 := (= #29 #158)
+#164 := (forall (vars (?v0 int) (?v1 int)) #161)
+#180 := (iff #164 #179)
+#177 := (iff #161 #176)
+#174 := (= #158 #173)
+#171 := (= #155 #170)
+#168 := (= #152 #167)
+#123 := (iff #79 #122)
+#120 := (iff #76 #119)
+#117 := (iff #18 #116)
+#118 := [rewrite]: #117
+#110 := (iff #16 #109)
+#111 := [rewrite]: #110
+#121 := [monotonicity #111 #118]: #120
+#113 := (iff #17 #112)
+#106 := (iff #15 #105)
+#107 := [rewrite]: #106
+#114 := [monotonicity #107 #111]: #113
+#124 := [monotonicity #114 #121]: #123
+#169 := [monotonicity #124]: #168
+#172 := [monotonicity #169]: #171
+#175 := [monotonicity #172]: #174
+#178 := [monotonicity #175]: #177
+#181 := [quant-intro #178]: #180
+#165 := (iff #37 #164)
+#162 := (iff #36 #161)
+#159 := (= #35 #158)
+#156 := (= #34 #155)
+#153 := (= #33 #152)
+#150 := (= #32 #147)
+#144 := (- #141)
+#148 := (= #144 #147)
+#149 := [rewrite]: #148
+#145 := (= #32 #144)
+#142 := (= #31 #141)
+#87 := (= #23 #86)
+#88 := [rewrite]: #87
+#84 := (= #22 #83)
+#85 := [rewrite]: #84
+#143 := [monotonicity #85 #88]: #142
+#146 := [monotonicity #143]: #145
+#151 := [trans #146 #149]: #150
+#80 := (iff #20 #79)
+#77 := (iff #19 #76)
+#78 := [rewrite]: #77
+#81 := [monotonicity #78]: #80
+#154 := [monotonicity #81 #151]: #153
+#157 := [monotonicity #154]: #156
+#160 := [monotonicity #157]: #159
+#163 := [monotonicity #160]: #162
+#166 := [quant-intro #163]: #165
+#183 := [trans #166 #181]: #182
+#140 := [asserted]: #37
+#184 := [mp #140 #183]: #179
+#291 := [mp~ #184 #275]: #179
+#341 := [mp #291 #340]: #338
+#372 := [mp #341 #371]: #369
+#856 := [mp #372 #855]: #851
+#676 := (not #851)
+#678 := (or #676 #838)
+#430 := (* -1::int 1::int)
+#514 := (* -1::int 0::int)
+#515 := (mod #514 #430)
+#516 := (+ #54 #515)
+#507 := (= #516 0::int)
+#518 := (mod 0::int 1::int)
+#519 := (* -1::int #518)
+#520 := (+ #54 #519)
+#517 := (= #520 0::int)
+#521 := (<= 1::int 0::int)
+#500 := (<= 0::int 0::int)
+#835 := (or #500 #521)
+#837 := (not #835)
+#494 := (>= 0::int 0::int)
+#624 := (or #521 #494)
+#831 := (not #624)
+#505 := (or #831 #837)
+#506 := (ite #505 #517 #507)
+#839 := (= 0::int 0::int)
+#840 := (ite #839 #838 #506)
+#841 := (= 0::int #54)
+#836 := (= 1::int 0::int)
+#842 := (ite #836 #841 #840)
+#679 := (or #676 #842)
+#680 := (iff #679 #678)
+#682 := (iff #678 #678)
+#683 := [rewrite]: #682
+#776 := (iff #842 #838)
+#625 := (ite false #838 #838)
+#780 := (iff #625 #838)
+#782 := [rewrite]: #780
+#772 := (iff #842 #625)
+#775 := (iff #840 #838)
+#1 := true
+#784 := (ite true #838 #838)
+#668 := (iff #784 #838)
+#627 := [rewrite]: #668
+#666 := (iff #840 #784)
+#783 := (iff #506 #838)
+#626 := (iff #506 #625)
+#794 := (iff #507 #838)
+#793 := (= #516 #54)
+#809 := (+ #54 0::int)
+#800 := (= #809 #54)
+#795 := [rewrite]: #800
+#786 := (= #516 #809)
+#791 := (= #515 0::int)
+#645 := (mod 0::int -1::int)
+#789 := (= #645 0::int)
+#790 := [rewrite]: #789
+#646 := (= #515 #645)
+#803 := (= #430 -1::int)
+#644 := [rewrite]: #803
+#522 := (= #514 0::int)
+#804 := [rewrite]: #522
+#788 := [monotonicity #804 #644]: #646
+#785 := [trans #788 #790]: #791
+#792 := [monotonicity #785]: #786
+#787 := [trans #792 #795]: #793
+#623 := [monotonicity #787]: #794
+#802 := (iff #517 #838)
+#796 := (= #520 #54)
+#798 := (= #520 #809)
+#808 := (= #519 0::int)
+#806 := (= #519 #514)
+#527 := (= #518 0::int)
+#528 := [rewrite]: #527
+#807 := [monotonicity #528]: #806
+#805 := [trans #807 #804]: #808
+#799 := [monotonicity #805]: #798
+#801 := [trans #799 #795]: #796
+#797 := [monotonicity #801]: #802
+#525 := (iff #505 false)
+#536 := (or false false)
+#811 := (iff #536 false)
+#815 := [rewrite]: #811
+#810 := (iff #505 #536)
+#813 := (iff #837 false)
+#819 := (not true)
+#826 := (iff #819 false)
+#828 := [rewrite]: #826
+#436 := (iff #837 #819)
+#541 := (iff #835 true)
+#830 := (or true false)
+#539 := (iff #830 true)
+#540 := [rewrite]: #539
+#816 := (iff #835 #830)
+#477 := (iff #521 false)
+#818 := [rewrite]: #477
+#829 := (iff #500 true)
+#827 := [rewrite]: #829
+#817 := [monotonicity #827 #818]: #816
+#542 := [trans #817 #540]: #541
+#812 := [monotonicity #542]: #436
+#535 := [trans #812 #828]: #813
+#468 := (iff #831 false)
+#825 := (iff #831 #819)
+#823 := (iff #624 true)
+#483 := (or false true)
+#821 := (iff #483 true)
+#822 := [rewrite]: #821
+#484 := (iff #624 #483)
+#481 := (iff #494 true)
+#482 := [rewrite]: #481
+#820 := [monotonicity #818 #482]: #484
+#824 := [trans #820 #822]: #823
+#463 := [monotonicity #824]: #825
+#469 := [trans #463 #828]: #468
+#814 := [monotonicity #469 #535]: #810
+#526 := [trans #814 #815]: #525
+#779 := [monotonicity #526 #797 #623]: #626
+#781 := [trans #779 #782]: #783
+#493 := (iff #839 true)
+#834 := [rewrite]: #493
+#667 := [monotonicity #834 #781]: #666
+#677 := [trans #667 #627]: #775
+#497 := (iff #841 #838)
+#833 := [rewrite]: #497
+#832 := (iff #836 false)
+#492 := [rewrite]: #832
+#773 := [monotonicity #492 #833 #677]: #772
+#661 := [trans #773 #782]: #776
+#681 := [monotonicity #661]: #680
+#684 := [trans #681 #683]: #680
+#672 := [quant-inst]: #679
+#777 := [mp #672 #684]: #678
+#712 := [unit-resolution #777 #856]: #838
+#708 := [monotonicity #712]: #57
+#58 := (not #57)
+#270 := [asserted]: #58
+[unit-resolution #270 #708]: false
+unsat
+121357a49f554c92c48529c5e8845192f6c0a8a5 302 0
+#2 := false
+decl f5 :: (-> int S2)
+#11 := 0::int
+#56 := (f5 0::int)
+decl f4 :: (-> int int int)
+#53 := 1::int
+#54 := (f4 1::int 1::int)
+#55 := (f5 #54)
+#57 := (= #55 #56)
+#505 := (= #54 0::int)
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#29 := (f4 #8 #9)
+#850 := (pattern #29)
+#82 := -1::int
+#86 := (* -1::int #9)
+#83 := (* -1::int #8)
+#141 := (mod #83 #86)
+#362 := (+ #29 #141)
+#363 := (= #362 0::int)
+#30 := (mod #8 #9)
+#359 := (* -1::int #30)
+#360 := (+ #29 #359)
+#361 := (= #360 0::int)
+#108 := (<= #9 0::int)
+#104 := (<= #8 0::int)
+#301 := (or #104 #108)
+#302 := (not #301)
+#115 := (>= #8 0::int)
+#284 := (or #108 #115)
+#285 := (not #284)
+#308 := (or #285 #302)
+#364 := (ite #308 #361 #363)
+#358 := (= #29 0::int)
+#12 := (= #8 0::int)
+#365 := (ite #12 #358 #364)
+#357 := (= #8 #29)
+#13 := (= #9 0::int)
+#366 := (ite #13 #357 #365)
+#851 := (forall (vars (?v0 int) (?v1 int)) (:pat #850) #366)
+#369 := (forall (vars (?v0 int) (?v1 int)) #366)
+#854 := (iff #369 #851)
+#852 := (iff #366 #366)
+#853 := [refl]: #852
+#855 := [quant-intro #853]: #854
+#147 := (* -1::int #141)
+#326 := (ite #308 #30 #147)
+#329 := (ite #12 0::int #326)
+#332 := (ite #13 #8 #329)
+#335 := (= #29 #332)
+#338 := (forall (vars (?v0 int) (?v1 int)) #335)
+#370 := (iff #338 #369)
+#367 := (iff #335 #366)
+#368 := [rewrite]: #367
+#371 := [quant-intro #368]: #370
+#116 := (not #115)
+#109 := (not #108)
+#119 := (and #109 #116)
+#105 := (not #104)
+#112 := (and #105 #109)
+#122 := (or #112 #119)
+#167 := (ite #122 #30 #147)
+#170 := (ite #12 0::int #167)
+#173 := (ite #13 #8 #170)
+#176 := (= #29 #173)
+#179 := (forall (vars (?v0 int) (?v1 int)) #176)
+#339 := (iff #179 #338)
+#336 := (iff #176 #335)
+#333 := (= #173 #332)
+#330 := (= #170 #329)
+#327 := (= #167 #326)
+#311 := (iff #122 #308)
+#305 := (or #302 #285)
+#309 := (iff #305 #308)
+#310 := [rewrite]: #309
+#306 := (iff #122 #305)
+#303 := (iff #119 #285)
+#304 := [rewrite]: #303
+#282 := (iff #112 #302)
+#283 := [rewrite]: #282
+#307 := [monotonicity #283 #304]: #306
+#312 := [trans #307 #310]: #311
+#328 := [monotonicity #312]: #327
+#331 := [monotonicity #328]: #330
+#334 := [monotonicity #331]: #333
+#337 := [monotonicity #334]: #336
+#340 := [quant-intro #337]: #339
+#274 := (~ #179 #179)
+#271 := (~ #176 #176)
+#290 := [refl]: #271
+#275 := [nnf-pos #290]: #274
+#23 := (- #9)
+#22 := (- #8)
+#31 := (mod #22 #23)
+#32 := (- #31)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#33 := (ite #20 #30 #32)
+#34 := (ite #12 0::int #33)
+#35 := (ite #13 #8 #34)
+#36 := (= #29 #35)
+#37 := (forall (vars (?v0 int) (?v1 int)) #36)
+#182 := (iff #37 #179)
+#76 := (and #16 #18)
+#79 := (or #17 #76)
+#152 := (ite #79 #30 #147)
+#155 := (ite #12 0::int #152)
+#158 := (ite #13 #8 #155)
+#161 := (= #29 #158)
+#164 := (forall (vars (?v0 int) (?v1 int)) #161)
+#180 := (iff #164 #179)
+#177 := (iff #161 #176)
+#174 := (= #158 #173)
+#171 := (= #155 #170)
+#168 := (= #152 #167)
+#123 := (iff #79 #122)
+#120 := (iff #76 #119)
+#117 := (iff #18 #116)
+#118 := [rewrite]: #117
+#110 := (iff #16 #109)
+#111 := [rewrite]: #110
+#121 := [monotonicity #111 #118]: #120
+#113 := (iff #17 #112)
+#106 := (iff #15 #105)
+#107 := [rewrite]: #106
+#114 := [monotonicity #107 #111]: #113
+#124 := [monotonicity #114 #121]: #123
+#169 := [monotonicity #124]: #168
+#172 := [monotonicity #169]: #171
+#175 := [monotonicity #172]: #174
+#178 := [monotonicity #175]: #177
+#181 := [quant-intro #178]: #180
+#165 := (iff #37 #164)
+#162 := (iff #36 #161)
+#159 := (= #35 #158)
+#156 := (= #34 #155)
+#153 := (= #33 #152)
+#150 := (= #32 #147)
+#144 := (- #141)
+#148 := (= #144 #147)
+#149 := [rewrite]: #148
+#145 := (= #32 #144)
+#142 := (= #31 #141)
+#87 := (= #23 #86)
+#88 := [rewrite]: #87
+#84 := (= #22 #83)
+#85 := [rewrite]: #84
+#143 := [monotonicity #85 #88]: #142
+#146 := [monotonicity #143]: #145
+#151 := [trans #146 #149]: #150
+#80 := (iff #20 #79)
+#77 := (iff #19 #76)
+#78 := [rewrite]: #77
+#81 := [monotonicity #78]: #80
+#154 := [monotonicity #81 #151]: #153
+#157 := [monotonicity #154]: #156
+#160 := [monotonicity #157]: #159
+#163 := [monotonicity #160]: #162
+#166 := [quant-intro #163]: #165
+#183 := [trans #166 #181]: #182
+#140 := [asserted]: #37
+#184 := [mp #140 #183]: #179
+#291 := [mp~ #184 #275]: #179
+#341 := [mp #291 #340]: #338
+#372 := [mp #341 #371]: #369
+#856 := [mp #372 #855]: #851
+#776 := (not #851)
+#661 := (or #776 #505)
+#430 := (* -1::int 1::int)
+#514 := (mod #430 #430)
+#515 := (+ #54 #514)
+#516 := (= #515 0::int)
+#507 := (mod 1::int 1::int)
+#518 := (* -1::int #507)
+#519 := (+ #54 #518)
+#520 := (= #519 0::int)
+#517 := (<= 1::int 0::int)
+#521 := (or #517 #517)
+#500 := (not #521)
+#835 := (>= 1::int 0::int)
+#837 := (or #517 #835)
+#494 := (not #837)
+#624 := (or #494 #500)
+#831 := (ite #624 #520 #516)
+#506 := (= 1::int 0::int)
+#838 := (ite #506 #505 #831)
+#839 := (= 1::int #54)
+#840 := (ite #506 #839 #838)
+#676 := (or #776 #840)
+#679 := (iff #676 #661)
+#680 := (iff #661 #661)
+#681 := [rewrite]: #680
+#772 := (iff #840 #505)
+#832 := (= #54 1::int)
+#667 := (ite false #832 #505)
+#775 := (iff #667 #505)
+#677 := [rewrite]: #775
+#668 := (iff #840 #667)
+#784 := (iff #838 #505)
+#779 := (ite false #505 #505)
+#783 := (iff #779 #505)
+#781 := [rewrite]: #783
+#780 := (iff #838 #779)
+#625 := (iff #831 #505)
+#1 := true
+#792 := (ite true #505 #505)
+#794 := (iff #792 #505)
+#623 := [rewrite]: #794
+#793 := (iff #831 #792)
+#785 := (iff #516 #505)
+#790 := (= #515 #54)
+#807 := (+ #54 0::int)
+#808 := (= #807 #54)
+#805 := [rewrite]: #808
+#788 := (= #515 #807)
+#645 := (= #514 0::int)
+#801 := (mod -1::int -1::int)
+#803 := (= #801 0::int)
+#644 := [rewrite]: #803
+#802 := (= #514 #801)
+#795 := (= #430 -1::int)
+#796 := [rewrite]: #795
+#797 := [monotonicity #796 #796]: #802
+#646 := [trans #797 #644]: #645
+#789 := [monotonicity #646]: #788
+#791 := [trans #789 #805]: #790
+#786 := [monotonicity #791]: #785
+#799 := (iff #520 #505)
+#809 := (= #519 #54)
+#522 := (= #519 #807)
+#528 := (= #518 0::int)
+#811 := (* -1::int 0::int)
+#526 := (= #811 0::int)
+#527 := [rewrite]: #526
+#815 := (= #518 #811)
+#810 := (= #507 0::int)
+#814 := [rewrite]: #810
+#525 := [monotonicity #814]: #815
+#806 := [trans #525 #527]: #528
+#804 := [monotonicity #806]: #522
+#798 := [trans #804 #805]: #809
+#800 := [monotonicity #798]: #799
+#535 := (iff #624 true)
+#477 := (or false true)
+#482 := (iff #477 true)
+#483 := [rewrite]: #482
+#812 := (iff #624 #477)
+#542 := (iff #500 true)
+#816 := (not false)
+#540 := (iff #816 true)
+#541 := [rewrite]: #540
+#817 := (iff #500 #816)
+#827 := (iff #521 false)
+#826 := (or false false)
+#469 := (iff #826 false)
+#829 := [rewrite]: #469
+#828 := (iff #521 #826)
+#497 := (iff #517 false)
+#833 := [rewrite]: #497
+#468 := [monotonicity #833 #833]: #828
+#830 := [trans #468 #829]: #827
+#539 := [monotonicity #830]: #817
+#436 := [trans #539 #541]: #542
+#825 := (iff #494 false)
+#821 := (not true)
+#824 := (iff #821 false)
+#819 := [rewrite]: #824
+#822 := (iff #494 #821)
+#484 := (iff #837 true)
+#818 := (iff #837 #477)
+#493 := (iff #835 true)
+#834 := [rewrite]: #493
+#481 := [monotonicity #833 #834]: #818
+#820 := [trans #481 #483]: #484
+#823 := [monotonicity #820]: #822
+#463 := [trans #823 #819]: #825
+#813 := [monotonicity #463 #436]: #812
+#536 := [trans #813 #483]: #535
+#787 := [monotonicity #536 #800 #786]: #793
+#626 := [trans #787 #623]: #625
+#841 := (iff #506 false)
+#836 := [rewrite]: #841
+#782 := [monotonicity #836 #626]: #780
+#666 := [trans #782 #781]: #784
+#842 := (iff #839 #832)
+#492 := [rewrite]: #842
+#627 := [monotonicity #836 #492 #666]: #668
+#773 := [trans #627 #677]: #772
+#672 := [monotonicity #773]: #679
+#682 := [trans #672 #681]: #679
+#678 := [quant-inst]: #676
+#683 := [mp #678 #682]: #661
+#703 := [unit-resolution #683 #856]: #505
+#699 := [monotonicity #703]: #57
+#58 := (not #57)
+#270 := [asserted]: #58
+[unit-resolution #270 #699]: false
+unsat
+2c035401970833955895c9fd35464fade49dcbf1 313 0
+#2 := false
+decl f5 :: (-> int S2)
+#11 := 0::int
+#57 := (f5 0::int)
+decl f4 :: (-> int int int)
+#54 := 1::int
+#53 := 3::int
+#55 := (f4 3::int 1::int)
+#56 := (f5 #55)
+#58 := (= #56 #57)
+#839 := (= #55 0::int)
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#29 := (f4 #8 #9)
+#851 := (pattern #29)
+#83 := -1::int
+#87 := (* -1::int #9)
+#84 := (* -1::int #8)
+#142 := (mod #84 #87)
+#363 := (+ #29 #142)
+#364 := (= #363 0::int)
+#30 := (mod #8 #9)
+#360 := (* -1::int #30)
+#361 := (+ #29 #360)
+#362 := (= #361 0::int)
+#109 := (<= #9 0::int)
+#105 := (<= #8 0::int)
+#302 := (or #105 #109)
+#303 := (not #302)
+#116 := (>= #8 0::int)
+#285 := (or #109 #116)
+#286 := (not #285)
+#309 := (or #286 #303)
+#365 := (ite #309 #362 #364)
+#359 := (= #29 0::int)
+#12 := (= #8 0::int)
+#366 := (ite #12 #359 #365)
+#358 := (= #8 #29)
+#13 := (= #9 0::int)
+#367 := (ite #13 #358 #366)
+#852 := (forall (vars (?v0 int) (?v1 int)) (:pat #851) #367)
+#370 := (forall (vars (?v0 int) (?v1 int)) #367)
+#855 := (iff #370 #852)
+#853 := (iff #367 #367)
+#854 := [refl]: #853
+#856 := [quant-intro #854]: #855
+#148 := (* -1::int #142)
+#327 := (ite #309 #30 #148)
+#330 := (ite #12 0::int #327)
+#333 := (ite #13 #8 #330)
+#336 := (= #29 #333)
+#339 := (forall (vars (?v0 int) (?v1 int)) #336)
+#371 := (iff #339 #370)
+#368 := (iff #336 #367)
+#369 := [rewrite]: #368
+#372 := [quant-intro #369]: #371
+#117 := (not #116)
+#110 := (not #109)
+#120 := (and #110 #117)
+#106 := (not #105)
+#113 := (and #106 #110)
+#123 := (or #113 #120)
+#168 := (ite #123 #30 #148)
+#171 := (ite #12 0::int #168)
+#174 := (ite #13 #8 #171)
+#177 := (= #29 #174)
+#180 := (forall (vars (?v0 int) (?v1 int)) #177)
+#340 := (iff #180 #339)
+#337 := (iff #177 #336)
+#334 := (= #174 #333)
+#331 := (= #171 #330)
+#328 := (= #168 #327)
+#312 := (iff #123 #309)
+#306 := (or #303 #286)
+#310 := (iff #306 #309)
+#311 := [rewrite]: #310
+#307 := (iff #123 #306)
+#304 := (iff #120 #286)
+#305 := [rewrite]: #304
+#283 := (iff #113 #303)
+#284 := [rewrite]: #283
+#308 := [monotonicity #284 #305]: #307
+#313 := [trans #308 #311]: #312
+#329 := [monotonicity #313]: #328
+#332 := [monotonicity #329]: #331
+#335 := [monotonicity #332]: #334
+#338 := [monotonicity #335]: #337
+#341 := [quant-intro #338]: #340
+#275 := (~ #180 #180)
+#272 := (~ #177 #177)
+#291 := [refl]: #272
+#276 := [nnf-pos #291]: #275
+#23 := (- #9)
+#22 := (- #8)
+#31 := (mod #22 #23)
+#32 := (- #31)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#33 := (ite #20 #30 #32)
+#34 := (ite #12 0::int #33)
+#35 := (ite #13 #8 #34)
+#36 := (= #29 #35)
+#37 := (forall (vars (?v0 int) (?v1 int)) #36)
+#183 := (iff #37 #180)
+#77 := (and #16 #18)
+#80 := (or #17 #77)
+#153 := (ite #80 #30 #148)
+#156 := (ite #12 0::int #153)
+#159 := (ite #13 #8 #156)
+#162 := (= #29 #159)
+#165 := (forall (vars (?v0 int) (?v1 int)) #162)
+#181 := (iff #165 #180)
+#178 := (iff #162 #177)
+#175 := (= #159 #174)
+#172 := (= #156 #171)
+#169 := (= #153 #168)
+#124 := (iff #80 #123)
+#121 := (iff #77 #120)
+#118 := (iff #18 #117)
+#119 := [rewrite]: #118
+#111 := (iff #16 #110)
+#112 := [rewrite]: #111
+#122 := [monotonicity #112 #119]: #121
+#114 := (iff #17 #113)
+#107 := (iff #15 #106)
+#108 := [rewrite]: #107
+#115 := [monotonicity #108 #112]: #114
+#125 := [monotonicity #115 #122]: #124
+#170 := [monotonicity #125]: #169
+#173 := [monotonicity #170]: #172
+#176 := [monotonicity #173]: #175
+#179 := [monotonicity #176]: #178
+#182 := [quant-intro #179]: #181
+#166 := (iff #37 #165)
+#163 := (iff #36 #162)
+#160 := (= #35 #159)
+#157 := (= #34 #156)
+#154 := (= #33 #153)
+#151 := (= #32 #148)
+#145 := (- #142)
+#149 := (= #145 #148)
+#150 := [rewrite]: #149
+#146 := (= #32 #145)
+#143 := (= #31 #142)
+#88 := (= #23 #87)
+#89 := [rewrite]: #88
+#85 := (= #22 #84)
+#86 := [rewrite]: #85
+#144 := [monotonicity #86 #89]: #143
+#147 := [monotonicity #144]: #146
+#152 := [trans #147 #150]: #151
+#81 := (iff #20 #80)
+#78 := (iff #19 #77)
+#79 := [rewrite]: #78
+#82 := [monotonicity #79]: #81
+#155 := [monotonicity #82 #152]: #154
+#158 := [monotonicity #155]: #157
+#161 := [monotonicity #158]: #160
+#164 := [monotonicity #161]: #163
+#167 := [quant-intro #164]: #166
+#184 := [trans #167 #182]: #183
+#141 := [asserted]: #37
+#185 := [mp #141 #184]: #180
+#292 := [mp~ #185 #276]: #180
+#342 := [mp #292 #341]: #339
+#373 := [mp #342 #372]: #370
+#857 := [mp #373 #856]: #852
+#685 := (not #852)
+#778 := (or #685 #839)
+#431 := (* -1::int 1::int)
+#515 := (* -1::int 3::int)
+#516 := (mod #515 #431)
+#517 := (+ #55 #516)
+#508 := (= #517 0::int)
+#519 := (mod 3::int 1::int)
+#520 := (* -1::int #519)
+#521 := (+ #55 #520)
+#518 := (= #521 0::int)
+#522 := (<= 1::int 0::int)
+#501 := (<= 3::int 0::int)
+#836 := (or #501 #522)
+#838 := (not #836)
+#495 := (>= 3::int 0::int)
+#625 := (or #522 #495)
+#832 := (not #625)
+#506 := (or #832 #838)
+#507 := (ite #506 #518 #508)
+#840 := (= 3::int 0::int)
+#841 := (ite #840 #839 #507)
+#842 := (= 3::int #55)
+#837 := (= 1::int 0::int)
+#843 := (ite #837 #842 #841)
+#775 := (or #685 #843)
+#766 := (iff #775 #778)
+#760 := (iff #778 #778)
+#757 := [rewrite]: #760
+#683 := (iff #843 #839)
+#834 := (= #55 3::int)
+#679 := (ite false #834 #839)
+#681 := (iff #679 #839)
+#682 := [rewrite]: #681
+#680 := (iff #843 #679)
+#662 := (iff #841 #839)
+#776 := (ite false #839 #839)
+#774 := (iff #776 #839)
+#777 := [rewrite]: #774
+#678 := (iff #841 #776)
+#669 := (iff #507 #839)
+#1 := true
+#784 := (ite true #839 #839)
+#667 := (iff #784 #839)
+#668 := [rewrite]: #667
+#782 := (iff #507 #784)
+#781 := (iff #508 #839)
+#627 := (= #517 #55)
+#800 := (+ #55 0::int)
+#797 := (= #800 #55)
+#802 := [rewrite]: #797
+#624 := (= #517 #800)
+#788 := (= #516 0::int)
+#646 := -3::int
+#792 := (mod -3::int -1::int)
+#793 := (= #792 0::int)
+#794 := [rewrite]: #793
+#786 := (= #516 #792)
+#790 := (= #431 -1::int)
+#791 := [rewrite]: #790
+#647 := (= #515 -3::int)
+#789 := [rewrite]: #647
+#787 := [monotonicity #789 #791]: #786
+#795 := [trans #787 #794]: #788
+#626 := [monotonicity #795]: #624
+#780 := [trans #626 #802]: #627
+#783 := [monotonicity #780]: #781
+#804 := (iff #518 #839)
+#803 := (= #521 #55)
+#801 := (= #521 #800)
+#810 := (= #520 0::int)
+#808 := (* -1::int 0::int)
+#809 := (= #808 0::int)
+#806 := [rewrite]: #809
+#523 := (= #520 #808)
+#529 := (= #519 0::int)
+#807 := [rewrite]: #529
+#805 := [monotonicity #807]: #523
+#799 := [trans #805 #806]: #810
+#796 := [monotonicity #799]: #801
+#798 := [trans #796 #802]: #803
+#645 := [monotonicity #798]: #804
+#527 := (iff #506 true)
+#485 := (or false true)
+#823 := (iff #485 true)
+#824 := [rewrite]: #823
+#816 := (iff #506 #485)
+#815 := (iff #838 true)
+#813 := (not false)
+#537 := (iff #813 true)
+#811 := [rewrite]: #537
+#814 := (iff #838 #813)
+#543 := (iff #836 false)
+#817 := (or false false)
+#541 := (iff #817 false)
+#542 := [rewrite]: #541
+#818 := (iff #836 #817)
+#819 := (iff #522 false)
+#482 := [rewrite]: #819
+#828 := (iff #501 false)
+#831 := [rewrite]: #828
+#540 := [monotonicity #831 #482]: #818
+#437 := [trans #540 #542]: #543
+#536 := [monotonicity #437]: #814
+#812 := [trans #536 #811]: #815
+#470 := (iff #832 false)
+#826 := (not true)
+#829 := (iff #826 false)
+#469 := [rewrite]: #829
+#464 := (iff #832 #826)
+#825 := (iff #625 true)
+#821 := (iff #625 #485)
+#483 := (iff #495 true)
+#484 := [rewrite]: #483
+#822 := [monotonicity #482 #484]: #821
+#820 := [trans #822 #824]: #825
+#827 := [monotonicity #820]: #464
+#830 := [trans #827 #469]: #470
+#526 := [monotonicity #830 #812]: #816
+#528 := [trans #526 #824]: #527
+#785 := [monotonicity #528 #645 #783]: #782
+#628 := [trans #785 #668]: #669
+#835 := (iff #840 false)
+#478 := [rewrite]: #835
+#773 := [monotonicity #478 #628]: #678
+#677 := [trans #773 #777]: #662
+#498 := (iff #842 #834)
+#494 := [rewrite]: #498
+#833 := (iff #837 false)
+#493 := [rewrite]: #833
+#673 := [monotonicity #493 #494 #677]: #680
+#684 := [trans #673 #682]: #683
+#768 := [monotonicity #684]: #766
+#759 := [trans #768 #757]: #766
+#779 := [quant-inst]: #775
+#769 := [mp #779 #759]: #778
+#701 := [unit-resolution #769 #857]: #839
+#702 := [monotonicity #701]: #58
+#59 := (not #58)
+#271 := [asserted]: #59
+[unit-resolution #271 #702]: false
+unsat
+80ae46a7e9c57ab25c6d7a036d7583b3c508d2eb 447 0
+#2 := false
+#11 := 0::int
+decl f6 :: (-> S2 int)
+decl f7 :: S2
+#53 := f7
+#54 := (f6 f7)
+#496 := (>= #54 0::int)
+decl f5 :: (-> int S2)
+#762 := (f5 #54)
+#708 := (f6 #762)
+#704 := (= #708 0::int)
+#655 := (not #704)
+#841 := (= #54 0::int)
+#674 := (not #841)
+#656 := (iff #674 #655)
+#653 := (iff #841 #704)
+#651 := (iff #704 #841)
+#649 := (= #708 #54)
+#643 := (= #762 f7)
+#763 := (= f7 #762)
+#38 := (:var 0 S2)
+#39 := (f6 #38)
+#859 := (pattern #39)
+#40 := (f5 #39)
+#188 := (= #38 #40)
+#860 := (forall (vars (?v0 S2)) (:pat #859) #188)
+#191 := (forall (vars (?v0 S2)) #188)
+#861 := (iff #191 #860)
+#863 := (iff #860 #860)
+#864 := [rewrite]: #863
+#862 := [rewrite]: #861
+#865 := [trans #862 #864]: #861
+#278 := (~ #191 #191)
+#294 := (~ #188 #188)
+#295 := [refl]: #294
+#279 := [nnf-pos #295]: #278
+#41 := (= #40 #38)
+#42 := (forall (vars (?v0 S2)) #41)
+#192 := (iff #42 #191)
+#189 := (iff #41 #188)
+#190 := [rewrite]: #189
+#193 := [quant-intro #190]: #192
+#187 := [asserted]: #42
+#196 := [mp #187 #193]: #191
+#296 := [mp~ #196 #279]: #191
+#866 := [mp #296 #865]: #860
+#759 := (not #860)
+#766 := (or #759 #763)
+#750 := [quant-inst]: #766
+#688 := [unit-resolution #750 #866]: #763
+#644 := [symm #688]: #643
+#650 := [monotonicity #644]: #649
+#652 := [monotonicity #650]: #651
+#654 := [symm #652]: #653
+#657 := [monotonicity #654]: #656
+#486 := (not #496)
+#673 := [hypothesis]: #486
+#677 := (or #674 #496)
+#687 := [th-lemma]: #677
+#667 := [unit-resolution #687 #673]: #674
+#658 := [mp #667 #657]: #655
+#689 := (or #496 #704)
+#9 := (:var 0 int)
+#44 := (f5 #9)
+#867 := (pattern #44)
+#211 := (>= #9 0::int)
+#45 := (f6 #44)
+#50 := (= #45 0::int)
+#261 := (or #50 #211)
+#874 := (forall (vars (?v0 int)) (:pat #867) #261)
+#266 := (forall (vars (?v0 int)) #261)
+#877 := (iff #266 #874)
+#875 := (iff #261 #261)
+#876 := [refl]: #875
+#878 := [quant-intro #876]: #877
+#282 := (~ #266 #266)
+#300 := (~ #261 #261)
+#301 := [refl]: #300
+#283 := [nnf-pos #301]: #282
+#49 := (< #9 0::int)
+#51 := (implies #49 #50)
+#52 := (forall (vars (?v0 int)) #51)
+#269 := (iff #52 #266)
+#232 := (= 0::int #45)
+#238 := (not #49)
+#239 := (or #238 #232)
+#244 := (forall (vars (?v0 int)) #239)
+#267 := (iff #244 #266)
+#264 := (iff #239 #261)
+#258 := (or #211 #50)
+#262 := (iff #258 #261)
+#263 := [rewrite]: #262
+#259 := (iff #239 #258)
+#256 := (iff #232 #50)
+#257 := [rewrite]: #256
+#254 := (iff #238 #211)
+#214 := (not #211)
+#249 := (not #214)
+#252 := (iff #249 #211)
+#253 := [rewrite]: #252
+#250 := (iff #238 #249)
+#247 := (iff #49 #214)
+#248 := [rewrite]: #247
+#251 := [monotonicity #248]: #250
+#255 := [trans #251 #253]: #254
+#260 := [monotonicity #255 #257]: #259
+#265 := [trans #260 #263]: #264
+#268 := [quant-intro #265]: #267
+#245 := (iff #52 #244)
+#242 := (iff #51 #239)
+#235 := (implies #49 #232)
+#240 := (iff #235 #239)
+#241 := [rewrite]: #240
+#236 := (iff #51 #235)
+#233 := (iff #50 #232)
+#234 := [rewrite]: #233
+#237 := [monotonicity #234]: #236
+#243 := [trans #237 #241]: #242
+#246 := [quant-intro #243]: #245
+#270 := [trans #246 #268]: #269
+#231 := [asserted]: #52
+#271 := [mp #231 #270]: #266
+#302 := [mp~ #271 #283]: #266
+#879 := [mp #302 #878]: #874
+#729 := (not #874)
+#671 := (or #729 #496 #704)
+#709 := (or #704 #496)
+#695 := (or #729 #709)
+#662 := (iff #695 #671)
+#691 := (or #729 #689)
+#672 := (iff #691 #671)
+#631 := [rewrite]: #672
+#697 := (iff #695 #691)
+#635 := (iff #709 #689)
+#690 := [rewrite]: #635
+#665 := [monotonicity #690]: #697
+#664 := [trans #665 #631]: #662
+#696 := [quant-inst]: #695
+#666 := [mp #696 #664]: #671
+#675 := [unit-resolution #666 #879]: #689
+#676 := [unit-resolution #675 #673]: #704
+#659 := [unit-resolution #676 #658]: false
+#660 := [lemma #659]: #496
+#502 := (<= #54 0::int)
+#830 := (not #502)
+#625 := (or #486 #830 #841)
+#684 := (not #625)
+decl f4 :: (-> int int int)
+#55 := 1::int
+#56 := (f4 #54 1::int)
+#840 := (= #56 0::int)
+#760 := (not #840)
+#58 := (f5 0::int)
+#57 := (f5 #56)
+#59 := (= #57 #58)
+#645 := [hypothesis]: #840
+#661 := [monotonicity #645]: #59
+#60 := (not #59)
+#272 := [asserted]: #60
+#622 := [unit-resolution #272 #661]: false
+#623 := [lemma #622]: #760
+#632 := (or #684 #840)
+#84 := -1::int
+#516 := (* -1::int #54)
+#809 := (mod #516 -1::int)
+#810 := (+ #56 #809)
+#800 := (= #810 0::int)
+#781 := (ite #625 #840 #800)
+#8 := (:var 1 int)
+#29 := (f4 #8 #9)
+#852 := (pattern #29)
+#88 := (* -1::int #9)
+#85 := (* -1::int #8)
+#143 := (mod #85 #88)
+#364 := (+ #29 #143)
+#365 := (= #364 0::int)
+#30 := (mod #8 #9)
+#361 := (* -1::int #30)
+#362 := (+ #29 #361)
+#363 := (= #362 0::int)
+#110 := (<= #9 0::int)
+#106 := (<= #8 0::int)
+#303 := (or #106 #110)
+#304 := (not #303)
+#117 := (>= #8 0::int)
+#286 := (or #110 #117)
+#287 := (not #286)
+#310 := (or #287 #304)
+#366 := (ite #310 #363 #365)
+#360 := (= #29 0::int)
+#12 := (= #8 0::int)
+#367 := (ite #12 #360 #366)
+#359 := (= #8 #29)
+#13 := (= #9 0::int)
+#368 := (ite #13 #359 #367)
+#853 := (forall (vars (?v0 int) (?v1 int)) (:pat #852) #368)
+#371 := (forall (vars (?v0 int) (?v1 int)) #368)
+#856 := (iff #371 #853)
+#854 := (iff #368 #368)
+#855 := [refl]: #854
+#857 := [quant-intro #855]: #856
+#149 := (* -1::int #143)
+#328 := (ite #310 #30 #149)
+#331 := (ite #12 0::int #328)
+#334 := (ite #13 #8 #331)
+#337 := (= #29 #334)
+#340 := (forall (vars (?v0 int) (?v1 int)) #337)
+#372 := (iff #340 #371)
+#369 := (iff #337 #368)
+#370 := [rewrite]: #369
+#373 := [quant-intro #370]: #372
+#118 := (not #117)
+#111 := (not #110)
+#121 := (and #111 #118)
+#107 := (not #106)
+#114 := (and #107 #111)
+#124 := (or #114 #121)
+#169 := (ite #124 #30 #149)
+#172 := (ite #12 0::int #169)
+#175 := (ite #13 #8 #172)
+#178 := (= #29 #175)
+#181 := (forall (vars (?v0 int) (?v1 int)) #178)
+#341 := (iff #181 #340)
+#338 := (iff #178 #337)
+#335 := (= #175 #334)
+#332 := (= #172 #331)
+#329 := (= #169 #328)
+#313 := (iff #124 #310)
+#307 := (or #304 #287)
+#311 := (iff #307 #310)
+#312 := [rewrite]: #311
+#308 := (iff #124 #307)
+#305 := (iff #121 #287)
+#306 := [rewrite]: #305
+#284 := (iff #114 #304)
+#285 := [rewrite]: #284
+#309 := [monotonicity #285 #306]: #308
+#314 := [trans #309 #312]: #313
+#330 := [monotonicity #314]: #329
+#333 := [monotonicity #330]: #332
+#336 := [monotonicity #333]: #335
+#339 := [monotonicity #336]: #338
+#342 := [quant-intro #339]: #341
+#276 := (~ #181 #181)
+#273 := (~ #178 #178)
+#292 := [refl]: #273
+#277 := [nnf-pos #292]: #276
+#23 := (- #9)
+#22 := (- #8)
+#31 := (mod #22 #23)
+#32 := (- #31)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#33 := (ite #20 #30 #32)
+#34 := (ite #12 0::int #33)
+#35 := (ite #13 #8 #34)
+#36 := (= #29 #35)
+#37 := (forall (vars (?v0 int) (?v1 int)) #36)
+#184 := (iff #37 #181)
+#78 := (and #16 #18)
+#81 := (or #17 #78)
+#154 := (ite #81 #30 #149)
+#157 := (ite #12 0::int #154)
+#160 := (ite #13 #8 #157)
+#163 := (= #29 #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 #81 #124)
+#122 := (iff #78 #121)
+#119 := (iff #18 #118)
+#120 := [rewrite]: #119
+#112 := (iff #16 #111)
+#113 := [rewrite]: #112
+#123 := [monotonicity #113 #120]: #122
+#115 := (iff #17 #114)
+#108 := (iff #15 #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 #37 #166)
+#164 := (iff #36 #163)
+#161 := (= #35 #160)
+#158 := (= #34 #157)
+#155 := (= #33 #154)
+#152 := (= #32 #149)
+#146 := (- #143)
+#150 := (= #146 #149)
+#151 := [rewrite]: #150
+#147 := (= #32 #146)
+#144 := (= #31 #143)
+#89 := (= #23 #88)
+#90 := [rewrite]: #89
+#86 := (= #22 #85)
+#87 := [rewrite]: #86
+#145 := [monotonicity #87 #90]: #144
+#148 := [monotonicity #145]: #147
+#153 := [trans #148 #151]: #152
+#82 := (iff #20 #81)
+#79 := (iff #19 #78)
+#80 := [rewrite]: #79
+#83 := [monotonicity #80]: #82
+#156 := [monotonicity #83 #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]: #37
+#186 := [mp #142 #185]: #181
+#293 := [mp~ #186 #277]: #181
+#343 := [mp #293 #342]: #340
+#374 := [mp #343 #373]: #371
+#858 := [mp #374 #857]: #853
+#786 := (not #853)
+#668 := (or #786 #781)
+#432 := (* -1::int 1::int)
+#517 := (mod #516 #432)
+#518 := (+ #56 #517)
+#509 := (= #518 0::int)
+#520 := (mod #54 1::int)
+#521 := (* -1::int #520)
+#522 := (+ #56 #521)
+#519 := (= #522 0::int)
+#523 := (<= 1::int 0::int)
+#837 := (or #502 #523)
+#839 := (not #837)
+#626 := (or #523 #496)
+#833 := (not #626)
+#507 := (or #833 #839)
+#508 := (ite #507 #519 #509)
+#842 := (ite #841 #840 #508)
+#843 := (= #54 #56)
+#838 := (= 1::int 0::int)
+#844 := (ite #838 #843 #842)
+#669 := (or #786 #844)
+#629 := (iff #669 #668)
+#679 := (iff #668 #668)
+#774 := [rewrite]: #679
+#785 := (iff #844 #781)
+#831 := (or #486 #830)
+#646 := (or #831 #841)
+#647 := (ite #646 #840 #800)
+#782 := (iff #647 #781)
+#627 := (iff #646 #625)
+#628 := [rewrite]: #627
+#784 := [monotonicity #628]: #782
+#789 := (iff #844 #647)
+#793 := (ite false #843 #647)
+#794 := (iff #793 #647)
+#795 := [rewrite]: #794
+#787 := (iff #844 #793)
+#791 := (iff #842 #647)
+#797 := (ite #831 #840 #800)
+#804 := (ite #841 #840 #797)
+#648 := (iff #804 #647)
+#790 := [rewrite]: #648
+#799 := (iff #842 #804)
+#798 := (iff #508 #797)
+#801 := (iff #509 #800)
+#807 := (= #518 #810)
+#524 := (= #517 #809)
+#530 := (= #432 -1::int)
+#808 := [rewrite]: #530
+#806 := [monotonicity #808]: #524
+#811 := [monotonicity #806]: #807
+#802 := [monotonicity #811]: #801
+#528 := (iff #519 #840)
+#817 := (= #522 #56)
+#537 := (+ #56 0::int)
+#816 := (= #537 #56)
+#813 := [rewrite]: #816
+#538 := (= #522 #537)
+#814 := (= #521 0::int)
+#541 := (* -1::int 0::int)
+#544 := (= #541 0::int)
+#438 := [rewrite]: #544
+#542 := (= #521 #541)
+#818 := (= #520 0::int)
+#819 := [rewrite]: #818
+#543 := [monotonicity #819]: #542
+#815 := [trans #543 #438]: #814
+#812 := [monotonicity #815]: #538
+#527 := [trans #812 #813]: #817
+#529 := [monotonicity #527]: #528
+#829 := (iff #507 #831)
+#470 := (iff #839 #830)
+#465 := (iff #837 #502)
+#824 := (or #502 false)
+#821 := (iff #824 #502)
+#827 := [rewrite]: #821
+#825 := (iff #837 #824)
+#499 := (iff #523 false)
+#835 := [rewrite]: #499
+#826 := [monotonicity #835]: #825
+#828 := [trans #826 #827]: #465
+#471 := [monotonicity #828]: #470
+#822 := (iff #833 #486)
+#484 := (iff #626 #496)
+#495 := (or false #496)
+#820 := (iff #495 #496)
+#483 := [rewrite]: #820
+#836 := (iff #626 #495)
+#479 := [monotonicity #835]: #836
+#485 := [trans #479 #483]: #484
+#823 := [monotonicity #485]: #822
+#832 := [monotonicity #823 #471]: #829
+#803 := [monotonicity #832 #529 #802]: #798
+#805 := [monotonicity #803]: #799
+#792 := [trans #805 #790]: #791
+#834 := (iff #838 false)
+#494 := [rewrite]: #834
+#788 := [monotonicity #494 #792]: #787
+#796 := [trans #788 #795]: #789
+#783 := [trans #796 #784]: #785
+#777 := [monotonicity #783]: #629
+#775 := [trans #777 #774]: #629
+#670 := [quant-inst]: #669
+#778 := [mp #670 #775]: #668
+#630 := [unit-resolution #778 #858]: #781
+#780 := (not #781)
+#767 := (or #780 #684 #840)
+#769 := [def-axiom]: #767
+#633 := [unit-resolution #769 #630]: #632
+#634 := [unit-resolution #633 #623]: #684
+#680 := (or #625 #502)
+#681 := [def-axiom]: #680
+#636 := [unit-resolution #681 #634]: #502
+#682 := (or #625 #674)
+#683 := [def-axiom]: #682
+#637 := [unit-resolution #683 #634]: #674
+#638 := (or #841 #830 #486)
+#639 := [th-lemma]: #638
+[unit-resolution #639 #637 #636 #660]: false
+unsat
+124e92a80da0b0c9c6ea2fb487ed064cc0c65420 305 0
+#2 := false
+decl f5 :: (-> int S2)
+#11 := 0::int
+#56 := (f5 0::int)
+decl f4 :: (-> int int int)
+#53 := 3::int
+#54 := (f4 0::int 3::int)
+#55 := (f5 #54)
+#57 := (= #55 #56)
+#507 := (= #54 0::int)
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#29 := (f4 #8 #9)
+#851 := (pattern #29)
+#82 := -1::int
+#86 := (* -1::int #9)
+#83 := (* -1::int #8)
+#141 := (mod #83 #86)
+#362 := (+ #29 #141)
+#363 := (= #362 0::int)
+#30 := (mod #8 #9)
+#359 := (* -1::int #30)
+#360 := (+ #29 #359)
+#361 := (= #360 0::int)
+#108 := (<= #9 0::int)
+#104 := (<= #8 0::int)
+#301 := (or #104 #108)
+#302 := (not #301)
+#115 := (>= #8 0::int)
+#284 := (or #108 #115)
+#285 := (not #284)
+#308 := (or #285 #302)
+#364 := (ite #308 #361 #363)
+#358 := (= #29 0::int)
+#12 := (= #8 0::int)
+#365 := (ite #12 #358 #364)
+#357 := (= #8 #29)
+#13 := (= #9 0::int)
+#366 := (ite #13 #357 #365)
+#852 := (forall (vars (?v0 int) (?v1 int)) (:pat #851) #366)
+#369 := (forall (vars (?v0 int) (?v1 int)) #366)
+#855 := (iff #369 #852)
+#853 := (iff #366 #366)
+#854 := [refl]: #853
+#856 := [quant-intro #854]: #855
+#147 := (* -1::int #141)
+#326 := (ite #308 #30 #147)
+#329 := (ite #12 0::int #326)
+#332 := (ite #13 #8 #329)
+#335 := (= #29 #332)
+#338 := (forall (vars (?v0 int) (?v1 int)) #335)
+#370 := (iff #338 #369)
+#367 := (iff #335 #366)
+#368 := [rewrite]: #367
+#371 := [quant-intro #368]: #370
+#116 := (not #115)
+#109 := (not #108)
+#119 := (and #109 #116)
+#105 := (not #104)
+#112 := (and #105 #109)
+#122 := (or #112 #119)
+#167 := (ite #122 #30 #147)
+#170 := (ite #12 0::int #167)
+#173 := (ite #13 #8 #170)
+#176 := (= #29 #173)
+#179 := (forall (vars (?v0 int) (?v1 int)) #176)
+#339 := (iff #179 #338)
+#336 := (iff #176 #335)
+#333 := (= #173 #332)
+#330 := (= #170 #329)
+#327 := (= #167 #326)
+#311 := (iff #122 #308)
+#305 := (or #302 #285)
+#309 := (iff #305 #308)
+#310 := [rewrite]: #309
+#306 := (iff #122 #305)
+#303 := (iff #119 #285)
+#304 := [rewrite]: #303
+#282 := (iff #112 #302)
+#283 := [rewrite]: #282
+#307 := [monotonicity #283 #304]: #306
+#312 := [trans #307 #310]: #311
+#328 := [monotonicity #312]: #327
+#331 := [monotonicity #328]: #330
+#334 := [monotonicity #331]: #333
+#337 := [monotonicity #334]: #336
+#340 := [quant-intro #337]: #339
+#274 := (~ #179 #179)
+#271 := (~ #176 #176)
+#290 := [refl]: #271
+#275 := [nnf-pos #290]: #274
+#23 := (- #9)
+#22 := (- #8)
+#31 := (mod #22 #23)
+#32 := (- #31)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#33 := (ite #20 #30 #32)
+#34 := (ite #12 0::int #33)
+#35 := (ite #13 #8 #34)
+#36 := (= #29 #35)
+#37 := (forall (vars (?v0 int) (?v1 int)) #36)
+#182 := (iff #37 #179)
+#76 := (and #16 #18)
+#79 := (or #17 #76)
+#152 := (ite #79 #30 #147)
+#155 := (ite #12 0::int #152)
+#158 := (ite #13 #8 #155)
+#161 := (= #29 #158)
+#164 := (forall (vars (?v0 int) (?v1 int)) #161)
+#180 := (iff #164 #179)
+#177 := (iff #161 #176)
+#174 := (= #158 #173)
+#171 := (= #155 #170)
+#168 := (= #152 #167)
+#123 := (iff #79 #122)
+#120 := (iff #76 #119)
+#117 := (iff #18 #116)
+#118 := [rewrite]: #117
+#110 := (iff #16 #109)
+#111 := [rewrite]: #110
+#121 := [monotonicity #111 #118]: #120
+#113 := (iff #17 #112)
+#106 := (iff #15 #105)
+#107 := [rewrite]: #106
+#114 := [monotonicity #107 #111]: #113
+#124 := [monotonicity #114 #121]: #123
+#169 := [monotonicity #124]: #168
+#172 := [monotonicity #169]: #171
+#175 := [monotonicity #172]: #174
+#178 := [monotonicity #175]: #177
+#181 := [quant-intro #178]: #180
+#165 := (iff #37 #164)
+#162 := (iff #36 #161)
+#159 := (= #35 #158)
+#156 := (= #34 #155)
+#153 := (= #33 #152)
+#150 := (= #32 #147)
+#144 := (- #141)
+#148 := (= #144 #147)
+#149 := [rewrite]: #148
+#145 := (= #32 #144)
+#142 := (= #31 #141)
+#87 := (= #23 #86)
+#88 := [rewrite]: #87
+#84 := (= #22 #83)
+#85 := [rewrite]: #84
+#143 := [monotonicity #85 #88]: #142
+#146 := [monotonicity #143]: #145
+#151 := [trans #146 #149]: #150
+#80 := (iff #20 #79)
+#77 := (iff #19 #76)
+#78 := [rewrite]: #77
+#81 := [monotonicity #78]: #80
+#154 := [monotonicity #81 #151]: #153
+#157 := [monotonicity #154]: #156
+#160 := [monotonicity #157]: #159
+#163 := [monotonicity #160]: #162
+#166 := [quant-intro #163]: #165
+#183 := [trans #166 #181]: #182
+#140 := [asserted]: #37
+#184 := [mp #140 #183]: #179
+#291 := [mp~ #184 #275]: #179
+#341 := [mp #291 #340]: #338
+#372 := [mp #341 #371]: #369
+#857 := [mp #372 #856]: #852
+#677 := (not #852)
+#679 := (or #677 #507)
+#430 := (* -1::int 3::int)
+#515 := (* -1::int 0::int)
+#516 := (mod #515 #430)
+#517 := (+ #54 #516)
+#508 := (= #517 0::int)
+#519 := (mod 0::int 3::int)
+#520 := (* -1::int #519)
+#448 := (+ #54 #520)
+#521 := (= #448 0::int)
+#518 := (<= 3::int 0::int)
+#522 := (<= 0::int 0::int)
+#501 := (or #522 #518)
+#836 := (not #501)
+#838 := (>= 0::int 0::int)
+#495 := (or #518 #838)
+#625 := (not #495)
+#832 := (or #625 #836)
+#506 := (ite #832 #521 #508)
+#839 := (= 0::int 0::int)
+#840 := (ite #839 #507 #506)
+#841 := (= 0::int #54)
+#842 := (= 3::int 0::int)
+#837 := (ite #842 #841 #840)
+#680 := (or #677 #837)
+#681 := (iff #680 #679)
+#683 := (iff #679 #679)
+#684 := [rewrite]: #683
+#777 := (iff #837 #507)
+#626 := (ite false #507 #507)
+#781 := (iff #626 #507)
+#783 := [rewrite]: #781
+#773 := (iff #837 #626)
+#776 := (iff #840 #507)
+#1 := true
+#785 := (ite true #507 #507)
+#669 := (iff #785 #507)
+#628 := [rewrite]: #669
+#667 := (iff #840 #785)
+#784 := (iff #506 #507)
+#627 := (iff #506 #626)
+#795 := (iff #508 #507)
+#794 := (= #517 #54)
+#806 := (+ #54 0::int)
+#800 := (= #806 #54)
+#801 := [rewrite]: #800
+#787 := (= #517 #806)
+#792 := (= #516 0::int)
+#798 := -3::int
+#646 := (mod 0::int -3::int)
+#790 := (= #646 0::int)
+#791 := [rewrite]: #790
+#647 := (= #516 #646)
+#804 := (= #430 -3::int)
+#645 := [rewrite]: #804
+#808 := (= #515 0::int)
+#523 := [rewrite]: #808
+#789 := [monotonicity #523 #645]: #647
+#786 := [trans #789 #791]: #792
+#793 := [monotonicity #786]: #787
+#788 := [trans #793 #801]: #794
+#624 := [monotonicity #788]: #795
+#802 := (iff #521 #507)
+#796 := (= #448 #54)
+#810 := (= #448 #806)
+#805 := (= #520 0::int)
+#529 := (= #520 #515)
+#527 := (= #519 0::int)
+#528 := [rewrite]: #527
+#807 := [monotonicity #528]: #529
+#809 := [trans #807 #523]: #805
+#799 := [monotonicity #809]: #810
+#797 := [trans #799 #801]: #796
+#803 := [monotonicity #797]: #802
+#816 := (iff #832 false)
+#536 := (or false false)
+#815 := (iff #536 false)
+#812 := [rewrite]: #815
+#537 := (iff #832 #536)
+#813 := (iff #836 false)
+#825 := (not true)
+#464 := (iff #825 false)
+#827 := [rewrite]: #464
+#543 := (iff #836 #825)
+#541 := (iff #501 true)
+#828 := (or true false)
+#818 := (iff #828 true)
+#540 := [rewrite]: #818
+#831 := (iff #501 #828)
+#835 := (iff #518 false)
+#478 := [rewrite]: #835
+#470 := (iff #522 true)
+#830 := [rewrite]: #470
+#817 := [monotonicity #830 #478]: #831
+#542 := [trans #817 #540]: #541
+#436 := [monotonicity #542]: #543
+#814 := [trans #436 #827]: #813
+#829 := (iff #625 false)
+#820 := (iff #625 #825)
+#823 := (iff #495 true)
+#483 := (or false true)
+#821 := (iff #483 true)
+#822 := [rewrite]: #821
+#484 := (iff #495 #483)
+#819 := (iff #838 true)
+#482 := [rewrite]: #819
+#485 := [monotonicity #478 #482]: #484
+#824 := [trans #485 #822]: #823
+#826 := [monotonicity #824]: #820
+#469 := [trans #826 #827]: #829
+#811 := [monotonicity #469 #814]: #537
+#526 := [trans #811 #812]: #816
+#780 := [monotonicity #526 #803 #624]: #627
+#782 := [trans #780 #783]: #784
+#834 := (iff #839 true)
+#494 := [rewrite]: #834
+#668 := [monotonicity #494 #782]: #667
+#678 := [trans #668 #628]: #776
+#493 := (iff #841 #507)
+#498 := [rewrite]: #493
+#843 := (iff #842 false)
+#833 := [rewrite]: #843
+#774 := [monotonicity #833 #498 #678]: #773
+#662 := [trans #774 #783]: #777
+#682 := [monotonicity #662]: #681
+#685 := [trans #682 #684]: #681
+#673 := [quant-inst]: #680
+#778 := [mp #673 #685]: #679
+#713 := [unit-resolution #778 #857]: #507
+#709 := [monotonicity #713]: #57
+#58 := (not #57)
+#270 := [asserted]: #58
+[unit-resolution #270 #709]: false
+unsat
+48fbc21a65dea108392ceeedb3a0cbae0d536aab 328 0
 #2 := false
 decl f5 :: (-> int S2)
 #53 := 1::int
@@ -19371,7 +19579,7 @@
 #271 := [asserted]: #59
 [unit-resolution #271 #708]: false
 unsat
-165a0539a08e682081735200b8ab8b693d5da15d 303 0
+30266f09c986583d8c5407f4e044da694063d38f 303 0
 #2 := false
 decl f5 :: (-> int S2)
 #11 := 0::int
@@ -19675,7 +19883,1213 @@
 #270 := [asserted]: #58
 [unit-resolution #270 #700]: false
 unsat
-1ea3fbc0fb4cdf962c25032ebbea931c454f09b7 848 0
+29a274d82d7dfcc17be98231a7a95a9ec14cd706 533 0
+#2 := false
+#55 := 3::int
+decl f6 :: (-> S2 int)
+decl f7 :: S2
+#53 := f7
+#54 := (f6 f7)
+#533 := (mod #54 3::int)
+#664 := (>= #533 3::int)
+#665 := (not #664)
+#1 := true
+#75 := [true-axiom]: true
+#674 := (or false #665)
+#635 := [th-lemma]: #674
+#636 := [unit-resolution #635 #75]: #665
+#11 := 0::int
+decl f5 :: (-> int S2)
+decl f4 :: (-> int int int)
+#56 := (f4 #54 3::int)
+#57 := (f5 #56)
+#58 := (f6 #57)
+#84 := -1::int
+#779 := (* -1::int #58)
+#763 := (+ #56 #779)
+#766 := (>= #763 0::int)
+#773 := (= #56 #58)
+#780 := (>= #56 0::int)
+#784 := (= #58 0::int)
+#649 := (not #784)
+#746 := (<= #58 0::int)
+#643 := (not #746)
+#276 := (>= #58 3::int)
+#59 := (< #58 3::int)
+#60 := (not #59)
+#284 := (iff #60 #276)
+#275 := (not #276)
+#279 := (not #275)
+#282 := (iff #279 #276)
+#283 := [rewrite]: #282
+#280 := (iff #60 #279)
+#277 := (iff #59 #275)
+#278 := [rewrite]: #277
+#281 := [monotonicity #278]: #280
+#285 := [trans #281 #283]: #284
+#272 := [asserted]: #60
+#286 := [mp #272 #285]: #276
+#645 := (or #643 #275)
+#646 := [th-lemma]: #645
+#647 := [unit-resolution #646 #286]: #643
+#650 := (or #649 #746)
+#651 := [th-lemma]: #650
+#652 := [unit-resolution #651 #647]: #649
+#9 := (:var 0 int)
+#44 := (f5 #9)
+#880 := (pattern #44)
+#211 := (>= #9 0::int)
+#45 := (f6 #44)
+#50 := (= #45 0::int)
+#261 := (or #50 #211)
+#887 := (forall (vars (?v0 int)) (:pat #880) #261)
+#266 := (forall (vars (?v0 int)) #261)
+#890 := (iff #266 #887)
+#888 := (iff #261 #261)
+#889 := [refl]: #888
+#891 := [quant-intro #889]: #890
+#294 := (~ #266 #266)
+#312 := (~ #261 #261)
+#313 := [refl]: #312
+#295 := [nnf-pos #313]: #294
+#49 := (< #9 0::int)
+#51 := (implies #49 #50)
+#52 := (forall (vars (?v0 int)) #51)
+#269 := (iff #52 #266)
+#232 := (= 0::int #45)
+#238 := (not #49)
+#239 := (or #238 #232)
+#244 := (forall (vars (?v0 int)) #239)
+#267 := (iff #244 #266)
+#264 := (iff #239 #261)
+#258 := (or #211 #50)
+#262 := (iff #258 #261)
+#263 := [rewrite]: #262
+#259 := (iff #239 #258)
+#256 := (iff #232 #50)
+#257 := [rewrite]: #256
+#254 := (iff #238 #211)
+#214 := (not #211)
+#249 := (not #214)
+#252 := (iff #249 #211)
+#253 := [rewrite]: #252
+#250 := (iff #238 #249)
+#247 := (iff #49 #214)
+#248 := [rewrite]: #247
+#251 := [monotonicity #248]: #250
+#255 := [trans #251 #253]: #254
+#260 := [monotonicity #255 #257]: #259
+#265 := [trans #260 #263]: #264
+#268 := [quant-intro #265]: #267
+#245 := (iff #52 #244)
+#242 := (iff #51 #239)
+#235 := (implies #49 #232)
+#240 := (iff #235 #239)
+#241 := [rewrite]: #240
+#236 := (iff #51 #235)
+#233 := (iff #50 #232)
+#234 := [rewrite]: #233
+#237 := [monotonicity #234]: #236
+#243 := [trans #237 #241]: #242
+#246 := [quant-intro #243]: #245
+#270 := [trans #246 #268]: #269
+#231 := [asserted]: #52
+#271 := [mp #231 #270]: #266
+#314 := [mp~ #271 #295]: #266
+#892 := [mp #314 #891]: #887
+#765 := (not #887)
+#770 := (or #765 #780 #784)
+#785 := (or #784 #780)
+#756 := (or #765 #785)
+#742 := (iff #756 #770)
+#767 := (or #780 #784)
+#759 := (or #765 #767)
+#758 := (iff #759 #770)
+#762 := [rewrite]: #758
+#760 := (iff #756 #759)
+#768 := (iff #785 #767)
+#769 := [rewrite]: #768
+#761 := [monotonicity #769]: #760
+#743 := [trans #761 #762]: #742
+#757 := [quant-inst]: #756
+#745 := [mp #757 #743]: #770
+#653 := [unit-resolution #745 #892 #652]: #780
+#782 := (not #780)
+#783 := (or #773 #782)
+#195 := (= #9 #45)
+#220 := (or #195 #214)
+#881 := (forall (vars (?v0 int)) (:pat #880) #220)
+#225 := (forall (vars (?v0 int)) #220)
+#884 := (iff #225 #881)
+#882 := (iff #220 #220)
+#883 := [refl]: #882
+#885 := [quant-intro #883]: #884
+#292 := (~ #225 #225)
+#309 := (~ #220 #220)
+#310 := [refl]: #309
+#293 := [nnf-pos #310]: #292
+#46 := (= #45 #9)
+#43 := (<= 0::int #9)
+#47 := (implies #43 #46)
+#48 := (forall (vars (?v0 int)) #47)
+#228 := (iff #48 #225)
+#202 := (not #43)
+#203 := (or #202 #195)
+#208 := (forall (vars (?v0 int)) #203)
+#226 := (iff #208 #225)
+#223 := (iff #203 #220)
+#217 := (or #214 #195)
+#221 := (iff #217 #220)
+#222 := [rewrite]: #221
+#218 := (iff #203 #217)
+#215 := (iff #202 #214)
+#212 := (iff #43 #211)
+#213 := [rewrite]: #212
+#216 := [monotonicity #213]: #215
+#219 := [monotonicity #216]: #218
+#224 := [trans #219 #222]: #223
+#227 := [quant-intro #224]: #226
+#209 := (iff #48 #208)
+#206 := (iff #47 #203)
+#199 := (implies #43 #195)
+#204 := (iff #199 #203)
+#205 := [rewrite]: #204
+#200 := (iff #47 #199)
+#197 := (iff #46 #195)
+#198 := [rewrite]: #197
+#201 := [monotonicity #198]: #200
+#207 := [trans #201 #205]: #206
+#210 := [quant-intro #207]: #209
+#229 := [trans #210 #227]: #228
+#194 := [asserted]: #48
+#230 := [mp #194 #229]: #225
+#311 := [mp~ #230 #293]: #225
+#886 := [mp #311 #885]: #881
+#781 := (not #881)
+#786 := (or #781 #773 #782)
+#775 := (or #781 #783)
+#777 := (iff #775 #786)
+#778 := [rewrite]: #777
+#776 := [quant-inst]: #775
+#772 := [mp #776 #778]: #786
+#654 := [unit-resolution #772 #886]: #783
+#637 := [unit-resolution #654 #653]: #773
+#655 := (not #773)
+#625 := (or #655 #766)
+#626 := [th-lemma]: #625
+#627 := [unit-resolution #626 #637]: #766
+#534 := (* -1::int #533)
+#462 := (+ #56 #534)
+#802 := (<= #462 0::int)
+#535 := (= #462 0::int)
+#556 := -3::int
+#529 := (* -1::int #54)
+#827 := (mod #529 -3::int)
+#551 := (+ #56 #827)
+#826 := (= #551 0::int)
+#852 := (>= #54 0::int)
+#498 := (not #852)
+#536 := (<= #54 0::int)
+#841 := (not #536)
+#845 := (or #841 #498)
+#541 := (ite #845 #535 #826)
+#521 := (= #56 0::int)
+#853 := (= #54 0::int)
+#821 := (ite #853 #521 #541)
+#8 := (:var 1 int)
+#29 := (f4 #8 #9)
+#865 := (pattern #29)
+#88 := (* -1::int #9)
+#85 := (* -1::int #8)
+#143 := (mod #85 #88)
+#376 := (+ #29 #143)
+#377 := (= #376 0::int)
+#30 := (mod #8 #9)
+#373 := (* -1::int #30)
+#374 := (+ #29 #373)
+#375 := (= #374 0::int)
+#110 := (<= #9 0::int)
+#106 := (<= #8 0::int)
+#315 := (or #106 #110)
+#316 := (not #315)
+#117 := (>= #8 0::int)
+#298 := (or #110 #117)
+#299 := (not #298)
+#322 := (or #299 #316)
+#378 := (ite #322 #375 #377)
+#372 := (= #29 0::int)
+#12 := (= #8 0::int)
+#379 := (ite #12 #372 #378)
+#371 := (= #8 #29)
+#13 := (= #9 0::int)
+#380 := (ite #13 #371 #379)
+#866 := (forall (vars (?v0 int) (?v1 int)) (:pat #865) #380)
+#383 := (forall (vars (?v0 int) (?v1 int)) #380)
+#869 := (iff #383 #866)
+#867 := (iff #380 #380)
+#868 := [refl]: #867
+#870 := [quant-intro #868]: #869
+#149 := (* -1::int #143)
+#340 := (ite #322 #30 #149)
+#343 := (ite #12 0::int #340)
+#346 := (ite #13 #8 #343)
+#349 := (= #29 #346)
+#352 := (forall (vars (?v0 int) (?v1 int)) #349)
+#384 := (iff #352 #383)
+#381 := (iff #349 #380)
+#382 := [rewrite]: #381
+#385 := [quant-intro #382]: #384
+#118 := (not #117)
+#111 := (not #110)
+#121 := (and #111 #118)
+#107 := (not #106)
+#114 := (and #107 #111)
+#124 := (or #114 #121)
+#169 := (ite #124 #30 #149)
+#172 := (ite #12 0::int #169)
+#175 := (ite #13 #8 #172)
+#178 := (= #29 #175)
+#181 := (forall (vars (?v0 int) (?v1 int)) #178)
+#353 := (iff #181 #352)
+#350 := (iff #178 #349)
+#347 := (= #175 #346)
+#344 := (= #172 #343)
+#341 := (= #169 #340)
+#325 := (iff #124 #322)
+#319 := (or #316 #299)
+#323 := (iff #319 #322)
+#324 := [rewrite]: #323
+#320 := (iff #124 #319)
+#317 := (iff #121 #299)
+#318 := [rewrite]: #317
+#296 := (iff #114 #316)
+#297 := [rewrite]: #296
+#321 := [monotonicity #297 #318]: #320
+#326 := [trans #321 #324]: #325
+#342 := [monotonicity #326]: #341
+#345 := [monotonicity #342]: #344
+#348 := [monotonicity #345]: #347
+#351 := [monotonicity #348]: #350
+#354 := [quant-intro #351]: #353
+#288 := (~ #181 #181)
+#273 := (~ #178 #178)
+#304 := [refl]: #273
+#289 := [nnf-pos #304]: #288
+#23 := (- #9)
+#22 := (- #8)
+#31 := (mod #22 #23)
+#32 := (- #31)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#33 := (ite #20 #30 #32)
+#34 := (ite #12 0::int #33)
+#35 := (ite #13 #8 #34)
+#36 := (= #29 #35)
+#37 := (forall (vars (?v0 int) (?v1 int)) #36)
+#184 := (iff #37 #181)
+#78 := (and #16 #18)
+#81 := (or #17 #78)
+#154 := (ite #81 #30 #149)
+#157 := (ite #12 0::int #154)
+#160 := (ite #13 #8 #157)
+#163 := (= #29 #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 #81 #124)
+#122 := (iff #78 #121)
+#119 := (iff #18 #118)
+#120 := [rewrite]: #119
+#112 := (iff #16 #111)
+#113 := [rewrite]: #112
+#123 := [monotonicity #113 #120]: #122
+#115 := (iff #17 #114)
+#108 := (iff #15 #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 #37 #166)
+#164 := (iff #36 #163)
+#161 := (= #35 #160)
+#158 := (= #34 #157)
+#155 := (= #33 #154)
+#152 := (= #32 #149)
+#146 := (- #143)
+#150 := (= #146 #149)
+#151 := [rewrite]: #150
+#147 := (= #32 #146)
+#144 := (= #31 #143)
+#89 := (= #23 #88)
+#90 := [rewrite]: #89
+#86 := (= #22 #85)
+#87 := [rewrite]: #86
+#145 := [monotonicity #87 #90]: #144
+#148 := [monotonicity #145]: #147
+#153 := [trans #148 #151]: #152
+#82 := (iff #20 #81)
+#79 := (iff #19 #78)
+#80 := [rewrite]: #79
+#83 := [monotonicity #80]: #82
+#156 := [monotonicity #83 #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]: #37
+#186 := [mp #142 #185]: #181
+#305 := [mp~ #186 #289]: #181
+#355 := [mp #305 #354]: #352
+#386 := [mp #355 #385]: #383
+#871 := [mp #386 #870]: #866
+#810 := (not #866)
+#811 := (or #810 #821)
+#444 := (* -1::int 3::int)
+#530 := (mod #529 #444)
+#531 := (+ #56 #530)
+#522 := (= #531 0::int)
+#532 := (<= 3::int 0::int)
+#515 := (or #536 #532)
+#850 := (not #515)
+#509 := (or #532 #852)
+#639 := (not #509)
+#846 := (or #639 #850)
+#520 := (ite #846 #535 #522)
+#854 := (ite #853 #521 #520)
+#855 := (= #54 #56)
+#856 := (= 3::int 0::int)
+#851 := (ite #856 #855 #854)
+#816 := (or #810 #851)
+#812 := (iff #816 #811)
+#659 := (iff #811 #811)
+#660 := [rewrite]: #659
+#814 := (iff #851 #821)
+#819 := (ite false #855 #821)
+#824 := (iff #819 #821)
+#813 := [rewrite]: #824
+#823 := (iff #851 #819)
+#822 := (iff #854 #821)
+#542 := (iff #520 #541)
+#830 := (iff #522 #826)
+#825 := (= #531 #551)
+#828 := (= #530 #827)
+#557 := (= #444 -3::int)
+#450 := [rewrite]: #557
+#550 := [monotonicity #450]: #828
+#829 := [monotonicity #550]: #825
+#540 := [monotonicity #829]: #830
+#554 := (iff #846 #845)
+#484 := (or #498 #841)
+#831 := (iff #484 #845)
+#832 := [rewrite]: #831
+#844 := (iff #846 #484)
+#843 := (iff #850 #841)
+#840 := (iff #515 #536)
+#836 := (or #536 false)
+#839 := (iff #836 #536)
+#834 := [rewrite]: #839
+#837 := (iff #515 #836)
+#507 := (iff #532 false)
+#512 := [rewrite]: #507
+#838 := [monotonicity #512]: #837
+#478 := [trans #838 #834]: #840
+#483 := [monotonicity #478]: #843
+#499 := (iff #639 #498)
+#496 := (iff #509 #852)
+#848 := (or false #852)
+#492 := (iff #848 #852)
+#833 := [rewrite]: #492
+#508 := (iff #509 #848)
+#849 := [monotonicity #512]: #508
+#497 := [trans #849 #833]: #496
+#835 := [monotonicity #497]: #499
+#842 := [monotonicity #835 #483]: #844
+#555 := [trans #842 #832]: #554
+#543 := [monotonicity #555 #540]: #542
+#537 := [monotonicity #543]: #822
+#857 := (iff #856 false)
+#847 := [rewrite]: #857
+#820 := [monotonicity #847 #537]: #823
+#815 := [trans #820 #813]: #814
+#818 := [monotonicity #815]: #812
+#661 := [trans #818 #660]: #812
+#817 := [quant-inst]: #816
+#803 := [mp #817 #661]: #811
+#628 := [unit-resolution #803 #871]: #821
+#692 := (not #853)
+#691 := (not #521)
+#633 := (iff #649 #691)
+#632 := (iff #784 #521)
+#630 := (iff #521 #784)
+#631 := [monotonicity #637]: #630
+#475 := [symm #631]: #632
+#629 := [monotonicity #475]: #633
+#634 := [mp #652 #629]: #691
+#790 := (not #821)
+#787 := (or #790 #692 #521)
+#788 := [def-axiom]: #787
+#613 := [unit-resolution #788 #634 #628]: #692
+#791 := (or #790 #853 #541)
+#676 := [def-axiom]: #791
+#616 := [unit-resolution #676 #613 #628]: #541
+#696 := (f5 #54)
+#748 := (f6 #696)
+#749 := (= #748 0::int)
+#614 := (not #749)
+#619 := (iff #692 #614)
+#617 := (iff #853 #749)
+#612 := (iff #749 #853)
+#621 := (= #748 #54)
+#622 := (= #696 f7)
+#697 := (= f7 #696)
+#38 := (:var 0 S2)
+#39 := (f6 #38)
+#872 := (pattern #39)
+#40 := (f5 #39)
+#188 := (= #38 #40)
+#873 := (forall (vars (?v0 S2)) (:pat #872) #188)
+#191 := (forall (vars (?v0 S2)) #188)
+#874 := (iff #191 #873)
+#876 := (iff #873 #873)
+#877 := [rewrite]: #876
+#875 := [rewrite]: #874
+#878 := [trans #875 #877]: #874
+#290 := (~ #191 #191)
+#306 := (~ #188 #188)
+#307 := [refl]: #306
+#291 := [nnf-pos #307]: #290
+#41 := (= #40 #38)
+#42 := (forall (vars (?v0 S2)) #41)
+#192 := (iff #42 #191)
+#189 := (iff #41 #188)
+#190 := [rewrite]: #189
+#193 := [quant-intro #190]: #192
+#187 := [asserted]: #42
+#196 := [mp #187 #193]: #191
+#308 := [mp~ #196 #291]: #191
+#879 := [mp #308 #878]: #873
+#792 := (not #873)
+#789 := (or #792 #697)
+#793 := [quant-inst]: #789
+#620 := [unit-resolution #793 #879]: #697
+#623 := [symm #620]: #622
+#624 := [monotonicity #623]: #621
+#615 := [monotonicity #624]: #612
+#618 := [symm #615]: #617
+#609 := [monotonicity #618]: #619
+#599 := [mp #613 #609]: #614
+#750 := (or #749 #852)
+#753 := (or #765 #749 #852)
+#754 := (or #765 #750)
+#755 := (iff #754 #753)
+#733 := [rewrite]: #755
+#744 := [quant-inst]: #754
+#734 := [mp #744 #733]: #753
+#601 := [unit-resolution #734 #892]: #750
+#602 := [unit-resolution #601 #599]: #852
+#605 := (or #853 #841 #498)
+#606 := [th-lemma]: #605
+#610 := [unit-resolution #606 #613]: #845
+#603 := [unit-resolution #610 #602]: #841
+#804 := (or #845 #536)
+#805 := [def-axiom]: #804
+#611 := [unit-resolution #805 #603]: #845
+#801 := (not #845)
+#641 := (not #541)
+#794 := (or #641 #801 #535)
+#795 := [def-axiom]: #794
+#604 := [unit-resolution #795 #611 #616]: #535
+#796 := (not #535)
+#607 := (or #796 #802)
+#608 := [th-lemma]: #607
+#594 := [unit-resolution #608 #604]: #802
+[th-lemma #594 #286 #627 #636]: false
+unsat
+914ab31d83892bfaa2a8848c9049be9e1ce2cb06 671 0
+#2 := false
+decl f6 :: (-> S2 int)
+decl f5 :: (-> int S2)
+decl f4 :: (-> int int int)
+#55 := 3::int
+decl f7 :: S2
+#53 := f7
+#54 := (f6 f7)
+#56 := (f4 #54 3::int)
+#57 := (f5 #56)
+#805 := (f6 #57)
+#806 := (= #56 #805)
+#476 := (not #806)
+#514 := (= #54 #56)
+#692 := (not #514)
+#473 := (iff #692 #476)
+#483 := (iff #514 #806)
+#490 := (iff #806 #514)
+#679 := (= #56 #54)
+#486 := (iff #679 #514)
+#489 := [commutativity]: #486
+#481 := (iff #806 #679)
+#494 := (= #805 #54)
+#58 := (= #57 f7)
+#274 := (= f7 #57)
+#286 := (>= #54 3::int)
+#11 := 0::int
+#870 := (= #56 0::int)
+#549 := (iff #870 #514)
+#547 := (iff #514 #870)
+#702 := (= 0::int #56)
+#545 := (iff #702 #870)
+#546 := [commutativity]: #545
+#541 := (iff #514 #702)
+#530 := (= #54 0::int)
+#285 := (not #286)
+#632 := [hypothesis]: #285
+#538 := (or #530 #286)
+#708 := 1::int
+#769 := (div #54 3::int)
+#628 := -2::int
+#618 := (* -2::int #769)
+#661 := (mod #54 3::int)
+#85 := -1::int
+#868 := (* -1::int #661)
+#619 := (+ #868 #618)
+#617 := (+ #54 #619)
+#620 := (>= #617 1::int)
+#627 := (not #620)
+#802 := (>= #56 0::int)
+#729 := (>= #661 0::int)
+#1 := true
+#76 := [true-axiom]: true
+#630 := (or false #729)
+#616 := [th-lemma]: #630
+#602 := [unit-resolution #616 #76]: #729
+#542 := (+ #56 #868)
+#817 := (>= #542 0::int)
+#543 := (= #542 0::int)
+#572 := -3::int
+#537 := (* -1::int #54)
+#851 := (mod #537 -3::int)
+#562 := (+ #56 #851)
+#565 := (= #562 0::int)
+#876 := (<= #54 0::int)
+#577 := (not #876)
+#873 := (>= #54 0::int)
+#863 := (not #873)
+#472 := (or #863 #577)
+#559 := (ite #472 #543 #565)
+#713 := (not #530)
+#604 := [hypothesis]: #713
+#606 := (or #530 #559)
+#842 := (ite #530 #870 #559)
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#29 := (f4 #8 #9)
+#887 := (pattern #29)
+#89 := (* -1::int #9)
+#86 := (* -1::int #8)
+#144 := (mod #86 #89)
+#398 := (+ #29 #144)
+#399 := (= #398 0::int)
+#30 := (mod #8 #9)
+#395 := (* -1::int #30)
+#396 := (+ #29 #395)
+#397 := (= #396 0::int)
+#111 := (<= #9 0::int)
+#107 := (<= #8 0::int)
+#337 := (or #107 #111)
+#338 := (not #337)
+#118 := (>= #8 0::int)
+#320 := (or #111 #118)
+#321 := (not #320)
+#344 := (or #321 #338)
+#400 := (ite #344 #397 #399)
+#394 := (= #29 0::int)
+#12 := (= #8 0::int)
+#401 := (ite #12 #394 #400)
+#393 := (= #8 #29)
+#13 := (= #9 0::int)
+#402 := (ite #13 #393 #401)
+#888 := (forall (vars (?v0 int) (?v1 int)) (:pat #887) #402)
+#405 := (forall (vars (?v0 int) (?v1 int)) #402)
+#891 := (iff #405 #888)
+#889 := (iff #402 #402)
+#890 := [refl]: #889
+#892 := [quant-intro #890]: #891
+#150 := (* -1::int #144)
+#362 := (ite #344 #30 #150)
+#365 := (ite #12 0::int #362)
+#368 := (ite #13 #8 #365)
+#371 := (= #29 #368)
+#374 := (forall (vars (?v0 int) (?v1 int)) #371)
+#406 := (iff #374 #405)
+#403 := (iff #371 #402)
+#404 := [rewrite]: #403
+#407 := [quant-intro #404]: #406
+#119 := (not #118)
+#112 := (not #111)
+#122 := (and #112 #119)
+#108 := (not #107)
+#115 := (and #108 #112)
+#125 := (or #115 #122)
+#170 := (ite #125 #30 #150)
+#173 := (ite #12 0::int #170)
+#176 := (ite #13 #8 #173)
+#179 := (= #29 #176)
+#182 := (forall (vars (?v0 int) (?v1 int)) #179)
+#375 := (iff #182 #374)
+#372 := (iff #179 #371)
+#369 := (= #176 #368)
+#366 := (= #173 #365)
+#363 := (= #170 #362)
+#347 := (iff #125 #344)
+#341 := (or #338 #321)
+#345 := (iff #341 #344)
+#346 := [rewrite]: #345
+#342 := (iff #125 #341)
+#339 := (iff #122 #321)
+#340 := [rewrite]: #339
+#318 := (iff #115 #338)
+#319 := [rewrite]: #318
+#343 := [monotonicity #319 #340]: #342
+#348 := [trans #343 #346]: #347
+#364 := [monotonicity #348]: #363
+#367 := [monotonicity #364]: #366
+#370 := [monotonicity #367]: #369
+#373 := [monotonicity #370]: #372
+#376 := [quant-intro #373]: #375
+#310 := (~ #182 #182)
+#307 := (~ #179 #179)
+#326 := [refl]: #307
+#311 := [nnf-pos #326]: #310
+#23 := (- #9)
+#22 := (- #8)
+#31 := (mod #22 #23)
+#32 := (- #31)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#33 := (ite #20 #30 #32)
+#34 := (ite #12 0::int #33)
+#35 := (ite #13 #8 #34)
+#36 := (= #29 #35)
+#37 := (forall (vars (?v0 int) (?v1 int)) #36)
+#185 := (iff #37 #182)
+#79 := (and #16 #18)
+#82 := (or #17 #79)
+#155 := (ite #82 #30 #150)
+#158 := (ite #12 0::int #155)
+#161 := (ite #13 #8 #158)
+#164 := (= #29 #161)
+#167 := (forall (vars (?v0 int) (?v1 int)) #164)
+#183 := (iff #167 #182)
+#180 := (iff #164 #179)
+#177 := (= #161 #176)
+#174 := (= #158 #173)
+#171 := (= #155 #170)
+#126 := (iff #82 #125)
+#123 := (iff #79 #122)
+#120 := (iff #18 #119)
+#121 := [rewrite]: #120
+#113 := (iff #16 #112)
+#114 := [rewrite]: #113
+#124 := [monotonicity #114 #121]: #123
+#116 := (iff #17 #115)
+#109 := (iff #15 #108)
+#110 := [rewrite]: #109
+#117 := [monotonicity #110 #114]: #116
+#127 := [monotonicity #117 #124]: #126
+#172 := [monotonicity #127]: #171
+#175 := [monotonicity #172]: #174
+#178 := [monotonicity #175]: #177
+#181 := [monotonicity #178]: #180
+#184 := [quant-intro #181]: #183
+#168 := (iff #37 #167)
+#165 := (iff #36 #164)
+#162 := (= #35 #161)
+#159 := (= #34 #158)
+#156 := (= #33 #155)
+#153 := (= #32 #150)
+#147 := (- #144)
+#151 := (= #147 #150)
+#152 := [rewrite]: #151
+#148 := (= #32 #147)
+#145 := (= #31 #144)
+#90 := (= #23 #89)
+#91 := [rewrite]: #90
+#87 := (= #22 #86)
+#88 := [rewrite]: #87
+#146 := [monotonicity #88 #91]: #145
+#149 := [monotonicity #146]: #148
+#154 := [trans #149 #152]: #153
+#83 := (iff #20 #82)
+#80 := (iff #19 #79)
+#81 := [rewrite]: #80
+#84 := [monotonicity #81]: #83
+#157 := [monotonicity #84 #154]: #156
+#160 := [monotonicity #157]: #159
+#163 := [monotonicity #160]: #162
+#166 := [monotonicity #163]: #165
+#169 := [quant-intro #166]: #168
+#186 := [trans #169 #184]: #185
+#143 := [asserted]: #37
+#187 := [mp #143 #186]: #182
+#327 := [mp~ #187 #311]: #182
+#377 := [mp #327 #376]: #374
+#408 := [mp #377 #407]: #405
+#893 := [mp #408 #892]: #888
+#840 := (not #888)
+#681 := (or #840 #842)
+#558 := (* -1::int 3::int)
+#872 := (mod #537 #558)
+#874 := (+ #56 #872)
+#531 := (= #874 0::int)
+#875 := (<= 3::int 0::int)
+#877 := (or #876 #875)
+#878 := (not #877)
+#879 := (or #875 #873)
+#869 := (not #879)
+#529 := (or #869 #878)
+#534 := (ite #529 #543 #531)
+#871 := (ite #530 #870 #534)
+#855 := (= 3::int 0::int)
+#518 := (ite #855 #514 #871)
+#682 := (or #840 #518)
+#825 := (iff #682 #681)
+#827 := (iff #681 #681)
+#828 := [rewrite]: #827
+#839 := (iff #518 #842)
+#836 := (ite false #514 #842)
+#833 := (iff #836 #842)
+#838 := [rewrite]: #833
+#837 := (iff #518 #836)
+#846 := (iff #871 #842)
+#841 := (iff #534 #559)
+#843 := (iff #531 #565)
+#563 := (= #874 #562)
+#848 := (= #872 #851)
+#573 := (= #558 -3::int)
+#847 := [rewrite]: #573
+#852 := [monotonicity #847]: #848
+#564 := [monotonicity #852]: #563
+#844 := [monotonicity #564]: #843
+#849 := (iff #529 #472)
+#578 := (iff #878 #577)
+#854 := (iff #877 #876)
+#506 := (or #876 false)
+#867 := (iff #506 #876)
+#853 := [rewrite]: #867
+#866 := (iff #877 #506)
+#521 := (iff #875 false)
+#857 := [rewrite]: #521
+#864 := [monotonicity #857]: #866
+#576 := [trans #864 #853]: #854
+#579 := [monotonicity #576]: #578
+#865 := (iff #869 #863)
+#862 := (iff #879 #873)
+#858 := (or false #873)
+#861 := (iff #858 #873)
+#856 := [rewrite]: #861
+#859 := (iff #879 #858)
+#860 := [monotonicity #857]: #859
+#500 := [trans #860 #856]: #862
+#505 := [monotonicity #500]: #865
+#850 := [monotonicity #505 #579]: #849
+#845 := [monotonicity #850 #844]: #841
+#835 := [monotonicity #845]: #846
+#519 := (iff #855 false)
+#520 := [rewrite]: #519
+#832 := [monotonicity #520 #835]: #837
+#834 := [trans #832 #838]: #839
+#826 := [monotonicity #834]: #825
+#822 := [trans #826 #828]: #825
+#683 := [quant-inst]: #682
+#823 := [mp #683 #822]: #681
+#605 := [unit-resolution #823 #893]: #842
+#698 := (not #842)
+#709 := (or #698 #530 #559)
+#717 := [def-axiom]: #709
+#607 := [unit-resolution #717 #605]: #606
+#608 := [unit-resolution #607 #604]: #559
+#811 := (f5 #54)
+#779 := (f6 #811)
+#781 := (= #779 0::int)
+#592 := (not #781)
+#594 := (iff #713 #592)
+#603 := (iff #530 #781)
+#613 := (iff #781 #530)
+#611 := (= #779 #54)
+#609 := (= #811 f7)
+#815 := (= f7 #811)
+#38 := (:var 0 S2)
+#39 := (f6 #38)
+#894 := (pattern #39)
+#40 := (f5 #39)
+#189 := (= #38 #40)
+#895 := (forall (vars (?v0 S2)) (:pat #894) #189)
+#192 := (forall (vars (?v0 S2)) #189)
+#896 := (iff #192 #895)
+#898 := (iff #895 #895)
+#899 := [rewrite]: #898
+#897 := [rewrite]: #896
+#900 := [trans #897 #899]: #896
+#312 := (~ #192 #192)
+#328 := (~ #189 #189)
+#329 := [refl]: #328
+#313 := [nnf-pos #329]: #312
+#41 := (= #40 #38)
+#42 := (forall (vars (?v0 S2)) #41)
+#193 := (iff #42 #192)
+#190 := (iff #41 #189)
+#191 := [rewrite]: #190
+#194 := [quant-intro #191]: #193
+#188 := [asserted]: #42
+#197 := [mp #188 #194]: #192
+#330 := [mp~ #197 #313]: #192
+#901 := [mp #330 #900]: #895
+#796 := (not #895)
+#793 := (or #796 #815)
+#795 := [quant-inst]: #793
+#685 := [unit-resolution #795 #901]: #815
+#610 := [symm #685]: #609
+#612 := [monotonicity #610]: #611
+#614 := [monotonicity #612]: #613
+#615 := [symm #614]: #603
+#595 := [monotonicity #615]: #594
+#596 := [mp #604 #595]: #592
+#768 := (or #781 #873)
+#44 := (f5 #9)
+#902 := (pattern #44)
+#212 := (>= #9 0::int)
+#45 := (f6 #44)
+#50 := (= #45 0::int)
+#262 := (or #50 #212)
+#909 := (forall (vars (?v0 int)) (:pat #902) #262)
+#267 := (forall (vars (?v0 int)) #262)
+#912 := (iff #267 #909)
+#910 := (iff #262 #262)
+#911 := [refl]: #910
+#913 := [quant-intro #911]: #912
+#316 := (~ #267 #267)
+#334 := (~ #262 #262)
+#335 := [refl]: #334
+#317 := [nnf-pos #335]: #316
+#49 := (< #9 0::int)
+#51 := (implies #49 #50)
+#52 := (forall (vars (?v0 int)) #51)
+#270 := (iff #52 #267)
+#233 := (= 0::int #45)
+#239 := (not #49)
+#240 := (or #239 #233)
+#245 := (forall (vars (?v0 int)) #240)
+#268 := (iff #245 #267)
+#265 := (iff #240 #262)
+#259 := (or #212 #50)
+#263 := (iff #259 #262)
+#264 := [rewrite]: #263
+#260 := (iff #240 #259)
+#257 := (iff #233 #50)
+#258 := [rewrite]: #257
+#255 := (iff #239 #212)
+#215 := (not #212)
+#250 := (not #215)
+#253 := (iff #250 #212)
+#254 := [rewrite]: #253
+#251 := (iff #239 #250)
+#248 := (iff #49 #215)
+#249 := [rewrite]: #248
+#252 := [monotonicity #249]: #251
+#256 := [trans #252 #254]: #255
+#261 := [monotonicity #256 #258]: #260
+#266 := [trans #261 #264]: #265
+#269 := [quant-intro #266]: #268
+#246 := (iff #52 #245)
+#243 := (iff #51 #240)
+#236 := (implies #49 #233)
+#241 := (iff #236 #240)
+#242 := [rewrite]: #241
+#237 := (iff #51 #236)
+#234 := (iff #50 #233)
+#235 := [rewrite]: #234
+#238 := [monotonicity #235]: #237
+#244 := [trans #238 #242]: #243
+#247 := [quant-intro #244]: #246
+#271 := [trans #247 #269]: #270
+#232 := [asserted]: #52
+#272 := [mp #232 #271]: #267
+#336 := [mp~ #272 #317]: #267
+#914 := [mp #336 #913]: #909
+#782 := (not #909)
+#771 := (or #782 #781 #873)
+#772 := (or #782 #768)
+#774 := (iff #772 #771)
+#775 := [rewrite]: #774
+#773 := [quant-inst]: #772
+#776 := [mp #773 #775]: #771
+#597 := [unit-resolution #776 #914]: #768
+#598 := [unit-resolution #597 #596]: #873
+#599 := (or #530 #577 #863)
+#600 := [th-lemma]: #599
+#593 := [unit-resolution #600 #598 #604]: #577
+#824 := (or #472 #876)
+#831 := [def-axiom]: #824
+#601 := [unit-resolution #831 #593]: #472
+#660 := (not #472)
+#818 := (not #559)
+#821 := (or #818 #660 #543)
+#703 := [def-axiom]: #821
+#586 := [unit-resolution #703 #601 #608]: #543
+#664 := (not #543)
+#587 := (or #664 #817)
+#588 := [th-lemma]: #587
+#590 := [unit-resolution #588 #586]: #817
+#591 := (not #729)
+#589 := (not #817)
+#580 := (or #802 #589 #591)
+#581 := [th-lemma]: #580
+#582 := [unit-resolution #581 #590 #602]: #802
+#816 := (<= #542 0::int)
+#574 := (or #664 #816)
+#583 := [th-lemma]: #574
+#584 := [unit-resolution #583 #586]: #816
+#804 := (not #802)
+#649 := (not #816)
+#626 := (or #627 #286 #649 #804)
+#625 := [hypothesis]: #802
+#747 := (* -3::int #769)
+#750 := (+ #868 #747)
+#751 := (+ #54 #750)
+#727 := (<= #751 0::int)
+#746 := (= #751 0::int)
+#657 := (or false #746)
+#695 := [th-lemma]: #657
+#680 := [unit-resolution #695 #76]: #746
+#658 := (not #746)
+#665 := (or #658 #727)
+#667 := [th-lemma]: #665
+#668 := [unit-resolution #667 #680]: #727
+#677 := [hypothesis]: #816
+#633 := [hypothesis]: #620
+#624 := [th-lemma #632 #633 #677 #668 #625]: false
+#629 := [lemma #624]: #626
+#575 := [unit-resolution #629 #584 #632 #582]: #627
+#666 := (>= #751 0::int)
+#640 := (or #658 #666)
+#636 := [th-lemma]: #640
+#641 := [unit-resolution #636 #680]: #666
+#711 := (* -1::int #56)
+#712 := (+ #54 #711)
+#722 := (<= #712 0::int)
+#560 := (not #722)
+#555 := (not #274)
+#557 := (or #555 #286)
+#299 := (iff #274 #286)
+#59 := (< #54 3::int)
+#60 := (not #59)
+#61 := (iff #58 #60)
+#304 := (iff #61 #299)
+#280 := (iff #60 #274)
+#302 := (iff #280 #299)
+#296 := (iff #286 #274)
+#300 := (iff #296 #299)
+#301 := [rewrite]: #300
+#297 := (iff #280 #296)
+#294 := (iff #60 #286)
+#289 := (not #285)
+#292 := (iff #289 #286)
+#293 := [rewrite]: #292
+#290 := (iff #60 #289)
+#287 := (iff #59 #285)
+#288 := [rewrite]: #287
+#291 := [monotonicity #288]: #290
+#295 := [trans #291 #293]: #294
+#298 := [monotonicity #295]: #297
+#303 := [trans #298 #301]: #302
+#283 := (iff #61 #280)
+#277 := (iff #274 #60)
+#281 := (iff #277 #280)
+#282 := [rewrite]: #281
+#278 := (iff #61 #277)
+#275 := (iff #58 #274)
+#276 := [rewrite]: #275
+#279 := [monotonicity #276]: #278
+#284 := [trans #279 #282]: #283
+#305 := [trans #284 #303]: #304
+#273 := [asserted]: #61
+#306 := [mp #273 #305]: #299
+#466 := (not #299)
+#556 := (or #555 #286 #466)
+#484 := [def-axiom]: #556
+#554 := [unit-resolution #484 #306]: #557
+#585 := [unit-resolution #554 #632]: #555
+#693 := (or #692 #274)
+#688 := (= #811 #57)
+#686 := (= #57 #811)
+#678 := [hypothesis]: #514
+#684 := [symm #678]: #679
+#687 := [monotonicity #684]: #686
+#689 := [symm #687]: #688
+#690 := [trans #685 #689]: #274
+#723 := [hypothesis]: #555
+#691 := [unit-resolution #723 #690]: false
+#694 := [lemma #691]: #693
+#566 := [unit-resolution #694 #585]: #692
+#707 := (>= #712 0::int)
+#669 := (* -1::int #769)
+#671 := (+ #868 #669)
+#637 := (+ #54 #671)
+#639 := (>= #637 0::int)
+#621 := (or #639 #876)
+#645 := (not #639)
+#634 := [hypothesis]: #645
+#706 := (>= #661 3::int)
+#730 := (not #706)
+#675 := (or false #730)
+#676 := [th-lemma]: #675
+#659 := [unit-resolution #676 #76]: #730
+#696 := [hypothesis]: #577
+#631 := [th-lemma #696 #659 #641 #634]: false
+#623 := [lemma #631]: #621
+#567 := [unit-resolution #623 #593]: #639
+#643 := (or #645 #649 #707)
+#642 := (not #707)
+#644 := [hypothesis]: #642
+#635 := [hypothesis]: #639
+#638 := [th-lemma #677 #635 #668 #644]: false
+#646 := [lemma #638]: #643
+#568 := [unit-resolution #646 #567 #584]: #707
+#569 := (or #514 #560 #642)
+#570 := [th-lemma]: #569
+#561 := [unit-resolution #570 #568 #566]: #560
+#571 := [th-lemma #590 #561 #641 #575]: false
+#540 := [lemma #571]: #538
+#697 := [unit-resolution #540 #632]: #530
+#532 := [monotonicity #697]: #541
+#548 := [trans #532 #546]: #547
+#539 := [symm #548]: #549
+#699 := (or #713 #870)
+#715 := (or #698 #713 #870)
+#716 := [def-axiom]: #715
+#701 := [unit-resolution #716 #605]: #699
+#728 := [unit-resolution #701 #697]: #870
+#550 := [mp #728 #539]: #514
+#533 := [unit-resolution #566 #550]: false
+#535 := [lemma #533]: #286
+#553 := (or #274 #285)
+#551 := (or #274 #285 #466)
+#552 := [def-axiom]: #551
+#544 := [unit-resolution #552 #306]: #553
+#496 := [unit-resolution #544 #535]: #274
+#498 := [symm #496]: #58
+#499 := [monotonicity #498]: #494
+#485 := [monotonicity #499]: #481
+#491 := [trans #485 #489]: #490
+#492 := [symm #491]: #483
+#474 := [monotonicity #492]: #473
+#524 := (or #577 #285)
+#525 := [th-lemma]: #524
+#526 := [unit-resolution #525 #535]: #577
+#527 := (or #713 #876)
+#516 := [th-lemma]: #527
+#528 := [unit-resolution #516 #526]: #713
+#509 := [unit-resolution #607 #528]: #559
+#511 := [unit-resolution #831 #526]: #472
+#512 := [unit-resolution #703 #511 #509]: #543
+#502 := [unit-resolution #583 #512]: #816
+#508 := (or #560 #285 #706 #649)
+#482 := [th-lemma]: #508
+#487 := [unit-resolution #482 #502 #535 #659]: #560
+#488 := (or #692 #722)
+#493 := [th-lemma]: #488
+#495 := [unit-resolution #493 #487]: #692
+#477 := [mp #495 #474]: #476
+#510 := [unit-resolution #588 #512]: #817
+#513 := (or #802 #591 #589)
+#501 := [th-lemma]: #513
+#503 := [unit-resolution #501 #510 #602]: #802
+#803 := (or #804 #806)
+#196 := (= #9 #45)
+#221 := (or #196 #215)
+#903 := (forall (vars (?v0 int)) (:pat #902) #221)
+#226 := (forall (vars (?v0 int)) #221)
+#906 := (iff #226 #903)
+#904 := (iff #221 #221)
+#905 := [refl]: #904
+#907 := [quant-intro #905]: #906
+#314 := (~ #226 #226)
+#331 := (~ #221 #221)
+#332 := [refl]: #331
+#315 := [nnf-pos #332]: #314
+#46 := (= #45 #9)
+#43 := (<= 0::int #9)
+#47 := (implies #43 #46)
+#48 := (forall (vars (?v0 int)) #47)
+#229 := (iff #48 #226)
+#203 := (not #43)
+#204 := (or #203 #196)
+#209 := (forall (vars (?v0 int)) #204)
+#227 := (iff #209 #226)
+#224 := (iff #204 #221)
+#218 := (or #215 #196)
+#222 := (iff #218 #221)
+#223 := [rewrite]: #222
+#219 := (iff #204 #218)
+#216 := (iff #203 #215)
+#213 := (iff #43 #212)
+#214 := [rewrite]: #213
+#217 := [monotonicity #214]: #216
+#220 := [monotonicity #217]: #219
+#225 := [trans #220 #223]: #224
+#228 := [quant-intro #225]: #227
+#210 := (iff #48 #209)
+#207 := (iff #47 #204)
+#200 := (implies #43 #196)
+#205 := (iff #200 #204)
+#206 := [rewrite]: #205
+#201 := (iff #47 #200)
+#198 := (iff #46 #196)
+#199 := [rewrite]: #198
+#202 := [monotonicity #199]: #201
+#208 := [trans #202 #206]: #207
+#211 := [quant-intro #208]: #210
+#230 := [trans #211 #228]: #229
+#195 := [asserted]: #48
+#231 := [mp #195 #230]: #226
+#333 := [mp~ #231 #315]: #226
+#908 := [mp #333 #907]: #903
+#798 := (not #903)
+#799 := (or #798 #804 #806)
+#807 := (or #806 #804)
+#800 := (or #798 #807)
+#790 := (iff #800 #799)
+#801 := (or #798 #803)
+#788 := (iff #801 #799)
+#789 := [rewrite]: #788
+#785 := (iff #800 #801)
+#808 := (iff #807 #803)
+#797 := [rewrite]: #808
+#786 := [monotonicity #797]: #785
+#791 := [trans #786 #789]: #790
+#794 := [quant-inst]: #800
+#787 := [mp #794 #791]: #799
+#504 := [unit-resolution #787 #908]: #803
+#507 := [unit-resolution #504 #503]: #806
+[unit-resolution #507 #477]: false
+unsat
+d5d2986a8efaefb5e38f30bc14afaad18d3334c6 848 0
 #2 := false
 decl f5 :: (-> int S2)
 decl f6 :: (-> S2 int)
@@ -20524,541 +21938,7 @@
 #1103 := [trans #935 #1102]: #300
 [unit-resolution #308 #1103]: false
 unsat
-c6db48e3ca2ab6d6ca0cc08571bd32ed207dc08a 533 0
-#2 := false
-#55 := 3::int
-decl f6 :: (-> S2 int)
-decl f7 :: S2
-#53 := f7
-#54 := (f6 f7)
-#533 := (mod #54 3::int)
-#664 := (>= #533 3::int)
-#665 := (not #664)
-#1 := true
-#75 := [true-axiom]: true
-#674 := (or false #665)
-#635 := [th-lemma]: #674
-#636 := [unit-resolution #635 #75]: #665
-#11 := 0::int
-decl f5 :: (-> int S2)
-decl f4 :: (-> int int int)
-#56 := (f4 #54 3::int)
-#57 := (f5 #56)
-#58 := (f6 #57)
-#84 := -1::int
-#779 := (* -1::int #58)
-#763 := (+ #56 #779)
-#766 := (>= #763 0::int)
-#773 := (= #56 #58)
-#780 := (>= #56 0::int)
-#784 := (= #58 0::int)
-#649 := (not #784)
-#746 := (<= #58 0::int)
-#643 := (not #746)
-#276 := (>= #58 3::int)
-#59 := (< #58 3::int)
-#60 := (not #59)
-#284 := (iff #60 #276)
-#275 := (not #276)
-#279 := (not #275)
-#282 := (iff #279 #276)
-#283 := [rewrite]: #282
-#280 := (iff #60 #279)
-#277 := (iff #59 #275)
-#278 := [rewrite]: #277
-#281 := [monotonicity #278]: #280
-#285 := [trans #281 #283]: #284
-#272 := [asserted]: #60
-#286 := [mp #272 #285]: #276
-#645 := (or #643 #275)
-#646 := [th-lemma]: #645
-#647 := [unit-resolution #646 #286]: #643
-#650 := (or #649 #746)
-#651 := [th-lemma]: #650
-#652 := [unit-resolution #651 #647]: #649
-#9 := (:var 0 int)
-#44 := (f5 #9)
-#880 := (pattern #44)
-#211 := (>= #9 0::int)
-#45 := (f6 #44)
-#50 := (= #45 0::int)
-#261 := (or #50 #211)
-#887 := (forall (vars (?v0 int)) (:pat #880) #261)
-#266 := (forall (vars (?v0 int)) #261)
-#890 := (iff #266 #887)
-#888 := (iff #261 #261)
-#889 := [refl]: #888
-#891 := [quant-intro #889]: #890
-#294 := (~ #266 #266)
-#312 := (~ #261 #261)
-#313 := [refl]: #312
-#295 := [nnf-pos #313]: #294
-#49 := (< #9 0::int)
-#51 := (implies #49 #50)
-#52 := (forall (vars (?v0 int)) #51)
-#269 := (iff #52 #266)
-#232 := (= 0::int #45)
-#238 := (not #49)
-#239 := (or #238 #232)
-#244 := (forall (vars (?v0 int)) #239)
-#267 := (iff #244 #266)
-#264 := (iff #239 #261)
-#258 := (or #211 #50)
-#262 := (iff #258 #261)
-#263 := [rewrite]: #262
-#259 := (iff #239 #258)
-#256 := (iff #232 #50)
-#257 := [rewrite]: #256
-#254 := (iff #238 #211)
-#214 := (not #211)
-#249 := (not #214)
-#252 := (iff #249 #211)
-#253 := [rewrite]: #252
-#250 := (iff #238 #249)
-#247 := (iff #49 #214)
-#248 := [rewrite]: #247
-#251 := [monotonicity #248]: #250
-#255 := [trans #251 #253]: #254
-#260 := [monotonicity #255 #257]: #259
-#265 := [trans #260 #263]: #264
-#268 := [quant-intro #265]: #267
-#245 := (iff #52 #244)
-#242 := (iff #51 #239)
-#235 := (implies #49 #232)
-#240 := (iff #235 #239)
-#241 := [rewrite]: #240
-#236 := (iff #51 #235)
-#233 := (iff #50 #232)
-#234 := [rewrite]: #233
-#237 := [monotonicity #234]: #236
-#243 := [trans #237 #241]: #242
-#246 := [quant-intro #243]: #245
-#270 := [trans #246 #268]: #269
-#231 := [asserted]: #52
-#271 := [mp #231 #270]: #266
-#314 := [mp~ #271 #295]: #266
-#892 := [mp #314 #891]: #887
-#765 := (not #887)
-#770 := (or #765 #780 #784)
-#785 := (or #784 #780)
-#756 := (or #765 #785)
-#742 := (iff #756 #770)
-#767 := (or #780 #784)
-#759 := (or #765 #767)
-#758 := (iff #759 #770)
-#762 := [rewrite]: #758
-#760 := (iff #756 #759)
-#768 := (iff #785 #767)
-#769 := [rewrite]: #768
-#761 := [monotonicity #769]: #760
-#743 := [trans #761 #762]: #742
-#757 := [quant-inst]: #756
-#745 := [mp #757 #743]: #770
-#653 := [unit-resolution #745 #892 #652]: #780
-#782 := (not #780)
-#783 := (or #773 #782)
-#195 := (= #9 #45)
-#220 := (or #195 #214)
-#881 := (forall (vars (?v0 int)) (:pat #880) #220)
-#225 := (forall (vars (?v0 int)) #220)
-#884 := (iff #225 #881)
-#882 := (iff #220 #220)
-#883 := [refl]: #882
-#885 := [quant-intro #883]: #884
-#292 := (~ #225 #225)
-#309 := (~ #220 #220)
-#310 := [refl]: #309
-#293 := [nnf-pos #310]: #292
-#46 := (= #45 #9)
-#43 := (<= 0::int #9)
-#47 := (implies #43 #46)
-#48 := (forall (vars (?v0 int)) #47)
-#228 := (iff #48 #225)
-#202 := (not #43)
-#203 := (or #202 #195)
-#208 := (forall (vars (?v0 int)) #203)
-#226 := (iff #208 #225)
-#223 := (iff #203 #220)
-#217 := (or #214 #195)
-#221 := (iff #217 #220)
-#222 := [rewrite]: #221
-#218 := (iff #203 #217)
-#215 := (iff #202 #214)
-#212 := (iff #43 #211)
-#213 := [rewrite]: #212
-#216 := [monotonicity #213]: #215
-#219 := [monotonicity #216]: #218
-#224 := [trans #219 #222]: #223
-#227 := [quant-intro #224]: #226
-#209 := (iff #48 #208)
-#206 := (iff #47 #203)
-#199 := (implies #43 #195)
-#204 := (iff #199 #203)
-#205 := [rewrite]: #204
-#200 := (iff #47 #199)
-#197 := (iff #46 #195)
-#198 := [rewrite]: #197
-#201 := [monotonicity #198]: #200
-#207 := [trans #201 #205]: #206
-#210 := [quant-intro #207]: #209
-#229 := [trans #210 #227]: #228
-#194 := [asserted]: #48
-#230 := [mp #194 #229]: #225
-#311 := [mp~ #230 #293]: #225
-#886 := [mp #311 #885]: #881
-#781 := (not #881)
-#786 := (or #781 #773 #782)
-#775 := (or #781 #783)
-#777 := (iff #775 #786)
-#778 := [rewrite]: #777
-#776 := [quant-inst]: #775
-#772 := [mp #776 #778]: #786
-#654 := [unit-resolution #772 #886]: #783
-#637 := [unit-resolution #654 #653]: #773
-#655 := (not #773)
-#625 := (or #655 #766)
-#626 := [th-lemma]: #625
-#627 := [unit-resolution #626 #637]: #766
-#534 := (* -1::int #533)
-#462 := (+ #56 #534)
-#802 := (<= #462 0::int)
-#535 := (= #462 0::int)
-#556 := -3::int
-#529 := (* -1::int #54)
-#827 := (mod #529 -3::int)
-#551 := (+ #56 #827)
-#826 := (= #551 0::int)
-#852 := (>= #54 0::int)
-#498 := (not #852)
-#536 := (<= #54 0::int)
-#841 := (not #536)
-#845 := (or #841 #498)
-#541 := (ite #845 #535 #826)
-#521 := (= #56 0::int)
-#853 := (= #54 0::int)
-#821 := (ite #853 #521 #541)
-#8 := (:var 1 int)
-#29 := (f4 #8 #9)
-#865 := (pattern #29)
-#88 := (* -1::int #9)
-#85 := (* -1::int #8)
-#143 := (mod #85 #88)
-#376 := (+ #29 #143)
-#377 := (= #376 0::int)
-#30 := (mod #8 #9)
-#373 := (* -1::int #30)
-#374 := (+ #29 #373)
-#375 := (= #374 0::int)
-#110 := (<= #9 0::int)
-#106 := (<= #8 0::int)
-#315 := (or #106 #110)
-#316 := (not #315)
-#117 := (>= #8 0::int)
-#298 := (or #110 #117)
-#299 := (not #298)
-#322 := (or #299 #316)
-#378 := (ite #322 #375 #377)
-#372 := (= #29 0::int)
-#12 := (= #8 0::int)
-#379 := (ite #12 #372 #378)
-#371 := (= #8 #29)
-#13 := (= #9 0::int)
-#380 := (ite #13 #371 #379)
-#866 := (forall (vars (?v0 int) (?v1 int)) (:pat #865) #380)
-#383 := (forall (vars (?v0 int) (?v1 int)) #380)
-#869 := (iff #383 #866)
-#867 := (iff #380 #380)
-#868 := [refl]: #867
-#870 := [quant-intro #868]: #869
-#149 := (* -1::int #143)
-#340 := (ite #322 #30 #149)
-#343 := (ite #12 0::int #340)
-#346 := (ite #13 #8 #343)
-#349 := (= #29 #346)
-#352 := (forall (vars (?v0 int) (?v1 int)) #349)
-#384 := (iff #352 #383)
-#381 := (iff #349 #380)
-#382 := [rewrite]: #381
-#385 := [quant-intro #382]: #384
-#118 := (not #117)
-#111 := (not #110)
-#121 := (and #111 #118)
-#107 := (not #106)
-#114 := (and #107 #111)
-#124 := (or #114 #121)
-#169 := (ite #124 #30 #149)
-#172 := (ite #12 0::int #169)
-#175 := (ite #13 #8 #172)
-#178 := (= #29 #175)
-#181 := (forall (vars (?v0 int) (?v1 int)) #178)
-#353 := (iff #181 #352)
-#350 := (iff #178 #349)
-#347 := (= #175 #346)
-#344 := (= #172 #343)
-#341 := (= #169 #340)
-#325 := (iff #124 #322)
-#319 := (or #316 #299)
-#323 := (iff #319 #322)
-#324 := [rewrite]: #323
-#320 := (iff #124 #319)
-#317 := (iff #121 #299)
-#318 := [rewrite]: #317
-#296 := (iff #114 #316)
-#297 := [rewrite]: #296
-#321 := [monotonicity #297 #318]: #320
-#326 := [trans #321 #324]: #325
-#342 := [monotonicity #326]: #341
-#345 := [monotonicity #342]: #344
-#348 := [monotonicity #345]: #347
-#351 := [monotonicity #348]: #350
-#354 := [quant-intro #351]: #353
-#288 := (~ #181 #181)
-#273 := (~ #178 #178)
-#304 := [refl]: #273
-#289 := [nnf-pos #304]: #288
-#23 := (- #9)
-#22 := (- #8)
-#31 := (mod #22 #23)
-#32 := (- #31)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#33 := (ite #20 #30 #32)
-#34 := (ite #12 0::int #33)
-#35 := (ite #13 #8 #34)
-#36 := (= #29 #35)
-#37 := (forall (vars (?v0 int) (?v1 int)) #36)
-#184 := (iff #37 #181)
-#78 := (and #16 #18)
-#81 := (or #17 #78)
-#154 := (ite #81 #30 #149)
-#157 := (ite #12 0::int #154)
-#160 := (ite #13 #8 #157)
-#163 := (= #29 #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 #81 #124)
-#122 := (iff #78 #121)
-#119 := (iff #18 #118)
-#120 := [rewrite]: #119
-#112 := (iff #16 #111)
-#113 := [rewrite]: #112
-#123 := [monotonicity #113 #120]: #122
-#115 := (iff #17 #114)
-#108 := (iff #15 #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 #37 #166)
-#164 := (iff #36 #163)
-#161 := (= #35 #160)
-#158 := (= #34 #157)
-#155 := (= #33 #154)
-#152 := (= #32 #149)
-#146 := (- #143)
-#150 := (= #146 #149)
-#151 := [rewrite]: #150
-#147 := (= #32 #146)
-#144 := (= #31 #143)
-#89 := (= #23 #88)
-#90 := [rewrite]: #89
-#86 := (= #22 #85)
-#87 := [rewrite]: #86
-#145 := [monotonicity #87 #90]: #144
-#148 := [monotonicity #145]: #147
-#153 := [trans #148 #151]: #152
-#82 := (iff #20 #81)
-#79 := (iff #19 #78)
-#80 := [rewrite]: #79
-#83 := [monotonicity #80]: #82
-#156 := [monotonicity #83 #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]: #37
-#186 := [mp #142 #185]: #181
-#305 := [mp~ #186 #289]: #181
-#355 := [mp #305 #354]: #352
-#386 := [mp #355 #385]: #383
-#871 := [mp #386 #870]: #866
-#810 := (not #866)
-#811 := (or #810 #821)
-#444 := (* -1::int 3::int)
-#530 := (mod #529 #444)
-#531 := (+ #56 #530)
-#522 := (= #531 0::int)
-#532 := (<= 3::int 0::int)
-#515 := (or #536 #532)
-#850 := (not #515)
-#509 := (or #532 #852)
-#639 := (not #509)
-#846 := (or #639 #850)
-#520 := (ite #846 #535 #522)
-#854 := (ite #853 #521 #520)
-#855 := (= #54 #56)
-#856 := (= 3::int 0::int)
-#851 := (ite #856 #855 #854)
-#816 := (or #810 #851)
-#812 := (iff #816 #811)
-#659 := (iff #811 #811)
-#660 := [rewrite]: #659
-#814 := (iff #851 #821)
-#819 := (ite false #855 #821)
-#824 := (iff #819 #821)
-#813 := [rewrite]: #824
-#823 := (iff #851 #819)
-#822 := (iff #854 #821)
-#542 := (iff #520 #541)
-#830 := (iff #522 #826)
-#825 := (= #531 #551)
-#828 := (= #530 #827)
-#557 := (= #444 -3::int)
-#450 := [rewrite]: #557
-#550 := [monotonicity #450]: #828
-#829 := [monotonicity #550]: #825
-#540 := [monotonicity #829]: #830
-#554 := (iff #846 #845)
-#484 := (or #498 #841)
-#831 := (iff #484 #845)
-#832 := [rewrite]: #831
-#844 := (iff #846 #484)
-#843 := (iff #850 #841)
-#840 := (iff #515 #536)
-#836 := (or #536 false)
-#839 := (iff #836 #536)
-#834 := [rewrite]: #839
-#837 := (iff #515 #836)
-#507 := (iff #532 false)
-#512 := [rewrite]: #507
-#838 := [monotonicity #512]: #837
-#478 := [trans #838 #834]: #840
-#483 := [monotonicity #478]: #843
-#499 := (iff #639 #498)
-#496 := (iff #509 #852)
-#848 := (or false #852)
-#492 := (iff #848 #852)
-#833 := [rewrite]: #492
-#508 := (iff #509 #848)
-#849 := [monotonicity #512]: #508
-#497 := [trans #849 #833]: #496
-#835 := [monotonicity #497]: #499
-#842 := [monotonicity #835 #483]: #844
-#555 := [trans #842 #832]: #554
-#543 := [monotonicity #555 #540]: #542
-#537 := [monotonicity #543]: #822
-#857 := (iff #856 false)
-#847 := [rewrite]: #857
-#820 := [monotonicity #847 #537]: #823
-#815 := [trans #820 #813]: #814
-#818 := [monotonicity #815]: #812
-#661 := [trans #818 #660]: #812
-#817 := [quant-inst]: #816
-#803 := [mp #817 #661]: #811
-#628 := [unit-resolution #803 #871]: #821
-#692 := (not #853)
-#691 := (not #521)
-#633 := (iff #649 #691)
-#632 := (iff #784 #521)
-#630 := (iff #521 #784)
-#631 := [monotonicity #637]: #630
-#475 := [symm #631]: #632
-#629 := [monotonicity #475]: #633
-#634 := [mp #652 #629]: #691
-#790 := (not #821)
-#787 := (or #790 #692 #521)
-#788 := [def-axiom]: #787
-#613 := [unit-resolution #788 #634 #628]: #692
-#791 := (or #790 #853 #541)
-#676 := [def-axiom]: #791
-#616 := [unit-resolution #676 #613 #628]: #541
-#696 := (f5 #54)
-#748 := (f6 #696)
-#749 := (= #748 0::int)
-#614 := (not #749)
-#619 := (iff #692 #614)
-#617 := (iff #853 #749)
-#612 := (iff #749 #853)
-#621 := (= #748 #54)
-#622 := (= #696 f7)
-#697 := (= f7 #696)
-#38 := (:var 0 S2)
-#39 := (f6 #38)
-#872 := (pattern #39)
-#40 := (f5 #39)
-#188 := (= #38 #40)
-#873 := (forall (vars (?v0 S2)) (:pat #872) #188)
-#191 := (forall (vars (?v0 S2)) #188)
-#874 := (iff #191 #873)
-#876 := (iff #873 #873)
-#877 := [rewrite]: #876
-#875 := [rewrite]: #874
-#878 := [trans #875 #877]: #874
-#290 := (~ #191 #191)
-#306 := (~ #188 #188)
-#307 := [refl]: #306
-#291 := [nnf-pos #307]: #290
-#41 := (= #40 #38)
-#42 := (forall (vars (?v0 S2)) #41)
-#192 := (iff #42 #191)
-#189 := (iff #41 #188)
-#190 := [rewrite]: #189
-#193 := [quant-intro #190]: #192
-#187 := [asserted]: #42
-#196 := [mp #187 #193]: #191
-#308 := [mp~ #196 #291]: #191
-#879 := [mp #308 #878]: #873
-#792 := (not #873)
-#789 := (or #792 #697)
-#793 := [quant-inst]: #789
-#620 := [unit-resolution #793 #879]: #697
-#623 := [symm #620]: #622
-#624 := [monotonicity #623]: #621
-#615 := [monotonicity #624]: #612
-#618 := [symm #615]: #617
-#609 := [monotonicity #618]: #619
-#599 := [mp #613 #609]: #614
-#750 := (or #749 #852)
-#753 := (or #765 #749 #852)
-#754 := (or #765 #750)
-#755 := (iff #754 #753)
-#733 := [rewrite]: #755
-#744 := [quant-inst]: #754
-#734 := [mp #744 #733]: #753
-#601 := [unit-resolution #734 #892]: #750
-#602 := [unit-resolution #601 #599]: #852
-#605 := (or #853 #841 #498)
-#606 := [th-lemma]: #605
-#610 := [unit-resolution #606 #613]: #845
-#603 := [unit-resolution #610 #602]: #841
-#804 := (or #845 #536)
-#805 := [def-axiom]: #804
-#611 := [unit-resolution #805 #603]: #845
-#801 := (not #845)
-#641 := (not #541)
-#794 := (or #641 #801 #535)
-#795 := [def-axiom]: #794
-#604 := [unit-resolution #795 #611 #616]: #535
-#796 := (not #535)
-#607 := (or #796 #802)
-#608 := [th-lemma]: #607
-#594 := [unit-resolution #608 #604]: #802
-[th-lemma #594 #286 #627 #636]: false
-unsat
-8342168df0506311d150a31ce587a3d2d101af50 961 0
+7c27e22543c792c36da02bc4168bcd3f7ce7432d 961 0
 #2 := false
 decl f5 :: (-> int S2)
 decl f6 :: (-> S2 int)
@@ -22020,679 +22900,7 @@
 #1225 := [trans #1224 #1218]: #301
 [unit-resolution #309 #1225]: false
 unsat
-c2c50da38f700cf1e41aeaaa2ecd24baa852fd7f 671 0
-#2 := false
-decl f6 :: (-> S2 int)
-decl f5 :: (-> int S2)
-decl f4 :: (-> int int int)
-#55 := 3::int
-decl f7 :: S2
-#53 := f7
-#54 := (f6 f7)
-#56 := (f4 #54 3::int)
-#57 := (f5 #56)
-#805 := (f6 #57)
-#806 := (= #56 #805)
-#476 := (not #806)
-#514 := (= #54 #56)
-#692 := (not #514)
-#473 := (iff #692 #476)
-#483 := (iff #514 #806)
-#490 := (iff #806 #514)
-#679 := (= #56 #54)
-#486 := (iff #679 #514)
-#489 := [commutativity]: #486
-#481 := (iff #806 #679)
-#494 := (= #805 #54)
-#58 := (= #57 f7)
-#274 := (= f7 #57)
-#286 := (>= #54 3::int)
-#11 := 0::int
-#870 := (= #56 0::int)
-#549 := (iff #870 #514)
-#547 := (iff #514 #870)
-#702 := (= 0::int #56)
-#545 := (iff #702 #870)
-#546 := [commutativity]: #545
-#541 := (iff #514 #702)
-#530 := (= #54 0::int)
-#285 := (not #286)
-#632 := [hypothesis]: #285
-#538 := (or #530 #286)
-#708 := 1::int
-#769 := (div #54 3::int)
-#628 := -2::int
-#618 := (* -2::int #769)
-#661 := (mod #54 3::int)
-#85 := -1::int
-#868 := (* -1::int #661)
-#619 := (+ #868 #618)
-#617 := (+ #54 #619)
-#620 := (>= #617 1::int)
-#627 := (not #620)
-#802 := (>= #56 0::int)
-#729 := (>= #661 0::int)
-#1 := true
-#76 := [true-axiom]: true
-#630 := (or false #729)
-#616 := [th-lemma]: #630
-#602 := [unit-resolution #616 #76]: #729
-#542 := (+ #56 #868)
-#817 := (>= #542 0::int)
-#543 := (= #542 0::int)
-#572 := -3::int
-#537 := (* -1::int #54)
-#851 := (mod #537 -3::int)
-#562 := (+ #56 #851)
-#565 := (= #562 0::int)
-#876 := (<= #54 0::int)
-#577 := (not #876)
-#873 := (>= #54 0::int)
-#863 := (not #873)
-#472 := (or #863 #577)
-#559 := (ite #472 #543 #565)
-#713 := (not #530)
-#604 := [hypothesis]: #713
-#606 := (or #530 #559)
-#842 := (ite #530 #870 #559)
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#29 := (f4 #8 #9)
-#887 := (pattern #29)
-#89 := (* -1::int #9)
-#86 := (* -1::int #8)
-#144 := (mod #86 #89)
-#398 := (+ #29 #144)
-#399 := (= #398 0::int)
-#30 := (mod #8 #9)
-#395 := (* -1::int #30)
-#396 := (+ #29 #395)
-#397 := (= #396 0::int)
-#111 := (<= #9 0::int)
-#107 := (<= #8 0::int)
-#337 := (or #107 #111)
-#338 := (not #337)
-#118 := (>= #8 0::int)
-#320 := (or #111 #118)
-#321 := (not #320)
-#344 := (or #321 #338)
-#400 := (ite #344 #397 #399)
-#394 := (= #29 0::int)
-#12 := (= #8 0::int)
-#401 := (ite #12 #394 #400)
-#393 := (= #8 #29)
-#13 := (= #9 0::int)
-#402 := (ite #13 #393 #401)
-#888 := (forall (vars (?v0 int) (?v1 int)) (:pat #887) #402)
-#405 := (forall (vars (?v0 int) (?v1 int)) #402)
-#891 := (iff #405 #888)
-#889 := (iff #402 #402)
-#890 := [refl]: #889
-#892 := [quant-intro #890]: #891
-#150 := (* -1::int #144)
-#362 := (ite #344 #30 #150)
-#365 := (ite #12 0::int #362)
-#368 := (ite #13 #8 #365)
-#371 := (= #29 #368)
-#374 := (forall (vars (?v0 int) (?v1 int)) #371)
-#406 := (iff #374 #405)
-#403 := (iff #371 #402)
-#404 := [rewrite]: #403
-#407 := [quant-intro #404]: #406
-#119 := (not #118)
-#112 := (not #111)
-#122 := (and #112 #119)
-#108 := (not #107)
-#115 := (and #108 #112)
-#125 := (or #115 #122)
-#170 := (ite #125 #30 #150)
-#173 := (ite #12 0::int #170)
-#176 := (ite #13 #8 #173)
-#179 := (= #29 #176)
-#182 := (forall (vars (?v0 int) (?v1 int)) #179)
-#375 := (iff #182 #374)
-#372 := (iff #179 #371)
-#369 := (= #176 #368)
-#366 := (= #173 #365)
-#363 := (= #170 #362)
-#347 := (iff #125 #344)
-#341 := (or #338 #321)
-#345 := (iff #341 #344)
-#346 := [rewrite]: #345
-#342 := (iff #125 #341)
-#339 := (iff #122 #321)
-#340 := [rewrite]: #339
-#318 := (iff #115 #338)
-#319 := [rewrite]: #318
-#343 := [monotonicity #319 #340]: #342
-#348 := [trans #343 #346]: #347
-#364 := [monotonicity #348]: #363
-#367 := [monotonicity #364]: #366
-#370 := [monotonicity #367]: #369
-#373 := [monotonicity #370]: #372
-#376 := [quant-intro #373]: #375
-#310 := (~ #182 #182)
-#307 := (~ #179 #179)
-#326 := [refl]: #307
-#311 := [nnf-pos #326]: #310
-#23 := (- #9)
-#22 := (- #8)
-#31 := (mod #22 #23)
-#32 := (- #31)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#33 := (ite #20 #30 #32)
-#34 := (ite #12 0::int #33)
-#35 := (ite #13 #8 #34)
-#36 := (= #29 #35)
-#37 := (forall (vars (?v0 int) (?v1 int)) #36)
-#185 := (iff #37 #182)
-#79 := (and #16 #18)
-#82 := (or #17 #79)
-#155 := (ite #82 #30 #150)
-#158 := (ite #12 0::int #155)
-#161 := (ite #13 #8 #158)
-#164 := (= #29 #161)
-#167 := (forall (vars (?v0 int) (?v1 int)) #164)
-#183 := (iff #167 #182)
-#180 := (iff #164 #179)
-#177 := (= #161 #176)
-#174 := (= #158 #173)
-#171 := (= #155 #170)
-#126 := (iff #82 #125)
-#123 := (iff #79 #122)
-#120 := (iff #18 #119)
-#121 := [rewrite]: #120
-#113 := (iff #16 #112)
-#114 := [rewrite]: #113
-#124 := [monotonicity #114 #121]: #123
-#116 := (iff #17 #115)
-#109 := (iff #15 #108)
-#110 := [rewrite]: #109
-#117 := [monotonicity #110 #114]: #116
-#127 := [monotonicity #117 #124]: #126
-#172 := [monotonicity #127]: #171
-#175 := [monotonicity #172]: #174
-#178 := [monotonicity #175]: #177
-#181 := [monotonicity #178]: #180
-#184 := [quant-intro #181]: #183
-#168 := (iff #37 #167)
-#165 := (iff #36 #164)
-#162 := (= #35 #161)
-#159 := (= #34 #158)
-#156 := (= #33 #155)
-#153 := (= #32 #150)
-#147 := (- #144)
-#151 := (= #147 #150)
-#152 := [rewrite]: #151
-#148 := (= #32 #147)
-#145 := (= #31 #144)
-#90 := (= #23 #89)
-#91 := [rewrite]: #90
-#87 := (= #22 #86)
-#88 := [rewrite]: #87
-#146 := [monotonicity #88 #91]: #145
-#149 := [monotonicity #146]: #148
-#154 := [trans #149 #152]: #153
-#83 := (iff #20 #82)
-#80 := (iff #19 #79)
-#81 := [rewrite]: #80
-#84 := [monotonicity #81]: #83
-#157 := [monotonicity #84 #154]: #156
-#160 := [monotonicity #157]: #159
-#163 := [monotonicity #160]: #162
-#166 := [monotonicity #163]: #165
-#169 := [quant-intro #166]: #168
-#186 := [trans #169 #184]: #185
-#143 := [asserted]: #37
-#187 := [mp #143 #186]: #182
-#327 := [mp~ #187 #311]: #182
-#377 := [mp #327 #376]: #374
-#408 := [mp #377 #407]: #405
-#893 := [mp #408 #892]: #888
-#840 := (not #888)
-#681 := (or #840 #842)
-#558 := (* -1::int 3::int)
-#872 := (mod #537 #558)
-#874 := (+ #56 #872)
-#531 := (= #874 0::int)
-#875 := (<= 3::int 0::int)
-#877 := (or #876 #875)
-#878 := (not #877)
-#879 := (or #875 #873)
-#869 := (not #879)
-#529 := (or #869 #878)
-#534 := (ite #529 #543 #531)
-#871 := (ite #530 #870 #534)
-#855 := (= 3::int 0::int)
-#518 := (ite #855 #514 #871)
-#682 := (or #840 #518)
-#825 := (iff #682 #681)
-#827 := (iff #681 #681)
-#828 := [rewrite]: #827
-#839 := (iff #518 #842)
-#836 := (ite false #514 #842)
-#833 := (iff #836 #842)
-#838 := [rewrite]: #833
-#837 := (iff #518 #836)
-#846 := (iff #871 #842)
-#841 := (iff #534 #559)
-#843 := (iff #531 #565)
-#563 := (= #874 #562)
-#848 := (= #872 #851)
-#573 := (= #558 -3::int)
-#847 := [rewrite]: #573
-#852 := [monotonicity #847]: #848
-#564 := [monotonicity #852]: #563
-#844 := [monotonicity #564]: #843
-#849 := (iff #529 #472)
-#578 := (iff #878 #577)
-#854 := (iff #877 #876)
-#506 := (or #876 false)
-#867 := (iff #506 #876)
-#853 := [rewrite]: #867
-#866 := (iff #877 #506)
-#521 := (iff #875 false)
-#857 := [rewrite]: #521
-#864 := [monotonicity #857]: #866
-#576 := [trans #864 #853]: #854
-#579 := [monotonicity #576]: #578
-#865 := (iff #869 #863)
-#862 := (iff #879 #873)
-#858 := (or false #873)
-#861 := (iff #858 #873)
-#856 := [rewrite]: #861
-#859 := (iff #879 #858)
-#860 := [monotonicity #857]: #859
-#500 := [trans #860 #856]: #862
-#505 := [monotonicity #500]: #865
-#850 := [monotonicity #505 #579]: #849
-#845 := [monotonicity #850 #844]: #841
-#835 := [monotonicity #845]: #846
-#519 := (iff #855 false)
-#520 := [rewrite]: #519
-#832 := [monotonicity #520 #835]: #837
-#834 := [trans #832 #838]: #839
-#826 := [monotonicity #834]: #825
-#822 := [trans #826 #828]: #825
-#683 := [quant-inst]: #682
-#823 := [mp #683 #822]: #681
-#605 := [unit-resolution #823 #893]: #842
-#698 := (not #842)
-#709 := (or #698 #530 #559)
-#717 := [def-axiom]: #709
-#607 := [unit-resolution #717 #605]: #606
-#608 := [unit-resolution #607 #604]: #559
-#811 := (f5 #54)
-#779 := (f6 #811)
-#781 := (= #779 0::int)
-#592 := (not #781)
-#594 := (iff #713 #592)
-#603 := (iff #530 #781)
-#613 := (iff #781 #530)
-#611 := (= #779 #54)
-#609 := (= #811 f7)
-#815 := (= f7 #811)
-#38 := (:var 0 S2)
-#39 := (f6 #38)
-#894 := (pattern #39)
-#40 := (f5 #39)
-#189 := (= #38 #40)
-#895 := (forall (vars (?v0 S2)) (:pat #894) #189)
-#192 := (forall (vars (?v0 S2)) #189)
-#896 := (iff #192 #895)
-#898 := (iff #895 #895)
-#899 := [rewrite]: #898
-#897 := [rewrite]: #896
-#900 := [trans #897 #899]: #896
-#312 := (~ #192 #192)
-#328 := (~ #189 #189)
-#329 := [refl]: #328
-#313 := [nnf-pos #329]: #312
-#41 := (= #40 #38)
-#42 := (forall (vars (?v0 S2)) #41)
-#193 := (iff #42 #192)
-#190 := (iff #41 #189)
-#191 := [rewrite]: #190
-#194 := [quant-intro #191]: #193
-#188 := [asserted]: #42
-#197 := [mp #188 #194]: #192
-#330 := [mp~ #197 #313]: #192
-#901 := [mp #330 #900]: #895
-#796 := (not #895)
-#793 := (or #796 #815)
-#795 := [quant-inst]: #793
-#685 := [unit-resolution #795 #901]: #815
-#610 := [symm #685]: #609
-#612 := [monotonicity #610]: #611
-#614 := [monotonicity #612]: #613
-#615 := [symm #614]: #603
-#595 := [monotonicity #615]: #594
-#596 := [mp #604 #595]: #592
-#768 := (or #781 #873)
-#44 := (f5 #9)
-#902 := (pattern #44)
-#212 := (>= #9 0::int)
-#45 := (f6 #44)
-#50 := (= #45 0::int)
-#262 := (or #50 #212)
-#909 := (forall (vars (?v0 int)) (:pat #902) #262)
-#267 := (forall (vars (?v0 int)) #262)
-#912 := (iff #267 #909)
-#910 := (iff #262 #262)
-#911 := [refl]: #910
-#913 := [quant-intro #911]: #912
-#316 := (~ #267 #267)
-#334 := (~ #262 #262)
-#335 := [refl]: #334
-#317 := [nnf-pos #335]: #316
-#49 := (< #9 0::int)
-#51 := (implies #49 #50)
-#52 := (forall (vars (?v0 int)) #51)
-#270 := (iff #52 #267)
-#233 := (= 0::int #45)
-#239 := (not #49)
-#240 := (or #239 #233)
-#245 := (forall (vars (?v0 int)) #240)
-#268 := (iff #245 #267)
-#265 := (iff #240 #262)
-#259 := (or #212 #50)
-#263 := (iff #259 #262)
-#264 := [rewrite]: #263
-#260 := (iff #240 #259)
-#257 := (iff #233 #50)
-#258 := [rewrite]: #257
-#255 := (iff #239 #212)
-#215 := (not #212)
-#250 := (not #215)
-#253 := (iff #250 #212)
-#254 := [rewrite]: #253
-#251 := (iff #239 #250)
-#248 := (iff #49 #215)
-#249 := [rewrite]: #248
-#252 := [monotonicity #249]: #251
-#256 := [trans #252 #254]: #255
-#261 := [monotonicity #256 #258]: #260
-#266 := [trans #261 #264]: #265
-#269 := [quant-intro #266]: #268
-#246 := (iff #52 #245)
-#243 := (iff #51 #240)
-#236 := (implies #49 #233)
-#241 := (iff #236 #240)
-#242 := [rewrite]: #241
-#237 := (iff #51 #236)
-#234 := (iff #50 #233)
-#235 := [rewrite]: #234
-#238 := [monotonicity #235]: #237
-#244 := [trans #238 #242]: #243
-#247 := [quant-intro #244]: #246
-#271 := [trans #247 #269]: #270
-#232 := [asserted]: #52
-#272 := [mp #232 #271]: #267
-#336 := [mp~ #272 #317]: #267
-#914 := [mp #336 #913]: #909
-#782 := (not #909)
-#771 := (or #782 #781 #873)
-#772 := (or #782 #768)
-#774 := (iff #772 #771)
-#775 := [rewrite]: #774
-#773 := [quant-inst]: #772
-#776 := [mp #773 #775]: #771
-#597 := [unit-resolution #776 #914]: #768
-#598 := [unit-resolution #597 #596]: #873
-#599 := (or #530 #577 #863)
-#600 := [th-lemma]: #599
-#593 := [unit-resolution #600 #598 #604]: #577
-#824 := (or #472 #876)
-#831 := [def-axiom]: #824
-#601 := [unit-resolution #831 #593]: #472
-#660 := (not #472)
-#818 := (not #559)
-#821 := (or #818 #660 #543)
-#703 := [def-axiom]: #821
-#586 := [unit-resolution #703 #601 #608]: #543
-#664 := (not #543)
-#587 := (or #664 #817)
-#588 := [th-lemma]: #587
-#590 := [unit-resolution #588 #586]: #817
-#591 := (not #729)
-#589 := (not #817)
-#580 := (or #802 #589 #591)
-#581 := [th-lemma]: #580
-#582 := [unit-resolution #581 #590 #602]: #802
-#816 := (<= #542 0::int)
-#574 := (or #664 #816)
-#583 := [th-lemma]: #574
-#584 := [unit-resolution #583 #586]: #816
-#804 := (not #802)
-#649 := (not #816)
-#626 := (or #627 #286 #649 #804)
-#625 := [hypothesis]: #802
-#747 := (* -3::int #769)
-#750 := (+ #868 #747)
-#751 := (+ #54 #750)
-#727 := (<= #751 0::int)
-#746 := (= #751 0::int)
-#657 := (or false #746)
-#695 := [th-lemma]: #657
-#680 := [unit-resolution #695 #76]: #746
-#658 := (not #746)
-#665 := (or #658 #727)
-#667 := [th-lemma]: #665
-#668 := [unit-resolution #667 #680]: #727
-#677 := [hypothesis]: #816
-#633 := [hypothesis]: #620
-#624 := [th-lemma #632 #633 #677 #668 #625]: false
-#629 := [lemma #624]: #626
-#575 := [unit-resolution #629 #584 #632 #582]: #627
-#666 := (>= #751 0::int)
-#640 := (or #658 #666)
-#636 := [th-lemma]: #640
-#641 := [unit-resolution #636 #680]: #666
-#711 := (* -1::int #56)
-#712 := (+ #54 #711)
-#722 := (<= #712 0::int)
-#560 := (not #722)
-#555 := (not #274)
-#557 := (or #555 #286)
-#299 := (iff #274 #286)
-#59 := (< #54 3::int)
-#60 := (not #59)
-#61 := (iff #58 #60)
-#304 := (iff #61 #299)
-#280 := (iff #60 #274)
-#302 := (iff #280 #299)
-#296 := (iff #286 #274)
-#300 := (iff #296 #299)
-#301 := [rewrite]: #300
-#297 := (iff #280 #296)
-#294 := (iff #60 #286)
-#289 := (not #285)
-#292 := (iff #289 #286)
-#293 := [rewrite]: #292
-#290 := (iff #60 #289)
-#287 := (iff #59 #285)
-#288 := [rewrite]: #287
-#291 := [monotonicity #288]: #290
-#295 := [trans #291 #293]: #294
-#298 := [monotonicity #295]: #297
-#303 := [trans #298 #301]: #302
-#283 := (iff #61 #280)
-#277 := (iff #274 #60)
-#281 := (iff #277 #280)
-#282 := [rewrite]: #281
-#278 := (iff #61 #277)
-#275 := (iff #58 #274)
-#276 := [rewrite]: #275
-#279 := [monotonicity #276]: #278
-#284 := [trans #279 #282]: #283
-#305 := [trans #284 #303]: #304
-#273 := [asserted]: #61
-#306 := [mp #273 #305]: #299
-#466 := (not #299)
-#556 := (or #555 #286 #466)
-#484 := [def-axiom]: #556
-#554 := [unit-resolution #484 #306]: #557
-#585 := [unit-resolution #554 #632]: #555
-#693 := (or #692 #274)
-#688 := (= #811 #57)
-#686 := (= #57 #811)
-#678 := [hypothesis]: #514
-#684 := [symm #678]: #679
-#687 := [monotonicity #684]: #686
-#689 := [symm #687]: #688
-#690 := [trans #685 #689]: #274
-#723 := [hypothesis]: #555
-#691 := [unit-resolution #723 #690]: false
-#694 := [lemma #691]: #693
-#566 := [unit-resolution #694 #585]: #692
-#707 := (>= #712 0::int)
-#669 := (* -1::int #769)
-#671 := (+ #868 #669)
-#637 := (+ #54 #671)
-#639 := (>= #637 0::int)
-#621 := (or #639 #876)
-#645 := (not #639)
-#634 := [hypothesis]: #645
-#706 := (>= #661 3::int)
-#730 := (not #706)
-#675 := (or false #730)
-#676 := [th-lemma]: #675
-#659 := [unit-resolution #676 #76]: #730
-#696 := [hypothesis]: #577
-#631 := [th-lemma #696 #659 #641 #634]: false
-#623 := [lemma #631]: #621
-#567 := [unit-resolution #623 #593]: #639
-#643 := (or #645 #649 #707)
-#642 := (not #707)
-#644 := [hypothesis]: #642
-#635 := [hypothesis]: #639
-#638 := [th-lemma #677 #635 #668 #644]: false
-#646 := [lemma #638]: #643
-#568 := [unit-resolution #646 #567 #584]: #707
-#569 := (or #514 #560 #642)
-#570 := [th-lemma]: #569
-#561 := [unit-resolution #570 #568 #566]: #560
-#571 := [th-lemma #590 #561 #641 #575]: false
-#540 := [lemma #571]: #538
-#697 := [unit-resolution #540 #632]: #530
-#532 := [monotonicity #697]: #541
-#548 := [trans #532 #546]: #547
-#539 := [symm #548]: #549
-#699 := (or #713 #870)
-#715 := (or #698 #713 #870)
-#716 := [def-axiom]: #715
-#701 := [unit-resolution #716 #605]: #699
-#728 := [unit-resolution #701 #697]: #870
-#550 := [mp #728 #539]: #514
-#533 := [unit-resolution #566 #550]: false
-#535 := [lemma #533]: #286
-#553 := (or #274 #285)
-#551 := (or #274 #285 #466)
-#552 := [def-axiom]: #551
-#544 := [unit-resolution #552 #306]: #553
-#496 := [unit-resolution #544 #535]: #274
-#498 := [symm #496]: #58
-#499 := [monotonicity #498]: #494
-#485 := [monotonicity #499]: #481
-#491 := [trans #485 #489]: #490
-#492 := [symm #491]: #483
-#474 := [monotonicity #492]: #473
-#524 := (or #577 #285)
-#525 := [th-lemma]: #524
-#526 := [unit-resolution #525 #535]: #577
-#527 := (or #713 #876)
-#516 := [th-lemma]: #527
-#528 := [unit-resolution #516 #526]: #713
-#509 := [unit-resolution #607 #528]: #559
-#511 := [unit-resolution #831 #526]: #472
-#512 := [unit-resolution #703 #511 #509]: #543
-#502 := [unit-resolution #583 #512]: #816
-#508 := (or #560 #285 #706 #649)
-#482 := [th-lemma]: #508
-#487 := [unit-resolution #482 #502 #535 #659]: #560
-#488 := (or #692 #722)
-#493 := [th-lemma]: #488
-#495 := [unit-resolution #493 #487]: #692
-#477 := [mp #495 #474]: #476
-#510 := [unit-resolution #588 #512]: #817
-#513 := (or #802 #591 #589)
-#501 := [th-lemma]: #513
-#503 := [unit-resolution #501 #510 #602]: #802
-#803 := (or #804 #806)
-#196 := (= #9 #45)
-#221 := (or #196 #215)
-#903 := (forall (vars (?v0 int)) (:pat #902) #221)
-#226 := (forall (vars (?v0 int)) #221)
-#906 := (iff #226 #903)
-#904 := (iff #221 #221)
-#905 := [refl]: #904
-#907 := [quant-intro #905]: #906
-#314 := (~ #226 #226)
-#331 := (~ #221 #221)
-#332 := [refl]: #331
-#315 := [nnf-pos #332]: #314
-#46 := (= #45 #9)
-#43 := (<= 0::int #9)
-#47 := (implies #43 #46)
-#48 := (forall (vars (?v0 int)) #47)
-#229 := (iff #48 #226)
-#203 := (not #43)
-#204 := (or #203 #196)
-#209 := (forall (vars (?v0 int)) #204)
-#227 := (iff #209 #226)
-#224 := (iff #204 #221)
-#218 := (or #215 #196)
-#222 := (iff #218 #221)
-#223 := [rewrite]: #222
-#219 := (iff #204 #218)
-#216 := (iff #203 #215)
-#213 := (iff #43 #212)
-#214 := [rewrite]: #213
-#217 := [monotonicity #214]: #216
-#220 := [monotonicity #217]: #219
-#225 := [trans #220 #223]: #224
-#228 := [quant-intro #225]: #227
-#210 := (iff #48 #209)
-#207 := (iff #47 #204)
-#200 := (implies #43 #196)
-#205 := (iff #200 #204)
-#206 := [rewrite]: #205
-#201 := (iff #47 #200)
-#198 := (iff #46 #196)
-#199 := [rewrite]: #198
-#202 := [monotonicity #199]: #201
-#208 := [trans #202 #206]: #207
-#211 := [quant-intro #208]: #210
-#230 := [trans #211 #228]: #229
-#195 := [asserted]: #48
-#231 := [mp #195 #230]: #226
-#333 := [mp~ #231 #315]: #226
-#908 := [mp #333 #907]: #903
-#798 := (not #903)
-#799 := (or #798 #804 #806)
-#807 := (or #806 #804)
-#800 := (or #798 #807)
-#790 := (iff #800 #799)
-#801 := (or #798 #803)
-#788 := (iff #801 #799)
-#789 := [rewrite]: #788
-#785 := (iff #800 #801)
-#808 := (iff #807 #803)
-#797 := [rewrite]: #808
-#786 := [monotonicity #797]: #785
-#791 := [trans #786 #789]: #790
-#794 := [quant-inst]: #800
-#787 := [mp #794 #791]: #799
-#504 := [unit-resolution #787 #908]: #803
-#507 := [unit-resolution #504 #503]: #806
-[unit-resolution #507 #477]: false
-unsat
-0acf022ca28d30081e73fadc3fdcfea21e99a21c 60 0
+85ac6a29deaf91400dbfbdf638baba5914d49e18 60 0
 #2 := false
 #13 := 0::int
 decl f4 :: (-> S2 int)
@@ -22753,7 +22961,7 @@
 #630 := [unit-resolution #632 #631]: #266
 [th-lemma #159 #629 #630]: false
 unsat
-aa7ead3dd8b4a36efac1c966a70956175fb7e15a 60 0
+ec431e8216f1d8b2f670cdb0885c1e07fa83e950 60 0
 #2 := false
 #13 := 0::int
 decl f4 :: (-> S2 int)
@@ -22814,7 +23022,7 @@
 #630 := [th-lemma]: #632
 [unit-resolution #630 #271 #159]: false
 unsat
-4da37f686a90a43bb4bcf93c60d1755d615df6ac 253 0
+db10b3baed96f5ad894a70d823fdba86b4614481 253 0
 #2 := false
 #13 := 0::int
 decl f4 :: (-> S2 int)
@@ -23068,129 +23276,7 @@
 #473 := [lemma #472]: #233
 [unit-resolution #631 #473 #487]: false
 unsat
-5536a80305cbc5c1c65da302174d985517f94ac4 60 0
-#2 := false
-#13 := 0::int
-decl f4 :: (-> S2 int)
-decl f5 :: S2
-#25 := f5
-#26 := (f4 f5)
-decl f6 :: S2
-#27 := f6
-#28 := (f4 f6)
-#137 := -1::int
-#138 := (* -1::int #28)
-#139 := (+ #26 #138)
-#140 := (<= #139 0::int)
-#143 := (ite #140 #28 #26)
-#149 := (* -1::int #143)
-#150 := (+ #26 #149)
-#151 := (<= #150 0::int)
-#156 := (not #151)
-#29 := (<= #26 #28)
-#30 := (ite #29 #28 #26)
-#31 := (<= #26 #30)
-#32 := (not #31)
-#157 := (iff #32 #156)
-#154 := (iff #31 #151)
-#146 := (<= #26 #143)
-#152 := (iff #146 #151)
-#153 := [rewrite]: #152
-#147 := (iff #31 #146)
-#144 := (= #30 #143)
-#141 := (iff #29 #140)
-#142 := [rewrite]: #141
-#145 := [monotonicity #142]: #144
-#148 := [monotonicity #145]: #147
-#155 := [trans #148 #153]: #154
-#158 := [monotonicity #155]: #157
-#134 := [asserted]: #32
-#159 := [mp #134 #158]: #156
-#317 := (= #26 #143)
-#232 := (not #140)
-#625 := [hypothesis]: #140
-#637 := (+ #28 #149)
-#280 := (<= #637 0::int)
-#231 := (= #28 #143)
-#318 := (or #232 #231)
-#319 := [def-axiom]: #318
-#626 := [unit-resolution #319 #625]: #231
-#627 := (not #231)
-#622 := (or #627 #280)
-#628 := [th-lemma]: #622
-#266 := [unit-resolution #628 #626]: #280
-#629 := [th-lemma #266 #159 #625]: false
-#631 := [lemma #629]: #232
-#310 := (or #140 #317)
-#321 := [def-axiom]: #310
-#271 := [unit-resolution #321 #631]: #317
-#272 := (not #317)
-#632 := (or #272 #151)
-#630 := [th-lemma]: #632
-[unit-resolution #630 #271 #159]: false
-unsat
-1f939e5dc55e38d665a6a7ffafbec37b57ae831f 60 0
-#2 := false
-#13 := 0::int
-decl f4 :: (-> S2 int)
-decl f6 :: S2
-#27 := f6
-#28 := (f4 f6)
-decl f5 :: S2
-#25 := f5
-#26 := (f4 f5)
-#137 := -1::int
-#140 := (* -1::int #28)
-#141 := (+ #26 #140)
-#139 := (>= #141 0::int)
-#143 := (ite #139 #26 #28)
-#149 := (* -1::int #143)
-#637 := (+ #28 #149)
-#628 := (<= #637 0::int)
-#317 := (= #28 #143)
-#232 := (not #139)
-#231 := (= #26 #143)
-#624 := (not #231)
-#150 := (+ #26 #149)
-#151 := (<= #150 0::int)
-#156 := (not #151)
-#29 := (<= #28 #26)
-#30 := (ite #29 #26 #28)
-#31 := (<= #26 #30)
-#32 := (not #31)
-#157 := (iff #32 #156)
-#154 := (iff #31 #151)
-#146 := (<= #26 #143)
-#152 := (iff #146 #151)
-#153 := [rewrite]: #152
-#147 := (iff #31 #146)
-#144 := (= #30 #143)
-#138 := (iff #29 #139)
-#142 := [rewrite]: #138
-#145 := [monotonicity #142]: #144
-#148 := [monotonicity #145]: #147
-#155 := [trans #148 #153]: #154
-#158 := [monotonicity #155]: #157
-#134 := [asserted]: #32
-#159 := [mp #134 #158]: #156
-#280 := [hypothesis]: #231
-#625 := (or #624 #151)
-#626 := [th-lemma]: #625
-#627 := [unit-resolution #626 #280 #159]: false
-#622 := [lemma #627]: #624
-#318 := (or #232 #231)
-#319 := [def-axiom]: #318
-#629 := [unit-resolution #319 #622]: #232
-#310 := (or #139 #317)
-#321 := [def-axiom]: #310
-#631 := [unit-resolution #321 #629]: #317
-#271 := (not #317)
-#272 := (or #271 #628)
-#632 := [th-lemma]: #272
-#630 := [unit-resolution #632 #631]: #628
-[th-lemma #159 #629 #630]: false
-unsat
-d3e83d573ed30cf5c8453245f4134046eb5c6487 103 0
+89bc3a6333da51700f7daac14eb84da9abc113a1 103 0
 #2 := false
 #13 := 0::int
 decl f4 :: (-> S2 int)
@@ -23294,7 +23380,404 @@
 #642 := [unit-resolution #636 #635]: #638
 [th-lemma #642 #193 #194]: false
 unsat
-e8bb28c85cd05d1e4eb92c6e0ebb5900ad2217ca 437 0
+c39739842d64e1dca61912be5a777a90f01e28fd 117 0
+#2 := false
+decl f3 :: (-> int S2)
+decl f4 :: (-> S2 int)
+decl f5 :: S2
+#25 := f5
+#26 := (f4 f5)
+decl f6 :: S2
+#27 := f6
+#28 := (f4 f6)
+#13 := 0::int
+#141 := -1::int
+#142 := (* -1::int #28)
+#143 := (+ #26 #142)
+#154 := (>= #143 0::int)
+#156 := (ite #154 #28 #26)
+#159 := (f3 #156)
+#144 := (<= #143 0::int)
+#147 := (ite #144 #26 #28)
+#150 := (f3 #147)
+#162 := (= #150 #159)
+#649 := (f3 #26)
+#558 := (= #649 #159)
+#556 := (= #159 #649)
+#564 := (= #156 #26)
+#259 := (= #26 #156)
+#436 := (f3 #28)
+#577 := (= #436 #159)
+#586 := (= #159 #436)
+#479 := (= #156 #28)
+#331 := (= #28 #156)
+#489 := (not #259)
+#584 := [hypothesis]: #489
+#312 := (or #154 #259)
+#647 := [def-axiom]: #312
+#588 := [unit-resolution #647 #584]: #154
+#332 := (not #154)
+#329 := (or #332 #331)
+#333 := [def-axiom]: #329
+#490 := [unit-resolution #333 #588]: #331
+#480 := [symm #490]: #479
+#590 := [monotonicity #480]: #586
+#579 := [symm #590]: #577
+#491 := (= #150 #436)
+#492 := (= #147 #28)
+#326 := (= #28 #147)
+#241 := (not #144)
+#496 := (or #241 #259)
+#473 := (= #26 #28)
+#585 := [hypothesis]: #144
+#488 := [th-lemma #588 #585]: #473
+#494 := [trans #488 #490]: #259
+#495 := [unit-resolution #584 #494]: false
+#589 := [lemma #495]: #496
+#439 := [unit-resolution #589 #584]: #241
+#319 := (or #144 #326)
+#330 := [def-axiom]: #319
+#587 := [unit-resolution #330 #439]: #326
+#493 := [symm #587]: #492
+#484 := [monotonicity #493]: #491
+#571 := [trans #484 #579]: #162
+#165 := (not #162)
+#32 := (<= #28 #26)
+#33 := (ite #32 #28 #26)
+#34 := (f3 #33)
+#29 := (<= #26 #28)
+#30 := (ite #29 #26 #28)
+#31 := (f3 #30)
+#35 := (= #31 #34)
+#36 := (not #35)
+#166 := (iff #36 #165)
+#163 := (iff #35 #162)
+#160 := (= #34 #159)
+#157 := (= #33 #156)
+#153 := (iff #32 #154)
+#155 := [rewrite]: #153
+#158 := [monotonicity #155]: #157
+#161 := [monotonicity #158]: #160
+#151 := (= #31 #150)
+#148 := (= #30 #147)
+#145 := (iff #29 #144)
+#146 := [rewrite]: #145
+#149 := [monotonicity #146]: #148
+#152 := [monotonicity #149]: #151
+#164 := [monotonicity #152 #161]: #163
+#167 := [monotonicity #164]: #166
+#138 := [asserted]: #36
+#168 := [mp #138 #167]: #165
+#568 := [unit-resolution #168 #571]: false
+#570 := [lemma #568]: #259
+#565 := [symm #570]: #564
+#557 := [monotonicity #565]: #556
+#555 := [symm #557]: #558
+#553 := (= #150 #649)
+#562 := (= #147 #26)
+#240 := (= #26 #147)
+#435 := [hypothesis]: #241
+#437 := (or #154 #144)
+#596 := [th-lemma]: #437
+#478 := [unit-resolution #596 #435]: #154
+#580 := [unit-resolution #333 #478]: #331
+#581 := [symm #580]: #479
+#572 := [monotonicity #581]: #586
+#573 := [symm #572]: #577
+#582 := [unit-resolution #330 #435]: #326
+#578 := [symm #582]: #492
+#583 := [monotonicity #578]: #491
+#574 := [trans #583 #573]: #162
+#575 := [unit-resolution #168 #574]: false
+#569 := [lemma #575]: #144
+#327 := (or #241 #240)
+#328 := [def-axiom]: #327
+#566 := [unit-resolution #328 #569]: #240
+#567 := [symm #566]: #562
+#554 := [monotonicity #567]: #553
+#559 := [trans #554 #555]: #162
+[unit-resolution #168 #559]: false
+unsat
+c3edc6cce39c9c380e8b07a4426540e2db018197 156 0
+#2 := false
+decl f3 :: (-> int S2)
+#13 := 0::int
+#30 := (f3 0::int)
+decl f4 :: (-> S2 int)
+decl f5 :: S2
+#25 := f5
+#26 := (f4 f5)
+#27 := (<= #26 0::int)
+#28 := (ite #27 #26 0::int)
+#29 := (f3 #28)
+#31 := (= #29 #30)
+#294 := (= #28 0::int)
+#299 := (f3 #26)
+#576 := (f4 #299)
+#577 := (= #576 0::int)
+#442 := (= #26 0::int)
+#567 := (not #294)
+#573 := [hypothesis]: #567
+#287 := (or #27 #294)
+#298 := [def-axiom]: #287
+#404 := [unit-resolution #298 #573]: #27
+#581 := (>= #26 0::int)
+#408 := (not #577)
+#556 := (iff #567 #408)
+#448 := (iff #294 #577)
+#565 := (= #28 #576)
+#564 := (= #26 #576)
+#561 := (= #576 #26)
+#568 := (= #299 f5)
+#227 := (= f5 #299)
+#8 := (:var 0 S2)
+#9 := (f4 #8)
+#624 := (pattern #9)
+#10 := (f3 #9)
+#50 := (= #8 #10)
+#625 := (forall (vars (?v0 S2)) (:pat #624) #50)
+#53 := (forall (vars (?v0 S2)) #50)
+#626 := (iff #53 #625)
+#628 := (iff #625 #625)
+#629 := [rewrite]: #628
+#627 := [rewrite]: #626
+#630 := [trans #627 #629]: #626
+#148 := (~ #53 #53)
+#146 := (~ #50 #50)
+#147 := [refl]: #146
+#149 := [nnf-pos #147]: #148
+#11 := (= #10 #8)
+#12 := (forall (vars (?v0 S2)) #11)
+#54 := (iff #12 #53)
+#51 := (iff #11 #50)
+#52 := [rewrite]: #51
+#55 := [quant-intro #52]: #54
+#49 := [asserted]: #12
+#58 := [mp #49 #55]: #53
+#137 := [mp~ #58 #149]: #53
+#631 := [mp #137 #630]: #625
+#301 := (not #625)
+#280 := (or #301 #227)
+#616 := [quant-inst]: #280
+#574 := [unit-resolution #616 #631]: #227
+#575 := [symm #574]: #568
+#563 := [monotonicity #575]: #561
+#562 := [symm #563]: #564
+#407 := (= #28 #26)
+#208 := (= #26 #28)
+#209 := (not #27)
+#295 := (or #209 #208)
+#296 := [def-axiom]: #295
+#406 := [unit-resolution #296 #404]: #208
+#560 := [symm #406]: #407
+#447 := [trans #560 #562]: #565
+#449 := [monotonicity #447]: #448
+#458 := [monotonicity #449]: #556
+#553 := [mp #573 #458]: #408
+#582 := (or #577 #581)
+#14 := (:var 0 int)
+#16 := (f3 #14)
+#632 := (pattern #16)
+#75 := (>= #14 0::int)
+#17 := (f4 #16)
+#22 := (= #17 0::int)
+#123 := (or #22 #75)
+#639 := (forall (vars (?v0 int)) (:pat #632) #123)
+#128 := (forall (vars (?v0 int)) #123)
+#642 := (iff #128 #639)
+#640 := (iff #123 #123)
+#641 := [refl]: #640
+#643 := [quant-intro #641]: #642
+#140 := (~ #128 #128)
+#152 := (~ #123 #123)
+#153 := [refl]: #152
+#141 := [nnf-pos #153]: #140
+#21 := (< #14 0::int)
+#23 := (implies #21 #22)
+#24 := (forall (vars (?v0 int)) #23)
+#131 := (iff #24 #128)
+#94 := (= 0::int #17)
+#100 := (not #21)
+#101 := (or #100 #94)
+#106 := (forall (vars (?v0 int)) #101)
+#129 := (iff #106 #128)
+#126 := (iff #101 #123)
+#120 := (or #75 #22)
+#124 := (iff #120 #123)
+#125 := [rewrite]: #124
+#121 := (iff #101 #120)
+#118 := (iff #94 #22)
+#119 := [rewrite]: #118
+#116 := (iff #100 #75)
+#76 := (not #75)
+#111 := (not #76)
+#114 := (iff #111 #75)
+#115 := [rewrite]: #114
+#112 := (iff #100 #111)
+#109 := (iff #21 #76)
+#110 := [rewrite]: #109
+#113 := [monotonicity #110]: #112
+#117 := [trans #113 #115]: #116
+#122 := [monotonicity #117 #119]: #121
+#127 := [trans #122 #125]: #126
+#130 := [quant-intro #127]: #129
+#107 := (iff #24 #106)
+#104 := (iff #23 #101)
+#97 := (implies #21 #94)
+#102 := (iff #97 #101)
+#103 := [rewrite]: #102
+#98 := (iff #23 #97)
+#95 := (iff #22 #94)
+#96 := [rewrite]: #95
+#99 := [monotonicity #96]: #98
+#105 := [trans #99 #103]: #104
+#108 := [quant-intro #105]: #107
+#132 := [trans #108 #130]: #131
+#93 := [asserted]: #24
+#133 := [mp #93 #132]: #128
+#154 := [mp~ #133 #141]: #128
+#644 := [mp #154 #643]: #639
+#614 := (not #639)
+#584 := (or #614 #577 #581)
+#425 := (or #614 #582)
+#427 := (iff #425 #584)
+#569 := [rewrite]: #427
+#426 := [quant-inst]: #425
+#570 := [mp #426 #569]: #584
+#554 := [unit-resolution #570 #644]: #582
+#557 := [unit-resolution #554 #553]: #581
+#457 := [th-lemma #557 #404]: #442
+#459 := [trans #563 #457]: #577
+#460 := [unit-resolution #553 #459]: false
+#453 := [lemma #460]: #294
+#583 := [monotonicity #453]: #31
+#32 := (not #31)
+#134 := [asserted]: #32
+[unit-resolution #134 #583]: false
+unsat
+d84df414dbdb37289879fadf547b4a4bc9c68951 60 0
+#2 := false
+#13 := 0::int
+decl f4 :: (-> S2 int)
+decl f5 :: S2
+#25 := f5
+#26 := (f4 f5)
+decl f6 :: S2
+#27 := f6
+#28 := (f4 f6)
+#137 := -1::int
+#138 := (* -1::int #28)
+#139 := (+ #26 #138)
+#140 := (<= #139 0::int)
+#143 := (ite #140 #28 #26)
+#149 := (* -1::int #143)
+#150 := (+ #26 #149)
+#151 := (<= #150 0::int)
+#156 := (not #151)
+#29 := (<= #26 #28)
+#30 := (ite #29 #28 #26)
+#31 := (<= #26 #30)
+#32 := (not #31)
+#157 := (iff #32 #156)
+#154 := (iff #31 #151)
+#146 := (<= #26 #143)
+#152 := (iff #146 #151)
+#153 := [rewrite]: #152
+#147 := (iff #31 #146)
+#144 := (= #30 #143)
+#141 := (iff #29 #140)
+#142 := [rewrite]: #141
+#145 := [monotonicity #142]: #144
+#148 := [monotonicity #145]: #147
+#155 := [trans #148 #153]: #154
+#158 := [monotonicity #155]: #157
+#134 := [asserted]: #32
+#159 := [mp #134 #158]: #156
+#317 := (= #26 #143)
+#232 := (not #140)
+#625 := [hypothesis]: #140
+#637 := (+ #28 #149)
+#280 := (<= #637 0::int)
+#231 := (= #28 #143)
+#318 := (or #232 #231)
+#319 := [def-axiom]: #318
+#626 := [unit-resolution #319 #625]: #231
+#627 := (not #231)
+#622 := (or #627 #280)
+#628 := [th-lemma]: #622
+#266 := [unit-resolution #628 #626]: #280
+#629 := [th-lemma #266 #159 #625]: false
+#631 := [lemma #629]: #232
+#310 := (or #140 #317)
+#321 := [def-axiom]: #310
+#271 := [unit-resolution #321 #631]: #317
+#272 := (not #317)
+#632 := (or #272 #151)
+#630 := [th-lemma]: #632
+[unit-resolution #630 #271 #159]: false
+unsat
+2abc216904909a8ab09db7d8ddd5116874c405b2 60 0
+#2 := false
+#13 := 0::int
+decl f4 :: (-> S2 int)
+decl f6 :: S2
+#27 := f6
+#28 := (f4 f6)
+decl f5 :: S2
+#25 := f5
+#26 := (f4 f5)
+#137 := -1::int
+#140 := (* -1::int #28)
+#141 := (+ #26 #140)
+#139 := (>= #141 0::int)
+#143 := (ite #139 #26 #28)
+#149 := (* -1::int #143)
+#637 := (+ #28 #149)
+#628 := (<= #637 0::int)
+#317 := (= #28 #143)
+#232 := (not #139)
+#231 := (= #26 #143)
+#624 := (not #231)
+#150 := (+ #26 #149)
+#151 := (<= #150 0::int)
+#156 := (not #151)
+#29 := (<= #28 #26)
+#30 := (ite #29 #26 #28)
+#31 := (<= #26 #30)
+#32 := (not #31)
+#157 := (iff #32 #156)
+#154 := (iff #31 #151)
+#146 := (<= #26 #143)
+#152 := (iff #146 #151)
+#153 := [rewrite]: #152
+#147 := (iff #31 #146)
+#144 := (= #30 #143)
+#138 := (iff #29 #139)
+#142 := [rewrite]: #138
+#145 := [monotonicity #142]: #144
+#148 := [monotonicity #145]: #147
+#155 := [trans #148 #153]: #154
+#158 := [monotonicity #155]: #157
+#134 := [asserted]: #32
+#159 := [mp #134 #158]: #156
+#280 := [hypothesis]: #231
+#625 := (or #624 #151)
+#626 := [th-lemma]: #625
+#627 := [unit-resolution #626 #280 #159]: false
+#622 := [lemma #627]: #624
+#318 := (or #232 #231)
+#319 := [def-axiom]: #318
+#629 := [unit-resolution #319 #622]: #232
+#310 := (or #139 #317)
+#321 := [def-axiom]: #310
+#631 := [unit-resolution #321 #629]: #317
+#271 := (not #317)
+#272 := (or #271 #628)
+#632 := [th-lemma]: #272
+#630 := [unit-resolution #632 #631]: #628
+[th-lemma #159 #629 #630]: false
+unsat
+14969492cac173e86b554790b2b1280f30825e8b 437 0
 #2 := false
 #13 := 0::int
 decl f4 :: (-> S2 int)
@@ -23732,282 +24215,7 @@
 #462 := [unit-resolution #492 #479]: #669
 [th-lemma #462 #528 #194 #475 #461]: false
 unsat
-ddf4ee3ffc7523dd708ffeab45ca4db621550c07 117 0
-#2 := false
-decl f3 :: (-> int S2)
-decl f4 :: (-> S2 int)
-decl f5 :: S2
-#25 := f5
-#26 := (f4 f5)
-decl f6 :: S2
-#27 := f6
-#28 := (f4 f6)
-#13 := 0::int
-#141 := -1::int
-#142 := (* -1::int #28)
-#143 := (+ #26 #142)
-#154 := (>= #143 0::int)
-#156 := (ite #154 #28 #26)
-#159 := (f3 #156)
-#144 := (<= #143 0::int)
-#147 := (ite #144 #26 #28)
-#150 := (f3 #147)
-#162 := (= #150 #159)
-#649 := (f3 #26)
-#558 := (= #649 #159)
-#556 := (= #159 #649)
-#564 := (= #156 #26)
-#259 := (= #26 #156)
-#436 := (f3 #28)
-#577 := (= #436 #159)
-#586 := (= #159 #436)
-#479 := (= #156 #28)
-#331 := (= #28 #156)
-#489 := (not #259)
-#584 := [hypothesis]: #489
-#312 := (or #154 #259)
-#647 := [def-axiom]: #312
-#588 := [unit-resolution #647 #584]: #154
-#332 := (not #154)
-#329 := (or #332 #331)
-#333 := [def-axiom]: #329
-#490 := [unit-resolution #333 #588]: #331
-#480 := [symm #490]: #479
-#590 := [monotonicity #480]: #586
-#579 := [symm #590]: #577
-#491 := (= #150 #436)
-#492 := (= #147 #28)
-#326 := (= #28 #147)
-#241 := (not #144)
-#496 := (or #241 #259)
-#473 := (= #26 #28)
-#585 := [hypothesis]: #144
-#488 := [th-lemma #588 #585]: #473
-#494 := [trans #488 #490]: #259
-#495 := [unit-resolution #584 #494]: false
-#589 := [lemma #495]: #496
-#439 := [unit-resolution #589 #584]: #241
-#319 := (or #144 #326)
-#330 := [def-axiom]: #319
-#587 := [unit-resolution #330 #439]: #326
-#493 := [symm #587]: #492
-#484 := [monotonicity #493]: #491
-#571 := [trans #484 #579]: #162
-#165 := (not #162)
-#32 := (<= #28 #26)
-#33 := (ite #32 #28 #26)
-#34 := (f3 #33)
-#29 := (<= #26 #28)
-#30 := (ite #29 #26 #28)
-#31 := (f3 #30)
-#35 := (= #31 #34)
-#36 := (not #35)
-#166 := (iff #36 #165)
-#163 := (iff #35 #162)
-#160 := (= #34 #159)
-#157 := (= #33 #156)
-#153 := (iff #32 #154)
-#155 := [rewrite]: #153
-#158 := [monotonicity #155]: #157
-#161 := [monotonicity #158]: #160
-#151 := (= #31 #150)
-#148 := (= #30 #147)
-#145 := (iff #29 #144)
-#146 := [rewrite]: #145
-#149 := [monotonicity #146]: #148
-#152 := [monotonicity #149]: #151
-#164 := [monotonicity #152 #161]: #163
-#167 := [monotonicity #164]: #166
-#138 := [asserted]: #36
-#168 := [mp #138 #167]: #165
-#568 := [unit-resolution #168 #571]: false
-#570 := [lemma #568]: #259
-#565 := [symm #570]: #564
-#557 := [monotonicity #565]: #556
-#555 := [symm #557]: #558
-#553 := (= #150 #649)
-#562 := (= #147 #26)
-#240 := (= #26 #147)
-#435 := [hypothesis]: #241
-#437 := (or #154 #144)
-#596 := [th-lemma]: #437
-#478 := [unit-resolution #596 #435]: #154
-#580 := [unit-resolution #333 #478]: #331
-#581 := [symm #580]: #479
-#572 := [monotonicity #581]: #586
-#573 := [symm #572]: #577
-#582 := [unit-resolution #330 #435]: #326
-#578 := [symm #582]: #492
-#583 := [monotonicity #578]: #491
-#574 := [trans #583 #573]: #162
-#575 := [unit-resolution #168 #574]: false
-#569 := [lemma #575]: #144
-#327 := (or #241 #240)
-#328 := [def-axiom]: #327
-#566 := [unit-resolution #328 #569]: #240
-#567 := [symm #566]: #562
-#554 := [monotonicity #567]: #553
-#559 := [trans #554 #555]: #162
-[unit-resolution #168 #559]: false
-unsat
-12109499f925774016387c1f6f51980ab6ede03b 156 0
-#2 := false
-decl f3 :: (-> int S2)
-#13 := 0::int
-#30 := (f3 0::int)
-decl f4 :: (-> S2 int)
-decl f5 :: S2
-#25 := f5
-#26 := (f4 f5)
-#27 := (<= #26 0::int)
-#28 := (ite #27 #26 0::int)
-#29 := (f3 #28)
-#31 := (= #29 #30)
-#294 := (= #28 0::int)
-#299 := (f3 #26)
-#576 := (f4 #299)
-#577 := (= #576 0::int)
-#442 := (= #26 0::int)
-#567 := (not #294)
-#573 := [hypothesis]: #567
-#287 := (or #27 #294)
-#298 := [def-axiom]: #287
-#404 := [unit-resolution #298 #573]: #27
-#581 := (>= #26 0::int)
-#408 := (not #577)
-#556 := (iff #567 #408)
-#448 := (iff #294 #577)
-#565 := (= #28 #576)
-#564 := (= #26 #576)
-#561 := (= #576 #26)
-#568 := (= #299 f5)
-#227 := (= f5 #299)
-#8 := (:var 0 S2)
-#9 := (f4 #8)
-#624 := (pattern #9)
-#10 := (f3 #9)
-#50 := (= #8 #10)
-#625 := (forall (vars (?v0 S2)) (:pat #624) #50)
-#53 := (forall (vars (?v0 S2)) #50)
-#626 := (iff #53 #625)
-#628 := (iff #625 #625)
-#629 := [rewrite]: #628
-#627 := [rewrite]: #626
-#630 := [trans #627 #629]: #626
-#148 := (~ #53 #53)
-#146 := (~ #50 #50)
-#147 := [refl]: #146
-#149 := [nnf-pos #147]: #148
-#11 := (= #10 #8)
-#12 := (forall (vars (?v0 S2)) #11)
-#54 := (iff #12 #53)
-#51 := (iff #11 #50)
-#52 := [rewrite]: #51
-#55 := [quant-intro #52]: #54
-#49 := [asserted]: #12
-#58 := [mp #49 #55]: #53
-#137 := [mp~ #58 #149]: #53
-#631 := [mp #137 #630]: #625
-#301 := (not #625)
-#280 := (or #301 #227)
-#616 := [quant-inst]: #280
-#574 := [unit-resolution #616 #631]: #227
-#575 := [symm #574]: #568
-#563 := [monotonicity #575]: #561
-#562 := [symm #563]: #564
-#407 := (= #28 #26)
-#208 := (= #26 #28)
-#209 := (not #27)
-#295 := (or #209 #208)
-#296 := [def-axiom]: #295
-#406 := [unit-resolution #296 #404]: #208
-#560 := [symm #406]: #407
-#447 := [trans #560 #562]: #565
-#449 := [monotonicity #447]: #448
-#458 := [monotonicity #449]: #556
-#553 := [mp #573 #458]: #408
-#582 := (or #577 #581)
-#14 := (:var 0 int)
-#16 := (f3 #14)
-#632 := (pattern #16)
-#75 := (>= #14 0::int)
-#17 := (f4 #16)
-#22 := (= #17 0::int)
-#123 := (or #22 #75)
-#639 := (forall (vars (?v0 int)) (:pat #632) #123)
-#128 := (forall (vars (?v0 int)) #123)
-#642 := (iff #128 #639)
-#640 := (iff #123 #123)
-#641 := [refl]: #640
-#643 := [quant-intro #641]: #642
-#140 := (~ #128 #128)
-#152 := (~ #123 #123)
-#153 := [refl]: #152
-#141 := [nnf-pos #153]: #140
-#21 := (< #14 0::int)
-#23 := (implies #21 #22)
-#24 := (forall (vars (?v0 int)) #23)
-#131 := (iff #24 #128)
-#94 := (= 0::int #17)
-#100 := (not #21)
-#101 := (or #100 #94)
-#106 := (forall (vars (?v0 int)) #101)
-#129 := (iff #106 #128)
-#126 := (iff #101 #123)
-#120 := (or #75 #22)
-#124 := (iff #120 #123)
-#125 := [rewrite]: #124
-#121 := (iff #101 #120)
-#118 := (iff #94 #22)
-#119 := [rewrite]: #118
-#116 := (iff #100 #75)
-#76 := (not #75)
-#111 := (not #76)
-#114 := (iff #111 #75)
-#115 := [rewrite]: #114
-#112 := (iff #100 #111)
-#109 := (iff #21 #76)
-#110 := [rewrite]: #109
-#113 := [monotonicity #110]: #112
-#117 := [trans #113 #115]: #116
-#122 := [monotonicity #117 #119]: #121
-#127 := [trans #122 #125]: #126
-#130 := [quant-intro #127]: #129
-#107 := (iff #24 #106)
-#104 := (iff #23 #101)
-#97 := (implies #21 #94)
-#102 := (iff #97 #101)
-#103 := [rewrite]: #102
-#98 := (iff #23 #97)
-#95 := (iff #22 #94)
-#96 := [rewrite]: #95
-#99 := [monotonicity #96]: #98
-#105 := [trans #99 #103]: #104
-#108 := [quant-intro #105]: #107
-#132 := [trans #108 #130]: #131
-#93 := [asserted]: #24
-#133 := [mp #93 #132]: #128
-#154 := [mp~ #133 #141]: #128
-#644 := [mp #154 #643]: #639
-#614 := (not #639)
-#584 := (or #614 #577 #581)
-#425 := (or #614 #582)
-#427 := (iff #425 #584)
-#569 := [rewrite]: #427
-#426 := [quant-inst]: #425
-#570 := [mp #426 #569]: #584
-#554 := [unit-resolution #570 #644]: #582
-#557 := [unit-resolution #554 #553]: #581
-#457 := [th-lemma #557 #404]: #442
-#459 := [trans #563 #457]: #577
-#460 := [unit-resolution #553 #459]: false
-#453 := [lemma #460]: #294
-#583 := [monotonicity #453]: #31
-#32 := (not #31)
-#134 := [asserted]: #32
-[unit-resolution #134 #583]: false
-unsat
-1a0992c511947512c28712c25329faef3db2a506 103 0
+452992a55139d640bc8969517cd3166d0f13bb37 103 0
 #2 := false
 #13 := 0::int
 decl f4 :: (-> S2 int)
@@ -24111,7 +24319,7 @@
 #642 := [unit-resolution #636 #635]: #639
 [th-lemma #642 #195 #194]: false
 unsat
-2cff16037fa28fd9b7fe01fda3fea3217a638862 119 0
+c2cc039657273f4c31549e108b87c8f33092d45f 119 0
 #2 := false
 decl f3 :: (-> int S2)
 decl f4 :: (-> S2 int)
@@ -24231,7 +24439,7 @@
 #540 := [trans #557 #539]: #162
 [unit-resolution #168 #540]: false
 unsat
-9e66dd3e51867d9c417bdcc6c8dae7bdcbe15ce9 164 0
+a6a23143d2240b5fcacf31f97f0b436ce20b01cf 164 0
 #2 := false
 decl f3 :: (-> int S2)
 decl f4 :: (-> S2 int)
@@ -24396,7 +24604,7 @@
 #142 := [mp #133 #139]: #137
 [unit-resolution #142 #586]: false
 unsat
-b26c16020695ec5cb482a06eafc452bf48579bac 139 0
+f5e2e20fb8c501f38b6d2f15c9fef9418df3a3f5 139 0
 #2 := false
 #13 := 0::int
 decl f4 :: (-> S2 int)
@@ -24536,7 +24744,7 @@
 #260 := [th-lemma]: #618
 [unit-resolution #260 #139 #617]: false
 unsat
-cc925ce2d7fa7e6045f376fe0b86848bf9965e04 75 0
+3413e63e105cab337755fe6950808ba664d2597a 75 0
 #2 := false
 decl f3 :: (-> int S2)
 #28 := 1::int
@@ -24612,7 +24820,7 @@
 #153 := [and-elim #151]: #32
 [unit-resolution #153 #339]: false
 unsat
-258c4ff3957e31cbdc91a889a36ababa8b18bb74 20 0
+7c99e4c706a5bfe7c4bac45e274e4bb5f5ba171b 20 0
 #2 := false
 decl f4 :: (-> S2 int)
 decl f5 :: S2
@@ -24633,7 +24841,7 @@
 #130 := [asserted]: #28
 [mp #130 #141]: false
 unsat
-b707e36223641f43f77b1aee28addc35f2bcee4b 319 0
+2c1b55c23f4e7de8da8d5f6f9bfa461e094d5fdc 319 0
 #2 := false
 #13 := 0::int
 decl f4 :: (-> S2 int)
@@ -24953,25 +25161,7 @@
 #507 := [unit-resolution #506 #504]: #623
 [th-lemma #507 #516 #163 #162]: false
 unsat
-631bb5c1eda2ba4c09e88b58a78c4ca019a52bd5 17 0
-#2 := false
-#8 := 0::int
-#9 := (= 0::int 0::int)
-#10 := (not #9)
-#35 := (iff #10 false)
-#1 := true
-#30 := (not true)
-#33 := (iff #30 false)
-#34 := [rewrite]: #33
-#31 := (iff #10 #30)
-#28 := (iff #9 true)
-#29 := [rewrite]: #28
-#32 := [monotonicity #29]: #31
-#36 := [trans #32 #34]: #35
-#27 := [asserted]: #10
-[mp #27 #36]: false
-unsat
-333d9adb9c916d1ff36a4481fd80329786125ad5 325 0
+b1d5ad1c199444dbf6fe73759c2c90d7c60b146e 325 0
 #2 := false
 #13 := 0::int
 decl f4 :: (-> S2 int)
@@ -25297,51 +25487,7 @@
 #515 := [th-lemma]: #519
 [unit-resolution #515 #510 #517]: false
 unsat
-28a49ab9e1e8afe831840a9d1d70432378c40524 25 0
-#2 := false
-#8 := 0::int
-#9 := (- 0::int)
-#10 := (= 0::int #9)
-#11 := (not #10)
-#43 := (iff #11 false)
-#1 := true
-#38 := (not true)
-#41 := (iff #38 false)
-#42 := [rewrite]: #41
-#39 := (iff #11 #38)
-#36 := (iff #10 true)
-#31 := (= 0::int 0::int)
-#34 := (iff #31 true)
-#35 := [rewrite]: #34
-#32 := (iff #10 #31)
-#29 := (= #9 0::int)
-#30 := [rewrite]: #29
-#33 := [monotonicity #30]: #32
-#37 := [trans #33 #35]: #36
-#40 := [monotonicity #37]: #39
-#44 := [trans #40 #42]: #43
-#28 := [asserted]: #11
-[mp #28 #44]: false
-unsat
-faf90dd136ed7151b245f7421706b221be33e497 17 0
-#2 := false
-#8 := 1::int
-#9 := (= 1::int 1::int)
-#10 := (not #9)
-#35 := (iff #10 false)
-#1 := true
-#30 := (not true)
-#33 := (iff #30 false)
-#34 := [rewrite]: #33
-#31 := (iff #10 #30)
-#28 := (iff #9 true)
-#29 := [rewrite]: #28
-#32 := [monotonicity #29]: #31
-#36 := [trans #32 #34]: #35
-#27 := [asserted]: #10
-[mp #27 #36]: false
-unsat
-00947a900acc4b06198fe82553f699a3348499ea 38 0
+a489de1c5d2012bec455302dfb4702c1a1df77d2 38 0
 #2 := false
 #13 := 0::int
 decl f4 :: (-> S2 int)
@@ -25380,42 +25526,7 @@
 #228 := [unit-resolution #313 #154]: #144
 [unit-resolution #228 #155]: false
 unsat
-7392687761c667823aab65a1bcb19e1202861610 34 0
-#2 := false
-#8 := 1::int
-#9 := (- 1::int)
-#10 := (= #9 1::int)
-#11 := (not #10)
-#12 := (not #11)
-#52 := (iff #12 false)
-#1 := true
-#47 := (not true)
-#50 := (iff #47 false)
-#51 := [rewrite]: #50
-#48 := (iff #12 #47)
-#45 := (iff #11 true)
-#40 := (not false)
-#43 := (iff #40 true)
-#44 := [rewrite]: #43
-#41 := (iff #11 #40)
-#38 := (iff #10 false)
-#30 := -1::int
-#33 := (= -1::int 1::int)
-#36 := (iff #33 false)
-#37 := [rewrite]: #36
-#34 := (iff #10 #33)
-#31 := (= #9 -1::int)
-#32 := [rewrite]: #31
-#35 := [monotonicity #32]: #34
-#39 := [trans #35 #37]: #38
-#42 := [monotonicity #39]: #41
-#46 := [trans #42 #44]: #45
-#49 := [monotonicity #46]: #48
-#53 := [trans #49 #51]: #52
-#29 := [asserted]: #12
-[mp #29 #53]: false
-unsat
-01e7532e7ea90063f4d8028d82907a60219f0691 292 0
+08e6db4ea57a4ed26d47431f93a880d0e622c6f9 292 0
 #2 := false
 #13 := 0::int
 decl f4 :: (-> S2 int)
@@ -25708,79 +25819,7 @@
 #583 := [unit-resolution #586 #503]: #657
 [th-lemma #583 #605 #592]: false
 unsat
-e368f2011d2d745eea4e08822ed13634e5b6be4e 40 0
-#2 := false
-#12 := 567::int
-#10 := 345::int
-#8 := 123::int
-#9 := (- 123::int)
-#11 := (+ #9 345::int)
-#13 := (< #11 567::int)
-#14 := (not #13)
-#58 := (iff #14 false)
-#38 := 222::int
-#43 := (< 222::int 567::int)
-#46 := (not #43)
-#56 := (iff #46 false)
-#1 := true
-#51 := (not true)
-#54 := (iff #51 false)
-#55 := [rewrite]: #54
-#52 := (iff #46 #51)
-#49 := (iff #43 true)
-#50 := [rewrite]: #49
-#53 := [monotonicity #50]: #52
-#57 := [trans #53 #55]: #56
-#47 := (iff #14 #46)
-#44 := (iff #13 #43)
-#41 := (= #11 222::int)
-#32 := -123::int
-#35 := (+ -123::int 345::int)
-#39 := (= #35 222::int)
-#40 := [rewrite]: #39
-#36 := (= #11 #35)
-#33 := (= #9 -123::int)
-#34 := [rewrite]: #33
-#37 := [monotonicity #34]: #36
-#42 := [trans #37 #40]: #41
-#45 := [monotonicity #42]: #44
-#48 := [monotonicity #45]: #47
-#59 := [trans #48 #57]: #58
-#31 := [asserted]: #14
-[mp #31 #59]: false
-unsat
-6856ce5c38958c154c5630f30070f1299de0a73a 30 0
-#2 := false
-#10 := 2345678901::int
-#8 := 123456789::int
-#9 := (- 123456789::int)
-#11 := (< #9 2345678901::int)
-#12 := (not #11)
-#48 := (iff #12 false)
-#30 := -123456789::int
-#33 := (< -123456789::int 2345678901::int)
-#36 := (not #33)
-#46 := (iff #36 false)
-#1 := true
-#41 := (not true)
-#44 := (iff #41 false)
-#45 := [rewrite]: #44
-#42 := (iff #36 #41)
-#39 := (iff #33 true)
-#40 := [rewrite]: #39
-#43 := [monotonicity #40]: #42
-#47 := [trans #43 #45]: #46
-#37 := (iff #12 #36)
-#34 := (iff #11 #33)
-#31 := (= #9 -123456789::int)
-#32 := [rewrite]: #31
-#35 := [monotonicity #32]: #34
-#38 := [monotonicity #35]: #37
-#49 := [trans #38 #47]: #48
-#29 := [asserted]: #12
-[mp #29 #49]: false
-unsat
-9af62203be53f0918acd3e5a19c29adeb325e821 11 0
+7a80c4532f5787abddc50ad2e7bc9838390acb73 11 0
 #2 := false
 decl f4 :: (-> S2 int)
 decl f5 :: S2
@@ -25792,34 +25831,7 @@
 #129 := [asserted]: #27
 [mp #129 #133]: false
 unsat
-5beff0398ac80bd4310abe8d1006f828914d1ddd 26 0
-#2 := false
-decl f3 :: int
-#8 := f3
-#9 := 0::int
-#10 := (+ f3 0::int)
-#11 := (= #10 f3)
-#12 := (not #11)
-#43 := (iff #12 false)
-#1 := true
-#38 := (not true)
-#41 := (iff #38 false)
-#42 := [rewrite]: #41
-#39 := (iff #12 #38)
-#36 := (iff #11 true)
-#31 := (= f3 f3)
-#34 := (iff #31 true)
-#35 := [rewrite]: #34
-#32 := (iff #11 #31)
-#30 := [rewrite]: #11
-#33 := [monotonicity #30]: #32
-#37 := [trans #33 #35]: #36
-#40 := [monotonicity #37]: #39
-#44 := [trans #40 #42]: #43
-#29 := [asserted]: #12
-[mp #29 #44]: false
-unsat
-ecbbc277377c6ea59671702660167f70e18bac2b 54 0
+a5e0b28483c053128bc0be113d20a82250433b74 54 0
 #2 := false
 #13 := 0::int
 decl f4 :: (-> S2 int)
@@ -25874,64 +25886,7 @@
 #171 := [and-elim #170]: #144
 [th-lemma #171 #172 #173]: false
 unsat
-89946d5d67b763e42f800ede863ab8bce34f2003 26 0
-#2 := false
-decl f3 :: int
-#9 := f3
-#8 := 0::int
-#10 := (+ 0::int f3)
-#11 := (= #10 f3)
-#12 := (not #11)
-#43 := (iff #12 false)
-#1 := true
-#38 := (not true)
-#41 := (iff #38 false)
-#42 := [rewrite]: #41
-#39 := (iff #12 #38)
-#36 := (iff #11 true)
-#31 := (= f3 f3)
-#34 := (iff #31 true)
-#35 := [rewrite]: #34
-#32 := (iff #11 #31)
-#30 := [rewrite]: #11
-#33 := [monotonicity #30]: #32
-#37 := [trans #33 #35]: #36
-#40 := [monotonicity #37]: #39
-#44 := [trans #40 #42]: #43
-#29 := [asserted]: #12
-[mp #29 #44]: false
-unsat
-bb5d039a38eba1786eb72150674dd17c7a6eed7e 29 0
-#2 := false
-decl f3 :: int
-#8 := f3
-decl f4 :: int
-#9 := f4
-#11 := (+ f4 f3)
-#10 := (+ f3 f4)
-#12 := (= #10 #11)
-#13 := (not #12)
-#45 := (iff #13 false)
-#1 := true
-#40 := (not true)
-#43 := (iff #40 false)
-#44 := [rewrite]: #43
-#41 := (iff #13 #40)
-#38 := (iff #12 true)
-#33 := (= #10 #10)
-#36 := (iff #33 true)
-#37 := [rewrite]: #36
-#34 := (iff #12 #33)
-#31 := (= #11 #10)
-#32 := [rewrite]: #31
-#35 := [monotonicity #32]: #34
-#39 := [trans #35 #37]: #38
-#42 := [monotonicity #39]: #41
-#46 := [trans #42 #44]: #45
-#30 := [asserted]: #13
-[mp #30 #46]: false
-unsat
-e761c0b6a7ce21273c17c2e2c5540ab035d56e47 55 0
+4d82a9a048129e417a9e0a1825ce1ac4532cabd0 55 0
 #2 := false
 #13 := 0::int
 decl f4 :: (-> S2 int)
@@ -25987,41 +25942,7 @@
 #172 := [and-elim #171]: #144
 [th-lemma #172 #173 #174]: false
 unsat
-6be423993ea88d7bfa5dfa04d4fc58ccd8bf39d4 33 0
-#2 := false
-decl f5 :: int
-#10 := f5
-decl f4 :: int
-#9 := f4
-decl f3 :: int
-#8 := f3
-#13 := (+ f3 f4)
-#14 := (+ #13 f5)
-#11 := (+ f4 f5)
-#12 := (+ f3 #11)
-#15 := (= #12 #14)
-#16 := (not #15)
-#48 := (iff #16 false)
-#1 := true
-#43 := (not true)
-#46 := (iff #43 false)
-#47 := [rewrite]: #46
-#44 := (iff #16 #43)
-#41 := (iff #15 true)
-#36 := (= #12 #12)
-#39 := (iff #36 true)
-#40 := [rewrite]: #39
-#37 := (iff #15 #36)
-#34 := (= #14 #12)
-#35 := [rewrite]: #34
-#38 := [monotonicity #35]: #37
-#42 := [trans #38 #40]: #41
-#45 := [monotonicity #42]: #44
-#49 := [trans #45 #47]: #48
-#33 := [asserted]: #16
-[mp #33 #49]: false
-unsat
-8ef23587881ccafe5fb8b0c94a42df18ed62c340 55 0
+839ca15e4db302dc5cd29e9aa43868c766ec1b81 55 0
 #2 := false
 #13 := 0::int
 decl f4 :: (-> S2 int)
@@ -26077,53 +25998,7 @@
 #173 := [and-elim #171]: #152
 [th-lemma #173 #172 #174]: false
 unsat
-f34f9158033293bcd8d531d1fd221cc3340dd1d4 45 0
-#2 := false
-decl f4 :: int
-#9 := f4
-#13 := (- f4)
-decl f3 :: int
-#8 := f3
-#14 := (= f3 #13)
-#11 := 0::int
-#10 := (+ f3 f4)
-#12 := (= #10 0::int)
-#15 := (iff #12 #14)
-#16 := (not #15)
-#62 := (iff #16 false)
-#47 := (not #12)
-#34 := -1::int
-#35 := (* -1::int f4)
-#38 := (= f3 #35)
-#48 := (iff #38 #47)
-#60 := (iff #48 false)
-#55 := (iff #12 #47)
-#58 := (iff #55 false)
-#59 := [rewrite]: #58
-#56 := (iff #48 #55)
-#53 := (iff #38 #12)
-#54 := [rewrite]: #53
-#57 := [monotonicity #54]: #56
-#61 := [trans #57 #59]: #60
-#51 := (iff #16 #48)
-#41 := (iff #12 #38)
-#44 := (not #41)
-#49 := (iff #44 #48)
-#50 := [rewrite]: #49
-#45 := (iff #16 #44)
-#42 := (iff #15 #41)
-#39 := (iff #14 #38)
-#36 := (= #13 #35)
-#37 := [rewrite]: #36
-#40 := [monotonicity #37]: #39
-#43 := [monotonicity #40]: #42
-#46 := [monotonicity #43]: #45
-#52 := [trans #46 #50]: #51
-#63 := [trans #52 #61]: #62
-#33 := [asserted]: #16
-[mp #33 #63]: false
-unsat
-3c933bbec35e3b1947f865f102b1cc06b618d6d5 61 0
+2fa07586f302aa9750784686f2cfa4df7ec4584e 61 0
 #2 := false
 #13 := 0::int
 decl f4 :: (-> S2 int)
@@ -26185,61 +26060,7 @@
 #179 := [and-elim #177]: #149
 [th-lemma #179 #178 #180]: false
 unsat
-581098ef0bfed59ccd21b34aa98652ee0c38c97c 26 0
-#2 := false
-#8 := 1::int
-#9 := (- 1::int)
-#10 := (= #9 #9)
-#11 := (not #10)
-#44 := (iff #11 false)
-#1 := true
-#39 := (not true)
-#42 := (iff #39 false)
-#43 := [rewrite]: #42
-#40 := (iff #11 #39)
-#37 := (iff #10 true)
-#29 := -1::int
-#32 := (= -1::int -1::int)
-#35 := (iff #32 true)
-#36 := [rewrite]: #35
-#33 := (iff #10 #32)
-#30 := (= #9 -1::int)
-#31 := [rewrite]: #30
-#34 := [monotonicity #31 #31]: #33
-#38 := [trans #34 #36]: #37
-#41 := [monotonicity #38]: #40
-#45 := [trans #41 #43]: #44
-#28 := [asserted]: #11
-[mp #28 #45]: false
-unsat
-627b9649fe4139aa10844958a49f92fc2ed56e6b 26 0
-#2 := false
-#8 := 3::int
-#9 := (- 3::int)
-#10 := (= #9 #9)
-#11 := (not #10)
-#44 := (iff #11 false)
-#1 := true
-#39 := (not true)
-#42 := (iff #39 false)
-#43 := [rewrite]: #42
-#40 := (iff #11 #39)
-#37 := (iff #10 true)
-#29 := -3::int
-#32 := (= -3::int -3::int)
-#35 := (iff #32 true)
-#36 := [rewrite]: #35
-#33 := (iff #10 #32)
-#30 := (= #9 -3::int)
-#31 := [rewrite]: #30
-#34 := [monotonicity #31 #31]: #33
-#38 := [trans #34 #36]: #37
-#41 := [monotonicity #38]: #40
-#45 := [trans #41 #43]: #44
-#28 := [asserted]: #11
-[mp #28 #45]: false
-unsat
-6f0da92e2f84bd612985ca7ec2cf4ab63063d3d9 53 0
+fe16712a1fa51f4a62675a12f177489494d4ff85 53 0
 #2 := false
 #13 := 0::int
 decl f4 :: (-> S2 int)
@@ -26293,7 +26114,394 @@
 #171 := [and-elim #169]: #149
 [th-lemma #171 #170 #172]: false
 unsat
-38476cf6298acd443b7e96455508697d68d94894 58 0
+88db0d3335480a81cfc5f118ecaef1678c52c57c 17 0
+#2 := false
+#8 := 0::int
+#9 := (= 0::int 0::int)
+#10 := (not #9)
+#35 := (iff #10 false)
+#1 := true
+#30 := (not true)
+#33 := (iff #30 false)
+#34 := [rewrite]: #33
+#31 := (iff #10 #30)
+#28 := (iff #9 true)
+#29 := [rewrite]: #28
+#32 := [monotonicity #29]: #31
+#36 := [trans #32 #34]: #35
+#27 := [asserted]: #10
+[mp #27 #36]: false
+unsat
+3bd4cf5a7808f0310fbc231a810dbf418d0f69cc 25 0
+#2 := false
+#8 := 0::int
+#9 := (- 0::int)
+#10 := (= 0::int #9)
+#11 := (not #10)
+#43 := (iff #11 false)
+#1 := true
+#38 := (not true)
+#41 := (iff #38 false)
+#42 := [rewrite]: #41
+#39 := (iff #11 #38)
+#36 := (iff #10 true)
+#31 := (= 0::int 0::int)
+#34 := (iff #31 true)
+#35 := [rewrite]: #34
+#32 := (iff #10 #31)
+#29 := (= #9 0::int)
+#30 := [rewrite]: #29
+#33 := [monotonicity #30]: #32
+#37 := [trans #33 #35]: #36
+#40 := [monotonicity #37]: #39
+#44 := [trans #40 #42]: #43
+#28 := [asserted]: #11
+[mp #28 #44]: false
+unsat
+575466500975f7b2e0bdbc552a0711af0164ab2f 17 0
+#2 := false
+#8 := 1::int
+#9 := (= 1::int 1::int)
+#10 := (not #9)
+#35 := (iff #10 false)
+#1 := true
+#30 := (not true)
+#33 := (iff #30 false)
+#34 := [rewrite]: #33
+#31 := (iff #10 #30)
+#28 := (iff #9 true)
+#29 := [rewrite]: #28
+#32 := [monotonicity #29]: #31
+#36 := [trans #32 #34]: #35
+#27 := [asserted]: #10
+[mp #27 #36]: false
+unsat
+a2fa5e02114a58b8161c5f17b5db7f1bd8de21cc 34 0
+#2 := false
+#8 := 1::int
+#9 := (- 1::int)
+#10 := (= #9 1::int)
+#11 := (not #10)
+#12 := (not #11)
+#52 := (iff #12 false)
+#1 := true
+#47 := (not true)
+#50 := (iff #47 false)
+#51 := [rewrite]: #50
+#48 := (iff #12 #47)
+#45 := (iff #11 true)
+#40 := (not false)
+#43 := (iff #40 true)
+#44 := [rewrite]: #43
+#41 := (iff #11 #40)
+#38 := (iff #10 false)
+#30 := -1::int
+#33 := (= -1::int 1::int)
+#36 := (iff #33 false)
+#37 := [rewrite]: #36
+#34 := (iff #10 #33)
+#31 := (= #9 -1::int)
+#32 := [rewrite]: #31
+#35 := [monotonicity #32]: #34
+#39 := [trans #35 #37]: #38
+#42 := [monotonicity #39]: #41
+#46 := [trans #42 #44]: #45
+#49 := [monotonicity #46]: #48
+#53 := [trans #49 #51]: #52
+#29 := [asserted]: #12
+[mp #29 #53]: false
+unsat
+833da25c34dad32196cafeb37ee24e9c4b09f5e6 40 0
+#2 := false
+#12 := 567::int
+#10 := 345::int
+#8 := 123::int
+#9 := (- 123::int)
+#11 := (+ #9 345::int)
+#13 := (< #11 567::int)
+#14 := (not #13)
+#58 := (iff #14 false)
+#38 := 222::int
+#43 := (< 222::int 567::int)
+#46 := (not #43)
+#56 := (iff #46 false)
+#1 := true
+#51 := (not true)
+#54 := (iff #51 false)
+#55 := [rewrite]: #54
+#52 := (iff #46 #51)
+#49 := (iff #43 true)
+#50 := [rewrite]: #49
+#53 := [monotonicity #50]: #52
+#57 := [trans #53 #55]: #56
+#47 := (iff #14 #46)
+#44 := (iff #13 #43)
+#41 := (= #11 222::int)
+#32 := -123::int
+#35 := (+ -123::int 345::int)
+#39 := (= #35 222::int)
+#40 := [rewrite]: #39
+#36 := (= #11 #35)
+#33 := (= #9 -123::int)
+#34 := [rewrite]: #33
+#37 := [monotonicity #34]: #36
+#42 := [trans #37 #40]: #41
+#45 := [monotonicity #42]: #44
+#48 := [monotonicity #45]: #47
+#59 := [trans #48 #57]: #58
+#31 := [asserted]: #14
+[mp #31 #59]: false
+unsat
+94a3c69e40ab8834fda1414aac36766b6e6d0448 30 0
+#2 := false
+#10 := 2345678901::int
+#8 := 123456789::int
+#9 := (- 123456789::int)
+#11 := (< #9 2345678901::int)
+#12 := (not #11)
+#48 := (iff #12 false)
+#30 := -123456789::int
+#33 := (< -123456789::int 2345678901::int)
+#36 := (not #33)
+#46 := (iff #36 false)
+#1 := true
+#41 := (not true)
+#44 := (iff #41 false)
+#45 := [rewrite]: #44
+#42 := (iff #36 #41)
+#39 := (iff #33 true)
+#40 := [rewrite]: #39
+#43 := [monotonicity #40]: #42
+#47 := [trans #43 #45]: #46
+#37 := (iff #12 #36)
+#34 := (iff #11 #33)
+#31 := (= #9 -123456789::int)
+#32 := [rewrite]: #31
+#35 := [monotonicity #32]: #34
+#38 := [monotonicity #35]: #37
+#49 := [trans #38 #47]: #48
+#29 := [asserted]: #12
+[mp #29 #49]: false
+unsat
+44da310bca35772c7c4ab7b6f9468b0cf95baf9e 26 0
+#2 := false
+decl f3 :: int
+#8 := f3
+#9 := 0::int
+#10 := (+ f3 0::int)
+#11 := (= #10 f3)
+#12 := (not #11)
+#43 := (iff #12 false)
+#1 := true
+#38 := (not true)
+#41 := (iff #38 false)
+#42 := [rewrite]: #41
+#39 := (iff #12 #38)
+#36 := (iff #11 true)
+#31 := (= f3 f3)
+#34 := (iff #31 true)
+#35 := [rewrite]: #34
+#32 := (iff #11 #31)
+#30 := [rewrite]: #11
+#33 := [monotonicity #30]: #32
+#37 := [trans #33 #35]: #36
+#40 := [monotonicity #37]: #39
+#44 := [trans #40 #42]: #43
+#29 := [asserted]: #12
+[mp #29 #44]: false
+unsat
+fed373ed1a62868b78e70bd0d9a3428669ee5457 26 0
+#2 := false
+decl f3 :: int
+#9 := f3
+#8 := 0::int
+#10 := (+ 0::int f3)
+#11 := (= #10 f3)
+#12 := (not #11)
+#43 := (iff #12 false)
+#1 := true
+#38 := (not true)
+#41 := (iff #38 false)
+#42 := [rewrite]: #41
+#39 := (iff #12 #38)
+#36 := (iff #11 true)
+#31 := (= f3 f3)
+#34 := (iff #31 true)
+#35 := [rewrite]: #34
+#32 := (iff #11 #31)
+#30 := [rewrite]: #11
+#33 := [monotonicity #30]: #32
+#37 := [trans #33 #35]: #36
+#40 := [monotonicity #37]: #39
+#44 := [trans #40 #42]: #43
+#29 := [asserted]: #12
+[mp #29 #44]: false
+unsat
+bb9d75a7a3abc7f60dd441b4f982645b6039e03e 29 0
+#2 := false
+decl f3 :: int
+#8 := f3
+decl f4 :: int
+#9 := f4
+#11 := (+ f4 f3)
+#10 := (+ f3 f4)
+#12 := (= #10 #11)
+#13 := (not #12)
+#45 := (iff #13 false)
+#1 := true
+#40 := (not true)
+#43 := (iff #40 false)
+#44 := [rewrite]: #43
+#41 := (iff #13 #40)
+#38 := (iff #12 true)
+#33 := (= #10 #10)
+#36 := (iff #33 true)
+#37 := [rewrite]: #36
+#34 := (iff #12 #33)
+#31 := (= #11 #10)
+#32 := [rewrite]: #31
+#35 := [monotonicity #32]: #34
+#39 := [trans #35 #37]: #38
+#42 := [monotonicity #39]: #41
+#46 := [trans #42 #44]: #45
+#30 := [asserted]: #13
+[mp #30 #46]: false
+unsat
+698ca3e8d5433877845f98d57033a6d63af62497 33 0
+#2 := false
+decl f5 :: int
+#10 := f5
+decl f4 :: int
+#9 := f4
+decl f3 :: int
+#8 := f3
+#13 := (+ f3 f4)
+#14 := (+ #13 f5)
+#11 := (+ f4 f5)
+#12 := (+ f3 #11)
+#15 := (= #12 #14)
+#16 := (not #15)
+#48 := (iff #16 false)
+#1 := true
+#43 := (not true)
+#46 := (iff #43 false)
+#47 := [rewrite]: #46
+#44 := (iff #16 #43)
+#41 := (iff #15 true)
+#36 := (= #12 #12)
+#39 := (iff #36 true)
+#40 := [rewrite]: #39
+#37 := (iff #15 #36)
+#34 := (= #14 #12)
+#35 := [rewrite]: #34
+#38 := [monotonicity #35]: #37
+#42 := [trans #38 #40]: #41
+#45 := [monotonicity #42]: #44
+#49 := [trans #45 #47]: #48
+#33 := [asserted]: #16
+[mp #33 #49]: false
+unsat
+41dc588e082dd0bed2354a8e37e9f3c6c05151c1 45 0
+#2 := false
+decl f4 :: int
+#9 := f4
+#13 := (- f4)
+decl f3 :: int
+#8 := f3
+#14 := (= f3 #13)
+#11 := 0::int
+#10 := (+ f3 f4)
+#12 := (= #10 0::int)
+#15 := (iff #12 #14)
+#16 := (not #15)
+#62 := (iff #16 false)
+#47 := (not #12)
+#34 := -1::int
+#35 := (* -1::int f4)
+#38 := (= f3 #35)
+#48 := (iff #38 #47)
+#60 := (iff #48 false)
+#55 := (iff #12 #47)
+#58 := (iff #55 false)
+#59 := [rewrite]: #58
+#56 := (iff #48 #55)
+#53 := (iff #38 #12)
+#54 := [rewrite]: #53
+#57 := [monotonicity #54]: #56
+#61 := [trans #57 #59]: #60
+#51 := (iff #16 #48)
+#41 := (iff #12 #38)
+#44 := (not #41)
+#49 := (iff #44 #48)
+#50 := [rewrite]: #49
+#45 := (iff #16 #44)
+#42 := (iff #15 #41)
+#39 := (iff #14 #38)
+#36 := (= #13 #35)
+#37 := [rewrite]: #36
+#40 := [monotonicity #37]: #39
+#43 := [monotonicity #40]: #42
+#46 := [monotonicity #43]: #45
+#52 := [trans #46 #50]: #51
+#63 := [trans #52 #61]: #62
+#33 := [asserted]: #16
+[mp #33 #63]: false
+unsat
+d7837998c869430b9e700142d2b377bead7796dc 26 0
+#2 := false
+#8 := 1::int
+#9 := (- 1::int)
+#10 := (= #9 #9)
+#11 := (not #10)
+#44 := (iff #11 false)
+#1 := true
+#39 := (not true)
+#42 := (iff #39 false)
+#43 := [rewrite]: #42
+#40 := (iff #11 #39)
+#37 := (iff #10 true)
+#29 := -1::int
+#32 := (= -1::int -1::int)
+#35 := (iff #32 true)
+#36 := [rewrite]: #35
+#33 := (iff #10 #32)
+#30 := (= #9 -1::int)
+#31 := [rewrite]: #30
+#34 := [monotonicity #31 #31]: #33
+#38 := [trans #34 #36]: #37
+#41 := [monotonicity #38]: #40
+#45 := [trans #41 #43]: #44
+#28 := [asserted]: #11
+[mp #28 #45]: false
+unsat
+70a72eb40efd415a7fc201441ac2d390324dc85f 26 0
+#2 := false
+#8 := 3::int
+#9 := (- 3::int)
+#10 := (= #9 #9)
+#11 := (not #10)
+#44 := (iff #11 false)
+#1 := true
+#39 := (not true)
+#42 := (iff #39 false)
+#43 := [rewrite]: #42
+#40 := (iff #11 #39)
+#37 := (iff #10 true)
+#29 := -3::int
+#32 := (= -3::int -3::int)
+#35 := (iff #32 true)
+#36 := [rewrite]: #35
+#33 := (iff #10 #32)
+#30 := (= #9 -3::int)
+#31 := [rewrite]: #30
+#34 := [monotonicity #31 #31]: #33
+#38 := [trans #34 #36]: #37
+#41 := [monotonicity #38]: #40
+#45 := [trans #41 #43]: #44
+#28 := [asserted]: #11
+[mp #28 #45]: false
+unsat
+c4720018b6df523588840d966fa6b5f79e529106 58 0
 #2 := false
 decl f3 :: int
 #8 := f3
@@ -26352,7 +26560,7 @@
 #31 := [asserted]: #14
 [mp #31 #77]: false
 unsat
-792b4b3859919262c72016c9acfe8ea3843e85d7 61 0
+c156df2cdcaafc0712bbd7f760cf7c1135b2a278 61 0
 #2 := false
 #8 := 0::int
 decl f3 :: int
@@ -26414,67 +26622,7 @@
 #31 := [asserted]: #14
 [mp #31 #79]: false
 unsat
-e55670f249d7d5f46be920cd9d904dea4e4fc39f 26 0
-#2 := false
-decl f3 :: int
-#8 := f3
-#9 := 0::int
-#10 := (- f3 0::int)
-#11 := (= #10 f3)
-#12 := (not #11)
-#43 := (iff #12 false)
-#1 := true
-#38 := (not true)
-#41 := (iff #38 false)
-#42 := [rewrite]: #41
-#39 := (iff #12 #38)
-#36 := (iff #11 true)
-#31 := (= f3 f3)
-#34 := (iff #31 true)
-#35 := [rewrite]: #34
-#32 := (iff #11 #31)
-#30 := [rewrite]: #11
-#33 := [monotonicity #30]: #32
-#37 := [trans #33 #35]: #36
-#40 := [monotonicity #37]: #39
-#44 := [trans #40 #42]: #43
-#29 := [asserted]: #12
-[mp #29 #44]: false
-unsat
-a90b5df92ad9a6706799b2ce3864229806b976dc 32 0
-#2 := false
-decl f3 :: int
-#9 := f3
-#11 := (- f3)
-#8 := 0::int
-#10 := (- 0::int f3)
-#12 := (= #10 #11)
-#13 := (not #12)
-#49 := (iff #13 false)
-#1 := true
-#44 := (not true)
-#47 := (iff #44 false)
-#48 := [rewrite]: #47
-#45 := (iff #13 #44)
-#42 := (iff #12 true)
-#31 := -1::int
-#32 := (* -1::int f3)
-#37 := (= #32 #32)
-#40 := (iff #37 true)
-#41 := [rewrite]: #40
-#38 := (iff #12 #37)
-#35 := (= #11 #32)
-#36 := [rewrite]: #35
-#33 := (= #10 #32)
-#34 := [rewrite]: #33
-#39 := [monotonicity #34 #36]: #38
-#43 := [trans #39 #41]: #42
-#46 := [monotonicity #43]: #45
-#50 := [trans #46 #48]: #49
-#30 := [asserted]: #13
-[mp #30 #50]: false
-unsat
-ad72886d0079cf60e87d335a960f64cba8248b20 61 0
+1c28009871674ce12da75bf1dced4812d4d59e07 61 0
 #2 := false
 decl f3 :: int
 #8 := f3
@@ -26536,7 +26684,67 @@
 #31 := [asserted]: #14
 [mp #31 #79]: false
 unsat
-d08f4e145bad8aa6777f75810df1459ec9077380 64 0
+fc81f629d09c01783c65b3ee17baad49669d55bf 26 0
+#2 := false
+decl f3 :: int
+#8 := f3
+#9 := 0::int
+#10 := (- f3 0::int)
+#11 := (= #10 f3)
+#12 := (not #11)
+#43 := (iff #12 false)
+#1 := true
+#38 := (not true)
+#41 := (iff #38 false)
+#42 := [rewrite]: #41
+#39 := (iff #12 #38)
+#36 := (iff #11 true)
+#31 := (= f3 f3)
+#34 := (iff #31 true)
+#35 := [rewrite]: #34
+#32 := (iff #11 #31)
+#30 := [rewrite]: #11
+#33 := [monotonicity #30]: #32
+#37 := [trans #33 #35]: #36
+#40 := [monotonicity #37]: #39
+#44 := [trans #40 #42]: #43
+#29 := [asserted]: #12
+[mp #29 #44]: false
+unsat
+aec51136d021a387f2d6872e725f19b898f18ef5 32 0
+#2 := false
+decl f3 :: int
+#9 := f3
+#11 := (- f3)
+#8 := 0::int
+#10 := (- 0::int f3)
+#12 := (= #10 #11)
+#13 := (not #12)
+#49 := (iff #13 false)
+#1 := true
+#44 := (not true)
+#47 := (iff #44 false)
+#48 := [rewrite]: #47
+#45 := (iff #13 #44)
+#42 := (iff #12 true)
+#31 := -1::int
+#32 := (* -1::int f3)
+#37 := (= #32 #32)
+#40 := (iff #37 true)
+#41 := [rewrite]: #40
+#38 := (iff #12 #37)
+#35 := (= #11 #32)
+#36 := [rewrite]: #35
+#33 := (= #10 #32)
+#34 := [rewrite]: #33
+#39 := [monotonicity #34 #36]: #38
+#43 := [trans #39 #41]: #42
+#46 := [monotonicity #43]: #45
+#50 := [trans #46 #48]: #49
+#30 := [asserted]: #13
+[mp #30 #50]: false
+unsat
+31fc2c351446233f2061fde4a4592b401c54f862 64 0
 #2 := false
 #12 := 0::int
 decl f4 :: int
@@ -26601,7 +26809,7 @@
 #32 := [asserted]: #15
 [mp #32 #81]: false
 unsat
-b835fac148452954f8fe0754df04f54bf681e269 44 0
+7e5077e88974e2fa3da11f2fc124b2fc85a4f520 44 0
 #2 := false
 decl f3 :: int
 #8 := f3
@@ -26646,7 +26854,7 @@
 #31 := [asserted]: #14
 [mp #31 #61]: false
 unsat
-988b478ff3b6a0c8396fcbd4c2a01d93e477e7a6 42 0
+891b65b932309e97c69155ee809ef85f6496b3c5 42 0
 #2 := false
 decl f3 :: int
 #8 := f3
@@ -26689,34 +26897,7 @@
 #31 := [asserted]: #14
 [mp #31 #59]: false
 unsat
-8d27e1414f3d037da5784d95941af1a199886e3b 26 0
-#2 := false
-#9 := 0::int
-decl f3 :: int
-#8 := f3
-#10 := (* f3 0::int)
-#11 := (= #10 0::int)
-#12 := (not #11)
-#43 := (iff #12 false)
-#1 := true
-#38 := (not true)
-#41 := (iff #38 false)
-#42 := [rewrite]: #41
-#39 := (iff #12 #38)
-#36 := (iff #11 true)
-#31 := (= 0::int 0::int)
-#34 := (iff #31 true)
-#35 := [rewrite]: #34
-#32 := (iff #11 #31)
-#30 := [rewrite]: #11
-#33 := [monotonicity #30]: #32
-#37 := [trans #33 #35]: #36
-#40 := [monotonicity #37]: #39
-#44 := [trans #40 #42]: #43
-#29 := [asserted]: #12
-[mp #29 #44]: false
-unsat
-49a025166be07b22ecb69da498e85645cad3436d 48 0
+e90209aecaafc9e5a8d65cc3bfe9923d664d1180 48 0
 #2 := false
 decl f5 :: int
 #11 := f5
@@ -26765,7 +26946,34 @@
 #33 := [asserted]: #16
 [mp #33 #64]: false
 unsat
-7a9044897deb1d6290b6a90dcd09885d5511c7e4 26 0
+65a9f6762bf7c73e050776f2f97205392ade9eb8 26 0
+#2 := false
+#9 := 0::int
+decl f3 :: int
+#8 := f3
+#10 := (* f3 0::int)
+#11 := (= #10 0::int)
+#12 := (not #11)
+#43 := (iff #12 false)
+#1 := true
+#38 := (not true)
+#41 := (iff #38 false)
+#42 := [rewrite]: #41
+#39 := (iff #12 #38)
+#36 := (iff #11 true)
+#31 := (= 0::int 0::int)
+#34 := (iff #31 true)
+#35 := [rewrite]: #34
+#32 := (iff #11 #31)
+#30 := [rewrite]: #11
+#33 := [monotonicity #30]: #32
+#37 := [trans #33 #35]: #36
+#40 := [monotonicity #37]: #39
+#44 := [trans #40 #42]: #43
+#29 := [asserted]: #12
+[mp #29 #44]: false
+unsat
+592736b964503a2e2a8b3d310866adf35fdeb70d 26 0
 #2 := false
 #8 := 0::int
 decl f3 :: int
@@ -26792,7 +27000,7 @@
 #29 := [asserted]: #12
 [mp #29 #44]: false
 unsat
-5eb0c3df4e96042f330237e1b919e90ecdb74f83 26 0
+89e44e496d1c9e19d85ee56ab3fb92079351f1b0 26 0
 #2 := false
 decl f3 :: int
 #8 := f3
@@ -26819,7 +27027,7 @@
 #29 := [asserted]: #12
 [mp #29 #44]: false
 unsat
-330d57794eed4517ab10b237d85c8054ec6a5ea1 26 0
+38200cf17f447eb7490b14a7927d61babf42be2c 26 0
 #2 := false
 decl f3 :: int
 #9 := f3
@@ -26846,7 +27054,7 @@
 #29 := [asserted]: #12
 [mp #29 #44]: false
 unsat
-118cf0819162003a4a6b17d8f67a00e396d7fe91 40 0
+f1f03cdf8e0332171bfb331e70236cd578ef1b1a 40 0
 #2 := false
 decl f3 :: int
 #8 := f3
@@ -26887,7 +27095,7 @@
 #31 := [asserted]: #14
 [mp #31 #58]: false
 unsat
-4971035b314acd7eb08bbed7d6905e357c5d571c 35 0
+dcb91b564372ef4df7bba10a66d7a6eecfdd1891 35 0
 #2 := false
 decl f3 :: int
 #10 := f3
@@ -26923,7 +27131,7 @@
 #31 := [asserted]: #14
 [mp #31 #53]: false
 unsat
-d8009fc7f359356bf7ea607592f41aaeaef5cced 28 0
+e58e9441ca33f62ea126f129a1302457840ee037 28 0
 #2 := false
 #8 := 3::int
 decl f3 :: int
@@ -26952,7 +27160,7 @@
 #30 := [asserted]: #13
 [mp #30 #46]: false
 unsat
-5f6da7a3c4fe819b8be2bc640a3f6427ec98f801 244 0
+dc9c4895f383c2bc2209a55e4f8478193dad358d 244 0
 #2 := false
 #11 := 0::int
 decl f3 :: (-> int int int)
@@ -27197,276 +27405,7 @@
 #683 := [mp #696 #682]: #695
 [unit-resolution #683 #742 #167]: false
 unsat
-c874d9a87d8800bf1a614a4635e058f59b457075 268 0
-#2 := false
-#11 := 0::int
-decl f4 :: (-> int int int)
-#38 := (f4 0::int 0::int)
-#39 := (= #38 0::int)
-#40 := (not #39)
-#167 := [asserted]: #40
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#29 := (f4 #8 #9)
-#743 := (pattern #29)
-#64 := -1::int
-#68 := (* -1::int #9)
-#65 := (* -1::int #8)
-#123 := (mod #65 #68)
-#254 := (+ #29 #123)
-#255 := (= #254 0::int)
-#30 := (mod #8 #9)
-#251 := (* -1::int #30)
-#252 := (+ #29 #251)
-#253 := (= #252 0::int)
-#90 := (<= #9 0::int)
-#86 := (<= #8 0::int)
-#193 := (or #86 #90)
-#194 := (not #193)
-#97 := (>= #8 0::int)
-#185 := (or #90 #97)
-#186 := (not #185)
-#200 := (or #186 #194)
-#256 := (ite #200 #253 #255)
-#250 := (= #29 0::int)
-#12 := (= #8 0::int)
-#257 := (ite #12 #250 #256)
-#249 := (= #8 #29)
-#13 := (= #9 0::int)
-#258 := (ite #13 #249 #257)
-#744 := (forall (vars (?v0 int) (?v1 int)) (:pat #743) #258)
-#261 := (forall (vars (?v0 int) (?v1 int)) #258)
-#747 := (iff #261 #744)
-#745 := (iff #258 #258)
-#746 := [refl]: #745
-#748 := [quant-intro #746]: #747
-#129 := (* -1::int #123)
-#218 := (ite #200 #30 #129)
-#221 := (ite #12 0::int #218)
-#224 := (ite #13 #8 #221)
-#227 := (= #29 #224)
-#230 := (forall (vars (?v0 int) (?v1 int)) #227)
-#262 := (iff #230 #261)
-#259 := (iff #227 #258)
-#260 := [rewrite]: #259
-#263 := [quant-intro #260]: #262
-#98 := (not #97)
-#91 := (not #90)
-#101 := (and #91 #98)
-#87 := (not #86)
-#94 := (and #87 #91)
-#104 := (or #94 #101)
-#149 := (ite #104 #30 #129)
-#152 := (ite #12 0::int #149)
-#155 := (ite #13 #8 #152)
-#158 := (= #29 #155)
-#161 := (forall (vars (?v0 int) (?v1 int)) #158)
-#231 := (iff #161 #230)
-#228 := (iff #158 #227)
-#225 := (= #155 #224)
-#222 := (= #152 #221)
-#219 := (= #149 #218)
-#203 := (iff #104 #200)
-#197 := (or #194 #186)
-#201 := (iff #197 #200)
-#202 := [rewrite]: #201
-#198 := (iff #104 #197)
-#195 := (iff #101 #186)
-#196 := [rewrite]: #195
-#183 := (iff #94 #194)
-#184 := [rewrite]: #183
-#199 := [monotonicity #184 #196]: #198
-#204 := [trans #199 #202]: #203
-#220 := [monotonicity #204]: #219
-#223 := [monotonicity #220]: #222
-#226 := [monotonicity #223]: #225
-#229 := [monotonicity #226]: #228
-#232 := [quant-intro #229]: #231
-#181 := (~ #161 #161)
-#178 := (~ #158 #158)
-#191 := [refl]: #178
-#182 := [nnf-pos #191]: #181
-#23 := (- #9)
-#22 := (- #8)
-#31 := (mod #22 #23)
-#32 := (- #31)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#33 := (ite #20 #30 #32)
-#34 := (ite #12 0::int #33)
-#35 := (ite #13 #8 #34)
-#36 := (= #29 #35)
-#37 := (forall (vars (?v0 int) (?v1 int)) #36)
-#164 := (iff #37 #161)
-#58 := (and #16 #18)
-#61 := (or #17 #58)
-#134 := (ite #61 #30 #129)
-#137 := (ite #12 0::int #134)
-#140 := (ite #13 #8 #137)
-#143 := (= #29 #140)
-#146 := (forall (vars (?v0 int) (?v1 int)) #143)
-#162 := (iff #146 #161)
-#159 := (iff #143 #158)
-#156 := (= #140 #155)
-#153 := (= #137 #152)
-#150 := (= #134 #149)
-#105 := (iff #61 #104)
-#102 := (iff #58 #101)
-#99 := (iff #18 #98)
-#100 := [rewrite]: #99
-#92 := (iff #16 #91)
-#93 := [rewrite]: #92
-#103 := [monotonicity #93 #100]: #102
-#95 := (iff #17 #94)
-#88 := (iff #15 #87)
-#89 := [rewrite]: #88
-#96 := [monotonicity #89 #93]: #95
-#106 := [monotonicity #96 #103]: #105
-#151 := [monotonicity #106]: #150
-#154 := [monotonicity #151]: #153
-#157 := [monotonicity #154]: #156
-#160 := [monotonicity #157]: #159
-#163 := [quant-intro #160]: #162
-#147 := (iff #37 #146)
-#144 := (iff #36 #143)
-#141 := (= #35 #140)
-#138 := (= #34 #137)
-#135 := (= #33 #134)
-#132 := (= #32 #129)
-#126 := (- #123)
-#130 := (= #126 #129)
-#131 := [rewrite]: #130
-#127 := (= #32 #126)
-#124 := (= #31 #123)
-#69 := (= #23 #68)
-#70 := [rewrite]: #69
-#66 := (= #22 #65)
-#67 := [rewrite]: #66
-#125 := [monotonicity #67 #70]: #124
-#128 := [monotonicity #125]: #127
-#133 := [trans #128 #131]: #132
-#62 := (iff #20 #61)
-#59 := (iff #19 #58)
-#60 := [rewrite]: #59
-#63 := [monotonicity #60]: #62
-#136 := [monotonicity #63 #133]: #135
-#139 := [monotonicity #136]: #138
-#142 := [monotonicity #139]: #141
-#145 := [monotonicity #142]: #144
-#148 := [quant-intro #145]: #147
-#165 := [trans #148 #163]: #164
-#122 := [asserted]: #37
-#166 := [mp #122 #165]: #161
-#192 := [mp~ #166 #182]: #161
-#233 := [mp #192 #232]: #230
-#264 := [mp #233 #263]: #261
-#749 := [mp #264 #748]: #744
-#690 := (not #744)
-#696 := (or #690 #39)
-#322 := (* -1::int 0::int)
-#407 := (mod #322 #322)
-#408 := (+ #38 #407)
-#409 := (= #408 0::int)
-#400 := (mod 0::int 0::int)
-#411 := (* -1::int #400)
-#412 := (+ #38 #411)
-#340 := (= #412 0::int)
-#413 := (<= 0::int 0::int)
-#410 := (or #413 #413)
-#414 := (not #410)
-#393 := (>= 0::int 0::int)
-#728 := (or #413 #393)
-#730 := (not #728)
-#387 := (or #730 #414)
-#517 := (ite #387 #340 #409)
-#724 := (= 0::int 0::int)
-#398 := (ite #724 #39 #517)
-#168 := (= 0::int #38)
-#399 := (ite #724 #168 #398)
-#537 := (or #690 #399)
-#539 := (iff #537 #696)
-#682 := (iff #696 #696)
-#683 := [rewrite]: #682
-#694 := (iff #399 #39)
-#1 := true
-#691 := (ite true #39 #39)
-#688 := (iff #691 #39)
-#689 := [rewrite]: #688
-#692 := (iff #399 #691)
-#698 := (iff #398 #39)
-#328 := (+ #38 #400)
-#428 := (= #328 0::int)
-#699 := (ite true #39 #428)
-#697 := (iff #699 #39)
-#701 := [rewrite]: #697
-#700 := (iff #398 #699)
-#420 := (iff #517 #428)
-#707 := (ite false #340 #428)
-#418 := (iff #707 #428)
-#419 := [rewrite]: #418
-#704 := (iff #517 #707)
-#429 := (iff #409 #428)
-#705 := (= #408 #328)
-#434 := (= #407 #400)
-#432 := (= #322 0::int)
-#433 := [rewrite]: #432
-#435 := [monotonicity #433 #433]: #434
-#706 := [monotonicity #435]: #705
-#703 := [monotonicity #706]: #429
-#709 := (iff #387 false)
-#361 := (or false false)
-#720 := (iff #361 false)
-#723 := [rewrite]: #720
-#362 := (iff #387 #361)
-#719 := (iff #414 false)
-#711 := (not true)
-#376 := (iff #711 false)
-#377 := [rewrite]: #376
-#718 := (iff #414 #711)
-#717 := (iff #410 true)
-#725 := (or true true)
-#726 := (iff #725 true)
-#386 := [rewrite]: #726
-#715 := (iff #410 #725)
-#733 := (iff #413 true)
-#734 := [rewrite]: #733
-#716 := [monotonicity #734 #734]: #715
-#712 := [trans #716 #386]: #717
-#356 := [monotonicity #712]: #718
-#721 := [trans #356 #377]: #719
-#713 := (iff #730 false)
-#374 := (iff #730 #711)
-#727 := (iff #728 true)
-#385 := (iff #728 #725)
-#729 := (iff #393 true)
-#735 := [rewrite]: #729
-#390 := [monotonicity #734 #735]: #385
-#370 := [trans #390 #386]: #727
-#375 := [monotonicity #370]: #374
-#714 := [trans #375 #377]: #713
-#722 := [monotonicity #714 #721]: #362
-#710 := [trans #722 #723]: #709
-#708 := [monotonicity #710 #703]: #704
-#421 := [trans #708 #419]: #420
-#731 := (iff #724 true)
-#732 := [rewrite]: #731
-#415 := [monotonicity #732 #421]: #700
-#702 := [trans #415 #701]: #698
-#174 := (iff #168 #39)
-#175 := [rewrite]: #174
-#693 := [monotonicity #732 #175 #702]: #692
-#695 := [trans #693 #689]: #694
-#681 := [monotonicity #695]: #539
-#684 := [trans #681 #683]: #539
-#538 := [quant-inst]: #537
-#678 := [mp #538 #684]: #696
-[unit-resolution #678 #749 #167]: false
-unsat
-3f2a64334c06cd9085e37e04663f0058aaac8b80 257 0
+6b27a316bc3cb1671937528bbb2aa4c5f16dd00c 257 0
 #2 := false
 #11 := 0::int
 decl f3 :: (-> int int int)
@@ -27724,7 +27663,7 @@
 #673 := [mp #687 #520]: #680
 [unit-resolution #673 #743 #168]: false
 unsat
-2feb38075fe02eeb2836b76bf8c76e21fff33262 281 0
+ac9468ec7d34e56bdf0ee161b89fe3a4afadc438 281 0
 #2 := false
 #11 := 0::int
 decl f3 :: (-> int int int)
@@ -28006,284 +27945,7 @@
 #667 := [mp #573 #670]: #572
 [unit-resolution #667 #742 #168]: false
 unsat
-6b83c7688301a86c79fbfd8522bd1686332a1eaf 276 0
-#2 := false
-decl f4 :: (-> int int int)
-#11 := 0::int
-decl f5 :: int
-#38 := f5
-#39 := (f4 f5 0::int)
-#169 := (= f5 #39)
-#172 := (not #169)
-#40 := (= #39 f5)
-#41 := (not #40)
-#173 := (iff #41 #172)
-#170 := (iff #40 #169)
-#171 := [rewrite]: #170
-#174 := [monotonicity #171]: #173
-#168 := [asserted]: #41
-#177 := [mp #168 #174]: #172
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#29 := (f4 #8 #9)
-#741 := (pattern #29)
-#65 := -1::int
-#69 := (* -1::int #9)
-#66 := (* -1::int #8)
-#124 := (mod #66 #69)
-#252 := (+ #29 #124)
-#253 := (= #252 0::int)
-#30 := (mod #8 #9)
-#249 := (* -1::int #30)
-#250 := (+ #29 #249)
-#251 := (= #250 0::int)
-#91 := (<= #9 0::int)
-#87 := (<= #8 0::int)
-#191 := (or #87 #91)
-#192 := (not #191)
-#98 := (>= #8 0::int)
-#183 := (or #91 #98)
-#184 := (not #183)
-#198 := (or #184 #192)
-#254 := (ite #198 #251 #253)
-#248 := (= #29 0::int)
-#12 := (= #8 0::int)
-#255 := (ite #12 #248 #254)
-#247 := (= #8 #29)
-#13 := (= #9 0::int)
-#256 := (ite #13 #247 #255)
-#742 := (forall (vars (?v0 int) (?v1 int)) (:pat #741) #256)
-#259 := (forall (vars (?v0 int) (?v1 int)) #256)
-#745 := (iff #259 #742)
-#743 := (iff #256 #256)
-#744 := [refl]: #743
-#746 := [quant-intro #744]: #745
-#130 := (* -1::int #124)
-#216 := (ite #198 #30 #130)
-#219 := (ite #12 0::int #216)
-#222 := (ite #13 #8 #219)
-#225 := (= #29 #222)
-#228 := (forall (vars (?v0 int) (?v1 int)) #225)
-#260 := (iff #228 #259)
-#257 := (iff #225 #256)
-#258 := [rewrite]: #257
-#261 := [quant-intro #258]: #260
-#99 := (not #98)
-#92 := (not #91)
-#102 := (and #92 #99)
-#88 := (not #87)
-#95 := (and #88 #92)
-#105 := (or #95 #102)
-#150 := (ite #105 #30 #130)
-#153 := (ite #12 0::int #150)
-#156 := (ite #13 #8 #153)
-#159 := (= #29 #156)
-#162 := (forall (vars (?v0 int) (?v1 int)) #159)
-#229 := (iff #162 #228)
-#226 := (iff #159 #225)
-#223 := (= #156 #222)
-#220 := (= #153 #219)
-#217 := (= #150 #216)
-#201 := (iff #105 #198)
-#195 := (or #192 #184)
-#199 := (iff #195 #198)
-#200 := [rewrite]: #199
-#196 := (iff #105 #195)
-#193 := (iff #102 #184)
-#194 := [rewrite]: #193
-#181 := (iff #95 #192)
-#182 := [rewrite]: #181
-#197 := [monotonicity #182 #194]: #196
-#202 := [trans #197 #200]: #201
-#218 := [monotonicity #202]: #217
-#221 := [monotonicity #218]: #220
-#224 := [monotonicity #221]: #223
-#227 := [monotonicity #224]: #226
-#230 := [quant-intro #227]: #229
-#179 := (~ #162 #162)
-#175 := (~ #159 #159)
-#189 := [refl]: #175
-#180 := [nnf-pos #189]: #179
-#23 := (- #9)
-#22 := (- #8)
-#31 := (mod #22 #23)
-#32 := (- #31)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#33 := (ite #20 #30 #32)
-#34 := (ite #12 0::int #33)
-#35 := (ite #13 #8 #34)
-#36 := (= #29 #35)
-#37 := (forall (vars (?v0 int) (?v1 int)) #36)
-#165 := (iff #37 #162)
-#59 := (and #16 #18)
-#62 := (or #17 #59)
-#135 := (ite #62 #30 #130)
-#138 := (ite #12 0::int #135)
-#141 := (ite #13 #8 #138)
-#144 := (= #29 #141)
-#147 := (forall (vars (?v0 int) (?v1 int)) #144)
-#163 := (iff #147 #162)
-#160 := (iff #144 #159)
-#157 := (= #141 #156)
-#154 := (= #138 #153)
-#151 := (= #135 #150)
-#106 := (iff #62 #105)
-#103 := (iff #59 #102)
-#100 := (iff #18 #99)
-#101 := [rewrite]: #100
-#93 := (iff #16 #92)
-#94 := [rewrite]: #93
-#104 := [monotonicity #94 #101]: #103
-#96 := (iff #17 #95)
-#89 := (iff #15 #88)
-#90 := [rewrite]: #89
-#97 := [monotonicity #90 #94]: #96
-#107 := [monotonicity #97 #104]: #106
-#152 := [monotonicity #107]: #151
-#155 := [monotonicity #152]: #154
-#158 := [monotonicity #155]: #157
-#161 := [monotonicity #158]: #160
-#164 := [quant-intro #161]: #163
-#148 := (iff #37 #147)
-#145 := (iff #36 #144)
-#142 := (= #35 #141)
-#139 := (= #34 #138)
-#136 := (= #33 #135)
-#133 := (= #32 #130)
-#127 := (- #124)
-#131 := (= #127 #130)
-#132 := [rewrite]: #131
-#128 := (= #32 #127)
-#125 := (= #31 #124)
-#70 := (= #23 #69)
-#71 := [rewrite]: #70
-#67 := (= #22 #66)
-#68 := [rewrite]: #67
-#126 := [monotonicity #68 #71]: #125
-#129 := [monotonicity #126]: #128
-#134 := [trans #129 #132]: #133
-#63 := (iff #20 #62)
-#60 := (iff #19 #59)
-#61 := [rewrite]: #60
-#64 := [monotonicity #61]: #63
-#137 := [monotonicity #64 #134]: #136
-#140 := [monotonicity #137]: #139
-#143 := [monotonicity #140]: #142
-#146 := [monotonicity #143]: #145
-#149 := [quant-intro #146]: #148
-#166 := [trans #149 #164]: #165
-#123 := [asserted]: #37
-#167 := [mp #123 #166]: #162
-#190 := [mp~ #167 #180]: #162
-#231 := [mp #190 #230]: #228
-#262 := [mp #231 #261]: #259
-#747 := [mp #262 #746]: #742
-#535 := (not #742)
-#536 := (or #535 #169)
-#320 := (* -1::int 0::int)
-#405 := (* -1::int f5)
-#406 := (mod #405 #320)
-#407 := (+ #39 #406)
-#398 := (= #407 0::int)
-#409 := (mod f5 0::int)
-#410 := (* -1::int #409)
-#338 := (+ #39 #410)
-#411 := (= #338 0::int)
-#408 := (<= 0::int 0::int)
-#412 := (<= f5 0::int)
-#391 := (or #412 #408)
-#726 := (not #391)
-#728 := (>= f5 0::int)
-#385 := (or #408 #728)
-#515 := (not #385)
-#722 := (or #515 #726)
-#396 := (ite #722 #411 #398)
-#397 := (= #39 0::int)
-#729 := (= f5 0::int)
-#730 := (ite #729 #397 #396)
-#731 := (= 0::int 0::int)
-#732 := (ite #731 #169 #730)
-#537 := (or #535 #732)
-#680 := (iff #537 #536)
-#682 := (iff #536 #536)
-#676 := [rewrite]: #682
-#688 := (iff #732 #169)
-#426 := (mod #405 0::int)
-#705 := (+ #39 #426)
-#416 := (= #705 0::int)
-#700 := (ite #729 #397 #416)
-#1 := true
-#691 := (ite true #169 #700)
-#692 := (iff #691 #169)
-#693 := [rewrite]: #692
-#686 := (iff #732 #691)
-#689 := (iff #730 #700)
-#699 := (iff #396 #416)
-#419 := (ite false #411 #416)
-#413 := (iff #419 #416)
-#695 := [rewrite]: #413
-#697 := (iff #396 #419)
-#417 := (iff #398 #416)
-#702 := (= #407 #705)
-#427 := (= #406 #426)
-#703 := (= #320 0::int)
-#704 := [rewrite]: #703
-#701 := [monotonicity #704]: #427
-#706 := [monotonicity #701]: #702
-#418 := [monotonicity #706]: #417
-#433 := (iff #722 false)
-#707 := (or false false)
-#431 := (iff #707 false)
-#432 := [rewrite]: #431
-#708 := (iff #722 #707)
-#718 := (iff #726 false)
-#373 := (not true)
-#711 := (iff #373 false)
-#712 := [rewrite]: #711
-#360 := (iff #726 #373)
-#719 := (iff #391 true)
-#715 := (or #412 true)
-#354 := (iff #715 true)
-#717 := [rewrite]: #354
-#710 := (iff #391 #715)
-#723 := (iff #408 true)
-#383 := [rewrite]: #723
-#716 := [monotonicity #383]: #710
-#359 := [trans #716 #717]: #719
-#720 := [monotonicity #359]: #360
-#721 := [trans #720 #712]: #718
-#713 := (iff #515 false)
-#374 := (iff #515 #373)
-#709 := (iff #385 true)
-#388 := (or true #728)
-#725 := (iff #388 true)
-#368 := [rewrite]: #725
-#724 := (iff #385 #388)
-#384 := [monotonicity #383]: #724
-#372 := [trans #384 #368]: #709
-#375 := [monotonicity #372]: #374
-#714 := [trans #375 #712]: #713
-#430 := [monotonicity #714 #721]: #708
-#326 := [trans #430 #432]: #433
-#698 := [monotonicity #326 #418]: #697
-#696 := [trans #698 #695]: #699
-#690 := [monotonicity #696]: #689
-#727 := (iff #731 true)
-#733 := [rewrite]: #727
-#687 := [monotonicity #733 #690]: #686
-#694 := [trans #687 #693]: #688
-#681 := [monotonicity #694]: #680
-#677 := [trans #681 #676]: #680
-#679 := [quant-inst]: #537
-#683 := [mp #679 #677]: #536
-[unit-resolution #683 #747 #177]: false
-unsat
-a0dd84d240eccf81ab65033fdd888802845765bd 281 0
+9c82ae76cec416fd404b909bc700a14311ef9064 281 0
 #2 := false
 #38 := 1::int
 decl f3 :: (-> int int int)
@@ -28565,306 +28227,7 @@
 #667 := [mp #573 #670]: #572
 [unit-resolution #667 #742 #168]: false
 unsat
-1028cda2400990738320cbc2fe031f4833efde19 298 0
-#2 := false
-#11 := 0::int
-decl f4 :: (-> int int int)
-#38 := 1::int
-#39 := (f4 0::int 1::int)
-#40 := (= #39 0::int)
-#41 := (not #40)
-#168 := [asserted]: #41
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#29 := (f4 #8 #9)
-#743 := (pattern #29)
-#65 := -1::int
-#69 := (* -1::int #9)
-#66 := (* -1::int #8)
-#124 := (mod #66 #69)
-#255 := (+ #29 #124)
-#256 := (= #255 0::int)
-#30 := (mod #8 #9)
-#252 := (* -1::int #30)
-#253 := (+ #29 #252)
-#254 := (= #253 0::int)
-#91 := (<= #9 0::int)
-#87 := (<= #8 0::int)
-#194 := (or #87 #91)
-#195 := (not #194)
-#98 := (>= #8 0::int)
-#186 := (or #91 #98)
-#187 := (not #186)
-#201 := (or #187 #195)
-#257 := (ite #201 #254 #256)
-#251 := (= #29 0::int)
-#12 := (= #8 0::int)
-#258 := (ite #12 #251 #257)
-#250 := (= #8 #29)
-#13 := (= #9 0::int)
-#259 := (ite #13 #250 #258)
-#744 := (forall (vars (?v0 int) (?v1 int)) (:pat #743) #259)
-#262 := (forall (vars (?v0 int) (?v1 int)) #259)
-#747 := (iff #262 #744)
-#745 := (iff #259 #259)
-#746 := [refl]: #745
-#748 := [quant-intro #746]: #747
-#130 := (* -1::int #124)
-#219 := (ite #201 #30 #130)
-#222 := (ite #12 0::int #219)
-#225 := (ite #13 #8 #222)
-#228 := (= #29 #225)
-#231 := (forall (vars (?v0 int) (?v1 int)) #228)
-#263 := (iff #231 #262)
-#260 := (iff #228 #259)
-#261 := [rewrite]: #260
-#264 := [quant-intro #261]: #263
-#99 := (not #98)
-#92 := (not #91)
-#102 := (and #92 #99)
-#88 := (not #87)
-#95 := (and #88 #92)
-#105 := (or #95 #102)
-#150 := (ite #105 #30 #130)
-#153 := (ite #12 0::int #150)
-#156 := (ite #13 #8 #153)
-#159 := (= #29 #156)
-#162 := (forall (vars (?v0 int) (?v1 int)) #159)
-#232 := (iff #162 #231)
-#229 := (iff #159 #228)
-#226 := (= #156 #225)
-#223 := (= #153 #222)
-#220 := (= #150 #219)
-#204 := (iff #105 #201)
-#198 := (or #195 #187)
-#202 := (iff #198 #201)
-#203 := [rewrite]: #202
-#199 := (iff #105 #198)
-#196 := (iff #102 #187)
-#197 := [rewrite]: #196
-#184 := (iff #95 #195)
-#185 := [rewrite]: #184
-#200 := [monotonicity #185 #197]: #199
-#205 := [trans #200 #203]: #204
-#221 := [monotonicity #205]: #220
-#224 := [monotonicity #221]: #223
-#227 := [monotonicity #224]: #226
-#230 := [monotonicity #227]: #229
-#233 := [quant-intro #230]: #232
-#182 := (~ #162 #162)
-#179 := (~ #159 #159)
-#192 := [refl]: #179
-#183 := [nnf-pos #192]: #182
-#23 := (- #9)
-#22 := (- #8)
-#31 := (mod #22 #23)
-#32 := (- #31)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#33 := (ite #20 #30 #32)
-#34 := (ite #12 0::int #33)
-#35 := (ite #13 #8 #34)
-#36 := (= #29 #35)
-#37 := (forall (vars (?v0 int) (?v1 int)) #36)
-#165 := (iff #37 #162)
-#59 := (and #16 #18)
-#62 := (or #17 #59)
-#135 := (ite #62 #30 #130)
-#138 := (ite #12 0::int #135)
-#141 := (ite #13 #8 #138)
-#144 := (= #29 #141)
-#147 := (forall (vars (?v0 int) (?v1 int)) #144)
-#163 := (iff #147 #162)
-#160 := (iff #144 #159)
-#157 := (= #141 #156)
-#154 := (= #138 #153)
-#151 := (= #135 #150)
-#106 := (iff #62 #105)
-#103 := (iff #59 #102)
-#100 := (iff #18 #99)
-#101 := [rewrite]: #100
-#93 := (iff #16 #92)
-#94 := [rewrite]: #93
-#104 := [monotonicity #94 #101]: #103
-#96 := (iff #17 #95)
-#89 := (iff #15 #88)
-#90 := [rewrite]: #89
-#97 := [monotonicity #90 #94]: #96
-#107 := [monotonicity #97 #104]: #106
-#152 := [monotonicity #107]: #151
-#155 := [monotonicity #152]: #154
-#158 := [monotonicity #155]: #157
-#161 := [monotonicity #158]: #160
-#164 := [quant-intro #161]: #163
-#148 := (iff #37 #147)
-#145 := (iff #36 #144)
-#142 := (= #35 #141)
-#139 := (= #34 #138)
-#136 := (= #33 #135)
-#133 := (= #32 #130)
-#127 := (- #124)
-#131 := (= #127 #130)
-#132 := [rewrite]: #131
-#128 := (= #32 #127)
-#125 := (= #31 #124)
-#70 := (= #23 #69)
-#71 := [rewrite]: #70
-#67 := (= #22 #66)
-#68 := [rewrite]: #67
-#126 := [monotonicity #68 #71]: #125
-#129 := [monotonicity #126]: #128
-#134 := [trans #129 #132]: #133
-#63 := (iff #20 #62)
-#60 := (iff #19 #59)
-#61 := [rewrite]: #60
-#64 := [monotonicity #61]: #63
-#137 := [monotonicity #64 #134]: #136
-#140 := [monotonicity #137]: #139
-#143 := [monotonicity #140]: #142
-#146 := [monotonicity #143]: #145
-#149 := [quant-intro #146]: #148
-#166 := [trans #149 #164]: #165
-#123 := [asserted]: #37
-#167 := [mp #123 #166]: #162
-#193 := [mp~ #167 #183]: #162
-#234 := [mp #193 #233]: #231
-#265 := [mp #234 #264]: #262
-#749 := [mp #265 #748]: #744
-#665 := (not #744)
-#666 := (or #665 #40)
-#323 := (* -1::int 1::int)
-#407 := (* -1::int 0::int)
-#408 := (mod #407 #323)
-#409 := (+ #39 #408)
-#400 := (= #409 0::int)
-#411 := (mod 0::int 1::int)
-#412 := (* -1::int #411)
-#413 := (+ #39 #412)
-#410 := (= #413 0::int)
-#414 := (<= 1::int 0::int)
-#393 := (<= 0::int 0::int)
-#728 := (or #393 #414)
-#730 := (not #728)
-#387 := (>= 0::int 0::int)
-#517 := (or #414 #387)
-#724 := (not #517)
-#398 := (or #724 #730)
-#399 := (ite #398 #410 #400)
-#731 := (= 0::int 0::int)
-#732 := (ite #731 #40 #399)
-#169 := (= 0::int #39)
-#733 := (= 1::int 0::int)
-#734 := (ite #733 #169 #732)
-#669 := (or #665 #734)
-#569 := (iff #669 #666)
-#572 := (iff #666 #666)
-#565 := [rewrite]: #572
-#668 := (iff #734 #40)
-#686 := (ite false #40 #40)
-#516 := (iff #686 #40)
-#518 := [rewrite]: #516
-#561 := (iff #734 #686)
-#559 := (iff #732 #40)
-#1 := true
-#673 := (ite true #40 #40)
-#674 := (iff #673 #40)
-#677 := [rewrite]: #674
-#675 := (iff #732 #673)
-#519 := (iff #399 #40)
-#680 := (iff #399 #686)
-#679 := (iff #400 #40)
-#684 := (= #409 #39)
-#415 := (+ #39 0::int)
-#698 := (= #415 #39)
-#702 := [rewrite]: #698
-#682 := (= #409 #415)
-#539 := (= #408 0::int)
-#695 := (mod 0::int -1::int)
-#537 := (= #695 0::int)
-#538 := [rewrite]: #537
-#690 := (= #408 #695)
-#689 := (= #323 -1::int)
-#694 := [rewrite]: #689
-#420 := (= #407 0::int)
-#421 := [rewrite]: #420
-#696 := [monotonicity #421 #694]: #690
-#681 := [trans #696 #538]: #539
-#683 := [monotonicity #681]: #682
-#678 := [trans #683 #702]: #684
-#685 := [monotonicity #678]: #679
-#693 := (iff #410 #40)
-#691 := (= #413 #39)
-#697 := (= #413 #415)
-#699 := (= #412 0::int)
-#418 := (= #412 #407)
-#704 := (= #411 0::int)
-#708 := [rewrite]: #704
-#419 := [monotonicity #708]: #418
-#700 := [trans #419 #421]: #699
-#701 := [monotonicity #700]: #697
-#692 := [trans #701 #702]: #691
-#688 := [monotonicity #692]: #693
-#703 := (iff #398 false)
-#329 := (or false false)
-#428 := (iff #329 false)
-#429 := [rewrite]: #428
-#705 := (iff #398 #329)
-#434 := (iff #730 false)
-#714 := (not true)
-#717 := (iff #714 false)
-#712 := [rewrite]: #717
-#432 := (iff #730 #714)
-#709 := (iff #728 true)
-#361 := (or true false)
-#720 := (iff #361 true)
-#723 := [rewrite]: #720
-#362 := (iff #728 #361)
-#390 := (iff #414 false)
-#726 := [rewrite]: #390
-#719 := (iff #393 true)
-#721 := [rewrite]: #719
-#722 := [monotonicity #721 #726]: #362
-#710 := [trans #722 #723]: #709
-#433 := [monotonicity #710]: #432
-#435 := [trans #433 #712]: #434
-#718 := (iff #724 false)
-#715 := (iff #724 #714)
-#377 := (iff #517 true)
-#370 := (or false true)
-#375 := (iff #370 true)
-#376 := [rewrite]: #375
-#711 := (iff #517 #370)
-#386 := (iff #387 true)
-#727 := [rewrite]: #386
-#374 := [monotonicity #726 #727]: #711
-#713 := [trans #374 #376]: #377
-#716 := [monotonicity #713]: #715
-#356 := [trans #716 #712]: #718
-#706 := [monotonicity #356 #435]: #705
-#707 := [trans #706 #429]: #703
-#687 := [monotonicity #707 #688 #685]: #680
-#672 := [trans #687 #518]: #519
-#725 := (iff #731 true)
-#385 := [rewrite]: #725
-#676 := [monotonicity #385 #672]: #675
-#560 := [trans #676 #677]: #559
-#175 := (iff #169 #40)
-#176 := [rewrite]: #175
-#729 := (iff #733 false)
-#735 := [rewrite]: #729
-#520 := [monotonicity #735 #176 #560]: #561
-#570 := [trans #520 #518]: #668
-#571 := [monotonicity #570]: #569
-#573 := [trans #571 #565]: #569
-#554 := [quant-inst]: #669
-#574 := [mp #554 #573]: #666
-[unit-resolution #574 #749 #168]: false
-unsat
-517d94dcf477960e98fa3a6ad1a7d3b1e41941d9 292 0
+892e25c8bdf423f86a8231d77f83b5c173fd91c0 292 0
 #2 := false
 #38 := 3::int
 decl f3 :: (-> int int int)
@@ -29157,304 +28520,7 @@
 #660 := [mp #653 #664]: #659
 [unit-resolution #660 #743 #169]: false
 unsat
-4c32c3169bc41583f37f35faf2502f1708b05bd0 296 0
-#2 := false
-#11 := 0::int
-decl f4 :: (-> int int int)
-#38 := 1::int
-#39 := (f4 1::int 1::int)
-#40 := (= #39 0::int)
-#41 := (not #40)
-#168 := [asserted]: #41
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#29 := (f4 #8 #9)
-#743 := (pattern #29)
-#65 := -1::int
-#69 := (* -1::int #9)
-#66 := (* -1::int #8)
-#124 := (mod #66 #69)
-#255 := (+ #29 #124)
-#256 := (= #255 0::int)
-#30 := (mod #8 #9)
-#252 := (* -1::int #30)
-#253 := (+ #29 #252)
-#254 := (= #253 0::int)
-#91 := (<= #9 0::int)
-#87 := (<= #8 0::int)
-#194 := (or #87 #91)
-#195 := (not #194)
-#98 := (>= #8 0::int)
-#186 := (or #91 #98)
-#187 := (not #186)
-#201 := (or #187 #195)
-#257 := (ite #201 #254 #256)
-#251 := (= #29 0::int)
-#12 := (= #8 0::int)
-#258 := (ite #12 #251 #257)
-#250 := (= #8 #29)
-#13 := (= #9 0::int)
-#259 := (ite #13 #250 #258)
-#744 := (forall (vars (?v0 int) (?v1 int)) (:pat #743) #259)
-#262 := (forall (vars (?v0 int) (?v1 int)) #259)
-#747 := (iff #262 #744)
-#745 := (iff #259 #259)
-#746 := [refl]: #745
-#748 := [quant-intro #746]: #747
-#130 := (* -1::int #124)
-#219 := (ite #201 #30 #130)
-#222 := (ite #12 0::int #219)
-#225 := (ite #13 #8 #222)
-#228 := (= #29 #225)
-#231 := (forall (vars (?v0 int) (?v1 int)) #228)
-#263 := (iff #231 #262)
-#260 := (iff #228 #259)
-#261 := [rewrite]: #260
-#264 := [quant-intro #261]: #263
-#99 := (not #98)
-#92 := (not #91)
-#102 := (and #92 #99)
-#88 := (not #87)
-#95 := (and #88 #92)
-#105 := (or #95 #102)
-#150 := (ite #105 #30 #130)
-#153 := (ite #12 0::int #150)
-#156 := (ite #13 #8 #153)
-#159 := (= #29 #156)
-#162 := (forall (vars (?v0 int) (?v1 int)) #159)
-#232 := (iff #162 #231)
-#229 := (iff #159 #228)
-#226 := (= #156 #225)
-#223 := (= #153 #222)
-#220 := (= #150 #219)
-#204 := (iff #105 #201)
-#198 := (or #195 #187)
-#202 := (iff #198 #201)
-#203 := [rewrite]: #202
-#199 := (iff #105 #198)
-#196 := (iff #102 #187)
-#197 := [rewrite]: #196
-#184 := (iff #95 #195)
-#185 := [rewrite]: #184
-#200 := [monotonicity #185 #197]: #199
-#205 := [trans #200 #203]: #204
-#221 := [monotonicity #205]: #220
-#224 := [monotonicity #221]: #223
-#227 := [monotonicity #224]: #226
-#230 := [monotonicity #227]: #229
-#233 := [quant-intro #230]: #232
-#182 := (~ #162 #162)
-#179 := (~ #159 #159)
-#192 := [refl]: #179
-#183 := [nnf-pos #192]: #182
-#23 := (- #9)
-#22 := (- #8)
-#31 := (mod #22 #23)
-#32 := (- #31)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#33 := (ite #20 #30 #32)
-#34 := (ite #12 0::int #33)
-#35 := (ite #13 #8 #34)
-#36 := (= #29 #35)
-#37 := (forall (vars (?v0 int) (?v1 int)) #36)
-#165 := (iff #37 #162)
-#59 := (and #16 #18)
-#62 := (or #17 #59)
-#135 := (ite #62 #30 #130)
-#138 := (ite #12 0::int #135)
-#141 := (ite #13 #8 #138)
-#144 := (= #29 #141)
-#147 := (forall (vars (?v0 int) (?v1 int)) #144)
-#163 := (iff #147 #162)
-#160 := (iff #144 #159)
-#157 := (= #141 #156)
-#154 := (= #138 #153)
-#151 := (= #135 #150)
-#106 := (iff #62 #105)
-#103 := (iff #59 #102)
-#100 := (iff #18 #99)
-#101 := [rewrite]: #100
-#93 := (iff #16 #92)
-#94 := [rewrite]: #93
-#104 := [monotonicity #94 #101]: #103
-#96 := (iff #17 #95)
-#89 := (iff #15 #88)
-#90 := [rewrite]: #89
-#97 := [monotonicity #90 #94]: #96
-#107 := [monotonicity #97 #104]: #106
-#152 := [monotonicity #107]: #151
-#155 := [monotonicity #152]: #154
-#158 := [monotonicity #155]: #157
-#161 := [monotonicity #158]: #160
-#164 := [quant-intro #161]: #163
-#148 := (iff #37 #147)
-#145 := (iff #36 #144)
-#142 := (= #35 #141)
-#139 := (= #34 #138)
-#136 := (= #33 #135)
-#133 := (= #32 #130)
-#127 := (- #124)
-#131 := (= #127 #130)
-#132 := [rewrite]: #131
-#128 := (= #32 #127)
-#125 := (= #31 #124)
-#70 := (= #23 #69)
-#71 := [rewrite]: #70
-#67 := (= #22 #66)
-#68 := [rewrite]: #67
-#126 := [monotonicity #68 #71]: #125
-#129 := [monotonicity #126]: #128
-#134 := [trans #129 #132]: #133
-#63 := (iff #20 #62)
-#60 := (iff #19 #59)
-#61 := [rewrite]: #60
-#64 := [monotonicity #61]: #63
-#137 := [monotonicity #64 #134]: #136
-#140 := [monotonicity #137]: #139
-#143 := [monotonicity #140]: #142
-#146 := [monotonicity #143]: #145
-#149 := [quant-intro #146]: #148
-#166 := [trans #149 #164]: #165
-#123 := [asserted]: #37
-#167 := [mp #123 #166]: #162
-#193 := [mp~ #167 #183]: #162
-#234 := [mp #193 #233]: #231
-#265 := [mp #234 #264]: #262
-#749 := [mp #265 #748]: #744
-#666 := (not #744)
-#669 := (or #666 #40)
-#323 := (* -1::int 1::int)
-#407 := (mod #323 #323)
-#408 := (+ #39 #407)
-#409 := (= #408 0::int)
-#400 := (mod 1::int 1::int)
-#411 := (* -1::int #400)
-#412 := (+ #39 #411)
-#413 := (= #412 0::int)
-#410 := (<= 1::int 0::int)
-#414 := (or #410 #410)
-#393 := (not #414)
-#728 := (>= 1::int 0::int)
-#730 := (or #410 #728)
-#387 := (not #730)
-#517 := (or #387 #393)
-#724 := (ite #517 #413 #409)
-#398 := (= 1::int 0::int)
-#399 := (ite #398 #40 #724)
-#731 := (= 1::int #39)
-#732 := (ite #398 #731 #399)
-#554 := (or #666 #732)
-#571 := (iff #554 #669)
-#565 := (iff #669 #669)
-#573 := [rewrite]: #565
-#570 := (iff #732 #40)
-#735 := (= #39 1::int)
-#559 := (ite false #735 #40)
-#520 := (iff #559 #40)
-#668 := [rewrite]: #520
-#560 := (iff #732 #559)
-#674 := (iff #399 #40)
-#519 := (ite false #40 #40)
-#675 := (iff #519 #40)
-#676 := [rewrite]: #675
-#672 := (iff #399 #519)
-#516 := (iff #724 #40)
-#1 := true
-#679 := (ite true #40 #40)
-#680 := (iff #679 #40)
-#687 := [rewrite]: #680
-#685 := (iff #724 #679)
-#684 := (iff #409 #40)
-#682 := (= #408 #39)
-#699 := (+ #39 0::int)
-#697 := (= #699 #39)
-#701 := [rewrite]: #697
-#539 := (= #408 #699)
-#537 := (= #407 0::int)
-#689 := (mod -1::int -1::int)
-#690 := (= #689 0::int)
-#696 := [rewrite]: #690
-#694 := (= #407 #689)
-#693 := (= #323 -1::int)
-#688 := [rewrite]: #693
-#695 := [monotonicity #688 #688]: #694
-#538 := [trans #695 #696]: #537
-#681 := [monotonicity #538]: #539
-#683 := [trans #681 #701]: #682
-#678 := [monotonicity #683]: #684
-#691 := (iff #413 #40)
-#698 := (= #412 #39)
-#700 := (= #412 #699)
-#420 := (= #411 0::int)
-#707 := (* -1::int 0::int)
-#418 := (= #707 0::int)
-#419 := [rewrite]: #418
-#704 := (= #411 #707)
-#429 := (= #400 0::int)
-#703 := [rewrite]: #429
-#708 := [monotonicity #703]: #704
-#421 := [trans #708 #419]: #420
-#415 := [monotonicity #421]: #700
-#702 := [trans #415 #701]: #698
-#692 := [monotonicity #702]: #691
-#706 := (iff #517 true)
-#727 := (or false true)
-#374 := (iff #727 true)
-#375 := [rewrite]: #374
-#329 := (iff #517 #727)
-#434 := (iff #393 true)
-#723 := (not false)
-#432 := (iff #723 true)
-#433 := [rewrite]: #432
-#709 := (iff #393 #723)
-#722 := (iff #414 false)
-#356 := (or false false)
-#361 := (iff #356 false)
-#362 := [rewrite]: #361
-#719 := (iff #414 #356)
-#385 := (iff #410 false)
-#390 := [rewrite]: #385
-#721 := [monotonicity #390 #390]: #719
-#720 := [trans #721 #362]: #722
-#710 := [monotonicity #720]: #709
-#435 := [trans #710 #433]: #434
-#712 := (iff #387 false)
-#713 := (not true)
-#716 := (iff #713 false)
-#717 := [rewrite]: #716
-#714 := (iff #387 #713)
-#376 := (iff #730 true)
-#370 := (iff #730 #727)
-#726 := (iff #728 true)
-#386 := [rewrite]: #726
-#711 := [monotonicity #390 #386]: #370
-#377 := [trans #711 #375]: #376
-#715 := [monotonicity #377]: #714
-#718 := [trans #715 #717]: #712
-#705 := [monotonicity #718 #435]: #329
-#428 := [trans #705 #375]: #706
-#686 := [monotonicity #428 #692 #678]: #685
-#518 := [trans #686 #687]: #516
-#733 := (iff #398 false)
-#734 := [rewrite]: #733
-#673 := [monotonicity #734 #518]: #672
-#677 := [trans #673 #676]: #674
-#729 := (iff #731 #735)
-#725 := [rewrite]: #729
-#561 := [monotonicity #734 #725 #677]: #560
-#665 := [trans #561 #668]: #570
-#572 := [monotonicity #665]: #571
-#574 := [trans #572 #573]: #571
-#569 := [quant-inst]: #554
-#575 := [mp #569 #574]: #669
-[unit-resolution #575 #749 #168]: false
-unsat
-7f4e87626d8d9f850815ab9f11ca7ece76ac9405 335 0
+dcc2c3c66de3e673b0cfeb392eb7d2e8ec04caf3 335 0
 #2 := false
 decl f3 :: (-> int int int)
 #39 := 1::int
@@ -29790,7 +28856,7 @@
 #500 := [mp #530 #497]: #170
 [unit-resolution #178 #500]: false
 unsat
-3dd3e5e067f99d1c90bd9e0ee3fa61701d0112dd 306 0
+2a60c54232ee9900c0a7254451730bed52200c08 306 0
 #2 := false
 #11 := 0::int
 decl f3 :: (-> int int int)
@@ -30097,315 +29163,7 @@
 #682 := [mp #588 #685]: #587
 [unit-resolution #682 #757 #193]: false
 unsat
-8c6b824970f94cadce3aa1f861ee0fe83c8cdcae 307 0
-#2 := false
-#11 := 0::int
-decl f4 :: (-> int int int)
-#39 := 1::int
-#38 := 3::int
-#40 := (f4 3::int 1::int)
-#41 := (= #40 0::int)
-#42 := (not #41)
-#169 := [asserted]: #42
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#29 := (f4 #8 #9)
-#744 := (pattern #29)
-#66 := -1::int
-#70 := (* -1::int #9)
-#67 := (* -1::int #8)
-#125 := (mod #67 #70)
-#256 := (+ #29 #125)
-#257 := (= #256 0::int)
-#30 := (mod #8 #9)
-#253 := (* -1::int #30)
-#254 := (+ #29 #253)
-#255 := (= #254 0::int)
-#92 := (<= #9 0::int)
-#88 := (<= #8 0::int)
-#195 := (or #88 #92)
-#196 := (not #195)
-#99 := (>= #8 0::int)
-#187 := (or #92 #99)
-#188 := (not #187)
-#202 := (or #188 #196)
-#258 := (ite #202 #255 #257)
-#252 := (= #29 0::int)
-#12 := (= #8 0::int)
-#259 := (ite #12 #252 #258)
-#251 := (= #8 #29)
-#13 := (= #9 0::int)
-#260 := (ite #13 #251 #259)
-#745 := (forall (vars (?v0 int) (?v1 int)) (:pat #744) #260)
-#263 := (forall (vars (?v0 int) (?v1 int)) #260)
-#748 := (iff #263 #745)
-#746 := (iff #260 #260)
-#747 := [refl]: #746
-#749 := [quant-intro #747]: #748
-#131 := (* -1::int #125)
-#220 := (ite #202 #30 #131)
-#223 := (ite #12 0::int #220)
-#226 := (ite #13 #8 #223)
-#229 := (= #29 #226)
-#232 := (forall (vars (?v0 int) (?v1 int)) #229)
-#264 := (iff #232 #263)
-#261 := (iff #229 #260)
-#262 := [rewrite]: #261
-#265 := [quant-intro #262]: #264
-#100 := (not #99)
-#93 := (not #92)
-#103 := (and #93 #100)
-#89 := (not #88)
-#96 := (and #89 #93)
-#106 := (or #96 #103)
-#151 := (ite #106 #30 #131)
-#154 := (ite #12 0::int #151)
-#157 := (ite #13 #8 #154)
-#160 := (= #29 #157)
-#163 := (forall (vars (?v0 int) (?v1 int)) #160)
-#233 := (iff #163 #232)
-#230 := (iff #160 #229)
-#227 := (= #157 #226)
-#224 := (= #154 #223)
-#221 := (= #151 #220)
-#205 := (iff #106 #202)
-#199 := (or #196 #188)
-#203 := (iff #199 #202)
-#204 := [rewrite]: #203
-#200 := (iff #106 #199)
-#197 := (iff #103 #188)
-#198 := [rewrite]: #197
-#185 := (iff #96 #196)
-#186 := [rewrite]: #185
-#201 := [monotonicity #186 #198]: #200
-#206 := [trans #201 #204]: #205
-#222 := [monotonicity #206]: #221
-#225 := [monotonicity #222]: #224
-#228 := [monotonicity #225]: #227
-#231 := [monotonicity #228]: #230
-#234 := [quant-intro #231]: #233
-#183 := (~ #163 #163)
-#180 := (~ #160 #160)
-#193 := [refl]: #180
-#184 := [nnf-pos #193]: #183
-#23 := (- #9)
-#22 := (- #8)
-#31 := (mod #22 #23)
-#32 := (- #31)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#33 := (ite #20 #30 #32)
-#34 := (ite #12 0::int #33)
-#35 := (ite #13 #8 #34)
-#36 := (= #29 #35)
-#37 := (forall (vars (?v0 int) (?v1 int)) #36)
-#166 := (iff #37 #163)
-#60 := (and #16 #18)
-#63 := (or #17 #60)
-#136 := (ite #63 #30 #131)
-#139 := (ite #12 0::int #136)
-#142 := (ite #13 #8 #139)
-#145 := (= #29 #142)
-#148 := (forall (vars (?v0 int) (?v1 int)) #145)
-#164 := (iff #148 #163)
-#161 := (iff #145 #160)
-#158 := (= #142 #157)
-#155 := (= #139 #154)
-#152 := (= #136 #151)
-#107 := (iff #63 #106)
-#104 := (iff #60 #103)
-#101 := (iff #18 #100)
-#102 := [rewrite]: #101
-#94 := (iff #16 #93)
-#95 := [rewrite]: #94
-#105 := [monotonicity #95 #102]: #104
-#97 := (iff #17 #96)
-#90 := (iff #15 #89)
-#91 := [rewrite]: #90
-#98 := [monotonicity #91 #95]: #97
-#108 := [monotonicity #98 #105]: #107
-#153 := [monotonicity #108]: #152
-#156 := [monotonicity #153]: #155
-#159 := [monotonicity #156]: #158
-#162 := [monotonicity #159]: #161
-#165 := [quant-intro #162]: #164
-#149 := (iff #37 #148)
-#146 := (iff #36 #145)
-#143 := (= #35 #142)
-#140 := (= #34 #139)
-#137 := (= #33 #136)
-#134 := (= #32 #131)
-#128 := (- #125)
-#132 := (= #128 #131)
-#133 := [rewrite]: #132
-#129 := (= #32 #128)
-#126 := (= #31 #125)
-#71 := (= #23 #70)
-#72 := [rewrite]: #71
-#68 := (= #22 #67)
-#69 := [rewrite]: #68
-#127 := [monotonicity #69 #72]: #126
-#130 := [monotonicity #127]: #129
-#135 := [trans #130 #133]: #134
-#64 := (iff #20 #63)
-#61 := (iff #19 #60)
-#62 := [rewrite]: #61
-#65 := [monotonicity #62]: #64
-#138 := [monotonicity #65 #135]: #137
-#141 := [monotonicity #138]: #140
-#144 := [monotonicity #141]: #143
-#147 := [monotonicity #144]: #146
-#150 := [quant-intro #147]: #149
-#167 := [trans #150 #165]: #166
-#124 := [asserted]: #37
-#168 := [mp #124 #167]: #163
-#194 := [mp~ #168 #184]: #163
-#235 := [mp #194 #234]: #232
-#266 := [mp #235 #265]: #263
-#750 := [mp #266 #749]: #745
-#577 := (not #745)
-#578 := (or #577 #41)
-#324 := (* -1::int 1::int)
-#408 := (* -1::int 3::int)
-#409 := (mod #408 #324)
-#410 := (+ #40 #409)
-#401 := (= #410 0::int)
-#412 := (mod 3::int 1::int)
-#413 := (* -1::int #412)
-#414 := (+ #40 #413)
-#411 := (= #414 0::int)
-#415 := (<= 1::int 0::int)
-#394 := (<= 3::int 0::int)
-#729 := (or #394 #415)
-#731 := (not #729)
-#388 := (>= 3::int 0::int)
-#518 := (or #415 #388)
-#725 := (not #518)
-#399 := (or #725 #731)
-#400 := (ite #399 #411 #401)
-#732 := (= 3::int 0::int)
-#733 := (ite #732 #41 #400)
-#734 := (= 3::int #40)
-#735 := (= 1::int 0::int)
-#730 := (ite #735 #734 #733)
-#671 := (or #577 #730)
-#672 := (iff #671 #578)
-#661 := (iff #578 #578)
-#653 := [rewrite]: #661
-#575 := (iff #730 #41)
-#391 := (= #40 3::int)
-#570 := (ite false #391 #41)
-#566 := (iff #570 #41)
-#574 := [rewrite]: #566
-#572 := (iff #730 #570)
-#670 := (iff #733 #41)
-#521 := (ite false #41 #41)
-#666 := (iff #521 #41)
-#667 := [rewrite]: #666
-#669 := (iff #733 #521)
-#561 := (iff #400 #41)
-#1 := true
-#676 := (ite true #41 #41)
-#678 := (iff #676 #41)
-#560 := [rewrite]: #678
-#677 := (iff #400 #676)
-#673 := (iff #401 #41)
-#519 := (= #410 #40)
-#692 := (+ #40 0::int)
-#689 := (= #692 #40)
-#690 := [rewrite]: #689
-#688 := (= #410 #692)
-#687 := (= #409 0::int)
-#538 := -3::int
-#684 := (mod -3::int -1::int)
-#680 := (= #684 0::int)
-#686 := [rewrite]: #680
-#685 := (= #409 #684)
-#682 := (= #324 -1::int)
-#683 := [rewrite]: #682
-#539 := (= #408 -3::int)
-#540 := [rewrite]: #539
-#679 := [monotonicity #540 #683]: #685
-#681 := [trans #679 #686]: #687
-#517 := [monotonicity #681]: #688
-#520 := [trans #517 #690]: #519
-#674 := [monotonicity #520]: #673
-#691 := (iff #411 #41)
-#695 := (= #414 #40)
-#693 := (= #414 #692)
-#699 := (= #413 0::int)
-#700 := (* -1::int 0::int)
-#698 := (= #700 0::int)
-#702 := [rewrite]: #698
-#701 := (= #413 #700)
-#421 := (= #412 0::int)
-#422 := [rewrite]: #421
-#416 := [monotonicity #422]: #701
-#703 := [trans #416 #702]: #699
-#694 := [monotonicity #703]: #693
-#696 := [trans #694 #690]: #695
-#697 := [monotonicity #696]: #691
-#419 := (iff #399 true)
-#377 := (or false true)
-#715 := (iff #377 true)
-#716 := [rewrite]: #715
-#705 := (iff #399 #377)
-#704 := (iff #731 true)
-#330 := (not false)
-#429 := (iff #330 true)
-#430 := [rewrite]: #429
-#706 := (iff #731 #330)
-#435 := (iff #729 false)
-#724 := (or false false)
-#433 := (iff #724 false)
-#434 := [rewrite]: #433
-#710 := (iff #729 #724)
-#371 := (iff #415 false)
-#712 := [rewrite]: #371
-#723 := (iff #394 false)
-#721 := [rewrite]: #723
-#711 := [monotonicity #721 #712]: #710
-#436 := [trans #711 #434]: #435
-#707 := [monotonicity #436]: #706
-#708 := [trans #707 #430]: #704
-#362 := (iff #725 false)
-#713 := (not true)
-#720 := (iff #713 false)
-#722 := [rewrite]: #720
-#719 := (iff #725 #713)
-#717 := (iff #518 true)
-#378 := (iff #518 #377)
-#375 := (iff #388 true)
-#376 := [rewrite]: #375
-#714 := [monotonicity #712 #376]: #378
-#718 := [trans #714 #716]: #717
-#357 := [monotonicity #718]: #719
-#363 := [trans #357 #722]: #362
-#709 := [monotonicity #363 #708]: #705
-#420 := [trans #709 #716]: #419
-#675 := [monotonicity #420 #697 #674]: #677
-#562 := [trans #675 #560]: #561
-#387 := (iff #732 false)
-#728 := [rewrite]: #387
-#571 := [monotonicity #728 #562]: #669
-#555 := [trans #571 #667]: #670
-#386 := (iff #734 #391)
-#727 := [rewrite]: #386
-#736 := (iff #735 false)
-#726 := [rewrite]: #736
-#573 := [monotonicity #726 #727 #555]: #572
-#576 := [trans #573 #574]: #575
-#659 := [monotonicity #576]: #672
-#650 := [trans #659 #653]: #672
-#668 := [quant-inst]: #671
-#652 := [mp #668 #650]: #578
-[unit-resolution #652 #750 #169]: false
-unsat
-7415b61df67d0382c7c8726284760c97b73cdf04 315 0
+923c4bec65f1ce0be6fa5d4ec7cf2b004a307f06 315 0
 #2 := false
 #66 := -1::int
 decl f3 :: (-> int int int)
@@ -30721,316 +29479,7 @@
 #678 := [mp #675 #677]: #686
 [unit-resolution #678 #757 #193]: false
 unsat
-b8e3350e9ffae8bc4261f75ecb8e0cef6e684921 308 0
-#2 := false
-#11 := 0::int
-decl f5 :: int
-#38 := f5
-#732 := (= f5 0::int)
-#573 := (not #732)
-#394 := (<= f5 0::int)
-#720 := (not #394)
-#388 := (>= f5 0::int)
-#377 := (not #388)
-#688 := (or #377 #720 #732)
-#575 := (not #688)
-#66 := -1::int
-#408 := (* -1::int f5)
-#700 := (mod #408 -1::int)
-decl f4 :: (-> int int int)
-#39 := 1::int
-#40 := (f4 f5 1::int)
-#698 := (+ #40 #700)
-#703 := (= #698 0::int)
-#41 := (= #40 0::int)
-#520 := (ite #688 #41 #703)
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#29 := (f4 #8 #9)
-#744 := (pattern #29)
-#70 := (* -1::int #9)
-#67 := (* -1::int #8)
-#125 := (mod #67 #70)
-#256 := (+ #29 #125)
-#257 := (= #256 0::int)
-#30 := (mod #8 #9)
-#253 := (* -1::int #30)
-#254 := (+ #29 #253)
-#255 := (= #254 0::int)
-#92 := (<= #9 0::int)
-#88 := (<= #8 0::int)
-#195 := (or #88 #92)
-#196 := (not #195)
-#99 := (>= #8 0::int)
-#187 := (or #92 #99)
-#188 := (not #187)
-#202 := (or #188 #196)
-#258 := (ite #202 #255 #257)
-#252 := (= #29 0::int)
-#12 := (= #8 0::int)
-#259 := (ite #12 #252 #258)
-#251 := (= #8 #29)
-#13 := (= #9 0::int)
-#260 := (ite #13 #251 #259)
-#745 := (forall (vars (?v0 int) (?v1 int)) (:pat #744) #260)
-#263 := (forall (vars (?v0 int) (?v1 int)) #260)
-#748 := (iff #263 #745)
-#746 := (iff #260 #260)
-#747 := [refl]: #746
-#749 := [quant-intro #747]: #748
-#131 := (* -1::int #125)
-#220 := (ite #202 #30 #131)
-#223 := (ite #12 0::int #220)
-#226 := (ite #13 #8 #223)
-#229 := (= #29 #226)
-#232 := (forall (vars (?v0 int) (?v1 int)) #229)
-#264 := (iff #232 #263)
-#261 := (iff #229 #260)
-#262 := [rewrite]: #261
-#265 := [quant-intro #262]: #264
-#100 := (not #99)
-#93 := (not #92)
-#103 := (and #93 #100)
-#89 := (not #88)
-#96 := (and #89 #93)
-#106 := (or #96 #103)
-#151 := (ite #106 #30 #131)
-#154 := (ite #12 0::int #151)
-#157 := (ite #13 #8 #154)
-#160 := (= #29 #157)
-#163 := (forall (vars (?v0 int) (?v1 int)) #160)
-#233 := (iff #163 #232)
-#230 := (iff #160 #229)
-#227 := (= #157 #226)
-#224 := (= #154 #223)
-#221 := (= #151 #220)
-#205 := (iff #106 #202)
-#199 := (or #196 #188)
-#203 := (iff #199 #202)
-#204 := [rewrite]: #203
-#200 := (iff #106 #199)
-#197 := (iff #103 #188)
-#198 := [rewrite]: #197
-#185 := (iff #96 #196)
-#186 := [rewrite]: #185
-#201 := [monotonicity #186 #198]: #200
-#206 := [trans #201 #204]: #205
-#222 := [monotonicity #206]: #221
-#225 := [monotonicity #222]: #224
-#228 := [monotonicity #225]: #227
-#231 := [monotonicity #228]: #230
-#234 := [quant-intro #231]: #233
-#183 := (~ #163 #163)
-#180 := (~ #160 #160)
-#193 := [refl]: #180
-#184 := [nnf-pos #193]: #183
-#23 := (- #9)
-#22 := (- #8)
-#31 := (mod #22 #23)
-#32 := (- #31)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#33 := (ite #20 #30 #32)
-#34 := (ite #12 0::int #33)
-#35 := (ite #13 #8 #34)
-#36 := (= #29 #35)
-#37 := (forall (vars (?v0 int) (?v1 int)) #36)
-#166 := (iff #37 #163)
-#60 := (and #16 #18)
-#63 := (or #17 #60)
-#136 := (ite #63 #30 #131)
-#139 := (ite #12 0::int #136)
-#142 := (ite #13 #8 #139)
-#145 := (= #29 #142)
-#148 := (forall (vars (?v0 int) (?v1 int)) #145)
-#164 := (iff #148 #163)
-#161 := (iff #145 #160)
-#158 := (= #142 #157)
-#155 := (= #139 #154)
-#152 := (= #136 #151)
-#107 := (iff #63 #106)
-#104 := (iff #60 #103)
-#101 := (iff #18 #100)
-#102 := [rewrite]: #101
-#94 := (iff #16 #93)
-#95 := [rewrite]: #94
-#105 := [monotonicity #95 #102]: #104
-#97 := (iff #17 #96)
-#90 := (iff #15 #89)
-#91 := [rewrite]: #90
-#98 := [monotonicity #91 #95]: #97
-#108 := [monotonicity #98 #105]: #107
-#153 := [monotonicity #108]: #152
-#156 := [monotonicity #153]: #155
-#159 := [monotonicity #156]: #158
-#162 := [monotonicity #159]: #161
-#165 := [quant-intro #162]: #164
-#149 := (iff #37 #148)
-#146 := (iff #36 #145)
-#143 := (= #35 #142)
-#140 := (= #34 #139)
-#137 := (= #33 #136)
-#134 := (= #32 #131)
-#128 := (- #125)
-#132 := (= #128 #131)
-#133 := [rewrite]: #132
-#129 := (= #32 #128)
-#126 := (= #31 #125)
-#71 := (= #23 #70)
-#72 := [rewrite]: #71
-#68 := (= #22 #67)
-#69 := [rewrite]: #68
-#127 := [monotonicity #69 #72]: #126
-#130 := [monotonicity #127]: #129
-#135 := [trans #130 #133]: #134
-#64 := (iff #20 #63)
-#61 := (iff #19 #60)
-#62 := [rewrite]: #61
-#65 := [monotonicity #62]: #64
-#138 := [monotonicity #65 #135]: #137
-#141 := [monotonicity #138]: #140
-#144 := [monotonicity #141]: #143
-#147 := [monotonicity #144]: #146
-#150 := [quant-intro #147]: #149
-#167 := [trans #150 #165]: #166
-#124 := [asserted]: #37
-#168 := [mp #124 #167]: #163
-#194 := [mp~ #168 #184]: #163
-#235 := [mp #194 #234]: #232
-#266 := [mp #235 #265]: #263
-#750 := [mp #266 #749]: #745
-#675 := (not #745)
-#678 := (or #675 #520)
-#324 := (* -1::int 1::int)
-#409 := (mod #408 #324)
-#410 := (+ #40 #409)
-#401 := (= #410 0::int)
-#412 := (mod f5 1::int)
-#413 := (* -1::int #412)
-#414 := (+ #40 #413)
-#411 := (= #414 0::int)
-#415 := (<= 1::int 0::int)
-#729 := (or #394 #415)
-#731 := (not #729)
-#518 := (or #415 #388)
-#725 := (not #518)
-#399 := (or #725 #731)
-#400 := (ite #399 #411 #401)
-#733 := (ite #732 #41 #400)
-#734 := (= f5 #40)
-#735 := (= 1::int 0::int)
-#730 := (ite #735 #734 #733)
-#560 := (or #675 #730)
-#562 := (iff #560 #678)
-#669 := (iff #678 #678)
-#571 := [rewrite]: #669
-#676 := (iff #730 #520)
-#363 := (or #377 #720)
-#697 := (or #363 #732)
-#538 := (ite #697 #41 #703)
-#673 := (iff #538 #520)
-#517 := (iff #697 #688)
-#519 := [rewrite]: #517
-#674 := [monotonicity #519]: #673
-#687 := (iff #730 #538)
-#684 := (ite false #734 #538)
-#680 := (iff #684 #538)
-#686 := [rewrite]: #680
-#685 := (iff #730 #684)
-#682 := (iff #733 #538)
-#694 := (ite #363 #41 #703)
-#695 := (ite #732 #41 #694)
-#539 := (iff #695 #538)
-#540 := [rewrite]: #539
-#696 := (iff #733 #695)
-#689 := (iff #400 #694)
-#692 := (iff #401 #703)
-#702 := (= #410 #698)
-#701 := (= #409 #700)
-#421 := (= #324 -1::int)
-#422 := [rewrite]: #421
-#416 := [monotonicity #422]: #701
-#699 := [monotonicity #416]: #702
-#693 := [monotonicity #699]: #692
-#419 := (iff #411 #41)
-#705 := (= #414 #40)
-#707 := (+ #40 0::int)
-#704 := (= #707 #40)
-#708 := [rewrite]: #704
-#429 := (= #414 #707)
-#330 := (= #413 0::int)
-#711 := (* -1::int 0::int)
-#435 := (= #711 0::int)
-#436 := [rewrite]: #435
-#433 := (= #413 #711)
-#724 := (= #412 0::int)
-#710 := [rewrite]: #724
-#434 := [monotonicity #710]: #433
-#706 := [trans #434 #436]: #330
-#430 := [monotonicity #706]: #429
-#709 := [trans #430 #708]: #705
-#420 := [monotonicity #709]: #419
-#723 := (iff #399 #363)
-#722 := (iff #731 #720)
-#719 := (iff #729 #394)
-#715 := (or #394 false)
-#718 := (iff #715 #394)
-#713 := [rewrite]: #718
-#716 := (iff #729 #715)
-#386 := (iff #415 false)
-#391 := [rewrite]: #386
-#717 := [monotonicity #391]: #716
-#357 := [trans #717 #713]: #719
-#362 := [monotonicity #357]: #722
-#378 := (iff #725 #377)
-#375 := (iff #518 #388)
-#727 := (or false #388)
-#371 := (iff #727 #388)
-#712 := [rewrite]: #371
-#387 := (iff #518 #727)
-#728 := [monotonicity #391]: #387
-#376 := [trans #728 #712]: #375
-#714 := [monotonicity #376]: #378
-#721 := [monotonicity #714 #362]: #723
-#690 := [monotonicity #721 #420 #693]: #689
-#691 := [monotonicity #690]: #696
-#683 := [trans #691 #540]: #682
-#736 := (iff #735 false)
-#726 := [rewrite]: #736
-#679 := [monotonicity #726 #683]: #685
-#681 := [trans #679 #686]: #687
-#677 := [trans #681 #674]: #676
-#521 := [monotonicity #677]: #562
-#666 := [trans #521 #571]: #562
-#561 := [quant-inst]: #560
-#667 := [mp #561 #666]: #678
-#660 := [unit-resolution #667 #750]: #520
-#668 := (not #520)
-#665 := (or #668 #575)
-#42 := (not #41)
-#169 := [asserted]: #42
-#672 := (or #668 #575 #41)
-#659 := [def-axiom]: #672
-#654 := [unit-resolution #659 #169]: #665
-#655 := [unit-resolution #654 #660]: #575
-#566 := (or #688 #573)
-#574 := [def-axiom]: #566
-#656 := [unit-resolution #574 #655]: #573
-#670 := (or #688 #388)
-#555 := [def-axiom]: #670
-#657 := [unit-resolution #555 #655]: #388
-#570 := (or #688 #394)
-#572 := [def-axiom]: #570
-#651 := [unit-resolution #572 #655]: #394
-#658 := (or #732 #720 #377)
-#642 := [th-lemma]: #658
-[unit-resolution #642 #651 #657 #656]: false
-unsat
-8385e60ed3af381853b64133bc38d95658f35f76 311 0
+f6b6632f2c5a6045abfce98be6bbc967a6c8ca72 311 0
 #2 := false
 #177 := -3::int
 decl f3 :: (-> int int int)
@@ -31342,325 +29791,7 @@
 #661 := [mp #659 #672]: #660
 [unit-resolution #661 #750 #188]: false
 unsat
-62cbad5f8230a32f1b29d792652e1e27f6a3a147 317 0
-#2 := false
-#11 := 0::int
-decl f4 :: (-> int int int)
-#66 := -1::int
-#172 := (f4 0::int -1::int)
-#175 := (= #172 0::int)
-#188 := (not #175)
-#38 := 1::int
-#39 := (- 1::int)
-#40 := (f4 0::int #39)
-#41 := (= #40 0::int)
-#42 := (not #41)
-#191 := (iff #42 #188)
-#178 := (= 0::int #172)
-#183 := (not #178)
-#189 := (iff #183 #188)
-#186 := (iff #178 #175)
-#187 := [rewrite]: #186
-#190 := [monotonicity #187]: #189
-#184 := (iff #42 #183)
-#181 := (iff #41 #178)
-#179 := (iff #175 #178)
-#180 := [rewrite]: #179
-#176 := (iff #41 #175)
-#173 := (= #40 #172)
-#170 := (= #39 -1::int)
-#171 := [rewrite]: #170
-#174 := [monotonicity #171]: #173
-#177 := [monotonicity #174]: #176
-#182 := [trans #177 #180]: #181
-#185 := [monotonicity #182]: #184
-#192 := [trans #185 #190]: #191
-#169 := [asserted]: #42
-#193 := [mp #169 #192]: #188
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#29 := (f4 #8 #9)
-#758 := (pattern #29)
-#70 := (* -1::int #9)
-#67 := (* -1::int #8)
-#125 := (mod #67 #70)
-#270 := (+ #29 #125)
-#271 := (= #270 0::int)
-#30 := (mod #8 #9)
-#267 := (* -1::int #30)
-#268 := (+ #29 #267)
-#269 := (= #268 0::int)
-#92 := (<= #9 0::int)
-#88 := (<= #8 0::int)
-#209 := (or #88 #92)
-#210 := (not #209)
-#99 := (>= #8 0::int)
-#201 := (or #92 #99)
-#202 := (not #201)
-#216 := (or #202 #210)
-#272 := (ite #216 #269 #271)
-#266 := (= #29 0::int)
-#12 := (= #8 0::int)
-#273 := (ite #12 #266 #272)
-#265 := (= #8 #29)
-#13 := (= #9 0::int)
-#274 := (ite #13 #265 #273)
-#759 := (forall (vars (?v0 int) (?v1 int)) (:pat #758) #274)
-#277 := (forall (vars (?v0 int) (?v1 int)) #274)
-#762 := (iff #277 #759)
-#760 := (iff #274 #274)
-#761 := [refl]: #760
-#763 := [quant-intro #761]: #762
-#131 := (* -1::int #125)
-#234 := (ite #216 #30 #131)
-#237 := (ite #12 0::int #234)
-#240 := (ite #13 #8 #237)
-#243 := (= #29 #240)
-#246 := (forall (vars (?v0 int) (?v1 int)) #243)
-#278 := (iff #246 #277)
-#275 := (iff #243 #274)
-#276 := [rewrite]: #275
-#279 := [quant-intro #276]: #278
-#100 := (not #99)
-#93 := (not #92)
-#103 := (and #93 #100)
-#89 := (not #88)
-#96 := (and #89 #93)
-#106 := (or #96 #103)
-#151 := (ite #106 #30 #131)
-#154 := (ite #12 0::int #151)
-#157 := (ite #13 #8 #154)
-#160 := (= #29 #157)
-#163 := (forall (vars (?v0 int) (?v1 int)) #160)
-#247 := (iff #163 #246)
-#244 := (iff #160 #243)
-#241 := (= #157 #240)
-#238 := (= #154 #237)
-#235 := (= #151 #234)
-#219 := (iff #106 #216)
-#213 := (or #210 #202)
-#217 := (iff #213 #216)
-#218 := [rewrite]: #217
-#214 := (iff #106 #213)
-#211 := (iff #103 #202)
-#212 := [rewrite]: #211
-#199 := (iff #96 #210)
-#200 := [rewrite]: #199
-#215 := [monotonicity #200 #212]: #214
-#220 := [trans #215 #218]: #219
-#236 := [monotonicity #220]: #235
-#239 := [monotonicity #236]: #238
-#242 := [monotonicity #239]: #241
-#245 := [monotonicity #242]: #244
-#248 := [quant-intro #245]: #247
-#197 := (~ #163 #163)
-#194 := (~ #160 #160)
-#207 := [refl]: #194
-#198 := [nnf-pos #207]: #197
-#23 := (- #9)
-#22 := (- #8)
-#31 := (mod #22 #23)
-#32 := (- #31)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#33 := (ite #20 #30 #32)
-#34 := (ite #12 0::int #33)
-#35 := (ite #13 #8 #34)
-#36 := (= #29 #35)
-#37 := (forall (vars (?v0 int) (?v1 int)) #36)
-#166 := (iff #37 #163)
-#60 := (and #16 #18)
-#63 := (or #17 #60)
-#136 := (ite #63 #30 #131)
-#139 := (ite #12 0::int #136)
-#142 := (ite #13 #8 #139)
-#145 := (= #29 #142)
-#148 := (forall (vars (?v0 int) (?v1 int)) #145)
-#164 := (iff #148 #163)
-#161 := (iff #145 #160)
-#158 := (= #142 #157)
-#155 := (= #139 #154)
-#152 := (= #136 #151)
-#107 := (iff #63 #106)
-#104 := (iff #60 #103)
-#101 := (iff #18 #100)
-#102 := [rewrite]: #101
-#94 := (iff #16 #93)
-#95 := [rewrite]: #94
-#105 := [monotonicity #95 #102]: #104
-#97 := (iff #17 #96)
-#90 := (iff #15 #89)
-#91 := [rewrite]: #90
-#98 := [monotonicity #91 #95]: #97
-#108 := [monotonicity #98 #105]: #107
-#153 := [monotonicity #108]: #152
-#156 := [monotonicity #153]: #155
-#159 := [monotonicity #156]: #158
-#162 := [monotonicity #159]: #161
-#165 := [quant-intro #162]: #164
-#149 := (iff #37 #148)
-#146 := (iff #36 #145)
-#143 := (= #35 #142)
-#140 := (= #34 #139)
-#137 := (= #33 #136)
-#134 := (= #32 #131)
-#128 := (- #125)
-#132 := (= #128 #131)
-#133 := [rewrite]: #132
-#129 := (= #32 #128)
-#126 := (= #31 #125)
-#71 := (= #23 #70)
-#72 := [rewrite]: #71
-#68 := (= #22 #67)
-#69 := [rewrite]: #68
-#127 := [monotonicity #69 #72]: #126
-#130 := [monotonicity #127]: #129
-#135 := [trans #130 #133]: #134
-#64 := (iff #20 #63)
-#61 := (iff #19 #60)
-#62 := [rewrite]: #61
-#65 := [monotonicity #62]: #64
-#138 := [monotonicity #65 #135]: #137
-#141 := [monotonicity #138]: #140
-#144 := [monotonicity #141]: #143
-#147 := [monotonicity #144]: #146
-#150 := [quant-intro #147]: #149
-#167 := [trans #150 #165]: #166
-#124 := [asserted]: #37
-#168 := [mp #124 #167]: #163
-#208 := [mp~ #168 #198]: #163
-#249 := [mp #208 #248]: #246
-#280 := [mp #249 #279]: #277
-#764 := [mp #280 #763]: #759
-#535 := (not #759)
-#683 := (or #535 #175)
-#338 := (* -1::int -1::int)
-#422 := (* -1::int 0::int)
-#423 := (mod #422 #338)
-#424 := (+ #172 #423)
-#415 := (= #424 0::int)
-#426 := (mod 0::int -1::int)
-#427 := (* -1::int #426)
-#428 := (+ #172 #427)
-#425 := (= #428 0::int)
-#429 := (<= -1::int 0::int)
-#408 := (<= 0::int 0::int)
-#743 := (or #408 #429)
-#745 := (not #743)
-#402 := (>= 0::int 0::int)
-#532 := (or #429 #402)
-#739 := (not #532)
-#413 := (or #739 #745)
-#414 := (ite #413 #425 #415)
-#746 := (= 0::int 0::int)
-#747 := (ite #746 #175 #414)
-#748 := (= -1::int 0::int)
-#749 := (ite #748 #178 #747)
-#585 := (or #535 #749)
-#681 := (iff #585 #683)
-#569 := (iff #683 #683)
-#584 := [rewrite]: #569
-#575 := (iff #749 #175)
-#693 := (ite false #175 #175)
-#701 := (iff #693 #175)
-#695 := [rewrite]: #701
-#692 := (iff #749 #693)
-#691 := (iff #747 #175)
-#1 := true
-#533 := (ite true #175 #175)
-#688 := (iff #533 #175)
-#690 := [rewrite]: #688
-#534 := (iff #747 #533)
-#702 := (iff #414 #175)
-#694 := (iff #414 #693)
-#698 := (iff #415 #175)
-#696 := (= #424 #172)
-#436 := (+ #172 0::int)
-#430 := (= #436 #172)
-#712 := [rewrite]: #430
-#553 := (= #424 #436)
-#711 := (= #423 0::int)
-#703 := (mod 0::int 1::int)
-#710 := (= #703 0::int)
-#705 := [rewrite]: #710
-#704 := (= #423 #703)
-#707 := (= #338 1::int)
-#708 := [rewrite]: #707
-#723 := (= #422 0::int)
-#433 := [rewrite]: #723
-#709 := [monotonicity #433 #708]: #704
-#552 := [trans #709 #705]: #711
-#554 := [monotonicity #552]: #553
-#697 := [trans #554 #712]: #696
-#699 := [monotonicity #697]: #698
-#717 := (iff #425 #175)
-#716 := (= #428 #172)
-#714 := (= #428 #436)
-#434 := (= #427 0::int)
-#722 := (= #427 #422)
-#444 := (= #426 0::int)
-#718 := [rewrite]: #444
-#719 := [monotonicity #718]: #722
-#435 := [trans #719 #433]: #434
-#715 := [monotonicity #435]: #714
-#713 := [trans #715 #712]: #716
-#706 := [monotonicity #713]: #717
-#721 := (iff #413 false)
-#448 := (or false false)
-#344 := (iff #448 false)
-#720 := [rewrite]: #344
-#449 := (iff #413 #448)
-#725 := (iff #745 false)
-#729 := (not true)
-#732 := (iff #729 false)
-#727 := [rewrite]: #732
-#738 := (iff #745 #729)
-#737 := (iff #743 true)
-#385 := (or true true)
-#390 := (iff #385 true)
-#391 := [rewrite]: #390
-#376 := (iff #743 #385)
-#405 := (iff #429 true)
-#741 := [rewrite]: #405
-#734 := (iff #408 true)
-#736 := [rewrite]: #734
-#377 := [monotonicity #736 #741]: #376
-#735 := [trans #377 #391]: #737
-#724 := [monotonicity #735]: #738
-#447 := [trans #724 #727]: #725
-#733 := (iff #739 false)
-#730 := (iff #739 #729)
-#392 := (iff #532 true)
-#726 := (iff #532 #385)
-#401 := (iff #402 true)
-#742 := [rewrite]: #401
-#389 := [monotonicity #741 #742]: #726
-#728 := [trans #389 #391]: #392
-#731 := [monotonicity #728]: #730
-#371 := [trans #731 #727]: #733
-#450 := [monotonicity #371 #447]: #449
-#443 := [trans #450 #720]: #721
-#700 := [monotonicity #443 #706 #699]: #694
-#531 := [trans #700 #695]: #702
-#740 := (iff #746 true)
-#400 := [rewrite]: #740
-#687 := [monotonicity #400 #531]: #534
-#689 := [trans #687 #690]: #691
-#744 := (iff #748 false)
-#750 := [rewrite]: #744
-#574 := [monotonicity #750 #187 #689]: #692
-#576 := [trans #574 #695]: #575
-#684 := [monotonicity #576]: #681
-#586 := [trans #684 #584]: #681
-#680 := [quant-inst]: #585
-#587 := [mp #680 #586]: #683
-[unit-resolution #587 #764 #193]: false
-unsat
-9598c0d2e42a8240df571681abaece65536e1221 362 0
+ef319beaaae1961d7e0fb146b0bd21608ccf741d 362 0
 #2 := false
 #11 := 0::int
 #39 := 1::int
@@ -32023,7 +30154,6498 @@
 #563 := [th-lemma]: #562
 [unit-resolution #563 #561 #633]: false
 unsat
-79ab94d82443340b024a2e044bfbebeb9d8e374b 327 0
+24bbfda15529f2b078998e7a57a330dfd637cef3 282 0
+#2 := false
+#11 := 0::int
+decl f3 :: (-> int int int)
+#38 := 3::int
+#39 := (f3 0::int 3::int)
+#40 := (= #39 0::int)
+#41 := (not #40)
+#168 := [asserted]: #41
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#10 := (f3 #8 #9)
+#737 := (pattern #10)
+#65 := -1::int
+#69 := (* -1::int #9)
+#66 := (* -1::int #8)
+#72 := (div #66 #69)
+#239 := (* -1::int #72)
+#240 := (+ #10 #239)
+#241 := (= #240 0::int)
+#21 := (div #8 #9)
+#236 := (* -1::int #21)
+#237 := (+ #10 #236)
+#238 := (= #237 0::int)
+#91 := (<= #9 0::int)
+#87 := (<= #8 0::int)
+#194 := (or #87 #91)
+#195 := (not #194)
+#98 := (>= #8 0::int)
+#186 := (or #91 #98)
+#187 := (not #186)
+#201 := (or #187 #195)
+#242 := (ite #201 #238 #241)
+#235 := (= #10 0::int)
+#13 := (= #9 0::int)
+#12 := (= #8 0::int)
+#14 := (or #12 #13)
+#243 := (ite #14 #235 #242)
+#738 := (forall (vars (?v0 int) (?v1 int)) (:pat #737) #243)
+#246 := (forall (vars (?v0 int) (?v1 int)) #243)
+#741 := (iff #246 #738)
+#739 := (iff #243 #243)
+#740 := [refl]: #739
+#742 := [quant-intro #740]: #741
+#206 := (ite #201 #21 #72)
+#209 := (ite #14 0::int #206)
+#212 := (= #10 #209)
+#215 := (forall (vars (?v0 int) (?v1 int)) #212)
+#247 := (iff #215 #246)
+#244 := (iff #212 #243)
+#245 := [rewrite]: #244
+#248 := [quant-intro #245]: #247
+#99 := (not #98)
+#92 := (not #91)
+#102 := (and #92 #99)
+#88 := (not #87)
+#95 := (and #88 #92)
+#105 := (or #95 #102)
+#108 := (ite #105 #21 #72)
+#111 := (ite #14 0::int #108)
+#114 := (= #10 #111)
+#117 := (forall (vars (?v0 int) (?v1 int)) #114)
+#216 := (iff #117 #215)
+#213 := (iff #114 #212)
+#210 := (= #111 #209)
+#207 := (= #108 #206)
+#204 := (iff #105 #201)
+#198 := (or #195 #187)
+#202 := (iff #198 #201)
+#203 := [rewrite]: #202
+#199 := (iff #105 #198)
+#196 := (iff #102 #187)
+#197 := [rewrite]: #196
+#184 := (iff #95 #195)
+#185 := [rewrite]: #184
+#200 := [monotonicity #185 #197]: #199
+#205 := [trans #200 #203]: #204
+#208 := [monotonicity #205]: #207
+#211 := [monotonicity #208]: #210
+#214 := [monotonicity #211]: #213
+#217 := [quant-intro #214]: #216
+#190 := (~ #117 #117)
+#188 := (~ #114 #114)
+#189 := [refl]: #188
+#191 := [nnf-pos #189]: #190
+#23 := (- #9)
+#22 := (- #8)
+#24 := (div #22 #23)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#25 := (ite #20 #21 #24)
+#26 := (ite #14 0::int #25)
+#27 := (= #10 #26)
+#28 := (forall (vars (?v0 int) (?v1 int)) #27)
+#120 := (iff #28 #117)
+#59 := (and #16 #18)
+#62 := (or #17 #59)
+#75 := (ite #62 #21 #72)
+#78 := (ite #14 0::int #75)
+#81 := (= #10 #78)
+#84 := (forall (vars (?v0 int) (?v1 int)) #81)
+#118 := (iff #84 #117)
+#115 := (iff #81 #114)
+#112 := (= #78 #111)
+#109 := (= #75 #108)
+#106 := (iff #62 #105)
+#103 := (iff #59 #102)
+#100 := (iff #18 #99)
+#101 := [rewrite]: #100
+#93 := (iff #16 #92)
+#94 := [rewrite]: #93
+#104 := [monotonicity #94 #101]: #103
+#96 := (iff #17 #95)
+#89 := (iff #15 #88)
+#90 := [rewrite]: #89
+#97 := [monotonicity #90 #94]: #96
+#107 := [monotonicity #97 #104]: #106
+#110 := [monotonicity #107]: #109
+#113 := [monotonicity #110]: #112
+#116 := [monotonicity #113]: #115
+#119 := [quant-intro #116]: #118
+#85 := (iff #28 #84)
+#82 := (iff #27 #81)
+#79 := (= #26 #78)
+#76 := (= #25 #75)
+#73 := (= #24 #72)
+#70 := (= #23 #69)
+#71 := [rewrite]: #70
+#67 := (= #22 #66)
+#68 := [rewrite]: #67
+#74 := [monotonicity #68 #71]: #73
+#63 := (iff #20 #62)
+#60 := (iff #19 #59)
+#61 := [rewrite]: #60
+#64 := [monotonicity #61]: #63
+#77 := [monotonicity #64 #74]: #76
+#80 := [monotonicity #77]: #79
+#83 := [monotonicity #80]: #82
+#86 := [quant-intro #83]: #85
+#121 := [trans #86 #119]: #120
+#58 := [asserted]: #28
+#122 := [mp #58 #121]: #117
+#181 := [mp~ #122 #191]: #117
+#218 := [mp #181 #217]: #215
+#249 := [mp #218 #248]: #246
+#743 := [mp #249 #742]: #738
+#572 := (not #738)
+#573 := (or #572 #40)
+#323 := (* -1::int 3::int)
+#408 := (* -1::int 0::int)
+#409 := (div #408 #323)
+#410 := (* -1::int #409)
+#401 := (+ #39 #410)
+#412 := (= #401 0::int)
+#413 := (div 0::int 3::int)
+#341 := (* -1::int #413)
+#414 := (+ #39 #341)
+#411 := (= #414 0::int)
+#415 := (<= 3::int 0::int)
+#394 := (<= 0::int 0::int)
+#729 := (or #394 #415)
+#731 := (not #729)
+#388 := (>= 0::int 0::int)
+#518 := (or #415 #388)
+#725 := (not #518)
+#399 := (or #725 #731)
+#400 := (ite #399 #411 #412)
+#732 := (= 3::int 0::int)
+#733 := (= 0::int 0::int)
+#734 := (or #733 #732)
+#735 := (ite #734 #40 #400)
+#566 := (or #572 #735)
+#575 := (iff #566 #573)
+#577 := (iff #573 #573)
+#578 := [rewrite]: #577
+#555 := (iff #735 #40)
+#1 := true
+#669 := (ite true #40 #40)
+#667 := (iff #669 #40)
+#670 := [rewrite]: #667
+#571 := (iff #735 #669)
+#562 := (iff #400 #40)
+#677 := (ite false #40 #40)
+#560 := (iff #677 #40)
+#561 := [rewrite]: #560
+#675 := (iff #400 #677)
+#674 := (iff #412 #40)
+#520 := (= #401 #39)
+#703 := (+ #39 0::int)
+#694 := (= #703 #39)
+#689 := [rewrite]: #694
+#517 := (= #401 #703)
+#681 := (= #410 0::int)
+#416 := (= #408 0::int)
+#698 := [rewrite]: #416
+#686 := (= #410 #408)
+#679 := (= #409 0::int)
+#697 := -3::int
+#540 := (div 0::int -3::int)
+#684 := (= #540 0::int)
+#685 := [rewrite]: #684
+#682 := (= #409 #540)
+#538 := (= #323 -3::int)
+#539 := [rewrite]: #538
+#683 := [monotonicity #698 #539]: #682
+#680 := [trans #683 #685]: #679
+#687 := [monotonicity #680]: #686
+#688 := [trans #687 #698]: #681
+#519 := [monotonicity #688]: #517
+#673 := [trans #519 #689]: #520
+#676 := [monotonicity #673]: #674
+#696 := (iff #411 #40)
+#690 := (= #414 #39)
+#692 := (= #414 #703)
+#702 := (= #341 0::int)
+#700 := (= #341 #408)
+#421 := (= #413 0::int)
+#422 := [rewrite]: #421
+#701 := [monotonicity #422]: #700
+#699 := [trans #701 #698]: #702
+#693 := [monotonicity #699]: #692
+#695 := [trans #693 #689]: #690
+#691 := [monotonicity #695]: #696
+#419 := (iff #399 false)
+#430 := (or false false)
+#705 := (iff #430 false)
+#709 := [rewrite]: #705
+#704 := (iff #399 #430)
+#707 := (iff #731 false)
+#720 := (not true)
+#363 := (iff #720 false)
+#723 := [rewrite]: #363
+#329 := (iff #731 #720)
+#435 := (iff #729 true)
+#391 := (or true false)
+#728 := (iff #391 true)
+#371 := [rewrite]: #728
+#433 := (iff #729 #391)
+#376 := (iff #415 false)
+#377 := [rewrite]: #376
+#710 := (iff #394 true)
+#711 := [rewrite]: #710
+#434 := [monotonicity #711 #377]: #433
+#436 := [trans #434 #371]: #435
+#706 := [monotonicity #436]: #329
+#429 := [trans #706 #723]: #707
+#721 := (iff #725 false)
+#722 := (iff #725 #720)
+#719 := (iff #518 true)
+#715 := (or false true)
+#718 := (iff #715 true)
+#713 := [rewrite]: #718
+#716 := (iff #518 #715)
+#378 := (iff #388 true)
+#714 := [rewrite]: #378
+#717 := [monotonicity #377 #714]: #716
+#357 := [trans #717 #713]: #719
+#362 := [monotonicity #357]: #722
+#724 := [trans #362 #723]: #721
+#708 := [monotonicity #724 #429]: #704
+#420 := [trans #708 #709]: #419
+#678 := [monotonicity #420 #691 #676]: #675
+#521 := [trans #678 #561]: #562
+#712 := (iff #734 true)
+#727 := (iff #734 #391)
+#726 := (iff #732 false)
+#386 := [rewrite]: #726
+#730 := (iff #733 true)
+#736 := [rewrite]: #730
+#387 := [monotonicity #736 #386]: #727
+#375 := [trans #387 #371]: #712
+#666 := [monotonicity #375 #521]: #571
+#570 := [trans #666 #670]: #555
+#576 := [monotonicity #570]: #575
+#671 := [trans #576 #578]: #575
+#574 := [quant-inst]: #566
+#668 := [mp #574 #671]: #573
+[unit-resolution #668 #743 #168]: false
+unsat
+12ef6b4ae9675610f423327a4c22119f401927ad 307 0
+#2 := false
+#11 := 0::int
+decl f3 :: (-> int int int)
+#170 := -3::int
+#173 := (f3 0::int -3::int)
+#176 := (= #173 0::int)
+#189 := (not #176)
+#38 := 3::int
+#39 := (- 3::int)
+#40 := (f3 0::int #39)
+#41 := (= #40 0::int)
+#42 := (not #41)
+#192 := (iff #42 #189)
+#179 := (= 0::int #173)
+#184 := (not #179)
+#190 := (iff #184 #189)
+#187 := (iff #179 #176)
+#188 := [rewrite]: #187
+#191 := [monotonicity #188]: #190
+#185 := (iff #42 #184)
+#182 := (iff #41 #179)
+#180 := (iff #176 #179)
+#181 := [rewrite]: #180
+#177 := (iff #41 #176)
+#174 := (= #40 #173)
+#171 := (= #39 -3::int)
+#172 := [rewrite]: #171
+#175 := [monotonicity #172]: #174
+#178 := [monotonicity #175]: #177
+#183 := [trans #178 #181]: #182
+#186 := [monotonicity #183]: #185
+#193 := [trans #186 #191]: #192
+#169 := [asserted]: #42
+#194 := [mp #169 #193]: #189
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#10 := (f3 #8 #9)
+#753 := (pattern #10)
+#66 := -1::int
+#70 := (* -1::int #9)
+#67 := (* -1::int #8)
+#73 := (div #67 #70)
+#255 := (* -1::int #73)
+#256 := (+ #10 #255)
+#257 := (= #256 0::int)
+#21 := (div #8 #9)
+#252 := (* -1::int #21)
+#253 := (+ #10 #252)
+#254 := (= #253 0::int)
+#92 := (<= #9 0::int)
+#88 := (<= #8 0::int)
+#210 := (or #88 #92)
+#211 := (not #210)
+#99 := (>= #8 0::int)
+#202 := (or #92 #99)
+#203 := (not #202)
+#217 := (or #203 #211)
+#258 := (ite #217 #254 #257)
+#251 := (= #10 0::int)
+#13 := (= #9 0::int)
+#12 := (= #8 0::int)
+#14 := (or #12 #13)
+#259 := (ite #14 #251 #258)
+#754 := (forall (vars (?v0 int) (?v1 int)) (:pat #753) #259)
+#262 := (forall (vars (?v0 int) (?v1 int)) #259)
+#757 := (iff #262 #754)
+#755 := (iff #259 #259)
+#756 := [refl]: #755
+#758 := [quant-intro #756]: #757
+#222 := (ite #217 #21 #73)
+#225 := (ite #14 0::int #222)
+#228 := (= #10 #225)
+#231 := (forall (vars (?v0 int) (?v1 int)) #228)
+#263 := (iff #231 #262)
+#260 := (iff #228 #259)
+#261 := [rewrite]: #260
+#264 := [quant-intro #261]: #263
+#100 := (not #99)
+#93 := (not #92)
+#103 := (and #93 #100)
+#89 := (not #88)
+#96 := (and #89 #93)
+#106 := (or #96 #103)
+#109 := (ite #106 #21 #73)
+#112 := (ite #14 0::int #109)
+#115 := (= #10 #112)
+#118 := (forall (vars (?v0 int) (?v1 int)) #115)
+#232 := (iff #118 #231)
+#229 := (iff #115 #228)
+#226 := (= #112 #225)
+#223 := (= #109 #222)
+#220 := (iff #106 #217)
+#214 := (or #211 #203)
+#218 := (iff #214 #217)
+#219 := [rewrite]: #218
+#215 := (iff #106 #214)
+#212 := (iff #103 #203)
+#213 := [rewrite]: #212
+#200 := (iff #96 #211)
+#201 := [rewrite]: #200
+#216 := [monotonicity #201 #213]: #215
+#221 := [trans #216 #219]: #220
+#224 := [monotonicity #221]: #223
+#227 := [monotonicity #224]: #226
+#230 := [monotonicity #227]: #229
+#233 := [quant-intro #230]: #232
+#206 := (~ #118 #118)
+#204 := (~ #115 #115)
+#205 := [refl]: #204
+#207 := [nnf-pos #205]: #206
+#23 := (- #9)
+#22 := (- #8)
+#24 := (div #22 #23)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#25 := (ite #20 #21 #24)
+#26 := (ite #14 0::int #25)
+#27 := (= #10 #26)
+#28 := (forall (vars (?v0 int) (?v1 int)) #27)
+#121 := (iff #28 #118)
+#60 := (and #16 #18)
+#63 := (or #17 #60)
+#76 := (ite #63 #21 #73)
+#79 := (ite #14 0::int #76)
+#82 := (= #10 #79)
+#85 := (forall (vars (?v0 int) (?v1 int)) #82)
+#119 := (iff #85 #118)
+#116 := (iff #82 #115)
+#113 := (= #79 #112)
+#110 := (= #76 #109)
+#107 := (iff #63 #106)
+#104 := (iff #60 #103)
+#101 := (iff #18 #100)
+#102 := [rewrite]: #101
+#94 := (iff #16 #93)
+#95 := [rewrite]: #94
+#105 := [monotonicity #95 #102]: #104
+#97 := (iff #17 #96)
+#90 := (iff #15 #89)
+#91 := [rewrite]: #90
+#98 := [monotonicity #91 #95]: #97
+#108 := [monotonicity #98 #105]: #107
+#111 := [monotonicity #108]: #110
+#114 := [monotonicity #111]: #113
+#117 := [monotonicity #114]: #116
+#120 := [quant-intro #117]: #119
+#86 := (iff #28 #85)
+#83 := (iff #27 #82)
+#80 := (= #26 #79)
+#77 := (= #25 #76)
+#74 := (= #24 #73)
+#71 := (= #23 #70)
+#72 := [rewrite]: #71
+#68 := (= #22 #67)
+#69 := [rewrite]: #68
+#75 := [monotonicity #69 #72]: #74
+#64 := (iff #20 #63)
+#61 := (iff #19 #60)
+#62 := [rewrite]: #61
+#65 := [monotonicity #62]: #64
+#78 := [monotonicity #65 #75]: #77
+#81 := [monotonicity #78]: #80
+#84 := [monotonicity #81]: #83
+#87 := [quant-intro #84]: #86
+#122 := [trans #87 #120]: #121
+#59 := [asserted]: #28
+#123 := [mp #59 #122]: #118
+#197 := [mp~ #123 #207]: #118
+#234 := [mp #197 #233]: #231
+#265 := [mp #234 #264]: #262
+#759 := [mp #265 #758]: #754
+#586 := (not #754)
+#588 := (or #586 #176)
+#339 := (* -1::int -3::int)
+#424 := (* -1::int 0::int)
+#425 := (div #424 #339)
+#426 := (* -1::int #425)
+#417 := (+ #173 #426)
+#428 := (= #417 0::int)
+#429 := (div 0::int -3::int)
+#357 := (* -1::int #429)
+#430 := (+ #173 #357)
+#427 := (= #430 0::int)
+#431 := (<= -3::int 0::int)
+#410 := (<= 0::int 0::int)
+#745 := (or #410 #431)
+#747 := (not #745)
+#404 := (>= 0::int 0::int)
+#534 := (or #431 #404)
+#741 := (not #534)
+#415 := (or #741 #747)
+#416 := (ite #415 #427 #428)
+#748 := (= -3::int 0::int)
+#749 := (= 0::int 0::int)
+#750 := (or #749 #748)
+#751 := (ite #750 #176 #416)
+#589 := (or #586 #751)
+#590 := (iff #589 #588)
+#592 := (iff #588 #588)
+#593 := [rewrite]: #592
+#686 := (iff #751 #176)
+#1 := true
+#537 := (ite true #176 #176)
+#682 := (iff #537 #176)
+#683 := [rewrite]: #682
+#685 := (iff #751 #537)
+#577 := (iff #416 #176)
+#692 := (ite false #176 #176)
+#694 := (iff #692 #176)
+#576 := [rewrite]: #694
+#693 := (iff #416 #692)
+#689 := (iff #428 #176)
+#535 := (= #417 #173)
+#719 := (+ #173 0::int)
+#710 := (= #719 #173)
+#705 := [rewrite]: #710
+#704 := (= #417 #719)
+#703 := (= #426 0::int)
+#432 := (= #424 0::int)
+#714 := [rewrite]: #432
+#696 := (= #426 #424)
+#701 := (= #425 0::int)
+#555 := (div 0::int 3::int)
+#699 := (= #555 0::int)
+#700 := [rewrite]: #699
+#556 := (= #425 #555)
+#713 := (= #339 3::int)
+#554 := [rewrite]: #713
+#698 := [monotonicity #714 #554]: #556
+#695 := [trans #698 #700]: #701
+#702 := [monotonicity #695]: #696
+#697 := [trans #702 #714]: #703
+#533 := [monotonicity #697]: #704
+#536 := [trans #533 #705]: #535
+#690 := [monotonicity #536]: #689
+#712 := (iff #427 #176)
+#706 := (= #430 #173)
+#708 := (= #430 #719)
+#718 := (= #357 0::int)
+#716 := (= #357 #424)
+#437 := (= #429 0::int)
+#438 := [rewrite]: #437
+#717 := [monotonicity #438]: #716
+#715 := [trans #717 #714]: #718
+#709 := [monotonicity #715]: #708
+#711 := [trans #709 #705]: #706
+#707 := [monotonicity #711]: #712
+#435 := (iff #415 false)
+#446 := (or false false)
+#721 := (iff #446 false)
+#725 := [rewrite]: #721
+#720 := (iff #415 #446)
+#723 := (iff #747 false)
+#736 := (not true)
+#379 := (iff #736 false)
+#739 := [rewrite]: #379
+#345 := (iff #747 #736)
+#451 := (iff #745 true)
+#731 := (or true true)
+#734 := (iff #731 true)
+#729 := [rewrite]: #734
+#449 := (iff #745 #731)
+#392 := (iff #431 true)
+#393 := [rewrite]: #392
+#726 := (iff #410 true)
+#727 := [rewrite]: #726
+#450 := [monotonicity #727 #393]: #449
+#452 := [trans #450 #729]: #451
+#722 := [monotonicity #452]: #345
+#445 := [trans #722 #739]: #723
+#737 := (iff #741 false)
+#738 := (iff #741 #736)
+#735 := (iff #534 true)
+#732 := (iff #534 #731)
+#394 := (iff #404 true)
+#730 := [rewrite]: #394
+#733 := [monotonicity #393 #730]: #732
+#373 := [trans #733 #729]: #735
+#378 := [monotonicity #373]: #738
+#740 := [trans #378 #739]: #737
+#724 := [monotonicity #740 #445]: #720
+#436 := [trans #724 #725]: #435
+#691 := [monotonicity #436 #707 #690]: #693
+#578 := [trans #691 #576]: #577
+#728 := (iff #750 true)
+#407 := (or true false)
+#744 := (iff #407 true)
+#387 := [rewrite]: #744
+#743 := (iff #750 #407)
+#742 := (iff #748 false)
+#402 := [rewrite]: #742
+#746 := (iff #749 true)
+#752 := [rewrite]: #746
+#403 := [monotonicity #752 #402]: #743
+#391 := [trans #403 #387]: #728
+#587 := [monotonicity #391 #578]: #685
+#571 := [trans #587 #683]: #686
+#591 := [monotonicity #571]: #590
+#594 := [trans #591 #593]: #590
+#582 := [quant-inst]: #589
+#687 := [mp #582 #594]: #588
+[unit-resolution #687 #759 #194]: false
+unsat
+35df6ca3b9d8297d2ae93fc18a5a2fbecd67e27f 296 0
+#2 := false
+#11 := 0::int
+decl f3 :: (-> int int int)
+#39 := 3::int
+#38 := 1::int
+#40 := (f3 1::int 3::int)
+#41 := (= #40 0::int)
+#42 := (not #41)
+#169 := [asserted]: #42
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#10 := (f3 #8 #9)
+#737 := (pattern #10)
+#66 := -1::int
+#70 := (* -1::int #9)
+#67 := (* -1::int #8)
+#73 := (div #67 #70)
+#240 := (* -1::int #73)
+#241 := (+ #10 #240)
+#242 := (= #241 0::int)
+#21 := (div #8 #9)
+#237 := (* -1::int #21)
+#238 := (+ #10 #237)
+#239 := (= #238 0::int)
+#92 := (<= #9 0::int)
+#88 := (<= #8 0::int)
+#195 := (or #88 #92)
+#196 := (not #195)
+#99 := (>= #8 0::int)
+#187 := (or #92 #99)
+#188 := (not #187)
+#202 := (or #188 #196)
+#243 := (ite #202 #239 #242)
+#236 := (= #10 0::int)
+#13 := (= #9 0::int)
+#12 := (= #8 0::int)
+#14 := (or #12 #13)
+#244 := (ite #14 #236 #243)
+#738 := (forall (vars (?v0 int) (?v1 int)) (:pat #737) #244)
+#247 := (forall (vars (?v0 int) (?v1 int)) #244)
+#741 := (iff #247 #738)
+#739 := (iff #244 #244)
+#740 := [refl]: #739
+#742 := [quant-intro #740]: #741
+#207 := (ite #202 #21 #73)
+#210 := (ite #14 0::int #207)
+#213 := (= #10 #210)
+#216 := (forall (vars (?v0 int) (?v1 int)) #213)
+#248 := (iff #216 #247)
+#245 := (iff #213 #244)
+#246 := [rewrite]: #245
+#249 := [quant-intro #246]: #248
+#100 := (not #99)
+#93 := (not #92)
+#103 := (and #93 #100)
+#89 := (not #88)
+#96 := (and #89 #93)
+#106 := (or #96 #103)
+#109 := (ite #106 #21 #73)
+#112 := (ite #14 0::int #109)
+#115 := (= #10 #112)
+#118 := (forall (vars (?v0 int) (?v1 int)) #115)
+#217 := (iff #118 #216)
+#214 := (iff #115 #213)
+#211 := (= #112 #210)
+#208 := (= #109 #207)
+#205 := (iff #106 #202)
+#199 := (or #196 #188)
+#203 := (iff #199 #202)
+#204 := [rewrite]: #203
+#200 := (iff #106 #199)
+#197 := (iff #103 #188)
+#198 := [rewrite]: #197
+#185 := (iff #96 #196)
+#186 := [rewrite]: #185
+#201 := [monotonicity #186 #198]: #200
+#206 := [trans #201 #204]: #205
+#209 := [monotonicity #206]: #208
+#212 := [monotonicity #209]: #211
+#215 := [monotonicity #212]: #214
+#218 := [quant-intro #215]: #217
+#191 := (~ #118 #118)
+#189 := (~ #115 #115)
+#190 := [refl]: #189
+#192 := [nnf-pos #190]: #191
+#23 := (- #9)
+#22 := (- #8)
+#24 := (div #22 #23)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#25 := (ite #20 #21 #24)
+#26 := (ite #14 0::int #25)
+#27 := (= #10 #26)
+#28 := (forall (vars (?v0 int) (?v1 int)) #27)
+#121 := (iff #28 #118)
+#60 := (and #16 #18)
+#63 := (or #17 #60)
+#76 := (ite #63 #21 #73)
+#79 := (ite #14 0::int #76)
+#82 := (= #10 #79)
+#85 := (forall (vars (?v0 int) (?v1 int)) #82)
+#119 := (iff #85 #118)
+#116 := (iff #82 #115)
+#113 := (= #79 #112)
+#110 := (= #76 #109)
+#107 := (iff #63 #106)
+#104 := (iff #60 #103)
+#101 := (iff #18 #100)
+#102 := [rewrite]: #101
+#94 := (iff #16 #93)
+#95 := [rewrite]: #94
+#105 := [monotonicity #95 #102]: #104
+#97 := (iff #17 #96)
+#90 := (iff #15 #89)
+#91 := [rewrite]: #90
+#98 := [monotonicity #91 #95]: #97
+#108 := [monotonicity #98 #105]: #107
+#111 := [monotonicity #108]: #110
+#114 := [monotonicity #111]: #113
+#117 := [monotonicity #114]: #116
+#120 := [quant-intro #117]: #119
+#86 := (iff #28 #85)
+#83 := (iff #27 #82)
+#80 := (= #26 #79)
+#77 := (= #25 #76)
+#74 := (= #24 #73)
+#71 := (= #23 #70)
+#72 := [rewrite]: #71
+#68 := (= #22 #67)
+#69 := [rewrite]: #68
+#75 := [monotonicity #69 #72]: #74
+#64 := (iff #20 #63)
+#61 := (iff #19 #60)
+#62 := [rewrite]: #61
+#65 := [monotonicity #62]: #64
+#78 := [monotonicity #65 #75]: #77
+#81 := [monotonicity #78]: #80
+#84 := [monotonicity #81]: #83
+#87 := [quant-intro #84]: #86
+#122 := [trans #87 #120]: #121
+#59 := [asserted]: #28
+#123 := [mp #59 #122]: #118
+#182 := [mp~ #123 #192]: #118
+#219 := [mp #182 #218]: #216
+#250 := [mp #219 #249]: #247
+#743 := [mp #250 #742]: #738
+#650 := (not #738)
+#652 := (or #650 #41)
+#324 := (* -1::int 3::int)
+#408 := (* -1::int 1::int)
+#409 := (div #408 #324)
+#410 := (* -1::int #409)
+#401 := (+ #40 #410)
+#412 := (= #401 0::int)
+#413 := (div 1::int 3::int)
+#414 := (* -1::int #413)
+#411 := (+ #40 #414)
+#415 := (= #411 0::int)
+#394 := (<= 3::int 0::int)
+#729 := (<= 1::int 0::int)
+#731 := (or #729 #394)
+#388 := (not #731)
+#518 := (>= 1::int 0::int)
+#725 := (or #394 #518)
+#399 := (not #725)
+#400 := (or #399 #388)
+#732 := (ite #400 #415 #412)
+#733 := (= 3::int 0::int)
+#734 := (= 1::int 0::int)
+#735 := (or #734 #733)
+#730 := (ite #735 #41 #732)
+#662 := (or #650 #730)
+#664 := (iff #662 #652)
+#665 := (iff #652 #652)
+#654 := [rewrite]: #665
+#661 := (iff #730 #41)
+#578 := (ite false #41 #41)
+#672 := (iff #578 #41)
+#659 := [rewrite]: #672
+#671 := (iff #730 #578)
+#576 := (iff #732 #41)
+#666 := (= #40 1::int)
+#1 := true
+#572 := (ite true #41 #666)
+#574 := (iff #572 #41)
+#575 := [rewrite]: #574
+#573 := (iff #732 #572)
+#555 := (iff #412 #666)
+#675 := (+ -1::int #40)
+#521 := (= #675 0::int)
+#667 := (iff #521 #666)
+#670 := [rewrite]: #667
+#669 := (iff #412 #521)
+#561 := (= #401 #675)
+#674 := (+ #40 -1::int)
+#678 := (= #674 #675)
+#560 := [rewrite]: #678
+#676 := (= #401 #674)
+#520 := (= #410 -1::int)
+#539 := (= #408 -1::int)
+#540 := [rewrite]: #539
+#517 := (= #410 #408)
+#681 := (= #409 1::int)
+#682 := -3::int
+#685 := (div -1::int -3::int)
+#686 := (= #685 1::int)
+#687 := [rewrite]: #686
+#679 := (= #409 #685)
+#683 := (= #324 -3::int)
+#684 := [rewrite]: #683
+#680 := [monotonicity #540 #684]: #679
+#688 := [trans #680 #687]: #681
+#519 := [monotonicity #688]: #517
+#673 := [trans #519 #540]: #520
+#677 := [monotonicity #673]: #676
+#562 := [trans #677 #560]: #561
+#571 := [monotonicity #562]: #669
+#570 := [trans #571 #670]: #555
+#697 := (iff #415 #41)
+#696 := (= #411 #40)
+#693 := (+ #40 0::int)
+#690 := (= #693 #40)
+#695 := [rewrite]: #690
+#694 := (= #411 #693)
+#703 := (= #414 0::int)
+#701 := (* -1::int 0::int)
+#702 := (= #701 0::int)
+#699 := [rewrite]: #702
+#416 := (= #414 #701)
+#422 := (= #413 0::int)
+#700 := [rewrite]: #422
+#698 := [monotonicity #700]: #416
+#692 := [trans #698 #699]: #703
+#689 := [monotonicity #692]: #694
+#691 := [trans #689 #695]: #696
+#538 := [monotonicity #691]: #697
+#420 := (iff #400 true)
+#716 := (or false true)
+#713 := (iff #716 true)
+#719 := [rewrite]: #713
+#709 := (iff #400 #716)
+#708 := (iff #388 true)
+#706 := (not false)
+#430 := (iff #706 true)
+#704 := [rewrite]: #430
+#707 := (iff #388 #706)
+#436 := (iff #731 false)
+#727 := (or false false)
+#371 := (iff #727 false)
+#712 := [rewrite]: #371
+#434 := (iff #731 #727)
+#377 := (iff #394 false)
+#378 := [rewrite]: #377
+#711 := (iff #729 false)
+#433 := [rewrite]: #711
+#435 := [monotonicity #433 #378]: #434
+#330 := [trans #435 #712]: #436
+#429 := [monotonicity #330]: #707
+#705 := [trans #429 #704]: #708
+#724 := (iff #399 false)
+#722 := (not true)
+#723 := (iff #722 false)
+#721 := [rewrite]: #723
+#362 := (iff #399 #722)
+#357 := (iff #725 true)
+#717 := (iff #725 #716)
+#714 := (iff #518 true)
+#715 := [rewrite]: #714
+#718 := [monotonicity #378 #715]: #717
+#720 := [trans #718 #719]: #357
+#363 := [monotonicity #720]: #362
+#710 := [trans #363 #721]: #724
+#419 := [monotonicity #710 #705]: #709
+#421 := [trans #419 #719]: #420
+#566 := [monotonicity #421 #538 #570]: #573
+#577 := [trans #566 #575]: #576
+#375 := (iff #735 false)
+#387 := (iff #735 #727)
+#386 := (iff #733 false)
+#391 := [rewrite]: #386
+#736 := (iff #734 false)
+#726 := [rewrite]: #736
+#728 := [monotonicity #726 #391]: #387
+#376 := [trans #728 #712]: #375
+#668 := [monotonicity #376 #577]: #671
+#653 := [trans #668 #659]: #661
+#660 := [monotonicity #653]: #664
+#655 := [trans #660 #654]: #664
+#663 := [quant-inst]: #662
+#656 := [mp #663 #655]: #652
+[unit-resolution #656 #743 #169]: false
+unsat
+d02187cb7fcefe99851febe02397cc4a2e6fdbff 286 0
+#2 := false
+#40 := 1::int
+decl f3 :: (-> int int int)
+#38 := 3::int
+#39 := (f3 3::int 3::int)
+#41 := (= #39 1::int)
+#42 := (not #41)
+#169 := [asserted]: #42
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#10 := (f3 #8 #9)
+#727 := (pattern #10)
+#11 := 0::int
+#66 := -1::int
+#70 := (* -1::int #9)
+#67 := (* -1::int #8)
+#73 := (div #67 #70)
+#230 := (* -1::int #73)
+#231 := (+ #10 #230)
+#232 := (= #231 0::int)
+#21 := (div #8 #9)
+#227 := (* -1::int #21)
+#228 := (+ #10 #227)
+#229 := (= #228 0::int)
+#92 := (<= #9 0::int)
+#88 := (<= #8 0::int)
+#185 := (or #88 #92)
+#186 := (not #185)
+#99 := (>= #8 0::int)
+#177 := (or #92 #99)
+#178 := (not #177)
+#192 := (or #178 #186)
+#233 := (ite #192 #229 #232)
+#226 := (= #10 0::int)
+#13 := (= #9 0::int)
+#12 := (= #8 0::int)
+#14 := (or #12 #13)
+#234 := (ite #14 #226 #233)
+#728 := (forall (vars (?v0 int) (?v1 int)) (:pat #727) #234)
+#237 := (forall (vars (?v0 int) (?v1 int)) #234)
+#731 := (iff #237 #728)
+#729 := (iff #234 #234)
+#730 := [refl]: #729
+#732 := [quant-intro #730]: #731
+#197 := (ite #192 #21 #73)
+#200 := (ite #14 0::int #197)
+#203 := (= #10 #200)
+#206 := (forall (vars (?v0 int) (?v1 int)) #203)
+#238 := (iff #206 #237)
+#235 := (iff #203 #234)
+#236 := [rewrite]: #235
+#239 := [quant-intro #236]: #238
+#100 := (not #99)
+#93 := (not #92)
+#103 := (and #93 #100)
+#89 := (not #88)
+#96 := (and #89 #93)
+#106 := (or #96 #103)
+#109 := (ite #106 #21 #73)
+#112 := (ite #14 0::int #109)
+#115 := (= #10 #112)
+#118 := (forall (vars (?v0 int) (?v1 int)) #115)
+#207 := (iff #118 #206)
+#204 := (iff #115 #203)
+#201 := (= #112 #200)
+#198 := (= #109 #197)
+#195 := (iff #106 #192)
+#189 := (or #186 #178)
+#193 := (iff #189 #192)
+#194 := [rewrite]: #193
+#190 := (iff #106 #189)
+#187 := (iff #103 #178)
+#188 := [rewrite]: #187
+#175 := (iff #96 #186)
+#176 := [rewrite]: #175
+#191 := [monotonicity #176 #188]: #190
+#196 := [trans #191 #194]: #195
+#199 := [monotonicity #196]: #198
+#202 := [monotonicity #199]: #201
+#205 := [monotonicity #202]: #204
+#208 := [quant-intro #205]: #207
+#181 := (~ #118 #118)
+#179 := (~ #115 #115)
+#180 := [refl]: #179
+#182 := [nnf-pos #180]: #181
+#23 := (- #9)
+#22 := (- #8)
+#24 := (div #22 #23)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#25 := (ite #20 #21 #24)
+#26 := (ite #14 0::int #25)
+#27 := (= #10 #26)
+#28 := (forall (vars (?v0 int) (?v1 int)) #27)
+#121 := (iff #28 #118)
+#60 := (and #16 #18)
+#63 := (or #17 #60)
+#76 := (ite #63 #21 #73)
+#79 := (ite #14 0::int #76)
+#82 := (= #10 #79)
+#85 := (forall (vars (?v0 int) (?v1 int)) #82)
+#119 := (iff #85 #118)
+#116 := (iff #82 #115)
+#113 := (= #79 #112)
+#110 := (= #76 #109)
+#107 := (iff #63 #106)
+#104 := (iff #60 #103)
+#101 := (iff #18 #100)
+#102 := [rewrite]: #101
+#94 := (iff #16 #93)
+#95 := [rewrite]: #94
+#105 := [monotonicity #95 #102]: #104
+#97 := (iff #17 #96)
+#90 := (iff #15 #89)
+#91 := [rewrite]: #90
+#98 := [monotonicity #91 #95]: #97
+#108 := [monotonicity #98 #105]: #107
+#111 := [monotonicity #108]: #110
+#114 := [monotonicity #111]: #113
+#117 := [monotonicity #114]: #116
+#120 := [quant-intro #117]: #119
+#86 := (iff #28 #85)
+#83 := (iff #27 #82)
+#80 := (= #26 #79)
+#77 := (= #25 #76)
+#74 := (= #24 #73)
+#71 := (= #23 #70)
+#72 := [rewrite]: #71
+#68 := (= #22 #67)
+#69 := [rewrite]: #68
+#75 := [monotonicity #69 #72]: #74
+#64 := (iff #20 #63)
+#61 := (iff #19 #60)
+#62 := [rewrite]: #61
+#65 := [monotonicity #62]: #64
+#78 := [monotonicity #65 #75]: #77
+#81 := [monotonicity #78]: #80
+#84 := [monotonicity #81]: #83
+#87 := [quant-intro #84]: #86
+#122 := [trans #87 #120]: #121
+#59 := [asserted]: #28
+#123 := [mp #59 #122]: #118
+#172 := [mp~ #123 #182]: #118
+#209 := [mp #172 #208]: #206
+#240 := [mp #209 #239]: #237
+#733 := [mp #240 #732]: #728
+#565 := (not #728)
+#566 := (or #565 #41)
+#314 := (* -1::int 3::int)
+#398 := (div #314 #314)
+#399 := (* -1::int #398)
+#400 := (+ #39 #399)
+#391 := (= #400 0::int)
+#402 := (div 3::int 3::int)
+#403 := (* -1::int #402)
+#404 := (+ #39 #403)
+#401 := (= #404 0::int)
+#405 := (<= 3::int 0::int)
+#384 := (or #405 #405)
+#719 := (not #384)
+#721 := (>= 3::int 0::int)
+#378 := (or #405 #721)
+#508 := (not #378)
+#715 := (or #508 #719)
+#389 := (ite #715 #401 #391)
+#390 := (= #39 0::int)
+#722 := (= 3::int 0::int)
+#723 := (or #722 #722)
+#724 := (ite #723 #390 #389)
+#567 := (or #565 #724)
+#661 := (iff #567 #566)
+#662 := (iff #566 #566)
+#649 := [rewrite]: #662
+#556 := (iff #724 #41)
+#660 := (ite false #390 #41)
+#562 := (iff #660 #41)
+#563 := [rewrite]: #562
+#545 := (iff #724 #660)
+#656 := (iff #389 #41)
+#1 := true
+#551 := (ite true #41 #41)
+#659 := (iff #551 #41)
+#561 := [rewrite]: #659
+#552 := (iff #389 #551)
+#668 := (iff #391 #41)
+#689 := (+ -1::int #39)
+#679 := (= #689 0::int)
+#686 := (iff #679 #41)
+#681 := [rewrite]: #686
+#667 := (iff #391 #679)
+#664 := (= #400 #689)
+#406 := (+ #39 -1::int)
+#693 := (= #406 #689)
+#682 := [rewrite]: #693
+#510 := (= #400 #406)
+#507 := (= #399 -1::int)
+#699 := (* -1::int 1::int)
+#411 := (= #699 -1::int)
+#412 := [rewrite]: #411
+#671 := (= #399 #699)
+#676 := (= #398 1::int)
+#529 := -3::int
+#673 := (div -3::int -3::int)
+#669 := (= #673 1::int)
+#670 := [rewrite]: #669
+#674 := (= #398 #673)
+#530 := (= #314 -3::int)
+#672 := [rewrite]: #530
+#675 := [monotonicity #672 #672]: #674
+#677 := [trans #675 #670]: #676
+#678 := [monotonicity #677]: #671
+#509 := [trans #678 #412]: #507
+#663 := [monotonicity #509]: #510
+#666 := [trans #663 #682]: #664
+#665 := [monotonicity #666]: #667
+#550 := [trans #665 #681]: #668
+#687 := (iff #401 #41)
+#680 := (iff #401 #679)
+#683 := (= #404 #689)
+#688 := (= #404 #406)
+#690 := (= #403 -1::int)
+#409 := (= #403 #699)
+#698 := (= #402 1::int)
+#695 := [rewrite]: #698
+#410 := [monotonicity #695]: #409
+#691 := [trans #410 #412]: #690
+#692 := [monotonicity #691]: #688
+#684 := [trans #692 #682]: #683
+#685 := [monotonicity #684]: #680
+#528 := [trans #685 #681]: #687
+#420 := (iff #715 true)
+#367 := (or false true)
+#705 := (iff #367 true)
+#706 := [rewrite]: #705
+#697 := (iff #715 #367)
+#320 := (iff #719 true)
+#701 := (not false)
+#425 := (iff #701 true)
+#426 := [rewrite]: #425
+#423 := (iff #719 #701)
+#714 := (iff #384 false)
+#726 := (or false false)
+#381 := (iff #726 false)
+#717 := [rewrite]: #381
+#713 := (iff #384 #726)
+#361 := (iff #405 false)
+#702 := [rewrite]: #361
+#711 := [monotonicity #702 #702]: #713
+#700 := [trans #711 #717]: #714
+#424 := [monotonicity #700]: #423
+#696 := [trans #424 #426]: #320
+#352 := (iff #508 false)
+#703 := (not true)
+#710 := (iff #703 false)
+#712 := [rewrite]: #710
+#709 := (iff #508 #703)
+#707 := (iff #378 true)
+#368 := (iff #378 #367)
+#365 := (iff #721 true)
+#366 := [rewrite]: #365
+#704 := [monotonicity #702 #366]: #368
+#708 := [trans #704 #706]: #707
+#347 := [monotonicity #708]: #709
+#353 := [trans #347 #712]: #352
+#419 := [monotonicity #353 #696]: #697
+#694 := [trans #419 #706]: #420
+#511 := [monotonicity #694 #528 #550]: #552
+#657 := [trans #511 #561]: #656
+#377 := (iff #723 false)
+#716 := (iff #723 #726)
+#725 := (iff #722 false)
+#720 := [rewrite]: #725
+#376 := [monotonicity #720 #720]: #716
+#718 := [trans #376 #717]: #377
+#560 := [monotonicity #718 #657]: #545
+#564 := [trans #560 #563]: #556
+#658 := [monotonicity #564]: #661
+#651 := [trans #658 #649]: #661
+#568 := [quant-inst]: #567
+#643 := [mp #568 #651]: #566
+[unit-resolution #643 #733 #169]: false
+unsat
+fc9ffeceefa2f1854d4c8adb404766a307c30ad9 310 0
+#2 := false
+#41 := 1::int
+decl f3 :: (-> int int int)
+#39 := 3::int
+#38 := 5::int
+#40 := (f3 5::int 3::int)
+#42 := (= #40 1::int)
+#43 := (not #42)
+#170 := [asserted]: #43
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#10 := (f3 #8 #9)
+#728 := (pattern #10)
+#11 := 0::int
+#67 := -1::int
+#71 := (* -1::int #9)
+#68 := (* -1::int #8)
+#74 := (div #68 #71)
+#231 := (* -1::int #74)
+#232 := (+ #10 #231)
+#233 := (= #232 0::int)
+#21 := (div #8 #9)
+#228 := (* -1::int #21)
+#229 := (+ #10 #228)
+#230 := (= #229 0::int)
+#93 := (<= #9 0::int)
+#89 := (<= #8 0::int)
+#186 := (or #89 #93)
+#187 := (not #186)
+#100 := (>= #8 0::int)
+#178 := (or #93 #100)
+#179 := (not #178)
+#193 := (or #179 #187)
+#234 := (ite #193 #230 #233)
+#227 := (= #10 0::int)
+#13 := (= #9 0::int)
+#12 := (= #8 0::int)
+#14 := (or #12 #13)
+#235 := (ite #14 #227 #234)
+#729 := (forall (vars (?v0 int) (?v1 int)) (:pat #728) #235)
+#238 := (forall (vars (?v0 int) (?v1 int)) #235)
+#732 := (iff #238 #729)
+#730 := (iff #235 #235)
+#731 := [refl]: #730
+#733 := [quant-intro #731]: #732
+#198 := (ite #193 #21 #74)
+#201 := (ite #14 0::int #198)
+#204 := (= #10 #201)
+#207 := (forall (vars (?v0 int) (?v1 int)) #204)
+#239 := (iff #207 #238)
+#236 := (iff #204 #235)
+#237 := [rewrite]: #236
+#240 := [quant-intro #237]: #239
+#101 := (not #100)
+#94 := (not #93)
+#104 := (and #94 #101)
+#90 := (not #89)
+#97 := (and #90 #94)
+#107 := (or #97 #104)
+#110 := (ite #107 #21 #74)
+#113 := (ite #14 0::int #110)
+#116 := (= #10 #113)
+#119 := (forall (vars (?v0 int) (?v1 int)) #116)
+#208 := (iff #119 #207)
+#205 := (iff #116 #204)
+#202 := (= #113 #201)
+#199 := (= #110 #198)
+#196 := (iff #107 #193)
+#190 := (or #187 #179)
+#194 := (iff #190 #193)
+#195 := [rewrite]: #194
+#191 := (iff #107 #190)
+#188 := (iff #104 #179)
+#189 := [rewrite]: #188
+#176 := (iff #97 #187)
+#177 := [rewrite]: #176
+#192 := [monotonicity #177 #189]: #191
+#197 := [trans #192 #195]: #196
+#200 := [monotonicity #197]: #199
+#203 := [monotonicity #200]: #202
+#206 := [monotonicity #203]: #205
+#209 := [quant-intro #206]: #208
+#182 := (~ #119 #119)
+#180 := (~ #116 #116)
+#181 := [refl]: #180
+#183 := [nnf-pos #181]: #182
+#23 := (- #9)
+#22 := (- #8)
+#24 := (div #22 #23)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#25 := (ite #20 #21 #24)
+#26 := (ite #14 0::int #25)
+#27 := (= #10 #26)
+#28 := (forall (vars (?v0 int) (?v1 int)) #27)
+#122 := (iff #28 #119)
+#61 := (and #16 #18)
+#64 := (or #17 #61)
+#77 := (ite #64 #21 #74)
+#80 := (ite #14 0::int #77)
+#83 := (= #10 #80)
+#86 := (forall (vars (?v0 int) (?v1 int)) #83)
+#120 := (iff #86 #119)
+#117 := (iff #83 #116)
+#114 := (= #80 #113)
+#111 := (= #77 #110)
+#108 := (iff #64 #107)
+#105 := (iff #61 #104)
+#102 := (iff #18 #101)
+#103 := [rewrite]: #102
+#95 := (iff #16 #94)
+#96 := [rewrite]: #95
+#106 := [monotonicity #96 #103]: #105
+#98 := (iff #17 #97)
+#91 := (iff #15 #90)
+#92 := [rewrite]: #91
+#99 := [monotonicity #92 #96]: #98
+#109 := [monotonicity #99 #106]: #108
+#112 := [monotonicity #109]: #111
+#115 := [monotonicity #112]: #114
+#118 := [monotonicity #115]: #117
+#121 := [quant-intro #118]: #120
+#87 := (iff #28 #86)
+#84 := (iff #27 #83)
+#81 := (= #26 #80)
+#78 := (= #25 #77)
+#75 := (= #24 #74)
+#72 := (= #23 #71)
+#73 := [rewrite]: #72
+#69 := (= #22 #68)
+#70 := [rewrite]: #69
+#76 := [monotonicity #70 #73]: #75
+#65 := (iff #20 #64)
+#62 := (iff #19 #61)
+#63 := [rewrite]: #62
+#66 := [monotonicity #63]: #65
+#79 := [monotonicity #66 #76]: #78
+#82 := [monotonicity #79]: #81
+#85 := [monotonicity #82]: #84
+#88 := [quant-intro #85]: #87
+#123 := [trans #88 #121]: #122
+#60 := [asserted]: #28
+#124 := [mp #60 #123]: #119
+#173 := [mp~ #124 #183]: #119
+#210 := [mp #173 #209]: #207
+#241 := [mp #210 #240]: #238
+#734 := [mp #241 #733]: #729
+#633 := (not #729)
+#634 := (or #633 #42)
+#315 := (* -1::int 3::int)
+#399 := (* -1::int 5::int)
+#400 := (div #399 #315)
+#401 := (* -1::int #400)
+#392 := (+ #40 #401)
+#403 := (= #392 0::int)
+#404 := (div 5::int 3::int)
+#405 := (* -1::int #404)
+#402 := (+ #40 #405)
+#406 := (= #402 0::int)
+#385 := (<= 3::int 0::int)
+#720 := (<= 5::int 0::int)
+#722 := (or #720 #385)
+#379 := (not #722)
+#509 := (>= 5::int 0::int)
+#716 := (or #385 #509)
+#390 := (not #716)
+#391 := (or #390 #379)
+#723 := (ite #391 #406 #403)
+#724 := (= #40 0::int)
+#725 := (= 3::int 0::int)
+#726 := (= 5::int 0::int)
+#721 := (or #726 #725)
+#727 := (ite #721 #724 #723)
+#636 := (or #633 #727)
+#638 := (iff #636 #634)
+#635 := (iff #634 #634)
+#640 := [rewrite]: #635
+#642 := (iff #727 #42)
+#656 := (ite false #724 #42)
+#647 := (iff #656 #42)
+#648 := [rewrite]: #647
+#645 := (iff #727 #656)
+#655 := (iff #723 #42)
+#665 := 2::int
+#662 := (= #40 2::int)
+#1 := true
+#644 := (ite true #42 #662)
+#653 := (iff #644 #42)
+#654 := [rewrite]: #653
+#641 := (iff #723 #644)
+#650 := (iff #403 #662)
+#512 := -2::int
+#563 := (+ -2::int #40)
+#567 := (= #563 0::int)
+#659 := (iff #567 #662)
+#663 := [rewrite]: #659
+#568 := (iff #403 #567)
+#565 := (= #392 #563)
+#661 := (+ #40 -2::int)
+#564 := (= #661 #563)
+#557 := [rewrite]: #564
+#546 := (= #392 #661)
+#657 := (= #401 -2::int)
+#551 := (* -1::int 2::int)
+#660 := (= #551 -2::int)
+#562 := [rewrite]: #660
+#552 := (= #401 #551)
+#666 := (= #400 2::int)
+#672 := -3::int
+#671 := -5::int
+#510 := (div -5::int -3::int)
+#667 := (= #510 2::int)
+#668 := [rewrite]: #667
+#511 := (= #400 #510)
+#679 := (= #315 -3::int)
+#508 := [rewrite]: #679
+#677 := (= #399 -5::int)
+#678 := [rewrite]: #677
+#664 := [monotonicity #678 #508]: #511
+#669 := [trans #664 #668]: #666
+#553 := [monotonicity #669]: #552
+#658 := [trans #553 #562]: #657
+#561 := [monotonicity #658]: #546
+#566 := [trans #561 #557]: #565
+#569 := [monotonicity #566]: #568
+#652 := [trans #569 #663]: #650
+#676 := (iff #406 #42)
+#686 := (+ -1::int #40)
+#530 := (= #686 0::int)
+#674 := (iff #530 #42)
+#675 := [rewrite]: #674
+#531 := (iff #406 #530)
+#688 := (= #402 #686)
+#685 := (+ #40 -1::int)
+#687 := (= #685 #686)
+#682 := [rewrite]: #687
+#680 := (= #402 #685)
+#683 := (= #405 -1::int)
+#407 := (* -1::int 1::int)
+#690 := (= #407 -1::int)
+#694 := [rewrite]: #690
+#689 := (= #405 #407)
+#691 := (= #404 1::int)
+#692 := [rewrite]: #691
+#693 := [monotonicity #692]: #689
+#684 := [trans #693 #694]: #683
+#681 := [monotonicity #684]: #680
+#529 := [trans #681 #682]: #688
+#673 := [monotonicity #529]: #531
+#670 := [trans #673 #675]: #676
+#412 := (iff #391 true)
+#708 := (or false true)
+#710 := (iff #708 true)
+#348 := [rewrite]: #710
+#410 := (iff #391 #708)
+#696 := (iff #379 true)
+#698 := (not false)
+#695 := (iff #698 true)
+#699 := [rewrite]: #695
+#420 := (iff #379 #698)
+#321 := (iff #722 false)
+#378 := (or false false)
+#703 := (iff #378 false)
+#366 := [rewrite]: #703
+#426 := (iff #722 #378)
+#369 := (iff #385 false)
+#705 := [rewrite]: #369
+#424 := (iff #720 false)
+#425 := [rewrite]: #424
+#427 := [monotonicity #425 #705]: #426
+#697 := [trans #427 #366]: #321
+#421 := [monotonicity #697]: #420
+#700 := [trans #421 #699]: #696
+#701 := (iff #390 false)
+#353 := (not true)
+#712 := (iff #353 false)
+#715 := [rewrite]: #712
+#354 := (iff #390 #353)
+#711 := (iff #716 true)
+#709 := (iff #716 #708)
+#706 := (iff #509 true)
+#707 := [rewrite]: #706
+#704 := [monotonicity #705 #707]: #709
+#713 := [trans #704 #348]: #711
+#714 := [monotonicity #713]: #354
+#702 := [trans #714 #715]: #701
+#411 := [monotonicity #702 #700]: #410
+#413 := [trans #411 #348]: #412
+#643 := [monotonicity #413 #670 #652]: #641
+#651 := [trans #643 #654]: #655
+#367 := (iff #721 false)
+#719 := (iff #721 #378)
+#382 := (iff #725 false)
+#718 := [rewrite]: #382
+#717 := (iff #726 false)
+#377 := [rewrite]: #717
+#362 := [monotonicity #377 #718]: #719
+#368 := [trans #362 #366]: #367
+#646 := [monotonicity #368 #651]: #645
+#649 := [trans #646 #648]: #642
+#639 := [monotonicity #649]: #638
+#626 := [trans #639 #640]: #638
+#637 := [quant-inst]: #636
+#627 := [mp #637 #626]: #634
+[unit-resolution #627 #734 #170]: false
+unsat
+334231bac1c10cb79d865cb2c86779cf7dd19631 324 0
+#2 := false
+#68 := -1::int
+decl f3 :: (-> int int int)
+#172 := -3::int
+#38 := 1::int
+#175 := (f3 1::int -3::int)
+#180 := (= #175 -1::int)
+#193 := (not #180)
+#42 := (- 1::int)
+#39 := 3::int
+#40 := (- 3::int)
+#41 := (f3 1::int #40)
+#43 := (= #41 #42)
+#44 := (not #43)
+#196 := (iff #44 #193)
+#183 := (= -1::int #175)
+#188 := (not #183)
+#194 := (iff #188 #193)
+#191 := (iff #183 #180)
+#192 := [rewrite]: #191
+#195 := [monotonicity #192]: #194
+#189 := (iff #44 #188)
+#186 := (iff #43 #183)
+#184 := (iff #180 #183)
+#185 := [rewrite]: #184
+#181 := (iff #43 #180)
+#178 := (= #42 -1::int)
+#179 := [rewrite]: #178
+#176 := (= #41 #175)
+#173 := (= #40 -3::int)
+#174 := [rewrite]: #173
+#177 := [monotonicity #174]: #176
+#182 := [monotonicity #177 #179]: #181
+#187 := [trans #182 #185]: #186
+#190 := [monotonicity #187]: #189
+#197 := [trans #190 #195]: #196
+#171 := [asserted]: #44
+#198 := [mp #171 #197]: #193
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#10 := (f3 #8 #9)
+#756 := (pattern #10)
+#11 := 0::int
+#72 := (* -1::int #9)
+#69 := (* -1::int #8)
+#75 := (div #69 #72)
+#259 := (* -1::int #75)
+#260 := (+ #10 #259)
+#261 := (= #260 0::int)
+#21 := (div #8 #9)
+#256 := (* -1::int #21)
+#257 := (+ #10 #256)
+#258 := (= #257 0::int)
+#94 := (<= #9 0::int)
+#90 := (<= #8 0::int)
+#214 := (or #90 #94)
+#215 := (not #214)
+#101 := (>= #8 0::int)
+#206 := (or #94 #101)
+#207 := (not #206)
+#221 := (or #207 #215)
+#262 := (ite #221 #258 #261)
+#255 := (= #10 0::int)
+#13 := (= #9 0::int)
+#12 := (= #8 0::int)
+#14 := (or #12 #13)
+#263 := (ite #14 #255 #262)
+#757 := (forall (vars (?v0 int) (?v1 int)) (:pat #756) #263)
+#266 := (forall (vars (?v0 int) (?v1 int)) #263)
+#760 := (iff #266 #757)
+#758 := (iff #263 #263)
+#759 := [refl]: #758
+#761 := [quant-intro #759]: #760
+#226 := (ite #221 #21 #75)
+#229 := (ite #14 0::int #226)
+#232 := (= #10 #229)
+#235 := (forall (vars (?v0 int) (?v1 int)) #232)
+#267 := (iff #235 #266)
+#264 := (iff #232 #263)
+#265 := [rewrite]: #264
+#268 := [quant-intro #265]: #267
+#102 := (not #101)
+#95 := (not #94)
+#105 := (and #95 #102)
+#91 := (not #90)
+#98 := (and #91 #95)
+#108 := (or #98 #105)
+#111 := (ite #108 #21 #75)
+#114 := (ite #14 0::int #111)
+#117 := (= #10 #114)
+#120 := (forall (vars (?v0 int) (?v1 int)) #117)
+#236 := (iff #120 #235)
+#233 := (iff #117 #232)
+#230 := (= #114 #229)
+#227 := (= #111 #226)
+#224 := (iff #108 #221)
+#218 := (or #215 #207)
+#222 := (iff #218 #221)
+#223 := [rewrite]: #222
+#219 := (iff #108 #218)
+#216 := (iff #105 #207)
+#217 := [rewrite]: #216
+#204 := (iff #98 #215)
+#205 := [rewrite]: #204
+#220 := [monotonicity #205 #217]: #219
+#225 := [trans #220 #223]: #224
+#228 := [monotonicity #225]: #227
+#231 := [monotonicity #228]: #230
+#234 := [monotonicity #231]: #233
+#237 := [quant-intro #234]: #236
+#210 := (~ #120 #120)
+#208 := (~ #117 #117)
+#209 := [refl]: #208
+#211 := [nnf-pos #209]: #210
+#23 := (- #9)
+#22 := (- #8)
+#24 := (div #22 #23)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#25 := (ite #20 #21 #24)
+#26 := (ite #14 0::int #25)
+#27 := (= #10 #26)
+#28 := (forall (vars (?v0 int) (?v1 int)) #27)
+#123 := (iff #28 #120)
+#62 := (and #16 #18)
+#65 := (or #17 #62)
+#78 := (ite #65 #21 #75)
+#81 := (ite #14 0::int #78)
+#84 := (= #10 #81)
+#87 := (forall (vars (?v0 int) (?v1 int)) #84)
+#121 := (iff #87 #120)
+#118 := (iff #84 #117)
+#115 := (= #81 #114)
+#112 := (= #78 #111)
+#109 := (iff #65 #108)
+#106 := (iff #62 #105)
+#103 := (iff #18 #102)
+#104 := [rewrite]: #103
+#96 := (iff #16 #95)
+#97 := [rewrite]: #96
+#107 := [monotonicity #97 #104]: #106
+#99 := (iff #17 #98)
+#92 := (iff #15 #91)
+#93 := [rewrite]: #92
+#100 := [monotonicity #93 #97]: #99
+#110 := [monotonicity #100 #107]: #109
+#113 := [monotonicity #110]: #112
+#116 := [monotonicity #113]: #115
+#119 := [monotonicity #116]: #118
+#122 := [quant-intro #119]: #121
+#88 := (iff #28 #87)
+#85 := (iff #27 #84)
+#82 := (= #26 #81)
+#79 := (= #25 #78)
+#76 := (= #24 #75)
+#73 := (= #23 #72)
+#74 := [rewrite]: #73
+#70 := (= #22 #69)
+#71 := [rewrite]: #70
+#77 := [monotonicity #71 #74]: #76
+#66 := (iff #20 #65)
+#63 := (iff #19 #62)
+#64 := [rewrite]: #63
+#67 := [monotonicity #64]: #66
+#80 := [monotonicity #67 #77]: #79
+#83 := [monotonicity #80]: #82
+#86 := [monotonicity #83]: #85
+#89 := [quant-intro #86]: #88
+#124 := [trans #89 #122]: #123
+#61 := [asserted]: #28
+#125 := [mp #61 #124]: #120
+#201 := [mp~ #125 #211]: #120
+#238 := [mp #201 #237]: #235
+#269 := [mp #238 #268]: #266
+#762 := [mp #269 #761]: #757
+#672 := (not #757)
+#669 := (or #672 #180)
+#343 := (* -1::int -3::int)
+#427 := (* -1::int 1::int)
+#428 := (div #427 #343)
+#429 := (* -1::int #428)
+#420 := (+ #175 #429)
+#431 := (= #420 0::int)
+#432 := (div 1::int -3::int)
+#433 := (* -1::int #432)
+#430 := (+ #175 #433)
+#434 := (= #430 0::int)
+#413 := (<= -3::int 0::int)
+#748 := (<= 1::int 0::int)
+#750 := (or #748 #413)
+#407 := (not #750)
+#537 := (>= 1::int 0::int)
+#744 := (or #413 #537)
+#418 := (not #744)
+#419 := (or #418 #407)
+#751 := (ite #419 #434 #431)
+#752 := (= #175 0::int)
+#753 := (= -3::int 0::int)
+#754 := (= 1::int 0::int)
+#749 := (or #754 #753)
+#755 := (ite #749 #752 #751)
+#671 := (or #672 #755)
+#682 := (iff #671 #669)
+#679 := (iff #669 #669)
+#684 := [rewrite]: #679
+#678 := (iff #755 #180)
+#585 := (ite false #752 #180)
+#595 := (iff #585 #180)
+#596 := [rewrite]: #595
+#687 := (iff #755 #585)
+#597 := (iff #751 #180)
+#593 := (iff #751 #585)
+#591 := (iff #431 #180)
+#580 := (+ 1::int #175)
+#685 := (= #580 0::int)
+#574 := (iff #685 #180)
+#589 := [rewrite]: #574
+#686 := (iff #431 #685)
+#688 := (= #420 #580)
+#694 := (+ #175 1::int)
+#581 := (= #694 #580)
+#540 := [rewrite]: #581
+#697 := (= #420 #694)
+#695 := (= #429 1::int)
+#536 := (* -1::int -1::int)
+#692 := (= #536 1::int)
+#693 := [rewrite]: #692
+#538 := (= #429 #536)
+#700 := (= #428 -1::int)
+#704 := (div -1::int 3::int)
+#705 := (= #704 -1::int)
+#706 := [rewrite]: #705
+#698 := (= #428 #704)
+#702 := (= #343 3::int)
+#703 := [rewrite]: #702
+#559 := (= #427 -1::int)
+#701 := [rewrite]: #559
+#699 := [monotonicity #701 #703]: #698
+#707 := [trans #699 #706]: #700
+#539 := [monotonicity #707]: #538
+#696 := [trans #539 #693]: #695
+#579 := [monotonicity #696]: #697
+#590 := [trans #579 #540]: #688
+#689 := [monotonicity #590]: #686
+#592 := [trans #689 #589]: #591
+#557 := (iff #434 #752)
+#710 := (= #430 #175)
+#713 := (+ #175 0::int)
+#714 := (= #713 #175)
+#715 := [rewrite]: #714
+#708 := (= #430 #713)
+#711 := (= #433 0::int)
+#435 := (* -1::int 0::int)
+#718 := (= #435 0::int)
+#722 := [rewrite]: #718
+#717 := (= #433 #435)
+#719 := (= #432 0::int)
+#720 := [rewrite]: #719
+#721 := [monotonicity #720]: #717
+#712 := [trans #721 #722]: #711
+#709 := [monotonicity #712]: #708
+#716 := [trans #709 #715]: #710
+#558 := [monotonicity #716]: #557
+#440 := (iff #419 false)
+#406 := (or false false)
+#731 := (iff #406 false)
+#394 := [rewrite]: #731
+#438 := (iff #419 #406)
+#724 := (iff #407 false)
+#1 := true
+#381 := (not true)
+#740 := (iff #381 false)
+#743 := [rewrite]: #740
+#723 := (iff #407 #381)
+#448 := (iff #750 true)
+#454 := (or false true)
+#725 := (iff #454 true)
+#726 := [rewrite]: #725
+#455 := (iff #750 #454)
+#397 := (iff #413 true)
+#733 := [rewrite]: #397
+#452 := (iff #748 false)
+#453 := [rewrite]: #452
+#349 := [monotonicity #453 #733]: #455
+#449 := [trans #349 #726]: #448
+#727 := [monotonicity #449]: #723
+#728 := [trans #727 #743]: #724
+#729 := (iff #418 false)
+#382 := (iff #418 #381)
+#739 := (iff #744 true)
+#736 := (or true true)
+#738 := (iff #736 true)
+#376 := [rewrite]: #738
+#737 := (iff #744 #736)
+#734 := (iff #537 true)
+#735 := [rewrite]: #734
+#732 := [monotonicity #733 #735]: #737
+#741 := [trans #732 #376]: #739
+#742 := [monotonicity #741]: #382
+#730 := [trans #742 #743]: #729
+#439 := [monotonicity #730 #728]: #438
+#441 := [trans #439 #394]: #440
+#594 := [monotonicity #441 #558 #592]: #593
+#690 := [trans #594 #596]: #597
+#395 := (iff #749 false)
+#747 := (iff #749 #406)
+#410 := (iff #753 false)
+#746 := [rewrite]: #410
+#745 := (iff #754 false)
+#405 := [rewrite]: #745
+#390 := [monotonicity #405 #746]: #747
+#396 := [trans #390 #394]: #395
+#691 := [monotonicity #396 #690]: #687
+#680 := [trans #691 #596]: #678
+#683 := [monotonicity #680]: #682
+#673 := [trans #683 #684]: #682
+#681 := [quant-inst]: #671
+#674 := [mp #681 #673]: #669
+[unit-resolution #674 #762 #198]: false
+unsat
+60a3e615568433502bc15c7d472c0ece644ee329 323 0
+#2 := false
+#68 := -1::int
+decl f3 :: (-> int int int)
+#172 := -3::int
+#38 := 3::int
+#175 := (f3 3::int -3::int)
+#180 := (= #175 -1::int)
+#193 := (not #180)
+#41 := 1::int
+#42 := (- 1::int)
+#39 := (- 3::int)
+#40 := (f3 3::int #39)
+#43 := (= #40 #42)
+#44 := (not #43)
+#196 := (iff #44 #193)
+#183 := (= -1::int #175)
+#188 := (not #183)
+#194 := (iff #188 #193)
+#191 := (iff #183 #180)
+#192 := [rewrite]: #191
+#195 := [monotonicity #192]: #194
+#189 := (iff #44 #188)
+#186 := (iff #43 #183)
+#184 := (iff #180 #183)
+#185 := [rewrite]: #184
+#181 := (iff #43 #180)
+#178 := (= #42 -1::int)
+#179 := [rewrite]: #178
+#176 := (= #40 #175)
+#173 := (= #39 -3::int)
+#174 := [rewrite]: #173
+#177 := [monotonicity #174]: #176
+#182 := [monotonicity #177 #179]: #181
+#187 := [trans #182 #185]: #186
+#190 := [monotonicity #187]: #189
+#197 := [trans #190 #195]: #196
+#171 := [asserted]: #44
+#198 := [mp #171 #197]: #193
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#10 := (f3 #8 #9)
+#756 := (pattern #10)
+#11 := 0::int
+#72 := (* -1::int #9)
+#69 := (* -1::int #8)
+#75 := (div #69 #72)
+#259 := (* -1::int #75)
+#260 := (+ #10 #259)
+#261 := (= #260 0::int)
+#21 := (div #8 #9)
+#256 := (* -1::int #21)
+#257 := (+ #10 #256)
+#258 := (= #257 0::int)
+#94 := (<= #9 0::int)
+#90 := (<= #8 0::int)
+#214 := (or #90 #94)
+#215 := (not #214)
+#101 := (>= #8 0::int)
+#206 := (or #94 #101)
+#207 := (not #206)
+#221 := (or #207 #215)
+#262 := (ite #221 #258 #261)
+#255 := (= #10 0::int)
+#13 := (= #9 0::int)
+#12 := (= #8 0::int)
+#14 := (or #12 #13)
+#263 := (ite #14 #255 #262)
+#757 := (forall (vars (?v0 int) (?v1 int)) (:pat #756) #263)
+#266 := (forall (vars (?v0 int) (?v1 int)) #263)
+#760 := (iff #266 #757)
+#758 := (iff #263 #263)
+#759 := [refl]: #758
+#761 := [quant-intro #759]: #760
+#226 := (ite #221 #21 #75)
+#229 := (ite #14 0::int #226)
+#232 := (= #10 #229)
+#235 := (forall (vars (?v0 int) (?v1 int)) #232)
+#267 := (iff #235 #266)
+#264 := (iff #232 #263)
+#265 := [rewrite]: #264
+#268 := [quant-intro #265]: #267
+#102 := (not #101)
+#95 := (not #94)
+#105 := (and #95 #102)
+#91 := (not #90)
+#98 := (and #91 #95)
+#108 := (or #98 #105)
+#111 := (ite #108 #21 #75)
+#114 := (ite #14 0::int #111)
+#117 := (= #10 #114)
+#120 := (forall (vars (?v0 int) (?v1 int)) #117)
+#236 := (iff #120 #235)
+#233 := (iff #117 #232)
+#230 := (= #114 #229)
+#227 := (= #111 #226)
+#224 := (iff #108 #221)
+#218 := (or #215 #207)
+#222 := (iff #218 #221)
+#223 := [rewrite]: #222
+#219 := (iff #108 #218)
+#216 := (iff #105 #207)
+#217 := [rewrite]: #216
+#204 := (iff #98 #215)
+#205 := [rewrite]: #204
+#220 := [monotonicity #205 #217]: #219
+#225 := [trans #220 #223]: #224
+#228 := [monotonicity #225]: #227
+#231 := [monotonicity #228]: #230
+#234 := [monotonicity #231]: #233
+#237 := [quant-intro #234]: #236
+#210 := (~ #120 #120)
+#208 := (~ #117 #117)
+#209 := [refl]: #208
+#211 := [nnf-pos #209]: #210
+#23 := (- #9)
+#22 := (- #8)
+#24 := (div #22 #23)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#25 := (ite #20 #21 #24)
+#26 := (ite #14 0::int #25)
+#27 := (= #10 #26)
+#28 := (forall (vars (?v0 int) (?v1 int)) #27)
+#123 := (iff #28 #120)
+#62 := (and #16 #18)
+#65 := (or #17 #62)
+#78 := (ite #65 #21 #75)
+#81 := (ite #14 0::int #78)
+#84 := (= #10 #81)
+#87 := (forall (vars (?v0 int) (?v1 int)) #84)
+#121 := (iff #87 #120)
+#118 := (iff #84 #117)
+#115 := (= #81 #114)
+#112 := (= #78 #111)
+#109 := (iff #65 #108)
+#106 := (iff #62 #105)
+#103 := (iff #18 #102)
+#104 := [rewrite]: #103
+#96 := (iff #16 #95)
+#97 := [rewrite]: #96
+#107 := [monotonicity #97 #104]: #106
+#99 := (iff #17 #98)
+#92 := (iff #15 #91)
+#93 := [rewrite]: #92
+#100 := [monotonicity #93 #97]: #99
+#110 := [monotonicity #100 #107]: #109
+#113 := [monotonicity #110]: #112
+#116 := [monotonicity #113]: #115
+#119 := [monotonicity #116]: #118
+#122 := [quant-intro #119]: #121
+#88 := (iff #28 #87)
+#85 := (iff #27 #84)
+#82 := (= #26 #81)
+#79 := (= #25 #78)
+#76 := (= #24 #75)
+#73 := (= #23 #72)
+#74 := [rewrite]: #73
+#70 := (= #22 #69)
+#71 := [rewrite]: #70
+#77 := [monotonicity #71 #74]: #76
+#66 := (iff #20 #65)
+#63 := (iff #19 #62)
+#64 := [rewrite]: #63
+#67 := [monotonicity #64]: #66
+#80 := [monotonicity #67 #77]: #79
+#83 := [monotonicity #80]: #82
+#86 := [monotonicity #83]: #85
+#89 := [quant-intro #86]: #88
+#124 := [trans #89 #122]: #123
+#61 := [asserted]: #28
+#125 := [mp #61 #124]: #120
+#201 := [mp~ #125 #211]: #120
+#238 := [mp #201 #237]: #235
+#269 := [mp #238 #268]: #266
+#762 := [mp #269 #761]: #757
+#680 := (not #757)
+#672 := (or #680 #180)
+#343 := (* -1::int -3::int)
+#427 := (* -1::int 3::int)
+#428 := (div #427 #343)
+#429 := (* -1::int #428)
+#420 := (+ #175 #429)
+#431 := (= #420 0::int)
+#432 := (div 3::int -3::int)
+#433 := (* -1::int #432)
+#430 := (+ #175 #433)
+#434 := (= #430 0::int)
+#413 := (<= -3::int 0::int)
+#748 := (<= 3::int 0::int)
+#750 := (or #748 #413)
+#407 := (not #750)
+#537 := (>= 3::int 0::int)
+#744 := (or #413 #537)
+#418 := (not #744)
+#419 := (or #418 #407)
+#751 := (ite #419 #434 #431)
+#752 := (= #175 0::int)
+#753 := (= -3::int 0::int)
+#754 := (= 3::int 0::int)
+#749 := (or #754 #753)
+#755 := (ite #749 #752 #751)
+#669 := (or #680 #755)
+#681 := (iff #669 #672)
+#683 := (iff #672 #672)
+#679 := [rewrite]: #683
+#691 := (iff #755 #180)
+#595 := (ite false #752 #180)
+#690 := (iff #595 #180)
+#687 := [rewrite]: #690
+#596 := (iff #755 #595)
+#593 := (iff #751 #180)
+#574 := (ite false #180 #180)
+#592 := (iff #574 #180)
+#585 := [rewrite]: #592
+#589 := (iff #751 #574)
+#686 := (iff #431 #180)
+#714 := (+ 1::int #175)
+#558 := (= #714 0::int)
+#702 := (iff #558 #180)
+#703 := [rewrite]: #702
+#590 := (iff #431 #558)
+#540 := (= #420 #714)
+#713 := (+ #175 1::int)
+#715 := (= #713 #714)
+#710 := [rewrite]: #715
+#580 := (= #420 #713)
+#697 := (= #429 1::int)
+#435 := (* -1::int -1::int)
+#718 := (= #435 1::int)
+#722 := [rewrite]: #718
+#696 := (= #429 #435)
+#693 := (= #428 -1::int)
+#707 := (div -3::int 3::int)
+#539 := (= #707 -1::int)
+#692 := [rewrite]: #539
+#536 := (= #428 #707)
+#706 := (= #343 3::int)
+#700 := [rewrite]: #706
+#699 := (= #427 -3::int)
+#705 := [rewrite]: #699
+#538 := [monotonicity #705 #700]: #536
+#695 := [trans #538 #692]: #693
+#694 := [monotonicity #695]: #696
+#579 := [trans #694 #722]: #697
+#581 := [monotonicity #579]: #580
+#688 := [trans #581 #710]: #540
+#685 := [monotonicity #688]: #590
+#689 := [trans #685 #703]: #686
+#704 := (iff #434 #180)
+#559 := (iff #434 #558)
+#716 := (= #430 #714)
+#708 := (= #430 #713)
+#711 := (= #433 1::int)
+#717 := (= #433 #435)
+#719 := (= #432 -1::int)
+#720 := [rewrite]: #719
+#721 := [monotonicity #720]: #717
+#712 := [trans #721 #722]: #711
+#709 := [monotonicity #712]: #708
+#557 := [trans #709 #710]: #716
+#701 := [monotonicity #557]: #559
+#698 := [trans #701 #703]: #704
+#440 := (iff #419 false)
+#406 := (or false false)
+#731 := (iff #406 false)
+#394 := [rewrite]: #731
+#438 := (iff #419 #406)
+#724 := (iff #407 false)
+#1 := true
+#381 := (not true)
+#740 := (iff #381 false)
+#743 := [rewrite]: #740
+#723 := (iff #407 #381)
+#448 := (iff #750 true)
+#454 := (or false true)
+#725 := (iff #454 true)
+#726 := [rewrite]: #725
+#455 := (iff #750 #454)
+#397 := (iff #413 true)
+#733 := [rewrite]: #397
+#452 := (iff #748 false)
+#453 := [rewrite]: #452
+#349 := [monotonicity #453 #733]: #455
+#449 := [trans #349 #726]: #448
+#727 := [monotonicity #449]: #723
+#728 := [trans #727 #743]: #724
+#729 := (iff #418 false)
+#382 := (iff #418 #381)
+#739 := (iff #744 true)
+#736 := (or true true)
+#738 := (iff #736 true)
+#376 := [rewrite]: #738
+#737 := (iff #744 #736)
+#734 := (iff #537 true)
+#735 := [rewrite]: #734
+#732 := [monotonicity #733 #735]: #737
+#741 := [trans #732 #376]: #739
+#742 := [monotonicity #741]: #382
+#730 := [trans #742 #743]: #729
+#439 := [monotonicity #730 #728]: #438
+#441 := [trans #439 #394]: #440
+#591 := [monotonicity #441 #698 #689]: #589
+#594 := [trans #591 #585]: #593
+#395 := (iff #749 false)
+#747 := (iff #749 #406)
+#410 := (iff #753 false)
+#746 := [rewrite]: #410
+#745 := (iff #754 false)
+#405 := [rewrite]: #745
+#390 := [monotonicity #405 #746]: #747
+#396 := [trans #390 #394]: #395
+#597 := [monotonicity #396 #594]: #596
+#678 := [trans #597 #687]: #691
+#682 := [monotonicity #678]: #681
+#684 := [trans #682 #679]: #681
+#671 := [quant-inst]: #669
+#673 := [mp #671 #684]: #672
+[unit-resolution #673 #762 #198]: false
+unsat
+0217fb3975a8ed7626e13b93029371f1151b43e4 326 0
+#2 := false
+#179 := -2::int
+decl f3 :: (-> int int int)
+#173 := -3::int
+#38 := 5::int
+#176 := (f3 5::int -3::int)
+#182 := (= #176 -2::int)
+#185 := (not #182)
+#42 := 2::int
+#43 := (- 2::int)
+#39 := 3::int
+#40 := (- 3::int)
+#41 := (f3 5::int #40)
+#44 := (= #41 #43)
+#45 := (not #44)
+#186 := (iff #45 #185)
+#183 := (iff #44 #182)
+#180 := (= #43 -2::int)
+#181 := [rewrite]: #180
+#177 := (= #41 #176)
+#174 := (= #40 -3::int)
+#175 := [rewrite]: #174
+#178 := [monotonicity #175]: #177
+#184 := [monotonicity #178 #181]: #183
+#187 := [monotonicity #184]: #186
+#172 := [asserted]: #45
+#190 := [mp #172 #187]: #185
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#10 := (f3 #8 #9)
+#747 := (pattern #10)
+#11 := 0::int
+#69 := -1::int
+#73 := (* -1::int #9)
+#70 := (* -1::int #8)
+#76 := (div #70 #73)
+#249 := (* -1::int #76)
+#250 := (+ #10 #249)
+#251 := (= #250 0::int)
+#21 := (div #8 #9)
+#246 := (* -1::int #21)
+#247 := (+ #10 #246)
+#248 := (= #247 0::int)
+#95 := (<= #9 0::int)
+#91 := (<= #8 0::int)
+#204 := (or #91 #95)
+#205 := (not #204)
+#102 := (>= #8 0::int)
+#196 := (or #95 #102)
+#197 := (not #196)
+#211 := (or #197 #205)
+#252 := (ite #211 #248 #251)
+#245 := (= #10 0::int)
+#13 := (= #9 0::int)
+#12 := (= #8 0::int)
+#14 := (or #12 #13)
+#253 := (ite #14 #245 #252)
+#748 := (forall (vars (?v0 int) (?v1 int)) (:pat #747) #253)
+#256 := (forall (vars (?v0 int) (?v1 int)) #253)
+#751 := (iff #256 #748)
+#749 := (iff #253 #253)
+#750 := [refl]: #749
+#752 := [quant-intro #750]: #751
+#216 := (ite #211 #21 #76)
+#219 := (ite #14 0::int #216)
+#222 := (= #10 #219)
+#225 := (forall (vars (?v0 int) (?v1 int)) #222)
+#257 := (iff #225 #256)
+#254 := (iff #222 #253)
+#255 := [rewrite]: #254
+#258 := [quant-intro #255]: #257
+#103 := (not #102)
+#96 := (not #95)
+#106 := (and #96 #103)
+#92 := (not #91)
+#99 := (and #92 #96)
+#109 := (or #99 #106)
+#112 := (ite #109 #21 #76)
+#115 := (ite #14 0::int #112)
+#118 := (= #10 #115)
+#121 := (forall (vars (?v0 int) (?v1 int)) #118)
+#226 := (iff #121 #225)
+#223 := (iff #118 #222)
+#220 := (= #115 #219)
+#217 := (= #112 #216)
+#214 := (iff #109 #211)
+#208 := (or #205 #197)
+#212 := (iff #208 #211)
+#213 := [rewrite]: #212
+#209 := (iff #109 #208)
+#206 := (iff #106 #197)
+#207 := [rewrite]: #206
+#194 := (iff #99 #205)
+#195 := [rewrite]: #194
+#210 := [monotonicity #195 #207]: #209
+#215 := [trans #210 #213]: #214
+#218 := [monotonicity #215]: #217
+#221 := [monotonicity #218]: #220
+#224 := [monotonicity #221]: #223
+#227 := [quant-intro #224]: #226
+#200 := (~ #121 #121)
+#198 := (~ #118 #118)
+#199 := [refl]: #198
+#201 := [nnf-pos #199]: #200
+#23 := (- #9)
+#22 := (- #8)
+#24 := (div #22 #23)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#25 := (ite #20 #21 #24)
+#26 := (ite #14 0::int #25)
+#27 := (= #10 #26)
+#28 := (forall (vars (?v0 int) (?v1 int)) #27)
+#124 := (iff #28 #121)
+#63 := (and #16 #18)
+#66 := (or #17 #63)
+#79 := (ite #66 #21 #76)
+#82 := (ite #14 0::int #79)
+#85 := (= #10 #82)
+#88 := (forall (vars (?v0 int) (?v1 int)) #85)
+#122 := (iff #88 #121)
+#119 := (iff #85 #118)
+#116 := (= #82 #115)
+#113 := (= #79 #112)
+#110 := (iff #66 #109)
+#107 := (iff #63 #106)
+#104 := (iff #18 #103)
+#105 := [rewrite]: #104
+#97 := (iff #16 #96)
+#98 := [rewrite]: #97
+#108 := [monotonicity #98 #105]: #107
+#100 := (iff #17 #99)
+#93 := (iff #15 #92)
+#94 := [rewrite]: #93
+#101 := [monotonicity #94 #98]: #100
+#111 := [monotonicity #101 #108]: #110
+#114 := [monotonicity #111]: #113
+#117 := [monotonicity #114]: #116
+#120 := [monotonicity #117]: #119
+#123 := [quant-intro #120]: #122
+#89 := (iff #28 #88)
+#86 := (iff #27 #85)
+#83 := (= #26 #82)
+#80 := (= #25 #79)
+#77 := (= #24 #76)
+#74 := (= #23 #73)
+#75 := [rewrite]: #74
+#71 := (= #22 #70)
+#72 := [rewrite]: #71
+#78 := [monotonicity #72 #75]: #77
+#67 := (iff #20 #66)
+#64 := (iff #19 #63)
+#65 := [rewrite]: #64
+#68 := [monotonicity #65]: #67
+#81 := [monotonicity #68 #78]: #80
+#84 := [monotonicity #81]: #83
+#87 := [monotonicity #84]: #86
+#90 := [quant-intro #87]: #89
+#125 := [trans #90 #123]: #124
+#62 := [asserted]: #28
+#126 := [mp #62 #125]: #121
+#191 := [mp~ #126 #201]: #121
+#228 := [mp #191 #227]: #225
+#259 := [mp #228 #258]: #256
+#753 := [mp #259 #752]: #748
+#667 := (not #748)
+#661 := (or #667 #182)
+#333 := (* -1::int -3::int)
+#418 := (* -1::int 5::int)
+#419 := (div #418 #333)
+#420 := (* -1::int #419)
+#411 := (+ #176 #420)
+#422 := (= #411 0::int)
+#423 := (div 5::int -3::int)
+#351 := (* -1::int #423)
+#424 := (+ #176 #351)
+#421 := (= #424 0::int)
+#425 := (<= -3::int 0::int)
+#404 := (<= 5::int 0::int)
+#739 := (or #404 #425)
+#741 := (not #739)
+#398 := (>= 5::int 0::int)
+#528 := (or #425 #398)
+#735 := (not #528)
+#409 := (or #735 #741)
+#410 := (ite #409 #421 #422)
+#742 := (= #176 0::int)
+#743 := (= -3::int 0::int)
+#744 := (= 5::int 0::int)
+#745 := (or #744 #743)
+#740 := (ite #745 #742 #410)
+#668 := (or #667 #740)
+#653 := (iff #668 #661)
+#656 := (iff #661 #661)
+#657 := [rewrite]: #656
+#665 := (iff #740 #182)
+#673 := (ite false #742 #182)
+#675 := (iff #673 #182)
+#664 := [rewrite]: #675
+#674 := (iff #740 #673)
+#662 := (iff #410 #182)
+#693 := (= #176 -1::int)
+#682 := (ite false #693 #182)
+#663 := (iff #682 #182)
+#660 := [rewrite]: #663
+#669 := (iff #410 #682)
+#681 := (iff #422 #182)
+#565 := (+ 2::int #176)
+#584 := (= #565 0::int)
+#587 := (iff #584 #182)
+#588 := [rewrite]: #587
+#585 := (iff #422 #584)
+#583 := (= #411 #565)
+#676 := (+ #176 2::int)
+#580 := (= #676 #565)
+#582 := [rewrite]: #580
+#677 := (= #411 #676)
+#679 := (= #420 2::int)
+#688 := (* -1::int -2::int)
+#572 := (= #688 2::int)
+#531 := [rewrite]: #572
+#570 := (= #420 #688)
+#687 := (= #419 -2::int)
+#696 := -5::int
+#529 := (div -5::int 3::int)
+#684 := (= #529 -2::int)
+#686 := [rewrite]: #684
+#530 := (= #419 #529)
+#698 := (= #333 3::int)
+#527 := [rewrite]: #698
+#697 := (= #418 -5::int)
+#691 := [rewrite]: #697
+#683 := [monotonicity #691 #527]: #530
+#685 := [trans #683 #686]: #687
+#571 := [monotonicity #685]: #570
+#581 := [trans #571 #531]: #679
+#680 := [monotonicity #581]: #677
+#576 := [trans #680 #582]: #583
+#586 := [monotonicity #576]: #585
+#678 := [trans #586 #588]: #681
+#689 := (iff #421 #693)
+#712 := 1::int
+#705 := (+ 1::int #176)
+#549 := (= #705 0::int)
+#694 := (iff #549 #693)
+#695 := [rewrite]: #694
+#550 := (iff #421 #549)
+#707 := (= #424 #705)
+#704 := (+ #176 1::int)
+#706 := (= #704 #705)
+#701 := [rewrite]: #706
+#699 := (= #424 #704)
+#702 := (= #351 1::int)
+#711 := (* -1::int -1::int)
+#709 := (= #711 1::int)
+#713 := [rewrite]: #709
+#426 := (= #351 #711)
+#432 := (= #423 -1::int)
+#710 := [rewrite]: #432
+#708 := [monotonicity #710]: #426
+#703 := [trans #708 #713]: #702
+#700 := [monotonicity #703]: #699
+#548 := [trans #700 #701]: #707
+#692 := [monotonicity #548]: #550
+#690 := [trans #692 #695]: #689
+#430 := (iff #409 false)
+#737 := (or false false)
+#381 := (iff #737 false)
+#722 := [rewrite]: #381
+#719 := (iff #409 #737)
+#718 := (iff #741 false)
+#1 := true
+#732 := (not true)
+#733 := (iff #732 false)
+#731 := [rewrite]: #733
+#440 := (iff #741 #732)
+#717 := (iff #739 true)
+#444 := (or false true)
+#339 := (iff #444 true)
+#716 := [rewrite]: #339
+#445 := (iff #739 #444)
+#387 := (iff #425 true)
+#388 := [rewrite]: #387
+#721 := (iff #404 false)
+#443 := [rewrite]: #721
+#446 := [monotonicity #443 #388]: #445
+#439 := [trans #446 #716]: #717
+#714 := [monotonicity #439]: #440
+#715 := [trans #714 #731]: #718
+#734 := (iff #735 false)
+#372 := (iff #735 #732)
+#367 := (iff #528 true)
+#726 := (or true true)
+#723 := (iff #726 true)
+#729 := [rewrite]: #723
+#727 := (iff #528 #726)
+#724 := (iff #398 true)
+#725 := [rewrite]: #724
+#728 := [monotonicity #388 #725]: #727
+#730 := [trans #728 #729]: #367
+#373 := [monotonicity #730]: #372
+#720 := [trans #373 #731]: #734
+#429 := [monotonicity #720 #715]: #719
+#431 := [trans #429 #722]: #430
+#671 := [monotonicity #431 #690 #678]: #669
+#672 := [trans #671 #660]: #662
+#385 := (iff #745 false)
+#397 := (iff #745 #737)
+#396 := (iff #743 false)
+#401 := [rewrite]: #396
+#746 := (iff #744 false)
+#736 := [rewrite]: #746
+#738 := [monotonicity #736 #401]: #397
+#386 := [trans #738 #722]: #385
+#670 := [monotonicity #386 #672]: #674
+#666 := [trans #670 #664]: #665
+#655 := [monotonicity #666]: #653
+#658 := [trans #655 #657]: #653
+#652 := [quant-inst]: #668
+#654 := [mp #652 #658]: #661
+[unit-resolution #654 #753 #190]: false
+unsat
+a5e7ca3a1a908b3cfa96261b62a5e5f564ff08b5 321 0
+#2 := false
+#67 := -1::int
+decl f3 :: (-> int int int)
+#40 := 3::int
+#173 := (f3 -1::int 3::int)
+#176 := (= #173 -1::int)
+#189 := (not #176)
+#38 := 1::int
+#39 := (- 1::int)
+#41 := (f3 #39 3::int)
+#42 := (= #41 #39)
+#43 := (not #42)
+#192 := (iff #43 #189)
+#179 := (= -1::int #173)
+#184 := (not #179)
+#190 := (iff #184 #189)
+#187 := (iff #179 #176)
+#188 := [rewrite]: #187
+#191 := [monotonicity #188]: #190
+#185 := (iff #43 #184)
+#182 := (iff #42 #179)
+#180 := (iff #176 #179)
+#181 := [rewrite]: #180
+#177 := (iff #42 #176)
+#171 := (= #39 -1::int)
+#172 := [rewrite]: #171
+#174 := (= #41 #173)
+#175 := [monotonicity #172]: #174
+#178 := [monotonicity #175 #172]: #177
+#183 := [trans #178 #181]: #182
+#186 := [monotonicity #183]: #185
+#193 := [trans #186 #191]: #192
+#170 := [asserted]: #43
+#194 := [mp #170 #193]: #189
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#10 := (f3 #8 #9)
+#752 := (pattern #10)
+#11 := 0::int
+#71 := (* -1::int #9)
+#68 := (* -1::int #8)
+#74 := (div #68 #71)
+#255 := (* -1::int #74)
+#256 := (+ #10 #255)
+#257 := (= #256 0::int)
+#21 := (div #8 #9)
+#252 := (* -1::int #21)
+#253 := (+ #10 #252)
+#254 := (= #253 0::int)
+#93 := (<= #9 0::int)
+#89 := (<= #8 0::int)
+#210 := (or #89 #93)
+#211 := (not #210)
+#100 := (>= #8 0::int)
+#202 := (or #93 #100)
+#203 := (not #202)
+#217 := (or #203 #211)
+#258 := (ite #217 #254 #257)
+#251 := (= #10 0::int)
+#13 := (= #9 0::int)
+#12 := (= #8 0::int)
+#14 := (or #12 #13)
+#259 := (ite #14 #251 #258)
+#753 := (forall (vars (?v0 int) (?v1 int)) (:pat #752) #259)
+#262 := (forall (vars (?v0 int) (?v1 int)) #259)
+#756 := (iff #262 #753)
+#754 := (iff #259 #259)
+#755 := [refl]: #754
+#757 := [quant-intro #755]: #756
+#222 := (ite #217 #21 #74)
+#225 := (ite #14 0::int #222)
+#228 := (= #10 #225)
+#231 := (forall (vars (?v0 int) (?v1 int)) #228)
+#263 := (iff #231 #262)
+#260 := (iff #228 #259)
+#261 := [rewrite]: #260
+#264 := [quant-intro #261]: #263
+#101 := (not #100)
+#94 := (not #93)
+#104 := (and #94 #101)
+#90 := (not #89)
+#97 := (and #90 #94)
+#107 := (or #97 #104)
+#110 := (ite #107 #21 #74)
+#113 := (ite #14 0::int #110)
+#116 := (= #10 #113)
+#119 := (forall (vars (?v0 int) (?v1 int)) #116)
+#232 := (iff #119 #231)
+#229 := (iff #116 #228)
+#226 := (= #113 #225)
+#223 := (= #110 #222)
+#220 := (iff #107 #217)
+#214 := (or #211 #203)
+#218 := (iff #214 #217)
+#219 := [rewrite]: #218
+#215 := (iff #107 #214)
+#212 := (iff #104 #203)
+#213 := [rewrite]: #212
+#200 := (iff #97 #211)
+#201 := [rewrite]: #200
+#216 := [monotonicity #201 #213]: #215
+#221 := [trans #216 #219]: #220
+#224 := [monotonicity #221]: #223
+#227 := [monotonicity #224]: #226
+#230 := [monotonicity #227]: #229
+#233 := [quant-intro #230]: #232
+#206 := (~ #119 #119)
+#204 := (~ #116 #116)
+#205 := [refl]: #204
+#207 := [nnf-pos #205]: #206
+#23 := (- #9)
+#22 := (- #8)
+#24 := (div #22 #23)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#25 := (ite #20 #21 #24)
+#26 := (ite #14 0::int #25)
+#27 := (= #10 #26)
+#28 := (forall (vars (?v0 int) (?v1 int)) #27)
+#122 := (iff #28 #119)
+#61 := (and #16 #18)
+#64 := (or #17 #61)
+#77 := (ite #64 #21 #74)
+#80 := (ite #14 0::int #77)
+#83 := (= #10 #80)
+#86 := (forall (vars (?v0 int) (?v1 int)) #83)
+#120 := (iff #86 #119)
+#117 := (iff #83 #116)
+#114 := (= #80 #113)
+#111 := (= #77 #110)
+#108 := (iff #64 #107)
+#105 := (iff #61 #104)
+#102 := (iff #18 #101)
+#103 := [rewrite]: #102
+#95 := (iff #16 #94)
+#96 := [rewrite]: #95
+#106 := [monotonicity #96 #103]: #105
+#98 := (iff #17 #97)
+#91 := (iff #15 #90)
+#92 := [rewrite]: #91
+#99 := [monotonicity #92 #96]: #98
+#109 := [monotonicity #99 #106]: #108
+#112 := [monotonicity #109]: #111
+#115 := [monotonicity #112]: #114
+#118 := [monotonicity #115]: #117
+#121 := [quant-intro #118]: #120
+#87 := (iff #28 #86)
+#84 := (iff #27 #83)
+#81 := (= #26 #80)
+#78 := (= #25 #77)
+#75 := (= #24 #74)
+#72 := (= #23 #71)
+#73 := [rewrite]: #72
+#69 := (= #22 #68)
+#70 := [rewrite]: #69
+#76 := [monotonicity #70 #73]: #75
+#65 := (iff #20 #64)
+#62 := (iff #19 #61)
+#63 := [rewrite]: #62
+#66 := [monotonicity #63]: #65
+#79 := [monotonicity #66 #76]: #78
+#82 := [monotonicity #79]: #81
+#85 := [monotonicity #82]: #84
+#88 := [quant-intro #85]: #87
+#123 := [trans #88 #121]: #122
+#60 := [asserted]: #28
+#124 := [mp #60 #123]: #119
+#197 := [mp~ #124 #207]: #119
+#234 := [mp #197 #233]: #231
+#265 := [mp #234 #264]: #262
+#758 := [mp #265 #757]: #753
+#665 := (not #753)
+#667 := (or #665 #176)
+#339 := (* -1::int 3::int)
+#423 := (* -1::int -1::int)
+#424 := (div #423 #339)
+#425 := (* -1::int #424)
+#416 := (+ #173 #425)
+#427 := (= #416 0::int)
+#428 := (div -1::int 3::int)
+#429 := (* -1::int #428)
+#426 := (+ #173 #429)
+#430 := (= #426 0::int)
+#409 := (<= 3::int 0::int)
+#744 := (<= -1::int 0::int)
+#746 := (or #744 #409)
+#403 := (not #746)
+#533 := (>= -1::int 0::int)
+#740 := (or #409 #533)
+#414 := (not #740)
+#415 := (or #414 #403)
+#747 := (ite #415 #430 #427)
+#748 := (= #173 0::int)
+#749 := (= 3::int 0::int)
+#750 := (= -1::int 0::int)
+#745 := (or #750 #749)
+#751 := (ite #745 #748 #747)
+#677 := (or #665 #751)
+#679 := (iff #677 #667)
+#680 := (iff #667 #667)
+#669 := [rewrite]: #680
+#676 := (iff #751 #176)
+#593 := (ite false #748 #176)
+#687 := (iff #593 #176)
+#674 := [rewrite]: #687
+#686 := (iff #751 #593)
+#591 := (iff #747 #176)
+#1 := true
+#587 := (ite true #176 #748)
+#589 := (iff #587 #176)
+#590 := [rewrite]: #589
+#588 := (iff #747 #587)
+#570 := (iff #427 #748)
+#682 := (= #416 #173)
+#577 := (+ #173 0::int)
+#586 := (= #577 #173)
+#681 := [rewrite]: #586
+#536 := (= #416 #577)
+#575 := (= #425 0::int)
+#689 := (* -1::int 0::int)
+#690 := (= #689 0::int)
+#693 := [rewrite]: #690
+#691 := (= #425 #689)
+#535 := (= #424 0::int)
+#694 := -3::int
+#702 := (div 1::int -3::int)
+#532 := (= #702 0::int)
+#534 := [rewrite]: #532
+#696 := (= #424 #702)
+#695 := (= #339 -3::int)
+#701 := [rewrite]: #695
+#717 := (= #423 1::int)
+#714 := [rewrite]: #717
+#703 := [monotonicity #714 #701]: #696
+#688 := [trans #703 #534]: #535
+#692 := [monotonicity #688]: #691
+#576 := [trans #692 #693]: #575
+#684 := [monotonicity #576]: #536
+#685 := [trans #684 #681]: #682
+#585 := [monotonicity #685]: #570
+#699 := (iff #430 #176)
+#705 := (+ 1::int #173)
+#553 := (= #705 0::int)
+#697 := (iff #553 #176)
+#698 := [rewrite]: #697
+#554 := (iff #430 #553)
+#706 := (= #426 #705)
+#708 := (+ #173 1::int)
+#710 := (= #708 #705)
+#711 := [rewrite]: #710
+#709 := (= #426 #708)
+#718 := (= #429 1::int)
+#431 := (= #429 #423)
+#715 := (= #428 -1::int)
+#716 := [rewrite]: #715
+#713 := [monotonicity #716]: #431
+#707 := [trans #713 #714]: #718
+#704 := [monotonicity #707]: #709
+#712 := [trans #704 #711]: #706
+#555 := [monotonicity #712]: #554
+#700 := [trans #555 #698]: #699
+#436 := (iff #415 true)
+#726 := (or true false)
+#450 := (iff #726 true)
+#451 := [rewrite]: #450
+#434 := (iff #415 #726)
+#720 := (iff #403 false)
+#722 := (not true)
+#719 := (iff #722 false)
+#723 := [rewrite]: #719
+#444 := (iff #403 #722)
+#345 := (iff #746 true)
+#448 := (iff #746 #726)
+#393 := (iff #409 false)
+#729 := [rewrite]: #393
+#739 := (iff #744 true)
+#725 := [rewrite]: #739
+#449 := [monotonicity #725 #729]: #448
+#721 := [trans #449 #451]: #345
+#445 := [monotonicity #721]: #444
+#724 := [trans #445 #723]: #720
+#738 := (iff #414 true)
+#372 := (not false)
+#377 := (iff #372 true)
+#378 := [rewrite]: #377
+#735 := (iff #414 #372)
+#728 := (iff #740 false)
+#402 := (or false false)
+#727 := (iff #402 false)
+#390 := [rewrite]: #727
+#732 := (iff #740 #402)
+#730 := (iff #533 false)
+#731 := [rewrite]: #730
+#733 := [monotonicity #729 #731]: #732
+#734 := [trans #733 #390]: #728
+#737 := [monotonicity #734]: #735
+#736 := [trans #737 #378]: #738
+#435 := [monotonicity #736 #724]: #434
+#437 := [trans #435 #451]: #436
+#581 := [monotonicity #437 #700 #585]: #588
+#592 := [trans #581 #590]: #591
+#391 := (iff #745 false)
+#743 := (iff #745 #402)
+#406 := (iff #749 false)
+#742 := [rewrite]: #406
+#741 := (iff #750 false)
+#401 := [rewrite]: #741
+#386 := [monotonicity #401 #742]: #743
+#392 := [trans #386 #390]: #391
+#683 := [monotonicity #392 #592]: #686
+#668 := [trans #683 #674]: #676
+#675 := [monotonicity #668]: #679
+#670 := [trans #675 #669]: #679
+#678 := [quant-inst]: #677
+#671 := [mp #678 #670]: #667
+[unit-resolution #671 #758 #194]: false
+unsat
+4a26f9179dc065943e8b39f87c50b753014abca8 323 0
+#2 := false
+#68 := -1::int
+decl f3 :: (-> int int int)
+#38 := 3::int
+#172 := -3::int
+#175 := (f3 -3::int 3::int)
+#180 := (= #175 -1::int)
+#193 := (not #180)
+#41 := 1::int
+#42 := (- 1::int)
+#39 := (- 3::int)
+#40 := (f3 #39 3::int)
+#43 := (= #40 #42)
+#44 := (not #43)
+#196 := (iff #44 #193)
+#183 := (= -1::int #175)
+#188 := (not #183)
+#194 := (iff #188 #193)
+#191 := (iff #183 #180)
+#192 := [rewrite]: #191
+#195 := [monotonicity #192]: #194
+#189 := (iff #44 #188)
+#186 := (iff #43 #183)
+#184 := (iff #180 #183)
+#185 := [rewrite]: #184
+#181 := (iff #43 #180)
+#178 := (= #42 -1::int)
+#179 := [rewrite]: #178
+#176 := (= #40 #175)
+#173 := (= #39 -3::int)
+#174 := [rewrite]: #173
+#177 := [monotonicity #174]: #176
+#182 := [monotonicity #177 #179]: #181
+#187 := [trans #182 #185]: #186
+#190 := [monotonicity #187]: #189
+#197 := [trans #190 #195]: #196
+#171 := [asserted]: #44
+#198 := [mp #171 #197]: #193
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#10 := (f3 #8 #9)
+#756 := (pattern #10)
+#11 := 0::int
+#72 := (* -1::int #9)
+#69 := (* -1::int #8)
+#75 := (div #69 #72)
+#259 := (* -1::int #75)
+#260 := (+ #10 #259)
+#261 := (= #260 0::int)
+#21 := (div #8 #9)
+#256 := (* -1::int #21)
+#257 := (+ #10 #256)
+#258 := (= #257 0::int)
+#94 := (<= #9 0::int)
+#90 := (<= #8 0::int)
+#214 := (or #90 #94)
+#215 := (not #214)
+#101 := (>= #8 0::int)
+#206 := (or #94 #101)
+#207 := (not #206)
+#221 := (or #207 #215)
+#262 := (ite #221 #258 #261)
+#255 := (= #10 0::int)
+#13 := (= #9 0::int)
+#12 := (= #8 0::int)
+#14 := (or #12 #13)
+#263 := (ite #14 #255 #262)
+#757 := (forall (vars (?v0 int) (?v1 int)) (:pat #756) #263)
+#266 := (forall (vars (?v0 int) (?v1 int)) #263)
+#760 := (iff #266 #757)
+#758 := (iff #263 #263)
+#759 := [refl]: #758
+#761 := [quant-intro #759]: #760
+#226 := (ite #221 #21 #75)
+#229 := (ite #14 0::int #226)
+#232 := (= #10 #229)
+#235 := (forall (vars (?v0 int) (?v1 int)) #232)
+#267 := (iff #235 #266)
+#264 := (iff #232 #263)
+#265 := [rewrite]: #264
+#268 := [quant-intro #265]: #267
+#102 := (not #101)
+#95 := (not #94)
+#105 := (and #95 #102)
+#91 := (not #90)
+#98 := (and #91 #95)
+#108 := (or #98 #105)
+#111 := (ite #108 #21 #75)
+#114 := (ite #14 0::int #111)
+#117 := (= #10 #114)
+#120 := (forall (vars (?v0 int) (?v1 int)) #117)
+#236 := (iff #120 #235)
+#233 := (iff #117 #232)
+#230 := (= #114 #229)
+#227 := (= #111 #226)
+#224 := (iff #108 #221)
+#218 := (or #215 #207)
+#222 := (iff #218 #221)
+#223 := [rewrite]: #222
+#219 := (iff #108 #218)
+#216 := (iff #105 #207)
+#217 := [rewrite]: #216
+#204 := (iff #98 #215)
+#205 := [rewrite]: #204
+#220 := [monotonicity #205 #217]: #219
+#225 := [trans #220 #223]: #224
+#228 := [monotonicity #225]: #227
+#231 := [monotonicity #228]: #230
+#234 := [monotonicity #231]: #233
+#237 := [quant-intro #234]: #236
+#210 := (~ #120 #120)
+#208 := (~ #117 #117)
+#209 := [refl]: #208
+#211 := [nnf-pos #209]: #210
+#23 := (- #9)
+#22 := (- #8)
+#24 := (div #22 #23)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#25 := (ite #20 #21 #24)
+#26 := (ite #14 0::int #25)
+#27 := (= #10 #26)
+#28 := (forall (vars (?v0 int) (?v1 int)) #27)
+#123 := (iff #28 #120)
+#62 := (and #16 #18)
+#65 := (or #17 #62)
+#78 := (ite #65 #21 #75)
+#81 := (ite #14 0::int #78)
+#84 := (= #10 #81)
+#87 := (forall (vars (?v0 int) (?v1 int)) #84)
+#121 := (iff #87 #120)
+#118 := (iff #84 #117)
+#115 := (= #81 #114)
+#112 := (= #78 #111)
+#109 := (iff #65 #108)
+#106 := (iff #62 #105)
+#103 := (iff #18 #102)
+#104 := [rewrite]: #103
+#96 := (iff #16 #95)
+#97 := [rewrite]: #96
+#107 := [monotonicity #97 #104]: #106
+#99 := (iff #17 #98)
+#92 := (iff #15 #91)
+#93 := [rewrite]: #92
+#100 := [monotonicity #93 #97]: #99
+#110 := [monotonicity #100 #107]: #109
+#113 := [monotonicity #110]: #112
+#116 := [monotonicity #113]: #115
+#119 := [monotonicity #116]: #118
+#122 := [quant-intro #119]: #121
+#88 := (iff #28 #87)
+#85 := (iff #27 #84)
+#82 := (= #26 #81)
+#79 := (= #25 #78)
+#76 := (= #24 #75)
+#73 := (= #23 #72)
+#74 := [rewrite]: #73
+#70 := (= #22 #69)
+#71 := [rewrite]: #70
+#77 := [monotonicity #71 #74]: #76
+#66 := (iff #20 #65)
+#63 := (iff #19 #62)
+#64 := [rewrite]: #63
+#67 := [monotonicity #64]: #66
+#80 := [monotonicity #67 #77]: #79
+#83 := [monotonicity #80]: #82
+#86 := [monotonicity #83]: #85
+#89 := [quant-intro #86]: #88
+#124 := [trans #89 #122]: #123
+#61 := [asserted]: #28
+#125 := [mp #61 #124]: #120
+#201 := [mp~ #125 #211]: #120
+#238 := [mp #201 #237]: #235
+#269 := [mp #238 #268]: #266
+#762 := [mp #269 #761]: #757
+#680 := (not #757)
+#672 := (or #680 #180)
+#343 := (* -1::int 3::int)
+#427 := (* -1::int -3::int)
+#428 := (div #427 #343)
+#429 := (* -1::int #428)
+#420 := (+ #175 #429)
+#431 := (= #420 0::int)
+#432 := (div -3::int 3::int)
+#433 := (* -1::int #432)
+#430 := (+ #175 #433)
+#434 := (= #430 0::int)
+#413 := (<= 3::int 0::int)
+#748 := (<= -3::int 0::int)
+#750 := (or #748 #413)
+#407 := (not #750)
+#537 := (>= -3::int 0::int)
+#744 := (or #413 #537)
+#418 := (not #744)
+#419 := (or #418 #407)
+#751 := (ite #419 #434 #431)
+#752 := (= #175 0::int)
+#753 := (= 3::int 0::int)
+#754 := (= -3::int 0::int)
+#749 := (or #754 #753)
+#755 := (ite #749 #752 #751)
+#669 := (or #680 #755)
+#681 := (iff #669 #672)
+#683 := (iff #672 #672)
+#679 := [rewrite]: #683
+#691 := (iff #755 #180)
+#595 := (ite false #752 #180)
+#690 := (iff #595 #180)
+#687 := [rewrite]: #690
+#596 := (iff #755 #595)
+#593 := (iff #751 #180)
+#1 := true
+#574 := (ite true #180 #180)
+#592 := (iff #574 #180)
+#585 := [rewrite]: #592
+#589 := (iff #751 #574)
+#686 := (iff #431 #180)
+#714 := (+ 1::int #175)
+#558 := (= #714 0::int)
+#702 := (iff #558 #180)
+#703 := [rewrite]: #702
+#590 := (iff #431 #558)
+#540 := (= #420 #714)
+#713 := (+ #175 1::int)
+#715 := (= #713 #714)
+#710 := [rewrite]: #715
+#580 := (= #420 #713)
+#697 := (= #429 1::int)
+#435 := (* -1::int -1::int)
+#718 := (= #435 1::int)
+#722 := [rewrite]: #718
+#696 := (= #429 #435)
+#693 := (= #428 -1::int)
+#707 := (div 3::int -3::int)
+#539 := (= #707 -1::int)
+#692 := [rewrite]: #539
+#536 := (= #428 #707)
+#706 := (= #343 -3::int)
+#700 := [rewrite]: #706
+#699 := (= #427 3::int)
+#705 := [rewrite]: #699
+#538 := [monotonicity #705 #700]: #536
+#695 := [trans #538 #692]: #693
+#694 := [monotonicity #695]: #696
+#579 := [trans #694 #722]: #697
+#581 := [monotonicity #579]: #580
+#688 := [trans #581 #710]: #540
+#685 := [monotonicity #688]: #590
+#689 := [trans #685 #703]: #686
+#704 := (iff #434 #180)
+#559 := (iff #434 #558)
+#716 := (= #430 #714)
+#708 := (= #430 #713)
+#711 := (= #433 1::int)
+#717 := (= #433 #435)
+#719 := (= #432 -1::int)
+#720 := [rewrite]: #719
+#721 := [monotonicity #720]: #717
+#712 := [trans #721 #722]: #711
+#709 := [monotonicity #712]: #708
+#557 := [trans #709 #710]: #716
+#701 := [monotonicity #557]: #559
+#698 := [trans #701 #703]: #704
+#440 := (iff #419 true)
+#730 := (or true false)
+#454 := (iff #730 true)
+#455 := [rewrite]: #454
+#438 := (iff #419 #730)
+#724 := (iff #407 false)
+#726 := (not true)
+#723 := (iff #726 false)
+#727 := [rewrite]: #723
+#448 := (iff #407 #726)
+#349 := (iff #750 true)
+#452 := (iff #750 #730)
+#397 := (iff #413 false)
+#733 := [rewrite]: #397
+#743 := (iff #748 true)
+#729 := [rewrite]: #743
+#453 := [monotonicity #729 #733]: #452
+#725 := [trans #453 #455]: #349
+#449 := [monotonicity #725]: #448
+#728 := [trans #449 #727]: #724
+#742 := (iff #418 true)
+#376 := (not false)
+#381 := (iff #376 true)
+#382 := [rewrite]: #381
+#739 := (iff #418 #376)
+#732 := (iff #744 false)
+#406 := (or false false)
+#731 := (iff #406 false)
+#394 := [rewrite]: #731
+#736 := (iff #744 #406)
+#734 := (iff #537 false)
+#735 := [rewrite]: #734
+#737 := [monotonicity #733 #735]: #736
+#738 := [trans #737 #394]: #732
+#741 := [monotonicity #738]: #739
+#740 := [trans #741 #382]: #742
+#439 := [monotonicity #740 #728]: #438
+#441 := [trans #439 #455]: #440
+#591 := [monotonicity #441 #698 #689]: #589
+#594 := [trans #591 #585]: #593
+#395 := (iff #749 false)
+#747 := (iff #749 #406)
+#410 := (iff #753 false)
+#746 := [rewrite]: #410
+#745 := (iff #754 false)
+#405 := [rewrite]: #745
+#390 := [monotonicity #405 #746]: #747
+#396 := [trans #390 #394]: #395
+#597 := [monotonicity #396 #594]: #596
+#678 := [trans #597 #687]: #691
+#682 := [monotonicity #678]: #681
+#684 := [trans #682 #679]: #681
+#671 := [quant-inst]: #669
+#673 := [mp #671 #684]: #672
+[unit-resolution #673 #762 #198]: false
+unsat
+aa5c848d75f817ce17fcfbcc43fdae3201eb9151 326 0
+#2 := false
+#179 := -2::int
+decl f3 :: (-> int int int)
+#40 := 3::int
+#173 := -5::int
+#176 := (f3 -5::int 3::int)
+#182 := (= #176 -2::int)
+#185 := (not #182)
+#42 := 2::int
+#43 := (- 2::int)
+#38 := 5::int
+#39 := (- 5::int)
+#41 := (f3 #39 3::int)
+#44 := (= #41 #43)
+#45 := (not #44)
+#186 := (iff #45 #185)
+#183 := (iff #44 #182)
+#180 := (= #43 -2::int)
+#181 := [rewrite]: #180
+#177 := (= #41 #176)
+#174 := (= #39 -5::int)
+#175 := [rewrite]: #174
+#178 := [monotonicity #175]: #177
+#184 := [monotonicity #178 #181]: #183
+#187 := [monotonicity #184]: #186
+#172 := [asserted]: #45
+#190 := [mp #172 #187]: #185
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#10 := (f3 #8 #9)
+#747 := (pattern #10)
+#11 := 0::int
+#69 := -1::int
+#73 := (* -1::int #9)
+#70 := (* -1::int #8)
+#76 := (div #70 #73)
+#249 := (* -1::int #76)
+#250 := (+ #10 #249)
+#251 := (= #250 0::int)
+#21 := (div #8 #9)
+#246 := (* -1::int #21)
+#247 := (+ #10 #246)
+#248 := (= #247 0::int)
+#95 := (<= #9 0::int)
+#91 := (<= #8 0::int)
+#204 := (or #91 #95)
+#205 := (not #204)
+#102 := (>= #8 0::int)
+#196 := (or #95 #102)
+#197 := (not #196)
+#211 := (or #197 #205)
+#252 := (ite #211 #248 #251)
+#245 := (= #10 0::int)
+#13 := (= #9 0::int)
+#12 := (= #8 0::int)
+#14 := (or #12 #13)
+#253 := (ite #14 #245 #252)
+#748 := (forall (vars (?v0 int) (?v1 int)) (:pat #747) #253)
+#256 := (forall (vars (?v0 int) (?v1 int)) #253)
+#751 := (iff #256 #748)
+#749 := (iff #253 #253)
+#750 := [refl]: #749
+#752 := [quant-intro #750]: #751
+#216 := (ite #211 #21 #76)
+#219 := (ite #14 0::int #216)
+#222 := (= #10 #219)
+#225 := (forall (vars (?v0 int) (?v1 int)) #222)
+#257 := (iff #225 #256)
+#254 := (iff #222 #253)
+#255 := [rewrite]: #254
+#258 := [quant-intro #255]: #257
+#103 := (not #102)
+#96 := (not #95)
+#106 := (and #96 #103)
+#92 := (not #91)
+#99 := (and #92 #96)
+#109 := (or #99 #106)
+#112 := (ite #109 #21 #76)
+#115 := (ite #14 0::int #112)
+#118 := (= #10 #115)
+#121 := (forall (vars (?v0 int) (?v1 int)) #118)
+#226 := (iff #121 #225)
+#223 := (iff #118 #222)
+#220 := (= #115 #219)
+#217 := (= #112 #216)
+#214 := (iff #109 #211)
+#208 := (or #205 #197)
+#212 := (iff #208 #211)
+#213 := [rewrite]: #212
+#209 := (iff #109 #208)
+#206 := (iff #106 #197)
+#207 := [rewrite]: #206
+#194 := (iff #99 #205)
+#195 := [rewrite]: #194
+#210 := [monotonicity #195 #207]: #209
+#215 := [trans #210 #213]: #214
+#218 := [monotonicity #215]: #217
+#221 := [monotonicity #218]: #220
+#224 := [monotonicity #221]: #223
+#227 := [quant-intro #224]: #226
+#200 := (~ #121 #121)
+#198 := (~ #118 #118)
+#199 := [refl]: #198
+#201 := [nnf-pos #199]: #200
+#23 := (- #9)
+#22 := (- #8)
+#24 := (div #22 #23)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#25 := (ite #20 #21 #24)
+#26 := (ite #14 0::int #25)
+#27 := (= #10 #26)
+#28 := (forall (vars (?v0 int) (?v1 int)) #27)
+#124 := (iff #28 #121)
+#63 := (and #16 #18)
+#66 := (or #17 #63)
+#79 := (ite #66 #21 #76)
+#82 := (ite #14 0::int #79)
+#85 := (= #10 #82)
+#88 := (forall (vars (?v0 int) (?v1 int)) #85)
+#122 := (iff #88 #121)
+#119 := (iff #85 #118)
+#116 := (= #82 #115)
+#113 := (= #79 #112)
+#110 := (iff #66 #109)
+#107 := (iff #63 #106)
+#104 := (iff #18 #103)
+#105 := [rewrite]: #104
+#97 := (iff #16 #96)
+#98 := [rewrite]: #97
+#108 := [monotonicity #98 #105]: #107
+#100 := (iff #17 #99)
+#93 := (iff #15 #92)
+#94 := [rewrite]: #93
+#101 := [monotonicity #94 #98]: #100
+#111 := [monotonicity #101 #108]: #110
+#114 := [monotonicity #111]: #113
+#117 := [monotonicity #114]: #116
+#120 := [monotonicity #117]: #119
+#123 := [quant-intro #120]: #122
+#89 := (iff #28 #88)
+#86 := (iff #27 #85)
+#83 := (= #26 #82)
+#80 := (= #25 #79)
+#77 := (= #24 #76)
+#74 := (= #23 #73)
+#75 := [rewrite]: #74
+#71 := (= #22 #70)
+#72 := [rewrite]: #71
+#78 := [monotonicity #72 #75]: #77
+#67 := (iff #20 #66)
+#64 := (iff #19 #63)
+#65 := [rewrite]: #64
+#68 := [monotonicity #65]: #67
+#81 := [monotonicity #68 #78]: #80
+#84 := [monotonicity #81]: #83
+#87 := [monotonicity #84]: #86
+#90 := [quant-intro #87]: #89
+#125 := [trans #90 #123]: #124
+#62 := [asserted]: #28
+#126 := [mp #62 #125]: #121
+#191 := [mp~ #126 #201]: #121
+#228 := [mp #191 #227]: #225
+#259 := [mp #228 #258]: #256
+#753 := [mp #259 #752]: #748
+#667 := (not #748)
+#661 := (or #667 #182)
+#333 := (* -1::int 3::int)
+#418 := (* -1::int -5::int)
+#419 := (div #418 #333)
+#420 := (* -1::int #419)
+#411 := (+ #176 #420)
+#422 := (= #411 0::int)
+#423 := (div -5::int 3::int)
+#351 := (* -1::int #423)
+#424 := (+ #176 #351)
+#421 := (= #424 0::int)
+#425 := (<= 3::int 0::int)
+#404 := (<= -5::int 0::int)
+#739 := (or #404 #425)
+#741 := (not #739)
+#398 := (>= -5::int 0::int)
+#528 := (or #425 #398)
+#735 := (not #528)
+#409 := (or #735 #741)
+#410 := (ite #409 #421 #422)
+#742 := (= #176 0::int)
+#743 := (= 3::int 0::int)
+#744 := (= -5::int 0::int)
+#745 := (or #744 #743)
+#740 := (ite #745 #742 #410)
+#668 := (or #667 #740)
+#653 := (iff #668 #661)
+#656 := (iff #661 #661)
+#657 := [rewrite]: #656
+#665 := (iff #740 #182)
+#673 := (ite false #742 #182)
+#675 := (iff #673 #182)
+#664 := [rewrite]: #675
+#674 := (iff #740 #673)
+#662 := (iff #410 #182)
+#586 := (= #176 -1::int)
+#1 := true
+#682 := (ite true #182 #586)
+#663 := (iff #682 #182)
+#660 := [rewrite]: #663
+#669 := (iff #410 #682)
+#681 := (iff #422 #586)
+#570 := 1::int
+#680 := (+ 1::int #176)
+#576 := (= #680 0::int)
+#587 := (iff #576 #586)
+#588 := [rewrite]: #587
+#584 := (iff #422 #576)
+#582 := (= #411 #680)
+#581 := (+ #176 1::int)
+#565 := (= #581 #680)
+#580 := [rewrite]: #565
+#676 := (= #411 #581)
+#531 := (= #420 1::int)
+#687 := (* -1::int -1::int)
+#571 := (= #687 1::int)
+#572 := [rewrite]: #571
+#685 := (= #420 #687)
+#684 := (= #419 -1::int)
+#696 := -3::int
+#698 := (div 5::int -3::int)
+#530 := (= #698 -1::int)
+#683 := [rewrite]: #530
+#527 := (= #419 #698)
+#697 := (= #333 -3::int)
+#691 := [rewrite]: #697
+#689 := (= #418 5::int)
+#690 := [rewrite]: #689
+#529 := [monotonicity #690 #691]: #527
+#686 := [trans #529 #683]: #684
+#688 := [monotonicity #686]: #685
+#679 := [trans #688 #572]: #531
+#677 := [monotonicity #679]: #676
+#583 := [trans #677 #580]: #582
+#585 := [monotonicity #583]: #584
+#678 := [trans #585 #588]: #681
+#694 := (iff #421 #182)
+#700 := (+ 2::int #176)
+#548 := (= #700 0::int)
+#692 := (iff #548 #182)
+#693 := [rewrite]: #692
+#549 := (iff #421 #548)
+#701 := (= #424 #700)
+#703 := (+ #176 2::int)
+#705 := (= #703 #700)
+#706 := [rewrite]: #705
+#704 := (= #424 #703)
+#713 := (= #351 2::int)
+#711 := (* -1::int -2::int)
+#712 := (= #711 2::int)
+#709 := [rewrite]: #712
+#426 := (= #351 #711)
+#432 := (= #423 -2::int)
+#710 := [rewrite]: #432
+#708 := [monotonicity #710]: #426
+#702 := [trans #708 #709]: #713
+#699 := [monotonicity #702]: #704
+#707 := [trans #699 #706]: #701
+#550 := [monotonicity #707]: #549
+#695 := [trans #550 #693]: #694
+#430 := (iff #409 true)
+#720 := (or true false)
+#444 := (iff #720 true)
+#445 := [rewrite]: #444
+#719 := (iff #409 #720)
+#718 := (iff #741 false)
+#716 := (not true)
+#440 := (iff #716 false)
+#714 := [rewrite]: #440
+#717 := (iff #741 #716)
+#446 := (iff #739 true)
+#721 := (iff #739 #720)
+#387 := (iff #425 false)
+#388 := [rewrite]: #387
+#731 := (iff #404 true)
+#734 := [rewrite]: #731
+#443 := [monotonicity #734 #388]: #721
+#339 := [trans #443 #445]: #446
+#439 := [monotonicity #339]: #717
+#715 := [trans #439 #714]: #718
+#373 := (iff #735 true)
+#729 := (not false)
+#732 := (iff #729 true)
+#372 := [rewrite]: #732
+#367 := (iff #735 #729)
+#728 := (iff #528 false)
+#737 := (or false false)
+#381 := (iff #737 false)
+#722 := [rewrite]: #381
+#726 := (iff #528 #737)
+#724 := (iff #398 false)
+#725 := [rewrite]: #724
+#727 := [monotonicity #388 #725]: #726
+#723 := [trans #727 #722]: #728
+#730 := [monotonicity #723]: #367
+#733 := [trans #730 #372]: #373
+#429 := [monotonicity #733 #715]: #719
+#431 := [trans #429 #445]: #430
+#671 := [monotonicity #431 #695 #678]: #669
+#672 := [trans #671 #660]: #662
+#385 := (iff #745 false)
+#397 := (iff #745 #737)
+#396 := (iff #743 false)
+#401 := [rewrite]: #396
+#746 := (iff #744 false)
+#736 := [rewrite]: #746
+#738 := [monotonicity #736 #401]: #397
+#386 := [trans #738 #722]: #385
+#670 := [monotonicity #386 #672]: #674
+#666 := [trans #670 #664]: #665
+#655 := [monotonicity #666]: #653
+#658 := [trans #655 #657]: #653
+#652 := [quant-inst]: #668
+#654 := [mp #652 #658]: #661
+[unit-resolution #654 #753 #190]: false
+unsat
+4ca477e5a43a16dacbb90f983ac5a8510c16d464 327 0
+#2 := false
+#11 := 0::int
+decl f3 :: (-> int int int)
+#174 := -3::int
+#68 := -1::int
+#177 := (f3 -1::int -3::int)
+#180 := (= #177 0::int)
+#193 := (not #180)
+#40 := 3::int
+#41 := (- 3::int)
+#38 := 1::int
+#39 := (- 1::int)
+#42 := (f3 #39 #41)
+#43 := (= #42 0::int)
+#44 := (not #43)
+#196 := (iff #44 #193)
+#183 := (= 0::int #177)
+#188 := (not #183)
+#194 := (iff #188 #193)
+#191 := (iff #183 #180)
+#192 := [rewrite]: #191
+#195 := [monotonicity #192]: #194
+#189 := (iff #44 #188)
+#186 := (iff #43 #183)
+#184 := (iff #180 #183)
+#185 := [rewrite]: #184
+#181 := (iff #43 #180)
+#178 := (= #42 #177)
+#175 := (= #41 -3::int)
+#176 := [rewrite]: #175
+#172 := (= #39 -1::int)
+#173 := [rewrite]: #172
+#179 := [monotonicity #173 #176]: #178
+#182 := [monotonicity #179]: #181
+#187 := [trans #182 #185]: #186
+#190 := [monotonicity #187]: #189
+#197 := [trans #190 #195]: #196
+#171 := [asserted]: #44
+#198 := [mp #171 #197]: #193
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#10 := (f3 #8 #9)
+#756 := (pattern #10)
+#72 := (* -1::int #9)
+#69 := (* -1::int #8)
+#75 := (div #69 #72)
+#259 := (* -1::int #75)
+#260 := (+ #10 #259)
+#261 := (= #260 0::int)
+#21 := (div #8 #9)
+#256 := (* -1::int #21)
+#257 := (+ #10 #256)
+#258 := (= #257 0::int)
+#94 := (<= #9 0::int)
+#90 := (<= #8 0::int)
+#214 := (or #90 #94)
+#215 := (not #214)
+#101 := (>= #8 0::int)
+#206 := (or #94 #101)
+#207 := (not #206)
+#221 := (or #207 #215)
+#262 := (ite #221 #258 #261)
+#255 := (= #10 0::int)
+#13 := (= #9 0::int)
+#12 := (= #8 0::int)
+#14 := (or #12 #13)
+#263 := (ite #14 #255 #262)
+#757 := (forall (vars (?v0 int) (?v1 int)) (:pat #756) #263)
+#266 := (forall (vars (?v0 int) (?v1 int)) #263)
+#760 := (iff #266 #757)
+#758 := (iff #263 #263)
+#759 := [refl]: #758
+#761 := [quant-intro #759]: #760
+#226 := (ite #221 #21 #75)
+#229 := (ite #14 0::int #226)
+#232 := (= #10 #229)
+#235 := (forall (vars (?v0 int) (?v1 int)) #232)
+#267 := (iff #235 #266)
+#264 := (iff #232 #263)
+#265 := [rewrite]: #264
+#268 := [quant-intro #265]: #267
+#102 := (not #101)
+#95 := (not #94)
+#105 := (and #95 #102)
+#91 := (not #90)
+#98 := (and #91 #95)
+#108 := (or #98 #105)
+#111 := (ite #108 #21 #75)
+#114 := (ite #14 0::int #111)
+#117 := (= #10 #114)
+#120 := (forall (vars (?v0 int) (?v1 int)) #117)
+#236 := (iff #120 #235)
+#233 := (iff #117 #232)
+#230 := (= #114 #229)
+#227 := (= #111 #226)
+#224 := (iff #108 #221)
+#218 := (or #215 #207)
+#222 := (iff #218 #221)
+#223 := [rewrite]: #222
+#219 := (iff #108 #218)
+#216 := (iff #105 #207)
+#217 := [rewrite]: #216
+#204 := (iff #98 #215)
+#205 := [rewrite]: #204
+#220 := [monotonicity #205 #217]: #219
+#225 := [trans #220 #223]: #224
+#228 := [monotonicity #225]: #227
+#231 := [monotonicity #228]: #230
+#234 := [monotonicity #231]: #233
+#237 := [quant-intro #234]: #236
+#210 := (~ #120 #120)
+#208 := (~ #117 #117)
+#209 := [refl]: #208
+#211 := [nnf-pos #209]: #210
+#23 := (- #9)
+#22 := (- #8)
+#24 := (div #22 #23)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#25 := (ite #20 #21 #24)
+#26 := (ite #14 0::int #25)
+#27 := (= #10 #26)
+#28 := (forall (vars (?v0 int) (?v1 int)) #27)
+#123 := (iff #28 #120)
+#62 := (and #16 #18)
+#65 := (or #17 #62)
+#78 := (ite #65 #21 #75)
+#81 := (ite #14 0::int #78)
+#84 := (= #10 #81)
+#87 := (forall (vars (?v0 int) (?v1 int)) #84)
+#121 := (iff #87 #120)
+#118 := (iff #84 #117)
+#115 := (= #81 #114)
+#112 := (= #78 #111)
+#109 := (iff #65 #108)
+#106 := (iff #62 #105)
+#103 := (iff #18 #102)
+#104 := [rewrite]: #103
+#96 := (iff #16 #95)
+#97 := [rewrite]: #96
+#107 := [monotonicity #97 #104]: #106
+#99 := (iff #17 #98)
+#92 := (iff #15 #91)
+#93 := [rewrite]: #92
+#100 := [monotonicity #93 #97]: #99
+#110 := [monotonicity #100 #107]: #109
+#113 := [monotonicity #110]: #112
+#116 := [monotonicity #113]: #115
+#119 := [monotonicity #116]: #118
+#122 := [quant-intro #119]: #121
+#88 := (iff #28 #87)
+#85 := (iff #27 #84)
+#82 := (= #26 #81)
+#79 := (= #25 #78)
+#76 := (= #24 #75)
+#73 := (= #23 #72)
+#74 := [rewrite]: #73
+#70 := (= #22 #69)
+#71 := [rewrite]: #70
+#77 := [monotonicity #71 #74]: #76
+#66 := (iff #20 #65)
+#63 := (iff #19 #62)
+#64 := [rewrite]: #63
+#67 := [monotonicity #64]: #66
+#80 := [monotonicity #67 #77]: #79
+#83 := [monotonicity #80]: #82
+#86 := [monotonicity #83]: #85
+#89 := [quant-intro #86]: #88
+#124 := [trans #89 #122]: #123
+#61 := [asserted]: #28
+#125 := [mp #61 #124]: #120
+#201 := [mp~ #125 #211]: #120
+#238 := [mp #201 #237]: #235
+#269 := [mp #238 #268]: #266
+#762 := [mp #269 #761]: #757
+#681 := (not #757)
+#682 := (or #681 #180)
+#343 := (* -1::int -3::int)
+#427 := (* -1::int -1::int)
+#428 := (div #427 #343)
+#429 := (* -1::int #428)
+#420 := (+ #177 #429)
+#431 := (= #420 0::int)
+#432 := (div -1::int -3::int)
+#433 := (* -1::int #432)
+#430 := (+ #177 #433)
+#434 := (= #430 0::int)
+#413 := (<= -3::int 0::int)
+#748 := (<= -1::int 0::int)
+#750 := (or #748 #413)
+#407 := (not #750)
+#537 := (>= -1::int 0::int)
+#744 := (or #413 #537)
+#418 := (not #744)
+#419 := (or #418 #407)
+#751 := (ite #419 #434 #431)
+#752 := (= -3::int 0::int)
+#753 := (= -1::int 0::int)
+#754 := (or #753 #752)
+#749 := (ite #754 #180 #751)
+#683 := (or #681 #749)
+#684 := (iff #683 #682)
+#674 := (iff #682 #682)
+#675 := [rewrite]: #674
+#669 := (iff #749 #180)
+#687 := (ite false #180 #180)
+#680 := (iff #687 #180)
+#672 := [rewrite]: #680
+#691 := (iff #749 #687)
+#597 := (iff #751 #180)
+#701 := (= #177 1::int)
+#585 := (ite false #701 #180)
+#595 := (iff #585 #180)
+#596 := [rewrite]: #595
+#593 := (iff #751 #585)
+#591 := (iff #431 #180)
+#574 := (= #420 #177)
+#688 := (+ #177 0::int)
+#686 := (= #688 #177)
+#689 := [rewrite]: #686
+#590 := (= #420 #688)
+#581 := (= #429 0::int)
+#696 := (* -1::int 0::int)
+#579 := (= #696 0::int)
+#580 := [rewrite]: #579
+#694 := (= #429 #696)
+#693 := (= #428 0::int)
+#707 := (div 1::int 3::int)
+#539 := (= #707 0::int)
+#692 := [rewrite]: #539
+#536 := (= #428 #707)
+#706 := (= #343 3::int)
+#700 := [rewrite]: #706
+#699 := (= #427 1::int)
+#705 := [rewrite]: #699
+#538 := [monotonicity #705 #700]: #536
+#695 := [trans #538 #692]: #693
+#697 := [monotonicity #695]: #694
+#540 := [trans #697 #580]: #581
+#685 := [monotonicity #540]: #590
+#589 := [trans #685 #689]: #574
+#592 := [monotonicity #589]: #591
+#704 := (iff #434 #701)
+#709 := (+ -1::int #177)
+#557 := (= #709 0::int)
+#702 := (iff #557 #701)
+#703 := [rewrite]: #702
+#558 := (iff #434 #557)
+#710 := (= #430 #709)
+#712 := (+ #177 -1::int)
+#714 := (= #712 #709)
+#715 := [rewrite]: #714
+#713 := (= #430 #712)
+#722 := (= #433 -1::int)
+#720 := (* -1::int 1::int)
+#721 := (= #720 -1::int)
+#718 := [rewrite]: #721
+#435 := (= #433 #720)
+#441 := (= #432 1::int)
+#719 := [rewrite]: #441
+#717 := [monotonicity #719]: #435
+#711 := [trans #717 #718]: #722
+#708 := [monotonicity #711]: #713
+#716 := [trans #708 #715]: #710
+#559 := [monotonicity #716]: #558
+#698 := [trans #559 #703]: #704
+#439 := (iff #419 false)
+#746 := (or false false)
+#390 := (iff #746 false)
+#731 := [rewrite]: #390
+#728 := (iff #419 #746)
+#727 := (iff #407 false)
+#1 := true
+#741 := (not true)
+#742 := (iff #741 false)
+#740 := [rewrite]: #742
+#449 := (iff #407 #741)
+#726 := (iff #750 true)
+#453 := (or true true)
+#349 := (iff #453 true)
+#725 := [rewrite]: #349
+#454 := (iff #750 #453)
+#396 := (iff #413 true)
+#397 := [rewrite]: #396
+#730 := (iff #748 true)
+#452 := [rewrite]: #730
+#455 := [monotonicity #452 #397]: #454
+#448 := [trans #455 #725]: #726
+#723 := [monotonicity #448]: #449
+#724 := [trans #723 #740]: #727
+#743 := (iff #418 false)
+#381 := (iff #418 #741)
+#376 := (iff #744 true)
+#735 := (or true false)
+#732 := (iff #735 true)
+#738 := [rewrite]: #732
+#736 := (iff #744 #735)
+#733 := (iff #537 false)
+#734 := [rewrite]: #733
+#737 := [monotonicity #397 #734]: #736
+#739 := [trans #737 #738]: #376
+#382 := [monotonicity #739]: #381
+#729 := [trans #382 #740]: #743
+#438 := [monotonicity #729 #724]: #728
+#440 := [trans #438 #731]: #439
+#594 := [monotonicity #440 #698 #592]: #593
+#690 := [trans #594 #596]: #597
+#394 := (iff #754 false)
+#406 := (iff #754 #746)
+#405 := (iff #752 false)
+#410 := [rewrite]: #405
+#755 := (iff #753 false)
+#745 := [rewrite]: #755
+#747 := [monotonicity #745 #410]: #406
+#395 := [trans #747 #731]: #394
+#678 := [monotonicity #395 #690]: #691
+#671 := [trans #678 #672]: #669
+#673 := [monotonicity #671]: #684
+#676 := [trans #673 #675]: #684
+#679 := [quant-inst]: #683
+#670 := [mp #679 #676]: #682
+[unit-resolution #670 #762 #198]: false
+unsat
+ddc5a6b304650a08be4d46afd02b8b25ea1c25a2 311 0
+#2 := false
+#41 := 1::int
+decl f3 :: (-> int int int)
+#171 := -3::int
+#174 := (f3 -3::int -3::int)
+#177 := (= #174 1::int)
+#190 := (not #177)
+#38 := 3::int
+#39 := (- 3::int)
+#40 := (f3 #39 #39)
+#42 := (= #40 1::int)
+#43 := (not #42)
+#193 := (iff #43 #190)
+#180 := (= 1::int #174)
+#185 := (not #180)
+#191 := (iff #185 #190)
+#188 := (iff #180 #177)
+#189 := [rewrite]: #188
+#192 := [monotonicity #189]: #191
+#186 := (iff #43 #185)
+#183 := (iff #42 #180)
+#181 := (iff #177 #180)
+#182 := [rewrite]: #181
+#178 := (iff #42 #177)
+#175 := (= #40 #174)
+#172 := (= #39 -3::int)
+#173 := [rewrite]: #172
+#176 := [monotonicity #173 #173]: #175
+#179 := [monotonicity #176]: #178
+#184 := [trans #179 #182]: #183
+#187 := [monotonicity #184]: #186
+#194 := [trans #187 #192]: #193
+#170 := [asserted]: #43
+#195 := [mp #170 #194]: #190
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#10 := (f3 #8 #9)
+#753 := (pattern #10)
+#11 := 0::int
+#67 := -1::int
+#71 := (* -1::int #9)
+#68 := (* -1::int #8)
+#74 := (div #68 #71)
+#256 := (* -1::int #74)
+#257 := (+ #10 #256)
+#258 := (= #257 0::int)
+#21 := (div #8 #9)
+#253 := (* -1::int #21)
+#254 := (+ #10 #253)
+#255 := (= #254 0::int)
+#93 := (<= #9 0::int)
+#89 := (<= #8 0::int)
+#211 := (or #89 #93)
+#212 := (not #211)
+#100 := (>= #8 0::int)
+#203 := (or #93 #100)
+#204 := (not #203)
+#218 := (or #204 #212)
+#259 := (ite #218 #255 #258)
+#252 := (= #10 0::int)
+#13 := (= #9 0::int)
+#12 := (= #8 0::int)
+#14 := (or #12 #13)
+#260 := (ite #14 #252 #259)
+#754 := (forall (vars (?v0 int) (?v1 int)) (:pat #753) #260)
+#263 := (forall (vars (?v0 int) (?v1 int)) #260)
+#757 := (iff #263 #754)
+#755 := (iff #260 #260)
+#756 := [refl]: #755
+#758 := [quant-intro #756]: #757
+#223 := (ite #218 #21 #74)
+#226 := (ite #14 0::int #223)
+#229 := (= #10 #226)
+#232 := (forall (vars (?v0 int) (?v1 int)) #229)
+#264 := (iff #232 #263)
+#261 := (iff #229 #260)
+#262 := [rewrite]: #261
+#265 := [quant-intro #262]: #264
+#101 := (not #100)
+#94 := (not #93)
+#104 := (and #94 #101)
+#90 := (not #89)
+#97 := (and #90 #94)
+#107 := (or #97 #104)
+#110 := (ite #107 #21 #74)
+#113 := (ite #14 0::int #110)
+#116 := (= #10 #113)
+#119 := (forall (vars (?v0 int) (?v1 int)) #116)
+#233 := (iff #119 #232)
+#230 := (iff #116 #229)
+#227 := (= #113 #226)
+#224 := (= #110 #223)
+#221 := (iff #107 #218)
+#215 := (or #212 #204)
+#219 := (iff #215 #218)
+#220 := [rewrite]: #219
+#216 := (iff #107 #215)
+#213 := (iff #104 #204)
+#214 := [rewrite]: #213
+#201 := (iff #97 #212)
+#202 := [rewrite]: #201
+#217 := [monotonicity #202 #214]: #216
+#222 := [trans #217 #220]: #221
+#225 := [monotonicity #222]: #224
+#228 := [monotonicity #225]: #227
+#231 := [monotonicity #228]: #230
+#234 := [quant-intro #231]: #233
+#207 := (~ #119 #119)
+#205 := (~ #116 #116)
+#206 := [refl]: #205
+#208 := [nnf-pos #206]: #207
+#23 := (- #9)
+#22 := (- #8)
+#24 := (div #22 #23)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#25 := (ite #20 #21 #24)
+#26 := (ite #14 0::int #25)
+#27 := (= #10 #26)
+#28 := (forall (vars (?v0 int) (?v1 int)) #27)
+#122 := (iff #28 #119)
+#61 := (and #16 #18)
+#64 := (or #17 #61)
+#77 := (ite #64 #21 #74)
+#80 := (ite #14 0::int #77)
+#83 := (= #10 #80)
+#86 := (forall (vars (?v0 int) (?v1 int)) #83)
+#120 := (iff #86 #119)
+#117 := (iff #83 #116)
+#114 := (= #80 #113)
+#111 := (= #77 #110)
+#108 := (iff #64 #107)
+#105 := (iff #61 #104)
+#102 := (iff #18 #101)
+#103 := [rewrite]: #102
+#95 := (iff #16 #94)
+#96 := [rewrite]: #95
+#106 := [monotonicity #96 #103]: #105
+#98 := (iff #17 #97)
+#91 := (iff #15 #90)
+#92 := [rewrite]: #91
+#99 := [monotonicity #92 #96]: #98
+#109 := [monotonicity #99 #106]: #108
+#112 := [monotonicity #109]: #111
+#115 := [monotonicity #112]: #114
+#118 := [monotonicity #115]: #117
+#121 := [quant-intro #118]: #120
+#87 := (iff #28 #86)
+#84 := (iff #27 #83)
+#81 := (= #26 #80)
+#78 := (= #25 #77)
+#75 := (= #24 #74)
+#72 := (= #23 #71)
+#73 := [rewrite]: #72
+#69 := (= #22 #68)
+#70 := [rewrite]: #69
+#76 := [monotonicity #70 #73]: #75
+#65 := (iff #20 #64)
+#62 := (iff #19 #61)
+#63 := [rewrite]: #62
+#66 := [monotonicity #63]: #65
+#79 := [monotonicity #66 #76]: #78
+#82 := [monotonicity #79]: #81
+#85 := [monotonicity #82]: #84
+#88 := [quant-intro #85]: #87
+#123 := [trans #88 #121]: #122
+#60 := [asserted]: #28
+#124 := [mp #60 #123]: #119
+#198 := [mp~ #124 #208]: #119
+#235 := [mp #198 #234]: #232
+#266 := [mp #235 #265]: #263
+#759 := [mp #266 #758]: #754
+#590 := (not #754)
+#591 := (or #590 #177)
+#340 := (* -1::int -3::int)
+#424 := (div #340 #340)
+#425 := (* -1::int #424)
+#426 := (+ #174 #425)
+#417 := (= #426 0::int)
+#428 := (div -3::int -3::int)
+#429 := (* -1::int #428)
+#430 := (+ #174 #429)
+#427 := (= #430 0::int)
+#431 := (<= -3::int 0::int)
+#410 := (or #431 #431)
+#745 := (not #410)
+#747 := (>= -3::int 0::int)
+#404 := (or #431 #747)
+#534 := (not #404)
+#741 := (or #534 #745)
+#415 := (ite #741 #427 #417)
+#416 := (= #174 0::int)
+#748 := (= -3::int 0::int)
+#749 := (or #748 #748)
+#750 := (ite #749 #416 #415)
+#592 := (or #590 #750)
+#594 := (iff #592 #591)
+#684 := (iff #591 #591)
+#688 := [rewrite]: #684
+#589 := (iff #750 #177)
+#683 := (ite false #416 #177)
+#586 := (iff #683 #177)
+#588 := [rewrite]: #586
+#686 := (iff #750 #683)
+#587 := (iff #415 #177)
+#576 := (ite false #177 #177)
+#537 := (iff #576 #177)
+#685 := [rewrite]: #537
+#577 := (iff #415 #576)
+#691 := (iff #417 #177)
+#715 := (+ -1::int #174)
+#705 := (= #715 0::int)
+#712 := (iff #705 #177)
+#707 := [rewrite]: #712
+#692 := (iff #417 #705)
+#689 := (= #426 #715)
+#432 := (+ #174 -1::int)
+#719 := (= #432 #715)
+#708 := [rewrite]: #719
+#535 := (= #426 #432)
+#704 := (= #425 -1::int)
+#725 := (* -1::int 1::int)
+#437 := (= #725 -1::int)
+#438 := [rewrite]: #437
+#703 := (= #425 #725)
+#696 := (= #424 1::int)
+#698 := (div 3::int 3::int)
+#701 := (= #698 1::int)
+#695 := [rewrite]: #701
+#699 := (= #424 #698)
+#555 := (= #340 3::int)
+#556 := [rewrite]: #555
+#700 := [monotonicity #556 #556]: #699
+#702 := [trans #700 #695]: #696
+#697 := [monotonicity #702]: #703
+#533 := [trans #697 #438]: #704
+#536 := [monotonicity #533]: #535
+#690 := [trans #536 #708]: #689
+#693 := [monotonicity #690]: #692
+#694 := [trans #693 #707]: #691
+#713 := (iff #427 #177)
+#706 := (iff #427 #705)
+#709 := (= #430 #715)
+#714 := (= #430 #432)
+#716 := (= #429 -1::int)
+#435 := (= #429 #725)
+#724 := (= #428 1::int)
+#721 := [rewrite]: #724
+#436 := [monotonicity #721]: #435
+#717 := [trans #436 #438]: #716
+#718 := [monotonicity #717]: #714
+#710 := [trans #718 #708]: #709
+#711 := [monotonicity #710]: #706
+#554 := [trans #711 #707]: #713
+#446 := (iff #741 false)
+#752 := (or false false)
+#407 := (iff #752 false)
+#743 := [rewrite]: #407
+#723 := (iff #741 #752)
+#346 := (iff #745 false)
+#1 := true
+#729 := (not true)
+#736 := (iff #729 false)
+#738 := [rewrite]: #736
+#451 := (iff #745 #729)
+#449 := (iff #410 true)
+#739 := (or true true)
+#726 := (iff #739 true)
+#727 := [rewrite]: #726
+#737 := (iff #410 #739)
+#387 := (iff #431 true)
+#728 := [rewrite]: #387
+#740 := [monotonicity #728 #728]: #737
+#450 := [trans #740 #727]: #449
+#452 := [monotonicity #450]: #451
+#722 := [trans #452 #738]: #346
+#378 := (iff #534 false)
+#735 := (iff #534 #729)
+#733 := (iff #404 true)
+#393 := (or true false)
+#731 := (iff #393 true)
+#732 := [rewrite]: #731
+#394 := (iff #404 #393)
+#391 := (iff #747 false)
+#392 := [rewrite]: #391
+#730 := [monotonicity #728 #392]: #394
+#734 := [trans #730 #732]: #733
+#373 := [monotonicity #734]: #735
+#379 := [trans #373 #738]: #378
+#445 := [monotonicity #379 #722]: #723
+#720 := [trans #445 #743]: #446
+#578 := [monotonicity #720 #554 #694]: #577
+#682 := [trans #578 #685]: #587
+#403 := (iff #749 false)
+#742 := (iff #749 #752)
+#751 := (iff #748 false)
+#746 := [rewrite]: #751
+#402 := [monotonicity #746 #746]: #742
+#744 := [trans #402 #743]: #403
+#571 := [monotonicity #744 #682]: #686
+#582 := [trans #571 #588]: #589
+#687 := [monotonicity #582]: #594
+#675 := [trans #687 #688]: #594
+#593 := [quant-inst]: #592
+#677 := [mp #593 #675]: #591
+[unit-resolution #677 #759 #195]: false
+unsat
+081d17418769caf8ceb0841229da368f6008d179 338 0
+#2 := false
+#43 := 1::int
+decl f3 :: (-> int int int)
+#176 := -3::int
+#173 := -5::int
+#179 := (f3 -5::int -3::int)
+#182 := (= #179 1::int)
+#195 := (not #182)
+#40 := 3::int
+#41 := (- 3::int)
+#38 := 5::int
+#39 := (- 5::int)
+#42 := (f3 #39 #41)
+#44 := (= #42 1::int)
+#45 := (not #44)
+#198 := (iff #45 #195)
+#185 := (= 1::int #179)
+#190 := (not #185)
+#196 := (iff #190 #195)
+#193 := (iff #185 #182)
+#194 := [rewrite]: #193
+#197 := [monotonicity #194]: #196
+#191 := (iff #45 #190)
+#188 := (iff #44 #185)
+#186 := (iff #182 #185)
+#187 := [rewrite]: #186
+#183 := (iff #44 #182)
+#180 := (= #42 #179)
+#177 := (= #41 -3::int)
+#178 := [rewrite]: #177
+#174 := (= #39 -5::int)
+#175 := [rewrite]: #174
+#181 := [monotonicity #175 #178]: #180
+#184 := [monotonicity #181]: #183
+#189 := [trans #184 #187]: #188
+#192 := [monotonicity #189]: #191
+#199 := [trans #192 #197]: #198
+#172 := [asserted]: #45
+#200 := [mp #172 #199]: #195
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#10 := (f3 #8 #9)
+#758 := (pattern #10)
+#11 := 0::int
+#69 := -1::int
+#73 := (* -1::int #9)
+#70 := (* -1::int #8)
+#76 := (div #70 #73)
+#261 := (* -1::int #76)
+#262 := (+ #10 #261)
+#263 := (= #262 0::int)
+#21 := (div #8 #9)
+#258 := (* -1::int #21)
+#259 := (+ #10 #258)
+#260 := (= #259 0::int)
+#95 := (<= #9 0::int)
+#91 := (<= #8 0::int)
+#216 := (or #91 #95)
+#217 := (not #216)
+#102 := (>= #8 0::int)
+#208 := (or #95 #102)
+#209 := (not #208)
+#223 := (or #209 #217)
+#264 := (ite #223 #260 #263)
+#257 := (= #10 0::int)
+#13 := (= #9 0::int)
+#12 := (= #8 0::int)
+#14 := (or #12 #13)
+#265 := (ite #14 #257 #264)
+#759 := (forall (vars (?v0 int) (?v1 int)) (:pat #758) #265)
+#268 := (forall (vars (?v0 int) (?v1 int)) #265)
+#762 := (iff #268 #759)
+#760 := (iff #265 #265)
+#761 := [refl]: #760
+#763 := [quant-intro #761]: #762
+#228 := (ite #223 #21 #76)
+#231 := (ite #14 0::int #228)
+#234 := (= #10 #231)
+#237 := (forall (vars (?v0 int) (?v1 int)) #234)
+#269 := (iff #237 #268)
+#266 := (iff #234 #265)
+#267 := [rewrite]: #266
+#270 := [quant-intro #267]: #269
+#103 := (not #102)
+#96 := (not #95)
+#106 := (and #96 #103)
+#92 := (not #91)
+#99 := (and #92 #96)
+#109 := (or #99 #106)
+#112 := (ite #109 #21 #76)
+#115 := (ite #14 0::int #112)
+#118 := (= #10 #115)
+#121 := (forall (vars (?v0 int) (?v1 int)) #118)
+#238 := (iff #121 #237)
+#235 := (iff #118 #234)
+#232 := (= #115 #231)
+#229 := (= #112 #228)
+#226 := (iff #109 #223)
+#220 := (or #217 #209)
+#224 := (iff #220 #223)
+#225 := [rewrite]: #224
+#221 := (iff #109 #220)
+#218 := (iff #106 #209)
+#219 := [rewrite]: #218
+#206 := (iff #99 #217)
+#207 := [rewrite]: #206
+#222 := [monotonicity #207 #219]: #221
+#227 := [trans #222 #225]: #226
+#230 := [monotonicity #227]: #229
+#233 := [monotonicity #230]: #232
+#236 := [monotonicity #233]: #235
+#239 := [quant-intro #236]: #238
+#212 := (~ #121 #121)
+#210 := (~ #118 #118)
+#211 := [refl]: #210
+#213 := [nnf-pos #211]: #212
+#23 := (- #9)
+#22 := (- #8)
+#24 := (div #22 #23)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#25 := (ite #20 #21 #24)
+#26 := (ite #14 0::int #25)
+#27 := (= #10 #26)
+#28 := (forall (vars (?v0 int) (?v1 int)) #27)
+#124 := (iff #28 #121)
+#63 := (and #16 #18)
+#66 := (or #17 #63)
+#79 := (ite #66 #21 #76)
+#82 := (ite #14 0::int #79)
+#85 := (= #10 #82)
+#88 := (forall (vars (?v0 int) (?v1 int)) #85)
+#122 := (iff #88 #121)
+#119 := (iff #85 #118)
+#116 := (= #82 #115)
+#113 := (= #79 #112)
+#110 := (iff #66 #109)
+#107 := (iff #63 #106)
+#104 := (iff #18 #103)
+#105 := [rewrite]: #104
+#97 := (iff #16 #96)
+#98 := [rewrite]: #97
+#108 := [monotonicity #98 #105]: #107
+#100 := (iff #17 #99)
+#93 := (iff #15 #92)
+#94 := [rewrite]: #93
+#101 := [monotonicity #94 #98]: #100
+#111 := [monotonicity #101 #108]: #110
+#114 := [monotonicity #111]: #113
+#117 := [monotonicity #114]: #116
+#120 := [monotonicity #117]: #119
+#123 := [quant-intro #120]: #122
+#89 := (iff #28 #88)
+#86 := (iff #27 #85)
+#83 := (= #26 #82)
+#80 := (= #25 #79)
+#77 := (= #24 #76)
+#74 := (= #23 #73)
+#75 := [rewrite]: #74
+#71 := (= #22 #70)
+#72 := [rewrite]: #71
+#78 := [monotonicity #72 #75]: #77
+#67 := (iff #20 #66)
+#64 := (iff #19 #63)
+#65 := [rewrite]: #64
+#68 := [monotonicity #65]: #67
+#81 := [monotonicity #68 #78]: #80
+#84 := [monotonicity #81]: #83
+#87 := [monotonicity #84]: #86
+#90 := [quant-intro #87]: #89
+#125 := [trans #90 #123]: #124
+#62 := [asserted]: #28
+#126 := [mp #62 #125]: #121
+#203 := [mp~ #126 #213]: #121
+#240 := [mp #203 #239]: #237
+#271 := [mp #240 #270]: #268
+#764 := [mp #271 #763]: #759
+#672 := (not #759)
+#679 := (or #672 #182)
+#345 := (* -1::int -3::int)
+#429 := (* -1::int -5::int)
+#430 := (div #429 #345)
+#431 := (* -1::int #430)
+#422 := (+ #179 #431)
+#433 := (= #422 0::int)
+#434 := (div -5::int -3::int)
+#435 := (* -1::int #434)
+#432 := (+ #179 #435)
+#436 := (= #432 0::int)
+#415 := (<= -3::int 0::int)
+#750 := (<= -5::int 0::int)
+#752 := (or #750 #415)
+#409 := (not #752)
+#539 := (>= -5::int 0::int)
+#746 := (or #415 #539)
+#420 := (not #746)
+#421 := (or #420 #409)
+#753 := (ite #421 #436 #433)
+#754 := (= #179 0::int)
+#755 := (= -3::int 0::int)
+#756 := (= -5::int 0::int)
+#751 := (or #756 #755)
+#757 := (ite #751 #754 #753)
+#663 := (or #672 #757)
+#666 := (iff #663 #679)
+#668 := (iff #679 #679)
+#669 := [rewrite]: #668
+#677 := (iff #757 #182)
+#685 := (ite false #754 #182)
+#675 := (iff #685 #182)
+#676 := [rewrite]: #675
+#681 := (iff #757 #685)
+#683 := (iff #753 #182)
+#721 := 2::int
+#706 := (= #179 2::int)
+#680 := (ite false #706 #182)
+#671 := (iff #680 #182)
+#673 := [rewrite]: #671
+#682 := (iff #753 #680)
+#689 := (iff #433 #182)
+#591 := (+ -1::int #179)
+#596 := (= #591 0::int)
+#599 := (iff #596 #182)
+#692 := [rewrite]: #599
+#597 := (iff #433 #596)
+#587 := (= #422 #591)
+#688 := (+ #179 -1::int)
+#593 := (= #688 #591)
+#594 := [rewrite]: #593
+#691 := (= #422 #688)
+#592 := (= #431 -1::int)
+#581 := (* -1::int 1::int)
+#542 := (= #581 -1::int)
+#690 := [rewrite]: #542
+#582 := (= #431 #581)
+#696 := (= #430 1::int)
+#541 := (div 5::int 3::int)
+#697 := (= #541 1::int)
+#698 := [rewrite]: #697
+#694 := (= #430 #541)
+#538 := (= #345 3::int)
+#540 := [rewrite]: #538
+#702 := (= #429 5::int)
+#709 := [rewrite]: #702
+#695 := [monotonicity #709 #540]: #694
+#699 := [trans #695 #698]: #696
+#583 := [monotonicity #699]: #582
+#687 := [trans #583 #690]: #592
+#576 := [monotonicity #687]: #691
+#595 := [trans #576 #594]: #587
+#598 := [monotonicity #595]: #597
+#693 := [trans #598 #692]: #689
+#707 := (iff #436 #706)
+#724 := -2::int
+#712 := (+ -2::int #179)
+#703 := (= #712 0::int)
+#700 := (iff #703 #706)
+#701 := [rewrite]: #700
+#704 := (iff #436 #703)
+#560 := (= #432 #712)
+#711 := (+ #179 -2::int)
+#718 := (= #711 #712)
+#559 := [rewrite]: #718
+#716 := (= #432 #711)
+#715 := (= #435 -2::int)
+#719 := (* -1::int 2::int)
+#713 := (= #719 -2::int)
+#714 := [rewrite]: #713
+#723 := (= #435 #719)
+#722 := (= #434 2::int)
+#437 := [rewrite]: #722
+#720 := [monotonicity #437]: #723
+#710 := [trans #720 #714]: #715
+#717 := [monotonicity #710]: #716
+#561 := [trans #717 #559]: #560
+#705 := [monotonicity #561]: #704
+#708 := [trans #705 #701]: #707
+#442 := (iff #421 false)
+#408 := (or false false)
+#733 := (iff #408 false)
+#396 := [rewrite]: #733
+#440 := (iff #421 #408)
+#726 := (iff #409 false)
+#1 := true
+#383 := (not true)
+#742 := (iff #383 false)
+#745 := [rewrite]: #742
+#725 := (iff #409 #383)
+#450 := (iff #752 true)
+#456 := (or true true)
+#727 := (iff #456 true)
+#728 := [rewrite]: #727
+#457 := (iff #752 #456)
+#399 := (iff #415 true)
+#735 := [rewrite]: #399
+#454 := (iff #750 true)
+#455 := [rewrite]: #454
+#351 := [monotonicity #455 #735]: #457
+#451 := [trans #351 #728]: #450
+#729 := [monotonicity #451]: #725
+#730 := [trans #729 #745]: #726
+#731 := (iff #420 false)
+#384 := (iff #420 #383)
+#741 := (iff #746 true)
+#738 := (or true false)
+#740 := (iff #738 true)
+#378 := [rewrite]: #740
+#739 := (iff #746 #738)
+#736 := (iff #539 false)
+#737 := [rewrite]: #736
+#734 := [monotonicity #735 #737]: #739
+#743 := [trans #734 #378]: #741
+#744 := [monotonicity #743]: #384
+#732 := [trans #744 #745]: #731
+#441 := [monotonicity #732 #730]: #440
+#443 := [trans #441 #396]: #442
+#674 := [monotonicity #443 #708 #693]: #682
+#684 := [trans #674 #673]: #683
+#397 := (iff #751 false)
+#749 := (iff #751 #408)
+#412 := (iff #755 false)
+#748 := [rewrite]: #412
+#747 := (iff #756 false)
+#407 := [rewrite]: #747
+#392 := [monotonicity #407 #748]: #749
+#398 := [trans #392 #396]: #397
+#686 := [monotonicity #398 #684]: #681
+#678 := [trans #686 #676]: #677
+#667 := [monotonicity #678]: #666
+#665 := [trans #667 #669]: #666
+#664 := [quant-inst]: #663
+#670 := [mp #664 #665]: #679
+[unit-resolution #670 #764 #200]: false
+unsat
+1a424f78adbf3f5b91865b15057a1edb6fd9f0c7 268 0
+#2 := false
+#11 := 0::int
+decl f4 :: (-> int int int)
+#38 := (f4 0::int 0::int)
+#39 := (= #38 0::int)
+#40 := (not #39)
+#167 := [asserted]: #40
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#29 := (f4 #8 #9)
+#743 := (pattern #29)
+#64 := -1::int
+#68 := (* -1::int #9)
+#65 := (* -1::int #8)
+#123 := (mod #65 #68)
+#254 := (+ #29 #123)
+#255 := (= #254 0::int)
+#30 := (mod #8 #9)
+#251 := (* -1::int #30)
+#252 := (+ #29 #251)
+#253 := (= #252 0::int)
+#90 := (<= #9 0::int)
+#86 := (<= #8 0::int)
+#193 := (or #86 #90)
+#194 := (not #193)
+#97 := (>= #8 0::int)
+#185 := (or #90 #97)
+#186 := (not #185)
+#200 := (or #186 #194)
+#256 := (ite #200 #253 #255)
+#250 := (= #29 0::int)
+#12 := (= #8 0::int)
+#257 := (ite #12 #250 #256)
+#249 := (= #8 #29)
+#13 := (= #9 0::int)
+#258 := (ite #13 #249 #257)
+#744 := (forall (vars (?v0 int) (?v1 int)) (:pat #743) #258)
+#261 := (forall (vars (?v0 int) (?v1 int)) #258)
+#747 := (iff #261 #744)
+#745 := (iff #258 #258)
+#746 := [refl]: #745
+#748 := [quant-intro #746]: #747
+#129 := (* -1::int #123)
+#218 := (ite #200 #30 #129)
+#221 := (ite #12 0::int #218)
+#224 := (ite #13 #8 #221)
+#227 := (= #29 #224)
+#230 := (forall (vars (?v0 int) (?v1 int)) #227)
+#262 := (iff #230 #261)
+#259 := (iff #227 #258)
+#260 := [rewrite]: #259
+#263 := [quant-intro #260]: #262
+#98 := (not #97)
+#91 := (not #90)
+#101 := (and #91 #98)
+#87 := (not #86)
+#94 := (and #87 #91)
+#104 := (or #94 #101)
+#149 := (ite #104 #30 #129)
+#152 := (ite #12 0::int #149)
+#155 := (ite #13 #8 #152)
+#158 := (= #29 #155)
+#161 := (forall (vars (?v0 int) (?v1 int)) #158)
+#231 := (iff #161 #230)
+#228 := (iff #158 #227)
+#225 := (= #155 #224)
+#222 := (= #152 #221)
+#219 := (= #149 #218)
+#203 := (iff #104 #200)
+#197 := (or #194 #186)
+#201 := (iff #197 #200)
+#202 := [rewrite]: #201
+#198 := (iff #104 #197)
+#195 := (iff #101 #186)
+#196 := [rewrite]: #195
+#183 := (iff #94 #194)
+#184 := [rewrite]: #183
+#199 := [monotonicity #184 #196]: #198
+#204 := [trans #199 #202]: #203
+#220 := [monotonicity #204]: #219
+#223 := [monotonicity #220]: #222
+#226 := [monotonicity #223]: #225
+#229 := [monotonicity #226]: #228
+#232 := [quant-intro #229]: #231
+#181 := (~ #161 #161)
+#178 := (~ #158 #158)
+#191 := [refl]: #178
+#182 := [nnf-pos #191]: #181
+#23 := (- #9)
+#22 := (- #8)
+#31 := (mod #22 #23)
+#32 := (- #31)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#33 := (ite #20 #30 #32)
+#34 := (ite #12 0::int #33)
+#35 := (ite #13 #8 #34)
+#36 := (= #29 #35)
+#37 := (forall (vars (?v0 int) (?v1 int)) #36)
+#164 := (iff #37 #161)
+#58 := (and #16 #18)
+#61 := (or #17 #58)
+#134 := (ite #61 #30 #129)
+#137 := (ite #12 0::int #134)
+#140 := (ite #13 #8 #137)
+#143 := (= #29 #140)
+#146 := (forall (vars (?v0 int) (?v1 int)) #143)
+#162 := (iff #146 #161)
+#159 := (iff #143 #158)
+#156 := (= #140 #155)
+#153 := (= #137 #152)
+#150 := (= #134 #149)
+#105 := (iff #61 #104)
+#102 := (iff #58 #101)
+#99 := (iff #18 #98)
+#100 := [rewrite]: #99
+#92 := (iff #16 #91)
+#93 := [rewrite]: #92
+#103 := [monotonicity #93 #100]: #102
+#95 := (iff #17 #94)
+#88 := (iff #15 #87)
+#89 := [rewrite]: #88
+#96 := [monotonicity #89 #93]: #95
+#106 := [monotonicity #96 #103]: #105
+#151 := [monotonicity #106]: #150
+#154 := [monotonicity #151]: #153
+#157 := [monotonicity #154]: #156
+#160 := [monotonicity #157]: #159
+#163 := [quant-intro #160]: #162
+#147 := (iff #37 #146)
+#144 := (iff #36 #143)
+#141 := (= #35 #140)
+#138 := (= #34 #137)
+#135 := (= #33 #134)
+#132 := (= #32 #129)
+#126 := (- #123)
+#130 := (= #126 #129)
+#131 := [rewrite]: #130
+#127 := (= #32 #126)
+#124 := (= #31 #123)
+#69 := (= #23 #68)
+#70 := [rewrite]: #69
+#66 := (= #22 #65)
+#67 := [rewrite]: #66
+#125 := [monotonicity #67 #70]: #124
+#128 := [monotonicity #125]: #127
+#133 := [trans #128 #131]: #132
+#62 := (iff #20 #61)
+#59 := (iff #19 #58)
+#60 := [rewrite]: #59
+#63 := [monotonicity #60]: #62
+#136 := [monotonicity #63 #133]: #135
+#139 := [monotonicity #136]: #138
+#142 := [monotonicity #139]: #141
+#145 := [monotonicity #142]: #144
+#148 := [quant-intro #145]: #147
+#165 := [trans #148 #163]: #164
+#122 := [asserted]: #37
+#166 := [mp #122 #165]: #161
+#192 := [mp~ #166 #182]: #161
+#233 := [mp #192 #232]: #230
+#264 := [mp #233 #263]: #261
+#749 := [mp #264 #748]: #744
+#690 := (not #744)
+#696 := (or #690 #39)
+#322 := (* -1::int 0::int)
+#407 := (mod #322 #322)
+#408 := (+ #38 #407)
+#409 := (= #408 0::int)
+#400 := (mod 0::int 0::int)
+#411 := (* -1::int #400)
+#412 := (+ #38 #411)
+#340 := (= #412 0::int)
+#413 := (<= 0::int 0::int)
+#410 := (or #413 #413)
+#414 := (not #410)
+#393 := (>= 0::int 0::int)
+#728 := (or #413 #393)
+#730 := (not #728)
+#387 := (or #730 #414)
+#517 := (ite #387 #340 #409)
+#724 := (= 0::int 0::int)
+#398 := (ite #724 #39 #517)
+#168 := (= 0::int #38)
+#399 := (ite #724 #168 #398)
+#537 := (or #690 #399)
+#539 := (iff #537 #696)
+#682 := (iff #696 #696)
+#683 := [rewrite]: #682
+#694 := (iff #399 #39)
+#1 := true
+#691 := (ite true #39 #39)
+#688 := (iff #691 #39)
+#689 := [rewrite]: #688
+#692 := (iff #399 #691)
+#698 := (iff #398 #39)
+#328 := (+ #38 #400)
+#428 := (= #328 0::int)
+#699 := (ite true #39 #428)
+#697 := (iff #699 #39)
+#701 := [rewrite]: #697
+#700 := (iff #398 #699)
+#420 := (iff #517 #428)
+#707 := (ite false #340 #428)
+#418 := (iff #707 #428)
+#419 := [rewrite]: #418
+#704 := (iff #517 #707)
+#429 := (iff #409 #428)
+#705 := (= #408 #328)
+#434 := (= #407 #400)
+#432 := (= #322 0::int)
+#433 := [rewrite]: #432
+#435 := [monotonicity #433 #433]: #434
+#706 := [monotonicity #435]: #705
+#703 := [monotonicity #706]: #429
+#709 := (iff #387 false)
+#361 := (or false false)
+#720 := (iff #361 false)
+#723 := [rewrite]: #720
+#362 := (iff #387 #361)
+#719 := (iff #414 false)
+#711 := (not true)
+#376 := (iff #711 false)
+#377 := [rewrite]: #376
+#718 := (iff #414 #711)
+#717 := (iff #410 true)
+#725 := (or true true)
+#726 := (iff #725 true)
+#386 := [rewrite]: #726
+#715 := (iff #410 #725)
+#733 := (iff #413 true)
+#734 := [rewrite]: #733
+#716 := [monotonicity #734 #734]: #715
+#712 := [trans #716 #386]: #717
+#356 := [monotonicity #712]: #718
+#721 := [trans #356 #377]: #719
+#713 := (iff #730 false)
+#374 := (iff #730 #711)
+#727 := (iff #728 true)
+#385 := (iff #728 #725)
+#729 := (iff #393 true)
+#735 := [rewrite]: #729
+#390 := [monotonicity #734 #735]: #385
+#370 := [trans #390 #386]: #727
+#375 := [monotonicity #370]: #374
+#714 := [trans #375 #377]: #713
+#722 := [monotonicity #714 #721]: #362
+#710 := [trans #722 #723]: #709
+#708 := [monotonicity #710 #703]: #704
+#421 := [trans #708 #419]: #420
+#731 := (iff #724 true)
+#732 := [rewrite]: #731
+#415 := [monotonicity #732 #421]: #700
+#702 := [trans #415 #701]: #698
+#174 := (iff #168 #39)
+#175 := [rewrite]: #174
+#693 := [monotonicity #732 #175 #702]: #692
+#695 := [trans #693 #689]: #694
+#681 := [monotonicity #695]: #539
+#684 := [trans #681 #683]: #539
+#538 := [quant-inst]: #537
+#678 := [mp #538 #684]: #696
+[unit-resolution #678 #749 #167]: false
+unsat
+d05667c302799334d71c6767e44df657452bd738 276 0
+#2 := false
+decl f4 :: (-> int int int)
+#11 := 0::int
+decl f5 :: int
+#38 := f5
+#39 := (f4 f5 0::int)
+#169 := (= f5 #39)
+#172 := (not #169)
+#40 := (= #39 f5)
+#41 := (not #40)
+#173 := (iff #41 #172)
+#170 := (iff #40 #169)
+#171 := [rewrite]: #170
+#174 := [monotonicity #171]: #173
+#168 := [asserted]: #41
+#177 := [mp #168 #174]: #172
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#29 := (f4 #8 #9)
+#741 := (pattern #29)
+#65 := -1::int
+#69 := (* -1::int #9)
+#66 := (* -1::int #8)
+#124 := (mod #66 #69)
+#252 := (+ #29 #124)
+#253 := (= #252 0::int)
+#30 := (mod #8 #9)
+#249 := (* -1::int #30)
+#250 := (+ #29 #249)
+#251 := (= #250 0::int)
+#91 := (<= #9 0::int)
+#87 := (<= #8 0::int)
+#191 := (or #87 #91)
+#192 := (not #191)
+#98 := (>= #8 0::int)
+#183 := (or #91 #98)
+#184 := (not #183)
+#198 := (or #184 #192)
+#254 := (ite #198 #251 #253)
+#248 := (= #29 0::int)
+#12 := (= #8 0::int)
+#255 := (ite #12 #248 #254)
+#247 := (= #8 #29)
+#13 := (= #9 0::int)
+#256 := (ite #13 #247 #255)
+#742 := (forall (vars (?v0 int) (?v1 int)) (:pat #741) #256)
+#259 := (forall (vars (?v0 int) (?v1 int)) #256)
+#745 := (iff #259 #742)
+#743 := (iff #256 #256)
+#744 := [refl]: #743
+#746 := [quant-intro #744]: #745
+#130 := (* -1::int #124)
+#216 := (ite #198 #30 #130)
+#219 := (ite #12 0::int #216)
+#222 := (ite #13 #8 #219)
+#225 := (= #29 #222)
+#228 := (forall (vars (?v0 int) (?v1 int)) #225)
+#260 := (iff #228 #259)
+#257 := (iff #225 #256)
+#258 := [rewrite]: #257
+#261 := [quant-intro #258]: #260
+#99 := (not #98)
+#92 := (not #91)
+#102 := (and #92 #99)
+#88 := (not #87)
+#95 := (and #88 #92)
+#105 := (or #95 #102)
+#150 := (ite #105 #30 #130)
+#153 := (ite #12 0::int #150)
+#156 := (ite #13 #8 #153)
+#159 := (= #29 #156)
+#162 := (forall (vars (?v0 int) (?v1 int)) #159)
+#229 := (iff #162 #228)
+#226 := (iff #159 #225)
+#223 := (= #156 #222)
+#220 := (= #153 #219)
+#217 := (= #150 #216)
+#201 := (iff #105 #198)
+#195 := (or #192 #184)
+#199 := (iff #195 #198)
+#200 := [rewrite]: #199
+#196 := (iff #105 #195)
+#193 := (iff #102 #184)
+#194 := [rewrite]: #193
+#181 := (iff #95 #192)
+#182 := [rewrite]: #181
+#197 := [monotonicity #182 #194]: #196
+#202 := [trans #197 #200]: #201
+#218 := [monotonicity #202]: #217
+#221 := [monotonicity #218]: #220
+#224 := [monotonicity #221]: #223
+#227 := [monotonicity #224]: #226
+#230 := [quant-intro #227]: #229
+#179 := (~ #162 #162)
+#175 := (~ #159 #159)
+#189 := [refl]: #175
+#180 := [nnf-pos #189]: #179
+#23 := (- #9)
+#22 := (- #8)
+#31 := (mod #22 #23)
+#32 := (- #31)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#33 := (ite #20 #30 #32)
+#34 := (ite #12 0::int #33)
+#35 := (ite #13 #8 #34)
+#36 := (= #29 #35)
+#37 := (forall (vars (?v0 int) (?v1 int)) #36)
+#165 := (iff #37 #162)
+#59 := (and #16 #18)
+#62 := (or #17 #59)
+#135 := (ite #62 #30 #130)
+#138 := (ite #12 0::int #135)
+#141 := (ite #13 #8 #138)
+#144 := (= #29 #141)
+#147 := (forall (vars (?v0 int) (?v1 int)) #144)
+#163 := (iff #147 #162)
+#160 := (iff #144 #159)
+#157 := (= #141 #156)
+#154 := (= #138 #153)
+#151 := (= #135 #150)
+#106 := (iff #62 #105)
+#103 := (iff #59 #102)
+#100 := (iff #18 #99)
+#101 := [rewrite]: #100
+#93 := (iff #16 #92)
+#94 := [rewrite]: #93
+#104 := [monotonicity #94 #101]: #103
+#96 := (iff #17 #95)
+#89 := (iff #15 #88)
+#90 := [rewrite]: #89
+#97 := [monotonicity #90 #94]: #96
+#107 := [monotonicity #97 #104]: #106
+#152 := [monotonicity #107]: #151
+#155 := [monotonicity #152]: #154
+#158 := [monotonicity #155]: #157
+#161 := [monotonicity #158]: #160
+#164 := [quant-intro #161]: #163
+#148 := (iff #37 #147)
+#145 := (iff #36 #144)
+#142 := (= #35 #141)
+#139 := (= #34 #138)
+#136 := (= #33 #135)
+#133 := (= #32 #130)
+#127 := (- #124)
+#131 := (= #127 #130)
+#132 := [rewrite]: #131
+#128 := (= #32 #127)
+#125 := (= #31 #124)
+#70 := (= #23 #69)
+#71 := [rewrite]: #70
+#67 := (= #22 #66)
+#68 := [rewrite]: #67
+#126 := [monotonicity #68 #71]: #125
+#129 := [monotonicity #126]: #128
+#134 := [trans #129 #132]: #133
+#63 := (iff #20 #62)
+#60 := (iff #19 #59)
+#61 := [rewrite]: #60
+#64 := [monotonicity #61]: #63
+#137 := [monotonicity #64 #134]: #136
+#140 := [monotonicity #137]: #139
+#143 := [monotonicity #140]: #142
+#146 := [monotonicity #143]: #145
+#149 := [quant-intro #146]: #148
+#166 := [trans #149 #164]: #165
+#123 := [asserted]: #37
+#167 := [mp #123 #166]: #162
+#190 := [mp~ #167 #180]: #162
+#231 := [mp #190 #230]: #228
+#262 := [mp #231 #261]: #259
+#747 := [mp #262 #746]: #742
+#535 := (not #742)
+#536 := (or #535 #169)
+#320 := (* -1::int 0::int)
+#405 := (* -1::int f5)
+#406 := (mod #405 #320)
+#407 := (+ #39 #406)
+#398 := (= #407 0::int)
+#409 := (mod f5 0::int)
+#410 := (* -1::int #409)
+#338 := (+ #39 #410)
+#411 := (= #338 0::int)
+#408 := (<= 0::int 0::int)
+#412 := (<= f5 0::int)
+#391 := (or #412 #408)
+#726 := (not #391)
+#728 := (>= f5 0::int)
+#385 := (or #408 #728)
+#515 := (not #385)
+#722 := (or #515 #726)
+#396 := (ite #722 #411 #398)
+#397 := (= #39 0::int)
+#729 := (= f5 0::int)
+#730 := (ite #729 #397 #396)
+#731 := (= 0::int 0::int)
+#732 := (ite #731 #169 #730)
+#537 := (or #535 #732)
+#680 := (iff #537 #536)
+#682 := (iff #536 #536)
+#676 := [rewrite]: #682
+#688 := (iff #732 #169)
+#426 := (mod #405 0::int)
+#705 := (+ #39 #426)
+#416 := (= #705 0::int)
+#700 := (ite #729 #397 #416)
+#1 := true
+#691 := (ite true #169 #700)
+#692 := (iff #691 #169)
+#693 := [rewrite]: #692
+#686 := (iff #732 #691)
+#689 := (iff #730 #700)
+#699 := (iff #396 #416)
+#419 := (ite false #411 #416)
+#413 := (iff #419 #416)
+#695 := [rewrite]: #413
+#697 := (iff #396 #419)
+#417 := (iff #398 #416)
+#702 := (= #407 #705)
+#427 := (= #406 #426)
+#703 := (= #320 0::int)
+#704 := [rewrite]: #703
+#701 := [monotonicity #704]: #427
+#706 := [monotonicity #701]: #702
+#418 := [monotonicity #706]: #417
+#433 := (iff #722 false)
+#707 := (or false false)
+#431 := (iff #707 false)
+#432 := [rewrite]: #431
+#708 := (iff #722 #707)
+#718 := (iff #726 false)
+#373 := (not true)
+#711 := (iff #373 false)
+#712 := [rewrite]: #711
+#360 := (iff #726 #373)
+#719 := (iff #391 true)
+#715 := (or #412 true)
+#354 := (iff #715 true)
+#717 := [rewrite]: #354
+#710 := (iff #391 #715)
+#723 := (iff #408 true)
+#383 := [rewrite]: #723
+#716 := [monotonicity #383]: #710
+#359 := [trans #716 #717]: #719
+#720 := [monotonicity #359]: #360
+#721 := [trans #720 #712]: #718
+#713 := (iff #515 false)
+#374 := (iff #515 #373)
+#709 := (iff #385 true)
+#388 := (or true #728)
+#725 := (iff #388 true)
+#368 := [rewrite]: #725
+#724 := (iff #385 #388)
+#384 := [monotonicity #383]: #724
+#372 := [trans #384 #368]: #709
+#375 := [monotonicity #372]: #374
+#714 := [trans #375 #712]: #713
+#430 := [monotonicity #714 #721]: #708
+#326 := [trans #430 #432]: #433
+#698 := [monotonicity #326 #418]: #697
+#696 := [trans #698 #695]: #699
+#690 := [monotonicity #696]: #689
+#727 := (iff #731 true)
+#733 := [rewrite]: #727
+#687 := [monotonicity #733 #690]: #686
+#694 := [trans #687 #693]: #688
+#681 := [monotonicity #694]: #680
+#677 := [trans #681 #676]: #680
+#679 := [quant-inst]: #537
+#683 := [mp #679 #677]: #536
+[unit-resolution #683 #747 #177]: false
+unsat
+07a66f14187da887bd1f0f4bb16d6c9a7023aeaf 298 0
+#2 := false
+#11 := 0::int
+decl f4 :: (-> int int int)
+#38 := 1::int
+#39 := (f4 0::int 1::int)
+#40 := (= #39 0::int)
+#41 := (not #40)
+#168 := [asserted]: #41
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#29 := (f4 #8 #9)
+#743 := (pattern #29)
+#65 := -1::int
+#69 := (* -1::int #9)
+#66 := (* -1::int #8)
+#124 := (mod #66 #69)
+#255 := (+ #29 #124)
+#256 := (= #255 0::int)
+#30 := (mod #8 #9)
+#252 := (* -1::int #30)
+#253 := (+ #29 #252)
+#254 := (= #253 0::int)
+#91 := (<= #9 0::int)
+#87 := (<= #8 0::int)
+#194 := (or #87 #91)
+#195 := (not #194)
+#98 := (>= #8 0::int)
+#186 := (or #91 #98)
+#187 := (not #186)
+#201 := (or #187 #195)
+#257 := (ite #201 #254 #256)
+#251 := (= #29 0::int)
+#12 := (= #8 0::int)
+#258 := (ite #12 #251 #257)
+#250 := (= #8 #29)
+#13 := (= #9 0::int)
+#259 := (ite #13 #250 #258)
+#744 := (forall (vars (?v0 int) (?v1 int)) (:pat #743) #259)
+#262 := (forall (vars (?v0 int) (?v1 int)) #259)
+#747 := (iff #262 #744)
+#745 := (iff #259 #259)
+#746 := [refl]: #745
+#748 := [quant-intro #746]: #747
+#130 := (* -1::int #124)
+#219 := (ite #201 #30 #130)
+#222 := (ite #12 0::int #219)
+#225 := (ite #13 #8 #222)
+#228 := (= #29 #225)
+#231 := (forall (vars (?v0 int) (?v1 int)) #228)
+#263 := (iff #231 #262)
+#260 := (iff #228 #259)
+#261 := [rewrite]: #260
+#264 := [quant-intro #261]: #263
+#99 := (not #98)
+#92 := (not #91)
+#102 := (and #92 #99)
+#88 := (not #87)
+#95 := (and #88 #92)
+#105 := (or #95 #102)
+#150 := (ite #105 #30 #130)
+#153 := (ite #12 0::int #150)
+#156 := (ite #13 #8 #153)
+#159 := (= #29 #156)
+#162 := (forall (vars (?v0 int) (?v1 int)) #159)
+#232 := (iff #162 #231)
+#229 := (iff #159 #228)
+#226 := (= #156 #225)
+#223 := (= #153 #222)
+#220 := (= #150 #219)
+#204 := (iff #105 #201)
+#198 := (or #195 #187)
+#202 := (iff #198 #201)
+#203 := [rewrite]: #202
+#199 := (iff #105 #198)
+#196 := (iff #102 #187)
+#197 := [rewrite]: #196
+#184 := (iff #95 #195)
+#185 := [rewrite]: #184
+#200 := [monotonicity #185 #197]: #199
+#205 := [trans #200 #203]: #204
+#221 := [monotonicity #205]: #220
+#224 := [monotonicity #221]: #223
+#227 := [monotonicity #224]: #226
+#230 := [monotonicity #227]: #229
+#233 := [quant-intro #230]: #232
+#182 := (~ #162 #162)
+#179 := (~ #159 #159)
+#192 := [refl]: #179
+#183 := [nnf-pos #192]: #182
+#23 := (- #9)
+#22 := (- #8)
+#31 := (mod #22 #23)
+#32 := (- #31)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#33 := (ite #20 #30 #32)
+#34 := (ite #12 0::int #33)
+#35 := (ite #13 #8 #34)
+#36 := (= #29 #35)
+#37 := (forall (vars (?v0 int) (?v1 int)) #36)
+#165 := (iff #37 #162)
+#59 := (and #16 #18)
+#62 := (or #17 #59)
+#135 := (ite #62 #30 #130)
+#138 := (ite #12 0::int #135)
+#141 := (ite #13 #8 #138)
+#144 := (= #29 #141)
+#147 := (forall (vars (?v0 int) (?v1 int)) #144)
+#163 := (iff #147 #162)
+#160 := (iff #144 #159)
+#157 := (= #141 #156)
+#154 := (= #138 #153)
+#151 := (= #135 #150)
+#106 := (iff #62 #105)
+#103 := (iff #59 #102)
+#100 := (iff #18 #99)
+#101 := [rewrite]: #100
+#93 := (iff #16 #92)
+#94 := [rewrite]: #93
+#104 := [monotonicity #94 #101]: #103
+#96 := (iff #17 #95)
+#89 := (iff #15 #88)
+#90 := [rewrite]: #89
+#97 := [monotonicity #90 #94]: #96
+#107 := [monotonicity #97 #104]: #106
+#152 := [monotonicity #107]: #151
+#155 := [monotonicity #152]: #154
+#158 := [monotonicity #155]: #157
+#161 := [monotonicity #158]: #160
+#164 := [quant-intro #161]: #163
+#148 := (iff #37 #147)
+#145 := (iff #36 #144)
+#142 := (= #35 #141)
+#139 := (= #34 #138)
+#136 := (= #33 #135)
+#133 := (= #32 #130)
+#127 := (- #124)
+#131 := (= #127 #130)
+#132 := [rewrite]: #131
+#128 := (= #32 #127)
+#125 := (= #31 #124)
+#70 := (= #23 #69)
+#71 := [rewrite]: #70
+#67 := (= #22 #66)
+#68 := [rewrite]: #67
+#126 := [monotonicity #68 #71]: #125
+#129 := [monotonicity #126]: #128
+#134 := [trans #129 #132]: #133
+#63 := (iff #20 #62)
+#60 := (iff #19 #59)
+#61 := [rewrite]: #60
+#64 := [monotonicity #61]: #63
+#137 := [monotonicity #64 #134]: #136
+#140 := [monotonicity #137]: #139
+#143 := [monotonicity #140]: #142
+#146 := [monotonicity #143]: #145
+#149 := [quant-intro #146]: #148
+#166 := [trans #149 #164]: #165
+#123 := [asserted]: #37
+#167 := [mp #123 #166]: #162
+#193 := [mp~ #167 #183]: #162
+#234 := [mp #193 #233]: #231
+#265 := [mp #234 #264]: #262
+#749 := [mp #265 #748]: #744
+#665 := (not #744)
+#666 := (or #665 #40)
+#323 := (* -1::int 1::int)
+#407 := (* -1::int 0::int)
+#408 := (mod #407 #323)
+#409 := (+ #39 #408)
+#400 := (= #409 0::int)
+#411 := (mod 0::int 1::int)
+#412 := (* -1::int #411)
+#413 := (+ #39 #412)
+#410 := (= #413 0::int)
+#414 := (<= 1::int 0::int)
+#393 := (<= 0::int 0::int)
+#728 := (or #393 #414)
+#730 := (not #728)
+#387 := (>= 0::int 0::int)
+#517 := (or #414 #387)
+#724 := (not #517)
+#398 := (or #724 #730)
+#399 := (ite #398 #410 #400)
+#731 := (= 0::int 0::int)
+#732 := (ite #731 #40 #399)
+#169 := (= 0::int #39)
+#733 := (= 1::int 0::int)
+#734 := (ite #733 #169 #732)
+#669 := (or #665 #734)
+#569 := (iff #669 #666)
+#572 := (iff #666 #666)
+#565 := [rewrite]: #572
+#668 := (iff #734 #40)
+#686 := (ite false #40 #40)
+#516 := (iff #686 #40)
+#518 := [rewrite]: #516
+#561 := (iff #734 #686)
+#559 := (iff #732 #40)
+#1 := true
+#673 := (ite true #40 #40)
+#674 := (iff #673 #40)
+#677 := [rewrite]: #674
+#675 := (iff #732 #673)
+#519 := (iff #399 #40)
+#680 := (iff #399 #686)
+#679 := (iff #400 #40)
+#684 := (= #409 #39)
+#415 := (+ #39 0::int)
+#698 := (= #415 #39)
+#702 := [rewrite]: #698
+#682 := (= #409 #415)
+#539 := (= #408 0::int)
+#695 := (mod 0::int -1::int)
+#537 := (= #695 0::int)
+#538 := [rewrite]: #537
+#690 := (= #408 #695)
+#689 := (= #323 -1::int)
+#694 := [rewrite]: #689
+#420 := (= #407 0::int)
+#421 := [rewrite]: #420
+#696 := [monotonicity #421 #694]: #690
+#681 := [trans #696 #538]: #539
+#683 := [monotonicity #681]: #682
+#678 := [trans #683 #702]: #684
+#685 := [monotonicity #678]: #679
+#693 := (iff #410 #40)
+#691 := (= #413 #39)
+#697 := (= #413 #415)
+#699 := (= #412 0::int)
+#418 := (= #412 #407)
+#704 := (= #411 0::int)
+#708 := [rewrite]: #704
+#419 := [monotonicity #708]: #418
+#700 := [trans #419 #421]: #699
+#701 := [monotonicity #700]: #697
+#692 := [trans #701 #702]: #691
+#688 := [monotonicity #692]: #693
+#703 := (iff #398 false)
+#329 := (or false false)
+#428 := (iff #329 false)
+#429 := [rewrite]: #428
+#705 := (iff #398 #329)
+#434 := (iff #730 false)
+#714 := (not true)
+#717 := (iff #714 false)
+#712 := [rewrite]: #717
+#432 := (iff #730 #714)
+#709 := (iff #728 true)
+#361 := (or true false)
+#720 := (iff #361 true)
+#723 := [rewrite]: #720
+#362 := (iff #728 #361)
+#390 := (iff #414 false)
+#726 := [rewrite]: #390
+#719 := (iff #393 true)
+#721 := [rewrite]: #719
+#722 := [monotonicity #721 #726]: #362
+#710 := [trans #722 #723]: #709
+#433 := [monotonicity #710]: #432
+#435 := [trans #433 #712]: #434
+#718 := (iff #724 false)
+#715 := (iff #724 #714)
+#377 := (iff #517 true)
+#370 := (or false true)
+#375 := (iff #370 true)
+#376 := [rewrite]: #375
+#711 := (iff #517 #370)
+#386 := (iff #387 true)
+#727 := [rewrite]: #386
+#374 := [monotonicity #726 #727]: #711
+#713 := [trans #374 #376]: #377
+#716 := [monotonicity #713]: #715
+#356 := [trans #716 #712]: #718
+#706 := [monotonicity #356 #435]: #705
+#707 := [trans #706 #429]: #703
+#687 := [monotonicity #707 #688 #685]: #680
+#672 := [trans #687 #518]: #519
+#725 := (iff #731 true)
+#385 := [rewrite]: #725
+#676 := [monotonicity #385 #672]: #675
+#560 := [trans #676 #677]: #559
+#175 := (iff #169 #40)
+#176 := [rewrite]: #175
+#729 := (iff #733 false)
+#735 := [rewrite]: #729
+#520 := [monotonicity #735 #176 #560]: #561
+#570 := [trans #520 #518]: #668
+#571 := [monotonicity #570]: #569
+#573 := [trans #571 #565]: #569
+#554 := [quant-inst]: #669
+#574 := [mp #554 #573]: #666
+[unit-resolution #574 #749 #168]: false
+unsat
+ed1abab981b05d680a1c92a497bf3305ff9e1f16 296 0
+#2 := false
+#11 := 0::int
+decl f4 :: (-> int int int)
+#38 := 1::int
+#39 := (f4 1::int 1::int)
+#40 := (= #39 0::int)
+#41 := (not #40)
+#168 := [asserted]: #41
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#29 := (f4 #8 #9)
+#743 := (pattern #29)
+#65 := -1::int
+#69 := (* -1::int #9)
+#66 := (* -1::int #8)
+#124 := (mod #66 #69)
+#255 := (+ #29 #124)
+#256 := (= #255 0::int)
+#30 := (mod #8 #9)
+#252 := (* -1::int #30)
+#253 := (+ #29 #252)
+#254 := (= #253 0::int)
+#91 := (<= #9 0::int)
+#87 := (<= #8 0::int)
+#194 := (or #87 #91)
+#195 := (not #194)
+#98 := (>= #8 0::int)
+#186 := (or #91 #98)
+#187 := (not #186)
+#201 := (or #187 #195)
+#257 := (ite #201 #254 #256)
+#251 := (= #29 0::int)
+#12 := (= #8 0::int)
+#258 := (ite #12 #251 #257)
+#250 := (= #8 #29)
+#13 := (= #9 0::int)
+#259 := (ite #13 #250 #258)
+#744 := (forall (vars (?v0 int) (?v1 int)) (:pat #743) #259)
+#262 := (forall (vars (?v0 int) (?v1 int)) #259)
+#747 := (iff #262 #744)
+#745 := (iff #259 #259)
+#746 := [refl]: #745
+#748 := [quant-intro #746]: #747
+#130 := (* -1::int #124)
+#219 := (ite #201 #30 #130)
+#222 := (ite #12 0::int #219)
+#225 := (ite #13 #8 #222)
+#228 := (= #29 #225)
+#231 := (forall (vars (?v0 int) (?v1 int)) #228)
+#263 := (iff #231 #262)
+#260 := (iff #228 #259)
+#261 := [rewrite]: #260
+#264 := [quant-intro #261]: #263
+#99 := (not #98)
+#92 := (not #91)
+#102 := (and #92 #99)
+#88 := (not #87)
+#95 := (and #88 #92)
+#105 := (or #95 #102)
+#150 := (ite #105 #30 #130)
+#153 := (ite #12 0::int #150)
+#156 := (ite #13 #8 #153)
+#159 := (= #29 #156)
+#162 := (forall (vars (?v0 int) (?v1 int)) #159)
+#232 := (iff #162 #231)
+#229 := (iff #159 #228)
+#226 := (= #156 #225)
+#223 := (= #153 #222)
+#220 := (= #150 #219)
+#204 := (iff #105 #201)
+#198 := (or #195 #187)
+#202 := (iff #198 #201)
+#203 := [rewrite]: #202
+#199 := (iff #105 #198)
+#196 := (iff #102 #187)
+#197 := [rewrite]: #196
+#184 := (iff #95 #195)
+#185 := [rewrite]: #184
+#200 := [monotonicity #185 #197]: #199
+#205 := [trans #200 #203]: #204
+#221 := [monotonicity #205]: #220
+#224 := [monotonicity #221]: #223
+#227 := [monotonicity #224]: #226
+#230 := [monotonicity #227]: #229
+#233 := [quant-intro #230]: #232
+#182 := (~ #162 #162)
+#179 := (~ #159 #159)
+#192 := [refl]: #179
+#183 := [nnf-pos #192]: #182
+#23 := (- #9)
+#22 := (- #8)
+#31 := (mod #22 #23)
+#32 := (- #31)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#33 := (ite #20 #30 #32)
+#34 := (ite #12 0::int #33)
+#35 := (ite #13 #8 #34)
+#36 := (= #29 #35)
+#37 := (forall (vars (?v0 int) (?v1 int)) #36)
+#165 := (iff #37 #162)
+#59 := (and #16 #18)
+#62 := (or #17 #59)
+#135 := (ite #62 #30 #130)
+#138 := (ite #12 0::int #135)
+#141 := (ite #13 #8 #138)
+#144 := (= #29 #141)
+#147 := (forall (vars (?v0 int) (?v1 int)) #144)
+#163 := (iff #147 #162)
+#160 := (iff #144 #159)
+#157 := (= #141 #156)
+#154 := (= #138 #153)
+#151 := (= #135 #150)
+#106 := (iff #62 #105)
+#103 := (iff #59 #102)
+#100 := (iff #18 #99)
+#101 := [rewrite]: #100
+#93 := (iff #16 #92)
+#94 := [rewrite]: #93
+#104 := [monotonicity #94 #101]: #103
+#96 := (iff #17 #95)
+#89 := (iff #15 #88)
+#90 := [rewrite]: #89
+#97 := [monotonicity #90 #94]: #96
+#107 := [monotonicity #97 #104]: #106
+#152 := [monotonicity #107]: #151
+#155 := [monotonicity #152]: #154
+#158 := [monotonicity #155]: #157
+#161 := [monotonicity #158]: #160
+#164 := [quant-intro #161]: #163
+#148 := (iff #37 #147)
+#145 := (iff #36 #144)
+#142 := (= #35 #141)
+#139 := (= #34 #138)
+#136 := (= #33 #135)
+#133 := (= #32 #130)
+#127 := (- #124)
+#131 := (= #127 #130)
+#132 := [rewrite]: #131
+#128 := (= #32 #127)
+#125 := (= #31 #124)
+#70 := (= #23 #69)
+#71 := [rewrite]: #70
+#67 := (= #22 #66)
+#68 := [rewrite]: #67
+#126 := [monotonicity #68 #71]: #125
+#129 := [monotonicity #126]: #128
+#134 := [trans #129 #132]: #133
+#63 := (iff #20 #62)
+#60 := (iff #19 #59)
+#61 := [rewrite]: #60
+#64 := [monotonicity #61]: #63
+#137 := [monotonicity #64 #134]: #136
+#140 := [monotonicity #137]: #139
+#143 := [monotonicity #140]: #142
+#146 := [monotonicity #143]: #145
+#149 := [quant-intro #146]: #148
+#166 := [trans #149 #164]: #165
+#123 := [asserted]: #37
+#167 := [mp #123 #166]: #162
+#193 := [mp~ #167 #183]: #162
+#234 := [mp #193 #233]: #231
+#265 := [mp #234 #264]: #262
+#749 := [mp #265 #748]: #744
+#666 := (not #744)
+#669 := (or #666 #40)
+#323 := (* -1::int 1::int)
+#407 := (mod #323 #323)
+#408 := (+ #39 #407)
+#409 := (= #408 0::int)
+#400 := (mod 1::int 1::int)
+#411 := (* -1::int #400)
+#412 := (+ #39 #411)
+#413 := (= #412 0::int)
+#410 := (<= 1::int 0::int)
+#414 := (or #410 #410)
+#393 := (not #414)
+#728 := (>= 1::int 0::int)
+#730 := (or #410 #728)
+#387 := (not #730)
+#517 := (or #387 #393)
+#724 := (ite #517 #413 #409)
+#398 := (= 1::int 0::int)
+#399 := (ite #398 #40 #724)
+#731 := (= 1::int #39)
+#732 := (ite #398 #731 #399)
+#554 := (or #666 #732)
+#571 := (iff #554 #669)
+#565 := (iff #669 #669)
+#573 := [rewrite]: #565
+#570 := (iff #732 #40)
+#735 := (= #39 1::int)
+#559 := (ite false #735 #40)
+#520 := (iff #559 #40)
+#668 := [rewrite]: #520
+#560 := (iff #732 #559)
+#674 := (iff #399 #40)
+#519 := (ite false #40 #40)
+#675 := (iff #519 #40)
+#676 := [rewrite]: #675
+#672 := (iff #399 #519)
+#516 := (iff #724 #40)
+#1 := true
+#679 := (ite true #40 #40)
+#680 := (iff #679 #40)
+#687 := [rewrite]: #680
+#685 := (iff #724 #679)
+#684 := (iff #409 #40)
+#682 := (= #408 #39)
+#699 := (+ #39 0::int)
+#697 := (= #699 #39)
+#701 := [rewrite]: #697
+#539 := (= #408 #699)
+#537 := (= #407 0::int)
+#689 := (mod -1::int -1::int)
+#690 := (= #689 0::int)
+#696 := [rewrite]: #690
+#694 := (= #407 #689)
+#693 := (= #323 -1::int)
+#688 := [rewrite]: #693
+#695 := [monotonicity #688 #688]: #694
+#538 := [trans #695 #696]: #537
+#681 := [monotonicity #538]: #539
+#683 := [trans #681 #701]: #682
+#678 := [monotonicity #683]: #684
+#691 := (iff #413 #40)
+#698 := (= #412 #39)
+#700 := (= #412 #699)
+#420 := (= #411 0::int)
+#707 := (* -1::int 0::int)
+#418 := (= #707 0::int)
+#419 := [rewrite]: #418
+#704 := (= #411 #707)
+#429 := (= #400 0::int)
+#703 := [rewrite]: #429
+#708 := [monotonicity #703]: #704
+#421 := [trans #708 #419]: #420
+#415 := [monotonicity #421]: #700
+#702 := [trans #415 #701]: #698
+#692 := [monotonicity #702]: #691
+#706 := (iff #517 true)
+#727 := (or false true)
+#374 := (iff #727 true)
+#375 := [rewrite]: #374
+#329 := (iff #517 #727)
+#434 := (iff #393 true)
+#723 := (not false)
+#432 := (iff #723 true)
+#433 := [rewrite]: #432
+#709 := (iff #393 #723)
+#722 := (iff #414 false)
+#356 := (or false false)
+#361 := (iff #356 false)
+#362 := [rewrite]: #361
+#719 := (iff #414 #356)
+#385 := (iff #410 false)
+#390 := [rewrite]: #385
+#721 := [monotonicity #390 #390]: #719
+#720 := [trans #721 #362]: #722
+#710 := [monotonicity #720]: #709
+#435 := [trans #710 #433]: #434
+#712 := (iff #387 false)
+#713 := (not true)
+#716 := (iff #713 false)
+#717 := [rewrite]: #716
+#714 := (iff #387 #713)
+#376 := (iff #730 true)
+#370 := (iff #730 #727)
+#726 := (iff #728 true)
+#386 := [rewrite]: #726
+#711 := [monotonicity #390 #386]: #370
+#377 := [trans #711 #375]: #376
+#715 := [monotonicity #377]: #714
+#718 := [trans #715 #717]: #712
+#705 := [monotonicity #718 #435]: #329
+#428 := [trans #705 #375]: #706
+#686 := [monotonicity #428 #692 #678]: #685
+#518 := [trans #686 #687]: #516
+#733 := (iff #398 false)
+#734 := [rewrite]: #733
+#673 := [monotonicity #734 #518]: #672
+#677 := [trans #673 #676]: #674
+#729 := (iff #731 #735)
+#725 := [rewrite]: #729
+#561 := [monotonicity #734 #725 #677]: #560
+#665 := [trans #561 #668]: #570
+#572 := [monotonicity #665]: #571
+#574 := [trans #572 #573]: #571
+#569 := [quant-inst]: #554
+#575 := [mp #569 #574]: #669
+[unit-resolution #575 #749 #168]: false
+unsat
+52cf99958a1402707eac2a505ecb8d6adc72de1a 307 0
+#2 := false
+#11 := 0::int
+decl f4 :: (-> int int int)
+#39 := 1::int
+#38 := 3::int
+#40 := (f4 3::int 1::int)
+#41 := (= #40 0::int)
+#42 := (not #41)
+#169 := [asserted]: #42
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#29 := (f4 #8 #9)
+#744 := (pattern #29)
+#66 := -1::int
+#70 := (* -1::int #9)
+#67 := (* -1::int #8)
+#125 := (mod #67 #70)
+#256 := (+ #29 #125)
+#257 := (= #256 0::int)
+#30 := (mod #8 #9)
+#253 := (* -1::int #30)
+#254 := (+ #29 #253)
+#255 := (= #254 0::int)
+#92 := (<= #9 0::int)
+#88 := (<= #8 0::int)
+#195 := (or #88 #92)
+#196 := (not #195)
+#99 := (>= #8 0::int)
+#187 := (or #92 #99)
+#188 := (not #187)
+#202 := (or #188 #196)
+#258 := (ite #202 #255 #257)
+#252 := (= #29 0::int)
+#12 := (= #8 0::int)
+#259 := (ite #12 #252 #258)
+#251 := (= #8 #29)
+#13 := (= #9 0::int)
+#260 := (ite #13 #251 #259)
+#745 := (forall (vars (?v0 int) (?v1 int)) (:pat #744) #260)
+#263 := (forall (vars (?v0 int) (?v1 int)) #260)
+#748 := (iff #263 #745)
+#746 := (iff #260 #260)
+#747 := [refl]: #746
+#749 := [quant-intro #747]: #748
+#131 := (* -1::int #125)
+#220 := (ite #202 #30 #131)
+#223 := (ite #12 0::int #220)
+#226 := (ite #13 #8 #223)
+#229 := (= #29 #226)
+#232 := (forall (vars (?v0 int) (?v1 int)) #229)
+#264 := (iff #232 #263)
+#261 := (iff #229 #260)
+#262 := [rewrite]: #261
+#265 := [quant-intro #262]: #264
+#100 := (not #99)
+#93 := (not #92)
+#103 := (and #93 #100)
+#89 := (not #88)
+#96 := (and #89 #93)
+#106 := (or #96 #103)
+#151 := (ite #106 #30 #131)
+#154 := (ite #12 0::int #151)
+#157 := (ite #13 #8 #154)
+#160 := (= #29 #157)
+#163 := (forall (vars (?v0 int) (?v1 int)) #160)
+#233 := (iff #163 #232)
+#230 := (iff #160 #229)
+#227 := (= #157 #226)
+#224 := (= #154 #223)
+#221 := (= #151 #220)
+#205 := (iff #106 #202)
+#199 := (or #196 #188)
+#203 := (iff #199 #202)
+#204 := [rewrite]: #203
+#200 := (iff #106 #199)
+#197 := (iff #103 #188)
+#198 := [rewrite]: #197
+#185 := (iff #96 #196)
+#186 := [rewrite]: #185
+#201 := [monotonicity #186 #198]: #200
+#206 := [trans #201 #204]: #205
+#222 := [monotonicity #206]: #221
+#225 := [monotonicity #222]: #224
+#228 := [monotonicity #225]: #227
+#231 := [monotonicity #228]: #230
+#234 := [quant-intro #231]: #233
+#183 := (~ #163 #163)
+#180 := (~ #160 #160)
+#193 := [refl]: #180
+#184 := [nnf-pos #193]: #183
+#23 := (- #9)
+#22 := (- #8)
+#31 := (mod #22 #23)
+#32 := (- #31)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#33 := (ite #20 #30 #32)
+#34 := (ite #12 0::int #33)
+#35 := (ite #13 #8 #34)
+#36 := (= #29 #35)
+#37 := (forall (vars (?v0 int) (?v1 int)) #36)
+#166 := (iff #37 #163)
+#60 := (and #16 #18)
+#63 := (or #17 #60)
+#136 := (ite #63 #30 #131)
+#139 := (ite #12 0::int #136)
+#142 := (ite #13 #8 #139)
+#145 := (= #29 #142)
+#148 := (forall (vars (?v0 int) (?v1 int)) #145)
+#164 := (iff #148 #163)
+#161 := (iff #145 #160)
+#158 := (= #142 #157)
+#155 := (= #139 #154)
+#152 := (= #136 #151)
+#107 := (iff #63 #106)
+#104 := (iff #60 #103)
+#101 := (iff #18 #100)
+#102 := [rewrite]: #101
+#94 := (iff #16 #93)
+#95 := [rewrite]: #94
+#105 := [monotonicity #95 #102]: #104
+#97 := (iff #17 #96)
+#90 := (iff #15 #89)
+#91 := [rewrite]: #90
+#98 := [monotonicity #91 #95]: #97
+#108 := [monotonicity #98 #105]: #107
+#153 := [monotonicity #108]: #152
+#156 := [monotonicity #153]: #155
+#159 := [monotonicity #156]: #158
+#162 := [monotonicity #159]: #161
+#165 := [quant-intro #162]: #164
+#149 := (iff #37 #148)
+#146 := (iff #36 #145)
+#143 := (= #35 #142)
+#140 := (= #34 #139)
+#137 := (= #33 #136)
+#134 := (= #32 #131)
+#128 := (- #125)
+#132 := (= #128 #131)
+#133 := [rewrite]: #132
+#129 := (= #32 #128)
+#126 := (= #31 #125)
+#71 := (= #23 #70)
+#72 := [rewrite]: #71
+#68 := (= #22 #67)
+#69 := [rewrite]: #68
+#127 := [monotonicity #69 #72]: #126
+#130 := [monotonicity #127]: #129
+#135 := [trans #130 #133]: #134
+#64 := (iff #20 #63)
+#61 := (iff #19 #60)
+#62 := [rewrite]: #61
+#65 := [monotonicity #62]: #64
+#138 := [monotonicity #65 #135]: #137
+#141 := [monotonicity #138]: #140
+#144 := [monotonicity #141]: #143
+#147 := [monotonicity #144]: #146
+#150 := [quant-intro #147]: #149
+#167 := [trans #150 #165]: #166
+#124 := [asserted]: #37
+#168 := [mp #124 #167]: #163
+#194 := [mp~ #168 #184]: #163
+#235 := [mp #194 #234]: #232
+#266 := [mp #235 #265]: #263
+#750 := [mp #266 #749]: #745
+#577 := (not #745)
+#578 := (or #577 #41)
+#324 := (* -1::int 1::int)
+#408 := (* -1::int 3::int)
+#409 := (mod #408 #324)
+#410 := (+ #40 #409)
+#401 := (= #410 0::int)
+#412 := (mod 3::int 1::int)
+#413 := (* -1::int #412)
+#414 := (+ #40 #413)
+#411 := (= #414 0::int)
+#415 := (<= 1::int 0::int)
+#394 := (<= 3::int 0::int)
+#729 := (or #394 #415)
+#731 := (not #729)
+#388 := (>= 3::int 0::int)
+#518 := (or #415 #388)
+#725 := (not #518)
+#399 := (or #725 #731)
+#400 := (ite #399 #411 #401)
+#732 := (= 3::int 0::int)
+#733 := (ite #732 #41 #400)
+#734 := (= 3::int #40)
+#735 := (= 1::int 0::int)
+#730 := (ite #735 #734 #733)
+#671 := (or #577 #730)
+#672 := (iff #671 #578)
+#661 := (iff #578 #578)
+#653 := [rewrite]: #661
+#575 := (iff #730 #41)
+#391 := (= #40 3::int)
+#570 := (ite false #391 #41)
+#566 := (iff #570 #41)
+#574 := [rewrite]: #566
+#572 := (iff #730 #570)
+#670 := (iff #733 #41)
+#521 := (ite false #41 #41)
+#666 := (iff #521 #41)
+#667 := [rewrite]: #666
+#669 := (iff #733 #521)
+#561 := (iff #400 #41)
+#1 := true
+#676 := (ite true #41 #41)
+#678 := (iff #676 #41)
+#560 := [rewrite]: #678
+#677 := (iff #400 #676)
+#673 := (iff #401 #41)
+#519 := (= #410 #40)
+#692 := (+ #40 0::int)
+#689 := (= #692 #40)
+#690 := [rewrite]: #689
+#688 := (= #410 #692)
+#687 := (= #409 0::int)
+#538 := -3::int
+#684 := (mod -3::int -1::int)
+#680 := (= #684 0::int)
+#686 := [rewrite]: #680
+#685 := (= #409 #684)
+#682 := (= #324 -1::int)
+#683 := [rewrite]: #682
+#539 := (= #408 -3::int)
+#540 := [rewrite]: #539
+#679 := [monotonicity #540 #683]: #685
+#681 := [trans #679 #686]: #687
+#517 := [monotonicity #681]: #688
+#520 := [trans #517 #690]: #519
+#674 := [monotonicity #520]: #673
+#691 := (iff #411 #41)
+#695 := (= #414 #40)
+#693 := (= #414 #692)
+#699 := (= #413 0::int)
+#700 := (* -1::int 0::int)
+#698 := (= #700 0::int)
+#702 := [rewrite]: #698
+#701 := (= #413 #700)
+#421 := (= #412 0::int)
+#422 := [rewrite]: #421
+#416 := [monotonicity #422]: #701
+#703 := [trans #416 #702]: #699
+#694 := [monotonicity #703]: #693
+#696 := [trans #694 #690]: #695
+#697 := [monotonicity #696]: #691
+#419 := (iff #399 true)
+#377 := (or false true)
+#715 := (iff #377 true)
+#716 := [rewrite]: #715
+#705 := (iff #399 #377)
+#704 := (iff #731 true)
+#330 := (not false)
+#429 := (iff #330 true)
+#430 := [rewrite]: #429
+#706 := (iff #731 #330)
+#435 := (iff #729 false)
+#724 := (or false false)
+#433 := (iff #724 false)
+#434 := [rewrite]: #433
+#710 := (iff #729 #724)
+#371 := (iff #415 false)
+#712 := [rewrite]: #371
+#723 := (iff #394 false)
+#721 := [rewrite]: #723
+#711 := [monotonicity #721 #712]: #710
+#436 := [trans #711 #434]: #435
+#707 := [monotonicity #436]: #706
+#708 := [trans #707 #430]: #704
+#362 := (iff #725 false)
+#713 := (not true)
+#720 := (iff #713 false)
+#722 := [rewrite]: #720
+#719 := (iff #725 #713)
+#717 := (iff #518 true)
+#378 := (iff #518 #377)
+#375 := (iff #388 true)
+#376 := [rewrite]: #375
+#714 := [monotonicity #712 #376]: #378
+#718 := [trans #714 #716]: #717
+#357 := [monotonicity #718]: #719
+#363 := [trans #357 #722]: #362
+#709 := [monotonicity #363 #708]: #705
+#420 := [trans #709 #716]: #419
+#675 := [monotonicity #420 #697 #674]: #677
+#562 := [trans #675 #560]: #561
+#387 := (iff #732 false)
+#728 := [rewrite]: #387
+#571 := [monotonicity #728 #562]: #669
+#555 := [trans #571 #667]: #670
+#386 := (iff #734 #391)
+#727 := [rewrite]: #386
+#736 := (iff #735 false)
+#726 := [rewrite]: #736
+#573 := [monotonicity #726 #727 #555]: #572
+#576 := [trans #573 #574]: #575
+#659 := [monotonicity #576]: #672
+#650 := [trans #659 #653]: #672
+#668 := [quant-inst]: #671
+#652 := [mp #668 #650]: #578
+[unit-resolution #652 #750 #169]: false
+unsat
+6d75bf2e8df3400e173e94c70d533cd85e2830a4 308 0
+#2 := false
+#11 := 0::int
+decl f5 :: int
+#38 := f5
+#732 := (= f5 0::int)
+#573 := (not #732)
+#394 := (<= f5 0::int)
+#720 := (not #394)
+#388 := (>= f5 0::int)
+#377 := (not #388)
+#688 := (or #377 #720 #732)
+#575 := (not #688)
+#66 := -1::int
+#408 := (* -1::int f5)
+#700 := (mod #408 -1::int)
+decl f4 :: (-> int int int)
+#39 := 1::int
+#40 := (f4 f5 1::int)
+#698 := (+ #40 #700)
+#703 := (= #698 0::int)
+#41 := (= #40 0::int)
+#520 := (ite #688 #41 #703)
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#29 := (f4 #8 #9)
+#744 := (pattern #29)
+#70 := (* -1::int #9)
+#67 := (* -1::int #8)
+#125 := (mod #67 #70)
+#256 := (+ #29 #125)
+#257 := (= #256 0::int)
+#30 := (mod #8 #9)
+#253 := (* -1::int #30)
+#254 := (+ #29 #253)
+#255 := (= #254 0::int)
+#92 := (<= #9 0::int)
+#88 := (<= #8 0::int)
+#195 := (or #88 #92)
+#196 := (not #195)
+#99 := (>= #8 0::int)
+#187 := (or #92 #99)
+#188 := (not #187)
+#202 := (or #188 #196)
+#258 := (ite #202 #255 #257)
+#252 := (= #29 0::int)
+#12 := (= #8 0::int)
+#259 := (ite #12 #252 #258)
+#251 := (= #8 #29)
+#13 := (= #9 0::int)
+#260 := (ite #13 #251 #259)
+#745 := (forall (vars (?v0 int) (?v1 int)) (:pat #744) #260)
+#263 := (forall (vars (?v0 int) (?v1 int)) #260)
+#748 := (iff #263 #745)
+#746 := (iff #260 #260)
+#747 := [refl]: #746
+#749 := [quant-intro #747]: #748
+#131 := (* -1::int #125)
+#220 := (ite #202 #30 #131)
+#223 := (ite #12 0::int #220)
+#226 := (ite #13 #8 #223)
+#229 := (= #29 #226)
+#232 := (forall (vars (?v0 int) (?v1 int)) #229)
+#264 := (iff #232 #263)
+#261 := (iff #229 #260)
+#262 := [rewrite]: #261
+#265 := [quant-intro #262]: #264
+#100 := (not #99)
+#93 := (not #92)
+#103 := (and #93 #100)
+#89 := (not #88)
+#96 := (and #89 #93)
+#106 := (or #96 #103)
+#151 := (ite #106 #30 #131)
+#154 := (ite #12 0::int #151)
+#157 := (ite #13 #8 #154)
+#160 := (= #29 #157)
+#163 := (forall (vars (?v0 int) (?v1 int)) #160)
+#233 := (iff #163 #232)
+#230 := (iff #160 #229)
+#227 := (= #157 #226)
+#224 := (= #154 #223)
+#221 := (= #151 #220)
+#205 := (iff #106 #202)
+#199 := (or #196 #188)
+#203 := (iff #199 #202)
+#204 := [rewrite]: #203
+#200 := (iff #106 #199)
+#197 := (iff #103 #188)
+#198 := [rewrite]: #197
+#185 := (iff #96 #196)
+#186 := [rewrite]: #185
+#201 := [monotonicity #186 #198]: #200
+#206 := [trans #201 #204]: #205
+#222 := [monotonicity #206]: #221
+#225 := [monotonicity #222]: #224
+#228 := [monotonicity #225]: #227
+#231 := [monotonicity #228]: #230
+#234 := [quant-intro #231]: #233
+#183 := (~ #163 #163)
+#180 := (~ #160 #160)
+#193 := [refl]: #180
+#184 := [nnf-pos #193]: #183
+#23 := (- #9)
+#22 := (- #8)
+#31 := (mod #22 #23)
+#32 := (- #31)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#33 := (ite #20 #30 #32)
+#34 := (ite #12 0::int #33)
+#35 := (ite #13 #8 #34)
+#36 := (= #29 #35)
+#37 := (forall (vars (?v0 int) (?v1 int)) #36)
+#166 := (iff #37 #163)
+#60 := (and #16 #18)
+#63 := (or #17 #60)
+#136 := (ite #63 #30 #131)
+#139 := (ite #12 0::int #136)
+#142 := (ite #13 #8 #139)
+#145 := (= #29 #142)
+#148 := (forall (vars (?v0 int) (?v1 int)) #145)
+#164 := (iff #148 #163)
+#161 := (iff #145 #160)
+#158 := (= #142 #157)
+#155 := (= #139 #154)
+#152 := (= #136 #151)
+#107 := (iff #63 #106)
+#104 := (iff #60 #103)
+#101 := (iff #18 #100)
+#102 := [rewrite]: #101
+#94 := (iff #16 #93)
+#95 := [rewrite]: #94
+#105 := [monotonicity #95 #102]: #104
+#97 := (iff #17 #96)
+#90 := (iff #15 #89)
+#91 := [rewrite]: #90
+#98 := [monotonicity #91 #95]: #97
+#108 := [monotonicity #98 #105]: #107
+#153 := [monotonicity #108]: #152
+#156 := [monotonicity #153]: #155
+#159 := [monotonicity #156]: #158
+#162 := [monotonicity #159]: #161
+#165 := [quant-intro #162]: #164
+#149 := (iff #37 #148)
+#146 := (iff #36 #145)
+#143 := (= #35 #142)
+#140 := (= #34 #139)
+#137 := (= #33 #136)
+#134 := (= #32 #131)
+#128 := (- #125)
+#132 := (= #128 #131)
+#133 := [rewrite]: #132
+#129 := (= #32 #128)
+#126 := (= #31 #125)
+#71 := (= #23 #70)
+#72 := [rewrite]: #71
+#68 := (= #22 #67)
+#69 := [rewrite]: #68
+#127 := [monotonicity #69 #72]: #126
+#130 := [monotonicity #127]: #129
+#135 := [trans #130 #133]: #134
+#64 := (iff #20 #63)
+#61 := (iff #19 #60)
+#62 := [rewrite]: #61
+#65 := [monotonicity #62]: #64
+#138 := [monotonicity #65 #135]: #137
+#141 := [monotonicity #138]: #140
+#144 := [monotonicity #141]: #143
+#147 := [monotonicity #144]: #146
+#150 := [quant-intro #147]: #149
+#167 := [trans #150 #165]: #166
+#124 := [asserted]: #37
+#168 := [mp #124 #167]: #163
+#194 := [mp~ #168 #184]: #163
+#235 := [mp #194 #234]: #232
+#266 := [mp #235 #265]: #263
+#750 := [mp #266 #749]: #745
+#675 := (not #745)
+#678 := (or #675 #520)
+#324 := (* -1::int 1::int)
+#409 := (mod #408 #324)
+#410 := (+ #40 #409)
+#401 := (= #410 0::int)
+#412 := (mod f5 1::int)
+#413 := (* -1::int #412)
+#414 := (+ #40 #413)
+#411 := (= #414 0::int)
+#415 := (<= 1::int 0::int)
+#729 := (or #394 #415)
+#731 := (not #729)
+#518 := (or #415 #388)
+#725 := (not #518)
+#399 := (or #725 #731)
+#400 := (ite #399 #411 #401)
+#733 := (ite #732 #41 #400)
+#734 := (= f5 #40)
+#735 := (= 1::int 0::int)
+#730 := (ite #735 #734 #733)
+#560 := (or #675 #730)
+#562 := (iff #560 #678)
+#669 := (iff #678 #678)
+#571 := [rewrite]: #669
+#676 := (iff #730 #520)
+#363 := (or #377 #720)
+#697 := (or #363 #732)
+#538 := (ite #697 #41 #703)
+#673 := (iff #538 #520)
+#517 := (iff #697 #688)
+#519 := [rewrite]: #517
+#674 := [monotonicity #519]: #673
+#687 := (iff #730 #538)
+#684 := (ite false #734 #538)
+#680 := (iff #684 #538)
+#686 := [rewrite]: #680
+#685 := (iff #730 #684)
+#682 := (iff #733 #538)
+#694 := (ite #363 #41 #703)
+#695 := (ite #732 #41 #694)
+#539 := (iff #695 #538)
+#540 := [rewrite]: #539
+#696 := (iff #733 #695)
+#689 := (iff #400 #694)
+#692 := (iff #401 #703)
+#702 := (= #410 #698)
+#701 := (= #409 #700)
+#421 := (= #324 -1::int)
+#422 := [rewrite]: #421
+#416 := [monotonicity #422]: #701
+#699 := [monotonicity #416]: #702
+#693 := [monotonicity #699]: #692
+#419 := (iff #411 #41)
+#705 := (= #414 #40)
+#707 := (+ #40 0::int)
+#704 := (= #707 #40)
+#708 := [rewrite]: #704
+#429 := (= #414 #707)
+#330 := (= #413 0::int)
+#711 := (* -1::int 0::int)
+#435 := (= #711 0::int)
+#436 := [rewrite]: #435
+#433 := (= #413 #711)
+#724 := (= #412 0::int)
+#710 := [rewrite]: #724
+#434 := [monotonicity #710]: #433
+#706 := [trans #434 #436]: #330
+#430 := [monotonicity #706]: #429
+#709 := [trans #430 #708]: #705
+#420 := [monotonicity #709]: #419
+#723 := (iff #399 #363)
+#722 := (iff #731 #720)
+#719 := (iff #729 #394)
+#715 := (or #394 false)
+#718 := (iff #715 #394)
+#713 := [rewrite]: #718
+#716 := (iff #729 #715)
+#386 := (iff #415 false)
+#391 := [rewrite]: #386
+#717 := [monotonicity #391]: #716
+#357 := [trans #717 #713]: #719
+#362 := [monotonicity #357]: #722
+#378 := (iff #725 #377)
+#375 := (iff #518 #388)
+#727 := (or false #388)
+#371 := (iff #727 #388)
+#712 := [rewrite]: #371
+#387 := (iff #518 #727)
+#728 := [monotonicity #391]: #387
+#376 := [trans #728 #712]: #375
+#714 := [monotonicity #376]: #378
+#721 := [monotonicity #714 #362]: #723
+#690 := [monotonicity #721 #420 #693]: #689
+#691 := [monotonicity #690]: #696
+#683 := [trans #691 #540]: #682
+#736 := (iff #735 false)
+#726 := [rewrite]: #736
+#679 := [monotonicity #726 #683]: #685
+#681 := [trans #679 #686]: #687
+#677 := [trans #681 #674]: #676
+#521 := [monotonicity #677]: #562
+#666 := [trans #521 #571]: #562
+#561 := [quant-inst]: #560
+#667 := [mp #561 #666]: #678
+#660 := [unit-resolution #667 #750]: #520
+#668 := (not #520)
+#665 := (or #668 #575)
+#42 := (not #41)
+#169 := [asserted]: #42
+#672 := (or #668 #575 #41)
+#659 := [def-axiom]: #672
+#654 := [unit-resolution #659 #169]: #665
+#655 := [unit-resolution #654 #660]: #575
+#566 := (or #688 #573)
+#574 := [def-axiom]: #566
+#656 := [unit-resolution #574 #655]: #573
+#670 := (or #688 #388)
+#555 := [def-axiom]: #670
+#657 := [unit-resolution #555 #655]: #388
+#570 := (or #688 #394)
+#572 := [def-axiom]: #570
+#651 := [unit-resolution #572 #655]: #394
+#658 := (or #732 #720 #377)
+#642 := [th-lemma]: #658
+[unit-resolution #642 #651 #657 #656]: false
+unsat
+a09ca425a49bbf59a7db8aba4b01dbd7473e9c09 317 0
+#2 := false
+#11 := 0::int
+decl f4 :: (-> int int int)
+#66 := -1::int
+#172 := (f4 0::int -1::int)
+#175 := (= #172 0::int)
+#188 := (not #175)
+#38 := 1::int
+#39 := (- 1::int)
+#40 := (f4 0::int #39)
+#41 := (= #40 0::int)
+#42 := (not #41)
+#191 := (iff #42 #188)
+#178 := (= 0::int #172)
+#183 := (not #178)
+#189 := (iff #183 #188)
+#186 := (iff #178 #175)
+#187 := [rewrite]: #186
+#190 := [monotonicity #187]: #189
+#184 := (iff #42 #183)
+#181 := (iff #41 #178)
+#179 := (iff #175 #178)
+#180 := [rewrite]: #179
+#176 := (iff #41 #175)
+#173 := (= #40 #172)
+#170 := (= #39 -1::int)
+#171 := [rewrite]: #170
+#174 := [monotonicity #171]: #173
+#177 := [monotonicity #174]: #176
+#182 := [trans #177 #180]: #181
+#185 := [monotonicity #182]: #184
+#192 := [trans #185 #190]: #191
+#169 := [asserted]: #42
+#193 := [mp #169 #192]: #188
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#29 := (f4 #8 #9)
+#758 := (pattern #29)
+#70 := (* -1::int #9)
+#67 := (* -1::int #8)
+#125 := (mod #67 #70)
+#270 := (+ #29 #125)
+#271 := (= #270 0::int)
+#30 := (mod #8 #9)
+#267 := (* -1::int #30)
+#268 := (+ #29 #267)
+#269 := (= #268 0::int)
+#92 := (<= #9 0::int)
+#88 := (<= #8 0::int)
+#209 := (or #88 #92)
+#210 := (not #209)
+#99 := (>= #8 0::int)
+#201 := (or #92 #99)
+#202 := (not #201)
+#216 := (or #202 #210)
+#272 := (ite #216 #269 #271)
+#266 := (= #29 0::int)
+#12 := (= #8 0::int)
+#273 := (ite #12 #266 #272)
+#265 := (= #8 #29)
+#13 := (= #9 0::int)
+#274 := (ite #13 #265 #273)
+#759 := (forall (vars (?v0 int) (?v1 int)) (:pat #758) #274)
+#277 := (forall (vars (?v0 int) (?v1 int)) #274)
+#762 := (iff #277 #759)
+#760 := (iff #274 #274)
+#761 := [refl]: #760
+#763 := [quant-intro #761]: #762
+#131 := (* -1::int #125)
+#234 := (ite #216 #30 #131)
+#237 := (ite #12 0::int #234)
+#240 := (ite #13 #8 #237)
+#243 := (= #29 #240)
+#246 := (forall (vars (?v0 int) (?v1 int)) #243)
+#278 := (iff #246 #277)
+#275 := (iff #243 #274)
+#276 := [rewrite]: #275
+#279 := [quant-intro #276]: #278
+#100 := (not #99)
+#93 := (not #92)
+#103 := (and #93 #100)
+#89 := (not #88)
+#96 := (and #89 #93)
+#106 := (or #96 #103)
+#151 := (ite #106 #30 #131)
+#154 := (ite #12 0::int #151)
+#157 := (ite #13 #8 #154)
+#160 := (= #29 #157)
+#163 := (forall (vars (?v0 int) (?v1 int)) #160)
+#247 := (iff #163 #246)
+#244 := (iff #160 #243)
+#241 := (= #157 #240)
+#238 := (= #154 #237)
+#235 := (= #151 #234)
+#219 := (iff #106 #216)
+#213 := (or #210 #202)
+#217 := (iff #213 #216)
+#218 := [rewrite]: #217
+#214 := (iff #106 #213)
+#211 := (iff #103 #202)
+#212 := [rewrite]: #211
+#199 := (iff #96 #210)
+#200 := [rewrite]: #199
+#215 := [monotonicity #200 #212]: #214
+#220 := [trans #215 #218]: #219
+#236 := [monotonicity #220]: #235
+#239 := [monotonicity #236]: #238
+#242 := [monotonicity #239]: #241
+#245 := [monotonicity #242]: #244
+#248 := [quant-intro #245]: #247
+#197 := (~ #163 #163)
+#194 := (~ #160 #160)
+#207 := [refl]: #194
+#198 := [nnf-pos #207]: #197
+#23 := (- #9)
+#22 := (- #8)
+#31 := (mod #22 #23)
+#32 := (- #31)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#33 := (ite #20 #30 #32)
+#34 := (ite #12 0::int #33)
+#35 := (ite #13 #8 #34)
+#36 := (= #29 #35)
+#37 := (forall (vars (?v0 int) (?v1 int)) #36)
+#166 := (iff #37 #163)
+#60 := (and #16 #18)
+#63 := (or #17 #60)
+#136 := (ite #63 #30 #131)
+#139 := (ite #12 0::int #136)
+#142 := (ite #13 #8 #139)
+#145 := (= #29 #142)
+#148 := (forall (vars (?v0 int) (?v1 int)) #145)
+#164 := (iff #148 #163)
+#161 := (iff #145 #160)
+#158 := (= #142 #157)
+#155 := (= #139 #154)
+#152 := (= #136 #151)
+#107 := (iff #63 #106)
+#104 := (iff #60 #103)
+#101 := (iff #18 #100)
+#102 := [rewrite]: #101
+#94 := (iff #16 #93)
+#95 := [rewrite]: #94
+#105 := [monotonicity #95 #102]: #104
+#97 := (iff #17 #96)
+#90 := (iff #15 #89)
+#91 := [rewrite]: #90
+#98 := [monotonicity #91 #95]: #97
+#108 := [monotonicity #98 #105]: #107
+#153 := [monotonicity #108]: #152
+#156 := [monotonicity #153]: #155
+#159 := [monotonicity #156]: #158
+#162 := [monotonicity #159]: #161
+#165 := [quant-intro #162]: #164
+#149 := (iff #37 #148)
+#146 := (iff #36 #145)
+#143 := (= #35 #142)
+#140 := (= #34 #139)
+#137 := (= #33 #136)
+#134 := (= #32 #131)
+#128 := (- #125)
+#132 := (= #128 #131)
+#133 := [rewrite]: #132
+#129 := (= #32 #128)
+#126 := (= #31 #125)
+#71 := (= #23 #70)
+#72 := [rewrite]: #71
+#68 := (= #22 #67)
+#69 := [rewrite]: #68
+#127 := [monotonicity #69 #72]: #126
+#130 := [monotonicity #127]: #129
+#135 := [trans #130 #133]: #134
+#64 := (iff #20 #63)
+#61 := (iff #19 #60)
+#62 := [rewrite]: #61
+#65 := [monotonicity #62]: #64
+#138 := [monotonicity #65 #135]: #137
+#141 := [monotonicity #138]: #140
+#144 := [monotonicity #141]: #143
+#147 := [monotonicity #144]: #146
+#150 := [quant-intro #147]: #149
+#167 := [trans #150 #165]: #166
+#124 := [asserted]: #37
+#168 := [mp #124 #167]: #163
+#208 := [mp~ #168 #198]: #163
+#249 := [mp #208 #248]: #246
+#280 := [mp #249 #279]: #277
+#764 := [mp #280 #763]: #759
+#535 := (not #759)
+#683 := (or #535 #175)
+#338 := (* -1::int -1::int)
+#422 := (* -1::int 0::int)
+#423 := (mod #422 #338)
+#424 := (+ #172 #423)
+#415 := (= #424 0::int)
+#426 := (mod 0::int -1::int)
+#427 := (* -1::int #426)
+#428 := (+ #172 #427)
+#425 := (= #428 0::int)
+#429 := (<= -1::int 0::int)
+#408 := (<= 0::int 0::int)
+#743 := (or #408 #429)
+#745 := (not #743)
+#402 := (>= 0::int 0::int)
+#532 := (or #429 #402)
+#739 := (not #532)
+#413 := (or #739 #745)
+#414 := (ite #413 #425 #415)
+#746 := (= 0::int 0::int)
+#747 := (ite #746 #175 #414)
+#748 := (= -1::int 0::int)
+#749 := (ite #748 #178 #747)
+#585 := (or #535 #749)
+#681 := (iff #585 #683)
+#569 := (iff #683 #683)
+#584 := [rewrite]: #569
+#575 := (iff #749 #175)
+#693 := (ite false #175 #175)
+#701 := (iff #693 #175)
+#695 := [rewrite]: #701
+#692 := (iff #749 #693)
+#691 := (iff #747 #175)
+#1 := true
+#533 := (ite true #175 #175)
+#688 := (iff #533 #175)
+#690 := [rewrite]: #688
+#534 := (iff #747 #533)
+#702 := (iff #414 #175)
+#694 := (iff #414 #693)
+#698 := (iff #415 #175)
+#696 := (= #424 #172)
+#436 := (+ #172 0::int)
+#430 := (= #436 #172)
+#712 := [rewrite]: #430
+#553 := (= #424 #436)
+#711 := (= #423 0::int)
+#703 := (mod 0::int 1::int)
+#710 := (= #703 0::int)
+#705 := [rewrite]: #710
+#704 := (= #423 #703)
+#707 := (= #338 1::int)
+#708 := [rewrite]: #707
+#723 := (= #422 0::int)
+#433 := [rewrite]: #723
+#709 := [monotonicity #433 #708]: #704
+#552 := [trans #709 #705]: #711
+#554 := [monotonicity #552]: #553
+#697 := [trans #554 #712]: #696
+#699 := [monotonicity #697]: #698
+#717 := (iff #425 #175)
+#716 := (= #428 #172)
+#714 := (= #428 #436)
+#434 := (= #427 0::int)
+#722 := (= #427 #422)
+#444 := (= #426 0::int)
+#718 := [rewrite]: #444
+#719 := [monotonicity #718]: #722
+#435 := [trans #719 #433]: #434
+#715 := [monotonicity #435]: #714
+#713 := [trans #715 #712]: #716
+#706 := [monotonicity #713]: #717
+#721 := (iff #413 false)
+#448 := (or false false)
+#344 := (iff #448 false)
+#720 := [rewrite]: #344
+#449 := (iff #413 #448)
+#725 := (iff #745 false)
+#729 := (not true)
+#732 := (iff #729 false)
+#727 := [rewrite]: #732
+#738 := (iff #745 #729)
+#737 := (iff #743 true)
+#385 := (or true true)
+#390 := (iff #385 true)
+#391 := [rewrite]: #390
+#376 := (iff #743 #385)
+#405 := (iff #429 true)
+#741 := [rewrite]: #405
+#734 := (iff #408 true)
+#736 := [rewrite]: #734
+#377 := [monotonicity #736 #741]: #376
+#735 := [trans #377 #391]: #737
+#724 := [monotonicity #735]: #738
+#447 := [trans #724 #727]: #725
+#733 := (iff #739 false)
+#730 := (iff #739 #729)
+#392 := (iff #532 true)
+#726 := (iff #532 #385)
+#401 := (iff #402 true)
+#742 := [rewrite]: #401
+#389 := [monotonicity #741 #742]: #726
+#728 := [trans #389 #391]: #392
+#731 := [monotonicity #728]: #730
+#371 := [trans #731 #727]: #733
+#450 := [monotonicity #371 #447]: #449
+#443 := [trans #450 #720]: #721
+#700 := [monotonicity #443 #706 #699]: #694
+#531 := [trans #700 #695]: #702
+#740 := (iff #746 true)
+#400 := [rewrite]: #740
+#687 := [monotonicity #400 #531]: #534
+#689 := [trans #687 #690]: #691
+#744 := (iff #748 false)
+#750 := [rewrite]: #744
+#574 := [monotonicity #750 #187 #689]: #692
+#576 := [trans #574 #695]: #575
+#684 := [monotonicity #576]: #681
+#586 := [trans #684 #584]: #681
+#680 := [quant-inst]: #585
+#587 := [mp #680 #586]: #683
+[unit-resolution #587 #764 #193]: false
+unsat
+07f15f3f61db6e3805237fa8c50e95e2181bdd46 327 0
 #2 := false
 #11 := 0::int
 decl f4 :: (-> int int int)
@@ -32351,598 +36973,7 @@
 #673 := [mp #590 #686]: #588
 [unit-resolution #673 #764 #193]: false
 unsat
-b94918a15776ebd094f6867ac20bcae7605bc260 282 0
-#2 := false
-#11 := 0::int
-decl f3 :: (-> int int int)
-#38 := 3::int
-#39 := (f3 0::int 3::int)
-#40 := (= #39 0::int)
-#41 := (not #40)
-#168 := [asserted]: #41
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#10 := (f3 #8 #9)
-#737 := (pattern #10)
-#65 := -1::int
-#69 := (* -1::int #9)
-#66 := (* -1::int #8)
-#72 := (div #66 #69)
-#239 := (* -1::int #72)
-#240 := (+ #10 #239)
-#241 := (= #240 0::int)
-#21 := (div #8 #9)
-#236 := (* -1::int #21)
-#237 := (+ #10 #236)
-#238 := (= #237 0::int)
-#91 := (<= #9 0::int)
-#87 := (<= #8 0::int)
-#194 := (or #87 #91)
-#195 := (not #194)
-#98 := (>= #8 0::int)
-#186 := (or #91 #98)
-#187 := (not #186)
-#201 := (or #187 #195)
-#242 := (ite #201 #238 #241)
-#235 := (= #10 0::int)
-#13 := (= #9 0::int)
-#12 := (= #8 0::int)
-#14 := (or #12 #13)
-#243 := (ite #14 #235 #242)
-#738 := (forall (vars (?v0 int) (?v1 int)) (:pat #737) #243)
-#246 := (forall (vars (?v0 int) (?v1 int)) #243)
-#741 := (iff #246 #738)
-#739 := (iff #243 #243)
-#740 := [refl]: #739
-#742 := [quant-intro #740]: #741
-#206 := (ite #201 #21 #72)
-#209 := (ite #14 0::int #206)
-#212 := (= #10 #209)
-#215 := (forall (vars (?v0 int) (?v1 int)) #212)
-#247 := (iff #215 #246)
-#244 := (iff #212 #243)
-#245 := [rewrite]: #244
-#248 := [quant-intro #245]: #247
-#99 := (not #98)
-#92 := (not #91)
-#102 := (and #92 #99)
-#88 := (not #87)
-#95 := (and #88 #92)
-#105 := (or #95 #102)
-#108 := (ite #105 #21 #72)
-#111 := (ite #14 0::int #108)
-#114 := (= #10 #111)
-#117 := (forall (vars (?v0 int) (?v1 int)) #114)
-#216 := (iff #117 #215)
-#213 := (iff #114 #212)
-#210 := (= #111 #209)
-#207 := (= #108 #206)
-#204 := (iff #105 #201)
-#198 := (or #195 #187)
-#202 := (iff #198 #201)
-#203 := [rewrite]: #202
-#199 := (iff #105 #198)
-#196 := (iff #102 #187)
-#197 := [rewrite]: #196
-#184 := (iff #95 #195)
-#185 := [rewrite]: #184
-#200 := [monotonicity #185 #197]: #199
-#205 := [trans #200 #203]: #204
-#208 := [monotonicity #205]: #207
-#211 := [monotonicity #208]: #210
-#214 := [monotonicity #211]: #213
-#217 := [quant-intro #214]: #216
-#190 := (~ #117 #117)
-#188 := (~ #114 #114)
-#189 := [refl]: #188
-#191 := [nnf-pos #189]: #190
-#23 := (- #9)
-#22 := (- #8)
-#24 := (div #22 #23)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#25 := (ite #20 #21 #24)
-#26 := (ite #14 0::int #25)
-#27 := (= #10 #26)
-#28 := (forall (vars (?v0 int) (?v1 int)) #27)
-#120 := (iff #28 #117)
-#59 := (and #16 #18)
-#62 := (or #17 #59)
-#75 := (ite #62 #21 #72)
-#78 := (ite #14 0::int #75)
-#81 := (= #10 #78)
-#84 := (forall (vars (?v0 int) (?v1 int)) #81)
-#118 := (iff #84 #117)
-#115 := (iff #81 #114)
-#112 := (= #78 #111)
-#109 := (= #75 #108)
-#106 := (iff #62 #105)
-#103 := (iff #59 #102)
-#100 := (iff #18 #99)
-#101 := [rewrite]: #100
-#93 := (iff #16 #92)
-#94 := [rewrite]: #93
-#104 := [monotonicity #94 #101]: #103
-#96 := (iff #17 #95)
-#89 := (iff #15 #88)
-#90 := [rewrite]: #89
-#97 := [monotonicity #90 #94]: #96
-#107 := [monotonicity #97 #104]: #106
-#110 := [monotonicity #107]: #109
-#113 := [monotonicity #110]: #112
-#116 := [monotonicity #113]: #115
-#119 := [quant-intro #116]: #118
-#85 := (iff #28 #84)
-#82 := (iff #27 #81)
-#79 := (= #26 #78)
-#76 := (= #25 #75)
-#73 := (= #24 #72)
-#70 := (= #23 #69)
-#71 := [rewrite]: #70
-#67 := (= #22 #66)
-#68 := [rewrite]: #67
-#74 := [monotonicity #68 #71]: #73
-#63 := (iff #20 #62)
-#60 := (iff #19 #59)
-#61 := [rewrite]: #60
-#64 := [monotonicity #61]: #63
-#77 := [monotonicity #64 #74]: #76
-#80 := [monotonicity #77]: #79
-#83 := [monotonicity #80]: #82
-#86 := [quant-intro #83]: #85
-#121 := [trans #86 #119]: #120
-#58 := [asserted]: #28
-#122 := [mp #58 #121]: #117
-#181 := [mp~ #122 #191]: #117
-#218 := [mp #181 #217]: #215
-#249 := [mp #218 #248]: #246
-#743 := [mp #249 #742]: #738
-#572 := (not #738)
-#573 := (or #572 #40)
-#323 := (* -1::int 3::int)
-#408 := (* -1::int 0::int)
-#409 := (div #408 #323)
-#410 := (* -1::int #409)
-#401 := (+ #39 #410)
-#412 := (= #401 0::int)
-#413 := (div 0::int 3::int)
-#341 := (* -1::int #413)
-#414 := (+ #39 #341)
-#411 := (= #414 0::int)
-#415 := (<= 3::int 0::int)
-#394 := (<= 0::int 0::int)
-#729 := (or #394 #415)
-#731 := (not #729)
-#388 := (>= 0::int 0::int)
-#518 := (or #415 #388)
-#725 := (not #518)
-#399 := (or #725 #731)
-#400 := (ite #399 #411 #412)
-#732 := (= 3::int 0::int)
-#733 := (= 0::int 0::int)
-#734 := (or #733 #732)
-#735 := (ite #734 #40 #400)
-#566 := (or #572 #735)
-#575 := (iff #566 #573)
-#577 := (iff #573 #573)
-#578 := [rewrite]: #577
-#555 := (iff #735 #40)
-#1 := true
-#669 := (ite true #40 #40)
-#667 := (iff #669 #40)
-#670 := [rewrite]: #667
-#571 := (iff #735 #669)
-#562 := (iff #400 #40)
-#677 := (ite false #40 #40)
-#560 := (iff #677 #40)
-#561 := [rewrite]: #560
-#675 := (iff #400 #677)
-#674 := (iff #412 #40)
-#520 := (= #401 #39)
-#703 := (+ #39 0::int)
-#694 := (= #703 #39)
-#689 := [rewrite]: #694
-#517 := (= #401 #703)
-#681 := (= #410 0::int)
-#416 := (= #408 0::int)
-#698 := [rewrite]: #416
-#686 := (= #410 #408)
-#679 := (= #409 0::int)
-#697 := -3::int
-#540 := (div 0::int -3::int)
-#684 := (= #540 0::int)
-#685 := [rewrite]: #684
-#682 := (= #409 #540)
-#538 := (= #323 -3::int)
-#539 := [rewrite]: #538
-#683 := [monotonicity #698 #539]: #682
-#680 := [trans #683 #685]: #679
-#687 := [monotonicity #680]: #686
-#688 := [trans #687 #698]: #681
-#519 := [monotonicity #688]: #517
-#673 := [trans #519 #689]: #520
-#676 := [monotonicity #673]: #674
-#696 := (iff #411 #40)
-#690 := (= #414 #39)
-#692 := (= #414 #703)
-#702 := (= #341 0::int)
-#700 := (= #341 #408)
-#421 := (= #413 0::int)
-#422 := [rewrite]: #421
-#701 := [monotonicity #422]: #700
-#699 := [trans #701 #698]: #702
-#693 := [monotonicity #699]: #692
-#695 := [trans #693 #689]: #690
-#691 := [monotonicity #695]: #696
-#419 := (iff #399 false)
-#430 := (or false false)
-#705 := (iff #430 false)
-#709 := [rewrite]: #705
-#704 := (iff #399 #430)
-#707 := (iff #731 false)
-#720 := (not true)
-#363 := (iff #720 false)
-#723 := [rewrite]: #363
-#329 := (iff #731 #720)
-#435 := (iff #729 true)
-#391 := (or true false)
-#728 := (iff #391 true)
-#371 := [rewrite]: #728
-#433 := (iff #729 #391)
-#376 := (iff #415 false)
-#377 := [rewrite]: #376
-#710 := (iff #394 true)
-#711 := [rewrite]: #710
-#434 := [monotonicity #711 #377]: #433
-#436 := [trans #434 #371]: #435
-#706 := [monotonicity #436]: #329
-#429 := [trans #706 #723]: #707
-#721 := (iff #725 false)
-#722 := (iff #725 #720)
-#719 := (iff #518 true)
-#715 := (or false true)
-#718 := (iff #715 true)
-#713 := [rewrite]: #718
-#716 := (iff #518 #715)
-#378 := (iff #388 true)
-#714 := [rewrite]: #378
-#717 := [monotonicity #377 #714]: #716
-#357 := [trans #717 #713]: #719
-#362 := [monotonicity #357]: #722
-#724 := [trans #362 #723]: #721
-#708 := [monotonicity #724 #429]: #704
-#420 := [trans #708 #709]: #419
-#678 := [monotonicity #420 #691 #676]: #675
-#521 := [trans #678 #561]: #562
-#712 := (iff #734 true)
-#727 := (iff #734 #391)
-#726 := (iff #732 false)
-#386 := [rewrite]: #726
-#730 := (iff #733 true)
-#736 := [rewrite]: #730
-#387 := [monotonicity #736 #386]: #727
-#375 := [trans #387 #371]: #712
-#666 := [monotonicity #375 #521]: #571
-#570 := [trans #666 #670]: #555
-#576 := [monotonicity #570]: #575
-#671 := [trans #576 #578]: #575
-#574 := [quant-inst]: #566
-#668 := [mp #574 #671]: #573
-[unit-resolution #668 #743 #168]: false
-unsat
-f28e01ed5231dd8c14a90a1eb336549610812dce 307 0
-#2 := false
-#11 := 0::int
-decl f3 :: (-> int int int)
-#170 := -3::int
-#173 := (f3 0::int -3::int)
-#176 := (= #173 0::int)
-#189 := (not #176)
-#38 := 3::int
-#39 := (- 3::int)
-#40 := (f3 0::int #39)
-#41 := (= #40 0::int)
-#42 := (not #41)
-#192 := (iff #42 #189)
-#179 := (= 0::int #173)
-#184 := (not #179)
-#190 := (iff #184 #189)
-#187 := (iff #179 #176)
-#188 := [rewrite]: #187
-#191 := [monotonicity #188]: #190
-#185 := (iff #42 #184)
-#182 := (iff #41 #179)
-#180 := (iff #176 #179)
-#181 := [rewrite]: #180
-#177 := (iff #41 #176)
-#174 := (= #40 #173)
-#171 := (= #39 -3::int)
-#172 := [rewrite]: #171
-#175 := [monotonicity #172]: #174
-#178 := [monotonicity #175]: #177
-#183 := [trans #178 #181]: #182
-#186 := [monotonicity #183]: #185
-#193 := [trans #186 #191]: #192
-#169 := [asserted]: #42
-#194 := [mp #169 #193]: #189
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#10 := (f3 #8 #9)
-#753 := (pattern #10)
-#66 := -1::int
-#70 := (* -1::int #9)
-#67 := (* -1::int #8)
-#73 := (div #67 #70)
-#255 := (* -1::int #73)
-#256 := (+ #10 #255)
-#257 := (= #256 0::int)
-#21 := (div #8 #9)
-#252 := (* -1::int #21)
-#253 := (+ #10 #252)
-#254 := (= #253 0::int)
-#92 := (<= #9 0::int)
-#88 := (<= #8 0::int)
-#210 := (or #88 #92)
-#211 := (not #210)
-#99 := (>= #8 0::int)
-#202 := (or #92 #99)
-#203 := (not #202)
-#217 := (or #203 #211)
-#258 := (ite #217 #254 #257)
-#251 := (= #10 0::int)
-#13 := (= #9 0::int)
-#12 := (= #8 0::int)
-#14 := (or #12 #13)
-#259 := (ite #14 #251 #258)
-#754 := (forall (vars (?v0 int) (?v1 int)) (:pat #753) #259)
-#262 := (forall (vars (?v0 int) (?v1 int)) #259)
-#757 := (iff #262 #754)
-#755 := (iff #259 #259)
-#756 := [refl]: #755
-#758 := [quant-intro #756]: #757
-#222 := (ite #217 #21 #73)
-#225 := (ite #14 0::int #222)
-#228 := (= #10 #225)
-#231 := (forall (vars (?v0 int) (?v1 int)) #228)
-#263 := (iff #231 #262)
-#260 := (iff #228 #259)
-#261 := [rewrite]: #260
-#264 := [quant-intro #261]: #263
-#100 := (not #99)
-#93 := (not #92)
-#103 := (and #93 #100)
-#89 := (not #88)
-#96 := (and #89 #93)
-#106 := (or #96 #103)
-#109 := (ite #106 #21 #73)
-#112 := (ite #14 0::int #109)
-#115 := (= #10 #112)
-#118 := (forall (vars (?v0 int) (?v1 int)) #115)
-#232 := (iff #118 #231)
-#229 := (iff #115 #228)
-#226 := (= #112 #225)
-#223 := (= #109 #222)
-#220 := (iff #106 #217)
-#214 := (or #211 #203)
-#218 := (iff #214 #217)
-#219 := [rewrite]: #218
-#215 := (iff #106 #214)
-#212 := (iff #103 #203)
-#213 := [rewrite]: #212
-#200 := (iff #96 #211)
-#201 := [rewrite]: #200
-#216 := [monotonicity #201 #213]: #215
-#221 := [trans #216 #219]: #220
-#224 := [monotonicity #221]: #223
-#227 := [monotonicity #224]: #226
-#230 := [monotonicity #227]: #229
-#233 := [quant-intro #230]: #232
-#206 := (~ #118 #118)
-#204 := (~ #115 #115)
-#205 := [refl]: #204
-#207 := [nnf-pos #205]: #206
-#23 := (- #9)
-#22 := (- #8)
-#24 := (div #22 #23)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#25 := (ite #20 #21 #24)
-#26 := (ite #14 0::int #25)
-#27 := (= #10 #26)
-#28 := (forall (vars (?v0 int) (?v1 int)) #27)
-#121 := (iff #28 #118)
-#60 := (and #16 #18)
-#63 := (or #17 #60)
-#76 := (ite #63 #21 #73)
-#79 := (ite #14 0::int #76)
-#82 := (= #10 #79)
-#85 := (forall (vars (?v0 int) (?v1 int)) #82)
-#119 := (iff #85 #118)
-#116 := (iff #82 #115)
-#113 := (= #79 #112)
-#110 := (= #76 #109)
-#107 := (iff #63 #106)
-#104 := (iff #60 #103)
-#101 := (iff #18 #100)
-#102 := [rewrite]: #101
-#94 := (iff #16 #93)
-#95 := [rewrite]: #94
-#105 := [monotonicity #95 #102]: #104
-#97 := (iff #17 #96)
-#90 := (iff #15 #89)
-#91 := [rewrite]: #90
-#98 := [monotonicity #91 #95]: #97
-#108 := [monotonicity #98 #105]: #107
-#111 := [monotonicity #108]: #110
-#114 := [monotonicity #111]: #113
-#117 := [monotonicity #114]: #116
-#120 := [quant-intro #117]: #119
-#86 := (iff #28 #85)
-#83 := (iff #27 #82)
-#80 := (= #26 #79)
-#77 := (= #25 #76)
-#74 := (= #24 #73)
-#71 := (= #23 #70)
-#72 := [rewrite]: #71
-#68 := (= #22 #67)
-#69 := [rewrite]: #68
-#75 := [monotonicity #69 #72]: #74
-#64 := (iff #20 #63)
-#61 := (iff #19 #60)
-#62 := [rewrite]: #61
-#65 := [monotonicity #62]: #64
-#78 := [monotonicity #65 #75]: #77
-#81 := [monotonicity #78]: #80
-#84 := [monotonicity #81]: #83
-#87 := [quant-intro #84]: #86
-#122 := [trans #87 #120]: #121
-#59 := [asserted]: #28
-#123 := [mp #59 #122]: #118
-#197 := [mp~ #123 #207]: #118
-#234 := [mp #197 #233]: #231
-#265 := [mp #234 #264]: #262
-#759 := [mp #265 #758]: #754
-#586 := (not #754)
-#588 := (or #586 #176)
-#339 := (* -1::int -3::int)
-#424 := (* -1::int 0::int)
-#425 := (div #424 #339)
-#426 := (* -1::int #425)
-#417 := (+ #173 #426)
-#428 := (= #417 0::int)
-#429 := (div 0::int -3::int)
-#357 := (* -1::int #429)
-#430 := (+ #173 #357)
-#427 := (= #430 0::int)
-#431 := (<= -3::int 0::int)
-#410 := (<= 0::int 0::int)
-#745 := (or #410 #431)
-#747 := (not #745)
-#404 := (>= 0::int 0::int)
-#534 := (or #431 #404)
-#741 := (not #534)
-#415 := (or #741 #747)
-#416 := (ite #415 #427 #428)
-#748 := (= -3::int 0::int)
-#749 := (= 0::int 0::int)
-#750 := (or #749 #748)
-#751 := (ite #750 #176 #416)
-#589 := (or #586 #751)
-#590 := (iff #589 #588)
-#592 := (iff #588 #588)
-#593 := [rewrite]: #592
-#686 := (iff #751 #176)
-#1 := true
-#537 := (ite true #176 #176)
-#682 := (iff #537 #176)
-#683 := [rewrite]: #682
-#685 := (iff #751 #537)
-#577 := (iff #416 #176)
-#692 := (ite false #176 #176)
-#694 := (iff #692 #176)
-#576 := [rewrite]: #694
-#693 := (iff #416 #692)
-#689 := (iff #428 #176)
-#535 := (= #417 #173)
-#719 := (+ #173 0::int)
-#710 := (= #719 #173)
-#705 := [rewrite]: #710
-#704 := (= #417 #719)
-#703 := (= #426 0::int)
-#432 := (= #424 0::int)
-#714 := [rewrite]: #432
-#696 := (= #426 #424)
-#701 := (= #425 0::int)
-#555 := (div 0::int 3::int)
-#699 := (= #555 0::int)
-#700 := [rewrite]: #699
-#556 := (= #425 #555)
-#713 := (= #339 3::int)
-#554 := [rewrite]: #713
-#698 := [monotonicity #714 #554]: #556
-#695 := [trans #698 #700]: #701
-#702 := [monotonicity #695]: #696
-#697 := [trans #702 #714]: #703
-#533 := [monotonicity #697]: #704
-#536 := [trans #533 #705]: #535
-#690 := [monotonicity #536]: #689
-#712 := (iff #427 #176)
-#706 := (= #430 #173)
-#708 := (= #430 #719)
-#718 := (= #357 0::int)
-#716 := (= #357 #424)
-#437 := (= #429 0::int)
-#438 := [rewrite]: #437
-#717 := [monotonicity #438]: #716
-#715 := [trans #717 #714]: #718
-#709 := [monotonicity #715]: #708
-#711 := [trans #709 #705]: #706
-#707 := [monotonicity #711]: #712
-#435 := (iff #415 false)
-#446 := (or false false)
-#721 := (iff #446 false)
-#725 := [rewrite]: #721
-#720 := (iff #415 #446)
-#723 := (iff #747 false)
-#736 := (not true)
-#379 := (iff #736 false)
-#739 := [rewrite]: #379
-#345 := (iff #747 #736)
-#451 := (iff #745 true)
-#731 := (or true true)
-#734 := (iff #731 true)
-#729 := [rewrite]: #734
-#449 := (iff #745 #731)
-#392 := (iff #431 true)
-#393 := [rewrite]: #392
-#726 := (iff #410 true)
-#727 := [rewrite]: #726
-#450 := [monotonicity #727 #393]: #449
-#452 := [trans #450 #729]: #451
-#722 := [monotonicity #452]: #345
-#445 := [trans #722 #739]: #723
-#737 := (iff #741 false)
-#738 := (iff #741 #736)
-#735 := (iff #534 true)
-#732 := (iff #534 #731)
-#394 := (iff #404 true)
-#730 := [rewrite]: #394
-#733 := [monotonicity #393 #730]: #732
-#373 := [trans #733 #729]: #735
-#378 := [monotonicity #373]: #738
-#740 := [trans #378 #739]: #737
-#724 := [monotonicity #740 #445]: #720
-#436 := [trans #724 #725]: #435
-#691 := [monotonicity #436 #707 #690]: #693
-#578 := [trans #691 #576]: #577
-#728 := (iff #750 true)
-#407 := (or true false)
-#744 := (iff #407 true)
-#387 := [rewrite]: #744
-#743 := (iff #750 #407)
-#742 := (iff #748 false)
-#402 := [rewrite]: #742
-#746 := (iff #749 true)
-#752 := [rewrite]: #746
-#403 := [monotonicity #752 #402]: #743
-#391 := [trans #403 #387]: #728
-#587 := [monotonicity #391 #578]: #685
-#571 := [trans #587 #683]: #686
-#591 := [monotonicity #571]: #590
-#594 := [trans #591 #593]: #590
-#582 := [quant-inst]: #589
-#687 := [mp #582 #594]: #588
-[unit-resolution #687 #759 #194]: false
-unsat
-e3ca3cfa3906cd959718a475a745a6d7221ec25b 329 0
+8144b6d093845ae8dee3931af66c9361d99b4453 329 0
 #2 := false
 #11 := 0::int
 decl f4 :: (-> int int int)
@@ -33272,304 +37303,7 @@
 #676 := [mp #592 #674]: #590
 [unit-resolution #676 #765 #194]: false
 unsat
-9c08922893cb633ba38530e84ba0b136e5cd1a81 296 0
-#2 := false
-#11 := 0::int
-decl f3 :: (-> int int int)
-#39 := 3::int
-#38 := 1::int
-#40 := (f3 1::int 3::int)
-#41 := (= #40 0::int)
-#42 := (not #41)
-#169 := [asserted]: #42
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#10 := (f3 #8 #9)
-#737 := (pattern #10)
-#66 := -1::int
-#70 := (* -1::int #9)
-#67 := (* -1::int #8)
-#73 := (div #67 #70)
-#240 := (* -1::int #73)
-#241 := (+ #10 #240)
-#242 := (= #241 0::int)
-#21 := (div #8 #9)
-#237 := (* -1::int #21)
-#238 := (+ #10 #237)
-#239 := (= #238 0::int)
-#92 := (<= #9 0::int)
-#88 := (<= #8 0::int)
-#195 := (or #88 #92)
-#196 := (not #195)
-#99 := (>= #8 0::int)
-#187 := (or #92 #99)
-#188 := (not #187)
-#202 := (or #188 #196)
-#243 := (ite #202 #239 #242)
-#236 := (= #10 0::int)
-#13 := (= #9 0::int)
-#12 := (= #8 0::int)
-#14 := (or #12 #13)
-#244 := (ite #14 #236 #243)
-#738 := (forall (vars (?v0 int) (?v1 int)) (:pat #737) #244)
-#247 := (forall (vars (?v0 int) (?v1 int)) #244)
-#741 := (iff #247 #738)
-#739 := (iff #244 #244)
-#740 := [refl]: #739
-#742 := [quant-intro #740]: #741
-#207 := (ite #202 #21 #73)
-#210 := (ite #14 0::int #207)
-#213 := (= #10 #210)
-#216 := (forall (vars (?v0 int) (?v1 int)) #213)
-#248 := (iff #216 #247)
-#245 := (iff #213 #244)
-#246 := [rewrite]: #245
-#249 := [quant-intro #246]: #248
-#100 := (not #99)
-#93 := (not #92)
-#103 := (and #93 #100)
-#89 := (not #88)
-#96 := (and #89 #93)
-#106 := (or #96 #103)
-#109 := (ite #106 #21 #73)
-#112 := (ite #14 0::int #109)
-#115 := (= #10 #112)
-#118 := (forall (vars (?v0 int) (?v1 int)) #115)
-#217 := (iff #118 #216)
-#214 := (iff #115 #213)
-#211 := (= #112 #210)
-#208 := (= #109 #207)
-#205 := (iff #106 #202)
-#199 := (or #196 #188)
-#203 := (iff #199 #202)
-#204 := [rewrite]: #203
-#200 := (iff #106 #199)
-#197 := (iff #103 #188)
-#198 := [rewrite]: #197
-#185 := (iff #96 #196)
-#186 := [rewrite]: #185
-#201 := [monotonicity #186 #198]: #200
-#206 := [trans #201 #204]: #205
-#209 := [monotonicity #206]: #208
-#212 := [monotonicity #209]: #211
-#215 := [monotonicity #212]: #214
-#218 := [quant-intro #215]: #217
-#191 := (~ #118 #118)
-#189 := (~ #115 #115)
-#190 := [refl]: #189
-#192 := [nnf-pos #190]: #191
-#23 := (- #9)
-#22 := (- #8)
-#24 := (div #22 #23)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#25 := (ite #20 #21 #24)
-#26 := (ite #14 0::int #25)
-#27 := (= #10 #26)
-#28 := (forall (vars (?v0 int) (?v1 int)) #27)
-#121 := (iff #28 #118)
-#60 := (and #16 #18)
-#63 := (or #17 #60)
-#76 := (ite #63 #21 #73)
-#79 := (ite #14 0::int #76)
-#82 := (= #10 #79)
-#85 := (forall (vars (?v0 int) (?v1 int)) #82)
-#119 := (iff #85 #118)
-#116 := (iff #82 #115)
-#113 := (= #79 #112)
-#110 := (= #76 #109)
-#107 := (iff #63 #106)
-#104 := (iff #60 #103)
-#101 := (iff #18 #100)
-#102 := [rewrite]: #101
-#94 := (iff #16 #93)
-#95 := [rewrite]: #94
-#105 := [monotonicity #95 #102]: #104
-#97 := (iff #17 #96)
-#90 := (iff #15 #89)
-#91 := [rewrite]: #90
-#98 := [monotonicity #91 #95]: #97
-#108 := [monotonicity #98 #105]: #107
-#111 := [monotonicity #108]: #110
-#114 := [monotonicity #111]: #113
-#117 := [monotonicity #114]: #116
-#120 := [quant-intro #117]: #119
-#86 := (iff #28 #85)
-#83 := (iff #27 #82)
-#80 := (= #26 #79)
-#77 := (= #25 #76)
-#74 := (= #24 #73)
-#71 := (= #23 #70)
-#72 := [rewrite]: #71
-#68 := (= #22 #67)
-#69 := [rewrite]: #68
-#75 := [monotonicity #69 #72]: #74
-#64 := (iff #20 #63)
-#61 := (iff #19 #60)
-#62 := [rewrite]: #61
-#65 := [monotonicity #62]: #64
-#78 := [monotonicity #65 #75]: #77
-#81 := [monotonicity #78]: #80
-#84 := [monotonicity #81]: #83
-#87 := [quant-intro #84]: #86
-#122 := [trans #87 #120]: #121
-#59 := [asserted]: #28
-#123 := [mp #59 #122]: #118
-#182 := [mp~ #123 #192]: #118
-#219 := [mp #182 #218]: #216
-#250 := [mp #219 #249]: #247
-#743 := [mp #250 #742]: #738
-#650 := (not #738)
-#652 := (or #650 #41)
-#324 := (* -1::int 3::int)
-#408 := (* -1::int 1::int)
-#409 := (div #408 #324)
-#410 := (* -1::int #409)
-#401 := (+ #40 #410)
-#412 := (= #401 0::int)
-#413 := (div 1::int 3::int)
-#414 := (* -1::int #413)
-#411 := (+ #40 #414)
-#415 := (= #411 0::int)
-#394 := (<= 3::int 0::int)
-#729 := (<= 1::int 0::int)
-#731 := (or #729 #394)
-#388 := (not #731)
-#518 := (>= 1::int 0::int)
-#725 := (or #394 #518)
-#399 := (not #725)
-#400 := (or #399 #388)
-#732 := (ite #400 #415 #412)
-#733 := (= 3::int 0::int)
-#734 := (= 1::int 0::int)
-#735 := (or #734 #733)
-#730 := (ite #735 #41 #732)
-#662 := (or #650 #730)
-#664 := (iff #662 #652)
-#665 := (iff #652 #652)
-#654 := [rewrite]: #665
-#661 := (iff #730 #41)
-#578 := (ite false #41 #41)
-#672 := (iff #578 #41)
-#659 := [rewrite]: #672
-#671 := (iff #730 #578)
-#576 := (iff #732 #41)
-#666 := (= #40 1::int)
-#1 := true
-#572 := (ite true #41 #666)
-#574 := (iff #572 #41)
-#575 := [rewrite]: #574
-#573 := (iff #732 #572)
-#555 := (iff #412 #666)
-#675 := (+ -1::int #40)
-#521 := (= #675 0::int)
-#667 := (iff #521 #666)
-#670 := [rewrite]: #667
-#669 := (iff #412 #521)
-#561 := (= #401 #675)
-#674 := (+ #40 -1::int)
-#678 := (= #674 #675)
-#560 := [rewrite]: #678
-#676 := (= #401 #674)
-#520 := (= #410 -1::int)
-#539 := (= #408 -1::int)
-#540 := [rewrite]: #539
-#517 := (= #410 #408)
-#681 := (= #409 1::int)
-#682 := -3::int
-#685 := (div -1::int -3::int)
-#686 := (= #685 1::int)
-#687 := [rewrite]: #686
-#679 := (= #409 #685)
-#683 := (= #324 -3::int)
-#684 := [rewrite]: #683
-#680 := [monotonicity #540 #684]: #679
-#688 := [trans #680 #687]: #681
-#519 := [monotonicity #688]: #517
-#673 := [trans #519 #540]: #520
-#677 := [monotonicity #673]: #676
-#562 := [trans #677 #560]: #561
-#571 := [monotonicity #562]: #669
-#570 := [trans #571 #670]: #555
-#697 := (iff #415 #41)
-#696 := (= #411 #40)
-#693 := (+ #40 0::int)
-#690 := (= #693 #40)
-#695 := [rewrite]: #690
-#694 := (= #411 #693)
-#703 := (= #414 0::int)
-#701 := (* -1::int 0::int)
-#702 := (= #701 0::int)
-#699 := [rewrite]: #702
-#416 := (= #414 #701)
-#422 := (= #413 0::int)
-#700 := [rewrite]: #422
-#698 := [monotonicity #700]: #416
-#692 := [trans #698 #699]: #703
-#689 := [monotonicity #692]: #694
-#691 := [trans #689 #695]: #696
-#538 := [monotonicity #691]: #697
-#420 := (iff #400 true)
-#716 := (or false true)
-#713 := (iff #716 true)
-#719 := [rewrite]: #713
-#709 := (iff #400 #716)
-#708 := (iff #388 true)
-#706 := (not false)
-#430 := (iff #706 true)
-#704 := [rewrite]: #430
-#707 := (iff #388 #706)
-#436 := (iff #731 false)
-#727 := (or false false)
-#371 := (iff #727 false)
-#712 := [rewrite]: #371
-#434 := (iff #731 #727)
-#377 := (iff #394 false)
-#378 := [rewrite]: #377
-#711 := (iff #729 false)
-#433 := [rewrite]: #711
-#435 := [monotonicity #433 #378]: #434
-#330 := [trans #435 #712]: #436
-#429 := [monotonicity #330]: #707
-#705 := [trans #429 #704]: #708
-#724 := (iff #399 false)
-#722 := (not true)
-#723 := (iff #722 false)
-#721 := [rewrite]: #723
-#362 := (iff #399 #722)
-#357 := (iff #725 true)
-#717 := (iff #725 #716)
-#714 := (iff #518 true)
-#715 := [rewrite]: #714
-#718 := [monotonicity #378 #715]: #717
-#720 := [trans #718 #719]: #357
-#363 := [monotonicity #720]: #362
-#710 := [trans #363 #721]: #724
-#419 := [monotonicity #710 #705]: #709
-#421 := [trans #419 #719]: #420
-#566 := [monotonicity #421 #538 #570]: #573
-#577 := [trans #566 #575]: #576
-#375 := (iff #735 false)
-#387 := (iff #735 #727)
-#386 := (iff #733 false)
-#391 := [rewrite]: #386
-#736 := (iff #734 false)
-#726 := [rewrite]: #736
-#728 := [monotonicity #726 #391]: #387
-#376 := [trans #728 #712]: #375
-#668 := [monotonicity #376 #577]: #671
-#653 := [trans #668 #659]: #661
-#660 := [monotonicity #653]: #664
-#655 := [trans #660 #654]: #664
-#663 := [quant-inst]: #662
-#656 := [mp #663 #655]: #652
-[unit-resolution #656 #743 #169]: false
-unsat
-296b085c95c2238dbf4db3cb90fb09e839ecea3d 306 0
+a20f91759967cf78b23640f832ca57a185975634 306 0
 #2 := false
 #11 := 0::int
 decl f4 :: (-> int int int)
@@ -33876,294 +37610,7 @@
 #693 := [mp #535 #690]: #532
 [unit-resolution #693 #765 #194]: false
 unsat
-4b6a61d48a9a0af02d6fd3725ddd2c3187d77b9e 286 0
-#2 := false
-#40 := 1::int
-decl f3 :: (-> int int int)
-#38 := 3::int
-#39 := (f3 3::int 3::int)
-#41 := (= #39 1::int)
-#42 := (not #41)
-#169 := [asserted]: #42
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#10 := (f3 #8 #9)
-#727 := (pattern #10)
-#11 := 0::int
-#66 := -1::int
-#70 := (* -1::int #9)
-#67 := (* -1::int #8)
-#73 := (div #67 #70)
-#230 := (* -1::int #73)
-#231 := (+ #10 #230)
-#232 := (= #231 0::int)
-#21 := (div #8 #9)
-#227 := (* -1::int #21)
-#228 := (+ #10 #227)
-#229 := (= #228 0::int)
-#92 := (<= #9 0::int)
-#88 := (<= #8 0::int)
-#185 := (or #88 #92)
-#186 := (not #185)
-#99 := (>= #8 0::int)
-#177 := (or #92 #99)
-#178 := (not #177)
-#192 := (or #178 #186)
-#233 := (ite #192 #229 #232)
-#226 := (= #10 0::int)
-#13 := (= #9 0::int)
-#12 := (= #8 0::int)
-#14 := (or #12 #13)
-#234 := (ite #14 #226 #233)
-#728 := (forall (vars (?v0 int) (?v1 int)) (:pat #727) #234)
-#237 := (forall (vars (?v0 int) (?v1 int)) #234)
-#731 := (iff #237 #728)
-#729 := (iff #234 #234)
-#730 := [refl]: #729
-#732 := [quant-intro #730]: #731
-#197 := (ite #192 #21 #73)
-#200 := (ite #14 0::int #197)
-#203 := (= #10 #200)
-#206 := (forall (vars (?v0 int) (?v1 int)) #203)
-#238 := (iff #206 #237)
-#235 := (iff #203 #234)
-#236 := [rewrite]: #235
-#239 := [quant-intro #236]: #238
-#100 := (not #99)
-#93 := (not #92)
-#103 := (and #93 #100)
-#89 := (not #88)
-#96 := (and #89 #93)
-#106 := (or #96 #103)
-#109 := (ite #106 #21 #73)
-#112 := (ite #14 0::int #109)
-#115 := (= #10 #112)
-#118 := (forall (vars (?v0 int) (?v1 int)) #115)
-#207 := (iff #118 #206)
-#204 := (iff #115 #203)
-#201 := (= #112 #200)
-#198 := (= #109 #197)
-#195 := (iff #106 #192)
-#189 := (or #186 #178)
-#193 := (iff #189 #192)
-#194 := [rewrite]: #193
-#190 := (iff #106 #189)
-#187 := (iff #103 #178)
-#188 := [rewrite]: #187
-#175 := (iff #96 #186)
-#176 := [rewrite]: #175
-#191 := [monotonicity #176 #188]: #190
-#196 := [trans #191 #194]: #195
-#199 := [monotonicity #196]: #198
-#202 := [monotonicity #199]: #201
-#205 := [monotonicity #202]: #204
-#208 := [quant-intro #205]: #207
-#181 := (~ #118 #118)
-#179 := (~ #115 #115)
-#180 := [refl]: #179
-#182 := [nnf-pos #180]: #181
-#23 := (- #9)
-#22 := (- #8)
-#24 := (div #22 #23)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#25 := (ite #20 #21 #24)
-#26 := (ite #14 0::int #25)
-#27 := (= #10 #26)
-#28 := (forall (vars (?v0 int) (?v1 int)) #27)
-#121 := (iff #28 #118)
-#60 := (and #16 #18)
-#63 := (or #17 #60)
-#76 := (ite #63 #21 #73)
-#79 := (ite #14 0::int #76)
-#82 := (= #10 #79)
-#85 := (forall (vars (?v0 int) (?v1 int)) #82)
-#119 := (iff #85 #118)
-#116 := (iff #82 #115)
-#113 := (= #79 #112)
-#110 := (= #76 #109)
-#107 := (iff #63 #106)
-#104 := (iff #60 #103)
-#101 := (iff #18 #100)
-#102 := [rewrite]: #101
-#94 := (iff #16 #93)
-#95 := [rewrite]: #94
-#105 := [monotonicity #95 #102]: #104
-#97 := (iff #17 #96)
-#90 := (iff #15 #89)
-#91 := [rewrite]: #90
-#98 := [monotonicity #91 #95]: #97
-#108 := [monotonicity #98 #105]: #107
-#111 := [monotonicity #108]: #110
-#114 := [monotonicity #111]: #113
-#117 := [monotonicity #114]: #116
-#120 := [quant-intro #117]: #119
-#86 := (iff #28 #85)
-#83 := (iff #27 #82)
-#80 := (= #26 #79)
-#77 := (= #25 #76)
-#74 := (= #24 #73)
-#71 := (= #23 #70)
-#72 := [rewrite]: #71
-#68 := (= #22 #67)
-#69 := [rewrite]: #68
-#75 := [monotonicity #69 #72]: #74
-#64 := (iff #20 #63)
-#61 := (iff #19 #60)
-#62 := [rewrite]: #61
-#65 := [monotonicity #62]: #64
-#78 := [monotonicity #65 #75]: #77
-#81 := [monotonicity #78]: #80
-#84 := [monotonicity #81]: #83
-#87 := [quant-intro #84]: #86
-#122 := [trans #87 #120]: #121
-#59 := [asserted]: #28
-#123 := [mp #59 #122]: #118
-#172 := [mp~ #123 #182]: #118
-#209 := [mp #172 #208]: #206
-#240 := [mp #209 #239]: #237
-#733 := [mp #240 #732]: #728
-#565 := (not #728)
-#566 := (or #565 #41)
-#314 := (* -1::int 3::int)
-#398 := (div #314 #314)
-#399 := (* -1::int #398)
-#400 := (+ #39 #399)
-#391 := (= #400 0::int)
-#402 := (div 3::int 3::int)
-#403 := (* -1::int #402)
-#404 := (+ #39 #403)
-#401 := (= #404 0::int)
-#405 := (<= 3::int 0::int)
-#384 := (or #405 #405)
-#719 := (not #384)
-#721 := (>= 3::int 0::int)
-#378 := (or #405 #721)
-#508 := (not #378)
-#715 := (or #508 #719)
-#389 := (ite #715 #401 #391)
-#390 := (= #39 0::int)
-#722 := (= 3::int 0::int)
-#723 := (or #722 #722)
-#724 := (ite #723 #390 #389)
-#567 := (or #565 #724)
-#661 := (iff #567 #566)
-#662 := (iff #566 #566)
-#649 := [rewrite]: #662
-#556 := (iff #724 #41)
-#660 := (ite false #390 #41)
-#562 := (iff #660 #41)
-#563 := [rewrite]: #562
-#545 := (iff #724 #660)
-#656 := (iff #389 #41)
-#1 := true
-#551 := (ite true #41 #41)
-#659 := (iff #551 #41)
-#561 := [rewrite]: #659
-#552 := (iff #389 #551)
-#668 := (iff #391 #41)
-#689 := (+ -1::int #39)
-#679 := (= #689 0::int)
-#686 := (iff #679 #41)
-#681 := [rewrite]: #686
-#667 := (iff #391 #679)
-#664 := (= #400 #689)
-#406 := (+ #39 -1::int)
-#693 := (= #406 #689)
-#682 := [rewrite]: #693
-#510 := (= #400 #406)
-#507 := (= #399 -1::int)
-#699 := (* -1::int 1::int)
-#411 := (= #699 -1::int)
-#412 := [rewrite]: #411
-#671 := (= #399 #699)
-#676 := (= #398 1::int)
-#529 := -3::int
-#673 := (div -3::int -3::int)
-#669 := (= #673 1::int)
-#670 := [rewrite]: #669
-#674 := (= #398 #673)
-#530 := (= #314 -3::int)
-#672 := [rewrite]: #530
-#675 := [monotonicity #672 #672]: #674
-#677 := [trans #675 #670]: #676
-#678 := [monotonicity #677]: #671
-#509 := [trans #678 #412]: #507
-#663 := [monotonicity #509]: #510
-#666 := [trans #663 #682]: #664
-#665 := [monotonicity #666]: #667
-#550 := [trans #665 #681]: #668
-#687 := (iff #401 #41)
-#680 := (iff #401 #679)
-#683 := (= #404 #689)
-#688 := (= #404 #406)
-#690 := (= #403 -1::int)
-#409 := (= #403 #699)
-#698 := (= #402 1::int)
-#695 := [rewrite]: #698
-#410 := [monotonicity #695]: #409
-#691 := [trans #410 #412]: #690
-#692 := [monotonicity #691]: #688
-#684 := [trans #692 #682]: #683
-#685 := [monotonicity #684]: #680
-#528 := [trans #685 #681]: #687
-#420 := (iff #715 true)
-#367 := (or false true)
-#705 := (iff #367 true)
-#706 := [rewrite]: #705
-#697 := (iff #715 #367)
-#320 := (iff #719 true)
-#701 := (not false)
-#425 := (iff #701 true)
-#426 := [rewrite]: #425
-#423 := (iff #719 #701)
-#714 := (iff #384 false)
-#726 := (or false false)
-#381 := (iff #726 false)
-#717 := [rewrite]: #381
-#713 := (iff #384 #726)
-#361 := (iff #405 false)
-#702 := [rewrite]: #361
-#711 := [monotonicity #702 #702]: #713
-#700 := [trans #711 #717]: #714
-#424 := [monotonicity #700]: #423
-#696 := [trans #424 #426]: #320
-#352 := (iff #508 false)
-#703 := (not true)
-#710 := (iff #703 false)
-#712 := [rewrite]: #710
-#709 := (iff #508 #703)
-#707 := (iff #378 true)
-#368 := (iff #378 #367)
-#365 := (iff #721 true)
-#366 := [rewrite]: #365
-#704 := [monotonicity #702 #366]: #368
-#708 := [trans #704 #706]: #707
-#347 := [monotonicity #708]: #709
-#353 := [trans #347 #712]: #352
-#419 := [monotonicity #353 #696]: #697
-#694 := [trans #419 #706]: #420
-#511 := [monotonicity #694 #528 #550]: #552
-#657 := [trans #511 #561]: #656
-#377 := (iff #723 false)
-#716 := (iff #723 #726)
-#725 := (iff #722 false)
-#720 := [rewrite]: #725
-#376 := [monotonicity #720 #720]: #716
-#718 := [trans #376 #717]: #377
-#560 := [monotonicity #718 #657]: #545
-#564 := [trans #560 #563]: #556
-#658 := [monotonicity #564]: #661
-#651 := [trans #658 #649]: #661
-#568 := [quant-inst]: #567
-#643 := [mp #568 #651]: #566
-[unit-resolution #643 #733 #169]: false
-unsat
-62ada4366d569f6c75cf58cfc0c1710258d4c728 299 0
+450a206c661160bbd83de0094374a847857dd804 299 0
 #2 := false
 #11 := 0::int
 decl f4 :: (-> int int int)
@@ -34463,318 +37910,7 @@
 #575 := [mp #555 #574]: #667
 [unit-resolution #575 #750 #168]: false
 unsat
-a30bbe93d2d8e059cff59983b5c369152f8b4c56 310 0
-#2 := false
-#41 := 1::int
-decl f3 :: (-> int int int)
-#39 := 3::int
-#38 := 5::int
-#40 := (f3 5::int 3::int)
-#42 := (= #40 1::int)
-#43 := (not #42)
-#170 := [asserted]: #43
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#10 := (f3 #8 #9)
-#728 := (pattern #10)
-#11 := 0::int
-#67 := -1::int
-#71 := (* -1::int #9)
-#68 := (* -1::int #8)
-#74 := (div #68 #71)
-#231 := (* -1::int #74)
-#232 := (+ #10 #231)
-#233 := (= #232 0::int)
-#21 := (div #8 #9)
-#228 := (* -1::int #21)
-#229 := (+ #10 #228)
-#230 := (= #229 0::int)
-#93 := (<= #9 0::int)
-#89 := (<= #8 0::int)
-#186 := (or #89 #93)
-#187 := (not #186)
-#100 := (>= #8 0::int)
-#178 := (or #93 #100)
-#179 := (not #178)
-#193 := (or #179 #187)
-#234 := (ite #193 #230 #233)
-#227 := (= #10 0::int)
-#13 := (= #9 0::int)
-#12 := (= #8 0::int)
-#14 := (or #12 #13)
-#235 := (ite #14 #227 #234)
-#729 := (forall (vars (?v0 int) (?v1 int)) (:pat #728) #235)
-#238 := (forall (vars (?v0 int) (?v1 int)) #235)
-#732 := (iff #238 #729)
-#730 := (iff #235 #235)
-#731 := [refl]: #730
-#733 := [quant-intro #731]: #732
-#198 := (ite #193 #21 #74)
-#201 := (ite #14 0::int #198)
-#204 := (= #10 #201)
-#207 := (forall (vars (?v0 int) (?v1 int)) #204)
-#239 := (iff #207 #238)
-#236 := (iff #204 #235)
-#237 := [rewrite]: #236
-#240 := [quant-intro #237]: #239
-#101 := (not #100)
-#94 := (not #93)
-#104 := (and #94 #101)
-#90 := (not #89)
-#97 := (and #90 #94)
-#107 := (or #97 #104)
-#110 := (ite #107 #21 #74)
-#113 := (ite #14 0::int #110)
-#116 := (= #10 #113)
-#119 := (forall (vars (?v0 int) (?v1 int)) #116)
-#208 := (iff #119 #207)
-#205 := (iff #116 #204)
-#202 := (= #113 #201)
-#199 := (= #110 #198)
-#196 := (iff #107 #193)
-#190 := (or #187 #179)
-#194 := (iff #190 #193)
-#195 := [rewrite]: #194
-#191 := (iff #107 #190)
-#188 := (iff #104 #179)
-#189 := [rewrite]: #188
-#176 := (iff #97 #187)
-#177 := [rewrite]: #176
-#192 := [monotonicity #177 #189]: #191
-#197 := [trans #192 #195]: #196
-#200 := [monotonicity #197]: #199
-#203 := [monotonicity #200]: #202
-#206 := [monotonicity #203]: #205
-#209 := [quant-intro #206]: #208
-#182 := (~ #119 #119)
-#180 := (~ #116 #116)
-#181 := [refl]: #180
-#183 := [nnf-pos #181]: #182
-#23 := (- #9)
-#22 := (- #8)
-#24 := (div #22 #23)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#25 := (ite #20 #21 #24)
-#26 := (ite #14 0::int #25)
-#27 := (= #10 #26)
-#28 := (forall (vars (?v0 int) (?v1 int)) #27)
-#122 := (iff #28 #119)
-#61 := (and #16 #18)
-#64 := (or #17 #61)
-#77 := (ite #64 #21 #74)
-#80 := (ite #14 0::int #77)
-#83 := (= #10 #80)
-#86 := (forall (vars (?v0 int) (?v1 int)) #83)
-#120 := (iff #86 #119)
-#117 := (iff #83 #116)
-#114 := (= #80 #113)
-#111 := (= #77 #110)
-#108 := (iff #64 #107)
-#105 := (iff #61 #104)
-#102 := (iff #18 #101)
-#103 := [rewrite]: #102
-#95 := (iff #16 #94)
-#96 := [rewrite]: #95
-#106 := [monotonicity #96 #103]: #105
-#98 := (iff #17 #97)
-#91 := (iff #15 #90)
-#92 := [rewrite]: #91
-#99 := [monotonicity #92 #96]: #98
-#109 := [monotonicity #99 #106]: #108
-#112 := [monotonicity #109]: #111
-#115 := [monotonicity #112]: #114
-#118 := [monotonicity #115]: #117
-#121 := [quant-intro #118]: #120
-#87 := (iff #28 #86)
-#84 := (iff #27 #83)
-#81 := (= #26 #80)
-#78 := (= #25 #77)
-#75 := (= #24 #74)
-#72 := (= #23 #71)
-#73 := [rewrite]: #72
-#69 := (= #22 #68)
-#70 := [rewrite]: #69
-#76 := [monotonicity #70 #73]: #75
-#65 := (iff #20 #64)
-#62 := (iff #19 #61)
-#63 := [rewrite]: #62
-#66 := [monotonicity #63]: #65
-#79 := [monotonicity #66 #76]: #78
-#82 := [monotonicity #79]: #81
-#85 := [monotonicity #82]: #84
-#88 := [quant-intro #85]: #87
-#123 := [trans #88 #121]: #122
-#60 := [asserted]: #28
-#124 := [mp #60 #123]: #119
-#173 := [mp~ #124 #183]: #119
-#210 := [mp #173 #209]: #207
-#241 := [mp #210 #240]: #238
-#734 := [mp #241 #733]: #729
-#633 := (not #729)
-#634 := (or #633 #42)
-#315 := (* -1::int 3::int)
-#399 := (* -1::int 5::int)
-#400 := (div #399 #315)
-#401 := (* -1::int #400)
-#392 := (+ #40 #401)
-#403 := (= #392 0::int)
-#404 := (div 5::int 3::int)
-#405 := (* -1::int #404)
-#402 := (+ #40 #405)
-#406 := (= #402 0::int)
-#385 := (<= 3::int 0::int)
-#720 := (<= 5::int 0::int)
-#722 := (or #720 #385)
-#379 := (not #722)
-#509 := (>= 5::int 0::int)
-#716 := (or #385 #509)
-#390 := (not #716)
-#391 := (or #390 #379)
-#723 := (ite #391 #406 #403)
-#724 := (= #40 0::int)
-#725 := (= 3::int 0::int)
-#726 := (= 5::int 0::int)
-#721 := (or #726 #725)
-#727 := (ite #721 #724 #723)
-#636 := (or #633 #727)
-#638 := (iff #636 #634)
-#635 := (iff #634 #634)
-#640 := [rewrite]: #635
-#642 := (iff #727 #42)
-#656 := (ite false #724 #42)
-#647 := (iff #656 #42)
-#648 := [rewrite]: #647
-#645 := (iff #727 #656)
-#655 := (iff #723 #42)
-#665 := 2::int
-#662 := (= #40 2::int)
-#1 := true
-#644 := (ite true #42 #662)
-#653 := (iff #644 #42)
-#654 := [rewrite]: #653
-#641 := (iff #723 #644)
-#650 := (iff #403 #662)
-#512 := -2::int
-#563 := (+ -2::int #40)
-#567 := (= #563 0::int)
-#659 := (iff #567 #662)
-#663 := [rewrite]: #659
-#568 := (iff #403 #567)
-#565 := (= #392 #563)
-#661 := (+ #40 -2::int)
-#564 := (= #661 #563)
-#557 := [rewrite]: #564
-#546 := (= #392 #661)
-#657 := (= #401 -2::int)
-#551 := (* -1::int 2::int)
-#660 := (= #551 -2::int)
-#562 := [rewrite]: #660
-#552 := (= #401 #551)
-#666 := (= #400 2::int)
-#672 := -3::int
-#671 := -5::int
-#510 := (div -5::int -3::int)
-#667 := (= #510 2::int)
-#668 := [rewrite]: #667
-#511 := (= #400 #510)
-#679 := (= #315 -3::int)
-#508 := [rewrite]: #679
-#677 := (= #399 -5::int)
-#678 := [rewrite]: #677
-#664 := [monotonicity #678 #508]: #511
-#669 := [trans #664 #668]: #666
-#553 := [monotonicity #669]: #552
-#658 := [trans #553 #562]: #657
-#561 := [monotonicity #658]: #546
-#566 := [trans #561 #557]: #565
-#569 := [monotonicity #566]: #568
-#652 := [trans #569 #663]: #650
-#676 := (iff #406 #42)
-#686 := (+ -1::int #40)
-#530 := (= #686 0::int)
-#674 := (iff #530 #42)
-#675 := [rewrite]: #674
-#531 := (iff #406 #530)
-#688 := (= #402 #686)
-#685 := (+ #40 -1::int)
-#687 := (= #685 #686)
-#682 := [rewrite]: #687
-#680 := (= #402 #685)
-#683 := (= #405 -1::int)
-#407 := (* -1::int 1::int)
-#690 := (= #407 -1::int)
-#694 := [rewrite]: #690
-#689 := (= #405 #407)
-#691 := (= #404 1::int)
-#692 := [rewrite]: #691
-#693 := [monotonicity #692]: #689
-#684 := [trans #693 #694]: #683
-#681 := [monotonicity #684]: #680
-#529 := [trans #681 #682]: #688
-#673 := [monotonicity #529]: #531
-#670 := [trans #673 #675]: #676
-#412 := (iff #391 true)
-#708 := (or false true)
-#710 := (iff #708 true)
-#348 := [rewrite]: #710
-#410 := (iff #391 #708)
-#696 := (iff #379 true)
-#698 := (not false)
-#695 := (iff #698 true)
-#699 := [rewrite]: #695
-#420 := (iff #379 #698)
-#321 := (iff #722 false)
-#378 := (or false false)
-#703 := (iff #378 false)
-#366 := [rewrite]: #703
-#426 := (iff #722 #378)
-#369 := (iff #385 false)
-#705 := [rewrite]: #369
-#424 := (iff #720 false)
-#425 := [rewrite]: #424
-#427 := [monotonicity #425 #705]: #426
-#697 := [trans #427 #366]: #321
-#421 := [monotonicity #697]: #420
-#700 := [trans #421 #699]: #696
-#701 := (iff #390 false)
-#353 := (not true)
-#712 := (iff #353 false)
-#715 := [rewrite]: #712
-#354 := (iff #390 #353)
-#711 := (iff #716 true)
-#709 := (iff #716 #708)
-#706 := (iff #509 true)
-#707 := [rewrite]: #706
-#704 := [monotonicity #705 #707]: #709
-#713 := [trans #704 #348]: #711
-#714 := [monotonicity #713]: #354
-#702 := [trans #714 #715]: #701
-#411 := [monotonicity #702 #700]: #410
-#413 := [trans #411 #348]: #412
-#643 := [monotonicity #413 #670 #652]: #641
-#651 := [trans #643 #654]: #655
-#367 := (iff #721 false)
-#719 := (iff #721 #378)
-#382 := (iff #725 false)
-#718 := [rewrite]: #382
-#717 := (iff #726 false)
-#377 := [rewrite]: #717
-#362 := [monotonicity #377 #718]: #719
-#368 := [trans #362 #366]: #367
-#646 := [monotonicity #368 #651]: #645
-#649 := [trans #646 #648]: #642
-#639 := [monotonicity #649]: #638
-#626 := [trans #639 #640]: #638
-#637 := [quant-inst]: #636
-#627 := [mp #637 #626]: #634
-[unit-resolution #627 #734 #170]: false
-unsat
-ce00a9a2ac98ee74eda5a6ff3af7cf5e8e01ed5f 318 0
+b8be49aa798ad99c4b186f0fa5e2759eac880cd8 318 0
 #2 := false
 #11 := 0::int
 decl f4 :: (-> int int int)
@@ -35093,656 +38229,7 @@
 #588 := [mp #587 #586]: #537
 [unit-resolution #588 #766 #194]: false
 unsat
-67d1507c84bf5b55736763e84724925e9337b87e 324 0
-#2 := false
-#68 := -1::int
-decl f3 :: (-> int int int)
-#172 := -3::int
-#38 := 1::int
-#175 := (f3 1::int -3::int)
-#180 := (= #175 -1::int)
-#193 := (not #180)
-#42 := (- 1::int)
-#39 := 3::int
-#40 := (- 3::int)
-#41 := (f3 1::int #40)
-#43 := (= #41 #42)
-#44 := (not #43)
-#196 := (iff #44 #193)
-#183 := (= -1::int #175)
-#188 := (not #183)
-#194 := (iff #188 #193)
-#191 := (iff #183 #180)
-#192 := [rewrite]: #191
-#195 := [monotonicity #192]: #194
-#189 := (iff #44 #188)
-#186 := (iff #43 #183)
-#184 := (iff #180 #183)
-#185 := [rewrite]: #184
-#181 := (iff #43 #180)
-#178 := (= #42 -1::int)
-#179 := [rewrite]: #178
-#176 := (= #41 #175)
-#173 := (= #40 -3::int)
-#174 := [rewrite]: #173
-#177 := [monotonicity #174]: #176
-#182 := [monotonicity #177 #179]: #181
-#187 := [trans #182 #185]: #186
-#190 := [monotonicity #187]: #189
-#197 := [trans #190 #195]: #196
-#171 := [asserted]: #44
-#198 := [mp #171 #197]: #193
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#10 := (f3 #8 #9)
-#756 := (pattern #10)
-#11 := 0::int
-#72 := (* -1::int #9)
-#69 := (* -1::int #8)
-#75 := (div #69 #72)
-#259 := (* -1::int #75)
-#260 := (+ #10 #259)
-#261 := (= #260 0::int)
-#21 := (div #8 #9)
-#256 := (* -1::int #21)
-#257 := (+ #10 #256)
-#258 := (= #257 0::int)
-#94 := (<= #9 0::int)
-#90 := (<= #8 0::int)
-#214 := (or #90 #94)
-#215 := (not #214)
-#101 := (>= #8 0::int)
-#206 := (or #94 #101)
-#207 := (not #206)
-#221 := (or #207 #215)
-#262 := (ite #221 #258 #261)
-#255 := (= #10 0::int)
-#13 := (= #9 0::int)
-#12 := (= #8 0::int)
-#14 := (or #12 #13)
-#263 := (ite #14 #255 #262)
-#757 := (forall (vars (?v0 int) (?v1 int)) (:pat #756) #263)
-#266 := (forall (vars (?v0 int) (?v1 int)) #263)
-#760 := (iff #266 #757)
-#758 := (iff #263 #263)
-#759 := [refl]: #758
-#761 := [quant-intro #759]: #760
-#226 := (ite #221 #21 #75)
-#229 := (ite #14 0::int #226)
-#232 := (= #10 #229)
-#235 := (forall (vars (?v0 int) (?v1 int)) #232)
-#267 := (iff #235 #266)
-#264 := (iff #232 #263)
-#265 := [rewrite]: #264
-#268 := [quant-intro #265]: #267
-#102 := (not #101)
-#95 := (not #94)
-#105 := (and #95 #102)
-#91 := (not #90)
-#98 := (and #91 #95)
-#108 := (or #98 #105)
-#111 := (ite #108 #21 #75)
-#114 := (ite #14 0::int #111)
-#117 := (= #10 #114)
-#120 := (forall (vars (?v0 int) (?v1 int)) #117)
-#236 := (iff #120 #235)
-#233 := (iff #117 #232)
-#230 := (= #114 #229)
-#227 := (= #111 #226)
-#224 := (iff #108 #221)
-#218 := (or #215 #207)
-#222 := (iff #218 #221)
-#223 := [rewrite]: #222
-#219 := (iff #108 #218)
-#216 := (iff #105 #207)
-#217 := [rewrite]: #216
-#204 := (iff #98 #215)
-#205 := [rewrite]: #204
-#220 := [monotonicity #205 #217]: #219
-#225 := [trans #220 #223]: #224
-#228 := [monotonicity #225]: #227
-#231 := [monotonicity #228]: #230
-#234 := [monotonicity #231]: #233
-#237 := [quant-intro #234]: #236
-#210 := (~ #120 #120)
-#208 := (~ #117 #117)
-#209 := [refl]: #208
-#211 := [nnf-pos #209]: #210
-#23 := (- #9)
-#22 := (- #8)
-#24 := (div #22 #23)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#25 := (ite #20 #21 #24)
-#26 := (ite #14 0::int #25)
-#27 := (= #10 #26)
-#28 := (forall (vars (?v0 int) (?v1 int)) #27)
-#123 := (iff #28 #120)
-#62 := (and #16 #18)
-#65 := (or #17 #62)
-#78 := (ite #65 #21 #75)
-#81 := (ite #14 0::int #78)
-#84 := (= #10 #81)
-#87 := (forall (vars (?v0 int) (?v1 int)) #84)
-#121 := (iff #87 #120)
-#118 := (iff #84 #117)
-#115 := (= #81 #114)
-#112 := (= #78 #111)
-#109 := (iff #65 #108)
-#106 := (iff #62 #105)
-#103 := (iff #18 #102)
-#104 := [rewrite]: #103
-#96 := (iff #16 #95)
-#97 := [rewrite]: #96
-#107 := [monotonicity #97 #104]: #106
-#99 := (iff #17 #98)
-#92 := (iff #15 #91)
-#93 := [rewrite]: #92
-#100 := [monotonicity #93 #97]: #99
-#110 := [monotonicity #100 #107]: #109
-#113 := [monotonicity #110]: #112
-#116 := [monotonicity #113]: #115
-#119 := [monotonicity #116]: #118
-#122 := [quant-intro #119]: #121
-#88 := (iff #28 #87)
-#85 := (iff #27 #84)
-#82 := (= #26 #81)
-#79 := (= #25 #78)
-#76 := (= #24 #75)
-#73 := (= #23 #72)
-#74 := [rewrite]: #73
-#70 := (= #22 #69)
-#71 := [rewrite]: #70
-#77 := [monotonicity #71 #74]: #76
-#66 := (iff #20 #65)
-#63 := (iff #19 #62)
-#64 := [rewrite]: #63
-#67 := [monotonicity #64]: #66
-#80 := [monotonicity #67 #77]: #79
-#83 := [monotonicity #80]: #82
-#86 := [monotonicity #83]: #85
-#89 := [quant-intro #86]: #88
-#124 := [trans #89 #122]: #123
-#61 := [asserted]: #28
-#125 := [mp #61 #124]: #120
-#201 := [mp~ #125 #211]: #120
-#238 := [mp #201 #237]: #235
-#269 := [mp #238 #268]: #266
-#762 := [mp #269 #761]: #757
-#672 := (not #757)
-#669 := (or #672 #180)
-#343 := (* -1::int -3::int)
-#427 := (* -1::int 1::int)
-#428 := (div #427 #343)
-#429 := (* -1::int #428)
-#420 := (+ #175 #429)
-#431 := (= #420 0::int)
-#432 := (div 1::int -3::int)
-#433 := (* -1::int #432)
-#430 := (+ #175 #433)
-#434 := (= #430 0::int)
-#413 := (<= -3::int 0::int)
-#748 := (<= 1::int 0::int)
-#750 := (or #748 #413)
-#407 := (not #750)
-#537 := (>= 1::int 0::int)
-#744 := (or #413 #537)
-#418 := (not #744)
-#419 := (or #418 #407)
-#751 := (ite #419 #434 #431)
-#752 := (= #175 0::int)
-#753 := (= -3::int 0::int)
-#754 := (= 1::int 0::int)
-#749 := (or #754 #753)
-#755 := (ite #749 #752 #751)
-#671 := (or #672 #755)
-#682 := (iff #671 #669)
-#679 := (iff #669 #669)
-#684 := [rewrite]: #679
-#678 := (iff #755 #180)
-#585 := (ite false #752 #180)
-#595 := (iff #585 #180)
-#596 := [rewrite]: #595
-#687 := (iff #755 #585)
-#597 := (iff #751 #180)
-#593 := (iff #751 #585)
-#591 := (iff #431 #180)
-#580 := (+ 1::int #175)
-#685 := (= #580 0::int)
-#574 := (iff #685 #180)
-#589 := [rewrite]: #574
-#686 := (iff #431 #685)
-#688 := (= #420 #580)
-#694 := (+ #175 1::int)
-#581 := (= #694 #580)
-#540 := [rewrite]: #581
-#697 := (= #420 #694)
-#695 := (= #429 1::int)
-#536 := (* -1::int -1::int)
-#692 := (= #536 1::int)
-#693 := [rewrite]: #692
-#538 := (= #429 #536)
-#700 := (= #428 -1::int)
-#704 := (div -1::int 3::int)
-#705 := (= #704 -1::int)
-#706 := [rewrite]: #705
-#698 := (= #428 #704)
-#702 := (= #343 3::int)
-#703 := [rewrite]: #702
-#559 := (= #427 -1::int)
-#701 := [rewrite]: #559
-#699 := [monotonicity #701 #703]: #698
-#707 := [trans #699 #706]: #700
-#539 := [monotonicity #707]: #538
-#696 := [trans #539 #693]: #695
-#579 := [monotonicity #696]: #697
-#590 := [trans #579 #540]: #688
-#689 := [monotonicity #590]: #686
-#592 := [trans #689 #589]: #591
-#557 := (iff #434 #752)
-#710 := (= #430 #175)
-#713 := (+ #175 0::int)
-#714 := (= #713 #175)
-#715 := [rewrite]: #714
-#708 := (= #430 #713)
-#711 := (= #433 0::int)
-#435 := (* -1::int 0::int)
-#718 := (= #435 0::int)
-#722 := [rewrite]: #718
-#717 := (= #433 #435)
-#719 := (= #432 0::int)
-#720 := [rewrite]: #719
-#721 := [monotonicity #720]: #717
-#712 := [trans #721 #722]: #711
-#709 := [monotonicity #712]: #708
-#716 := [trans #709 #715]: #710
-#558 := [monotonicity #716]: #557
-#440 := (iff #419 false)
-#406 := (or false false)
-#731 := (iff #406 false)
-#394 := [rewrite]: #731
-#438 := (iff #419 #406)
-#724 := (iff #407 false)
-#1 := true
-#381 := (not true)
-#740 := (iff #381 false)
-#743 := [rewrite]: #740
-#723 := (iff #407 #381)
-#448 := (iff #750 true)
-#454 := (or false true)
-#725 := (iff #454 true)
-#726 := [rewrite]: #725
-#455 := (iff #750 #454)
-#397 := (iff #413 true)
-#733 := [rewrite]: #397
-#452 := (iff #748 false)
-#453 := [rewrite]: #452
-#349 := [monotonicity #453 #733]: #455
-#449 := [trans #349 #726]: #448
-#727 := [monotonicity #449]: #723
-#728 := [trans #727 #743]: #724
-#729 := (iff #418 false)
-#382 := (iff #418 #381)
-#739 := (iff #744 true)
-#736 := (or true true)
-#738 := (iff #736 true)
-#376 := [rewrite]: #738
-#737 := (iff #744 #736)
-#734 := (iff #537 true)
-#735 := [rewrite]: #734
-#732 := [monotonicity #733 #735]: #737
-#741 := [trans #732 #376]: #739
-#742 := [monotonicity #741]: #382
-#730 := [trans #742 #743]: #729
-#439 := [monotonicity #730 #728]: #438
-#441 := [trans #439 #394]: #440
-#594 := [monotonicity #441 #558 #592]: #593
-#690 := [trans #594 #596]: #597
-#395 := (iff #749 false)
-#747 := (iff #749 #406)
-#410 := (iff #753 false)
-#746 := [rewrite]: #410
-#745 := (iff #754 false)
-#405 := [rewrite]: #745
-#390 := [monotonicity #405 #746]: #747
-#396 := [trans #390 #394]: #395
-#691 := [monotonicity #396 #690]: #687
-#680 := [trans #691 #596]: #678
-#683 := [monotonicity #680]: #682
-#673 := [trans #683 #684]: #682
-#681 := [quant-inst]: #671
-#674 := [mp #681 #673]: #669
-[unit-resolution #674 #762 #198]: false
-unsat
-3ca836a50902cedb9598c64c13667ab84779478c 323 0
-#2 := false
-#68 := -1::int
-decl f3 :: (-> int int int)
-#172 := -3::int
-#38 := 3::int
-#175 := (f3 3::int -3::int)
-#180 := (= #175 -1::int)
-#193 := (not #180)
-#41 := 1::int
-#42 := (- 1::int)
-#39 := (- 3::int)
-#40 := (f3 3::int #39)
-#43 := (= #40 #42)
-#44 := (not #43)
-#196 := (iff #44 #193)
-#183 := (= -1::int #175)
-#188 := (not #183)
-#194 := (iff #188 #193)
-#191 := (iff #183 #180)
-#192 := [rewrite]: #191
-#195 := [monotonicity #192]: #194
-#189 := (iff #44 #188)
-#186 := (iff #43 #183)
-#184 := (iff #180 #183)
-#185 := [rewrite]: #184
-#181 := (iff #43 #180)
-#178 := (= #42 -1::int)
-#179 := [rewrite]: #178
-#176 := (= #40 #175)
-#173 := (= #39 -3::int)
-#174 := [rewrite]: #173
-#177 := [monotonicity #174]: #176
-#182 := [monotonicity #177 #179]: #181
-#187 := [trans #182 #185]: #186
-#190 := [monotonicity #187]: #189
-#197 := [trans #190 #195]: #196
-#171 := [asserted]: #44
-#198 := [mp #171 #197]: #193
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#10 := (f3 #8 #9)
-#756 := (pattern #10)
-#11 := 0::int
-#72 := (* -1::int #9)
-#69 := (* -1::int #8)
-#75 := (div #69 #72)
-#259 := (* -1::int #75)
-#260 := (+ #10 #259)
-#261 := (= #260 0::int)
-#21 := (div #8 #9)
-#256 := (* -1::int #21)
-#257 := (+ #10 #256)
-#258 := (= #257 0::int)
-#94 := (<= #9 0::int)
-#90 := (<= #8 0::int)
-#214 := (or #90 #94)
-#215 := (not #214)
-#101 := (>= #8 0::int)
-#206 := (or #94 #101)
-#207 := (not #206)
-#221 := (or #207 #215)
-#262 := (ite #221 #258 #261)
-#255 := (= #10 0::int)
-#13 := (= #9 0::int)
-#12 := (= #8 0::int)
-#14 := (or #12 #13)
-#263 := (ite #14 #255 #262)
-#757 := (forall (vars (?v0 int) (?v1 int)) (:pat #756) #263)
-#266 := (forall (vars (?v0 int) (?v1 int)) #263)
-#760 := (iff #266 #757)
-#758 := (iff #263 #263)
-#759 := [refl]: #758
-#761 := [quant-intro #759]: #760
-#226 := (ite #221 #21 #75)
-#229 := (ite #14 0::int #226)
-#232 := (= #10 #229)
-#235 := (forall (vars (?v0 int) (?v1 int)) #232)
-#267 := (iff #235 #266)
-#264 := (iff #232 #263)
-#265 := [rewrite]: #264
-#268 := [quant-intro #265]: #267
-#102 := (not #101)
-#95 := (not #94)
-#105 := (and #95 #102)
-#91 := (not #90)
-#98 := (and #91 #95)
-#108 := (or #98 #105)
-#111 := (ite #108 #21 #75)
-#114 := (ite #14 0::int #111)
-#117 := (= #10 #114)
-#120 := (forall (vars (?v0 int) (?v1 int)) #117)
-#236 := (iff #120 #235)
-#233 := (iff #117 #232)
-#230 := (= #114 #229)
-#227 := (= #111 #226)
-#224 := (iff #108 #221)
-#218 := (or #215 #207)
-#222 := (iff #218 #221)
-#223 := [rewrite]: #222
-#219 := (iff #108 #218)
-#216 := (iff #105 #207)
-#217 := [rewrite]: #216
-#204 := (iff #98 #215)
-#205 := [rewrite]: #204
-#220 := [monotonicity #205 #217]: #219
-#225 := [trans #220 #223]: #224
-#228 := [monotonicity #225]: #227
-#231 := [monotonicity #228]: #230
-#234 := [monotonicity #231]: #233
-#237 := [quant-intro #234]: #236
-#210 := (~ #120 #120)
-#208 := (~ #117 #117)
-#209 := [refl]: #208
-#211 := [nnf-pos #209]: #210
-#23 := (- #9)
-#22 := (- #8)
-#24 := (div #22 #23)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#25 := (ite #20 #21 #24)
-#26 := (ite #14 0::int #25)
-#27 := (= #10 #26)
-#28 := (forall (vars (?v0 int) (?v1 int)) #27)
-#123 := (iff #28 #120)
-#62 := (and #16 #18)
-#65 := (or #17 #62)
-#78 := (ite #65 #21 #75)
-#81 := (ite #14 0::int #78)
-#84 := (= #10 #81)
-#87 := (forall (vars (?v0 int) (?v1 int)) #84)
-#121 := (iff #87 #120)
-#118 := (iff #84 #117)
-#115 := (= #81 #114)
-#112 := (= #78 #111)
-#109 := (iff #65 #108)
-#106 := (iff #62 #105)
-#103 := (iff #18 #102)
-#104 := [rewrite]: #103
-#96 := (iff #16 #95)
-#97 := [rewrite]: #96
-#107 := [monotonicity #97 #104]: #106
-#99 := (iff #17 #98)
-#92 := (iff #15 #91)
-#93 := [rewrite]: #92
-#100 := [monotonicity #93 #97]: #99
-#110 := [monotonicity #100 #107]: #109
-#113 := [monotonicity #110]: #112
-#116 := [monotonicity #113]: #115
-#119 := [monotonicity #116]: #118
-#122 := [quant-intro #119]: #121
-#88 := (iff #28 #87)
-#85 := (iff #27 #84)
-#82 := (= #26 #81)
-#79 := (= #25 #78)
-#76 := (= #24 #75)
-#73 := (= #23 #72)
-#74 := [rewrite]: #73
-#70 := (= #22 #69)
-#71 := [rewrite]: #70
-#77 := [monotonicity #71 #74]: #76
-#66 := (iff #20 #65)
-#63 := (iff #19 #62)
-#64 := [rewrite]: #63
-#67 := [monotonicity #64]: #66
-#80 := [monotonicity #67 #77]: #79
-#83 := [monotonicity #80]: #82
-#86 := [monotonicity #83]: #85
-#89 := [quant-intro #86]: #88
-#124 := [trans #89 #122]: #123
-#61 := [asserted]: #28
-#125 := [mp #61 #124]: #120
-#201 := [mp~ #125 #211]: #120
-#238 := [mp #201 #237]: #235
-#269 := [mp #238 #268]: #266
-#762 := [mp #269 #761]: #757
-#680 := (not #757)
-#672 := (or #680 #180)
-#343 := (* -1::int -3::int)
-#427 := (* -1::int 3::int)
-#428 := (div #427 #343)
-#429 := (* -1::int #428)
-#420 := (+ #175 #429)
-#431 := (= #420 0::int)
-#432 := (div 3::int -3::int)
-#433 := (* -1::int #432)
-#430 := (+ #175 #433)
-#434 := (= #430 0::int)
-#413 := (<= -3::int 0::int)
-#748 := (<= 3::int 0::int)
-#750 := (or #748 #413)
-#407 := (not #750)
-#537 := (>= 3::int 0::int)
-#744 := (or #413 #537)
-#418 := (not #744)
-#419 := (or #418 #407)
-#751 := (ite #419 #434 #431)
-#752 := (= #175 0::int)
-#753 := (= -3::int 0::int)
-#754 := (= 3::int 0::int)
-#749 := (or #754 #753)
-#755 := (ite #749 #752 #751)
-#669 := (or #680 #755)
-#681 := (iff #669 #672)
-#683 := (iff #672 #672)
-#679 := [rewrite]: #683
-#691 := (iff #755 #180)
-#595 := (ite false #752 #180)
-#690 := (iff #595 #180)
-#687 := [rewrite]: #690
-#596 := (iff #755 #595)
-#593 := (iff #751 #180)
-#574 := (ite false #180 #180)
-#592 := (iff #574 #180)
-#585 := [rewrite]: #592
-#589 := (iff #751 #574)
-#686 := (iff #431 #180)
-#714 := (+ 1::int #175)
-#558 := (= #714 0::int)
-#702 := (iff #558 #180)
-#703 := [rewrite]: #702
-#590 := (iff #431 #558)
-#540 := (= #420 #714)
-#713 := (+ #175 1::int)
-#715 := (= #713 #714)
-#710 := [rewrite]: #715
-#580 := (= #420 #713)
-#697 := (= #429 1::int)
-#435 := (* -1::int -1::int)
-#718 := (= #435 1::int)
-#722 := [rewrite]: #718
-#696 := (= #429 #435)
-#693 := (= #428 -1::int)
-#707 := (div -3::int 3::int)
-#539 := (= #707 -1::int)
-#692 := [rewrite]: #539
-#536 := (= #428 #707)
-#706 := (= #343 3::int)
-#700 := [rewrite]: #706
-#699 := (= #427 -3::int)
-#705 := [rewrite]: #699
-#538 := [monotonicity #705 #700]: #536
-#695 := [trans #538 #692]: #693
-#694 := [monotonicity #695]: #696
-#579 := [trans #694 #722]: #697
-#581 := [monotonicity #579]: #580
-#688 := [trans #581 #710]: #540
-#685 := [monotonicity #688]: #590
-#689 := [trans #685 #703]: #686
-#704 := (iff #434 #180)
-#559 := (iff #434 #558)
-#716 := (= #430 #714)
-#708 := (= #430 #713)
-#711 := (= #433 1::int)
-#717 := (= #433 #435)
-#719 := (= #432 -1::int)
-#720 := [rewrite]: #719
-#721 := [monotonicity #720]: #717
-#712 := [trans #721 #722]: #711
-#709 := [monotonicity #712]: #708
-#557 := [trans #709 #710]: #716
-#701 := [monotonicity #557]: #559
-#698 := [trans #701 #703]: #704
-#440 := (iff #419 false)
-#406 := (or false false)
-#731 := (iff #406 false)
-#394 := [rewrite]: #731
-#438 := (iff #419 #406)
-#724 := (iff #407 false)
-#1 := true
-#381 := (not true)
-#740 := (iff #381 false)
-#743 := [rewrite]: #740
-#723 := (iff #407 #381)
-#448 := (iff #750 true)
-#454 := (or false true)
-#725 := (iff #454 true)
-#726 := [rewrite]: #725
-#455 := (iff #750 #454)
-#397 := (iff #413 true)
-#733 := [rewrite]: #397
-#452 := (iff #748 false)
-#453 := [rewrite]: #452
-#349 := [monotonicity #453 #733]: #455
-#449 := [trans #349 #726]: #448
-#727 := [monotonicity #449]: #723
-#728 := [trans #727 #743]: #724
-#729 := (iff #418 false)
-#382 := (iff #418 #381)
-#739 := (iff #744 true)
-#736 := (or true true)
-#738 := (iff #736 true)
-#376 := [rewrite]: #738
-#737 := (iff #744 #736)
-#734 := (iff #537 true)
-#735 := [rewrite]: #734
-#732 := [monotonicity #733 #735]: #737
-#741 := [trans #732 #376]: #739
-#742 := [monotonicity #741]: #382
-#730 := [trans #742 #743]: #729
-#439 := [monotonicity #730 #728]: #438
-#441 := [trans #439 #394]: #440
-#591 := [monotonicity #441 #698 #689]: #589
-#594 := [trans #591 #585]: #593
-#395 := (iff #749 false)
-#747 := (iff #749 #406)
-#410 := (iff #753 false)
-#746 := [rewrite]: #410
-#745 := (iff #754 false)
-#405 := [rewrite]: #745
-#390 := [monotonicity #405 #746]: #747
-#396 := [trans #390 #394]: #395
-#597 := [monotonicity #396 #594]: #596
-#678 := [trans #597 #687]: #691
-#682 := [monotonicity #678]: #681
-#684 := [trans #682 #679]: #681
-#671 := [quant-inst]: #669
-#673 := [mp #671 #684]: #672
-[unit-resolution #673 #762 #198]: false
-unsat
-7e60c8e870227efa11986e096bca8fb9f1d5067a 322 0
+ed5e3568452fe6600500ab5071290e80a7175922 322 0
 #2 := false
 #38 := 1::int
 decl f4 :: (-> int int int)
@@ -36065,334 +38552,7 @@
 #642 := [mp #654 #658]: #660
 [unit-resolution #642 #750 #169]: false
 unsat
-4e54628370c67b364f92814a1e587bd252e7851a 326 0
-#2 := false
-#179 := -2::int
-decl f3 :: (-> int int int)
-#173 := -3::int
-#38 := 5::int
-#176 := (f3 5::int -3::int)
-#182 := (= #176 -2::int)
-#185 := (not #182)
-#42 := 2::int
-#43 := (- 2::int)
-#39 := 3::int
-#40 := (- 3::int)
-#41 := (f3 5::int #40)
-#44 := (= #41 #43)
-#45 := (not #44)
-#186 := (iff #45 #185)
-#183 := (iff #44 #182)
-#180 := (= #43 -2::int)
-#181 := [rewrite]: #180
-#177 := (= #41 #176)
-#174 := (= #40 -3::int)
-#175 := [rewrite]: #174
-#178 := [monotonicity #175]: #177
-#184 := [monotonicity #178 #181]: #183
-#187 := [monotonicity #184]: #186
-#172 := [asserted]: #45
-#190 := [mp #172 #187]: #185
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#10 := (f3 #8 #9)
-#747 := (pattern #10)
-#11 := 0::int
-#69 := -1::int
-#73 := (* -1::int #9)
-#70 := (* -1::int #8)
-#76 := (div #70 #73)
-#249 := (* -1::int #76)
-#250 := (+ #10 #249)
-#251 := (= #250 0::int)
-#21 := (div #8 #9)
-#246 := (* -1::int #21)
-#247 := (+ #10 #246)
-#248 := (= #247 0::int)
-#95 := (<= #9 0::int)
-#91 := (<= #8 0::int)
-#204 := (or #91 #95)
-#205 := (not #204)
-#102 := (>= #8 0::int)
-#196 := (or #95 #102)
-#197 := (not #196)
-#211 := (or #197 #205)
-#252 := (ite #211 #248 #251)
-#245 := (= #10 0::int)
-#13 := (= #9 0::int)
-#12 := (= #8 0::int)
-#14 := (or #12 #13)
-#253 := (ite #14 #245 #252)
-#748 := (forall (vars (?v0 int) (?v1 int)) (:pat #747) #253)
-#256 := (forall (vars (?v0 int) (?v1 int)) #253)
-#751 := (iff #256 #748)
-#749 := (iff #253 #253)
-#750 := [refl]: #749
-#752 := [quant-intro #750]: #751
-#216 := (ite #211 #21 #76)
-#219 := (ite #14 0::int #216)
-#222 := (= #10 #219)
-#225 := (forall (vars (?v0 int) (?v1 int)) #222)
-#257 := (iff #225 #256)
-#254 := (iff #222 #253)
-#255 := [rewrite]: #254
-#258 := [quant-intro #255]: #257
-#103 := (not #102)
-#96 := (not #95)
-#106 := (and #96 #103)
-#92 := (not #91)
-#99 := (and #92 #96)
-#109 := (or #99 #106)
-#112 := (ite #109 #21 #76)
-#115 := (ite #14 0::int #112)
-#118 := (= #10 #115)
-#121 := (forall (vars (?v0 int) (?v1 int)) #118)
-#226 := (iff #121 #225)
-#223 := (iff #118 #222)
-#220 := (= #115 #219)
-#217 := (= #112 #216)
-#214 := (iff #109 #211)
-#208 := (or #205 #197)
-#212 := (iff #208 #211)
-#213 := [rewrite]: #212
-#209 := (iff #109 #208)
-#206 := (iff #106 #197)
-#207 := [rewrite]: #206
-#194 := (iff #99 #205)
-#195 := [rewrite]: #194
-#210 := [monotonicity #195 #207]: #209
-#215 := [trans #210 #213]: #214
-#218 := [monotonicity #215]: #217
-#221 := [monotonicity #218]: #220
-#224 := [monotonicity #221]: #223
-#227 := [quant-intro #224]: #226
-#200 := (~ #121 #121)
-#198 := (~ #118 #118)
-#199 := [refl]: #198
-#201 := [nnf-pos #199]: #200
-#23 := (- #9)
-#22 := (- #8)
-#24 := (div #22 #23)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#25 := (ite #20 #21 #24)
-#26 := (ite #14 0::int #25)
-#27 := (= #10 #26)
-#28 := (forall (vars (?v0 int) (?v1 int)) #27)
-#124 := (iff #28 #121)
-#63 := (and #16 #18)
-#66 := (or #17 #63)
-#79 := (ite #66 #21 #76)
-#82 := (ite #14 0::int #79)
-#85 := (= #10 #82)
-#88 := (forall (vars (?v0 int) (?v1 int)) #85)
-#122 := (iff #88 #121)
-#119 := (iff #85 #118)
-#116 := (= #82 #115)
-#113 := (= #79 #112)
-#110 := (iff #66 #109)
-#107 := (iff #63 #106)
-#104 := (iff #18 #103)
-#105 := [rewrite]: #104
-#97 := (iff #16 #96)
-#98 := [rewrite]: #97
-#108 := [monotonicity #98 #105]: #107
-#100 := (iff #17 #99)
-#93 := (iff #15 #92)
-#94 := [rewrite]: #93
-#101 := [monotonicity #94 #98]: #100
-#111 := [monotonicity #101 #108]: #110
-#114 := [monotonicity #111]: #113
-#117 := [monotonicity #114]: #116
-#120 := [monotonicity #117]: #119
-#123 := [quant-intro #120]: #122
-#89 := (iff #28 #88)
-#86 := (iff #27 #85)
-#83 := (= #26 #82)
-#80 := (= #25 #79)
-#77 := (= #24 #76)
-#74 := (= #23 #73)
-#75 := [rewrite]: #74
-#71 := (= #22 #70)
-#72 := [rewrite]: #71
-#78 := [monotonicity #72 #75]: #77
-#67 := (iff #20 #66)
-#64 := (iff #19 #63)
-#65 := [rewrite]: #64
-#68 := [monotonicity #65]: #67
-#81 := [monotonicity #68 #78]: #80
-#84 := [monotonicity #81]: #83
-#87 := [monotonicity #84]: #86
-#90 := [quant-intro #87]: #89
-#125 := [trans #90 #123]: #124
-#62 := [asserted]: #28
-#126 := [mp #62 #125]: #121
-#191 := [mp~ #126 #201]: #121
-#228 := [mp #191 #227]: #225
-#259 := [mp #228 #258]: #256
-#753 := [mp #259 #752]: #748
-#667 := (not #748)
-#661 := (or #667 #182)
-#333 := (* -1::int -3::int)
-#418 := (* -1::int 5::int)
-#419 := (div #418 #333)
-#420 := (* -1::int #419)
-#411 := (+ #176 #420)
-#422 := (= #411 0::int)
-#423 := (div 5::int -3::int)
-#351 := (* -1::int #423)
-#424 := (+ #176 #351)
-#421 := (= #424 0::int)
-#425 := (<= -3::int 0::int)
-#404 := (<= 5::int 0::int)
-#739 := (or #404 #425)
-#741 := (not #739)
-#398 := (>= 5::int 0::int)
-#528 := (or #425 #398)
-#735 := (not #528)
-#409 := (or #735 #741)
-#410 := (ite #409 #421 #422)
-#742 := (= #176 0::int)
-#743 := (= -3::int 0::int)
-#744 := (= 5::int 0::int)
-#745 := (or #744 #743)
-#740 := (ite #745 #742 #410)
-#668 := (or #667 #740)
-#653 := (iff #668 #661)
-#656 := (iff #661 #661)
-#657 := [rewrite]: #656
-#665 := (iff #740 #182)
-#673 := (ite false #742 #182)
-#675 := (iff #673 #182)
-#664 := [rewrite]: #675
-#674 := (iff #740 #673)
-#662 := (iff #410 #182)
-#693 := (= #176 -1::int)
-#682 := (ite false #693 #182)
-#663 := (iff #682 #182)
-#660 := [rewrite]: #663
-#669 := (iff #410 #682)
-#681 := (iff #422 #182)
-#565 := (+ 2::int #176)
-#584 := (= #565 0::int)
-#587 := (iff #584 #182)
-#588 := [rewrite]: #587
-#585 := (iff #422 #584)
-#583 := (= #411 #565)
-#676 := (+ #176 2::int)
-#580 := (= #676 #565)
-#582 := [rewrite]: #580
-#677 := (= #411 #676)
-#679 := (= #420 2::int)
-#688 := (* -1::int -2::int)
-#572 := (= #688 2::int)
-#531 := [rewrite]: #572
-#570 := (= #420 #688)
-#687 := (= #419 -2::int)
-#696 := -5::int
-#529 := (div -5::int 3::int)
-#684 := (= #529 -2::int)
-#686 := [rewrite]: #684
-#530 := (= #419 #529)
-#698 := (= #333 3::int)
-#527 := [rewrite]: #698
-#697 := (= #418 -5::int)
-#691 := [rewrite]: #697
-#683 := [monotonicity #691 #527]: #530
-#685 := [trans #683 #686]: #687
-#571 := [monotonicity #685]: #570
-#581 := [trans #571 #531]: #679
-#680 := [monotonicity #581]: #677
-#576 := [trans #680 #582]: #583
-#586 := [monotonicity #576]: #585
-#678 := [trans #586 #588]: #681
-#689 := (iff #421 #693)
-#712 := 1::int
-#705 := (+ 1::int #176)
-#549 := (= #705 0::int)
-#694 := (iff #549 #693)
-#695 := [rewrite]: #694
-#550 := (iff #421 #549)
-#707 := (= #424 #705)
-#704 := (+ #176 1::int)
-#706 := (= #704 #705)
-#701 := [rewrite]: #706
-#699 := (= #424 #704)
-#702 := (= #351 1::int)
-#711 := (* -1::int -1::int)
-#709 := (= #711 1::int)
-#713 := [rewrite]: #709
-#426 := (= #351 #711)
-#432 := (= #423 -1::int)
-#710 := [rewrite]: #432
-#708 := [monotonicity #710]: #426
-#703 := [trans #708 #713]: #702
-#700 := [monotonicity #703]: #699
-#548 := [trans #700 #701]: #707
-#692 := [monotonicity #548]: #550
-#690 := [trans #692 #695]: #689
-#430 := (iff #409 false)
-#737 := (or false false)
-#381 := (iff #737 false)
-#722 := [rewrite]: #381
-#719 := (iff #409 #737)
-#718 := (iff #741 false)
-#1 := true
-#732 := (not true)
-#733 := (iff #732 false)
-#731 := [rewrite]: #733
-#440 := (iff #741 #732)
-#717 := (iff #739 true)
-#444 := (or false true)
-#339 := (iff #444 true)
-#716 := [rewrite]: #339
-#445 := (iff #739 #444)
-#387 := (iff #425 true)
-#388 := [rewrite]: #387
-#721 := (iff #404 false)
-#443 := [rewrite]: #721
-#446 := [monotonicity #443 #388]: #445
-#439 := [trans #446 #716]: #717
-#714 := [monotonicity #439]: #440
-#715 := [trans #714 #731]: #718
-#734 := (iff #735 false)
-#372 := (iff #735 #732)
-#367 := (iff #528 true)
-#726 := (or true true)
-#723 := (iff #726 true)
-#729 := [rewrite]: #723
-#727 := (iff #528 #726)
-#724 := (iff #398 true)
-#725 := [rewrite]: #724
-#728 := [monotonicity #388 #725]: #727
-#730 := [trans #728 #729]: #367
-#373 := [monotonicity #730]: #372
-#720 := [trans #373 #731]: #734
-#429 := [monotonicity #720 #715]: #719
-#431 := [trans #429 #722]: #430
-#671 := [monotonicity #431 #690 #678]: #669
-#672 := [trans #671 #660]: #662
-#385 := (iff #745 false)
-#397 := (iff #745 #737)
-#396 := (iff #743 false)
-#401 := [rewrite]: #396
-#746 := (iff #744 false)
-#736 := [rewrite]: #746
-#738 := [monotonicity #736 #401]: #397
-#386 := [trans #738 #722]: #385
-#670 := [monotonicity #386 #672]: #674
-#666 := [trans #670 #664]: #665
-#655 := [monotonicity #666]: #653
-#658 := [trans #655 #657]: #653
-#652 := [quant-inst]: #668
-#654 := [mp #652 #658]: #661
-[unit-resolution #654 #753 #190]: false
-unsat
-6560bfd674a12139ed8cf76157a9702921110fe8 297 0
+1897b65589b01a36e8c901ddcdce3eb69e66fd4d 297 0
 #2 := false
 #11 := 0::int
 decl f4 :: (-> int int int)
@@ -36690,329 +38850,7 @@
 #576 := [mp #570 #575]: #670
 [unit-resolution #576 #750 #168]: false
 unsat
-d0a25a18442fdd27e302706e9c8d8ea55968f062 321 0
-#2 := false
-#67 := -1::int
-decl f3 :: (-> int int int)
-#40 := 3::int
-#173 := (f3 -1::int 3::int)
-#176 := (= #173 -1::int)
-#189 := (not #176)
-#38 := 1::int
-#39 := (- 1::int)
-#41 := (f3 #39 3::int)
-#42 := (= #41 #39)
-#43 := (not #42)
-#192 := (iff #43 #189)
-#179 := (= -1::int #173)
-#184 := (not #179)
-#190 := (iff #184 #189)
-#187 := (iff #179 #176)
-#188 := [rewrite]: #187
-#191 := [monotonicity #188]: #190
-#185 := (iff #43 #184)
-#182 := (iff #42 #179)
-#180 := (iff #176 #179)
-#181 := [rewrite]: #180
-#177 := (iff #42 #176)
-#171 := (= #39 -1::int)
-#172 := [rewrite]: #171
-#174 := (= #41 #173)
-#175 := [monotonicity #172]: #174
-#178 := [monotonicity #175 #172]: #177
-#183 := [trans #178 #181]: #182
-#186 := [monotonicity #183]: #185
-#193 := [trans #186 #191]: #192
-#170 := [asserted]: #43
-#194 := [mp #170 #193]: #189
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#10 := (f3 #8 #9)
-#752 := (pattern #10)
-#11 := 0::int
-#71 := (* -1::int #9)
-#68 := (* -1::int #8)
-#74 := (div #68 #71)
-#255 := (* -1::int #74)
-#256 := (+ #10 #255)
-#257 := (= #256 0::int)
-#21 := (div #8 #9)
-#252 := (* -1::int #21)
-#253 := (+ #10 #252)
-#254 := (= #253 0::int)
-#93 := (<= #9 0::int)
-#89 := (<= #8 0::int)
-#210 := (or #89 #93)
-#211 := (not #210)
-#100 := (>= #8 0::int)
-#202 := (or #93 #100)
-#203 := (not #202)
-#217 := (or #203 #211)
-#258 := (ite #217 #254 #257)
-#251 := (= #10 0::int)
-#13 := (= #9 0::int)
-#12 := (= #8 0::int)
-#14 := (or #12 #13)
-#259 := (ite #14 #251 #258)
-#753 := (forall (vars (?v0 int) (?v1 int)) (:pat #752) #259)
-#262 := (forall (vars (?v0 int) (?v1 int)) #259)
-#756 := (iff #262 #753)
-#754 := (iff #259 #259)
-#755 := [refl]: #754
-#757 := [quant-intro #755]: #756
-#222 := (ite #217 #21 #74)
-#225 := (ite #14 0::int #222)
-#228 := (= #10 #225)
-#231 := (forall (vars (?v0 int) (?v1 int)) #228)
-#263 := (iff #231 #262)
-#260 := (iff #228 #259)
-#261 := [rewrite]: #260
-#264 := [quant-intro #261]: #263
-#101 := (not #100)
-#94 := (not #93)
-#104 := (and #94 #101)
-#90 := (not #89)
-#97 := (and #90 #94)
-#107 := (or #97 #104)
-#110 := (ite #107 #21 #74)
-#113 := (ite #14 0::int #110)
-#116 := (= #10 #113)
-#119 := (forall (vars (?v0 int) (?v1 int)) #116)
-#232 := (iff #119 #231)
-#229 := (iff #116 #228)
-#226 := (= #113 #225)
-#223 := (= #110 #222)
-#220 := (iff #107 #217)
-#214 := (or #211 #203)
-#218 := (iff #214 #217)
-#219 := [rewrite]: #218
-#215 := (iff #107 #214)
-#212 := (iff #104 #203)
-#213 := [rewrite]: #212
-#200 := (iff #97 #211)
-#201 := [rewrite]: #200
-#216 := [monotonicity #201 #213]: #215
-#221 := [trans #216 #219]: #220
-#224 := [monotonicity #221]: #223
-#227 := [monotonicity #224]: #226
-#230 := [monotonicity #227]: #229
-#233 := [quant-intro #230]: #232
-#206 := (~ #119 #119)
-#204 := (~ #116 #116)
-#205 := [refl]: #204
-#207 := [nnf-pos #205]: #206
-#23 := (- #9)
-#22 := (- #8)
-#24 := (div #22 #23)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#25 := (ite #20 #21 #24)
-#26 := (ite #14 0::int #25)
-#27 := (= #10 #26)
-#28 := (forall (vars (?v0 int) (?v1 int)) #27)
-#122 := (iff #28 #119)
-#61 := (and #16 #18)
-#64 := (or #17 #61)
-#77 := (ite #64 #21 #74)
-#80 := (ite #14 0::int #77)
-#83 := (= #10 #80)
-#86 := (forall (vars (?v0 int) (?v1 int)) #83)
-#120 := (iff #86 #119)
-#117 := (iff #83 #116)
-#114 := (= #80 #113)
-#111 := (= #77 #110)
-#108 := (iff #64 #107)
-#105 := (iff #61 #104)
-#102 := (iff #18 #101)
-#103 := [rewrite]: #102
-#95 := (iff #16 #94)
-#96 := [rewrite]: #95
-#106 := [monotonicity #96 #103]: #105
-#98 := (iff #17 #97)
-#91 := (iff #15 #90)
-#92 := [rewrite]: #91
-#99 := [monotonicity #92 #96]: #98
-#109 := [monotonicity #99 #106]: #108
-#112 := [monotonicity #109]: #111
-#115 := [monotonicity #112]: #114
-#118 := [monotonicity #115]: #117
-#121 := [quant-intro #118]: #120
-#87 := (iff #28 #86)
-#84 := (iff #27 #83)
-#81 := (= #26 #80)
-#78 := (= #25 #77)
-#75 := (= #24 #74)
-#72 := (= #23 #71)
-#73 := [rewrite]: #72
-#69 := (= #22 #68)
-#70 := [rewrite]: #69
-#76 := [monotonicity #70 #73]: #75
-#65 := (iff #20 #64)
-#62 := (iff #19 #61)
-#63 := [rewrite]: #62
-#66 := [monotonicity #63]: #65
-#79 := [monotonicity #66 #76]: #78
-#82 := [monotonicity #79]: #81
-#85 := [monotonicity #82]: #84
-#88 := [quant-intro #85]: #87
-#123 := [trans #88 #121]: #122
-#60 := [asserted]: #28
-#124 := [mp #60 #123]: #119
-#197 := [mp~ #124 #207]: #119
-#234 := [mp #197 #233]: #231
-#265 := [mp #234 #264]: #262
-#758 := [mp #265 #757]: #753
-#665 := (not #753)
-#667 := (or #665 #176)
-#339 := (* -1::int 3::int)
-#423 := (* -1::int -1::int)
-#424 := (div #423 #339)
-#425 := (* -1::int #424)
-#416 := (+ #173 #425)
-#427 := (= #416 0::int)
-#428 := (div -1::int 3::int)
-#429 := (* -1::int #428)
-#426 := (+ #173 #429)
-#430 := (= #426 0::int)
-#409 := (<= 3::int 0::int)
-#744 := (<= -1::int 0::int)
-#746 := (or #744 #409)
-#403 := (not #746)
-#533 := (>= -1::int 0::int)
-#740 := (or #409 #533)
-#414 := (not #740)
-#415 := (or #414 #403)
-#747 := (ite #415 #430 #427)
-#748 := (= #173 0::int)
-#749 := (= 3::int 0::int)
-#750 := (= -1::int 0::int)
-#745 := (or #750 #749)
-#751 := (ite #745 #748 #747)
-#677 := (or #665 #751)
-#679 := (iff #677 #667)
-#680 := (iff #667 #667)
-#669 := [rewrite]: #680
-#676 := (iff #751 #176)
-#593 := (ite false #748 #176)
-#687 := (iff #593 #176)
-#674 := [rewrite]: #687
-#686 := (iff #751 #593)
-#591 := (iff #747 #176)
-#1 := true
-#587 := (ite true #176 #748)
-#589 := (iff #587 #176)
-#590 := [rewrite]: #589
-#588 := (iff #747 #587)
-#570 := (iff #427 #748)
-#682 := (= #416 #173)
-#577 := (+ #173 0::int)
-#586 := (= #577 #173)
-#681 := [rewrite]: #586
-#536 := (= #416 #577)
-#575 := (= #425 0::int)
-#689 := (* -1::int 0::int)
-#690 := (= #689 0::int)
-#693 := [rewrite]: #690
-#691 := (= #425 #689)
-#535 := (= #424 0::int)
-#694 := -3::int
-#702 := (div 1::int -3::int)
-#532 := (= #702 0::int)
-#534 := [rewrite]: #532
-#696 := (= #424 #702)
-#695 := (= #339 -3::int)
-#701 := [rewrite]: #695
-#717 := (= #423 1::int)
-#714 := [rewrite]: #717
-#703 := [monotonicity #714 #701]: #696
-#688 := [trans #703 #534]: #535
-#692 := [monotonicity #688]: #691
-#576 := [trans #692 #693]: #575
-#684 := [monotonicity #576]: #536
-#685 := [trans #684 #681]: #682
-#585 := [monotonicity #685]: #570
-#699 := (iff #430 #176)
-#705 := (+ 1::int #173)
-#553 := (= #705 0::int)
-#697 := (iff #553 #176)
-#698 := [rewrite]: #697
-#554 := (iff #430 #553)
-#706 := (= #426 #705)
-#708 := (+ #173 1::int)
-#710 := (= #708 #705)
-#711 := [rewrite]: #710
-#709 := (= #426 #708)
-#718 := (= #429 1::int)
-#431 := (= #429 #423)
-#715 := (= #428 -1::int)
-#716 := [rewrite]: #715
-#713 := [monotonicity #716]: #431
-#707 := [trans #713 #714]: #718
-#704 := [monotonicity #707]: #709
-#712 := [trans #704 #711]: #706
-#555 := [monotonicity #712]: #554
-#700 := [trans #555 #698]: #699
-#436 := (iff #415 true)
-#726 := (or true false)
-#450 := (iff #726 true)
-#451 := [rewrite]: #450
-#434 := (iff #415 #726)
-#720 := (iff #403 false)
-#722 := (not true)
-#719 := (iff #722 false)
-#723 := [rewrite]: #719
-#444 := (iff #403 #722)
-#345 := (iff #746 true)
-#448 := (iff #746 #726)
-#393 := (iff #409 false)
-#729 := [rewrite]: #393
-#739 := (iff #744 true)
-#725 := [rewrite]: #739
-#449 := [monotonicity #725 #729]: #448
-#721 := [trans #449 #451]: #345
-#445 := [monotonicity #721]: #444
-#724 := [trans #445 #723]: #720
-#738 := (iff #414 true)
-#372 := (not false)
-#377 := (iff #372 true)
-#378 := [rewrite]: #377
-#735 := (iff #414 #372)
-#728 := (iff #740 false)
-#402 := (or false false)
-#727 := (iff #402 false)
-#390 := [rewrite]: #727
-#732 := (iff #740 #402)
-#730 := (iff #533 false)
-#731 := [rewrite]: #730
-#733 := [monotonicity #729 #731]: #732
-#734 := [trans #733 #390]: #728
-#737 := [monotonicity #734]: #735
-#736 := [trans #737 #378]: #738
-#435 := [monotonicity #736 #724]: #434
-#437 := [trans #435 #451]: #436
-#581 := [monotonicity #437 #700 #585]: #588
-#592 := [trans #581 #590]: #591
-#391 := (iff #745 false)
-#743 := (iff #745 #402)
-#406 := (iff #749 false)
-#742 := [rewrite]: #406
-#741 := (iff #750 false)
-#401 := [rewrite]: #741
-#386 := [monotonicity #401 #742]: #743
-#392 := [trans #386 #390]: #391
-#683 := [monotonicity #392 #592]: #686
-#668 := [trans #683 #674]: #676
-#675 := [monotonicity #668]: #679
-#670 := [trans #675 #669]: #679
-#678 := [quant-inst]: #677
-#671 := [mp #678 #670]: #667
-[unit-resolution #671 #758 #194]: false
-unsat
-9f593513c6223913ee1444c7c2d179bd416f9b73 328 0
+0950f059c8c59b021ddf8b08664ab965652bd6ab 328 0
 #2 := false
 #41 := 2::int
 decl f4 :: (-> int int int)
@@ -37341,331 +39179,7 @@
 #641 := [mp #635 #636]: #650
 [unit-resolution #641 #742 #170]: false
 unsat
-502812e8687a6ca679f5bc63f75291a9fdb3db70 323 0
-#2 := false
-#68 := -1::int
-decl f3 :: (-> int int int)
-#38 := 3::int
-#172 := -3::int
-#175 := (f3 -3::int 3::int)
-#180 := (= #175 -1::int)
-#193 := (not #180)
-#41 := 1::int
-#42 := (- 1::int)
-#39 := (- 3::int)
-#40 := (f3 #39 3::int)
-#43 := (= #40 #42)
-#44 := (not #43)
-#196 := (iff #44 #193)
-#183 := (= -1::int #175)
-#188 := (not #183)
-#194 := (iff #188 #193)
-#191 := (iff #183 #180)
-#192 := [rewrite]: #191
-#195 := [monotonicity #192]: #194
-#189 := (iff #44 #188)
-#186 := (iff #43 #183)
-#184 := (iff #180 #183)
-#185 := [rewrite]: #184
-#181 := (iff #43 #180)
-#178 := (= #42 -1::int)
-#179 := [rewrite]: #178
-#176 := (= #40 #175)
-#173 := (= #39 -3::int)
-#174 := [rewrite]: #173
-#177 := [monotonicity #174]: #176
-#182 := [monotonicity #177 #179]: #181
-#187 := [trans #182 #185]: #186
-#190 := [monotonicity #187]: #189
-#197 := [trans #190 #195]: #196
-#171 := [asserted]: #44
-#198 := [mp #171 #197]: #193
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#10 := (f3 #8 #9)
-#756 := (pattern #10)
-#11 := 0::int
-#72 := (* -1::int #9)
-#69 := (* -1::int #8)
-#75 := (div #69 #72)
-#259 := (* -1::int #75)
-#260 := (+ #10 #259)
-#261 := (= #260 0::int)
-#21 := (div #8 #9)
-#256 := (* -1::int #21)
-#257 := (+ #10 #256)
-#258 := (= #257 0::int)
-#94 := (<= #9 0::int)
-#90 := (<= #8 0::int)
-#214 := (or #90 #94)
-#215 := (not #214)
-#101 := (>= #8 0::int)
-#206 := (or #94 #101)
-#207 := (not #206)
-#221 := (or #207 #215)
-#262 := (ite #221 #258 #261)
-#255 := (= #10 0::int)
-#13 := (= #9 0::int)
-#12 := (= #8 0::int)
-#14 := (or #12 #13)
-#263 := (ite #14 #255 #262)
-#757 := (forall (vars (?v0 int) (?v1 int)) (:pat #756) #263)
-#266 := (forall (vars (?v0 int) (?v1 int)) #263)
-#760 := (iff #266 #757)
-#758 := (iff #263 #263)
-#759 := [refl]: #758
-#761 := [quant-intro #759]: #760
-#226 := (ite #221 #21 #75)
-#229 := (ite #14 0::int #226)
-#232 := (= #10 #229)
-#235 := (forall (vars (?v0 int) (?v1 int)) #232)
-#267 := (iff #235 #266)
-#264 := (iff #232 #263)
-#265 := [rewrite]: #264
-#268 := [quant-intro #265]: #267
-#102 := (not #101)
-#95 := (not #94)
-#105 := (and #95 #102)
-#91 := (not #90)
-#98 := (and #91 #95)
-#108 := (or #98 #105)
-#111 := (ite #108 #21 #75)
-#114 := (ite #14 0::int #111)
-#117 := (= #10 #114)
-#120 := (forall (vars (?v0 int) (?v1 int)) #117)
-#236 := (iff #120 #235)
-#233 := (iff #117 #232)
-#230 := (= #114 #229)
-#227 := (= #111 #226)
-#224 := (iff #108 #221)
-#218 := (or #215 #207)
-#222 := (iff #218 #221)
-#223 := [rewrite]: #222
-#219 := (iff #108 #218)
-#216 := (iff #105 #207)
-#217 := [rewrite]: #216
-#204 := (iff #98 #215)
-#205 := [rewrite]: #204
-#220 := [monotonicity #205 #217]: #219
-#225 := [trans #220 #223]: #224
-#228 := [monotonicity #225]: #227
-#231 := [monotonicity #228]: #230
-#234 := [monotonicity #231]: #233
-#237 := [quant-intro #234]: #236
-#210 := (~ #120 #120)
-#208 := (~ #117 #117)
-#209 := [refl]: #208
-#211 := [nnf-pos #209]: #210
-#23 := (- #9)
-#22 := (- #8)
-#24 := (div #22 #23)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#25 := (ite #20 #21 #24)
-#26 := (ite #14 0::int #25)
-#27 := (= #10 #26)
-#28 := (forall (vars (?v0 int) (?v1 int)) #27)
-#123 := (iff #28 #120)
-#62 := (and #16 #18)
-#65 := (or #17 #62)
-#78 := (ite #65 #21 #75)
-#81 := (ite #14 0::int #78)
-#84 := (= #10 #81)
-#87 := (forall (vars (?v0 int) (?v1 int)) #84)
-#121 := (iff #87 #120)
-#118 := (iff #84 #117)
-#115 := (= #81 #114)
-#112 := (= #78 #111)
-#109 := (iff #65 #108)
-#106 := (iff #62 #105)
-#103 := (iff #18 #102)
-#104 := [rewrite]: #103
-#96 := (iff #16 #95)
-#97 := [rewrite]: #96
-#107 := [monotonicity #97 #104]: #106
-#99 := (iff #17 #98)
-#92 := (iff #15 #91)
-#93 := [rewrite]: #92
-#100 := [monotonicity #93 #97]: #99
-#110 := [monotonicity #100 #107]: #109
-#113 := [monotonicity #110]: #112
-#116 := [monotonicity #113]: #115
-#119 := [monotonicity #116]: #118
-#122 := [quant-intro #119]: #121
-#88 := (iff #28 #87)
-#85 := (iff #27 #84)
-#82 := (= #26 #81)
-#79 := (= #25 #78)
-#76 := (= #24 #75)
-#73 := (= #23 #72)
-#74 := [rewrite]: #73
-#70 := (= #22 #69)
-#71 := [rewrite]: #70
-#77 := [monotonicity #71 #74]: #76
-#66 := (iff #20 #65)
-#63 := (iff #19 #62)
-#64 := [rewrite]: #63
-#67 := [monotonicity #64]: #66
-#80 := [monotonicity #67 #77]: #79
-#83 := [monotonicity #80]: #82
-#86 := [monotonicity #83]: #85
-#89 := [quant-intro #86]: #88
-#124 := [trans #89 #122]: #123
-#61 := [asserted]: #28
-#125 := [mp #61 #124]: #120
-#201 := [mp~ #125 #211]: #120
-#238 := [mp #201 #237]: #235
-#269 := [mp #238 #268]: #266
-#762 := [mp #269 #761]: #757
-#680 := (not #757)
-#672 := (or #680 #180)
-#343 := (* -1::int 3::int)
-#427 := (* -1::int -3::int)
-#428 := (div #427 #343)
-#429 := (* -1::int #428)
-#420 := (+ #175 #429)
-#431 := (= #420 0::int)
-#432 := (div -3::int 3::int)
-#433 := (* -1::int #432)
-#430 := (+ #175 #433)
-#434 := (= #430 0::int)
-#413 := (<= 3::int 0::int)
-#748 := (<= -3::int 0::int)
-#750 := (or #748 #413)
-#407 := (not #750)
-#537 := (>= -3::int 0::int)
-#744 := (or #413 #537)
-#418 := (not #744)
-#419 := (or #418 #407)
-#751 := (ite #419 #434 #431)
-#752 := (= #175 0::int)
-#753 := (= 3::int 0::int)
-#754 := (= -3::int 0::int)
-#749 := (or #754 #753)
-#755 := (ite #749 #752 #751)
-#669 := (or #680 #755)
-#681 := (iff #669 #672)
-#683 := (iff #672 #672)
-#679 := [rewrite]: #683
-#691 := (iff #755 #180)
-#595 := (ite false #752 #180)
-#690 := (iff #595 #180)
-#687 := [rewrite]: #690
-#596 := (iff #755 #595)
-#593 := (iff #751 #180)
-#1 := true
-#574 := (ite true #180 #180)
-#592 := (iff #574 #180)
-#585 := [rewrite]: #592
-#589 := (iff #751 #574)
-#686 := (iff #431 #180)
-#714 := (+ 1::int #175)
-#558 := (= #714 0::int)
-#702 := (iff #558 #180)
-#703 := [rewrite]: #702
-#590 := (iff #431 #558)
-#540 := (= #420 #714)
-#713 := (+ #175 1::int)
-#715 := (= #713 #714)
-#710 := [rewrite]: #715
-#580 := (= #420 #713)
-#697 := (= #429 1::int)
-#435 := (* -1::int -1::int)
-#718 := (= #435 1::int)
-#722 := [rewrite]: #718
-#696 := (= #429 #435)
-#693 := (= #428 -1::int)
-#707 := (div 3::int -3::int)
-#539 := (= #707 -1::int)
-#692 := [rewrite]: #539
-#536 := (= #428 #707)
-#706 := (= #343 -3::int)
-#700 := [rewrite]: #706
-#699 := (= #427 3::int)
-#705 := [rewrite]: #699
-#538 := [monotonicity #705 #700]: #536
-#695 := [trans #538 #692]: #693
-#694 := [monotonicity #695]: #696
-#579 := [trans #694 #722]: #697
-#581 := [monotonicity #579]: #580
-#688 := [trans #581 #710]: #540
-#685 := [monotonicity #688]: #590
-#689 := [trans #685 #703]: #686
-#704 := (iff #434 #180)
-#559 := (iff #434 #558)
-#716 := (= #430 #714)
-#708 := (= #430 #713)
-#711 := (= #433 1::int)
-#717 := (= #433 #435)
-#719 := (= #432 -1::int)
-#720 := [rewrite]: #719
-#721 := [monotonicity #720]: #717
-#712 := [trans #721 #722]: #711
-#709 := [monotonicity #712]: #708
-#557 := [trans #709 #710]: #716
-#701 := [monotonicity #557]: #559
-#698 := [trans #701 #703]: #704
-#440 := (iff #419 true)
-#730 := (or true false)
-#454 := (iff #730 true)
-#455 := [rewrite]: #454
-#438 := (iff #419 #730)
-#724 := (iff #407 false)
-#726 := (not true)
-#723 := (iff #726 false)
-#727 := [rewrite]: #723
-#448 := (iff #407 #726)
-#349 := (iff #750 true)
-#452 := (iff #750 #730)
-#397 := (iff #413 false)
-#733 := [rewrite]: #397
-#743 := (iff #748 true)
-#729 := [rewrite]: #743
-#453 := [monotonicity #729 #733]: #452
-#725 := [trans #453 #455]: #349
-#449 := [monotonicity #725]: #448
-#728 := [trans #449 #727]: #724
-#742 := (iff #418 true)
-#376 := (not false)
-#381 := (iff #376 true)
-#382 := [rewrite]: #381
-#739 := (iff #418 #376)
-#732 := (iff #744 false)
-#406 := (or false false)
-#731 := (iff #406 false)
-#394 := [rewrite]: #731
-#736 := (iff #744 #406)
-#734 := (iff #537 false)
-#735 := [rewrite]: #734
-#737 := [monotonicity #733 #735]: #736
-#738 := [trans #737 #394]: #732
-#741 := [monotonicity #738]: #739
-#740 := [trans #741 #382]: #742
-#439 := [monotonicity #740 #728]: #438
-#441 := [trans #439 #455]: #440
-#591 := [monotonicity #441 #698 #689]: #589
-#594 := [trans #591 #585]: #593
-#395 := (iff #749 false)
-#747 := (iff #749 #406)
-#410 := (iff #753 false)
-#746 := [rewrite]: #410
-#745 := (iff #754 false)
-#405 := [rewrite]: #745
-#390 := [monotonicity #405 #746]: #747
-#396 := [trans #390 #394]: #395
-#597 := [monotonicity #396 #594]: #596
-#678 := [trans #597 #687]: #691
-#682 := [monotonicity #678]: #681
-#684 := [trans #682 #679]: #681
-#671 := [quant-inst]: #669
-#673 := [mp #671 #684]: #672
-[unit-resolution #673 #762 #198]: false
-unsat
-9617f71299996e3e5076f71a16d3a6ef69c7fc1f 335 0
+f77523eb025c744a899f70384fe23771c2e2df49 335 0
 #2 := false
 #179 := -2::int
 decl f4 :: (-> int int int)
@@ -38001,662 +39515,7 @@
 #666 := [mp #673 #665]: #671
 [unit-resolution #666 #759 #190]: false
 unsat
-ca3dae13e761e693f76c918c923ae5c553f3cdb1 326 0
-#2 := false
-#179 := -2::int
-decl f3 :: (-> int int int)
-#40 := 3::int
-#173 := -5::int
-#176 := (f3 -5::int 3::int)
-#182 := (= #176 -2::int)
-#185 := (not #182)
-#42 := 2::int
-#43 := (- 2::int)
-#38 := 5::int
-#39 := (- 5::int)
-#41 := (f3 #39 3::int)
-#44 := (= #41 #43)
-#45 := (not #44)
-#186 := (iff #45 #185)
-#183 := (iff #44 #182)
-#180 := (= #43 -2::int)
-#181 := [rewrite]: #180
-#177 := (= #41 #176)
-#174 := (= #39 -5::int)
-#175 := [rewrite]: #174
-#178 := [monotonicity #175]: #177
-#184 := [monotonicity #178 #181]: #183
-#187 := [monotonicity #184]: #186
-#172 := [asserted]: #45
-#190 := [mp #172 #187]: #185
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#10 := (f3 #8 #9)
-#747 := (pattern #10)
-#11 := 0::int
-#69 := -1::int
-#73 := (* -1::int #9)
-#70 := (* -1::int #8)
-#76 := (div #70 #73)
-#249 := (* -1::int #76)
-#250 := (+ #10 #249)
-#251 := (= #250 0::int)
-#21 := (div #8 #9)
-#246 := (* -1::int #21)
-#247 := (+ #10 #246)
-#248 := (= #247 0::int)
-#95 := (<= #9 0::int)
-#91 := (<= #8 0::int)
-#204 := (or #91 #95)
-#205 := (not #204)
-#102 := (>= #8 0::int)
-#196 := (or #95 #102)
-#197 := (not #196)
-#211 := (or #197 #205)
-#252 := (ite #211 #248 #251)
-#245 := (= #10 0::int)
-#13 := (= #9 0::int)
-#12 := (= #8 0::int)
-#14 := (or #12 #13)
-#253 := (ite #14 #245 #252)
-#748 := (forall (vars (?v0 int) (?v1 int)) (:pat #747) #253)
-#256 := (forall (vars (?v0 int) (?v1 int)) #253)
-#751 := (iff #256 #748)
-#749 := (iff #253 #253)
-#750 := [refl]: #749
-#752 := [quant-intro #750]: #751
-#216 := (ite #211 #21 #76)
-#219 := (ite #14 0::int #216)
-#222 := (= #10 #219)
-#225 := (forall (vars (?v0 int) (?v1 int)) #222)
-#257 := (iff #225 #256)
-#254 := (iff #222 #253)
-#255 := [rewrite]: #254
-#258 := [quant-intro #255]: #257
-#103 := (not #102)
-#96 := (not #95)
-#106 := (and #96 #103)
-#92 := (not #91)
-#99 := (and #92 #96)
-#109 := (or #99 #106)
-#112 := (ite #109 #21 #76)
-#115 := (ite #14 0::int #112)
-#118 := (= #10 #115)
-#121 := (forall (vars (?v0 int) (?v1 int)) #118)
-#226 := (iff #121 #225)
-#223 := (iff #118 #222)
-#220 := (= #115 #219)
-#217 := (= #112 #216)
-#214 := (iff #109 #211)
-#208 := (or #205 #197)
-#212 := (iff #208 #211)
-#213 := [rewrite]: #212
-#209 := (iff #109 #208)
-#206 := (iff #106 #197)
-#207 := [rewrite]: #206
-#194 := (iff #99 #205)
-#195 := [rewrite]: #194
-#210 := [monotonicity #195 #207]: #209
-#215 := [trans #210 #213]: #214
-#218 := [monotonicity #215]: #217
-#221 := [monotonicity #218]: #220
-#224 := [monotonicity #221]: #223
-#227 := [quant-intro #224]: #226
-#200 := (~ #121 #121)
-#198 := (~ #118 #118)
-#199 := [refl]: #198
-#201 := [nnf-pos #199]: #200
-#23 := (- #9)
-#22 := (- #8)
-#24 := (div #22 #23)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#25 := (ite #20 #21 #24)
-#26 := (ite #14 0::int #25)
-#27 := (= #10 #26)
-#28 := (forall (vars (?v0 int) (?v1 int)) #27)
-#124 := (iff #28 #121)
-#63 := (and #16 #18)
-#66 := (or #17 #63)
-#79 := (ite #66 #21 #76)
-#82 := (ite #14 0::int #79)
-#85 := (= #10 #82)
-#88 := (forall (vars (?v0 int) (?v1 int)) #85)
-#122 := (iff #88 #121)
-#119 := (iff #85 #118)
-#116 := (= #82 #115)
-#113 := (= #79 #112)
-#110 := (iff #66 #109)
-#107 := (iff #63 #106)
-#104 := (iff #18 #103)
-#105 := [rewrite]: #104
-#97 := (iff #16 #96)
-#98 := [rewrite]: #97
-#108 := [monotonicity #98 #105]: #107
-#100 := (iff #17 #99)
-#93 := (iff #15 #92)
-#94 := [rewrite]: #93
-#101 := [monotonicity #94 #98]: #100
-#111 := [monotonicity #101 #108]: #110
-#114 := [monotonicity #111]: #113
-#117 := [monotonicity #114]: #116
-#120 := [monotonicity #117]: #119
-#123 := [quant-intro #120]: #122
-#89 := (iff #28 #88)
-#86 := (iff #27 #85)
-#83 := (= #26 #82)
-#80 := (= #25 #79)
-#77 := (= #24 #76)
-#74 := (= #23 #73)
-#75 := [rewrite]: #74
-#71 := (= #22 #70)
-#72 := [rewrite]: #71
-#78 := [monotonicity #72 #75]: #77
-#67 := (iff #20 #66)
-#64 := (iff #19 #63)
-#65 := [rewrite]: #64
-#68 := [monotonicity #65]: #67
-#81 := [monotonicity #68 #78]: #80
-#84 := [monotonicity #81]: #83
-#87 := [monotonicity #84]: #86
-#90 := [quant-intro #87]: #89
-#125 := [trans #90 #123]: #124
-#62 := [asserted]: #28
-#126 := [mp #62 #125]: #121
-#191 := [mp~ #126 #201]: #121
-#228 := [mp #191 #227]: #225
-#259 := [mp #228 #258]: #256
-#753 := [mp #259 #752]: #748
-#667 := (not #748)
-#661 := (or #667 #182)
-#333 := (* -1::int 3::int)
-#418 := (* -1::int -5::int)
-#419 := (div #418 #333)
-#420 := (* -1::int #419)
-#411 := (+ #176 #420)
-#422 := (= #411 0::int)
-#423 := (div -5::int 3::int)
-#351 := (* -1::int #423)
-#424 := (+ #176 #351)
-#421 := (= #424 0::int)
-#425 := (<= 3::int 0::int)
-#404 := (<= -5::int 0::int)
-#739 := (or #404 #425)
-#741 := (not #739)
-#398 := (>= -5::int 0::int)
-#528 := (or #425 #398)
-#735 := (not #528)
-#409 := (or #735 #741)
-#410 := (ite #409 #421 #422)
-#742 := (= #176 0::int)
-#743 := (= 3::int 0::int)
-#744 := (= -5::int 0::int)
-#745 := (or #744 #743)
-#740 := (ite #745 #742 #410)
-#668 := (or #667 #740)
-#653 := (iff #668 #661)
-#656 := (iff #661 #661)
-#657 := [rewrite]: #656
-#665 := (iff #740 #182)
-#673 := (ite false #742 #182)
-#675 := (iff #673 #182)
-#664 := [rewrite]: #675
-#674 := (iff #740 #673)
-#662 := (iff #410 #182)
-#586 := (= #176 -1::int)
-#1 := true
-#682 := (ite true #182 #586)
-#663 := (iff #682 #182)
-#660 := [rewrite]: #663
-#669 := (iff #410 #682)
-#681 := (iff #422 #586)
-#570 := 1::int
-#680 := (+ 1::int #176)
-#576 := (= #680 0::int)
-#587 := (iff #576 #586)
-#588 := [rewrite]: #587
-#584 := (iff #422 #576)
-#582 := (= #411 #680)
-#581 := (+ #176 1::int)
-#565 := (= #581 #680)
-#580 := [rewrite]: #565
-#676 := (= #411 #581)
-#531 := (= #420 1::int)
-#687 := (* -1::int -1::int)
-#571 := (= #687 1::int)
-#572 := [rewrite]: #571
-#685 := (= #420 #687)
-#684 := (= #419 -1::int)
-#696 := -3::int
-#698 := (div 5::int -3::int)
-#530 := (= #698 -1::int)
-#683 := [rewrite]: #530
-#527 := (= #419 #698)
-#697 := (= #333 -3::int)
-#691 := [rewrite]: #697
-#689 := (= #418 5::int)
-#690 := [rewrite]: #689
-#529 := [monotonicity #690 #691]: #527
-#686 := [trans #529 #683]: #684
-#688 := [monotonicity #686]: #685
-#679 := [trans #688 #572]: #531
-#677 := [monotonicity #679]: #676
-#583 := [trans #677 #580]: #582
-#585 := [monotonicity #583]: #584
-#678 := [trans #585 #588]: #681
-#694 := (iff #421 #182)
-#700 := (+ 2::int #176)
-#548 := (= #700 0::int)
-#692 := (iff #548 #182)
-#693 := [rewrite]: #692
-#549 := (iff #421 #548)
-#701 := (= #424 #700)
-#703 := (+ #176 2::int)
-#705 := (= #703 #700)
-#706 := [rewrite]: #705
-#704 := (= #424 #703)
-#713 := (= #351 2::int)
-#711 := (* -1::int -2::int)
-#712 := (= #711 2::int)
-#709 := [rewrite]: #712
-#426 := (= #351 #711)
-#432 := (= #423 -2::int)
-#710 := [rewrite]: #432
-#708 := [monotonicity #710]: #426
-#702 := [trans #708 #709]: #713
-#699 := [monotonicity #702]: #704
-#707 := [trans #699 #706]: #701
-#550 := [monotonicity #707]: #549
-#695 := [trans #550 #693]: #694
-#430 := (iff #409 true)
-#720 := (or true false)
-#444 := (iff #720 true)
-#445 := [rewrite]: #444
-#719 := (iff #409 #720)
-#718 := (iff #741 false)
-#716 := (not true)
-#440 := (iff #716 false)
-#714 := [rewrite]: #440
-#717 := (iff #741 #716)
-#446 := (iff #739 true)
-#721 := (iff #739 #720)
-#387 := (iff #425 false)
-#388 := [rewrite]: #387
-#731 := (iff #404 true)
-#734 := [rewrite]: #731
-#443 := [monotonicity #734 #388]: #721
-#339 := [trans #443 #445]: #446
-#439 := [monotonicity #339]: #717
-#715 := [trans #439 #714]: #718
-#373 := (iff #735 true)
-#729 := (not false)
-#732 := (iff #729 true)
-#372 := [rewrite]: #732
-#367 := (iff #735 #729)
-#728 := (iff #528 false)
-#737 := (or false false)
-#381 := (iff #737 false)
-#722 := [rewrite]: #381
-#726 := (iff #528 #737)
-#724 := (iff #398 false)
-#725 := [rewrite]: #724
-#727 := [monotonicity #388 #725]: #726
-#723 := [trans #727 #722]: #728
-#730 := [monotonicity #723]: #367
-#733 := [trans #730 #372]: #373
-#429 := [monotonicity #733 #715]: #719
-#431 := [trans #429 #445]: #430
-#671 := [monotonicity #431 #695 #678]: #669
-#672 := [trans #671 #660]: #662
-#385 := (iff #745 false)
-#397 := (iff #745 #737)
-#396 := (iff #743 false)
-#401 := [rewrite]: #396
-#746 := (iff #744 false)
-#736 := [rewrite]: #746
-#738 := [monotonicity #736 #401]: #397
-#386 := [trans #738 #722]: #385
-#670 := [monotonicity #386 #672]: #674
-#666 := [trans #670 #664]: #665
-#655 := [monotonicity #666]: #653
-#658 := [trans #655 #657]: #653
-#652 := [quant-inst]: #668
-#654 := [mp #652 #658]: #661
-[unit-resolution #654 #753 #190]: false
-unsat
-3370d946d5a916a6204a00a35f20bfddd3071fdf 327 0
-#2 := false
-#11 := 0::int
-decl f3 :: (-> int int int)
-#174 := -3::int
-#68 := -1::int
-#177 := (f3 -1::int -3::int)
-#180 := (= #177 0::int)
-#193 := (not #180)
-#40 := 3::int
-#41 := (- 3::int)
-#38 := 1::int
-#39 := (- 1::int)
-#42 := (f3 #39 #41)
-#43 := (= #42 0::int)
-#44 := (not #43)
-#196 := (iff #44 #193)
-#183 := (= 0::int #177)
-#188 := (not #183)
-#194 := (iff #188 #193)
-#191 := (iff #183 #180)
-#192 := [rewrite]: #191
-#195 := [monotonicity #192]: #194
-#189 := (iff #44 #188)
-#186 := (iff #43 #183)
-#184 := (iff #180 #183)
-#185 := [rewrite]: #184
-#181 := (iff #43 #180)
-#178 := (= #42 #177)
-#175 := (= #41 -3::int)
-#176 := [rewrite]: #175
-#172 := (= #39 -1::int)
-#173 := [rewrite]: #172
-#179 := [monotonicity #173 #176]: #178
-#182 := [monotonicity #179]: #181
-#187 := [trans #182 #185]: #186
-#190 := [monotonicity #187]: #189
-#197 := [trans #190 #195]: #196
-#171 := [asserted]: #44
-#198 := [mp #171 #197]: #193
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#10 := (f3 #8 #9)
-#756 := (pattern #10)
-#72 := (* -1::int #9)
-#69 := (* -1::int #8)
-#75 := (div #69 #72)
-#259 := (* -1::int #75)
-#260 := (+ #10 #259)
-#261 := (= #260 0::int)
-#21 := (div #8 #9)
-#256 := (* -1::int #21)
-#257 := (+ #10 #256)
-#258 := (= #257 0::int)
-#94 := (<= #9 0::int)
-#90 := (<= #8 0::int)
-#214 := (or #90 #94)
-#215 := (not #214)
-#101 := (>= #8 0::int)
-#206 := (or #94 #101)
-#207 := (not #206)
-#221 := (or #207 #215)
-#262 := (ite #221 #258 #261)
-#255 := (= #10 0::int)
-#13 := (= #9 0::int)
-#12 := (= #8 0::int)
-#14 := (or #12 #13)
-#263 := (ite #14 #255 #262)
-#757 := (forall (vars (?v0 int) (?v1 int)) (:pat #756) #263)
-#266 := (forall (vars (?v0 int) (?v1 int)) #263)
-#760 := (iff #266 #757)
-#758 := (iff #263 #263)
-#759 := [refl]: #758
-#761 := [quant-intro #759]: #760
-#226 := (ite #221 #21 #75)
-#229 := (ite #14 0::int #226)
-#232 := (= #10 #229)
-#235 := (forall (vars (?v0 int) (?v1 int)) #232)
-#267 := (iff #235 #266)
-#264 := (iff #232 #263)
-#265 := [rewrite]: #264
-#268 := [quant-intro #265]: #267
-#102 := (not #101)
-#95 := (not #94)
-#105 := (and #95 #102)
-#91 := (not #90)
-#98 := (and #91 #95)
-#108 := (or #98 #105)
-#111 := (ite #108 #21 #75)
-#114 := (ite #14 0::int #111)
-#117 := (= #10 #114)
-#120 := (forall (vars (?v0 int) (?v1 int)) #117)
-#236 := (iff #120 #235)
-#233 := (iff #117 #232)
-#230 := (= #114 #229)
-#227 := (= #111 #226)
-#224 := (iff #108 #221)
-#218 := (or #215 #207)
-#222 := (iff #218 #221)
-#223 := [rewrite]: #222
-#219 := (iff #108 #218)
-#216 := (iff #105 #207)
-#217 := [rewrite]: #216
-#204 := (iff #98 #215)
-#205 := [rewrite]: #204
-#220 := [monotonicity #205 #217]: #219
-#225 := [trans #220 #223]: #224
-#228 := [monotonicity #225]: #227
-#231 := [monotonicity #228]: #230
-#234 := [monotonicity #231]: #233
-#237 := [quant-intro #234]: #236
-#210 := (~ #120 #120)
-#208 := (~ #117 #117)
-#209 := [refl]: #208
-#211 := [nnf-pos #209]: #210
-#23 := (- #9)
-#22 := (- #8)
-#24 := (div #22 #23)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#25 := (ite #20 #21 #24)
-#26 := (ite #14 0::int #25)
-#27 := (= #10 #26)
-#28 := (forall (vars (?v0 int) (?v1 int)) #27)
-#123 := (iff #28 #120)
-#62 := (and #16 #18)
-#65 := (or #17 #62)
-#78 := (ite #65 #21 #75)
-#81 := (ite #14 0::int #78)
-#84 := (= #10 #81)
-#87 := (forall (vars (?v0 int) (?v1 int)) #84)
-#121 := (iff #87 #120)
-#118 := (iff #84 #117)
-#115 := (= #81 #114)
-#112 := (= #78 #111)
-#109 := (iff #65 #108)
-#106 := (iff #62 #105)
-#103 := (iff #18 #102)
-#104 := [rewrite]: #103
-#96 := (iff #16 #95)
-#97 := [rewrite]: #96
-#107 := [monotonicity #97 #104]: #106
-#99 := (iff #17 #98)
-#92 := (iff #15 #91)
-#93 := [rewrite]: #92
-#100 := [monotonicity #93 #97]: #99
-#110 := [monotonicity #100 #107]: #109
-#113 := [monotonicity #110]: #112
-#116 := [monotonicity #113]: #115
-#119 := [monotonicity #116]: #118
-#122 := [quant-intro #119]: #121
-#88 := (iff #28 #87)
-#85 := (iff #27 #84)
-#82 := (= #26 #81)
-#79 := (= #25 #78)
-#76 := (= #24 #75)
-#73 := (= #23 #72)
-#74 := [rewrite]: #73
-#70 := (= #22 #69)
-#71 := [rewrite]: #70
-#77 := [monotonicity #71 #74]: #76
-#66 := (iff #20 #65)
-#63 := (iff #19 #62)
-#64 := [rewrite]: #63
-#67 := [monotonicity #64]: #66
-#80 := [monotonicity #67 #77]: #79
-#83 := [monotonicity #80]: #82
-#86 := [monotonicity #83]: #85
-#89 := [quant-intro #86]: #88
-#124 := [trans #89 #122]: #123
-#61 := [asserted]: #28
-#125 := [mp #61 #124]: #120
-#201 := [mp~ #125 #211]: #120
-#238 := [mp #201 #237]: #235
-#269 := [mp #238 #268]: #266
-#762 := [mp #269 #761]: #757
-#681 := (not #757)
-#682 := (or #681 #180)
-#343 := (* -1::int -3::int)
-#427 := (* -1::int -1::int)
-#428 := (div #427 #343)
-#429 := (* -1::int #428)
-#420 := (+ #177 #429)
-#431 := (= #420 0::int)
-#432 := (div -1::int -3::int)
-#433 := (* -1::int #432)
-#430 := (+ #177 #433)
-#434 := (= #430 0::int)
-#413 := (<= -3::int 0::int)
-#748 := (<= -1::int 0::int)
-#750 := (or #748 #413)
-#407 := (not #750)
-#537 := (>= -1::int 0::int)
-#744 := (or #413 #537)
-#418 := (not #744)
-#419 := (or #418 #407)
-#751 := (ite #419 #434 #431)
-#752 := (= -3::int 0::int)
-#753 := (= -1::int 0::int)
-#754 := (or #753 #752)
-#749 := (ite #754 #180 #751)
-#683 := (or #681 #749)
-#684 := (iff #683 #682)
-#674 := (iff #682 #682)
-#675 := [rewrite]: #674
-#669 := (iff #749 #180)
-#687 := (ite false #180 #180)
-#680 := (iff #687 #180)
-#672 := [rewrite]: #680
-#691 := (iff #749 #687)
-#597 := (iff #751 #180)
-#701 := (= #177 1::int)
-#585 := (ite false #701 #180)
-#595 := (iff #585 #180)
-#596 := [rewrite]: #595
-#593 := (iff #751 #585)
-#591 := (iff #431 #180)
-#574 := (= #420 #177)
-#688 := (+ #177 0::int)
-#686 := (= #688 #177)
-#689 := [rewrite]: #686
-#590 := (= #420 #688)
-#581 := (= #429 0::int)
-#696 := (* -1::int 0::int)
-#579 := (= #696 0::int)
-#580 := [rewrite]: #579
-#694 := (= #429 #696)
-#693 := (= #428 0::int)
-#707 := (div 1::int 3::int)
-#539 := (= #707 0::int)
-#692 := [rewrite]: #539
-#536 := (= #428 #707)
-#706 := (= #343 3::int)
-#700 := [rewrite]: #706
-#699 := (= #427 1::int)
-#705 := [rewrite]: #699
-#538 := [monotonicity #705 #700]: #536
-#695 := [trans #538 #692]: #693
-#697 := [monotonicity #695]: #694
-#540 := [trans #697 #580]: #581
-#685 := [monotonicity #540]: #590
-#589 := [trans #685 #689]: #574
-#592 := [monotonicity #589]: #591
-#704 := (iff #434 #701)
-#709 := (+ -1::int #177)
-#557 := (= #709 0::int)
-#702 := (iff #557 #701)
-#703 := [rewrite]: #702
-#558 := (iff #434 #557)
-#710 := (= #430 #709)
-#712 := (+ #177 -1::int)
-#714 := (= #712 #709)
-#715 := [rewrite]: #714
-#713 := (= #430 #712)
-#722 := (= #433 -1::int)
-#720 := (* -1::int 1::int)
-#721 := (= #720 -1::int)
-#718 := [rewrite]: #721
-#435 := (= #433 #720)
-#441 := (= #432 1::int)
-#719 := [rewrite]: #441
-#717 := [monotonicity #719]: #435
-#711 := [trans #717 #718]: #722
-#708 := [monotonicity #711]: #713
-#716 := [trans #708 #715]: #710
-#559 := [monotonicity #716]: #558
-#698 := [trans #559 #703]: #704
-#439 := (iff #419 false)
-#746 := (or false false)
-#390 := (iff #746 false)
-#731 := [rewrite]: #390
-#728 := (iff #419 #746)
-#727 := (iff #407 false)
-#1 := true
-#741 := (not true)
-#742 := (iff #741 false)
-#740 := [rewrite]: #742
-#449 := (iff #407 #741)
-#726 := (iff #750 true)
-#453 := (or true true)
-#349 := (iff #453 true)
-#725 := [rewrite]: #349
-#454 := (iff #750 #453)
-#396 := (iff #413 true)
-#397 := [rewrite]: #396
-#730 := (iff #748 true)
-#452 := [rewrite]: #730
-#455 := [monotonicity #452 #397]: #454
-#448 := [trans #455 #725]: #726
-#723 := [monotonicity #448]: #449
-#724 := [trans #723 #740]: #727
-#743 := (iff #418 false)
-#381 := (iff #418 #741)
-#376 := (iff #744 true)
-#735 := (or true false)
-#732 := (iff #735 true)
-#738 := [rewrite]: #732
-#736 := (iff #744 #735)
-#733 := (iff #537 false)
-#734 := [rewrite]: #733
-#737 := [monotonicity #397 #734]: #736
-#739 := [trans #737 #738]: #376
-#382 := [monotonicity #739]: #381
-#729 := [trans #382 #740]: #743
-#438 := [monotonicity #729 #724]: #728
-#440 := [trans #438 #731]: #439
-#594 := [monotonicity #440 #698 #592]: #593
-#690 := [trans #594 #596]: #597
-#394 := (iff #754 false)
-#406 := (iff #754 #746)
-#405 := (iff #752 false)
-#410 := [rewrite]: #405
-#755 := (iff #753 false)
-#745 := [rewrite]: #755
-#747 := [monotonicity #745 #410]: #406
-#395 := [trans #747 #731]: #394
-#678 := [monotonicity #395 #690]: #691
-#671 := [trans #678 #672]: #669
-#673 := [monotonicity #671]: #684
-#676 := [trans #673 #675]: #684
-#679 := [quant-inst]: #683
-#670 := [mp #679 #676]: #682
-[unit-resolution #670 #762 #198]: false
-unsat
-e0123afab77616719c99ab38af30fadddb3d3514 328 0
+51326f007e56e93f50afc5ba2d03581cdc697a2d 328 0
 #2 := false
 #11 := 0::int
 decl f4 :: (-> int int int)
@@ -38985,319 +39844,7 @@
 #688 := [mp #591 #684]: #582
 [unit-resolution #688 #766 #194]: false
 unsat
-c88cce63d6e0a186daa8cd66ece19fb97c1244ca 311 0
-#2 := false
-#41 := 1::int
-decl f3 :: (-> int int int)
-#171 := -3::int
-#174 := (f3 -3::int -3::int)
-#177 := (= #174 1::int)
-#190 := (not #177)
-#38 := 3::int
-#39 := (- 3::int)
-#40 := (f3 #39 #39)
-#42 := (= #40 1::int)
-#43 := (not #42)
-#193 := (iff #43 #190)
-#180 := (= 1::int #174)
-#185 := (not #180)
-#191 := (iff #185 #190)
-#188 := (iff #180 #177)
-#189 := [rewrite]: #188
-#192 := [monotonicity #189]: #191
-#186 := (iff #43 #185)
-#183 := (iff #42 #180)
-#181 := (iff #177 #180)
-#182 := [rewrite]: #181
-#178 := (iff #42 #177)
-#175 := (= #40 #174)
-#172 := (= #39 -3::int)
-#173 := [rewrite]: #172
-#176 := [monotonicity #173 #173]: #175
-#179 := [monotonicity #176]: #178
-#184 := [trans #179 #182]: #183
-#187 := [monotonicity #184]: #186
-#194 := [trans #187 #192]: #193
-#170 := [asserted]: #43
-#195 := [mp #170 #194]: #190
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#10 := (f3 #8 #9)
-#753 := (pattern #10)
-#11 := 0::int
-#67 := -1::int
-#71 := (* -1::int #9)
-#68 := (* -1::int #8)
-#74 := (div #68 #71)
-#256 := (* -1::int #74)
-#257 := (+ #10 #256)
-#258 := (= #257 0::int)
-#21 := (div #8 #9)
-#253 := (* -1::int #21)
-#254 := (+ #10 #253)
-#255 := (= #254 0::int)
-#93 := (<= #9 0::int)
-#89 := (<= #8 0::int)
-#211 := (or #89 #93)
-#212 := (not #211)
-#100 := (>= #8 0::int)
-#203 := (or #93 #100)
-#204 := (not #203)
-#218 := (or #204 #212)
-#259 := (ite #218 #255 #258)
-#252 := (= #10 0::int)
-#13 := (= #9 0::int)
-#12 := (= #8 0::int)
-#14 := (or #12 #13)
-#260 := (ite #14 #252 #259)
-#754 := (forall (vars (?v0 int) (?v1 int)) (:pat #753) #260)
-#263 := (forall (vars (?v0 int) (?v1 int)) #260)
-#757 := (iff #263 #754)
-#755 := (iff #260 #260)
-#756 := [refl]: #755
-#758 := [quant-intro #756]: #757
-#223 := (ite #218 #21 #74)
-#226 := (ite #14 0::int #223)
-#229 := (= #10 #226)
-#232 := (forall (vars (?v0 int) (?v1 int)) #229)
-#264 := (iff #232 #263)
-#261 := (iff #229 #260)
-#262 := [rewrite]: #261
-#265 := [quant-intro #262]: #264
-#101 := (not #100)
-#94 := (not #93)
-#104 := (and #94 #101)
-#90 := (not #89)
-#97 := (and #90 #94)
-#107 := (or #97 #104)
-#110 := (ite #107 #21 #74)
-#113 := (ite #14 0::int #110)
-#116 := (= #10 #113)
-#119 := (forall (vars (?v0 int) (?v1 int)) #116)
-#233 := (iff #119 #232)
-#230 := (iff #116 #229)
-#227 := (= #113 #226)
-#224 := (= #110 #223)
-#221 := (iff #107 #218)
-#215 := (or #212 #204)
-#219 := (iff #215 #218)
-#220 := [rewrite]: #219
-#216 := (iff #107 #215)
-#213 := (iff #104 #204)
-#214 := [rewrite]: #213
-#201 := (iff #97 #212)
-#202 := [rewrite]: #201
-#217 := [monotonicity #202 #214]: #216
-#222 := [trans #217 #220]: #221
-#225 := [monotonicity #222]: #224
-#228 := [monotonicity #225]: #227
-#231 := [monotonicity #228]: #230
-#234 := [quant-intro #231]: #233
-#207 := (~ #119 #119)
-#205 := (~ #116 #116)
-#206 := [refl]: #205
-#208 := [nnf-pos #206]: #207
-#23 := (- #9)
-#22 := (- #8)
-#24 := (div #22 #23)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#25 := (ite #20 #21 #24)
-#26 := (ite #14 0::int #25)
-#27 := (= #10 #26)
-#28 := (forall (vars (?v0 int) (?v1 int)) #27)
-#122 := (iff #28 #119)
-#61 := (and #16 #18)
-#64 := (or #17 #61)
-#77 := (ite #64 #21 #74)
-#80 := (ite #14 0::int #77)
-#83 := (= #10 #80)
-#86 := (forall (vars (?v0 int) (?v1 int)) #83)
-#120 := (iff #86 #119)
-#117 := (iff #83 #116)
-#114 := (= #80 #113)
-#111 := (= #77 #110)
-#108 := (iff #64 #107)
-#105 := (iff #61 #104)
-#102 := (iff #18 #101)
-#103 := [rewrite]: #102
-#95 := (iff #16 #94)
-#96 := [rewrite]: #95
-#106 := [monotonicity #96 #103]: #105
-#98 := (iff #17 #97)
-#91 := (iff #15 #90)
-#92 := [rewrite]: #91
-#99 := [monotonicity #92 #96]: #98
-#109 := [monotonicity #99 #106]: #108
-#112 := [monotonicity #109]: #111
-#115 := [monotonicity #112]: #114
-#118 := [monotonicity #115]: #117
-#121 := [quant-intro #118]: #120
-#87 := (iff #28 #86)
-#84 := (iff #27 #83)
-#81 := (= #26 #80)
-#78 := (= #25 #77)
-#75 := (= #24 #74)
-#72 := (= #23 #71)
-#73 := [rewrite]: #72
-#69 := (= #22 #68)
-#70 := [rewrite]: #69
-#76 := [monotonicity #70 #73]: #75
-#65 := (iff #20 #64)
-#62 := (iff #19 #61)
-#63 := [rewrite]: #62
-#66 := [monotonicity #63]: #65
-#79 := [monotonicity #66 #76]: #78
-#82 := [monotonicity #79]: #81
-#85 := [monotonicity #82]: #84
-#88 := [quant-intro #85]: #87
-#123 := [trans #88 #121]: #122
-#60 := [asserted]: #28
-#124 := [mp #60 #123]: #119
-#198 := [mp~ #124 #208]: #119
-#235 := [mp #198 #234]: #232
-#266 := [mp #235 #265]: #263
-#759 := [mp #266 #758]: #754
-#590 := (not #754)
-#591 := (or #590 #177)
-#340 := (* -1::int -3::int)
-#424 := (div #340 #340)
-#425 := (* -1::int #424)
-#426 := (+ #174 #425)
-#417 := (= #426 0::int)
-#428 := (div -3::int -3::int)
-#429 := (* -1::int #428)
-#430 := (+ #174 #429)
-#427 := (= #430 0::int)
-#431 := (<= -3::int 0::int)
-#410 := (or #431 #431)
-#745 := (not #410)
-#747 := (>= -3::int 0::int)
-#404 := (or #431 #747)
-#534 := (not #404)
-#741 := (or #534 #745)
-#415 := (ite #741 #427 #417)
-#416 := (= #174 0::int)
-#748 := (= -3::int 0::int)
-#749 := (or #748 #748)
-#750 := (ite #749 #416 #415)
-#592 := (or #590 #750)
-#594 := (iff #592 #591)
-#684 := (iff #591 #591)
-#688 := [rewrite]: #684
-#589 := (iff #750 #177)
-#683 := (ite false #416 #177)
-#586 := (iff #683 #177)
-#588 := [rewrite]: #586
-#686 := (iff #750 #683)
-#587 := (iff #415 #177)
-#576 := (ite false #177 #177)
-#537 := (iff #576 #177)
-#685 := [rewrite]: #537
-#577 := (iff #415 #576)
-#691 := (iff #417 #177)
-#715 := (+ -1::int #174)
-#705 := (= #715 0::int)
-#712 := (iff #705 #177)
-#707 := [rewrite]: #712
-#692 := (iff #417 #705)
-#689 := (= #426 #715)
-#432 := (+ #174 -1::int)
-#719 := (= #432 #715)
-#708 := [rewrite]: #719
-#535 := (= #426 #432)
-#704 := (= #425 -1::int)
-#725 := (* -1::int 1::int)
-#437 := (= #725 -1::int)
-#438 := [rewrite]: #437
-#703 := (= #425 #725)
-#696 := (= #424 1::int)
-#698 := (div 3::int 3::int)
-#701 := (= #698 1::int)
-#695 := [rewrite]: #701
-#699 := (= #424 #698)
-#555 := (= #340 3::int)
-#556 := [rewrite]: #555
-#700 := [monotonicity #556 #556]: #699
-#702 := [trans #700 #695]: #696
-#697 := [monotonicity #702]: #703
-#533 := [trans #697 #438]: #704
-#536 := [monotonicity #533]: #535
-#690 := [trans #536 #708]: #689
-#693 := [monotonicity #690]: #692
-#694 := [trans #693 #707]: #691
-#713 := (iff #427 #177)
-#706 := (iff #427 #705)
-#709 := (= #430 #715)
-#714 := (= #430 #432)
-#716 := (= #429 -1::int)
-#435 := (= #429 #725)
-#724 := (= #428 1::int)
-#721 := [rewrite]: #724
-#436 := [monotonicity #721]: #435
-#717 := [trans #436 #438]: #716
-#718 := [monotonicity #717]: #714
-#710 := [trans #718 #708]: #709
-#711 := [monotonicity #710]: #706
-#554 := [trans #711 #707]: #713
-#446 := (iff #741 false)
-#752 := (or false false)
-#407 := (iff #752 false)
-#743 := [rewrite]: #407
-#723 := (iff #741 #752)
-#346 := (iff #745 false)
-#1 := true
-#729 := (not true)
-#736 := (iff #729 false)
-#738 := [rewrite]: #736
-#451 := (iff #745 #729)
-#449 := (iff #410 true)
-#739 := (or true true)
-#726 := (iff #739 true)
-#727 := [rewrite]: #726
-#737 := (iff #410 #739)
-#387 := (iff #431 true)
-#728 := [rewrite]: #387
-#740 := [monotonicity #728 #728]: #737
-#450 := [trans #740 #727]: #449
-#452 := [monotonicity #450]: #451
-#722 := [trans #452 #738]: #346
-#378 := (iff #534 false)
-#735 := (iff #534 #729)
-#733 := (iff #404 true)
-#393 := (or true false)
-#731 := (iff #393 true)
-#732 := [rewrite]: #731
-#394 := (iff #404 #393)
-#391 := (iff #747 false)
-#392 := [rewrite]: #391
-#730 := [monotonicity #728 #392]: #394
-#734 := [trans #730 #732]: #733
-#373 := [monotonicity #734]: #735
-#379 := [trans #373 #738]: #378
-#445 := [monotonicity #379 #722]: #723
-#720 := [trans #445 #743]: #446
-#578 := [monotonicity #720 #554 #694]: #577
-#682 := [trans #578 #685]: #587
-#403 := (iff #749 false)
-#742 := (iff #749 #752)
-#751 := (iff #748 false)
-#746 := [rewrite]: #751
-#402 := [monotonicity #746 #746]: #742
-#744 := [trans #402 #743]: #403
-#571 := [monotonicity #744 #682]: #686
-#582 := [trans #571 #588]: #589
-#687 := [monotonicity #582]: #594
-#675 := [trans #687 #688]: #594
-#593 := [quant-inst]: #592
-#677 := [mp #593 #675]: #591
-[unit-resolution #677 #759 #195]: false
-unsat
-9635e0bc08ee97ed1256a38a5dbd874adffbbe2e 356 0
+0f74c41c938e9378ba24d8b61b7a4d0b50210370 356 0
 #2 := false
 #69 := -1::int
 decl f4 :: (-> int int int)
@@ -39654,346 +40201,7 @@
 #669 := [mp #663 #664]: #678
 [unit-resolution #669 #770 #199]: false
 unsat
-d7d527dac30acb7be44977e66eb01c6ffc6a5300 338 0
-#2 := false
-#43 := 1::int
-decl f3 :: (-> int int int)
-#176 := -3::int
-#173 := -5::int
-#179 := (f3 -5::int -3::int)
-#182 := (= #179 1::int)
-#195 := (not #182)
-#40 := 3::int
-#41 := (- 3::int)
-#38 := 5::int
-#39 := (- 5::int)
-#42 := (f3 #39 #41)
-#44 := (= #42 1::int)
-#45 := (not #44)
-#198 := (iff #45 #195)
-#185 := (= 1::int #179)
-#190 := (not #185)
-#196 := (iff #190 #195)
-#193 := (iff #185 #182)
-#194 := [rewrite]: #193
-#197 := [monotonicity #194]: #196
-#191 := (iff #45 #190)
-#188 := (iff #44 #185)
-#186 := (iff #182 #185)
-#187 := [rewrite]: #186
-#183 := (iff #44 #182)
-#180 := (= #42 #179)
-#177 := (= #41 -3::int)
-#178 := [rewrite]: #177
-#174 := (= #39 -5::int)
-#175 := [rewrite]: #174
-#181 := [monotonicity #175 #178]: #180
-#184 := [monotonicity #181]: #183
-#189 := [trans #184 #187]: #188
-#192 := [monotonicity #189]: #191
-#199 := [trans #192 #197]: #198
-#172 := [asserted]: #45
-#200 := [mp #172 #199]: #195
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#10 := (f3 #8 #9)
-#758 := (pattern #10)
-#11 := 0::int
-#69 := -1::int
-#73 := (* -1::int #9)
-#70 := (* -1::int #8)
-#76 := (div #70 #73)
-#261 := (* -1::int #76)
-#262 := (+ #10 #261)
-#263 := (= #262 0::int)
-#21 := (div #8 #9)
-#258 := (* -1::int #21)
-#259 := (+ #10 #258)
-#260 := (= #259 0::int)
-#95 := (<= #9 0::int)
-#91 := (<= #8 0::int)
-#216 := (or #91 #95)
-#217 := (not #216)
-#102 := (>= #8 0::int)
-#208 := (or #95 #102)
-#209 := (not #208)
-#223 := (or #209 #217)
-#264 := (ite #223 #260 #263)
-#257 := (= #10 0::int)
-#13 := (= #9 0::int)
-#12 := (= #8 0::int)
-#14 := (or #12 #13)
-#265 := (ite #14 #257 #264)
-#759 := (forall (vars (?v0 int) (?v1 int)) (:pat #758) #265)
-#268 := (forall (vars (?v0 int) (?v1 int)) #265)
-#762 := (iff #268 #759)
-#760 := (iff #265 #265)
-#761 := [refl]: #760
-#763 := [quant-intro #761]: #762
-#228 := (ite #223 #21 #76)
-#231 := (ite #14 0::int #228)
-#234 := (= #10 #231)
-#237 := (forall (vars (?v0 int) (?v1 int)) #234)
-#269 := (iff #237 #268)
-#266 := (iff #234 #265)
-#267 := [rewrite]: #266
-#270 := [quant-intro #267]: #269
-#103 := (not #102)
-#96 := (not #95)
-#106 := (and #96 #103)
-#92 := (not #91)
-#99 := (and #92 #96)
-#109 := (or #99 #106)
-#112 := (ite #109 #21 #76)
-#115 := (ite #14 0::int #112)
-#118 := (= #10 #115)
-#121 := (forall (vars (?v0 int) (?v1 int)) #118)
-#238 := (iff #121 #237)
-#235 := (iff #118 #234)
-#232 := (= #115 #231)
-#229 := (= #112 #228)
-#226 := (iff #109 #223)
-#220 := (or #217 #209)
-#224 := (iff #220 #223)
-#225 := [rewrite]: #224
-#221 := (iff #109 #220)
-#218 := (iff #106 #209)
-#219 := [rewrite]: #218
-#206 := (iff #99 #217)
-#207 := [rewrite]: #206
-#222 := [monotonicity #207 #219]: #221
-#227 := [trans #222 #225]: #226
-#230 := [monotonicity #227]: #229
-#233 := [monotonicity #230]: #232
-#236 := [monotonicity #233]: #235
-#239 := [quant-intro #236]: #238
-#212 := (~ #121 #121)
-#210 := (~ #118 #118)
-#211 := [refl]: #210
-#213 := [nnf-pos #211]: #212
-#23 := (- #9)
-#22 := (- #8)
-#24 := (div #22 #23)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#25 := (ite #20 #21 #24)
-#26 := (ite #14 0::int #25)
-#27 := (= #10 #26)
-#28 := (forall (vars (?v0 int) (?v1 int)) #27)
-#124 := (iff #28 #121)
-#63 := (and #16 #18)
-#66 := (or #17 #63)
-#79 := (ite #66 #21 #76)
-#82 := (ite #14 0::int #79)
-#85 := (= #10 #82)
-#88 := (forall (vars (?v0 int) (?v1 int)) #85)
-#122 := (iff #88 #121)
-#119 := (iff #85 #118)
-#116 := (= #82 #115)
-#113 := (= #79 #112)
-#110 := (iff #66 #109)
-#107 := (iff #63 #106)
-#104 := (iff #18 #103)
-#105 := [rewrite]: #104
-#97 := (iff #16 #96)
-#98 := [rewrite]: #97
-#108 := [monotonicity #98 #105]: #107
-#100 := (iff #17 #99)
-#93 := (iff #15 #92)
-#94 := [rewrite]: #93
-#101 := [monotonicity #94 #98]: #100
-#111 := [monotonicity #101 #108]: #110
-#114 := [monotonicity #111]: #113
-#117 := [monotonicity #114]: #116
-#120 := [monotonicity #117]: #119
-#123 := [quant-intro #120]: #122
-#89 := (iff #28 #88)
-#86 := (iff #27 #85)
-#83 := (= #26 #82)
-#80 := (= #25 #79)
-#77 := (= #24 #76)
-#74 := (= #23 #73)
-#75 := [rewrite]: #74
-#71 := (= #22 #70)
-#72 := [rewrite]: #71
-#78 := [monotonicity #72 #75]: #77
-#67 := (iff #20 #66)
-#64 := (iff #19 #63)
-#65 := [rewrite]: #64
-#68 := [monotonicity #65]: #67
-#81 := [monotonicity #68 #78]: #80
-#84 := [monotonicity #81]: #83
-#87 := [monotonicity #84]: #86
-#90 := [quant-intro #87]: #89
-#125 := [trans #90 #123]: #124
-#62 := [asserted]: #28
-#126 := [mp #62 #125]: #121
-#203 := [mp~ #126 #213]: #121
-#240 := [mp #203 #239]: #237
-#271 := [mp #240 #270]: #268
-#764 := [mp #271 #763]: #759
-#672 := (not #759)
-#679 := (or #672 #182)
-#345 := (* -1::int -3::int)
-#429 := (* -1::int -5::int)
-#430 := (div #429 #345)
-#431 := (* -1::int #430)
-#422 := (+ #179 #431)
-#433 := (= #422 0::int)
-#434 := (div -5::int -3::int)
-#435 := (* -1::int #434)
-#432 := (+ #179 #435)
-#436 := (= #432 0::int)
-#415 := (<= -3::int 0::int)
-#750 := (<= -5::int 0::int)
-#752 := (or #750 #415)
-#409 := (not #752)
-#539 := (>= -5::int 0::int)
-#746 := (or #415 #539)
-#420 := (not #746)
-#421 := (or #420 #409)
-#753 := (ite #421 #436 #433)
-#754 := (= #179 0::int)
-#755 := (= -3::int 0::int)
-#756 := (= -5::int 0::int)
-#751 := (or #756 #755)
-#757 := (ite #751 #754 #753)
-#663 := (or #672 #757)
-#666 := (iff #663 #679)
-#668 := (iff #679 #679)
-#669 := [rewrite]: #668
-#677 := (iff #757 #182)
-#685 := (ite false #754 #182)
-#675 := (iff #685 #182)
-#676 := [rewrite]: #675
-#681 := (iff #757 #685)
-#683 := (iff #753 #182)
-#721 := 2::int
-#706 := (= #179 2::int)
-#680 := (ite false #706 #182)
-#671 := (iff #680 #182)
-#673 := [rewrite]: #671
-#682 := (iff #753 #680)
-#689 := (iff #433 #182)
-#591 := (+ -1::int #179)
-#596 := (= #591 0::int)
-#599 := (iff #596 #182)
-#692 := [rewrite]: #599
-#597 := (iff #433 #596)
-#587 := (= #422 #591)
-#688 := (+ #179 -1::int)
-#593 := (= #688 #591)
-#594 := [rewrite]: #593
-#691 := (= #422 #688)
-#592 := (= #431 -1::int)
-#581 := (* -1::int 1::int)
-#542 := (= #581 -1::int)
-#690 := [rewrite]: #542
-#582 := (= #431 #581)
-#696 := (= #430 1::int)
-#541 := (div 5::int 3::int)
-#697 := (= #541 1::int)
-#698 := [rewrite]: #697
-#694 := (= #430 #541)
-#538 := (= #345 3::int)
-#540 := [rewrite]: #538
-#702 := (= #429 5::int)
-#709 := [rewrite]: #702
-#695 := [monotonicity #709 #540]: #694
-#699 := [trans #695 #698]: #696
-#583 := [monotonicity #699]: #582
-#687 := [trans #583 #690]: #592
-#576 := [monotonicity #687]: #691
-#595 := [trans #576 #594]: #587
-#598 := [monotonicity #595]: #597
-#693 := [trans #598 #692]: #689
-#707 := (iff #436 #706)
-#724 := -2::int
-#712 := (+ -2::int #179)
-#703 := (= #712 0::int)
-#700 := (iff #703 #706)
-#701 := [rewrite]: #700
-#704 := (iff #436 #703)
-#560 := (= #432 #712)
-#711 := (+ #179 -2::int)
-#718 := (= #711 #712)
-#559 := [rewrite]: #718
-#716 := (= #432 #711)
-#715 := (= #435 -2::int)
-#719 := (* -1::int 2::int)
-#713 := (= #719 -2::int)
-#714 := [rewrite]: #713
-#723 := (= #435 #719)
-#722 := (= #434 2::int)
-#437 := [rewrite]: #722
-#720 := [monotonicity #437]: #723
-#710 := [trans #720 #714]: #715
-#717 := [monotonicity #710]: #716
-#561 := [trans #717 #559]: #560
-#705 := [monotonicity #561]: #704
-#708 := [trans #705 #701]: #707
-#442 := (iff #421 false)
-#408 := (or false false)
-#733 := (iff #408 false)
-#396 := [rewrite]: #733
-#440 := (iff #421 #408)
-#726 := (iff #409 false)
-#1 := true
-#383 := (not true)
-#742 := (iff #383 false)
-#745 := [rewrite]: #742
-#725 := (iff #409 #383)
-#450 := (iff #752 true)
-#456 := (or true true)
-#727 := (iff #456 true)
-#728 := [rewrite]: #727
-#457 := (iff #752 #456)
-#399 := (iff #415 true)
-#735 := [rewrite]: #399
-#454 := (iff #750 true)
-#455 := [rewrite]: #454
-#351 := [monotonicity #455 #735]: #457
-#451 := [trans #351 #728]: #450
-#729 := [monotonicity #451]: #725
-#730 := [trans #729 #745]: #726
-#731 := (iff #420 false)
-#384 := (iff #420 #383)
-#741 := (iff #746 true)
-#738 := (or true false)
-#740 := (iff #738 true)
-#378 := [rewrite]: #740
-#739 := (iff #746 #738)
-#736 := (iff #539 false)
-#737 := [rewrite]: #736
-#734 := [monotonicity #735 #737]: #739
-#743 := [trans #734 #378]: #741
-#744 := [monotonicity #743]: #384
-#732 := [trans #744 #745]: #731
-#441 := [monotonicity #732 #730]: #440
-#443 := [trans #441 #396]: #442
-#674 := [monotonicity #443 #708 #693]: #682
-#684 := [trans #674 #673]: #683
-#397 := (iff #751 false)
-#749 := (iff #751 #408)
-#412 := (iff #755 false)
-#748 := [rewrite]: #412
-#747 := (iff #756 false)
-#407 := [rewrite]: #747
-#392 := [monotonicity #407 #748]: #749
-#398 := [trans #392 #396]: #397
-#686 := [monotonicity #398 #684]: #681
-#678 := [trans #686 #676]: #677
-#667 := [monotonicity #678]: #666
-#665 := [trans #667 #669]: #666
-#664 := [quant-inst]: #663
-#670 := [mp #664 #665]: #679
-[unit-resolution #670 #764 #200]: false
-unsat
-23ee07330031fda7a649c24c72b346abba5a15ed 350 0
+3c5a12d92626ff8ea2f3e7c2d04bebd45ef00916 350 0
 #2 := false
 #42 := 2::int
 decl f4 :: (-> int int int)
@@ -40344,7 +40552,2386 @@
 #664 := [mp #674 #663]: #673
 [unit-resolution #664 #766 #195]: false
 unsat
-5a36ea40ec6f132383fe7e0f161b9c3e56510423 592 0
+9ee8d0dc4ede7b68c19a265d2c4cda80156ee20b 331 0
+#2 := false
+#11 := 0::int
+decl f4 :: (-> int int int)
+#38 := 3::int
+#170 := -3::int
+#173 := (f4 -3::int 3::int)
+#176 := (= #173 0::int)
+#189 := (not #176)
+#39 := (- 3::int)
+#40 := (f4 #39 3::int)
+#41 := (= #40 0::int)
+#42 := (not #41)
+#192 := (iff #42 #189)
+#179 := (= 0::int #173)
+#184 := (not #179)
+#190 := (iff #184 #189)
+#187 := (iff #179 #176)
+#188 := [rewrite]: #187
+#191 := [monotonicity #188]: #190
+#185 := (iff #42 #184)
+#182 := (iff #41 #179)
+#180 := (iff #176 #179)
+#181 := [rewrite]: #180
+#177 := (iff #41 #176)
+#174 := (= #40 #173)
+#171 := (= #39 -3::int)
+#172 := [rewrite]: #171
+#175 := [monotonicity #172]: #174
+#178 := [monotonicity #175]: #177
+#183 := [trans #178 #181]: #182
+#186 := [monotonicity #183]: #185
+#193 := [trans #186 #191]: #192
+#169 := [asserted]: #42
+#194 := [mp #169 #193]: #189
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#29 := (f4 #8 #9)
+#760 := (pattern #29)
+#66 := -1::int
+#70 := (* -1::int #9)
+#67 := (* -1::int #8)
+#125 := (mod #67 #70)
+#271 := (+ #29 #125)
+#272 := (= #271 0::int)
+#30 := (mod #8 #9)
+#268 := (* -1::int #30)
+#269 := (+ #29 #268)
+#270 := (= #269 0::int)
+#92 := (<= #9 0::int)
+#88 := (<= #8 0::int)
+#210 := (or #88 #92)
+#211 := (not #210)
+#99 := (>= #8 0::int)
+#202 := (or #92 #99)
+#203 := (not #202)
+#217 := (or #203 #211)
+#273 := (ite #217 #270 #272)
+#267 := (= #29 0::int)
+#12 := (= #8 0::int)
+#274 := (ite #12 #267 #273)
+#266 := (= #8 #29)
+#13 := (= #9 0::int)
+#275 := (ite #13 #266 #274)
+#761 := (forall (vars (?v0 int) (?v1 int)) (:pat #760) #275)
+#278 := (forall (vars (?v0 int) (?v1 int)) #275)
+#764 := (iff #278 #761)
+#762 := (iff #275 #275)
+#763 := [refl]: #762
+#765 := [quant-intro #763]: #764
+#131 := (* -1::int #125)
+#235 := (ite #217 #30 #131)
+#238 := (ite #12 0::int #235)
+#241 := (ite #13 #8 #238)
+#244 := (= #29 #241)
+#247 := (forall (vars (?v0 int) (?v1 int)) #244)
+#279 := (iff #247 #278)
+#276 := (iff #244 #275)
+#277 := [rewrite]: #276
+#280 := [quant-intro #277]: #279
+#100 := (not #99)
+#93 := (not #92)
+#103 := (and #93 #100)
+#89 := (not #88)
+#96 := (and #89 #93)
+#106 := (or #96 #103)
+#151 := (ite #106 #30 #131)
+#154 := (ite #12 0::int #151)
+#157 := (ite #13 #8 #154)
+#160 := (= #29 #157)
+#163 := (forall (vars (?v0 int) (?v1 int)) #160)
+#248 := (iff #163 #247)
+#245 := (iff #160 #244)
+#242 := (= #157 #241)
+#239 := (= #154 #238)
+#236 := (= #151 #235)
+#220 := (iff #106 #217)
+#214 := (or #211 #203)
+#218 := (iff #214 #217)
+#219 := [rewrite]: #218
+#215 := (iff #106 #214)
+#212 := (iff #103 #203)
+#213 := [rewrite]: #212
+#200 := (iff #96 #211)
+#201 := [rewrite]: #200
+#216 := [monotonicity #201 #213]: #215
+#221 := [trans #216 #219]: #220
+#237 := [monotonicity #221]: #236
+#240 := [monotonicity #237]: #239
+#243 := [monotonicity #240]: #242
+#246 := [monotonicity #243]: #245
+#249 := [quant-intro #246]: #248
+#198 := (~ #163 #163)
+#195 := (~ #160 #160)
+#208 := [refl]: #195
+#199 := [nnf-pos #208]: #198
+#23 := (- #9)
+#22 := (- #8)
+#31 := (mod #22 #23)
+#32 := (- #31)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#33 := (ite #20 #30 #32)
+#34 := (ite #12 0::int #33)
+#35 := (ite #13 #8 #34)
+#36 := (= #29 #35)
+#37 := (forall (vars (?v0 int) (?v1 int)) #36)
+#166 := (iff #37 #163)
+#60 := (and #16 #18)
+#63 := (or #17 #60)
+#136 := (ite #63 #30 #131)
+#139 := (ite #12 0::int #136)
+#142 := (ite #13 #8 #139)
+#145 := (= #29 #142)
+#148 := (forall (vars (?v0 int) (?v1 int)) #145)
+#164 := (iff #148 #163)
+#161 := (iff #145 #160)
+#158 := (= #142 #157)
+#155 := (= #139 #154)
+#152 := (= #136 #151)
+#107 := (iff #63 #106)
+#104 := (iff #60 #103)
+#101 := (iff #18 #100)
+#102 := [rewrite]: #101
+#94 := (iff #16 #93)
+#95 := [rewrite]: #94
+#105 := [monotonicity #95 #102]: #104
+#97 := (iff #17 #96)
+#90 := (iff #15 #89)
+#91 := [rewrite]: #90
+#98 := [monotonicity #91 #95]: #97
+#108 := [monotonicity #98 #105]: #107
+#153 := [monotonicity #108]: #152
+#156 := [monotonicity #153]: #155
+#159 := [monotonicity #156]: #158
+#162 := [monotonicity #159]: #161
+#165 := [quant-intro #162]: #164
+#149 := (iff #37 #148)
+#146 := (iff #36 #145)
+#143 := (= #35 #142)
+#140 := (= #34 #139)
+#137 := (= #33 #136)
+#134 := (= #32 #131)
+#128 := (- #125)
+#132 := (= #128 #131)
+#133 := [rewrite]: #132
+#129 := (= #32 #128)
+#126 := (= #31 #125)
+#71 := (= #23 #70)
+#72 := [rewrite]: #71
+#68 := (= #22 #67)
+#69 := [rewrite]: #68
+#127 := [monotonicity #69 #72]: #126
+#130 := [monotonicity #127]: #129
+#135 := [trans #130 #133]: #134
+#64 := (iff #20 #63)
+#61 := (iff #19 #60)
+#62 := [rewrite]: #61
+#65 := [monotonicity #62]: #64
+#138 := [monotonicity #65 #135]: #137
+#141 := [monotonicity #138]: #140
+#144 := [monotonicity #141]: #143
+#147 := [monotonicity #144]: #146
+#150 := [quant-intro #147]: #149
+#167 := [trans #150 #165]: #166
+#124 := [asserted]: #37
+#168 := [mp #124 #167]: #163
+#209 := [mp~ #168 #199]: #163
+#250 := [mp #209 #249]: #247
+#281 := [mp #250 #280]: #278
+#766 := [mp #281 #765]: #761
+#591 := (not #761)
+#592 := (or #591 #176)
+#339 := (* -1::int 3::int)
+#424 := (* -1::int -3::int)
+#425 := (mod #424 #339)
+#426 := (+ #173 #425)
+#417 := (= #426 0::int)
+#428 := (mod -3::int 3::int)
+#429 := (* -1::int #428)
+#357 := (+ #173 #429)
+#430 := (= #357 0::int)
+#427 := (<= 3::int 0::int)
+#431 := (<= -3::int 0::int)
+#410 := (or #431 #427)
+#745 := (not #410)
+#747 := (>= -3::int 0::int)
+#404 := (or #427 #747)
+#534 := (not #404)
+#741 := (or #534 #745)
+#415 := (ite #741 #430 #417)
+#416 := (= -3::int 0::int)
+#748 := (ite #416 #176 #415)
+#749 := (= -3::int #173)
+#750 := (= 3::int 0::int)
+#751 := (ite #750 #749 #748)
+#593 := (or #591 #751)
+#687 := (iff #593 #592)
+#688 := (iff #592 #592)
+#675 := [rewrite]: #688
+#582 := (iff #751 #176)
+#402 := (= #173 -3::int)
+#686 := (ite false #402 #176)
+#588 := (iff #686 #176)
+#589 := [rewrite]: #588
+#571 := (iff #751 #686)
+#682 := (iff #748 #176)
+#577 := (ite false #176 #176)
+#685 := (iff #577 #176)
+#587 := [rewrite]: #685
+#578 := (iff #748 #577)
+#694 := (iff #415 #176)
+#1 := true
+#689 := (ite true #176 #176)
+#693 := (iff #689 #176)
+#691 := [rewrite]: #693
+#690 := (iff #415 #689)
+#535 := (iff #417 #176)
+#704 := (= #426 #173)
+#719 := (+ #173 0::int)
+#710 := (= #719 #173)
+#705 := [rewrite]: #710
+#703 := (= #426 #719)
+#696 := (= #425 0::int)
+#698 := (mod 3::int -3::int)
+#701 := (= #698 0::int)
+#695 := [rewrite]: #701
+#699 := (= #425 #698)
+#555 := (= #339 -3::int)
+#556 := [rewrite]: #555
+#713 := (= #424 3::int)
+#554 := [rewrite]: #713
+#700 := [monotonicity #554 #556]: #699
+#702 := [trans #700 #695]: #696
+#697 := [monotonicity #702]: #703
+#533 := [trans #697 #705]: #704
+#536 := [monotonicity #533]: #535
+#712 := (iff #430 #176)
+#706 := (= #357 #173)
+#708 := (= #357 #719)
+#718 := (= #429 0::int)
+#438 := (* -1::int 0::int)
+#432 := (= #438 0::int)
+#714 := [rewrite]: #432
+#716 := (= #429 #438)
+#436 := (= #428 0::int)
+#437 := [rewrite]: #436
+#717 := [monotonicity #437]: #716
+#715 := [trans #717 #714]: #718
+#709 := [monotonicity #715]: #708
+#711 := [trans #709 #705]: #706
+#707 := [monotonicity #711]: #712
+#725 := (iff #741 true)
+#737 := (or true false)
+#727 := (iff #737 true)
+#449 := [rewrite]: #727
+#724 := (iff #741 #737)
+#446 := (iff #745 false)
+#452 := (not true)
+#723 := (iff #452 false)
+#445 := [rewrite]: #723
+#345 := (iff #745 #452)
+#450 := (iff #410 true)
+#740 := (iff #410 #737)
+#744 := (iff #427 false)
+#387 := [rewrite]: #744
+#379 := (iff #431 true)
+#739 := [rewrite]: #379
+#726 := [monotonicity #739 #387]: #740
+#451 := [trans #726 #449]: #450
+#722 := [monotonicity #451]: #345
+#720 := [trans #722 #445]: #446
+#738 := (iff #534 true)
+#734 := (not false)
+#373 := (iff #734 true)
+#736 := [rewrite]: #373
+#729 := (iff #534 #734)
+#732 := (iff #404 false)
+#392 := (or false false)
+#730 := (iff #392 false)
+#731 := [rewrite]: #730
+#393 := (iff #404 #392)
+#728 := (iff #747 false)
+#391 := [rewrite]: #728
+#394 := [monotonicity #387 #391]: #393
+#733 := [trans #394 #731]: #732
+#735 := [monotonicity #733]: #729
+#378 := [trans #735 #736]: #738
+#721 := [monotonicity #378 #720]: #724
+#435 := [trans #721 #449]: #725
+#692 := [monotonicity #435 #707 #536]: #690
+#576 := [trans #692 #691]: #694
+#743 := (iff #416 false)
+#403 := [rewrite]: #743
+#537 := [monotonicity #403 #576]: #578
+#683 := [trans #537 #587]: #682
+#742 := (iff #749 #402)
+#407 := [rewrite]: #742
+#746 := (iff #750 false)
+#752 := [rewrite]: #746
+#586 := [monotonicity #752 #407 #683]: #571
+#590 := [trans #586 #589]: #582
+#684 := [monotonicity #590]: #687
+#677 := [trans #684 #675]: #687
+#594 := [quant-inst]: #593
+#669 := [mp #594 #677]: #592
+[unit-resolution #669 #766 #194]: false
+unsat
+0dff6836610ad24fd350792fdba18f67fffe78ae 353 0
+#2 := false
+#42 := 1::int
+decl f4 :: (-> int int int)
+#40 := 3::int
+#172 := -5::int
+#175 := (f4 -5::int 3::int)
+#178 := (= #175 1::int)
+#191 := (not #178)
+#38 := 5::int
+#39 := (- 5::int)
+#41 := (f4 #39 3::int)
+#43 := (= #41 1::int)
+#44 := (not #43)
+#194 := (iff #44 #191)
+#181 := (= 1::int #175)
+#186 := (not #181)
+#192 := (iff #186 #191)
+#189 := (iff #181 #178)
+#190 := [rewrite]: #189
+#193 := [monotonicity #190]: #192
+#187 := (iff #44 #186)
+#184 := (iff #43 #181)
+#182 := (iff #178 #181)
+#183 := [rewrite]: #182
+#179 := (iff #43 #178)
+#176 := (= #41 #175)
+#173 := (= #39 -5::int)
+#174 := [rewrite]: #173
+#177 := [monotonicity #174]: #176
+#180 := [monotonicity #177]: #179
+#185 := [trans #180 #183]: #184
+#188 := [monotonicity #185]: #187
+#195 := [trans #188 #193]: #194
+#171 := [asserted]: #44
+#196 := [mp #171 #195]: #191
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#29 := (f4 #8 #9)
+#761 := (pattern #29)
+#11 := 0::int
+#68 := -1::int
+#72 := (* -1::int #9)
+#69 := (* -1::int #8)
+#127 := (mod #69 #72)
+#273 := (+ #29 #127)
+#274 := (= #273 0::int)
+#30 := (mod #8 #9)
+#270 := (* -1::int #30)
+#271 := (+ #29 #270)
+#272 := (= #271 0::int)
+#94 := (<= #9 0::int)
+#90 := (<= #8 0::int)
+#212 := (or #90 #94)
+#213 := (not #212)
+#101 := (>= #8 0::int)
+#204 := (or #94 #101)
+#205 := (not #204)
+#219 := (or #205 #213)
+#275 := (ite #219 #272 #274)
+#269 := (= #29 0::int)
+#12 := (= #8 0::int)
+#276 := (ite #12 #269 #275)
+#268 := (= #8 #29)
+#13 := (= #9 0::int)
+#277 := (ite #13 #268 #276)
+#762 := (forall (vars (?v0 int) (?v1 int)) (:pat #761) #277)
+#280 := (forall (vars (?v0 int) (?v1 int)) #277)
+#765 := (iff #280 #762)
+#763 := (iff #277 #277)
+#764 := [refl]: #763
+#766 := [quant-intro #764]: #765
+#133 := (* -1::int #127)
+#237 := (ite #219 #30 #133)
+#240 := (ite #12 0::int #237)
+#243 := (ite #13 #8 #240)
+#246 := (= #29 #243)
+#249 := (forall (vars (?v0 int) (?v1 int)) #246)
+#281 := (iff #249 #280)
+#278 := (iff #246 #277)
+#279 := [rewrite]: #278
+#282 := [quant-intro #279]: #281
+#102 := (not #101)
+#95 := (not #94)
+#105 := (and #95 #102)
+#91 := (not #90)
+#98 := (and #91 #95)
+#108 := (or #98 #105)
+#153 := (ite #108 #30 #133)
+#156 := (ite #12 0::int #153)
+#159 := (ite #13 #8 #156)
+#162 := (= #29 #159)
+#165 := (forall (vars (?v0 int) (?v1 int)) #162)
+#250 := (iff #165 #249)
+#247 := (iff #162 #246)
+#244 := (= #159 #243)
+#241 := (= #156 #240)
+#238 := (= #153 #237)
+#222 := (iff #108 #219)
+#216 := (or #213 #205)
+#220 := (iff #216 #219)
+#221 := [rewrite]: #220
+#217 := (iff #108 #216)
+#214 := (iff #105 #205)
+#215 := [rewrite]: #214
+#202 := (iff #98 #213)
+#203 := [rewrite]: #202
+#218 := [monotonicity #203 #215]: #217
+#223 := [trans #218 #221]: #222
+#239 := [monotonicity #223]: #238
+#242 := [monotonicity #239]: #241
+#245 := [monotonicity #242]: #244
+#248 := [monotonicity #245]: #247
+#251 := [quant-intro #248]: #250
+#200 := (~ #165 #165)
+#197 := (~ #162 #162)
+#210 := [refl]: #197
+#201 := [nnf-pos #210]: #200
+#23 := (- #9)
+#22 := (- #8)
+#31 := (mod #22 #23)
+#32 := (- #31)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#33 := (ite #20 #30 #32)
+#34 := (ite #12 0::int #33)
+#35 := (ite #13 #8 #34)
+#36 := (= #29 #35)
+#37 := (forall (vars (?v0 int) (?v1 int)) #36)
+#168 := (iff #37 #165)
+#62 := (and #16 #18)
+#65 := (or #17 #62)
+#138 := (ite #65 #30 #133)
+#141 := (ite #12 0::int #138)
+#144 := (ite #13 #8 #141)
+#147 := (= #29 #144)
+#150 := (forall (vars (?v0 int) (?v1 int)) #147)
+#166 := (iff #150 #165)
+#163 := (iff #147 #162)
+#160 := (= #144 #159)
+#157 := (= #141 #156)
+#154 := (= #138 #153)
+#109 := (iff #65 #108)
+#106 := (iff #62 #105)
+#103 := (iff #18 #102)
+#104 := [rewrite]: #103
+#96 := (iff #16 #95)
+#97 := [rewrite]: #96
+#107 := [monotonicity #97 #104]: #106
+#99 := (iff #17 #98)
+#92 := (iff #15 #91)
+#93 := [rewrite]: #92
+#100 := [monotonicity #93 #97]: #99
+#110 := [monotonicity #100 #107]: #109
+#155 := [monotonicity #110]: #154
+#158 := [monotonicity #155]: #157
+#161 := [monotonicity #158]: #160
+#164 := [monotonicity #161]: #163
+#167 := [quant-intro #164]: #166
+#151 := (iff #37 #150)
+#148 := (iff #36 #147)
+#145 := (= #35 #144)
+#142 := (= #34 #141)
+#139 := (= #33 #138)
+#136 := (= #32 #133)
+#130 := (- #127)
+#134 := (= #130 #133)
+#135 := [rewrite]: #134
+#131 := (= #32 #130)
+#128 := (= #31 #127)
+#73 := (= #23 #72)
+#74 := [rewrite]: #73
+#70 := (= #22 #69)
+#71 := [rewrite]: #70
+#129 := [monotonicity #71 #74]: #128
+#132 := [monotonicity #129]: #131
+#137 := [trans #132 #135]: #136
+#66 := (iff #20 #65)
+#63 := (iff #19 #62)
+#64 := [rewrite]: #63
+#67 := [monotonicity #64]: #66
+#140 := [monotonicity #67 #137]: #139
+#143 := [monotonicity #140]: #142
+#146 := [monotonicity #143]: #145
+#149 := [monotonicity #146]: #148
+#152 := [quant-intro #149]: #151
+#169 := [trans #152 #167]: #168
+#126 := [asserted]: #37
+#170 := [mp #126 #169]: #165
+#211 := [mp~ #170 #201]: #165
+#252 := [mp #211 #251]: #249
+#283 := [mp #252 #282]: #280
+#767 := [mp #283 #766]: #762
+#668 := (not #762)
+#675 := (or #668 #178)
+#341 := (* -1::int 3::int)
+#425 := (* -1::int -5::int)
+#426 := (mod #425 #341)
+#427 := (+ #175 #426)
+#418 := (= #427 0::int)
+#429 := (mod -5::int 3::int)
+#430 := (* -1::int #429)
+#431 := (+ #175 #430)
+#428 := (= #431 0::int)
+#432 := (<= 3::int 0::int)
+#411 := (<= -5::int 0::int)
+#746 := (or #411 #432)
+#748 := (not #746)
+#405 := (>= -5::int 0::int)
+#535 := (or #432 #405)
+#742 := (not #535)
+#416 := (or #742 #748)
+#417 := (ite #416 #428 #418)
+#749 := (= #175 0::int)
+#750 := (= -5::int 0::int)
+#751 := (ite #750 #749 #417)
+#752 := (= -5::int #175)
+#747 := (= 3::int 0::int)
+#753 := (ite #747 #752 #751)
+#659 := (or #668 #753)
+#662 := (iff #659 #675)
+#664 := (iff #675 #675)
+#665 := [rewrite]: #664
+#673 := (iff #753 #178)
+#744 := (= #175 -5::int)
+#681 := (ite false #744 #178)
+#671 := (iff #681 #178)
+#672 := [rewrite]: #671
+#677 := (iff #753 #681)
+#679 := (iff #751 #178)
+#676 := (ite false #749 #178)
+#667 := (iff #676 #178)
+#669 := [rewrite]: #667
+#678 := (iff #751 #676)
+#685 := (iff #417 #178)
+#572 := -2::int
+#587 := (= #175 -2::int)
+#1 := true
+#592 := (ite true #178 #587)
+#595 := (iff #592 #178)
+#688 := [rewrite]: #595
+#593 := (iff #417 #592)
+#583 := (iff #418 #587)
+#537 := 2::int
+#578 := (+ 2::int #175)
+#683 := (= #578 0::int)
+#589 := (iff #683 #587)
+#590 := [rewrite]: #589
+#684 := (iff #418 #683)
+#686 := (= #427 #578)
+#692 := (+ #175 2::int)
+#579 := (= #692 #578)
+#538 := [rewrite]: #579
+#695 := (= #427 #692)
+#693 := (= #426 2::int)
+#703 := -3::int
+#705 := (mod 5::int -3::int)
+#690 := (= #705 2::int)
+#691 := [rewrite]: #690
+#534 := (= #426 #705)
+#704 := (= #341 -3::int)
+#698 := [rewrite]: #704
+#696 := (= #425 5::int)
+#697 := [rewrite]: #696
+#536 := [monotonicity #697 #698]: #534
+#694 := [trans #536 #691]: #693
+#577 := [monotonicity #694]: #695
+#588 := [trans #577 #538]: #686
+#687 := [monotonicity #588]: #684
+#591 := [trans #687 #590]: #583
+#701 := (iff #428 #178)
+#707 := (+ -1::int #175)
+#555 := (= #707 0::int)
+#699 := (iff #555 #178)
+#700 := [rewrite]: #699
+#556 := (iff #428 #555)
+#708 := (= #431 #707)
+#710 := (+ #175 -1::int)
+#712 := (= #710 #707)
+#713 := [rewrite]: #712
+#711 := (= #431 #710)
+#720 := (= #430 -1::int)
+#718 := (* -1::int 1::int)
+#719 := (= #718 -1::int)
+#716 := [rewrite]: #719
+#433 := (= #430 #718)
+#439 := (= #429 1::int)
+#717 := [rewrite]: #439
+#715 := [monotonicity #717]: #433
+#709 := [trans #715 #716]: #720
+#706 := [monotonicity #709]: #711
+#714 := [trans #706 #713]: #708
+#557 := [monotonicity #714]: #556
+#702 := [trans #557 #700]: #701
+#437 := (iff #416 true)
+#727 := (or true false)
+#451 := (iff #727 true)
+#452 := [rewrite]: #451
+#726 := (iff #416 #727)
+#725 := (iff #748 false)
+#723 := (not true)
+#447 := (iff #723 false)
+#721 := [rewrite]: #447
+#724 := (iff #748 #723)
+#453 := (iff #746 true)
+#728 := (iff #746 #727)
+#729 := (iff #432 false)
+#392 := [rewrite]: #729
+#738 := (iff #411 true)
+#741 := [rewrite]: #738
+#450 := [monotonicity #741 #392]: #728
+#347 := [trans #450 #452]: #453
+#446 := [monotonicity #347]: #724
+#722 := [trans #446 #721]: #725
+#380 := (iff #742 true)
+#736 := (not false)
+#739 := (iff #736 true)
+#379 := [rewrite]: #739
+#374 := (iff #742 #736)
+#735 := (iff #535 false)
+#395 := (or false false)
+#733 := (iff #395 false)
+#734 := [rewrite]: #733
+#731 := (iff #535 #395)
+#393 := (iff #405 false)
+#394 := [rewrite]: #393
+#732 := [monotonicity #392 #394]: #731
+#730 := [trans #732 #734]: #735
+#737 := [monotonicity #730]: #374
+#740 := [trans #737 #379]: #380
+#436 := [monotonicity #740 #722]: #726
+#438 := [trans #436 #452]: #437
+#594 := [monotonicity #438 #702 #591]: #593
+#689 := [trans #594 #688]: #685
+#745 := (iff #750 false)
+#388 := [rewrite]: #745
+#670 := [monotonicity #388 #689]: #678
+#680 := [trans #670 #669]: #679
+#408 := (iff #752 #744)
+#404 := [rewrite]: #408
+#743 := (iff #747 false)
+#403 := [rewrite]: #743
+#682 := [monotonicity #403 #404 #680]: #677
+#674 := [trans #682 #672]: #673
+#663 := [monotonicity #674]: #662
+#661 := [trans #663 #665]: #662
+#660 := [quant-inst]: #659
+#666 := [mp #660 #661]: #675
+[unit-resolution #666 #767 #196]: false
+unsat
+2131919aacf2b81143bdcdf048ab7c7b571f4dab 350 0
+#2 := false
+#68 := -1::int
+decl f4 :: (-> int int int)
+#174 := -3::int
+#177 := (f4 -1::int -3::int)
+#180 := (= #177 -1::int)
+#193 := (not #180)
+#38 := 1::int
+#39 := (- 1::int)
+#40 := 3::int
+#41 := (- 3::int)
+#42 := (f4 #39 #41)
+#43 := (= #42 #39)
+#44 := (not #43)
+#196 := (iff #44 #193)
+#183 := (= -1::int #177)
+#188 := (not #183)
+#194 := (iff #188 #193)
+#191 := (iff #183 #180)
+#192 := [rewrite]: #191
+#195 := [monotonicity #192]: #194
+#189 := (iff #44 #188)
+#186 := (iff #43 #183)
+#184 := (iff #180 #183)
+#185 := [rewrite]: #184
+#181 := (iff #43 #180)
+#172 := (= #39 -1::int)
+#173 := [rewrite]: #172
+#178 := (= #42 #177)
+#175 := (= #41 -3::int)
+#176 := [rewrite]: #175
+#179 := [monotonicity #173 #176]: #178
+#182 := [monotonicity #179 #173]: #181
+#187 := [trans #182 #185]: #186
+#190 := [monotonicity #187]: #189
+#197 := [trans #190 #195]: #196
+#171 := [asserted]: #44
+#198 := [mp #171 #197]: #193
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#29 := (f4 #8 #9)
+#763 := (pattern #29)
+#11 := 0::int
+#72 := (* -1::int #9)
+#69 := (* -1::int #8)
+#127 := (mod #69 #72)
+#275 := (+ #29 #127)
+#276 := (= #275 0::int)
+#30 := (mod #8 #9)
+#272 := (* -1::int #30)
+#273 := (+ #29 #272)
+#274 := (= #273 0::int)
+#94 := (<= #9 0::int)
+#90 := (<= #8 0::int)
+#214 := (or #90 #94)
+#215 := (not #214)
+#101 := (>= #8 0::int)
+#206 := (or #94 #101)
+#207 := (not #206)
+#221 := (or #207 #215)
+#277 := (ite #221 #274 #276)
+#271 := (= #29 0::int)
+#12 := (= #8 0::int)
+#278 := (ite #12 #271 #277)
+#270 := (= #8 #29)
+#13 := (= #9 0::int)
+#279 := (ite #13 #270 #278)
+#764 := (forall (vars (?v0 int) (?v1 int)) (:pat #763) #279)
+#282 := (forall (vars (?v0 int) (?v1 int)) #279)
+#767 := (iff #282 #764)
+#765 := (iff #279 #279)
+#766 := [refl]: #765
+#768 := [quant-intro #766]: #767
+#133 := (* -1::int #127)
+#239 := (ite #221 #30 #133)
+#242 := (ite #12 0::int #239)
+#245 := (ite #13 #8 #242)
+#248 := (= #29 #245)
+#251 := (forall (vars (?v0 int) (?v1 int)) #248)
+#283 := (iff #251 #282)
+#280 := (iff #248 #279)
+#281 := [rewrite]: #280
+#284 := [quant-intro #281]: #283
+#102 := (not #101)
+#95 := (not #94)
+#105 := (and #95 #102)
+#91 := (not #90)
+#98 := (and #91 #95)
+#108 := (or #98 #105)
+#153 := (ite #108 #30 #133)
+#156 := (ite #12 0::int #153)
+#159 := (ite #13 #8 #156)
+#162 := (= #29 #159)
+#165 := (forall (vars (?v0 int) (?v1 int)) #162)
+#252 := (iff #165 #251)
+#249 := (iff #162 #248)
+#246 := (= #159 #245)
+#243 := (= #156 #242)
+#240 := (= #153 #239)
+#224 := (iff #108 #221)
+#218 := (or #215 #207)
+#222 := (iff #218 #221)
+#223 := [rewrite]: #222
+#219 := (iff #108 #218)
+#216 := (iff #105 #207)
+#217 := [rewrite]: #216
+#204 := (iff #98 #215)
+#205 := [rewrite]: #204
+#220 := [monotonicity #205 #217]: #219
+#225 := [trans #220 #223]: #224
+#241 := [monotonicity #225]: #240
+#244 := [monotonicity #241]: #243
+#247 := [monotonicity #244]: #246
+#250 := [monotonicity #247]: #249
+#253 := [quant-intro #250]: #252
+#202 := (~ #165 #165)
+#199 := (~ #162 #162)
+#212 := [refl]: #199
+#203 := [nnf-pos #212]: #202
+#23 := (- #9)
+#22 := (- #8)
+#31 := (mod #22 #23)
+#32 := (- #31)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#33 := (ite #20 #30 #32)
+#34 := (ite #12 0::int #33)
+#35 := (ite #13 #8 #34)
+#36 := (= #29 #35)
+#37 := (forall (vars (?v0 int) (?v1 int)) #36)
+#168 := (iff #37 #165)
+#62 := (and #16 #18)
+#65 := (or #17 #62)
+#138 := (ite #65 #30 #133)
+#141 := (ite #12 0::int #138)
+#144 := (ite #13 #8 #141)
+#147 := (= #29 #144)
+#150 := (forall (vars (?v0 int) (?v1 int)) #147)
+#166 := (iff #150 #165)
+#163 := (iff #147 #162)
+#160 := (= #144 #159)
+#157 := (= #141 #156)
+#154 := (= #138 #153)
+#109 := (iff #65 #108)
+#106 := (iff #62 #105)
+#103 := (iff #18 #102)
+#104 := [rewrite]: #103
+#96 := (iff #16 #95)
+#97 := [rewrite]: #96
+#107 := [monotonicity #97 #104]: #106
+#99 := (iff #17 #98)
+#92 := (iff #15 #91)
+#93 := [rewrite]: #92
+#100 := [monotonicity #93 #97]: #99
+#110 := [monotonicity #100 #107]: #109
+#155 := [monotonicity #110]: #154
+#158 := [monotonicity #155]: #157
+#161 := [monotonicity #158]: #160
+#164 := [monotonicity #161]: #163
+#167 := [quant-intro #164]: #166
+#151 := (iff #37 #150)
+#148 := (iff #36 #147)
+#145 := (= #35 #144)
+#142 := (= #34 #141)
+#139 := (= #33 #138)
+#136 := (= #32 #133)
+#130 := (- #127)
+#134 := (= #130 #133)
+#135 := [rewrite]: #134
+#131 := (= #32 #130)
+#128 := (= #31 #127)
+#73 := (= #23 #72)
+#74 := [rewrite]: #73
+#70 := (= #22 #69)
+#71 := [rewrite]: #70
+#129 := [monotonicity #71 #74]: #128
+#132 := [monotonicity #129]: #131
+#137 := [trans #132 #135]: #136
+#66 := (iff #20 #65)
+#63 := (iff #19 #62)
+#64 := [rewrite]: #63
+#67 := [monotonicity #64]: #66
+#140 := [monotonicity #67 #137]: #139
+#143 := [monotonicity #140]: #142
+#146 := [monotonicity #143]: #145
+#149 := [monotonicity #146]: #148
+#152 := [quant-intro #149]: #151
+#169 := [trans #152 #167]: #168
+#126 := [asserted]: #37
+#170 := [mp #126 #169]: #165
+#213 := [mp~ #170 #203]: #165
+#254 := [mp #213 #253]: #251
+#285 := [mp #254 #284]: #282
+#769 := [mp #285 #768]: #764
+#684 := (not #764)
+#673 := (or #684 #180)
+#343 := (* -1::int -3::int)
+#427 := (* -1::int -1::int)
+#428 := (mod #427 #343)
+#429 := (+ #177 #428)
+#420 := (= #429 0::int)
+#431 := (mod -1::int -3::int)
+#432 := (* -1::int #431)
+#433 := (+ #177 #432)
+#430 := (= #433 0::int)
+#434 := (<= -3::int 0::int)
+#413 := (<= -1::int 0::int)
+#748 := (or #413 #434)
+#750 := (not #748)
+#407 := (>= -1::int 0::int)
+#537 := (or #434 #407)
+#744 := (not #537)
+#418 := (or #744 #750)
+#419 := (ite #418 #430 #420)
+#751 := (= #177 0::int)
+#752 := (= -1::int 0::int)
+#753 := (ite #752 #751 #419)
+#754 := (= -3::int 0::int)
+#749 := (ite #754 #183 #753)
+#674 := (or #684 #749)
+#676 := (iff #674 #673)
+#677 := (iff #673 #673)
+#661 := [rewrite]: #677
+#683 := (iff #749 #180)
+#672 := (ite false #180 #180)
+#681 := (iff #672 #180)
+#682 := [rewrite]: #681
+#669 := (iff #749 #672)
+#678 := (iff #753 #180)
+#596 := (ite false #751 #180)
+#687 := (iff #596 #180)
+#691 := [rewrite]: #687
+#597 := (iff #753 #596)
+#594 := (iff #419 #180)
+#728 := 2::int
+#558 := (= #177 2::int)
+#589 := (ite false #558 #180)
+#585 := (iff #589 #180)
+#593 := [rewrite]: #585
+#591 := (iff #419 #589)
+#689 := (iff #420 #180)
+#694 := (+ 1::int #177)
+#540 := (= #694 0::int)
+#685 := (iff #540 #180)
+#686 := [rewrite]: #685
+#688 := (iff #420 #540)
+#580 := (= #429 #694)
+#693 := (+ #177 1::int)
+#697 := (= #693 #694)
+#579 := [rewrite]: #697
+#695 := (= #429 #693)
+#539 := (= #428 1::int)
+#706 := (mod 1::int 3::int)
+#536 := (= #706 1::int)
+#538 := [rewrite]: #536
+#700 := (= #428 #706)
+#699 := (= #343 3::int)
+#705 := [rewrite]: #699
+#704 := (= #427 1::int)
+#698 := [rewrite]: #704
+#707 := [monotonicity #698 #705]: #700
+#692 := [trans #707 #538]: #539
+#696 := [monotonicity #692]: #695
+#581 := [trans #696 #579]: #580
+#590 := [monotonicity #581]: #688
+#574 := [trans #590 #686]: #689
+#702 := (iff #430 #558)
+#720 := -2::int
+#713 := (+ -2::int #177)
+#710 := (= #713 0::int)
+#559 := (iff #710 #558)
+#701 := [rewrite]: #559
+#716 := (iff #430 #710)
+#714 := (= #433 #713)
+#722 := (+ #177 -2::int)
+#708 := (= #722 #713)
+#709 := [rewrite]: #708
+#711 := (= #433 #722)
+#721 := (= #432 -2::int)
+#440 := (* -1::int 2::int)
+#435 := (= #440 -2::int)
+#717 := [rewrite]: #435
+#441 := (= #432 #440)
+#438 := (= #431 2::int)
+#439 := [rewrite]: #438
+#719 := [monotonicity #439]: #441
+#718 := [trans #719 #717]: #721
+#712 := [monotonicity #718]: #711
+#715 := [trans #712 #709]: #714
+#557 := [monotonicity #715]: #716
+#703 := [trans #557 #701]: #702
+#727 := (iff #418 false)
+#725 := (or false false)
+#449 := (iff #725 false)
+#723 := [rewrite]: #449
+#726 := (iff #418 #725)
+#455 := (iff #750 false)
+#1 := true
+#735 := (not true)
+#732 := (iff #735 false)
+#738 := [rewrite]: #732
+#453 := (iff #750 #735)
+#730 := (iff #748 true)
+#382 := (or true true)
+#743 := (iff #382 true)
+#729 := [rewrite]: #743
+#742 := (iff #748 #382)
+#746 := (iff #434 true)
+#406 := [rewrite]: #746
+#741 := (iff #413 true)
+#381 := [rewrite]: #741
+#740 := [monotonicity #381 #406]: #742
+#452 := [trans #740 #729]: #730
+#454 := [monotonicity #452]: #453
+#349 := [trans #454 #738]: #455
+#376 := (iff #744 false)
+#736 := (iff #744 #735)
+#733 := (iff #537 true)
+#731 := (or true false)
+#396 := (iff #731 true)
+#397 := [rewrite]: #396
+#394 := (iff #537 #731)
+#747 := (iff #407 false)
+#390 := [rewrite]: #747
+#395 := [monotonicity #406 #390]: #394
+#734 := [trans #395 #397]: #733
+#737 := [monotonicity #734]: #736
+#739 := [trans #737 #738]: #376
+#448 := [monotonicity #739 #349]: #726
+#724 := [trans #448 #723]: #727
+#592 := [monotonicity #724 #703 #574]: #591
+#595 := [trans #592 #593]: #594
+#405 := (iff #752 false)
+#410 := [rewrite]: #405
+#690 := [monotonicity #410 #595]: #597
+#680 := [trans #690 #691]: #678
+#755 := (iff #754 false)
+#745 := [rewrite]: #755
+#671 := [monotonicity #745 #192 #680]: #669
+#679 := [trans #671 #682]: #683
+#670 := [monotonicity #679]: #676
+#662 := [trans #670 #661]: #676
+#675 := [quant-inst]: #674
+#664 := [mp #675 #662]: #673
+[unit-resolution #664 #769 #198]: false
+unsat
+f185b775c0ed207acb3897bf916bd950213b79f9 319 0
+#2 := false
+#11 := 0::int
+decl f4 :: (-> int int int)
+#170 := -3::int
+#173 := (f4 -3::int -3::int)
+#176 := (= #173 0::int)
+#189 := (not #176)
+#38 := 3::int
+#39 := (- 3::int)
+#40 := (f4 #39 #39)
+#41 := (= #40 0::int)
+#42 := (not #41)
+#192 := (iff #42 #189)
+#179 := (= 0::int #173)
+#184 := (not #179)
+#190 := (iff #184 #189)
+#187 := (iff #179 #176)
+#188 := [rewrite]: #187
+#191 := [monotonicity #188]: #190
+#185 := (iff #42 #184)
+#182 := (iff #41 #179)
+#180 := (iff #176 #179)
+#181 := [rewrite]: #180
+#177 := (iff #41 #176)
+#174 := (= #40 #173)
+#171 := (= #39 -3::int)
+#172 := [rewrite]: #171
+#175 := [monotonicity #172 #172]: #174
+#178 := [monotonicity #175]: #177
+#183 := [trans #178 #181]: #182
+#186 := [monotonicity #183]: #185
+#193 := [trans #186 #191]: #192
+#169 := [asserted]: #42
+#194 := [mp #169 #193]: #189
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#29 := (f4 #8 #9)
+#760 := (pattern #29)
+#66 := -1::int
+#70 := (* -1::int #9)
+#67 := (* -1::int #8)
+#125 := (mod #67 #70)
+#271 := (+ #29 #125)
+#272 := (= #271 0::int)
+#30 := (mod #8 #9)
+#268 := (* -1::int #30)
+#269 := (+ #29 #268)
+#270 := (= #269 0::int)
+#92 := (<= #9 0::int)
+#88 := (<= #8 0::int)
+#210 := (or #88 #92)
+#211 := (not #210)
+#99 := (>= #8 0::int)
+#202 := (or #92 #99)
+#203 := (not #202)
+#217 := (or #203 #211)
+#273 := (ite #217 #270 #272)
+#267 := (= #29 0::int)
+#12 := (= #8 0::int)
+#274 := (ite #12 #267 #273)
+#266 := (= #8 #29)
+#13 := (= #9 0::int)
+#275 := (ite #13 #266 #274)
+#761 := (forall (vars (?v0 int) (?v1 int)) (:pat #760) #275)
+#278 := (forall (vars (?v0 int) (?v1 int)) #275)
+#764 := (iff #278 #761)
+#762 := (iff #275 #275)
+#763 := [refl]: #762
+#765 := [quant-intro #763]: #764
+#131 := (* -1::int #125)
+#235 := (ite #217 #30 #131)
+#238 := (ite #12 0::int #235)
+#241 := (ite #13 #8 #238)
+#244 := (= #29 #241)
+#247 := (forall (vars (?v0 int) (?v1 int)) #244)
+#279 := (iff #247 #278)
+#276 := (iff #244 #275)
+#277 := [rewrite]: #276
+#280 := [quant-intro #277]: #279
+#100 := (not #99)
+#93 := (not #92)
+#103 := (and #93 #100)
+#89 := (not #88)
+#96 := (and #89 #93)
+#106 := (or #96 #103)
+#151 := (ite #106 #30 #131)
+#154 := (ite #12 0::int #151)
+#157 := (ite #13 #8 #154)
+#160 := (= #29 #157)
+#163 := (forall (vars (?v0 int) (?v1 int)) #160)
+#248 := (iff #163 #247)
+#245 := (iff #160 #244)
+#242 := (= #157 #241)
+#239 := (= #154 #238)
+#236 := (= #151 #235)
+#220 := (iff #106 #217)
+#214 := (or #211 #203)
+#218 := (iff #214 #217)
+#219 := [rewrite]: #218
+#215 := (iff #106 #214)
+#212 := (iff #103 #203)
+#213 := [rewrite]: #212
+#200 := (iff #96 #211)
+#201 := [rewrite]: #200
+#216 := [monotonicity #201 #213]: #215
+#221 := [trans #216 #219]: #220
+#237 := [monotonicity #221]: #236
+#240 := [monotonicity #237]: #239
+#243 := [monotonicity #240]: #242
+#246 := [monotonicity #243]: #245
+#249 := [quant-intro #246]: #248
+#198 := (~ #163 #163)
+#195 := (~ #160 #160)
+#208 := [refl]: #195
+#199 := [nnf-pos #208]: #198
+#23 := (- #9)
+#22 := (- #8)
+#31 := (mod #22 #23)
+#32 := (- #31)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#33 := (ite #20 #30 #32)
+#34 := (ite #12 0::int #33)
+#35 := (ite #13 #8 #34)
+#36 := (= #29 #35)
+#37 := (forall (vars (?v0 int) (?v1 int)) #36)
+#166 := (iff #37 #163)
+#60 := (and #16 #18)
+#63 := (or #17 #60)
+#136 := (ite #63 #30 #131)
+#139 := (ite #12 0::int #136)
+#142 := (ite #13 #8 #139)
+#145 := (= #29 #142)
+#148 := (forall (vars (?v0 int) (?v1 int)) #145)
+#164 := (iff #148 #163)
+#161 := (iff #145 #160)
+#158 := (= #142 #157)
+#155 := (= #139 #154)
+#152 := (= #136 #151)
+#107 := (iff #63 #106)
+#104 := (iff #60 #103)
+#101 := (iff #18 #100)
+#102 := [rewrite]: #101
+#94 := (iff #16 #93)
+#95 := [rewrite]: #94
+#105 := [monotonicity #95 #102]: #104
+#97 := (iff #17 #96)
+#90 := (iff #15 #89)
+#91 := [rewrite]: #90
+#98 := [monotonicity #91 #95]: #97
+#108 := [monotonicity #98 #105]: #107
+#153 := [monotonicity #108]: #152
+#156 := [monotonicity #153]: #155
+#159 := [monotonicity #156]: #158
+#162 := [monotonicity #159]: #161
+#165 := [quant-intro #162]: #164
+#149 := (iff #37 #148)
+#146 := (iff #36 #145)
+#143 := (= #35 #142)
+#140 := (= #34 #139)
+#137 := (= #33 #136)
+#134 := (= #32 #131)
+#128 := (- #125)
+#132 := (= #128 #131)
+#133 := [rewrite]: #132
+#129 := (= #32 #128)
+#126 := (= #31 #125)
+#71 := (= #23 #70)
+#72 := [rewrite]: #71
+#68 := (= #22 #67)
+#69 := [rewrite]: #68
+#127 := [monotonicity #69 #72]: #126
+#130 := [monotonicity #127]: #129
+#135 := [trans #130 #133]: #134
+#64 := (iff #20 #63)
+#61 := (iff #19 #60)
+#62 := [rewrite]: #61
+#65 := [monotonicity #62]: #64
+#138 := [monotonicity #65 #135]: #137
+#141 := [monotonicity #138]: #140
+#144 := [monotonicity #141]: #143
+#147 := [monotonicity #144]: #146
+#150 := [quant-intro #147]: #149
+#167 := [trans #150 #165]: #166
+#124 := [asserted]: #37
+#168 := [mp #124 #167]: #163
+#209 := [mp~ #168 #199]: #163
+#250 := [mp #209 #249]: #247
+#281 := [mp #250 #280]: #278
+#766 := [mp #281 #765]: #761
+#537 := (not #761)
+#685 := (or #537 #176)
+#339 := (* -1::int -3::int)
+#424 := (mod #339 #339)
+#425 := (+ #173 #424)
+#426 := (= #425 0::int)
+#417 := (mod -3::int -3::int)
+#428 := (* -1::int #417)
+#429 := (+ #173 #428)
+#357 := (= #429 0::int)
+#430 := (<= -3::int 0::int)
+#427 := (or #430 #430)
+#431 := (not #427)
+#410 := (>= -3::int 0::int)
+#745 := (or #430 #410)
+#747 := (not #745)
+#404 := (or #747 #431)
+#534 := (ite #404 #357 #426)
+#741 := (= -3::int 0::int)
+#415 := (ite #741 #176 #534)
+#416 := (= -3::int #173)
+#748 := (ite #741 #416 #415)
+#587 := (or #537 #748)
+#683 := (iff #587 #685)
+#571 := (iff #685 #685)
+#586 := [rewrite]: #571
+#577 := (iff #748 #176)
+#746 := (= #173 -3::int)
+#692 := (ite false #746 #176)
+#694 := (iff #692 #176)
+#576 := [rewrite]: #694
+#693 := (iff #748 #692)
+#689 := (iff #415 #176)
+#695 := (ite false #176 #176)
+#703 := (iff #695 #176)
+#697 := [rewrite]: #703
+#535 := (iff #415 #695)
+#704 := (iff #534 #176)
+#696 := (iff #534 #695)
+#700 := (iff #426 #176)
+#698 := (= #425 #173)
+#438 := (+ #173 0::int)
+#432 := (= #438 #173)
+#714 := [rewrite]: #432
+#555 := (= #425 #438)
+#713 := (= #424 0::int)
+#705 := (mod 3::int 3::int)
+#712 := (= #705 0::int)
+#707 := [rewrite]: #712
+#706 := (= #424 #705)
+#709 := (= #339 3::int)
+#710 := [rewrite]: #709
+#711 := [monotonicity #710 #710]: #706
+#554 := [trans #711 #707]: #713
+#556 := [monotonicity #554]: #555
+#699 := [trans #556 #714]: #698
+#701 := [monotonicity #699]: #700
+#719 := (iff #357 #176)
+#718 := (= #429 #173)
+#716 := (= #429 #438)
+#436 := (= #428 0::int)
+#720 := (* -1::int 0::int)
+#725 := (= #720 0::int)
+#435 := [rewrite]: #725
+#724 := (= #428 #720)
+#445 := (= #417 0::int)
+#446 := [rewrite]: #445
+#721 := [monotonicity #446]: #724
+#437 := [trans #721 #435]: #436
+#717 := [monotonicity #437]: #716
+#715 := [trans #717 #714]: #718
+#708 := [monotonicity #715]: #719
+#722 := (iff #404 false)
+#449 := (or false false)
+#452 := (iff #449 false)
+#345 := [rewrite]: #452
+#450 := (iff #404 #449)
+#726 := (iff #431 false)
+#1 := true
+#394 := (not true)
+#732 := (iff #394 false)
+#733 := [rewrite]: #732
+#737 := (iff #431 #394)
+#379 := (iff #427 true)
+#735 := (or true true)
+#738 := (iff #735 true)
+#378 := [rewrite]: #738
+#373 := (iff #427 #735)
+#742 := (iff #430 true)
+#402 := [rewrite]: #742
+#736 := [monotonicity #402 #402]: #373
+#739 := [trans #736 #378]: #379
+#740 := [monotonicity #739]: #737
+#727 := [trans #740 #733]: #726
+#734 := (iff #747 false)
+#730 := (iff #747 #394)
+#392 := (iff #745 true)
+#403 := (or true false)
+#728 := (iff #403 true)
+#391 := [rewrite]: #728
+#744 := (iff #745 #403)
+#407 := (iff #410 false)
+#743 := [rewrite]: #407
+#387 := [monotonicity #402 #743]: #744
+#393 := [trans #387 #391]: #392
+#731 := [monotonicity #393]: #730
+#729 := [trans #731 #733]: #734
+#451 := [monotonicity #729 #727]: #450
+#723 := [trans #451 #345]: #722
+#702 := [monotonicity #723 #708 #701]: #696
+#533 := [trans #702 #697]: #704
+#749 := (iff #741 false)
+#750 := [rewrite]: #749
+#536 := [monotonicity #750 #533]: #535
+#690 := [trans #536 #697]: #689
+#751 := (iff #416 #746)
+#752 := [rewrite]: #751
+#691 := [monotonicity #750 #752 #690]: #693
+#578 := [trans #691 #576]: #577
+#686 := [monotonicity #578]: #683
+#588 := [trans #686 #586]: #683
+#682 := [quant-inst]: #587
+#589 := [mp #682 #588]: #685
+[unit-resolution #589 #766 #194]: false
+unsat
+fc9f250b0540aecbc3a008bfe364829a1228d272 347 0
+#2 := false
+#183 := -2::int
+decl f4 :: (-> int int int)
+#177 := -3::int
+#174 := -5::int
+#180 := (f4 -5::int -3::int)
+#186 := (= #180 -2::int)
+#189 := (not #186)
+#43 := 2::int
+#44 := (- 2::int)
+#40 := 3::int
+#41 := (- 3::int)
+#38 := 5::int
+#39 := (- 5::int)
+#42 := (f4 #39 #41)
+#45 := (= #42 #44)
+#46 := (not #45)
+#190 := (iff #46 #189)
+#187 := (iff #45 #186)
+#184 := (= #44 -2::int)
+#185 := [rewrite]: #184
+#181 := (= #42 #180)
+#178 := (= #41 -3::int)
+#179 := [rewrite]: #178
+#175 := (= #39 -5::int)
+#176 := [rewrite]: #175
+#182 := [monotonicity #176 #179]: #181
+#188 := [monotonicity #182 #185]: #187
+#191 := [monotonicity #188]: #190
+#173 := [asserted]: #46
+#194 := [mp #173 #191]: #189
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#29 := (f4 #8 #9)
+#758 := (pattern #29)
+#11 := 0::int
+#70 := -1::int
+#74 := (* -1::int #9)
+#71 := (* -1::int #8)
+#129 := (mod #71 #74)
+#269 := (+ #29 #129)
+#270 := (= #269 0::int)
+#30 := (mod #8 #9)
+#266 := (* -1::int #30)
+#267 := (+ #29 #266)
+#268 := (= #267 0::int)
+#96 := (<= #9 0::int)
+#92 := (<= #8 0::int)
+#208 := (or #92 #96)
+#209 := (not #208)
+#103 := (>= #8 0::int)
+#200 := (or #96 #103)
+#201 := (not #200)
+#215 := (or #201 #209)
+#271 := (ite #215 #268 #270)
+#265 := (= #29 0::int)
+#12 := (= #8 0::int)
+#272 := (ite #12 #265 #271)
+#264 := (= #8 #29)
+#13 := (= #9 0::int)
+#273 := (ite #13 #264 #272)
+#759 := (forall (vars (?v0 int) (?v1 int)) (:pat #758) #273)
+#276 := (forall (vars (?v0 int) (?v1 int)) #273)
+#762 := (iff #276 #759)
+#760 := (iff #273 #273)
+#761 := [refl]: #760
+#763 := [quant-intro #761]: #762
+#135 := (* -1::int #129)
+#233 := (ite #215 #30 #135)
+#236 := (ite #12 0::int #233)
+#239 := (ite #13 #8 #236)
+#242 := (= #29 #239)
+#245 := (forall (vars (?v0 int) (?v1 int)) #242)
+#277 := (iff #245 #276)
+#274 := (iff #242 #273)
+#275 := [rewrite]: #274
+#278 := [quant-intro #275]: #277
+#104 := (not #103)
+#97 := (not #96)
+#107 := (and #97 #104)
+#93 := (not #92)
+#100 := (and #93 #97)
+#110 := (or #100 #107)
+#155 := (ite #110 #30 #135)
+#158 := (ite #12 0::int #155)
+#161 := (ite #13 #8 #158)
+#164 := (= #29 #161)
+#167 := (forall (vars (?v0 int) (?v1 int)) #164)
+#246 := (iff #167 #245)
+#243 := (iff #164 #242)
+#240 := (= #161 #239)
+#237 := (= #158 #236)
+#234 := (= #155 #233)
+#218 := (iff #110 #215)
+#212 := (or #209 #201)
+#216 := (iff #212 #215)
+#217 := [rewrite]: #216
+#213 := (iff #110 #212)
+#210 := (iff #107 #201)
+#211 := [rewrite]: #210
+#198 := (iff #100 #209)
+#199 := [rewrite]: #198
+#214 := [monotonicity #199 #211]: #213
+#219 := [trans #214 #217]: #218
+#235 := [monotonicity #219]: #234
+#238 := [monotonicity #235]: #237
+#241 := [monotonicity #238]: #240
+#244 := [monotonicity #241]: #243
+#247 := [quant-intro #244]: #246
+#196 := (~ #167 #167)
+#192 := (~ #164 #164)
+#206 := [refl]: #192
+#197 := [nnf-pos #206]: #196
+#23 := (- #9)
+#22 := (- #8)
+#31 := (mod #22 #23)
+#32 := (- #31)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#33 := (ite #20 #30 #32)
+#34 := (ite #12 0::int #33)
+#35 := (ite #13 #8 #34)
+#36 := (= #29 #35)
+#37 := (forall (vars (?v0 int) (?v1 int)) #36)
+#170 := (iff #37 #167)
+#64 := (and #16 #18)
+#67 := (or #17 #64)
+#140 := (ite #67 #30 #135)
+#143 := (ite #12 0::int #140)
+#146 := (ite #13 #8 #143)
+#149 := (= #29 #146)
+#152 := (forall (vars (?v0 int) (?v1 int)) #149)
+#168 := (iff #152 #167)
+#165 := (iff #149 #164)
+#162 := (= #146 #161)
+#159 := (= #143 #158)
+#156 := (= #140 #155)
+#111 := (iff #67 #110)
+#108 := (iff #64 #107)
+#105 := (iff #18 #104)
+#106 := [rewrite]: #105
+#98 := (iff #16 #97)
+#99 := [rewrite]: #98
+#109 := [monotonicity #99 #106]: #108
+#101 := (iff #17 #100)
+#94 := (iff #15 #93)
+#95 := [rewrite]: #94
+#102 := [monotonicity #95 #99]: #101
+#112 := [monotonicity #102 #109]: #111
+#157 := [monotonicity #112]: #156
+#160 := [monotonicity #157]: #159
+#163 := [monotonicity #160]: #162
+#166 := [monotonicity #163]: #165
+#169 := [quant-intro #166]: #168
+#153 := (iff #37 #152)
+#150 := (iff #36 #149)
+#147 := (= #35 #146)
+#144 := (= #34 #143)
+#141 := (= #33 #140)
+#138 := (= #32 #135)
+#132 := (- #129)
+#136 := (= #132 #135)
+#137 := [rewrite]: #136
+#133 := (= #32 #132)
+#130 := (= #31 #129)
+#75 := (= #23 #74)
+#76 := [rewrite]: #75
+#72 := (= #22 #71)
+#73 := [rewrite]: #72
+#131 := [monotonicity #73 #76]: #130
+#134 := [monotonicity #131]: #133
+#139 := [trans #134 #137]: #138
+#68 := (iff #20 #67)
+#65 := (iff #19 #64)
+#66 := [rewrite]: #65
+#69 := [monotonicity #66]: #68
+#142 := [monotonicity #69 #139]: #141
+#145 := [monotonicity #142]: #144
+#148 := [monotonicity #145]: #147
+#151 := [monotonicity #148]: #150
+#154 := [quant-intro #151]: #153
+#171 := [trans #154 #169]: #170
+#128 := [asserted]: #37
+#172 := [mp #128 #171]: #167
+#207 := [mp~ #172 #197]: #167
+#248 := [mp #207 #247]: #245
+#279 := [mp #248 #278]: #276
+#764 := [mp #279 #763]: #759
+#669 := (not #759)
+#670 := (or #669 #186)
+#337 := (* -1::int -3::int)
+#422 := (* -1::int -5::int)
+#423 := (mod #422 #337)
+#424 := (+ #180 #423)
+#415 := (= #424 0::int)
+#426 := (mod -5::int -3::int)
+#427 := (* -1::int #426)
+#355 := (+ #180 #427)
+#428 := (= #355 0::int)
+#425 := (<= -3::int 0::int)
+#429 := (<= -5::int 0::int)
+#408 := (or #429 #425)
+#743 := (not #408)
+#745 := (>= -5::int 0::int)
+#402 := (or #425 #745)
+#532 := (not #402)
+#739 := (or #532 #743)
+#413 := (ite #739 #428 #415)
+#414 := (= #180 0::int)
+#746 := (= -5::int 0::int)
+#747 := (ite #746 #414 #413)
+#748 := (= -5::int #180)
+#749 := (= -3::int 0::int)
+#744 := (ite #749 #748 #747)
+#671 := (or #669 #744)
+#672 := (iff #671 #670)
+#657 := (iff #670 #670)
+#659 := [rewrite]: #657
+#679 := (iff #744 #186)
+#405 := (= #180 -5::int)
+#666 := (ite false #405 #186)
+#678 := (iff #666 #186)
+#674 := [rewrite]: #678
+#676 := (iff #744 #666)
+#667 := (iff #747 #186)
+#685 := (ite false #414 #186)
+#673 := (iff #685 #186)
+#675 := [rewrite]: #673
+#682 := (iff #747 #685)
+#591 := (iff #413 #186)
+#435 := 1::int
+#696 := (= #180 1::int)
+#587 := (ite false #696 #186)
+#589 := (iff #587 #186)
+#590 := [rewrite]: #589
+#580 := (iff #413 #587)
+#584 := (iff #415 #186)
+#574 := (+ 2::int #180)
+#585 := (= #574 0::int)
+#684 := (iff #585 #186)
+#569 := [rewrite]: #684
+#680 := (iff #415 #585)
+#535 := (= #424 #574)
+#691 := (+ #180 2::int)
+#575 := (= #691 #574)
+#576 := [rewrite]: #575
+#689 := (= #424 #691)
+#688 := (= #423 2::int)
+#702 := (mod 5::int 3::int)
+#534 := (= #702 2::int)
+#687 := [rewrite]: #534
+#531 := (= #423 #702)
+#701 := (= #337 3::int)
+#695 := [rewrite]: #701
+#694 := (= #422 5::int)
+#700 := [rewrite]: #694
+#533 := [monotonicity #700 #695]: #531
+#690 := [trans #533 #687]: #688
+#692 := [monotonicity #690]: #689
+#683 := [trans #692 #576]: #535
+#681 := [monotonicity #683]: #680
+#586 := [trans #681 #569]: #584
+#699 := (iff #428 #696)
+#704 := (+ -1::int #180)
+#552 := (= #704 0::int)
+#697 := (iff #552 #696)
+#698 := [rewrite]: #697
+#553 := (iff #428 #552)
+#705 := (= #355 #704)
+#707 := (+ #180 -1::int)
+#709 := (= #707 #704)
+#710 := [rewrite]: #709
+#708 := (= #355 #707)
+#717 := (= #427 -1::int)
+#715 := (* -1::int 1::int)
+#716 := (= #715 -1::int)
+#713 := [rewrite]: #716
+#430 := (= #427 #715)
+#436 := (= #426 1::int)
+#714 := [rewrite]: #436
+#712 := [monotonicity #714]: #430
+#706 := [trans #712 #713]: #717
+#703 := [monotonicity #706]: #708
+#711 := [trans #703 #710]: #705
+#554 := [monotonicity #711]: #553
+#693 := [trans #554 #698]: #699
+#433 := (iff #739 false)
+#444 := (or false false)
+#719 := (iff #444 false)
+#723 := [rewrite]: #719
+#718 := (iff #739 #444)
+#721 := (iff #743 false)
+#1 := true
+#727 := (not true)
+#734 := (iff #727 false)
+#736 := [rewrite]: #734
+#343 := (iff #743 #727)
+#449 := (iff #408 true)
+#738 := (or true true)
+#447 := (iff #738 true)
+#448 := [rewrite]: #447
+#724 := (iff #408 #738)
+#385 := (iff #425 true)
+#726 := [rewrite]: #385
+#737 := (iff #429 true)
+#735 := [rewrite]: #737
+#725 := [monotonicity #735 #726]: #724
+#450 := [trans #725 #448]: #449
+#720 := [monotonicity #450]: #343
+#443 := [trans #720 #736]: #721
+#376 := (iff #532 false)
+#733 := (iff #532 #727)
+#731 := (iff #402 true)
+#391 := (or true false)
+#729 := (iff #391 true)
+#730 := [rewrite]: #729
+#392 := (iff #402 #391)
+#389 := (iff #745 false)
+#390 := [rewrite]: #389
+#728 := [monotonicity #726 #390]: #392
+#732 := [trans #728 #730]: #731
+#371 := [monotonicity #732]: #733
+#377 := [trans #371 #736]: #376
+#722 := [monotonicity #377 #443]: #718
+#434 := [trans #722 #723]: #433
+#588 := [monotonicity #434 #693 #586]: #580
+#592 := [trans #588 #590]: #591
+#401 := (iff #746 false)
+#742 := [rewrite]: #401
+#686 := [monotonicity #742 #592]: #682
+#664 := [trans #686 #675]: #667
+#400 := (iff #748 #405)
+#741 := [rewrite]: #400
+#750 := (iff #749 false)
+#740 := [rewrite]: #750
+#677 := [monotonicity #740 #741 #664]: #676
+#668 := [trans #677 #674]: #679
+#656 := [monotonicity #668]: #672
+#660 := [trans #656 #659]: #672
+#665 := [quant-inst]: #671
+#661 := [mp #665 #660]: #670
+[unit-resolution #661 #764 #194]: false
+unsat
+66247d826364718443f0bc07b82495d982768eb5 327 0
+#2 := false
+#11 := 0::int
+decl f5 :: int
+#38 := f5
+#418 := (<= f5 0::int)
+#734 := (>= f5 0::int)
+#380 := (not #734)
+#723 := (not #418)
+#727 := (or #723 #380)
+#690 := (not #727)
+#39 := 3::int
+#415 := (mod f5 3::int)
+#66 := -1::int
+#416 := (* -1::int #415)
+decl f4 :: (-> int int int)
+#40 := (f4 f5 3::int)
+#344 := (+ #40 #416)
+#417 := (= #344 0::int)
+#563 := (not #417)
+#520 := (<= #344 0::int)
+#623 := (not #520)
+#631 := (>= #415 3::int)
+#632 := (not #631)
+#1 := true
+#57 := [true-axiom]: true
+#619 := (or false #632)
+#620 := [th-lemma]: #619
+#621 := [unit-resolution #620 #57]: #632
+#622 := [hypothesis]: #520
+#173 := (>= #40 3::int)
+#41 := (< #40 3::int)
+#42 := (not #41)
+#181 := (iff #42 #173)
+#172 := (not #173)
+#176 := (not #172)
+#179 := (iff #176 #173)
+#180 := [rewrite]: #179
+#177 := (iff #42 #176)
+#174 := (iff #41 #172)
+#175 := [rewrite]: #174
+#178 := [monotonicity #175]: #177
+#182 := [trans #178 #180]: #181
+#169 := [asserted]: #42
+#183 := [mp #169 #182]: #173
+#617 := [th-lemma #183 #622 #621]: false
+#609 := [lemma #617]: #623
+#626 := (or #563 #520)
+#637 := [th-lemma]: #626
+#615 := [unit-resolution #637 #609]: #563
+#614 := (or #690 #417)
+#438 := -3::int
+#411 := (* -1::int f5)
+#709 := (mod #411 -3::int)
+#433 := (+ #40 #709)
+#708 := (= #433 0::int)
+#423 := (ite #727 #417 #708)
+#403 := (= #40 0::int)
+#735 := (= f5 0::int)
+#703 := (ite #735 #403 #423)
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#29 := (f4 #8 #9)
+#747 := (pattern #29)
+#70 := (* -1::int #9)
+#67 := (* -1::int #8)
+#125 := (mod #67 #70)
+#258 := (+ #29 #125)
+#259 := (= #258 0::int)
+#30 := (mod #8 #9)
+#255 := (* -1::int #30)
+#256 := (+ #29 #255)
+#257 := (= #256 0::int)
+#92 := (<= #9 0::int)
+#88 := (<= #8 0::int)
+#197 := (or #88 #92)
+#198 := (not #197)
+#99 := (>= #8 0::int)
+#189 := (or #92 #99)
+#190 := (not #189)
+#204 := (or #190 #198)
+#260 := (ite #204 #257 #259)
+#254 := (= #29 0::int)
+#12 := (= #8 0::int)
+#261 := (ite #12 #254 #260)
+#253 := (= #8 #29)
+#13 := (= #9 0::int)
+#262 := (ite #13 #253 #261)
+#748 := (forall (vars (?v0 int) (?v1 int)) (:pat #747) #262)
+#265 := (forall (vars (?v0 int) (?v1 int)) #262)
+#751 := (iff #265 #748)
+#749 := (iff #262 #262)
+#750 := [refl]: #749
+#752 := [quant-intro #750]: #751
+#131 := (* -1::int #125)
+#222 := (ite #204 #30 #131)
+#225 := (ite #12 0::int #222)
+#228 := (ite #13 #8 #225)
+#231 := (= #29 #228)
+#234 := (forall (vars (?v0 int) (?v1 int)) #231)
+#266 := (iff #234 #265)
+#263 := (iff #231 #262)
+#264 := [rewrite]: #263
+#267 := [quant-intro #264]: #266
+#100 := (not #99)
+#93 := (not #92)
+#103 := (and #93 #100)
+#89 := (not #88)
+#96 := (and #89 #93)
+#106 := (or #96 #103)
+#151 := (ite #106 #30 #131)
+#154 := (ite #12 0::int #151)
+#157 := (ite #13 #8 #154)
+#160 := (= #29 #157)
+#163 := (forall (vars (?v0 int) (?v1 int)) #160)
+#235 := (iff #163 #234)
+#232 := (iff #160 #231)
+#229 := (= #157 #228)
+#226 := (= #154 #225)
+#223 := (= #151 #222)
+#207 := (iff #106 #204)
+#201 := (or #198 #190)
+#205 := (iff #201 #204)
+#206 := [rewrite]: #205
+#202 := (iff #106 #201)
+#199 := (iff #103 #190)
+#200 := [rewrite]: #199
+#187 := (iff #96 #198)
+#188 := [rewrite]: #187
+#203 := [monotonicity #188 #200]: #202
+#208 := [trans #203 #206]: #207
+#224 := [monotonicity #208]: #223
+#227 := [monotonicity #224]: #226
+#230 := [monotonicity #227]: #229
+#233 := [monotonicity #230]: #232
+#236 := [quant-intro #233]: #235
+#185 := (~ #163 #163)
+#170 := (~ #160 #160)
+#195 := [refl]: #170
+#186 := [nnf-pos #195]: #185
+#23 := (- #9)
+#22 := (- #8)
+#31 := (mod #22 #23)
+#32 := (- #31)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#33 := (ite #20 #30 #32)
+#34 := (ite #12 0::int #33)
+#35 := (ite #13 #8 #34)
+#36 := (= #29 #35)
+#37 := (forall (vars (?v0 int) (?v1 int)) #36)
+#166 := (iff #37 #163)
+#60 := (and #16 #18)
+#63 := (or #17 #60)
+#136 := (ite #63 #30 #131)
+#139 := (ite #12 0::int #136)
+#142 := (ite #13 #8 #139)
+#145 := (= #29 #142)
+#148 := (forall (vars (?v0 int) (?v1 int)) #145)
+#164 := (iff #148 #163)
+#161 := (iff #145 #160)
+#158 := (= #142 #157)
+#155 := (= #139 #154)
+#152 := (= #136 #151)
+#107 := (iff #63 #106)
+#104 := (iff #60 #103)
+#101 := (iff #18 #100)
+#102 := [rewrite]: #101
+#94 := (iff #16 #93)
+#95 := [rewrite]: #94
+#105 := [monotonicity #95 #102]: #104
+#97 := (iff #17 #96)
+#90 := (iff #15 #89)
+#91 := [rewrite]: #90
+#98 := [monotonicity #91 #95]: #97
+#108 := [monotonicity #98 #105]: #107
+#153 := [monotonicity #108]: #152
+#156 := [monotonicity #153]: #155
+#159 := [monotonicity #156]: #158
+#162 := [monotonicity #159]: #161
+#165 := [quant-intro #162]: #164
+#149 := (iff #37 #148)
+#146 := (iff #36 #145)
+#143 := (= #35 #142)
+#140 := (= #34 #139)
+#137 := (= #33 #136)
+#134 := (= #32 #131)
+#128 := (- #125)
+#132 := (= #128 #131)
+#133 := [rewrite]: #132
+#129 := (= #32 #128)
+#126 := (= #31 #125)
+#71 := (= #23 #70)
+#72 := [rewrite]: #71
+#68 := (= #22 #67)
+#69 := [rewrite]: #68
+#127 := [monotonicity #69 #72]: #126
+#130 := [monotonicity #127]: #129
+#135 := [trans #130 #133]: #134
+#64 := (iff #20 #63)
+#61 := (iff #19 #60)
+#62 := [rewrite]: #61
+#65 := [monotonicity #62]: #64
+#138 := [monotonicity #65 #135]: #137
+#141 := [monotonicity #138]: #140
+#144 := [monotonicity #141]: #143
+#147 := [monotonicity #144]: #146
+#150 := [quant-intro #147]: #149
+#167 := [trans #150 #165]: #166
+#124 := [asserted]: #37
+#168 := [mp #124 #167]: #163
+#196 := [mp~ #168 #186]: #163
+#237 := [mp #196 #236]: #234
+#268 := [mp #237 #267]: #265
+#753 := [mp #268 #752]: #748
+#692 := (not #748)
+#693 := (or #692 #703)
+#326 := (* -1::int 3::int)
+#412 := (mod #411 #326)
+#413 := (+ #40 #412)
+#404 := (= #413 0::int)
+#414 := (<= 3::int 0::int)
+#397 := (or #418 #414)
+#732 := (not #397)
+#391 := (or #414 #734)
+#521 := (not #391)
+#728 := (or #521 #732)
+#402 := (ite #728 #417 #404)
+#736 := (ite #735 #403 #402)
+#737 := (= f5 #40)
+#738 := (= 3::int 0::int)
+#733 := (ite #738 #737 #736)
+#698 := (or #692 #733)
+#694 := (iff #698 #693)
+#541 := (iff #693 #693)
+#542 := [rewrite]: #541
+#696 := (iff #733 #703)
+#701 := (ite false #737 #703)
+#706 := (iff #701 #703)
+#695 := [rewrite]: #706
+#705 := (iff #733 #701)
+#704 := (iff #736 #703)
+#424 := (iff #402 #423)
+#712 := (iff #404 #708)
+#707 := (= #413 #433)
+#710 := (= #412 #709)
+#439 := (= #326 -3::int)
+#332 := [rewrite]: #439
+#432 := [monotonicity #332]: #710
+#711 := [monotonicity #432]: #707
+#422 := [monotonicity #711]: #712
+#436 := (iff #728 #727)
+#366 := (or #380 #723)
+#713 := (iff #366 #727)
+#714 := [rewrite]: #713
+#726 := (iff #728 #366)
+#725 := (iff #732 #723)
+#722 := (iff #397 #418)
+#718 := (or #418 false)
+#721 := (iff #718 #418)
+#716 := [rewrite]: #721
+#719 := (iff #397 #718)
+#389 := (iff #414 false)
+#394 := [rewrite]: #389
+#720 := [monotonicity #394]: #719
+#360 := [trans #720 #716]: #722
+#365 := [monotonicity #360]: #725
+#381 := (iff #521 #380)
+#378 := (iff #391 #734)
+#730 := (or false #734)
+#374 := (iff #730 #734)
+#715 := [rewrite]: #374
+#390 := (iff #391 #730)
+#731 := [monotonicity #394]: #390
+#379 := [trans #731 #715]: #378
+#717 := [monotonicity #379]: #381
+#724 := [monotonicity #717 #365]: #726
+#437 := [trans #724 #714]: #436
+#425 := [monotonicity #437 #422]: #424
+#419 := [monotonicity #425]: #704
+#739 := (iff #738 false)
+#729 := [rewrite]: #739
+#702 := [monotonicity #729 #419]: #705
+#697 := [trans #702 #695]: #696
+#700 := [monotonicity #697]: #694
+#543 := [trans #700 #542]: #694
+#699 := [quant-inst]: #698
+#685 := [mp #699 #543]: #693
+#616 := [unit-resolution #685 #753]: #703
+#670 := (not #735)
+#669 := (not #703)
+#611 := (or #669 #670)
+#576 := (not #403)
+#686 := (<= #40 0::int)
+#618 := (not #686)
+#648 := (or #618 #172)
+#649 := [th-lemma]: #648
+#605 := [unit-resolution #649 #183]: #618
+#606 := (or #576 #686)
+#607 := [th-lemma]: #606
+#610 := [unit-resolution #607 #605]: #576
+#673 := (or #669 #670 #403)
+#558 := [def-axiom]: #673
+#612 := [unit-resolution #558 #610]: #611
+#613 := [unit-resolution #612 #616]: #670
+#573 := (or #669 #735 #423)
+#575 := [def-axiom]: #573
+#608 := [unit-resolution #575 #613 #616]: #423
+#677 := (not #423)
+#679 := (or #677 #690 #417)
+#680 := [def-axiom]: #679
+#587 := [unit-resolution #680 #608]: #614
+#588 := [unit-resolution #587 #615]: #690
+#688 := (or #727 #418)
+#682 := [def-axiom]: #688
+#589 := [unit-resolution #682 #588]: #418
+#683 := (or #727 #734)
+#689 := [def-axiom]: #683
+#482 := [unit-resolution #689 #588]: #734
+#593 := (or #735 #723 #380)
+#594 := [th-lemma]: #593
+#595 := [unit-resolution #594 #613]: #727
+[unit-resolution #595 #482 #589]: false
+unsat
+e48a7e8698a7ec4071541fd4399337a2b0f4fdb6 345 0
+#2 := false
+#39 := 3::int
+decl f4 :: (-> int int int)
+decl f5 :: int
+#38 := f5
+#40 := (f4 f5 3::int)
+#441 := (mod #40 3::int)
+#657 := (>= #441 3::int)
+#658 := (not #657)
+#1 := true
+#59 := [true-axiom]: true
+#647 := (or false #658)
+#642 := [th-lemma]: #647
+#648 := [unit-resolution #642 #59]: #658
+#11 := 0::int
+#68 := -1::int
+#436 := (* -1::int #40)
+#600 := (+ f5 #436)
+#601 := (<= #600 0::int)
+#172 := (= f5 #40)
+#188 := (>= f5 3::int)
+#187 := (not #188)
+#178 := (not #172)
+#194 := (or #178 #187)
+#199 := (not #194)
+#42 := (< f5 3::int)
+#41 := (= #40 f5)
+#43 := (implies #41 #42)
+#44 := (not #43)
+#202 := (iff #44 #199)
+#179 := (or #42 #178)
+#184 := (not #179)
+#200 := (iff #184 #199)
+#197 := (iff #179 #194)
+#191 := (or #187 #178)
+#195 := (iff #191 #194)
+#196 := [rewrite]: #195
+#192 := (iff #179 #191)
+#189 := (iff #42 #187)
+#190 := [rewrite]: #189
+#193 := [monotonicity #190]: #192
+#198 := [trans #193 #196]: #197
+#201 := [monotonicity #198]: #200
+#185 := (iff #44 #184)
+#182 := (iff #43 #179)
+#175 := (implies #172 #42)
+#180 := (iff #175 #179)
+#181 := [rewrite]: #180
+#176 := (iff #43 #175)
+#173 := (iff #41 #172)
+#174 := [rewrite]: #173
+#177 := [monotonicity #174]: #176
+#183 := [trans #177 #181]: #182
+#186 := [monotonicity #183]: #185
+#203 := [trans #186 #201]: #202
+#171 := [asserted]: #44
+#204 := [mp #171 #203]: #199
+#205 := [not-or-elim #204]: #172
+#634 := (or #178 #601)
+#630 := [th-lemma]: #634
+#631 := [unit-resolution #630 #205]: #601
+#206 := [not-or-elim #204]: #188
+#438 := (f4 #40 3::int)
+#602 := (* -1::int #438)
+#603 := (+ #40 #602)
+#604 := (<= #603 0::int)
+#763 := (= #40 #438)
+#635 := (= #438 #40)
+#632 := [symm #205]: #41
+#636 := [monotonicity #632]: #635
+#637 := [symm #636]: #763
+#638 := (not #763)
+#633 := (or #638 #604)
+#639 := [th-lemma]: #633
+#612 := [unit-resolution #639 #637]: #604
+#369 := (* -1::int #441)
+#442 := (+ #438 #369)
+#707 := (<= #442 0::int)
+#439 := (= #442 0::int)
+#738 := -3::int
+#462 := (mod #436 -3::int)
+#357 := (+ #438 #462)
+#457 := (= #357 0::int)
+#422 := (<= #40 0::int)
+#750 := (not #422)
+#416 := (>= #40 0::int)
+#406 := (not #416)
+#751 := (or #406 #750)
+#736 := (ite #751 #439 #457)
+#760 := (= #438 0::int)
+#761 := (= #40 0::int)
+#447 := (ite #761 #760 #736)
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#29 := (f4 #8 #9)
+#772 := (pattern #29)
+#72 := (* -1::int #9)
+#69 := (* -1::int #8)
+#127 := (mod #69 #72)
+#283 := (+ #29 #127)
+#284 := (= #283 0::int)
+#30 := (mod #8 #9)
+#280 := (* -1::int #30)
+#281 := (+ #29 #280)
+#282 := (= #281 0::int)
+#94 := (<= #9 0::int)
+#90 := (<= #8 0::int)
+#224 := (or #90 #94)
+#225 := (not #224)
+#101 := (>= #8 0::int)
+#216 := (or #94 #101)
+#217 := (not #216)
+#229 := (or #217 #225)
+#285 := (ite #229 #282 #284)
+#279 := (= #29 0::int)
+#12 := (= #8 0::int)
+#286 := (ite #12 #279 #285)
+#278 := (= #8 #29)
+#13 := (= #9 0::int)
+#287 := (ite #13 #278 #286)
+#773 := (forall (vars (?v0 int) (?v1 int)) (:pat #772) #287)
+#290 := (forall (vars (?v0 int) (?v1 int)) #287)
+#776 := (iff #290 #773)
+#774 := (iff #287 #287)
+#775 := [refl]: #774
+#777 := [quant-intro #775]: #776
+#133 := (* -1::int #127)
+#247 := (ite #229 #30 #133)
+#250 := (ite #12 0::int #247)
+#253 := (ite #13 #8 #250)
+#256 := (= #29 #253)
+#259 := (forall (vars (?v0 int) (?v1 int)) #256)
+#291 := (iff #259 #290)
+#288 := (iff #256 #287)
+#289 := [rewrite]: #288
+#292 := [quant-intro #289]: #291
+#102 := (not #101)
+#95 := (not #94)
+#105 := (and #95 #102)
+#91 := (not #90)
+#98 := (and #91 #95)
+#108 := (or #98 #105)
+#153 := (ite #108 #30 #133)
+#156 := (ite #12 0::int #153)
+#159 := (ite #13 #8 #156)
+#162 := (= #29 #159)
+#165 := (forall (vars (?v0 int) (?v1 int)) #162)
+#260 := (iff #165 #259)
+#257 := (iff #162 #256)
+#254 := (= #159 #253)
+#251 := (= #156 #250)
+#248 := (= #153 #247)
+#232 := (iff #108 #229)
+#226 := (or #225 #217)
+#230 := (iff #226 #229)
+#231 := [rewrite]: #230
+#227 := (iff #108 #226)
+#214 := (iff #105 #217)
+#215 := [rewrite]: #214
+#212 := (iff #98 #225)
+#213 := [rewrite]: #212
+#228 := [monotonicity #213 #215]: #227
+#233 := [trans #228 #231]: #232
+#249 := [monotonicity #233]: #248
+#252 := [monotonicity #249]: #251
+#255 := [monotonicity #252]: #254
+#258 := [monotonicity #255]: #257
+#261 := [quant-intro #258]: #260
+#210 := (~ #165 #165)
+#207 := (~ #162 #162)
+#222 := [refl]: #207
+#211 := [nnf-pos #222]: #210
+#23 := (- #9)
+#22 := (- #8)
+#31 := (mod #22 #23)
+#32 := (- #31)
+#16 := (< 0::int #9)
+#18 := (< #8 0::int)
+#19 := (and #18 #16)
+#15 := (< 0::int #8)
+#17 := (and #15 #16)
+#20 := (or #17 #19)
+#33 := (ite #20 #30 #32)
+#34 := (ite #12 0::int #33)
+#35 := (ite #13 #8 #34)
+#36 := (= #29 #35)
+#37 := (forall (vars (?v0 int) (?v1 int)) #36)
+#168 := (iff #37 #165)
+#62 := (and #16 #18)
+#65 := (or #17 #62)
+#138 := (ite #65 #30 #133)
+#141 := (ite #12 0::int #138)
+#144 := (ite #13 #8 #141)
+#147 := (= #29 #144)
+#150 := (forall (vars (?v0 int) (?v1 int)) #147)
+#166 := (iff #150 #165)
+#163 := (iff #147 #162)
+#160 := (= #144 #159)
+#157 := (= #141 #156)
+#154 := (= #138 #153)
+#109 := (iff #65 #108)
+#106 := (iff #62 #105)
+#103 := (iff #18 #102)
+#104 := [rewrite]: #103
+#96 := (iff #16 #95)
+#97 := [rewrite]: #96
+#107 := [monotonicity #97 #104]: #106
+#99 := (iff #17 #98)
+#92 := (iff #15 #91)
+#93 := [rewrite]: #92
+#100 := [monotonicity #93 #97]: #99
+#110 := [monotonicity #100 #107]: #109
+#155 := [monotonicity #110]: #154
+#158 := [monotonicity #155]: #157
+#161 := [monotonicity #158]: #160
+#164 := [monotonicity #161]: #163
+#167 := [quant-intro #164]: #166
+#151 := (iff #37 #150)
+#148 := (iff #36 #147)
+#145 := (= #35 #144)
+#142 := (= #34 #141)
+#139 := (= #33 #138)
+#136 := (= #32 #133)
+#130 := (- #127)
+#134 := (= #130 #133)
+#135 := [rewrite]: #134
+#131 := (= #32 #130)
+#128 := (= #31 #127)
+#73 := (= #23 #72)
+#74 := [rewrite]: #73
+#70 := (= #22 #69)
+#71 := [rewrite]: #70
+#129 := [monotonicity #71 #74]: #128
+#132 := [monotonicity #129]: #131
+#137 := [trans #132 #135]: #136
+#66 := (iff #20 #65)
+#63 := (iff #19 #62)
+#64 := [rewrite]: #63
+#67 := [monotonicity #64]: #66
+#140 := [monotonicity #67 #137]: #139
+#143 := [monotonicity #140]: #142
+#146 := [monotonicity #143]: #145
+#149 := [monotonicity #146]: #148
+#152 := [quant-intro #149]: #151
+#169 := [trans #152 #167]: #168
+#126 := [asserted]: #37
+#170 := [mp #126 #169]: #165
+#223 := [mp~ #170 #211]: #165
+#262 := [mp #223 #261]: #259
+#293 := [mp #262 #292]: #290
+#778 := [mp #293 #777]: #773
+#731 := (not #773)
+#720 := (or #731 #447)
+#351 := (* -1::int 3::int)
+#437 := (mod #436 #351)
+#429 := (+ #438 #437)
+#440 := (= #429 0::int)
+#443 := (<= 3::int 0::int)
+#757 := (or #422 #443)
+#759 := (not #757)
+#546 := (or #443 #416)
+#753 := (not #546)
+#427 := (or #753 #759)
+#428 := (ite #427 #439 #440)
+#762 := (ite #761 #760 #428)
+#758 := (= 3::int 0::int)
+#764 := (ite #758 #763 #762)
+#721 := (or #731 #764)
+#717 := (iff #721 #720)
+#723 := (iff #720 #720)
+#724 := [rewrite]: #723
+#730 := (iff #764 #447)
+#450 := (ite false #763 #447)
+#444 := (iff #450 #447)
+#726 := [rewrite]: #444
+#728 := (iff #764 #450)
+#448 := (iff #762 #447)
+#733 := (iff #428 #736)
+#458 := (iff #440 #457)
+#734 := (= #429 #357)
+#463 := (= #437 #462)
+#739 := (= #351 -3::int)
+#461 := [rewrite]: #739
+#464 := [monotonicity #461]: #463
+#735 := [monotonicity #464]: #734
+#732 := [monotonicity #735]: #458
+#749 := (iff #427 #751)
+#390 := (iff #759 #750)
+#385 := (iff #757 #422)
+#744 := (or #422 false)
+#741 := (iff #744 #422)
+#747 := [rewrite]: #741
+#745 := (iff #757 #744)
+#419 := (iff #443 false)
+#755 := [rewrite]: #419
+#746 := [monotonicity #755]: #745
+#748 := [trans #746 #747]: #385
+#391 := [monotonicity #748]: #390
+#742 := (iff #753 #406)
+#404 := (iff #546 #416)
+#415 := (or false #416)
+#740 := (iff #415 #416)
+#403 := [rewrite]: #740
+#756 := (iff #546 #415)
+#399 := [monotonicity #755]: #756
+#405 := [trans #399 #403]: #404
+#743 := [monotonicity #405]: #742
+#752 := [monotonicity #743 #391]: #749
+#737 := [monotonicity #752 #732]: #733
+#449 := [monotonicity #737]: #448
+#754 := (iff #758 false)
+#414 := [rewrite]: #754
+#729 := [monotonicity #414 #449]: #728
+#727 := [trans #729 #726]: #730
+#718 := [monotonicity #727]: #717
+#719 := [trans #718 #724]: #717
+#722 := [quant-inst]: #721
+#725 := [mp #722 #719]: #720
+#613 := [unit-resolution #725 #778]: #447
+#589 := (not #761)
+#614 := (not #601)
+#507 := (or #750 #187 #614)
+#618 := [th-lemma]: #507
+#619 := [unit-resolution #618 #206 #631]: #750
+#620 := (or #589 #422)
+#625 := [th-lemma]: #620
+#621 := [unit-resolution #625 #619]: #589
+#588 := (not #447)
+#697 := (or #588 #761 #736)
+#599 := [def-axiom]: #697
+#622 := [unit-resolution #599 #621 #613]: #736
+#568 := (or #751 #422)
+#710 := [def-axiom]: #568
+#623 := [unit-resolution #710 #619]: #751
+#711 := (not #751)
+#709 := (not #736)
+#716 := (or #709 #711 #439)
+#545 := [def-axiom]: #716
+#626 := [unit-resolution #545 #623 #622]: #439
+#701 := (not #439)
+#627 := (or #701 #707)
+#628 := [th-lemma]: #627
+#624 := [unit-resolution #628 #626]: #707
+[th-lemma #624 #612 #206 #631 #648]: false
+unsat
+ab1c3c43720ded895d2a82ffbaf56e6367dbc9b8 592 0
 #2 := false
 #11 := 0::int
 decl f3 :: (-> int int int)
@@ -40937,693 +43524,7 @@
 #362 := [unit-resolution #397 #361]: #600
 [th-lemma #362 #415 #382 #378]: false
 unsat
-faecfb2bed668dee6da551d02f1c23504bfe42c5 331 0
-#2 := false
-#11 := 0::int
-decl f4 :: (-> int int int)
-#38 := 3::int
-#170 := -3::int
-#173 := (f4 -3::int 3::int)
-#176 := (= #173 0::int)
-#189 := (not #176)
-#39 := (- 3::int)
-#40 := (f4 #39 3::int)
-#41 := (= #40 0::int)
-#42 := (not #41)
-#192 := (iff #42 #189)
-#179 := (= 0::int #173)
-#184 := (not #179)
-#190 := (iff #184 #189)
-#187 := (iff #179 #176)
-#188 := [rewrite]: #187
-#191 := [monotonicity #188]: #190
-#185 := (iff #42 #184)
-#182 := (iff #41 #179)
-#180 := (iff #176 #179)
-#181 := [rewrite]: #180
-#177 := (iff #41 #176)
-#174 := (= #40 #173)
-#171 := (= #39 -3::int)
-#172 := [rewrite]: #171
-#175 := [monotonicity #172]: #174
-#178 := [monotonicity #175]: #177
-#183 := [trans #178 #181]: #182
-#186 := [monotonicity #183]: #185
-#193 := [trans #186 #191]: #192
-#169 := [asserted]: #42
-#194 := [mp #169 #193]: #189
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#29 := (f4 #8 #9)
-#760 := (pattern #29)
-#66 := -1::int
-#70 := (* -1::int #9)
-#67 := (* -1::int #8)
-#125 := (mod #67 #70)
-#271 := (+ #29 #125)
-#272 := (= #271 0::int)
-#30 := (mod #8 #9)
-#268 := (* -1::int #30)
-#269 := (+ #29 #268)
-#270 := (= #269 0::int)
-#92 := (<= #9 0::int)
-#88 := (<= #8 0::int)
-#210 := (or #88 #92)
-#211 := (not #210)
-#99 := (>= #8 0::int)
-#202 := (or #92 #99)
-#203 := (not #202)
-#217 := (or #203 #211)
-#273 := (ite #217 #270 #272)
-#267 := (= #29 0::int)
-#12 := (= #8 0::int)
-#274 := (ite #12 #267 #273)
-#266 := (= #8 #29)
-#13 := (= #9 0::int)
-#275 := (ite #13 #266 #274)
-#761 := (forall (vars (?v0 int) (?v1 int)) (:pat #760) #275)
-#278 := (forall (vars (?v0 int) (?v1 int)) #275)
-#764 := (iff #278 #761)
-#762 := (iff #275 #275)
-#763 := [refl]: #762
-#765 := [quant-intro #763]: #764
-#131 := (* -1::int #125)
-#235 := (ite #217 #30 #131)
-#238 := (ite #12 0::int #235)
-#241 := (ite #13 #8 #238)
-#244 := (= #29 #241)
-#247 := (forall (vars (?v0 int) (?v1 int)) #244)
-#279 := (iff #247 #278)
-#276 := (iff #244 #275)
-#277 := [rewrite]: #276
-#280 := [quant-intro #277]: #279
-#100 := (not #99)
-#93 := (not #92)
-#103 := (and #93 #100)
-#89 := (not #88)
-#96 := (and #89 #93)
-#106 := (or #96 #103)
-#151 := (ite #106 #30 #131)
-#154 := (ite #12 0::int #151)
-#157 := (ite #13 #8 #154)
-#160 := (= #29 #157)
-#163 := (forall (vars (?v0 int) (?v1 int)) #160)
-#248 := (iff #163 #247)
-#245 := (iff #160 #244)
-#242 := (= #157 #241)
-#239 := (= #154 #238)
-#236 := (= #151 #235)
-#220 := (iff #106 #217)
-#214 := (or #211 #203)
-#218 := (iff #214 #217)
-#219 := [rewrite]: #218
-#215 := (iff #106 #214)
-#212 := (iff #103 #203)
-#213 := [rewrite]: #212
-#200 := (iff #96 #211)
-#201 := [rewrite]: #200
-#216 := [monotonicity #201 #213]: #215
-#221 := [trans #216 #219]: #220
-#237 := [monotonicity #221]: #236
-#240 := [monotonicity #237]: #239
-#243 := [monotonicity #240]: #242
-#246 := [monotonicity #243]: #245
-#249 := [quant-intro #246]: #248
-#198 := (~ #163 #163)
-#195 := (~ #160 #160)
-#208 := [refl]: #195
-#199 := [nnf-pos #208]: #198
-#23 := (- #9)
-#22 := (- #8)
-#31 := (mod #22 #23)
-#32 := (- #31)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#33 := (ite #20 #30 #32)
-#34 := (ite #12 0::int #33)
-#35 := (ite #13 #8 #34)
-#36 := (= #29 #35)
-#37 := (forall (vars (?v0 int) (?v1 int)) #36)
-#166 := (iff #37 #163)
-#60 := (and #16 #18)
-#63 := (or #17 #60)
-#136 := (ite #63 #30 #131)
-#139 := (ite #12 0::int #136)
-#142 := (ite #13 #8 #139)
-#145 := (= #29 #142)
-#148 := (forall (vars (?v0 int) (?v1 int)) #145)
-#164 := (iff #148 #163)
-#161 := (iff #145 #160)
-#158 := (= #142 #157)
-#155 := (= #139 #154)
-#152 := (= #136 #151)
-#107 := (iff #63 #106)
-#104 := (iff #60 #103)
-#101 := (iff #18 #100)
-#102 := [rewrite]: #101
-#94 := (iff #16 #93)
-#95 := [rewrite]: #94
-#105 := [monotonicity #95 #102]: #104
-#97 := (iff #17 #96)
-#90 := (iff #15 #89)
-#91 := [rewrite]: #90
-#98 := [monotonicity #91 #95]: #97
-#108 := [monotonicity #98 #105]: #107
-#153 := [monotonicity #108]: #152
-#156 := [monotonicity #153]: #155
-#159 := [monotonicity #156]: #158
-#162 := [monotonicity #159]: #161
-#165 := [quant-intro #162]: #164
-#149 := (iff #37 #148)
-#146 := (iff #36 #145)
-#143 := (= #35 #142)
-#140 := (= #34 #139)
-#137 := (= #33 #136)
-#134 := (= #32 #131)
-#128 := (- #125)
-#132 := (= #128 #131)
-#133 := [rewrite]: #132
-#129 := (= #32 #128)
-#126 := (= #31 #125)
-#71 := (= #23 #70)
-#72 := [rewrite]: #71
-#68 := (= #22 #67)
-#69 := [rewrite]: #68
-#127 := [monotonicity #69 #72]: #126
-#130 := [monotonicity #127]: #129
-#135 := [trans #130 #133]: #134
-#64 := (iff #20 #63)
-#61 := (iff #19 #60)
-#62 := [rewrite]: #61
-#65 := [monotonicity #62]: #64
-#138 := [monotonicity #65 #135]: #137
-#141 := [monotonicity #138]: #140
-#144 := [monotonicity #141]: #143
-#147 := [monotonicity #144]: #146
-#150 := [quant-intro #147]: #149
-#167 := [trans #150 #165]: #166
-#124 := [asserted]: #37
-#168 := [mp #124 #167]: #163
-#209 := [mp~ #168 #199]: #163
-#250 := [mp #209 #249]: #247
-#281 := [mp #250 #280]: #278
-#766 := [mp #281 #765]: #761
-#591 := (not #761)
-#592 := (or #591 #176)
-#339 := (* -1::int 3::int)
-#424 := (* -1::int -3::int)
-#425 := (mod #424 #339)
-#426 := (+ #173 #425)
-#417 := (= #426 0::int)
-#428 := (mod -3::int 3::int)
-#429 := (* -1::int #428)
-#357 := (+ #173 #429)
-#430 := (= #357 0::int)
-#427 := (<= 3::int 0::int)
-#431 := (<= -3::int 0::int)
-#410 := (or #431 #427)
-#745 := (not #410)
-#747 := (>= -3::int 0::int)
-#404 := (or #427 #747)
-#534 := (not #404)
-#741 := (or #534 #745)
-#415 := (ite #741 #430 #417)
-#416 := (= -3::int 0::int)
-#748 := (ite #416 #176 #415)
-#749 := (= -3::int #173)
-#750 := (= 3::int 0::int)
-#751 := (ite #750 #749 #748)
-#593 := (or #591 #751)
-#687 := (iff #593 #592)
-#688 := (iff #592 #592)
-#675 := [rewrite]: #688
-#582 := (iff #751 #176)
-#402 := (= #173 -3::int)
-#686 := (ite false #402 #176)
-#588 := (iff #686 #176)
-#589 := [rewrite]: #588
-#571 := (iff #751 #686)
-#682 := (iff #748 #176)
-#577 := (ite false #176 #176)
-#685 := (iff #577 #176)
-#587 := [rewrite]: #685
-#578 := (iff #748 #577)
-#694 := (iff #415 #176)
-#1 := true
-#689 := (ite true #176 #176)
-#693 := (iff #689 #176)
-#691 := [rewrite]: #693
-#690 := (iff #415 #689)
-#535 := (iff #417 #176)
-#704 := (= #426 #173)
-#719 := (+ #173 0::int)
-#710 := (= #719 #173)
-#705 := [rewrite]: #710
-#703 := (= #426 #719)
-#696 := (= #425 0::int)
-#698 := (mod 3::int -3::int)
-#701 := (= #698 0::int)
-#695 := [rewrite]: #701
-#699 := (= #425 #698)
-#555 := (= #339 -3::int)
-#556 := [rewrite]: #555
-#713 := (= #424 3::int)
-#554 := [rewrite]: #713
-#700 := [monotonicity #554 #556]: #699
-#702 := [trans #700 #695]: #696
-#697 := [monotonicity #702]: #703
-#533 := [trans #697 #705]: #704
-#536 := [monotonicity #533]: #535
-#712 := (iff #430 #176)
-#706 := (= #357 #173)
-#708 := (= #357 #719)
-#718 := (= #429 0::int)
-#438 := (* -1::int 0::int)
-#432 := (= #438 0::int)
-#714 := [rewrite]: #432
-#716 := (= #429 #438)
-#436 := (= #428 0::int)
-#437 := [rewrite]: #436
-#717 := [monotonicity #437]: #716
-#715 := [trans #717 #714]: #718
-#709 := [monotonicity #715]: #708
-#711 := [trans #709 #705]: #706
-#707 := [monotonicity #711]: #712
-#725 := (iff #741 true)
-#737 := (or true false)
-#727 := (iff #737 true)
-#449 := [rewrite]: #727
-#724 := (iff #741 #737)
-#446 := (iff #745 false)
-#452 := (not true)
-#723 := (iff #452 false)
-#445 := [rewrite]: #723
-#345 := (iff #745 #452)
-#450 := (iff #410 true)
-#740 := (iff #410 #737)
-#744 := (iff #427 false)
-#387 := [rewrite]: #744
-#379 := (iff #431 true)
-#739 := [rewrite]: #379
-#726 := [monotonicity #739 #387]: #740
-#451 := [trans #726 #449]: #450
-#722 := [monotonicity #451]: #345
-#720 := [trans #722 #445]: #446
-#738 := (iff #534 true)
-#734 := (not false)
-#373 := (iff #734 true)
-#736 := [rewrite]: #373
-#729 := (iff #534 #734)
-#732 := (iff #404 false)
-#392 := (or false false)
-#730 := (iff #392 false)
-#731 := [rewrite]: #730
-#393 := (iff #404 #392)
-#728 := (iff #747 false)
-#391 := [rewrite]: #728
-#394 := [monotonicity #387 #391]: #393
-#733 := [trans #394 #731]: #732
-#735 := [monotonicity #733]: #729
-#378 := [trans #735 #736]: #738
-#721 := [monotonicity #378 #720]: #724
-#435 := [trans #721 #449]: #725
-#692 := [monotonicity #435 #707 #536]: #690
-#576 := [trans #692 #691]: #694
-#743 := (iff #416 false)
-#403 := [rewrite]: #743
-#537 := [monotonicity #403 #576]: #578
-#683 := [trans #537 #587]: #682
-#742 := (iff #749 #402)
-#407 := [rewrite]: #742
-#746 := (iff #750 false)
-#752 := [rewrite]: #746
-#586 := [monotonicity #752 #407 #683]: #571
-#590 := [trans #586 #589]: #582
-#684 := [monotonicity #590]: #687
-#677 := [trans #684 #675]: #687
-#594 := [quant-inst]: #593
-#669 := [mp #594 #677]: #592
-[unit-resolution #669 #766 #194]: false
-unsat
-1cc348013db3bd37921bfa60ffc6651c1d4acce9 353 0
-#2 := false
-#42 := 1::int
-decl f4 :: (-> int int int)
-#40 := 3::int
-#172 := -5::int
-#175 := (f4 -5::int 3::int)
-#178 := (= #175 1::int)
-#191 := (not #178)
-#38 := 5::int
-#39 := (- 5::int)
-#41 := (f4 #39 3::int)
-#43 := (= #41 1::int)
-#44 := (not #43)
-#194 := (iff #44 #191)
-#181 := (= 1::int #175)
-#186 := (not #181)
-#192 := (iff #186 #191)
-#189 := (iff #181 #178)
-#190 := [rewrite]: #189
-#193 := [monotonicity #190]: #192
-#187 := (iff #44 #186)
-#184 := (iff #43 #181)
-#182 := (iff #178 #181)
-#183 := [rewrite]: #182
-#179 := (iff #43 #178)
-#176 := (= #41 #175)
-#173 := (= #39 -5::int)
-#174 := [rewrite]: #173
-#177 := [monotonicity #174]: #176
-#180 := [monotonicity #177]: #179
-#185 := [trans #180 #183]: #184
-#188 := [monotonicity #185]: #187
-#195 := [trans #188 #193]: #194
-#171 := [asserted]: #44
-#196 := [mp #171 #195]: #191
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#29 := (f4 #8 #9)
-#761 := (pattern #29)
-#11 := 0::int
-#68 := -1::int
-#72 := (* -1::int #9)
-#69 := (* -1::int #8)
-#127 := (mod #69 #72)
-#273 := (+ #29 #127)
-#274 := (= #273 0::int)
-#30 := (mod #8 #9)
-#270 := (* -1::int #30)
-#271 := (+ #29 #270)
-#272 := (= #271 0::int)
-#94 := (<= #9 0::int)
-#90 := (<= #8 0::int)
-#212 := (or #90 #94)
-#213 := (not #212)
-#101 := (>= #8 0::int)
-#204 := (or #94 #101)
-#205 := (not #204)
-#219 := (or #205 #213)
-#275 := (ite #219 #272 #274)
-#269 := (= #29 0::int)
-#12 := (= #8 0::int)
-#276 := (ite #12 #269 #275)
-#268 := (= #8 #29)
-#13 := (= #9 0::int)
-#277 := (ite #13 #268 #276)
-#762 := (forall (vars (?v0 int) (?v1 int)) (:pat #761) #277)
-#280 := (forall (vars (?v0 int) (?v1 int)) #277)
-#765 := (iff #280 #762)
-#763 := (iff #277 #277)
-#764 := [refl]: #763
-#766 := [quant-intro #764]: #765
-#133 := (* -1::int #127)
-#237 := (ite #219 #30 #133)
-#240 := (ite #12 0::int #237)
-#243 := (ite #13 #8 #240)
-#246 := (= #29 #243)
-#249 := (forall (vars (?v0 int) (?v1 int)) #246)
-#281 := (iff #249 #280)
-#278 := (iff #246 #277)
-#279 := [rewrite]: #278
-#282 := [quant-intro #279]: #281
-#102 := (not #101)
-#95 := (not #94)
-#105 := (and #95 #102)
-#91 := (not #90)
-#98 := (and #91 #95)
-#108 := (or #98 #105)
-#153 := (ite #108 #30 #133)
-#156 := (ite #12 0::int #153)
-#159 := (ite #13 #8 #156)
-#162 := (= #29 #159)
-#165 := (forall (vars (?v0 int) (?v1 int)) #162)
-#250 := (iff #165 #249)
-#247 := (iff #162 #246)
-#244 := (= #159 #243)
-#241 := (= #156 #240)
-#238 := (= #153 #237)
-#222 := (iff #108 #219)
-#216 := (or #213 #205)
-#220 := (iff #216 #219)
-#221 := [rewrite]: #220
-#217 := (iff #108 #216)
-#214 := (iff #105 #205)
-#215 := [rewrite]: #214
-#202 := (iff #98 #213)
-#203 := [rewrite]: #202
-#218 := [monotonicity #203 #215]: #217
-#223 := [trans #218 #221]: #222
-#239 := [monotonicity #223]: #238
-#242 := [monotonicity #239]: #241
-#245 := [monotonicity #242]: #244
-#248 := [monotonicity #245]: #247
-#251 := [quant-intro #248]: #250
-#200 := (~ #165 #165)
-#197 := (~ #162 #162)
-#210 := [refl]: #197
-#201 := [nnf-pos #210]: #200
-#23 := (- #9)
-#22 := (- #8)
-#31 := (mod #22 #23)
-#32 := (- #31)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#33 := (ite #20 #30 #32)
-#34 := (ite #12 0::int #33)
-#35 := (ite #13 #8 #34)
-#36 := (= #29 #35)
-#37 := (forall (vars (?v0 int) (?v1 int)) #36)
-#168 := (iff #37 #165)
-#62 := (and #16 #18)
-#65 := (or #17 #62)
-#138 := (ite #65 #30 #133)
-#141 := (ite #12 0::int #138)
-#144 := (ite #13 #8 #141)
-#147 := (= #29 #144)
-#150 := (forall (vars (?v0 int) (?v1 int)) #147)
-#166 := (iff #150 #165)
-#163 := (iff #147 #162)
-#160 := (= #144 #159)
-#157 := (= #141 #156)
-#154 := (= #138 #153)
-#109 := (iff #65 #108)
-#106 := (iff #62 #105)
-#103 := (iff #18 #102)
-#104 := [rewrite]: #103
-#96 := (iff #16 #95)
-#97 := [rewrite]: #96
-#107 := [monotonicity #97 #104]: #106
-#99 := (iff #17 #98)
-#92 := (iff #15 #91)
-#93 := [rewrite]: #92
-#100 := [monotonicity #93 #97]: #99
-#110 := [monotonicity #100 #107]: #109
-#155 := [monotonicity #110]: #154
-#158 := [monotonicity #155]: #157
-#161 := [monotonicity #158]: #160
-#164 := [monotonicity #161]: #163
-#167 := [quant-intro #164]: #166
-#151 := (iff #37 #150)
-#148 := (iff #36 #147)
-#145 := (= #35 #144)
-#142 := (= #34 #141)
-#139 := (= #33 #138)
-#136 := (= #32 #133)
-#130 := (- #127)
-#134 := (= #130 #133)
-#135 := [rewrite]: #134
-#131 := (= #32 #130)
-#128 := (= #31 #127)
-#73 := (= #23 #72)
-#74 := [rewrite]: #73
-#70 := (= #22 #69)
-#71 := [rewrite]: #70
-#129 := [monotonicity #71 #74]: #128
-#132 := [monotonicity #129]: #131
-#137 := [trans #132 #135]: #136
-#66 := (iff #20 #65)
-#63 := (iff #19 #62)
-#64 := [rewrite]: #63
-#67 := [monotonicity #64]: #66
-#140 := [monotonicity #67 #137]: #139
-#143 := [monotonicity #140]: #142
-#146 := [monotonicity #143]: #145
-#149 := [monotonicity #146]: #148
-#152 := [quant-intro #149]: #151
-#169 := [trans #152 #167]: #168
-#126 := [asserted]: #37
-#170 := [mp #126 #169]: #165
-#211 := [mp~ #170 #201]: #165
-#252 := [mp #211 #251]: #249
-#283 := [mp #252 #282]: #280
-#767 := [mp #283 #766]: #762
-#668 := (not #762)
-#675 := (or #668 #178)
-#341 := (* -1::int 3::int)
-#425 := (* -1::int -5::int)
-#426 := (mod #425 #341)
-#427 := (+ #175 #426)
-#418 := (= #427 0::int)
-#429 := (mod -5::int 3::int)
-#430 := (* -1::int #429)
-#431 := (+ #175 #430)
-#428 := (= #431 0::int)
-#432 := (<= 3::int 0::int)
-#411 := (<= -5::int 0::int)
-#746 := (or #411 #432)
-#748 := (not #746)
-#405 := (>= -5::int 0::int)
-#535 := (or #432 #405)
-#742 := (not #535)
-#416 := (or #742 #748)
-#417 := (ite #416 #428 #418)
-#749 := (= #175 0::int)
-#750 := (= -5::int 0::int)
-#751 := (ite #750 #749 #417)
-#752 := (= -5::int #175)
-#747 := (= 3::int 0::int)
-#753 := (ite #747 #752 #751)
-#659 := (or #668 #753)
-#662 := (iff #659 #675)
-#664 := (iff #675 #675)
-#665 := [rewrite]: #664
-#673 := (iff #753 #178)
-#744 := (= #175 -5::int)
-#681 := (ite false #744 #178)
-#671 := (iff #681 #178)
-#672 := [rewrite]: #671
-#677 := (iff #753 #681)
-#679 := (iff #751 #178)
-#676 := (ite false #749 #178)
-#667 := (iff #676 #178)
-#669 := [rewrite]: #667
-#678 := (iff #751 #676)
-#685 := (iff #417 #178)
-#572 := -2::int
-#587 := (= #175 -2::int)
-#1 := true
-#592 := (ite true #178 #587)
-#595 := (iff #592 #178)
-#688 := [rewrite]: #595
-#593 := (iff #417 #592)
-#583 := (iff #418 #587)
-#537 := 2::int
-#578 := (+ 2::int #175)
-#683 := (= #578 0::int)
-#589 := (iff #683 #587)
-#590 := [rewrite]: #589
-#684 := (iff #418 #683)
-#686 := (= #427 #578)
-#692 := (+ #175 2::int)
-#579 := (= #692 #578)
-#538 := [rewrite]: #579
-#695 := (= #427 #692)
-#693 := (= #426 2::int)
-#703 := -3::int
-#705 := (mod 5::int -3::int)
-#690 := (= #705 2::int)
-#691 := [rewrite]: #690
-#534 := (= #426 #705)
-#704 := (= #341 -3::int)
-#698 := [rewrite]: #704
-#696 := (= #425 5::int)
-#697 := [rewrite]: #696
-#536 := [monotonicity #697 #698]: #534
-#694 := [trans #536 #691]: #693
-#577 := [monotonicity #694]: #695
-#588 := [trans #577 #538]: #686
-#687 := [monotonicity #588]: #684
-#591 := [trans #687 #590]: #583
-#701 := (iff #428 #178)
-#707 := (+ -1::int #175)
-#555 := (= #707 0::int)
-#699 := (iff #555 #178)
-#700 := [rewrite]: #699
-#556 := (iff #428 #555)
-#708 := (= #431 #707)
-#710 := (+ #175 -1::int)
-#712 := (= #710 #707)
-#713 := [rewrite]: #712
-#711 := (= #431 #710)
-#720 := (= #430 -1::int)
-#718 := (* -1::int 1::int)
-#719 := (= #718 -1::int)
-#716 := [rewrite]: #719
-#433 := (= #430 #718)
-#439 := (= #429 1::int)
-#717 := [rewrite]: #439
-#715 := [monotonicity #717]: #433
-#709 := [trans #715 #716]: #720
-#706 := [monotonicity #709]: #711
-#714 := [trans #706 #713]: #708
-#557 := [monotonicity #714]: #556
-#702 := [trans #557 #700]: #701
-#437 := (iff #416 true)
-#727 := (or true false)
-#451 := (iff #727 true)
-#452 := [rewrite]: #451
-#726 := (iff #416 #727)
-#725 := (iff #748 false)
-#723 := (not true)
-#447 := (iff #723 false)
-#721 := [rewrite]: #447
-#724 := (iff #748 #723)
-#453 := (iff #746 true)
-#728 := (iff #746 #727)
-#729 := (iff #432 false)
-#392 := [rewrite]: #729
-#738 := (iff #411 true)
-#741 := [rewrite]: #738
-#450 := [monotonicity #741 #392]: #728
-#347 := [trans #450 #452]: #453
-#446 := [monotonicity #347]: #724
-#722 := [trans #446 #721]: #725
-#380 := (iff #742 true)
-#736 := (not false)
-#739 := (iff #736 true)
-#379 := [rewrite]: #739
-#374 := (iff #742 #736)
-#735 := (iff #535 false)
-#395 := (or false false)
-#733 := (iff #395 false)
-#734 := [rewrite]: #733
-#731 := (iff #535 #395)
-#393 := (iff #405 false)
-#394 := [rewrite]: #393
-#732 := [monotonicity #392 #394]: #731
-#730 := [trans #732 #734]: #735
-#737 := [monotonicity #730]: #374
-#740 := [trans #737 #379]: #380
-#436 := [monotonicity #740 #722]: #726
-#438 := [trans #436 #452]: #437
-#594 := [monotonicity #438 #702 #591]: #593
-#689 := [trans #594 #688]: #685
-#745 := (iff #750 false)
-#388 := [rewrite]: #745
-#670 := [monotonicity #388 #689]: #678
-#680 := [trans #670 #669]: #679
-#408 := (iff #752 #744)
-#404 := [rewrite]: #408
-#743 := (iff #747 false)
-#403 := [rewrite]: #743
-#682 := [monotonicity #403 #404 #680]: #677
-#674 := [trans #682 #672]: #673
-#663 := [monotonicity #674]: #662
-#661 := [trans #663 #665]: #662
-#660 := [quant-inst]: #659
-#666 := [mp #660 #661]: #675
-[unit-resolution #666 #767 #196]: false
-unsat
-0cfeb0d355451ec30e86083321e953491e3cc226 549 0
+f7d21861165159375bc45e6364e84fc0183b7f2c 549 0
 #2 := false
 #11 := 0::int
 decl f3 :: (-> int int int)
@@ -42173,358 +44074,7 @@
 #478 := [unit-resolution #470 #507]: #751
 [th-lemma #478 #476 #486 #472]: false
 unsat
-2a90f79ab6d055bafcbc76d83737ad593ffdc667 350 0
-#2 := false
-#68 := -1::int
-decl f4 :: (-> int int int)
-#174 := -3::int
-#177 := (f4 -1::int -3::int)
-#180 := (= #177 -1::int)
-#193 := (not #180)
-#38 := 1::int
-#39 := (- 1::int)
-#40 := 3::int
-#41 := (- 3::int)
-#42 := (f4 #39 #41)
-#43 := (= #42 #39)
-#44 := (not #43)
-#196 := (iff #44 #193)
-#183 := (= -1::int #177)
-#188 := (not #183)
-#194 := (iff #188 #193)
-#191 := (iff #183 #180)
-#192 := [rewrite]: #191
-#195 := [monotonicity #192]: #194
-#189 := (iff #44 #188)
-#186 := (iff #43 #183)
-#184 := (iff #180 #183)
-#185 := [rewrite]: #184
-#181 := (iff #43 #180)
-#172 := (= #39 -1::int)
-#173 := [rewrite]: #172
-#178 := (= #42 #177)
-#175 := (= #41 -3::int)
-#176 := [rewrite]: #175
-#179 := [monotonicity #173 #176]: #178
-#182 := [monotonicity #179 #173]: #181
-#187 := [trans #182 #185]: #186
-#190 := [monotonicity #187]: #189
-#197 := [trans #190 #195]: #196
-#171 := [asserted]: #44
-#198 := [mp #171 #197]: #193
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#29 := (f4 #8 #9)
-#763 := (pattern #29)
-#11 := 0::int
-#72 := (* -1::int #9)
-#69 := (* -1::int #8)
-#127 := (mod #69 #72)
-#275 := (+ #29 #127)
-#276 := (= #275 0::int)
-#30 := (mod #8 #9)
-#272 := (* -1::int #30)
-#273 := (+ #29 #272)
-#274 := (= #273 0::int)
-#94 := (<= #9 0::int)
-#90 := (<= #8 0::int)
-#214 := (or #90 #94)
-#215 := (not #214)
-#101 := (>= #8 0::int)
-#206 := (or #94 #101)
-#207 := (not #206)
-#221 := (or #207 #215)
-#277 := (ite #221 #274 #276)
-#271 := (= #29 0::int)
-#12 := (= #8 0::int)
-#278 := (ite #12 #271 #277)
-#270 := (= #8 #29)
-#13 := (= #9 0::int)
-#279 := (ite #13 #270 #278)
-#764 := (forall (vars (?v0 int) (?v1 int)) (:pat #763) #279)
-#282 := (forall (vars (?v0 int) (?v1 int)) #279)
-#767 := (iff #282 #764)
-#765 := (iff #279 #279)
-#766 := [refl]: #765
-#768 := [quant-intro #766]: #767
-#133 := (* -1::int #127)
-#239 := (ite #221 #30 #133)
-#242 := (ite #12 0::int #239)
-#245 := (ite #13 #8 #242)
-#248 := (= #29 #245)
-#251 := (forall (vars (?v0 int) (?v1 int)) #248)
-#283 := (iff #251 #282)
-#280 := (iff #248 #279)
-#281 := [rewrite]: #280
-#284 := [quant-intro #281]: #283
-#102 := (not #101)
-#95 := (not #94)
-#105 := (and #95 #102)
-#91 := (not #90)
-#98 := (and #91 #95)
-#108 := (or #98 #105)
-#153 := (ite #108 #30 #133)
-#156 := (ite #12 0::int #153)
-#159 := (ite #13 #8 #156)
-#162 := (= #29 #159)
-#165 := (forall (vars (?v0 int) (?v1 int)) #162)
-#252 := (iff #165 #251)
-#249 := (iff #162 #248)
-#246 := (= #159 #245)
-#243 := (= #156 #242)
-#240 := (= #153 #239)
-#224 := (iff #108 #221)
-#218 := (or #215 #207)
-#222 := (iff #218 #221)
-#223 := [rewrite]: #222
-#219 := (iff #108 #218)
-#216 := (iff #105 #207)
-#217 := [rewrite]: #216
-#204 := (iff #98 #215)
-#205 := [rewrite]: #204
-#220 := [monotonicity #205 #217]: #219
-#225 := [trans #220 #223]: #224
-#241 := [monotonicity #225]: #240
-#244 := [monotonicity #241]: #243
-#247 := [monotonicity #244]: #246
-#250 := [monotonicity #247]: #249
-#253 := [quant-intro #250]: #252
-#202 := (~ #165 #165)
-#199 := (~ #162 #162)
-#212 := [refl]: #199
-#203 := [nnf-pos #212]: #202
-#23 := (- #9)
-#22 := (- #8)
-#31 := (mod #22 #23)
-#32 := (- #31)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#33 := (ite #20 #30 #32)
-#34 := (ite #12 0::int #33)
-#35 := (ite #13 #8 #34)
-#36 := (= #29 #35)
-#37 := (forall (vars (?v0 int) (?v1 int)) #36)
-#168 := (iff #37 #165)
-#62 := (and #16 #18)
-#65 := (or #17 #62)
-#138 := (ite #65 #30 #133)
-#141 := (ite #12 0::int #138)
-#144 := (ite #13 #8 #141)
-#147 := (= #29 #144)
-#150 := (forall (vars (?v0 int) (?v1 int)) #147)
-#166 := (iff #150 #165)
-#163 := (iff #147 #162)
-#160 := (= #144 #159)
-#157 := (= #141 #156)
-#154 := (= #138 #153)
-#109 := (iff #65 #108)
-#106 := (iff #62 #105)
-#103 := (iff #18 #102)
-#104 := [rewrite]: #103
-#96 := (iff #16 #95)
-#97 := [rewrite]: #96
-#107 := [monotonicity #97 #104]: #106
-#99 := (iff #17 #98)
-#92 := (iff #15 #91)
-#93 := [rewrite]: #92
-#100 := [monotonicity #93 #97]: #99
-#110 := [monotonicity #100 #107]: #109
-#155 := [monotonicity #110]: #154
-#158 := [monotonicity #155]: #157
-#161 := [monotonicity #158]: #160
-#164 := [monotonicity #161]: #163
-#167 := [quant-intro #164]: #166
-#151 := (iff #37 #150)
-#148 := (iff #36 #147)
-#145 := (= #35 #144)
-#142 := (= #34 #141)
-#139 := (= #33 #138)
-#136 := (= #32 #133)
-#130 := (- #127)
-#134 := (= #130 #133)
-#135 := [rewrite]: #134
-#131 := (= #32 #130)
-#128 := (= #31 #127)
-#73 := (= #23 #72)
-#74 := [rewrite]: #73
-#70 := (= #22 #69)
-#71 := [rewrite]: #70
-#129 := [monotonicity #71 #74]: #128
-#132 := [monotonicity #129]: #131
-#137 := [trans #132 #135]: #136
-#66 := (iff #20 #65)
-#63 := (iff #19 #62)
-#64 := [rewrite]: #63
-#67 := [monotonicity #64]: #66
-#140 := [monotonicity #67 #137]: #139
-#143 := [monotonicity #140]: #142
-#146 := [monotonicity #143]: #145
-#149 := [monotonicity #146]: #148
-#152 := [quant-intro #149]: #151
-#169 := [trans #152 #167]: #168
-#126 := [asserted]: #37
-#170 := [mp #126 #169]: #165
-#213 := [mp~ #170 #203]: #165
-#254 := [mp #213 #253]: #251
-#285 := [mp #254 #284]: #282
-#769 := [mp #285 #768]: #764
-#684 := (not #764)
-#673 := (or #684 #180)
-#343 := (* -1::int -3::int)
-#427 := (* -1::int -1::int)
-#428 := (mod #427 #343)
-#429 := (+ #177 #428)
-#420 := (= #429 0::int)
-#431 := (mod -1::int -3::int)
-#432 := (* -1::int #431)
-#433 := (+ #177 #432)
-#430 := (= #433 0::int)
-#434 := (<= -3::int 0::int)
-#413 := (<= -1::int 0::int)
-#748 := (or #413 #434)
-#750 := (not #748)
-#407 := (>= -1::int 0::int)
-#537 := (or #434 #407)
-#744 := (not #537)
-#418 := (or #744 #750)
-#419 := (ite #418 #430 #420)
-#751 := (= #177 0::int)
-#752 := (= -1::int 0::int)
-#753 := (ite #752 #751 #419)
-#754 := (= -3::int 0::int)
-#749 := (ite #754 #183 #753)
-#674 := (or #684 #749)
-#676 := (iff #674 #673)
-#677 := (iff #673 #673)
-#661 := [rewrite]: #677
-#683 := (iff #749 #180)
-#672 := (ite false #180 #180)
-#681 := (iff #672 #180)
-#682 := [rewrite]: #681
-#669 := (iff #749 #672)
-#678 := (iff #753 #180)
-#596 := (ite false #751 #180)
-#687 := (iff #596 #180)
-#691 := [rewrite]: #687
-#597 := (iff #753 #596)
-#594 := (iff #419 #180)
-#728 := 2::int
-#558 := (= #177 2::int)
-#589 := (ite false #558 #180)
-#585 := (iff #589 #180)
-#593 := [rewrite]: #585
-#591 := (iff #419 #589)
-#689 := (iff #420 #180)
-#694 := (+ 1::int #177)
-#540 := (= #694 0::int)
-#685 := (iff #540 #180)
-#686 := [rewrite]: #685
-#688 := (iff #420 #540)
-#580 := (= #429 #694)
-#693 := (+ #177 1::int)
-#697 := (= #693 #694)
-#579 := [rewrite]: #697
-#695 := (= #429 #693)
-#539 := (= #428 1::int)
-#706 := (mod 1::int 3::int)
-#536 := (= #706 1::int)
-#538 := [rewrite]: #536
-#700 := (= #428 #706)
-#699 := (= #343 3::int)
-#705 := [rewrite]: #699
-#704 := (= #427 1::int)
-#698 := [rewrite]: #704
-#707 := [monotonicity #698 #705]: #700
-#692 := [trans #707 #538]: #539
-#696 := [monotonicity #692]: #695
-#581 := [trans #696 #579]: #580
-#590 := [monotonicity #581]: #688
-#574 := [trans #590 #686]: #689
-#702 := (iff #430 #558)
-#720 := -2::int
-#713 := (+ -2::int #177)
-#710 := (= #713 0::int)
-#559 := (iff #710 #558)
-#701 := [rewrite]: #559
-#716 := (iff #430 #710)
-#714 := (= #433 #713)
-#722 := (+ #177 -2::int)
-#708 := (= #722 #713)
-#709 := [rewrite]: #708
-#711 := (= #433 #722)
-#721 := (= #432 -2::int)
-#440 := (* -1::int 2::int)
-#435 := (= #440 -2::int)
-#717 := [rewrite]: #435
-#441 := (= #432 #440)
-#438 := (= #431 2::int)
-#439 := [rewrite]: #438
-#719 := [monotonicity #439]: #441
-#718 := [trans #719 #717]: #721
-#712 := [monotonicity #718]: #711
-#715 := [trans #712 #709]: #714
-#557 := [monotonicity #715]: #716
-#703 := [trans #557 #701]: #702
-#727 := (iff #418 false)
-#725 := (or false false)
-#449 := (iff #725 false)
-#723 := [rewrite]: #449
-#726 := (iff #418 #725)
-#455 := (iff #750 false)
-#1 := true
-#735 := (not true)
-#732 := (iff #735 false)
-#738 := [rewrite]: #732
-#453 := (iff #750 #735)
-#730 := (iff #748 true)
-#382 := (or true true)
-#743 := (iff #382 true)
-#729 := [rewrite]: #743
-#742 := (iff #748 #382)
-#746 := (iff #434 true)
-#406 := [rewrite]: #746
-#741 := (iff #413 true)
-#381 := [rewrite]: #741
-#740 := [monotonicity #381 #406]: #742
-#452 := [trans #740 #729]: #730
-#454 := [monotonicity #452]: #453
-#349 := [trans #454 #738]: #455
-#376 := (iff #744 false)
-#736 := (iff #744 #735)
-#733 := (iff #537 true)
-#731 := (or true false)
-#396 := (iff #731 true)
-#397 := [rewrite]: #396
-#394 := (iff #537 #731)
-#747 := (iff #407 false)
-#390 := [rewrite]: #747
-#395 := [monotonicity #406 #390]: #394
-#734 := [trans #395 #397]: #733
-#737 := [monotonicity #734]: #736
-#739 := [trans #737 #738]: #376
-#448 := [monotonicity #739 #349]: #726
-#724 := [trans #448 #723]: #727
-#592 := [monotonicity #724 #703 #574]: #591
-#595 := [trans #592 #593]: #594
-#405 := (iff #752 false)
-#410 := [rewrite]: #405
-#690 := [monotonicity #410 #595]: #597
-#680 := [trans #690 #691]: #678
-#755 := (iff #754 false)
-#745 := [rewrite]: #755
-#671 := [monotonicity #745 #192 #680]: #669
-#679 := [trans #671 #682]: #683
-#670 := [monotonicity #679]: #676
-#662 := [trans #670 #661]: #676
-#675 := [quant-inst]: #674
-#664 := [mp #675 #662]: #673
-[unit-resolution #664 #769 #198]: false
-unsat
-15bd1be394385145d20cb6ec18ac38668f075cbe 75 0
+b8a677f41acfea6ca1bd28f53fd89f2c10758abc 75 0
 #2 := false
 #8 := 0::int
 decl f3 :: int
@@ -42600,327 +44150,7 @@
 #96 := [unit-resolution #95 #92]: #90
 [th-lemma #89 #70 #96]: false
 unsat
-7f8f90a6f1c8936c3bb2040faca497f7f7f61ec8 319 0
-#2 := false
-#11 := 0::int
-decl f4 :: (-> int int int)
-#170 := -3::int
-#173 := (f4 -3::int -3::int)
-#176 := (= #173 0::int)
-#189 := (not #176)
-#38 := 3::int
-#39 := (- 3::int)
-#40 := (f4 #39 #39)
-#41 := (= #40 0::int)
-#42 := (not #41)
-#192 := (iff #42 #189)
-#179 := (= 0::int #173)
-#184 := (not #179)
-#190 := (iff #184 #189)
-#187 := (iff #179 #176)
-#188 := [rewrite]: #187
-#191 := [monotonicity #188]: #190
-#185 := (iff #42 #184)
-#182 := (iff #41 #179)
-#180 := (iff #176 #179)
-#181 := [rewrite]: #180
-#177 := (iff #41 #176)
-#174 := (= #40 #173)
-#171 := (= #39 -3::int)
-#172 := [rewrite]: #171
-#175 := [monotonicity #172 #172]: #174
-#178 := [monotonicity #175]: #177
-#183 := [trans #178 #181]: #182
-#186 := [monotonicity #183]: #185
-#193 := [trans #186 #191]: #192
-#169 := [asserted]: #42
-#194 := [mp #169 #193]: #189
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#29 := (f4 #8 #9)
-#760 := (pattern #29)
-#66 := -1::int
-#70 := (* -1::int #9)
-#67 := (* -1::int #8)
-#125 := (mod #67 #70)
-#271 := (+ #29 #125)
-#272 := (= #271 0::int)
-#30 := (mod #8 #9)
-#268 := (* -1::int #30)
-#269 := (+ #29 #268)
-#270 := (= #269 0::int)
-#92 := (<= #9 0::int)
-#88 := (<= #8 0::int)
-#210 := (or #88 #92)
-#211 := (not #210)
-#99 := (>= #8 0::int)
-#202 := (or #92 #99)
-#203 := (not #202)
-#217 := (or #203 #211)
-#273 := (ite #217 #270 #272)
-#267 := (= #29 0::int)
-#12 := (= #8 0::int)
-#274 := (ite #12 #267 #273)
-#266 := (= #8 #29)
-#13 := (= #9 0::int)
-#275 := (ite #13 #266 #274)
-#761 := (forall (vars (?v0 int) (?v1 int)) (:pat #760) #275)
-#278 := (forall (vars (?v0 int) (?v1 int)) #275)
-#764 := (iff #278 #761)
-#762 := (iff #275 #275)
-#763 := [refl]: #762
-#765 := [quant-intro #763]: #764
-#131 := (* -1::int #125)
-#235 := (ite #217 #30 #131)
-#238 := (ite #12 0::int #235)
-#241 := (ite #13 #8 #238)
-#244 := (= #29 #241)
-#247 := (forall (vars (?v0 int) (?v1 int)) #244)
-#279 := (iff #247 #278)
-#276 := (iff #244 #275)
-#277 := [rewrite]: #276
-#280 := [quant-intro #277]: #279
-#100 := (not #99)
-#93 := (not #92)
-#103 := (and #93 #100)
-#89 := (not #88)
-#96 := (and #89 #93)
-#106 := (or #96 #103)
-#151 := (ite #106 #30 #131)
-#154 := (ite #12 0::int #151)
-#157 := (ite #13 #8 #154)
-#160 := (= #29 #157)
-#163 := (forall (vars (?v0 int) (?v1 int)) #160)
-#248 := (iff #163 #247)
-#245 := (iff #160 #244)
-#242 := (= #157 #241)
-#239 := (= #154 #238)
-#236 := (= #151 #235)
-#220 := (iff #106 #217)
-#214 := (or #211 #203)
-#218 := (iff #214 #217)
-#219 := [rewrite]: #218
-#215 := (iff #106 #214)
-#212 := (iff #103 #203)
-#213 := [rewrite]: #212
-#200 := (iff #96 #211)
-#201 := [rewrite]: #200
-#216 := [monotonicity #201 #213]: #215
-#221 := [trans #216 #219]: #220
-#237 := [monotonicity #221]: #236
-#240 := [monotonicity #237]: #239
-#243 := [monotonicity #240]: #242
-#246 := [monotonicity #243]: #245
-#249 := [quant-intro #246]: #248
-#198 := (~ #163 #163)
-#195 := (~ #160 #160)
-#208 := [refl]: #195
-#199 := [nnf-pos #208]: #198
-#23 := (- #9)
-#22 := (- #8)
-#31 := (mod #22 #23)
-#32 := (- #31)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#33 := (ite #20 #30 #32)
-#34 := (ite #12 0::int #33)
-#35 := (ite #13 #8 #34)
-#36 := (= #29 #35)
-#37 := (forall (vars (?v0 int) (?v1 int)) #36)
-#166 := (iff #37 #163)
-#60 := (and #16 #18)
-#63 := (or #17 #60)
-#136 := (ite #63 #30 #131)
-#139 := (ite #12 0::int #136)
-#142 := (ite #13 #8 #139)
-#145 := (= #29 #142)
-#148 := (forall (vars (?v0 int) (?v1 int)) #145)
-#164 := (iff #148 #163)
-#161 := (iff #145 #160)
-#158 := (= #142 #157)
-#155 := (= #139 #154)
-#152 := (= #136 #151)
-#107 := (iff #63 #106)
-#104 := (iff #60 #103)
-#101 := (iff #18 #100)
-#102 := [rewrite]: #101
-#94 := (iff #16 #93)
-#95 := [rewrite]: #94
-#105 := [monotonicity #95 #102]: #104
-#97 := (iff #17 #96)
-#90 := (iff #15 #89)
-#91 := [rewrite]: #90
-#98 := [monotonicity #91 #95]: #97
-#108 := [monotonicity #98 #105]: #107
-#153 := [monotonicity #108]: #152
-#156 := [monotonicity #153]: #155
-#159 := [monotonicity #156]: #158
-#162 := [monotonicity #159]: #161
-#165 := [quant-intro #162]: #164
-#149 := (iff #37 #148)
-#146 := (iff #36 #145)
-#143 := (= #35 #142)
-#140 := (= #34 #139)
-#137 := (= #33 #136)
-#134 := (= #32 #131)
-#128 := (- #125)
-#132 := (= #128 #131)
-#133 := [rewrite]: #132
-#129 := (= #32 #128)
-#126 := (= #31 #125)
-#71 := (= #23 #70)
-#72 := [rewrite]: #71
-#68 := (= #22 #67)
-#69 := [rewrite]: #68
-#127 := [monotonicity #69 #72]: #126
-#130 := [monotonicity #127]: #129
-#135 := [trans #130 #133]: #134
-#64 := (iff #20 #63)
-#61 := (iff #19 #60)
-#62 := [rewrite]: #61
-#65 := [monotonicity #62]: #64
-#138 := [monotonicity #65 #135]: #137
-#141 := [monotonicity #138]: #140
-#144 := [monotonicity #141]: #143
-#147 := [monotonicity #144]: #146
-#150 := [quant-intro #147]: #149
-#167 := [trans #150 #165]: #166
-#124 := [asserted]: #37
-#168 := [mp #124 #167]: #163
-#209 := [mp~ #168 #199]: #163
-#250 := [mp #209 #249]: #247
-#281 := [mp #250 #280]: #278
-#766 := [mp #281 #765]: #761
-#537 := (not #761)
-#685 := (or #537 #176)
-#339 := (* -1::int -3::int)
-#424 := (mod #339 #339)
-#425 := (+ #173 #424)
-#426 := (= #425 0::int)
-#417 := (mod -3::int -3::int)
-#428 := (* -1::int #417)
-#429 := (+ #173 #428)
-#357 := (= #429 0::int)
-#430 := (<= -3::int 0::int)
-#427 := (or #430 #430)
-#431 := (not #427)
-#410 := (>= -3::int 0::int)
-#745 := (or #430 #410)
-#747 := (not #745)
-#404 := (or #747 #431)
-#534 := (ite #404 #357 #426)
-#741 := (= -3::int 0::int)
-#415 := (ite #741 #176 #534)
-#416 := (= -3::int #173)
-#748 := (ite #741 #416 #415)
-#587 := (or #537 #748)
-#683 := (iff #587 #685)
-#571 := (iff #685 #685)
-#586 := [rewrite]: #571
-#577 := (iff #748 #176)
-#746 := (= #173 -3::int)
-#692 := (ite false #746 #176)
-#694 := (iff #692 #176)
-#576 := [rewrite]: #694
-#693 := (iff #748 #692)
-#689 := (iff #415 #176)
-#695 := (ite false #176 #176)
-#703 := (iff #695 #176)
-#697 := [rewrite]: #703
-#535 := (iff #415 #695)
-#704 := (iff #534 #176)
-#696 := (iff #534 #695)
-#700 := (iff #426 #176)
-#698 := (= #425 #173)
-#438 := (+ #173 0::int)
-#432 := (= #438 #173)
-#714 := [rewrite]: #432
-#555 := (= #425 #438)
-#713 := (= #424 0::int)
-#705 := (mod 3::int 3::int)
-#712 := (= #705 0::int)
-#707 := [rewrite]: #712
-#706 := (= #424 #705)
-#709 := (= #339 3::int)
-#710 := [rewrite]: #709
-#711 := [monotonicity #710 #710]: #706
-#554 := [trans #711 #707]: #713
-#556 := [monotonicity #554]: #555
-#699 := [trans #556 #714]: #698
-#701 := [monotonicity #699]: #700
-#719 := (iff #357 #176)
-#718 := (= #429 #173)
-#716 := (= #429 #438)
-#436 := (= #428 0::int)
-#720 := (* -1::int 0::int)
-#725 := (= #720 0::int)
-#435 := [rewrite]: #725
-#724 := (= #428 #720)
-#445 := (= #417 0::int)
-#446 := [rewrite]: #445
-#721 := [monotonicity #446]: #724
-#437 := [trans #721 #435]: #436
-#717 := [monotonicity #437]: #716
-#715 := [trans #717 #714]: #718
-#708 := [monotonicity #715]: #719
-#722 := (iff #404 false)
-#449 := (or false false)
-#452 := (iff #449 false)
-#345 := [rewrite]: #452
-#450 := (iff #404 #449)
-#726 := (iff #431 false)
-#1 := true
-#394 := (not true)
-#732 := (iff #394 false)
-#733 := [rewrite]: #732
-#737 := (iff #431 #394)
-#379 := (iff #427 true)
-#735 := (or true true)
-#738 := (iff #735 true)
-#378 := [rewrite]: #738
-#373 := (iff #427 #735)
-#742 := (iff #430 true)
-#402 := [rewrite]: #742
-#736 := [monotonicity #402 #402]: #373
-#739 := [trans #736 #378]: #379
-#740 := [monotonicity #739]: #737
-#727 := [trans #740 #733]: #726
-#734 := (iff #747 false)
-#730 := (iff #747 #394)
-#392 := (iff #745 true)
-#403 := (or true false)
-#728 := (iff #403 true)
-#391 := [rewrite]: #728
-#744 := (iff #745 #403)
-#407 := (iff #410 false)
-#743 := [rewrite]: #407
-#387 := [monotonicity #402 #743]: #744
-#393 := [trans #387 #391]: #392
-#731 := [monotonicity #393]: #730
-#729 := [trans #731 #733]: #734
-#451 := [monotonicity #729 #727]: #450
-#723 := [trans #451 #345]: #722
-#702 := [monotonicity #723 #708 #701]: #696
-#533 := [trans #702 #697]: #704
-#749 := (iff #741 false)
-#750 := [rewrite]: #749
-#536 := [monotonicity #750 #533]: #535
-#690 := [trans #536 #697]: #689
-#751 := (iff #416 #746)
-#752 := [rewrite]: #751
-#691 := [monotonicity #750 #752 #690]: #693
-#578 := [trans #691 #576]: #577
-#686 := [monotonicity #578]: #683
-#588 := [trans #686 #586]: #683
-#682 := [quant-inst]: #587
-#589 := [mp #682 #588]: #685
-[unit-resolution #589 #766 #194]: false
-unsat
-98ce2e7fe8c73154e060ef0707079589f8062204 132 0
+957272aa2f77e61687a16c7aa490ba2664908958 132 0
 #2 := false
 #9 := 0::int
 decl f3 :: int
@@ -43053,7 +44283,7 @@
 #162 := [th-lemma]: #161
 [unit-resolution #162 #160 #143]: false
 unsat
-11cbd971427456e31403121d2898ec25b367a8f3 103 0
+9485d306079aa8cdc08952543a8ac86391d45fc0 103 0
 #2 := false
 #8 := 0::int
 decl f3 :: int
@@ -43157,7 +44387,7 @@
 #124 := [unit-resolution #123 #120]: #115
 [th-lemma #110 #124]: false
 unsat
-08b44e1a6f35749f6a110c211434791c0d3e5dd5 149 0
+443d726f158e992756000aa2462469d5494d62ee 149 0
 #2 := false
 #9 := 0::int
 decl f3 :: int
@@ -43307,7 +44537,7 @@
 #170 := [unit-resolution #169 #155]: #97
 [th-lemma #159 #170 #167]: false
 unsat
-603464f42ca695d3e1441697b9257ae6111df711 114 0
+5d75c12dca04a93af891e80a12133cb48b0e8ce3 114 0
 #2 := false
 #9 := 0::int
 decl f3 :: int
@@ -43422,355 +44652,7 @@
 #143 := [unit-resolution #142 #139]: #137
 [th-lemma #136 #134 #143]: false
 unsat
-6b1006b88fea0865d9f28ce9c2aca40c518bab7b 347 0
-#2 := false
-#183 := -2::int
-decl f4 :: (-> int int int)
-#177 := -3::int
-#174 := -5::int
-#180 := (f4 -5::int -3::int)
-#186 := (= #180 -2::int)
-#189 := (not #186)
-#43 := 2::int
-#44 := (- 2::int)
-#40 := 3::int
-#41 := (- 3::int)
-#38 := 5::int
-#39 := (- 5::int)
-#42 := (f4 #39 #41)
-#45 := (= #42 #44)
-#46 := (not #45)
-#190 := (iff #46 #189)
-#187 := (iff #45 #186)
-#184 := (= #44 -2::int)
-#185 := [rewrite]: #184
-#181 := (= #42 #180)
-#178 := (= #41 -3::int)
-#179 := [rewrite]: #178
-#175 := (= #39 -5::int)
-#176 := [rewrite]: #175
-#182 := [monotonicity #176 #179]: #181
-#188 := [monotonicity #182 #185]: #187
-#191 := [monotonicity #188]: #190
-#173 := [asserted]: #46
-#194 := [mp #173 #191]: #189
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#29 := (f4 #8 #9)
-#758 := (pattern #29)
-#11 := 0::int
-#70 := -1::int
-#74 := (* -1::int #9)
-#71 := (* -1::int #8)
-#129 := (mod #71 #74)
-#269 := (+ #29 #129)
-#270 := (= #269 0::int)
-#30 := (mod #8 #9)
-#266 := (* -1::int #30)
-#267 := (+ #29 #266)
-#268 := (= #267 0::int)
-#96 := (<= #9 0::int)
-#92 := (<= #8 0::int)
-#208 := (or #92 #96)
-#209 := (not #208)
-#103 := (>= #8 0::int)
-#200 := (or #96 #103)
-#201 := (not #200)
-#215 := (or #201 #209)
-#271 := (ite #215 #268 #270)
-#265 := (= #29 0::int)
-#12 := (= #8 0::int)
-#272 := (ite #12 #265 #271)
-#264 := (= #8 #29)
-#13 := (= #9 0::int)
-#273 := (ite #13 #264 #272)
-#759 := (forall (vars (?v0 int) (?v1 int)) (:pat #758) #273)
-#276 := (forall (vars (?v0 int) (?v1 int)) #273)
-#762 := (iff #276 #759)
-#760 := (iff #273 #273)
-#761 := [refl]: #760
-#763 := [quant-intro #761]: #762
-#135 := (* -1::int #129)
-#233 := (ite #215 #30 #135)
-#236 := (ite #12 0::int #233)
-#239 := (ite #13 #8 #236)
-#242 := (= #29 #239)
-#245 := (forall (vars (?v0 int) (?v1 int)) #242)
-#277 := (iff #245 #276)
-#274 := (iff #242 #273)
-#275 := [rewrite]: #274
-#278 := [quant-intro #275]: #277
-#104 := (not #103)
-#97 := (not #96)
-#107 := (and #97 #104)
-#93 := (not #92)
-#100 := (and #93 #97)
-#110 := (or #100 #107)
-#155 := (ite #110 #30 #135)
-#158 := (ite #12 0::int #155)
-#161 := (ite #13 #8 #158)
-#164 := (= #29 #161)
-#167 := (forall (vars (?v0 int) (?v1 int)) #164)
-#246 := (iff #167 #245)
-#243 := (iff #164 #242)
-#240 := (= #161 #239)
-#237 := (= #158 #236)
-#234 := (= #155 #233)
-#218 := (iff #110 #215)
-#212 := (or #209 #201)
-#216 := (iff #212 #215)
-#217 := [rewrite]: #216
-#213 := (iff #110 #212)
-#210 := (iff #107 #201)
-#211 := [rewrite]: #210
-#198 := (iff #100 #209)
-#199 := [rewrite]: #198
-#214 := [monotonicity #199 #211]: #213
-#219 := [trans #214 #217]: #218
-#235 := [monotonicity #219]: #234
-#238 := [monotonicity #235]: #237
-#241 := [monotonicity #238]: #240
-#244 := [monotonicity #241]: #243
-#247 := [quant-intro #244]: #246
-#196 := (~ #167 #167)
-#192 := (~ #164 #164)
-#206 := [refl]: #192
-#197 := [nnf-pos #206]: #196
-#23 := (- #9)
-#22 := (- #8)
-#31 := (mod #22 #23)
-#32 := (- #31)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#33 := (ite #20 #30 #32)
-#34 := (ite #12 0::int #33)
-#35 := (ite #13 #8 #34)
-#36 := (= #29 #35)
-#37 := (forall (vars (?v0 int) (?v1 int)) #36)
-#170 := (iff #37 #167)
-#64 := (and #16 #18)
-#67 := (or #17 #64)
-#140 := (ite #67 #30 #135)
-#143 := (ite #12 0::int #140)
-#146 := (ite #13 #8 #143)
-#149 := (= #29 #146)
-#152 := (forall (vars (?v0 int) (?v1 int)) #149)
-#168 := (iff #152 #167)
-#165 := (iff #149 #164)
-#162 := (= #146 #161)
-#159 := (= #143 #158)
-#156 := (= #140 #155)
-#111 := (iff #67 #110)
-#108 := (iff #64 #107)
-#105 := (iff #18 #104)
-#106 := [rewrite]: #105
-#98 := (iff #16 #97)
-#99 := [rewrite]: #98
-#109 := [monotonicity #99 #106]: #108
-#101 := (iff #17 #100)
-#94 := (iff #15 #93)
-#95 := [rewrite]: #94
-#102 := [monotonicity #95 #99]: #101
-#112 := [monotonicity #102 #109]: #111
-#157 := [monotonicity #112]: #156
-#160 := [monotonicity #157]: #159
-#163 := [monotonicity #160]: #162
-#166 := [monotonicity #163]: #165
-#169 := [quant-intro #166]: #168
-#153 := (iff #37 #152)
-#150 := (iff #36 #149)
-#147 := (= #35 #146)
-#144 := (= #34 #143)
-#141 := (= #33 #140)
-#138 := (= #32 #135)
-#132 := (- #129)
-#136 := (= #132 #135)
-#137 := [rewrite]: #136
-#133 := (= #32 #132)
-#130 := (= #31 #129)
-#75 := (= #23 #74)
-#76 := [rewrite]: #75
-#72 := (= #22 #71)
-#73 := [rewrite]: #72
-#131 := [monotonicity #73 #76]: #130
-#134 := [monotonicity #131]: #133
-#139 := [trans #134 #137]: #138
-#68 := (iff #20 #67)
-#65 := (iff #19 #64)
-#66 := [rewrite]: #65
-#69 := [monotonicity #66]: #68
-#142 := [monotonicity #69 #139]: #141
-#145 := [monotonicity #142]: #144
-#148 := [monotonicity #145]: #147
-#151 := [monotonicity #148]: #150
-#154 := [quant-intro #151]: #153
-#171 := [trans #154 #169]: #170
-#128 := [asserted]: #37
-#172 := [mp #128 #171]: #167
-#207 := [mp~ #172 #197]: #167
-#248 := [mp #207 #247]: #245
-#279 := [mp #248 #278]: #276
-#764 := [mp #279 #763]: #759
-#669 := (not #759)
-#670 := (or #669 #186)
-#337 := (* -1::int -3::int)
-#422 := (* -1::int -5::int)
-#423 := (mod #422 #337)
-#424 := (+ #180 #423)
-#415 := (= #424 0::int)
-#426 := (mod -5::int -3::int)
-#427 := (* -1::int #426)
-#355 := (+ #180 #427)
-#428 := (= #355 0::int)
-#425 := (<= -3::int 0::int)
-#429 := (<= -5::int 0::int)
-#408 := (or #429 #425)
-#743 := (not #408)
-#745 := (>= -5::int 0::int)
-#402 := (or #425 #745)
-#532 := (not #402)
-#739 := (or #532 #743)
-#413 := (ite #739 #428 #415)
-#414 := (= #180 0::int)
-#746 := (= -5::int 0::int)
-#747 := (ite #746 #414 #413)
-#748 := (= -5::int #180)
-#749 := (= -3::int 0::int)
-#744 := (ite #749 #748 #747)
-#671 := (or #669 #744)
-#672 := (iff #671 #670)
-#657 := (iff #670 #670)
-#659 := [rewrite]: #657
-#679 := (iff #744 #186)
-#405 := (= #180 -5::int)
-#666 := (ite false #405 #186)
-#678 := (iff #666 #186)
-#674 := [rewrite]: #678
-#676 := (iff #744 #666)
-#667 := (iff #747 #186)
-#685 := (ite false #414 #186)
-#673 := (iff #685 #186)
-#675 := [rewrite]: #673
-#682 := (iff #747 #685)
-#591 := (iff #413 #186)
-#435 := 1::int
-#696 := (= #180 1::int)
-#587 := (ite false #696 #186)
-#589 := (iff #587 #186)
-#590 := [rewrite]: #589
-#580 := (iff #413 #587)
-#584 := (iff #415 #186)
-#574 := (+ 2::int #180)
-#585 := (= #574 0::int)
-#684 := (iff #585 #186)
-#569 := [rewrite]: #684
-#680 := (iff #415 #585)
-#535 := (= #424 #574)
-#691 := (+ #180 2::int)
-#575 := (= #691 #574)
-#576 := [rewrite]: #575
-#689 := (= #424 #691)
-#688 := (= #423 2::int)
-#702 := (mod 5::int 3::int)
-#534 := (= #702 2::int)
-#687 := [rewrite]: #534
-#531 := (= #423 #702)
-#701 := (= #337 3::int)
-#695 := [rewrite]: #701
-#694 := (= #422 5::int)
-#700 := [rewrite]: #694
-#533 := [monotonicity #700 #695]: #531
-#690 := [trans #533 #687]: #688
-#692 := [monotonicity #690]: #689
-#683 := [trans #692 #576]: #535
-#681 := [monotonicity #683]: #680
-#586 := [trans #681 #569]: #584
-#699 := (iff #428 #696)
-#704 := (+ -1::int #180)
-#552 := (= #704 0::int)
-#697 := (iff #552 #696)
-#698 := [rewrite]: #697
-#553 := (iff #428 #552)
-#705 := (= #355 #704)
-#707 := (+ #180 -1::int)
-#709 := (= #707 #704)
-#710 := [rewrite]: #709
-#708 := (= #355 #707)
-#717 := (= #427 -1::int)
-#715 := (* -1::int 1::int)
-#716 := (= #715 -1::int)
-#713 := [rewrite]: #716
-#430 := (= #427 #715)
-#436 := (= #426 1::int)
-#714 := [rewrite]: #436
-#712 := [monotonicity #714]: #430
-#706 := [trans #712 #713]: #717
-#703 := [monotonicity #706]: #708
-#711 := [trans #703 #710]: #705
-#554 := [monotonicity #711]: #553
-#693 := [trans #554 #698]: #699
-#433 := (iff #739 false)
-#444 := (or false false)
-#719 := (iff #444 false)
-#723 := [rewrite]: #719
-#718 := (iff #739 #444)
-#721 := (iff #743 false)
-#1 := true
-#727 := (not true)
-#734 := (iff #727 false)
-#736 := [rewrite]: #734
-#343 := (iff #743 #727)
-#449 := (iff #408 true)
-#738 := (or true true)
-#447 := (iff #738 true)
-#448 := [rewrite]: #447
-#724 := (iff #408 #738)
-#385 := (iff #425 true)
-#726 := [rewrite]: #385
-#737 := (iff #429 true)
-#735 := [rewrite]: #737
-#725 := [monotonicity #735 #726]: #724
-#450 := [trans #725 #448]: #449
-#720 := [monotonicity #450]: #343
-#443 := [trans #720 #736]: #721
-#376 := (iff #532 false)
-#733 := (iff #532 #727)
-#731 := (iff #402 true)
-#391 := (or true false)
-#729 := (iff #391 true)
-#730 := [rewrite]: #729
-#392 := (iff #402 #391)
-#389 := (iff #745 false)
-#390 := [rewrite]: #389
-#728 := [monotonicity #726 #390]: #392
-#732 := [trans #728 #730]: #731
-#371 := [monotonicity #732]: #733
-#377 := [trans #371 #736]: #376
-#722 := [monotonicity #377 #443]: #718
-#434 := [trans #722 #723]: #433
-#588 := [monotonicity #434 #693 #586]: #580
-#592 := [trans #588 #590]: #591
-#401 := (iff #746 false)
-#742 := [rewrite]: #401
-#686 := [monotonicity #742 #592]: #682
-#664 := [trans #686 #675]: #667
-#400 := (iff #748 #405)
-#741 := [rewrite]: #400
-#750 := (iff #749 false)
-#740 := [rewrite]: #750
-#677 := [monotonicity #740 #741 #664]: #676
-#668 := [trans #677 #674]: #679
-#656 := [monotonicity #668]: #672
-#660 := [trans #656 #659]: #672
-#665 := [quant-inst]: #671
-#661 := [mp #665 #660]: #670
-[unit-resolution #661 #764 #194]: false
-unsat
-dce856f3e21ae2f97ea031dc9fe8a9a3d8bcd965 57 0
+1c29e683939b29daf6c0c39d6dc5211f90494d2b 57 0
 #2 := false
 #36 := 0::int
 decl f4 :: int
@@ -43828,7 +44710,7 @@
 #77 := [th-lemma]: #76
 [unit-resolution #77 #74 #56]: false
 unsat
-1c2c19656480fea9eac0d64b2d7713bbaf1ce18e 57 0
+f5bdd04ade0ff62ddedb44ca71a4ab32cea43ed1 57 0
 #2 := false
 #36 := 0::int
 decl f4 :: int
@@ -43886,7 +44768,7 @@
 #77 := [unit-resolution #76 #73]: #71
 [th-lemma #77 #56 #72]: false
 unsat
-dcce7f1c4cfeb0855ad108c9c837348dc081133f 103 0
+db7c87847633d861c99b1fec84d68fb0617d9c27 103 0
 #2 := false
 #45 := 0::int
 decl f5 :: int
@@ -43990,335 +44872,7 @@
 #122 := [unit-resolution #121 #118]: #115
 [th-lemma #122 #92 #93]: false
 unsat
-cc54796722d31f22331ee38cd4d2e526cb5b071c 327 0
-#2 := false
-#11 := 0::int
-decl f5 :: int
-#38 := f5
-#418 := (<= f5 0::int)
-#734 := (>= f5 0::int)
-#380 := (not #734)
-#723 := (not #418)
-#727 := (or #723 #380)
-#690 := (not #727)
-#39 := 3::int
-#415 := (mod f5 3::int)
-#66 := -1::int
-#416 := (* -1::int #415)
-decl f4 :: (-> int int int)
-#40 := (f4 f5 3::int)
-#344 := (+ #40 #416)
-#417 := (= #344 0::int)
-#563 := (not #417)
-#520 := (<= #344 0::int)
-#623 := (not #520)
-#631 := (>= #415 3::int)
-#632 := (not #631)
-#1 := true
-#57 := [true-axiom]: true
-#619 := (or false #632)
-#620 := [th-lemma]: #619
-#621 := [unit-resolution #620 #57]: #632
-#622 := [hypothesis]: #520
-#173 := (>= #40 3::int)
-#41 := (< #40 3::int)
-#42 := (not #41)
-#181 := (iff #42 #173)
-#172 := (not #173)
-#176 := (not #172)
-#179 := (iff #176 #173)
-#180 := [rewrite]: #179
-#177 := (iff #42 #176)
-#174 := (iff #41 #172)
-#175 := [rewrite]: #174
-#178 := [monotonicity #175]: #177
-#182 := [trans #178 #180]: #181
-#169 := [asserted]: #42
-#183 := [mp #169 #182]: #173
-#617 := [th-lemma #183 #622 #621]: false
-#609 := [lemma #617]: #623
-#626 := (or #563 #520)
-#637 := [th-lemma]: #626
-#615 := [unit-resolution #637 #609]: #563
-#614 := (or #690 #417)
-#438 := -3::int
-#411 := (* -1::int f5)
-#709 := (mod #411 -3::int)
-#433 := (+ #40 #709)
-#708 := (= #433 0::int)
-#423 := (ite #727 #417 #708)
-#403 := (= #40 0::int)
-#735 := (= f5 0::int)
-#703 := (ite #735 #403 #423)
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#29 := (f4 #8 #9)
-#747 := (pattern #29)
-#70 := (* -1::int #9)
-#67 := (* -1::int #8)
-#125 := (mod #67 #70)
-#258 := (+ #29 #125)
-#259 := (= #258 0::int)
-#30 := (mod #8 #9)
-#255 := (* -1::int #30)
-#256 := (+ #29 #255)
-#257 := (= #256 0::int)
-#92 := (<= #9 0::int)
-#88 := (<= #8 0::int)
-#197 := (or #88 #92)
-#198 := (not #197)
-#99 := (>= #8 0::int)
-#189 := (or #92 #99)
-#190 := (not #189)
-#204 := (or #190 #198)
-#260 := (ite #204 #257 #259)
-#254 := (= #29 0::int)
-#12 := (= #8 0::int)
-#261 := (ite #12 #254 #260)
-#253 := (= #8 #29)
-#13 := (= #9 0::int)
-#262 := (ite #13 #253 #261)
-#748 := (forall (vars (?v0 int) (?v1 int)) (:pat #747) #262)
-#265 := (forall (vars (?v0 int) (?v1 int)) #262)
-#751 := (iff #265 #748)
-#749 := (iff #262 #262)
-#750 := [refl]: #749
-#752 := [quant-intro #750]: #751
-#131 := (* -1::int #125)
-#222 := (ite #204 #30 #131)
-#225 := (ite #12 0::int #222)
-#228 := (ite #13 #8 #225)
-#231 := (= #29 #228)
-#234 := (forall (vars (?v0 int) (?v1 int)) #231)
-#266 := (iff #234 #265)
-#263 := (iff #231 #262)
-#264 := [rewrite]: #263
-#267 := [quant-intro #264]: #266
-#100 := (not #99)
-#93 := (not #92)
-#103 := (and #93 #100)
-#89 := (not #88)
-#96 := (and #89 #93)
-#106 := (or #96 #103)
-#151 := (ite #106 #30 #131)
-#154 := (ite #12 0::int #151)
-#157 := (ite #13 #8 #154)
-#160 := (= #29 #157)
-#163 := (forall (vars (?v0 int) (?v1 int)) #160)
-#235 := (iff #163 #234)
-#232 := (iff #160 #231)
-#229 := (= #157 #228)
-#226 := (= #154 #225)
-#223 := (= #151 #222)
-#207 := (iff #106 #204)
-#201 := (or #198 #190)
-#205 := (iff #201 #204)
-#206 := [rewrite]: #205
-#202 := (iff #106 #201)
-#199 := (iff #103 #190)
-#200 := [rewrite]: #199
-#187 := (iff #96 #198)
-#188 := [rewrite]: #187
-#203 := [monotonicity #188 #200]: #202
-#208 := [trans #203 #206]: #207
-#224 := [monotonicity #208]: #223
-#227 := [monotonicity #224]: #226
-#230 := [monotonicity #227]: #229
-#233 := [monotonicity #230]: #232
-#236 := [quant-intro #233]: #235
-#185 := (~ #163 #163)
-#170 := (~ #160 #160)
-#195 := [refl]: #170
-#186 := [nnf-pos #195]: #185
-#23 := (- #9)
-#22 := (- #8)
-#31 := (mod #22 #23)
-#32 := (- #31)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#33 := (ite #20 #30 #32)
-#34 := (ite #12 0::int #33)
-#35 := (ite #13 #8 #34)
-#36 := (= #29 #35)
-#37 := (forall (vars (?v0 int) (?v1 int)) #36)
-#166 := (iff #37 #163)
-#60 := (and #16 #18)
-#63 := (or #17 #60)
-#136 := (ite #63 #30 #131)
-#139 := (ite #12 0::int #136)
-#142 := (ite #13 #8 #139)
-#145 := (= #29 #142)
-#148 := (forall (vars (?v0 int) (?v1 int)) #145)
-#164 := (iff #148 #163)
-#161 := (iff #145 #160)
-#158 := (= #142 #157)
-#155 := (= #139 #154)
-#152 := (= #136 #151)
-#107 := (iff #63 #106)
-#104 := (iff #60 #103)
-#101 := (iff #18 #100)
-#102 := [rewrite]: #101
-#94 := (iff #16 #93)
-#95 := [rewrite]: #94
-#105 := [monotonicity #95 #102]: #104
-#97 := (iff #17 #96)
-#90 := (iff #15 #89)
-#91 := [rewrite]: #90
-#98 := [monotonicity #91 #95]: #97
-#108 := [monotonicity #98 #105]: #107
-#153 := [monotonicity #108]: #152
-#156 := [monotonicity #153]: #155
-#159 := [monotonicity #156]: #158
-#162 := [monotonicity #159]: #161
-#165 := [quant-intro #162]: #164
-#149 := (iff #37 #148)
-#146 := (iff #36 #145)
-#143 := (= #35 #142)
-#140 := (= #34 #139)
-#137 := (= #33 #136)
-#134 := (= #32 #131)
-#128 := (- #125)
-#132 := (= #128 #131)
-#133 := [rewrite]: #132
-#129 := (= #32 #128)
-#126 := (= #31 #125)
-#71 := (= #23 #70)
-#72 := [rewrite]: #71
-#68 := (= #22 #67)
-#69 := [rewrite]: #68
-#127 := [monotonicity #69 #72]: #126
-#130 := [monotonicity #127]: #129
-#135 := [trans #130 #133]: #134
-#64 := (iff #20 #63)
-#61 := (iff #19 #60)
-#62 := [rewrite]: #61
-#65 := [monotonicity #62]: #64
-#138 := [monotonicity #65 #135]: #137
-#141 := [monotonicity #138]: #140
-#144 := [monotonicity #141]: #143
-#147 := [monotonicity #144]: #146
-#150 := [quant-intro #147]: #149
-#167 := [trans #150 #165]: #166
-#124 := [asserted]: #37
-#168 := [mp #124 #167]: #163
-#196 := [mp~ #168 #186]: #163
-#237 := [mp #196 #236]: #234
-#268 := [mp #237 #267]: #265
-#753 := [mp #268 #752]: #748
-#692 := (not #748)
-#693 := (or #692 #703)
-#326 := (* -1::int 3::int)
-#412 := (mod #411 #326)
-#413 := (+ #40 #412)
-#404 := (= #413 0::int)
-#414 := (<= 3::int 0::int)
-#397 := (or #418 #414)
-#732 := (not #397)
-#391 := (or #414 #734)
-#521 := (not #391)
-#728 := (or #521 #732)
-#402 := (ite #728 #417 #404)
-#736 := (ite #735 #403 #402)
-#737 := (= f5 #40)
-#738 := (= 3::int 0::int)
-#733 := (ite #738 #737 #736)
-#698 := (or #692 #733)
-#694 := (iff #698 #693)
-#541 := (iff #693 #693)
-#542 := [rewrite]: #541
-#696 := (iff #733 #703)
-#701 := (ite false #737 #703)
-#706 := (iff #701 #703)
-#695 := [rewrite]: #706
-#705 := (iff #733 #701)
-#704 := (iff #736 #703)
-#424 := (iff #402 #423)
-#712 := (iff #404 #708)
-#707 := (= #413 #433)
-#710 := (= #412 #709)
-#439 := (= #326 -3::int)
-#332 := [rewrite]: #439
-#432 := [monotonicity #332]: #710
-#711 := [monotonicity #432]: #707
-#422 := [monotonicity #711]: #712
-#436 := (iff #728 #727)
-#366 := (or #380 #723)
-#713 := (iff #366 #727)
-#714 := [rewrite]: #713
-#726 := (iff #728 #366)
-#725 := (iff #732 #723)
-#722 := (iff #397 #418)
-#718 := (or #418 false)
-#721 := (iff #718 #418)
-#716 := [rewrite]: #721
-#719 := (iff #397 #718)
-#389 := (iff #414 false)
-#394 := [rewrite]: #389
-#720 := [monotonicity #394]: #719
-#360 := [trans #720 #716]: #722
-#365 := [monotonicity #360]: #725
-#381 := (iff #521 #380)
-#378 := (iff #391 #734)
-#730 := (or false #734)
-#374 := (iff #730 #734)
-#715 := [rewrite]: #374
-#390 := (iff #391 #730)
-#731 := [monotonicity #394]: #390
-#379 := [trans #731 #715]: #378
-#717 := [monotonicity #379]: #381
-#724 := [monotonicity #717 #365]: #726
-#437 := [trans #724 #714]: #436
-#425 := [monotonicity #437 #422]: #424
-#419 := [monotonicity #425]: #704
-#739 := (iff #738 false)
-#729 := [rewrite]: #739
-#702 := [monotonicity #729 #419]: #705
-#697 := [trans #702 #695]: #696
-#700 := [monotonicity #697]: #694
-#543 := [trans #700 #542]: #694
-#699 := [quant-inst]: #698
-#685 := [mp #699 #543]: #693
-#616 := [unit-resolution #685 #753]: #703
-#670 := (not #735)
-#669 := (not #703)
-#611 := (or #669 #670)
-#576 := (not #403)
-#686 := (<= #40 0::int)
-#618 := (not #686)
-#648 := (or #618 #172)
-#649 := [th-lemma]: #648
-#605 := [unit-resolution #649 #183]: #618
-#606 := (or #576 #686)
-#607 := [th-lemma]: #606
-#610 := [unit-resolution #607 #605]: #576
-#673 := (or #669 #670 #403)
-#558 := [def-axiom]: #673
-#612 := [unit-resolution #558 #610]: #611
-#613 := [unit-resolution #612 #616]: #670
-#573 := (or #669 #735 #423)
-#575 := [def-axiom]: #573
-#608 := [unit-resolution #575 #613 #616]: #423
-#677 := (not #423)
-#679 := (or #677 #690 #417)
-#680 := [def-axiom]: #679
-#587 := [unit-resolution #680 #608]: #614
-#588 := [unit-resolution #587 #615]: #690
-#688 := (or #727 #418)
-#682 := [def-axiom]: #688
-#589 := [unit-resolution #682 #588]: #418
-#683 := (or #727 #734)
-#689 := [def-axiom]: #683
-#482 := [unit-resolution #689 #588]: #734
-#593 := (or #735 #723 #380)
-#594 := [th-lemma]: #593
-#595 := [unit-resolution #594 #613]: #727
-[unit-resolution #595 #482 #589]: false
-unsat
-ce21500c8b03b5e1a394643a4c361d6ebb12e85f 86 0
+30ec056750b76c9c44be73825eb8959f04eaf2af 86 0
 #2 := false
 decl f3 :: int
 #8 := f3
@@ -44405,7 +44959,7 @@
 #105 := [trans #104 #83]: #51
 [unit-resolution #57 #105]: false
 unsat
-22be385fa1e47d4d76e059f3bfa8188c80f22cdd 68 0
+87cffe59a43ecd3ffcda40df06578c44d2a9feac 68 0
 #2 := false
 #8 := 0::int
 decl f3 :: int
@@ -44474,7 +45028,7 @@
 #88 := [trans #87 #84]: #13
 [unit-resolution #71 #88]: false
 unsat
-2253342327dba5f019f3a86fda5fc5ff24024127 124 0
+90c0dd925b0006ab13047c20a6f0e8d277ec4cdf 124 0
 #2 := false
 #13 := 0::int
 decl f4 :: int
@@ -44599,7 +45153,7 @@
 #145 := [unit-resolution #117 #144]: #106
 [th-lemma #131 #143 #141 #85 #145]: false
 unsat
-8a7449660323be981d8338388af366ba304fb872 57 0
+de67b37b9c3eba477099ff5d9b09e840de376762 57 0
 #2 := false
 #36 := 0::int
 decl f4 :: int
@@ -44657,353 +45211,7 @@
 #77 := [unit-resolution #76 #73]: #70
 [th-lemma #77 #56 #72]: false
 unsat
-1c20abf656adf6300dff05d39209b84754f973b9 345 0
-#2 := false
-#39 := 3::int
-decl f4 :: (-> int int int)
-decl f5 :: int
-#38 := f5
-#40 := (f4 f5 3::int)
-#441 := (mod #40 3::int)
-#657 := (>= #441 3::int)
-#658 := (not #657)
-#1 := true
-#59 := [true-axiom]: true
-#647 := (or false #658)
-#642 := [th-lemma]: #647
-#648 := [unit-resolution #642 #59]: #658
-#11 := 0::int
-#68 := -1::int
-#436 := (* -1::int #40)
-#600 := (+ f5 #436)
-#601 := (<= #600 0::int)
-#172 := (= f5 #40)
-#188 := (>= f5 3::int)
-#187 := (not #188)
-#178 := (not #172)
-#194 := (or #178 #187)
-#199 := (not #194)
-#42 := (< f5 3::int)
-#41 := (= #40 f5)
-#43 := (implies #41 #42)
-#44 := (not #43)
-#202 := (iff #44 #199)
-#179 := (or #42 #178)
-#184 := (not #179)
-#200 := (iff #184 #199)
-#197 := (iff #179 #194)
-#191 := (or #187 #178)
-#195 := (iff #191 #194)
-#196 := [rewrite]: #195
-#192 := (iff #179 #191)
-#189 := (iff #42 #187)
-#190 := [rewrite]: #189
-#193 := [monotonicity #190]: #192
-#198 := [trans #193 #196]: #197
-#201 := [monotonicity #198]: #200
-#185 := (iff #44 #184)
-#182 := (iff #43 #179)
-#175 := (implies #172 #42)
-#180 := (iff #175 #179)
-#181 := [rewrite]: #180
-#176 := (iff #43 #175)
-#173 := (iff #41 #172)
-#174 := [rewrite]: #173
-#177 := [monotonicity #174]: #176
-#183 := [trans #177 #181]: #182
-#186 := [monotonicity #183]: #185
-#203 := [trans #186 #201]: #202
-#171 := [asserted]: #44
-#204 := [mp #171 #203]: #199
-#205 := [not-or-elim #204]: #172
-#634 := (or #178 #601)
-#630 := [th-lemma]: #634
-#631 := [unit-resolution #630 #205]: #601
-#206 := [not-or-elim #204]: #188
-#438 := (f4 #40 3::int)
-#602 := (* -1::int #438)
-#603 := (+ #40 #602)
-#604 := (<= #603 0::int)
-#763 := (= #40 #438)
-#635 := (= #438 #40)
-#632 := [symm #205]: #41
-#636 := [monotonicity #632]: #635
-#637 := [symm #636]: #763
-#638 := (not #763)
-#633 := (or #638 #604)
-#639 := [th-lemma]: #633
-#612 := [unit-resolution #639 #637]: #604
-#369 := (* -1::int #441)
-#442 := (+ #438 #369)
-#707 := (<= #442 0::int)
-#439 := (= #442 0::int)
-#738 := -3::int
-#462 := (mod #436 -3::int)
-#357 := (+ #438 #462)
-#457 := (= #357 0::int)
-#422 := (<= #40 0::int)
-#750 := (not #422)
-#416 := (>= #40 0::int)
-#406 := (not #416)
-#751 := (or #406 #750)
-#736 := (ite #751 #439 #457)
-#760 := (= #438 0::int)
-#761 := (= #40 0::int)
-#447 := (ite #761 #760 #736)
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#29 := (f4 #8 #9)
-#772 := (pattern #29)
-#72 := (* -1::int #9)
-#69 := (* -1::int #8)
-#127 := (mod #69 #72)
-#283 := (+ #29 #127)
-#284 := (= #283 0::int)
-#30 := (mod #8 #9)
-#280 := (* -1::int #30)
-#281 := (+ #29 #280)
-#282 := (= #281 0::int)
-#94 := (<= #9 0::int)
-#90 := (<= #8 0::int)
-#224 := (or #90 #94)
-#225 := (not #224)
-#101 := (>= #8 0::int)
-#216 := (or #94 #101)
-#217 := (not #216)
-#229 := (or #217 #225)
-#285 := (ite #229 #282 #284)
-#279 := (= #29 0::int)
-#12 := (= #8 0::int)
-#286 := (ite #12 #279 #285)
-#278 := (= #8 #29)
-#13 := (= #9 0::int)
-#287 := (ite #13 #278 #286)
-#773 := (forall (vars (?v0 int) (?v1 int)) (:pat #772) #287)
-#290 := (forall (vars (?v0 int) (?v1 int)) #287)
-#776 := (iff #290 #773)
-#774 := (iff #287 #287)
-#775 := [refl]: #774
-#777 := [quant-intro #775]: #776
-#133 := (* -1::int #127)
-#247 := (ite #229 #30 #133)
-#250 := (ite #12 0::int #247)
-#253 := (ite #13 #8 #250)
-#256 := (= #29 #253)
-#259 := (forall (vars (?v0 int) (?v1 int)) #256)
-#291 := (iff #259 #290)
-#288 := (iff #256 #287)
-#289 := [rewrite]: #288
-#292 := [quant-intro #289]: #291
-#102 := (not #101)
-#95 := (not #94)
-#105 := (and #95 #102)
-#91 := (not #90)
-#98 := (and #91 #95)
-#108 := (or #98 #105)
-#153 := (ite #108 #30 #133)
-#156 := (ite #12 0::int #153)
-#159 := (ite #13 #8 #156)
-#162 := (= #29 #159)
-#165 := (forall (vars (?v0 int) (?v1 int)) #162)
-#260 := (iff #165 #259)
-#257 := (iff #162 #256)
-#254 := (= #159 #253)
-#251 := (= #156 #250)
-#248 := (= #153 #247)
-#232 := (iff #108 #229)
-#226 := (or #225 #217)
-#230 := (iff #226 #229)
-#231 := [rewrite]: #230
-#227 := (iff #108 #226)
-#214 := (iff #105 #217)
-#215 := [rewrite]: #214
-#212 := (iff #98 #225)
-#213 := [rewrite]: #212
-#228 := [monotonicity #213 #215]: #227
-#233 := [trans #228 #231]: #232
-#249 := [monotonicity #233]: #248
-#252 := [monotonicity #249]: #251
-#255 := [monotonicity #252]: #254
-#258 := [monotonicity #255]: #257
-#261 := [quant-intro #258]: #260
-#210 := (~ #165 #165)
-#207 := (~ #162 #162)
-#222 := [refl]: #207
-#211 := [nnf-pos #222]: #210
-#23 := (- #9)
-#22 := (- #8)
-#31 := (mod #22 #23)
-#32 := (- #31)
-#16 := (< 0::int #9)
-#18 := (< #8 0::int)
-#19 := (and #18 #16)
-#15 := (< 0::int #8)
-#17 := (and #15 #16)
-#20 := (or #17 #19)
-#33 := (ite #20 #30 #32)
-#34 := (ite #12 0::int #33)
-#35 := (ite #13 #8 #34)
-#36 := (= #29 #35)
-#37 := (forall (vars (?v0 int) (?v1 int)) #36)
-#168 := (iff #37 #165)
-#62 := (and #16 #18)
-#65 := (or #17 #62)
-#138 := (ite #65 #30 #133)
-#141 := (ite #12 0::int #138)
-#144 := (ite #13 #8 #141)
-#147 := (= #29 #144)
-#150 := (forall (vars (?v0 int) (?v1 int)) #147)
-#166 := (iff #150 #165)
-#163 := (iff #147 #162)
-#160 := (= #144 #159)
-#157 := (= #141 #156)
-#154 := (= #138 #153)
-#109 := (iff #65 #108)
-#106 := (iff #62 #105)
-#103 := (iff #18 #102)
-#104 := [rewrite]: #103
-#96 := (iff #16 #95)
-#97 := [rewrite]: #96
-#107 := [monotonicity #97 #104]: #106
-#99 := (iff #17 #98)
-#92 := (iff #15 #91)
-#93 := [rewrite]: #92
-#100 := [monotonicity #93 #97]: #99
-#110 := [monotonicity #100 #107]: #109
-#155 := [monotonicity #110]: #154
-#158 := [monotonicity #155]: #157
-#161 := [monotonicity #158]: #160
-#164 := [monotonicity #161]: #163
-#167 := [quant-intro #164]: #166
-#151 := (iff #37 #150)
-#148 := (iff #36 #147)
-#145 := (= #35 #144)
-#142 := (= #34 #141)
-#139 := (= #33 #138)
-#136 := (= #32 #133)
-#130 := (- #127)
-#134 := (= #130 #133)
-#135 := [rewrite]: #134
-#131 := (= #32 #130)
-#128 := (= #31 #127)
-#73 := (= #23 #72)
-#74 := [rewrite]: #73
-#70 := (= #22 #69)
-#71 := [rewrite]: #70
-#129 := [monotonicity #71 #74]: #128
-#132 := [monotonicity #129]: #131
-#137 := [trans #132 #135]: #136
-#66 := (iff #20 #65)
-#63 := (iff #19 #62)
-#64 := [rewrite]: #63
-#67 := [monotonicity #64]: #66
-#140 := [monotonicity #67 #137]: #139
-#143 := [monotonicity #140]: #142
-#146 := [monotonicity #143]: #145
-#149 := [monotonicity #146]: #148
-#152 := [quant-intro #149]: #151
-#169 := [trans #152 #167]: #168
-#126 := [asserted]: #37
-#170 := [mp #126 #169]: #165
-#223 := [mp~ #170 #211]: #165
-#262 := [mp #223 #261]: #259
-#293 := [mp #262 #292]: #290
-#778 := [mp #293 #777]: #773
-#731 := (not #773)
-#720 := (or #731 #447)
-#351 := (* -1::int 3::int)
-#437 := (mod #436 #351)
-#429 := (+ #438 #437)
-#440 := (= #429 0::int)
-#443 := (<= 3::int 0::int)
-#757 := (or #422 #443)
-#759 := (not #757)
-#546 := (or #443 #416)
-#753 := (not #546)
-#427 := (or #753 #759)
-#428 := (ite #427 #439 #440)
-#762 := (ite #761 #760 #428)
-#758 := (= 3::int 0::int)
-#764 := (ite #758 #763 #762)
-#721 := (or #731 #764)
-#717 := (iff #721 #720)
-#723 := (iff #720 #720)
-#724 := [rewrite]: #723
-#730 := (iff #764 #447)
-#450 := (ite false #763 #447)
-#444 := (iff #450 #447)
-#726 := [rewrite]: #444
-#728 := (iff #764 #450)
-#448 := (iff #762 #447)
-#733 := (iff #428 #736)
-#458 := (iff #440 #457)
-#734 := (= #429 #357)
-#463 := (= #437 #462)
-#739 := (= #351 -3::int)
-#461 := [rewrite]: #739
-#464 := [monotonicity #461]: #463
-#735 := [monotonicity #464]: #734
-#732 := [monotonicity #735]: #458
-#749 := (iff #427 #751)
-#390 := (iff #759 #750)
-#385 := (iff #757 #422)
-#744 := (or #422 false)
-#741 := (iff #744 #422)
-#747 := [rewrite]: #741
-#745 := (iff #757 #744)
-#419 := (iff #443 false)
-#755 := [rewrite]: #419
-#746 := [monotonicity #755]: #745
-#748 := [trans #746 #747]: #385
-#391 := [monotonicity #748]: #390
-#742 := (iff #753 #406)
-#404 := (iff #546 #416)
-#415 := (or false #416)
-#740 := (iff #415 #416)
-#403 := [rewrite]: #740
-#756 := (iff #546 #415)
-#399 := [monotonicity #755]: #756
-#405 := [trans #399 #403]: #404
-#743 := [monotonicity #405]: #742
-#752 := [monotonicity #743 #391]: #749
-#737 := [monotonicity #752 #732]: #733
-#449 := [monotonicity #737]: #448
-#754 := (iff #758 false)
-#414 := [rewrite]: #754
-#729 := [monotonicity #414 #449]: #728
-#727 := [trans #729 #726]: #730
-#718 := [monotonicity #727]: #717
-#719 := [trans #718 #724]: #717
-#722 := [quant-inst]: #721
-#725 := [mp #722 #719]: #720
-#613 := [unit-resolution #725 #778]: #447
-#589 := (not #761)
-#614 := (not #601)
-#507 := (or #750 #187 #614)
-#618 := [th-lemma]: #507
-#619 := [unit-resolution #618 #206 #631]: #750
-#620 := (or #589 #422)
-#625 := [th-lemma]: #620
-#621 := [unit-resolution #625 #619]: #589
-#588 := (not #447)
-#697 := (or #588 #761 #736)
-#599 := [def-axiom]: #697
-#622 := [unit-resolution #599 #621 #613]: #736
-#568 := (or #751 #422)
-#710 := [def-axiom]: #568
-#623 := [unit-resolution #710 #619]: #751
-#711 := (not #751)
-#709 := (not #736)
-#716 := (or #709 #711 #439)
-#545 := [def-axiom]: #716
-#626 := [unit-resolution #545 #623 #622]: #439
-#701 := (not #439)
-#627 := (or #701 #707)
-#628 := [th-lemma]: #627
-#624 := [unit-resolution #628 #626]: #707
-[th-lemma #624 #612 #206 #631 #648]: false
-unsat
-4dce9532aca5ad04bb32a324214a8d6e10933595 57 0
+8f0b2c755e25c9b0ef5750035ec56d21d3c6b32f 57 0
 #2 := false
 #35 := 0::int
 decl f4 :: int
@@ -45061,7 +45269,7 @@
 #77 := [th-lemma]: #76
 [unit-resolution #77 #74 #56]: false
 unsat
-4c651ded3b3a1d80978c76e6d7d096efa4221c9f 103 0
+ba050e782d9473f46d4f759400e69d34d257264e 103 0
 #2 := false
 #45 := 0::int
 decl f3 :: int
@@ -45165,7 +45373,7 @@
 #122 := [unit-resolution #121 #118]: #116
 [th-lemma #122 #91 #93]: false
 unsat
-4dcddda0c75e28e85f0009ab3200f08d892d02c4 88 0
+9c2cfdc304fee997e8f612dd9123abd86a877f29 88 0
 #2 := false
 decl f4 :: int
 #9 := f4
@@ -45254,7 +45462,7 @@
 #109 := [trans #108 #83]: #51
 [unit-resolution #57 #109]: false
 unsat
-600697e3c100a020914718001e57a63394741230 66 0
+66ca5933e35ee3faa0336268f65c331335330f57 66 0
 #2 := false
 decl f3 :: int
 #9 := f3
@@ -45321,7 +45529,7 @@
 #86 := [trans #83 #85]: #33
 [unit-resolution #69 #86]: false
 unsat
-45b4f832b01805ab14c570832adc97242b473045 228 0
+a176c257f1c89ffea971e08639242e005fb1222f 228 0
 #2 := false
 #9 := 0::int
 decl f4 :: int
@@ -45550,7 +45758,7 @@
 #250 := [unit-resolution #196 #249]: #190
 [th-lemma #236 #250 #131 #234 #248 #238]: false
 unsat
-49cc983d173a58df0ca800b897e7609030c1adbd 56 0
+efdfb7635dd41e99f117d8463d9d75f3897207dd 56 0
 #2 := false
 #11 := 1::int
 decl f3 :: int
@@ -45607,7 +45815,7 @@
 #71 := [not-or-elim #69]: #70
 [unit-resolution #71 #85]: false
 unsat
-838b399a3a6da31b0c058aef87c3e393bc9c198f 18 0
+6a196ac012aba8c7b41b1d8437b14fc3fb23e5a4 18 0
 #2 := false
 decl f3 :: int
 #8 := f3
@@ -45626,25 +45834,7 @@
 #27 := [asserted]: #10
 [mp #27 #38]: false
 unsat
-f401bbb35719168186be7c7fdc36b726ddad542e 17 0
-#2 := false
-#8 := 0::real
-#9 := (= 0::real 0::real)
-#10 := (not #9)
-#35 := (iff #10 false)
-#1 := true
-#30 := (not true)
-#33 := (iff #30 false)
-#34 := [rewrite]: #33
-#31 := (iff #10 #30)
-#28 := (iff #9 true)
-#29 := [rewrite]: #28
-#32 := [monotonicity #29]: #31
-#36 := [trans #32 #34]: #35
-#27 := [asserted]: #10
-[mp #27 #36]: false
-unsat
-acd9efc2e486913895049a85d947303a715afaf4 51 0
+4afe860c901fd901bf34dc83b799fd986f9e4155 51 0
 #2 := false
 decl f4 :: int
 #9 := f4
@@ -45696,7 +45886,7 @@
 #33 := [asserted]: #16
 [mp #33 #68]: false
 unsat
-876d4808395b1353460558d27a1ff7e23ea1a045 56 0
+5551b99ea77d386f007b1b1979a0b6b379c3a347 56 0
 #2 := false
 decl f4 :: int
 #9 := f4
@@ -45753,51 +45943,7 @@
 #33 := [asserted]: #16
 [mp #33 #73]: false
 unsat
-2c7d006b46aef7c76688a4b4a4964ad07b6656a1 25 0
-#2 := false
-#8 := 0::real
-#9 := (- 0::real)
-#10 := (= 0::real #9)
-#11 := (not #10)
-#43 := (iff #11 false)
-#1 := true
-#38 := (not true)
-#41 := (iff #38 false)
-#42 := [rewrite]: #41
-#39 := (iff #11 #38)
-#36 := (iff #10 true)
-#31 := (= 0::real 0::real)
-#34 := (iff #31 true)
-#35 := [rewrite]: #34
-#32 := (iff #10 #31)
-#29 := (= #9 0::real)
-#30 := [rewrite]: #29
-#33 := [monotonicity #30]: #32
-#37 := [trans #33 #35]: #36
-#40 := [monotonicity #37]: #39
-#44 := [trans #40 #42]: #43
-#28 := [asserted]: #11
-[mp #28 #44]: false
-unsat
-70737154521415199c9b18902d90c0ff2d6966da 17 0
-#2 := false
-#8 := 1::real
-#9 := (= 1::real 1::real)
-#10 := (not #9)
-#35 := (iff #10 false)
-#1 := true
-#30 := (not true)
-#33 := (iff #30 false)
-#34 := [rewrite]: #33
-#31 := (iff #10 #30)
-#28 := (iff #9 true)
-#29 := [rewrite]: #28
-#32 := [monotonicity #29]: #31
-#36 := [trans #32 #34]: #35
-#27 := [asserted]: #10
-[mp #27 #36]: false
-unsat
-e1c4625af215a329500d17ad89e6dc8f577b02b1 52 0
+76fdd03dce324fb45070b1e1983f0a282cda6fa6 52 0
 #2 := false
 #40 := 0::int
 decl f4 :: int
@@ -45850,42 +45996,7 @@
 #77 := [unit-resolution #76 #65]: #53
 [unit-resolution #77 #67]: false
 unsat
-079e8a190b645549c5a2d16ab104adb956315901 34 0
-#2 := false
-#8 := 1::real
-#9 := (- 1::real)
-#10 := (= #9 1::real)
-#11 := (not #10)
-#12 := (not #11)
-#52 := (iff #12 false)
-#1 := true
-#47 := (not true)
-#50 := (iff #47 false)
-#51 := [rewrite]: #50
-#48 := (iff #12 #47)
-#45 := (iff #11 true)
-#40 := (not false)
-#43 := (iff #40 true)
-#44 := [rewrite]: #43
-#41 := (iff #11 #40)
-#38 := (iff #10 false)
-#30 := -1::real
-#33 := (= -1::real 1::real)
-#36 := (iff #33 false)
-#37 := [rewrite]: #36
-#34 := (iff #10 #33)
-#31 := (= #9 -1::real)
-#32 := [rewrite]: #31
-#35 := [monotonicity #32]: #34
-#39 := [trans #35 #37]: #38
-#42 := [monotonicity #39]: #41
-#46 := [trans #42 #44]: #45
-#49 := [monotonicity #46]: #48
-#53 := [trans #49 #51]: #52
-#29 := [asserted]: #12
-[mp #29 #53]: false
-unsat
-2eb8b7f54e63226cb1a9ab4c59e18e9b90fe4ed0 73 0
+550f7d68b0836924510ab12e5b3c4ed9f6c794bb 73 0
 #2 := false
 #58 := 0::int
 decl f4 :: int
@@ -45959,26 +46070,7 @@
 #78 := [unit-resolution #82 #74]: #83
 [unit-resolution #78 #92 #89]: false
 unsat
-a2184c17e30d6d1fdfb1b99141d590277924af4f 18 0
-#2 := false
-#9 := 1::real
-#8 := 0::real
-#10 := (< 0::real 1::real)
-#11 := (not #10)
-#38 := (iff #11 false)
-#1 := true
-#33 := (not true)
-#36 := (iff #33 false)
-#37 := [rewrite]: #36
-#34 := (iff #11 #33)
-#31 := (iff #10 true)
-#32 := [rewrite]: #31
-#35 := [monotonicity #32]: #34
-#39 := [trans #35 #37]: #38
-#28 := [asserted]: #11
-[mp #28 #39]: false
-unsat
-56d5fed825ed136d8c662448f5f823a6b191adf2 15 0
+65c77124096528f9cb7c1cb6b89625b62f5eeb2a 15 0
 #2 := false
 decl f3 :: int
 #8 := f3
@@ -45994,26 +46086,7 @@
 #28 := [asserted]: #11
 [mp #28 #34]: false
 unsat
-ed3ca239863adeeb7925561fb018b520b18a23a9 18 0
-#2 := false
-#9 := 1::real
-#8 := 0::real
-#10 := (<= 0::real 1::real)
-#11 := (not #10)
-#38 := (iff #11 false)
-#1 := true
-#33 := (not true)
-#36 := (iff #33 false)
-#37 := [rewrite]: #36
-#34 := (iff #11 #33)
-#31 := (iff #10 true)
-#32 := [rewrite]: #31
-#35 := [monotonicity #32]: #34
-#39 := [trans #35 #37]: #38
-#28 := [asserted]: #11
-[mp #28 #39]: false
-unsat
-1969eccfcb8fcfbee552d3423834a8c5c81b7a3e 75 0
+8aa9d7db6e905a29f4a4232a6a82f51c63e2641b 75 0
 #2 := false
 #53 := 0::int
 decl f4 :: int
@@ -46089,67 +46162,7 @@
 #89 := [not-or-elim #88]: #54
 [th-lemma #89 #91 #92]: false
 unsat
-88a6a03ef194abfd27e9e9c71cd5d14dc85a3e0c 40 0
-#2 := false
-#12 := 567::real
-#10 := 345::real
-#8 := 123::real
-#9 := (- 123::real)
-#11 := (+ #9 345::real)
-#13 := (< #11 567::real)
-#14 := (not #13)
-#58 := (iff #14 false)
-#38 := 222::real
-#43 := (< 222::real 567::real)
-#46 := (not #43)
-#56 := (iff #46 false)
-#1 := true
-#51 := (not true)
-#54 := (iff #51 false)
-#55 := [rewrite]: #54
-#52 := (iff #46 #51)
-#49 := (iff #43 true)
-#50 := [rewrite]: #49
-#53 := [monotonicity #50]: #52
-#57 := [trans #53 #55]: #56
-#47 := (iff #14 #46)
-#44 := (iff #13 #43)
-#41 := (= #11 222::real)
-#32 := -123::real
-#35 := (+ -123::real 345::real)
-#39 := (= #35 222::real)
-#40 := [rewrite]: #39
-#36 := (= #11 #35)
-#33 := (= #9 -123::real)
-#34 := [rewrite]: #33
-#37 := [monotonicity #34]: #36
-#42 := [trans #37 #40]: #41
-#45 := [monotonicity #42]: #44
-#48 := [monotonicity #45]: #47
-#59 := [trans #48 #57]: #58
-#31 := [asserted]: #14
-[mp #31 #59]: false
-unsat
-15b2b70e02e33b3fd62ad125d3d090b78e843ef0 18 0
-#2 := false
-#9 := 2345678901::real
-#8 := 123456789::real
-#10 := (< 123456789::real 2345678901::real)
-#11 := (not #10)
-#38 := (iff #11 false)
-#1 := true
-#33 := (not true)
-#36 := (iff #33 false)
-#37 := [rewrite]: #36
-#34 := (iff #11 #33)
-#31 := (iff #10 true)
-#32 := [rewrite]: #31
-#35 := [monotonicity #32]: #34
-#39 := [trans #35 #37]: #38
-#28 := [asserted]: #11
-[mp #28 #39]: false
-unsat
-1f807fc1d963d1157d011baf2bc15be87a543e47 80 0
+625ca444862ba21a1fa0a52470751952d4f5acee 80 0
 #2 := false
 #53 := 0::int
 decl f4 :: int
@@ -46230,38 +46243,7 @@
 #94 := [not-or-elim #93]: #54
 [th-lemma #94 #96 #97]: false
 unsat
-99d44ac5b1eabf28408c90ba22548007cb115846 30 0
-#2 := false
-#10 := 2345678901::real
-#8 := 123456789::real
-#9 := (- 123456789::real)
-#11 := (< #9 2345678901::real)
-#12 := (not #11)
-#48 := (iff #12 false)
-#30 := -123456789::real
-#33 := (< -123456789::real 2345678901::real)
-#36 := (not #33)
-#46 := (iff #36 false)
-#1 := true
-#41 := (not true)
-#44 := (iff #41 false)
-#45 := [rewrite]: #44
-#42 := (iff #36 #41)
-#39 := (iff #33 true)
-#40 := [rewrite]: #39
-#43 := [monotonicity #40]: #42
-#47 := [trans #43 #45]: #46
-#37 := (iff #12 #36)
-#34 := (iff #11 #33)
-#31 := (= #9 -123456789::real)
-#32 := [rewrite]: #31
-#35 := [monotonicity #32]: #34
-#38 := [monotonicity #35]: #37
-#49 := [trans #38 #47]: #48
-#29 := [asserted]: #12
-[mp #29 #49]: false
-unsat
-285e04b09127df030f3fb9fc56707e4ac5536d20 80 0
+59582f385e59c3a49a302bb1b9701c92bb368fac 80 0
 #2 := false
 #52 := 0::int
 decl f4 :: int
@@ -46342,61 +46324,7 @@
 #96 := [not-or-elim #93]: #95
 [th-lemma #96 #94 #97]: false
 unsat
-1e2efa99a6b6da8267e97fbabfbfc4ad6d75e298 26 0
-#2 := false
-decl f3 :: real
-#8 := f3
-#9 := 0::real
-#10 := (+ f3 0::real)
-#11 := (= #10 f3)
-#12 := (not #11)
-#43 := (iff #12 false)
-#1 := true
-#38 := (not true)
-#41 := (iff #38 false)
-#42 := [rewrite]: #41
-#39 := (iff #12 #38)
-#36 := (iff #11 true)
-#31 := (= f3 f3)
-#34 := (iff #31 true)
-#35 := [rewrite]: #34
-#32 := (iff #11 #31)
-#30 := [rewrite]: #11
-#33 := [monotonicity #30]: #32
-#37 := [trans #33 #35]: #36
-#40 := [monotonicity #37]: #39
-#44 := [trans #40 #42]: #43
-#29 := [asserted]: #12
-[mp #29 #44]: false
-unsat
-db8ff96061c803374c98cc69bff61f209e6704b2 26 0
-#2 := false
-decl f3 :: real
-#9 := f3
-#8 := 0::real
-#10 := (+ 0::real f3)
-#11 := (= #10 f3)
-#12 := (not #11)
-#43 := (iff #12 false)
-#1 := true
-#38 := (not true)
-#41 := (iff #38 false)
-#42 := [rewrite]: #41
-#39 := (iff #12 #38)
-#36 := (iff #11 true)
-#31 := (= f3 f3)
-#34 := (iff #31 true)
-#35 := [rewrite]: #34
-#32 := (iff #11 #31)
-#30 := [rewrite]: #11
-#33 := [monotonicity #30]: #32
-#37 := [trans #33 #35]: #36
-#40 := [monotonicity #37]: #39
-#44 := [trans #40 #42]: #43
-#29 := [asserted]: #12
-[mp #29 #44]: false
-unsat
-ef932f399da076ed00e03e4da73316a7d197a3fc 85 0
+6dfe89c45e68ba796109229b1c54d23603c64ded 85 0
 #2 := false
 #52 := 0::int
 decl f4 :: int
@@ -46482,37 +46410,7 @@
 #101 := [not-or-elim #99]: #66
 [th-lemma #101 #100 #102]: false
 unsat
-0e8698b45bf05888a03321a1f80727dbab79ce92 29 0
-#2 := false
-decl f3 :: real
-#8 := f3
-decl f4 :: real
-#9 := f4
-#11 := (+ f4 f3)
-#10 := (+ f3 f4)
-#12 := (= #10 #11)
-#13 := (not #12)
-#45 := (iff #13 false)
-#1 := true
-#40 := (not true)
-#43 := (iff #40 false)
-#44 := [rewrite]: #43
-#41 := (iff #13 #40)
-#38 := (iff #12 true)
-#33 := (= #10 #10)
-#36 := (iff #33 true)
-#37 := [rewrite]: #36
-#34 := (iff #12 #33)
-#31 := (= #11 #10)
-#32 := [rewrite]: #31
-#35 := [monotonicity #32]: #34
-#39 := [trans #35 #37]: #38
-#42 := [monotonicity #39]: #41
-#46 := [trans #42 #44]: #45
-#30 := [asserted]: #13
-[mp #30 #46]: false
-unsat
-fa0e4c03539b13ece5f7ecd99f1cc759989c1bcf 69 0
+650b9ceb8c2cf801e6ceee5375af886f013abe2c 69 0
 #2 := false
 #44 := 0::int
 decl f5 :: int
@@ -46582,7 +46480,317 @@
 #85 := [and-elim #83]: #53
 [th-lemma #85 #84 #86]: false
 unsat
-38433e3faa18ccf7e7f7256e4aa4dd42dcf43860 33 0
+0eeabd2aea1625cea2b9a9ade1eee621e6dfb534 17 0
+#2 := false
+#8 := 0::real
+#9 := (= 0::real 0::real)
+#10 := (not #9)
+#35 := (iff #10 false)
+#1 := true
+#30 := (not true)
+#33 := (iff #30 false)
+#34 := [rewrite]: #33
+#31 := (iff #10 #30)
+#28 := (iff #9 true)
+#29 := [rewrite]: #28
+#32 := [monotonicity #29]: #31
+#36 := [trans #32 #34]: #35
+#27 := [asserted]: #10
+[mp #27 #36]: false
+unsat
+582ef11d89a67dd029d6b29272ca6688d8312bd7 25 0
+#2 := false
+#8 := 0::real
+#9 := (- 0::real)
+#10 := (= 0::real #9)
+#11 := (not #10)
+#43 := (iff #11 false)
+#1 := true
+#38 := (not true)
+#41 := (iff #38 false)
+#42 := [rewrite]: #41
+#39 := (iff #11 #38)
+#36 := (iff #10 true)
+#31 := (= 0::real 0::real)
+#34 := (iff #31 true)
+#35 := [rewrite]: #34
+#32 := (iff #10 #31)
+#29 := (= #9 0::real)
+#30 := [rewrite]: #29
+#33 := [monotonicity #30]: #32
+#37 := [trans #33 #35]: #36
+#40 := [monotonicity #37]: #39
+#44 := [trans #40 #42]: #43
+#28 := [asserted]: #11
+[mp #28 #44]: false
+unsat
+4de73c1caf8d03027271f9dfec62b23debf65d9b 17 0
+#2 := false
+#8 := 1::real
+#9 := (= 1::real 1::real)
+#10 := (not #9)
+#35 := (iff #10 false)
+#1 := true
+#30 := (not true)
+#33 := (iff #30 false)
+#34 := [rewrite]: #33
+#31 := (iff #10 #30)
+#28 := (iff #9 true)
+#29 := [rewrite]: #28
+#32 := [monotonicity #29]: #31
+#36 := [trans #32 #34]: #35
+#27 := [asserted]: #10
+[mp #27 #36]: false
+unsat
+793c5962cbb9973917b8d8c66a2d2a3d94594803 34 0
+#2 := false
+#8 := 1::real
+#9 := (- 1::real)
+#10 := (= #9 1::real)
+#11 := (not #10)
+#12 := (not #11)
+#52 := (iff #12 false)
+#1 := true
+#47 := (not true)
+#50 := (iff #47 false)
+#51 := [rewrite]: #50
+#48 := (iff #12 #47)
+#45 := (iff #11 true)
+#40 := (not false)
+#43 := (iff #40 true)
+#44 := [rewrite]: #43
+#41 := (iff #11 #40)
+#38 := (iff #10 false)
+#30 := -1::real
+#33 := (= -1::real 1::real)
+#36 := (iff #33 false)
+#37 := [rewrite]: #36
+#34 := (iff #10 #33)
+#31 := (= #9 -1::real)
+#32 := [rewrite]: #31
+#35 := [monotonicity #32]: #34
+#39 := [trans #35 #37]: #38
+#42 := [monotonicity #39]: #41
+#46 := [trans #42 #44]: #45
+#49 := [monotonicity #46]: #48
+#53 := [trans #49 #51]: #52
+#29 := [asserted]: #12
+[mp #29 #53]: false
+unsat
+8f6bb26e23c49df219a9fc6920a8bbad56872cba 18 0
+#2 := false
+#9 := 1::real
+#8 := 0::real
+#10 := (< 0::real 1::real)
+#11 := (not #10)
+#38 := (iff #11 false)
+#1 := true
+#33 := (not true)
+#36 := (iff #33 false)
+#37 := [rewrite]: #36
+#34 := (iff #11 #33)
+#31 := (iff #10 true)
+#32 := [rewrite]: #31
+#35 := [monotonicity #32]: #34
+#39 := [trans #35 #37]: #38
+#28 := [asserted]: #11
+[mp #28 #39]: false
+unsat
+6edbad6a7b2935c8d856f03c54b56bf43891a380 18 0
+#2 := false
+#9 := 1::real
+#8 := 0::real
+#10 := (<= 0::real 1::real)
+#11 := (not #10)
+#38 := (iff #11 false)
+#1 := true
+#33 := (not true)
+#36 := (iff #33 false)
+#37 := [rewrite]: #36
+#34 := (iff #11 #33)
+#31 := (iff #10 true)
+#32 := [rewrite]: #31
+#35 := [monotonicity #32]: #34
+#39 := [trans #35 #37]: #38
+#28 := [asserted]: #11
+[mp #28 #39]: false
+unsat
+a64bea89744a3af0ed560a5c871536de3fcb2ecf 40 0
+#2 := false
+#12 := 567::real
+#10 := 345::real
+#8 := 123::real
+#9 := (- 123::real)
+#11 := (+ #9 345::real)
+#13 := (< #11 567::real)
+#14 := (not #13)
+#58 := (iff #14 false)
+#38 := 222::real
+#43 := (< 222::real 567::real)
+#46 := (not #43)
+#56 := (iff #46 false)
+#1 := true
+#51 := (not true)
+#54 := (iff #51 false)
+#55 := [rewrite]: #54
+#52 := (iff #46 #51)
+#49 := (iff #43 true)
+#50 := [rewrite]: #49
+#53 := [monotonicity #50]: #52
+#57 := [trans #53 #55]: #56
+#47 := (iff #14 #46)
+#44 := (iff #13 #43)
+#41 := (= #11 222::real)
+#32 := -123::real
+#35 := (+ -123::real 345::real)
+#39 := (= #35 222::real)
+#40 := [rewrite]: #39
+#36 := (= #11 #35)
+#33 := (= #9 -123::real)
+#34 := [rewrite]: #33
+#37 := [monotonicity #34]: #36
+#42 := [trans #37 #40]: #41
+#45 := [monotonicity #42]: #44
+#48 := [monotonicity #45]: #47
+#59 := [trans #48 #57]: #58
+#31 := [asserted]: #14
+[mp #31 #59]: false
+unsat
+a79ee55daec933b360adb88ed0db37156e2f7821 18 0
+#2 := false
+#9 := 2345678901::real
+#8 := 123456789::real
+#10 := (< 123456789::real 2345678901::real)
+#11 := (not #10)
+#38 := (iff #11 false)
+#1 := true
+#33 := (not true)
+#36 := (iff #33 false)
+#37 := [rewrite]: #36
+#34 := (iff #11 #33)
+#31 := (iff #10 true)
+#32 := [rewrite]: #31
+#35 := [monotonicity #32]: #34
+#39 := [trans #35 #37]: #38
+#28 := [asserted]: #11
+[mp #28 #39]: false
+unsat
+4b51b4eed1dc4e4e6099fbf906b70822230d5e3f 30 0
+#2 := false
+#10 := 2345678901::real
+#8 := 123456789::real
+#9 := (- 123456789::real)
+#11 := (< #9 2345678901::real)
+#12 := (not #11)
+#48 := (iff #12 false)
+#30 := -123456789::real
+#33 := (< -123456789::real 2345678901::real)
+#36 := (not #33)
+#46 := (iff #36 false)
+#1 := true
+#41 := (not true)
+#44 := (iff #41 false)
+#45 := [rewrite]: #44
+#42 := (iff #36 #41)
+#39 := (iff #33 true)
+#40 := [rewrite]: #39
+#43 := [monotonicity #40]: #42
+#47 := [trans #43 #45]: #46
+#37 := (iff #12 #36)
+#34 := (iff #11 #33)
+#31 := (= #9 -123456789::real)
+#32 := [rewrite]: #31
+#35 := [monotonicity #32]: #34
+#38 := [monotonicity #35]: #37
+#49 := [trans #38 #47]: #48
+#29 := [asserted]: #12
+[mp #29 #49]: false
+unsat
+209cf5d4457acae6abf17187fd6e89c259187362 26 0
+#2 := false
+decl f3 :: real
+#8 := f3
+#9 := 0::real
+#10 := (+ f3 0::real)
+#11 := (= #10 f3)
+#12 := (not #11)
+#43 := (iff #12 false)
+#1 := true
+#38 := (not true)
+#41 := (iff #38 false)
+#42 := [rewrite]: #41
+#39 := (iff #12 #38)
+#36 := (iff #11 true)
+#31 := (= f3 f3)
+#34 := (iff #31 true)
+#35 := [rewrite]: #34
+#32 := (iff #11 #31)
+#30 := [rewrite]: #11
+#33 := [monotonicity #30]: #32
+#37 := [trans #33 #35]: #36
+#40 := [monotonicity #37]: #39
+#44 := [trans #40 #42]: #43
+#29 := [asserted]: #12
+[mp #29 #44]: false
+unsat
+97dafa0bba507529233c05cc33dd1519b4e28e4b 26 0
+#2 := false
+decl f3 :: real
+#9 := f3
+#8 := 0::real
+#10 := (+ 0::real f3)
+#11 := (= #10 f3)
+#12 := (not #11)
+#43 := (iff #12 false)
+#1 := true
+#38 := (not true)
+#41 := (iff #38 false)
+#42 := [rewrite]: #41
+#39 := (iff #12 #38)
+#36 := (iff #11 true)
+#31 := (= f3 f3)
+#34 := (iff #31 true)
+#35 := [rewrite]: #34
+#32 := (iff #11 #31)
+#30 := [rewrite]: #11
+#33 := [monotonicity #30]: #32
+#37 := [trans #33 #35]: #36
+#40 := [monotonicity #37]: #39
+#44 := [trans #40 #42]: #43
+#29 := [asserted]: #12
+[mp #29 #44]: false
+unsat
+393f5dbb0db04527f22218d40363df111099baf5 29 0
+#2 := false
+decl f3 :: real
+#8 := f3
+decl f4 :: real
+#9 := f4
+#11 := (+ f4 f3)
+#10 := (+ f3 f4)
+#12 := (= #10 #11)
+#13 := (not #12)
+#45 := (iff #13 false)
+#1 := true
+#40 := (not true)
+#43 := (iff #40 false)
+#44 := [rewrite]: #43
+#41 := (iff #13 #40)
+#38 := (iff #12 true)
+#33 := (= #10 #10)
+#36 := (iff #33 true)
+#37 := [rewrite]: #36
+#34 := (iff #12 #33)
+#31 := (= #11 #10)
+#32 := [rewrite]: #31
+#35 := [monotonicity #32]: #34
+#39 := [trans #35 #37]: #38
+#42 := [monotonicity #39]: #41
+#46 := [trans #42 #44]: #45
+#30 := [asserted]: #13
+[mp #30 #46]: false
+unsat
+08018537eda07e4214b99470b7aa75706e7cd28c 33 0
 #2 := false
 decl f5 :: real
 #10 := f5
@@ -46616,7 +46824,7 @@
 #33 := [asserted]: #16
 [mp #33 #49]: false
 unsat
-1fdb35a1399271a8f0f610e1f353642ad440e0a5 45 0
+afbfb539cd215cee99e6f9076de9621f41db722f 45 0
 #2 := false
 decl f4 :: real
 #9 := f4
@@ -46662,7 +46870,7 @@
 #33 := [asserted]: #16
 [mp #33 #63]: false
 unsat
-1dd629fda226d261e70599846e7cb0aed5bce0e4 58 0
+4a497c9573e0d4cf02c29e65ea385ac694a5d2a9 58 0
 #2 := false
 decl f3 :: real
 #8 := f3
@@ -46721,34 +46929,7 @@
 #31 := [asserted]: #14
 [mp #31 #77]: false
 unsat
-f05f2d2d486f345aec8eda31ab857db74ae27876 26 0
-#2 := false
-decl f3 :: real
-#8 := f3
-#9 := 0::real
-#10 := (- f3 0::real)
-#11 := (= #10 f3)
-#12 := (not #11)
-#43 := (iff #12 false)
-#1 := true
-#38 := (not true)
-#41 := (iff #38 false)
-#42 := [rewrite]: #41
-#39 := (iff #12 #38)
-#36 := (iff #11 true)
-#31 := (= f3 f3)
-#34 := (iff #31 true)
-#35 := [rewrite]: #34
-#32 := (iff #11 #31)
-#30 := [rewrite]: #11
-#33 := [monotonicity #30]: #32
-#37 := [trans #33 #35]: #36
-#40 := [monotonicity #37]: #39
-#44 := [trans #40 #42]: #43
-#29 := [asserted]: #12
-[mp #29 #44]: false
-unsat
-56d5d1ebf2a3e86b0eef89fb438ab0199a14cb1f 61 0
+5310d952762e8bd0d9db1c0127d9100532f42cf4 61 0
 #2 := false
 #8 := 0::real
 decl f3 :: real
@@ -46810,7 +46991,96 @@
 #31 := [asserted]: #14
 [mp #31 #79]: false
 unsat
-790f41c10a6a91834ac605d1e5a736ba38cac6ca 32 0
+aed50da01a066fa4f6c8e7af4a461d101bc09de0 61 0
+#2 := false
+decl f3 :: real
+#8 := f3
+#11 := (- f3)
+#9 := 0::real
+#12 := (< 0::real #11)
+#10 := (< f3 0::real)
+#13 := (implies #10 #12)
+#14 := (not #13)
+#78 := (iff #14 false)
+#32 := -1::real
+#33 := (* -1::real f3)
+#36 := (< 0::real #33)
+#42 := (not #10)
+#43 := (or #42 #36)
+#48 := (not #43)
+#76 := (iff #48 false)
+#1 := true
+#71 := (not true)
+#74 := (iff #71 false)
+#75 := [rewrite]: #74
+#72 := (iff #48 #71)
+#69 := (iff #43 true)
+#51 := (>= f3 0::real)
+#52 := (not #51)
+#64 := (or #51 #52)
+#67 := (iff #64 true)
+#68 := [rewrite]: #67
+#65 := (iff #43 #64)
+#62 := (iff #36 #52)
+#63 := [rewrite]: #62
+#60 := (iff #42 #51)
+#55 := (not #52)
+#58 := (iff #55 #51)
+#59 := [rewrite]: #58
+#56 := (iff #42 #55)
+#53 := (iff #10 #52)
+#54 := [rewrite]: #53
+#57 := [monotonicity #54]: #56
+#61 := [trans #57 #59]: #60
+#66 := [monotonicity #61 #63]: #65
+#70 := [trans #66 #68]: #69
+#73 := [monotonicity #70]: #72
+#77 := [trans #73 #75]: #76
+#49 := (iff #14 #48)
+#46 := (iff #13 #43)
+#39 := (implies #10 #36)
+#44 := (iff #39 #43)
+#45 := [rewrite]: #44
+#40 := (iff #13 #39)
+#37 := (iff #12 #36)
+#34 := (= #11 #33)
+#35 := [rewrite]: #34
+#38 := [monotonicity #35]: #37
+#41 := [monotonicity #38]: #40
+#47 := [trans #41 #45]: #46
+#50 := [monotonicity #47]: #49
+#79 := [trans #50 #77]: #78
+#31 := [asserted]: #14
+[mp #31 #79]: false
+unsat
+4c33a6158d9a09229cc4bfd8f383004f3cebc48a 26 0
+#2 := false
+decl f3 :: real
+#8 := f3
+#9 := 0::real
+#10 := (- f3 0::real)
+#11 := (= #10 f3)
+#12 := (not #11)
+#43 := (iff #12 false)
+#1 := true
+#38 := (not true)
+#41 := (iff #38 false)
+#42 := [rewrite]: #41
+#39 := (iff #12 #38)
+#36 := (iff #11 true)
+#31 := (= f3 f3)
+#34 := (iff #31 true)
+#35 := [rewrite]: #34
+#32 := (iff #11 #31)
+#30 := [rewrite]: #11
+#33 := [monotonicity #30]: #32
+#37 := [trans #33 #35]: #36
+#40 := [monotonicity #37]: #39
+#44 := [trans #40 #42]: #43
+#29 := [asserted]: #12
+[mp #29 #44]: false
+unsat
+5fa448e93a50f684d9003c71f7eb6c60b8908f77 32 0
 #2 := false
 decl f3 :: real
 #9 := f3
@@ -46843,7 +47113,7 @@
 #30 := [asserted]: #13
 [mp #30 #50]: false
 unsat
-45b60930fea61c4a8f7b103e68e903e1ceab7852 64 0
+abc85cf86b0b1c8a34cd74d5da49e570c77ddea0 64 0
 #2 := false
 #12 := 0::real
 decl f4 :: real
@@ -46908,69 +47178,7 @@
 #32 := [asserted]: #15
 [mp #32 #81]: false
 unsat
-60ae5a86685e4f3dd9f231aaafe937bfd9ceb4bc 61 0
-#2 := false
-decl f3 :: real
-#8 := f3
-#11 := (- f3)
-#9 := 0::real
-#12 := (< 0::real #11)
-#10 := (< f3 0::real)
-#13 := (implies #10 #12)
-#14 := (not #13)
-#78 := (iff #14 false)
-#32 := -1::real
-#33 := (* -1::real f3)
-#36 := (< 0::real #33)
-#42 := (not #10)
-#43 := (or #42 #36)
-#48 := (not #43)
-#76 := (iff #48 false)
-#1 := true
-#71 := (not true)
-#74 := (iff #71 false)
-#75 := [rewrite]: #74
-#72 := (iff #48 #71)
-#69 := (iff #43 true)
-#51 := (>= f3 0::real)
-#52 := (not #51)
-#64 := (or #51 #52)
-#67 := (iff #64 true)
-#68 := [rewrite]: #67
-#65 := (iff #43 #64)
-#62 := (iff #36 #52)
-#63 := [rewrite]: #62
-#60 := (iff #42 #51)
-#55 := (not #52)
-#58 := (iff #55 #51)
-#59 := [rewrite]: #58
-#56 := (iff #42 #55)
-#53 := (iff #10 #52)
-#54 := [rewrite]: #53
-#57 := [monotonicity #54]: #56
-#61 := [trans #57 #59]: #60
-#66 := [monotonicity #61 #63]: #65
-#70 := [trans #66 #68]: #69
-#73 := [monotonicity #70]: #72
-#77 := [trans #73 #75]: #76
-#49 := (iff #14 #48)
-#46 := (iff #13 #43)
-#39 := (implies #10 #36)
-#44 := (iff #39 #43)
-#45 := [rewrite]: #44
-#40 := (iff #13 #39)
-#37 := (iff #12 #36)
-#34 := (= #11 #33)
-#35 := [rewrite]: #34
-#38 := [monotonicity #35]: #37
-#41 := [monotonicity #38]: #40
-#47 := [trans #41 #45]: #46
-#50 := [monotonicity #47]: #49
-#79 := [trans #50 #77]: #78
-#31 := [asserted]: #14
-[mp #31 #79]: false
-unsat
-390f065bbd860ad0bb40feb02d88be98544fabc4 44 0
+8084abd03f885814b6c145eb3fdc41f5a2cc85f8 44 0
 #2 := false
 decl f3 :: real
 #8 := f3
@@ -47015,7 +47223,7 @@
 #31 := [asserted]: #14
 [mp #31 #61]: false
 unsat
-ff168f63f9897fde96bfb7de5db7da56a8ec5202 42 0
+9afe5717b3e2c9027579ba3e8d9d66db44c6cbda 42 0
 #2 := false
 decl f3 :: real
 #8 := f3
@@ -47058,7 +47266,7 @@
 #31 := [asserted]: #14
 [mp #31 #59]: false
 unsat
-fbdc4cb05a56d63893610ccb3b08ec35be0de896 48 0
+a1710b56b11072c592931adbf944a5bd3bc6efc2 48 0
 #2 := false
 decl f5 :: real
 #11 := f5
@@ -47107,7 +47315,7 @@
 #33 := [asserted]: #16
 [mp #33 #64]: false
 unsat
-35cb5c6f18d527568affbd6fa3b8106b1a97d4c5 30 0
+4611c2bccf9693379d8f36635022f4216eaaa919 30 0
 #2 := false
 #8 := 1::real
 #9 := 2::real
@@ -47138,44 +47346,7 @@
 #29 := [asserted]: #12
 [mp #29 #49]: false
 unsat
-a66448e31ec3061bc2656a02b5d410722684294d 36 0
-#2 := false
-decl f3 :: real
-#9 := f3
-#8 := 3::real
-#10 := (/ f3 3::real)
-#11 := (* 3::real #10)
-#12 := (= #11 f3)
-#13 := (not #12)
-#53 := (iff #13 false)
-#1 := true
-#48 := (not true)
-#51 := (iff #48 false)
-#52 := [rewrite]: #51
-#49 := (iff #13 #48)
-#46 := (iff #12 true)
-#41 := (= f3 f3)
-#44 := (iff #41 true)
-#45 := [rewrite]: #44
-#42 := (iff #12 #41)
-#31 := 1/3::real
-#32 := (* 1/3::real f3)
-#35 := (* 3::real #32)
-#38 := (= #35 f3)
-#39 := [rewrite]: #38
-#36 := (= #11 #35)
-#33 := (= #10 #32)
-#34 := [rewrite]: #33
-#37 := [monotonicity #34]: #36
-#40 := [trans #37 #39]: #12
-#43 := [monotonicity #40]: #42
-#47 := [trans #43 #45]: #46
-#50 := [monotonicity #47]: #49
-#54 := [trans #50 #52]: #53
-#30 := [asserted]: #13
-[mp #30 #54]: false
-unsat
-3037834a6c55cfbe96598ac3d4b462835eeaed7d 27 0
+300cf7e266e08d39c05d69e60b19d0ac09a08a2e 27 0
 #2 := false
 #9 := 3::real
 #8 := 1::real
@@ -47203,43 +47374,7 @@
 #29 := [asserted]: #12
 [mp #29 #46]: false
 unsat
-0c0a26d5b20997d02798707aedc1fa07591ae1bf 35 0
-#2 := false
-decl f3 :: real
-#8 := f3
-#9 := 3::real
-#10 := (* f3 3::real)
-#11 := (/ #10 3::real)
-#12 := (= #11 f3)
-#13 := (not #12)
-#52 := (iff #13 false)
-#1 := true
-#47 := (not true)
-#50 := (iff #47 false)
-#51 := [rewrite]: #50
-#48 := (iff #13 #47)
-#45 := (iff #12 true)
-#40 := (= f3 f3)
-#43 := (iff #40 true)
-#44 := [rewrite]: #43
-#41 := (iff #12 #40)
-#31 := (* 3::real f3)
-#34 := (/ #31 3::real)
-#37 := (= #34 f3)
-#38 := [rewrite]: #37
-#35 := (= #11 #34)
-#32 := (= #10 #31)
-#33 := [rewrite]: #32
-#36 := [monotonicity #33]: #35
-#39 := [trans #36 #38]: #12
-#42 := [monotonicity #39]: #41
-#46 := [trans #42 #44]: #45
-#49 := [monotonicity #46]: #48
-#53 := [trans #49 #51]: #52
-#30 := [asserted]: #13
-[mp #30 #53]: false
-unsat
-c0ca03e227f816842315b007cba7acf006e07203 48 0
+b59d35169428d1cd490a83e17435ecfc93225870 48 0
 #2 := false
 #9 := 3::real
 #8 := 1::real
@@ -47288,7 +47423,310 @@
 #32 := [asserted]: #15
 [mp #32 #67]: false
 unsat
-167b73607533e079718467537de3ca170ac6b4fc 62 0
+1c222157f628411945e7e47ef57130c5ba4ad155 36 0
+#2 := false
+#10 := 3::real
+#8 := 1::real
+#9 := (- 1::real)
+#11 := (/ #9 3::real)
+#12 := (= #11 #11)
+#13 := (not #12)
+#54 := (iff #13 false)
+#1 := true
+#49 := (not true)
+#52 := (iff #49 false)
+#53 := [rewrite]: #52
+#50 := (iff #13 #49)
+#47 := (iff #12 true)
+#37 := -1/3::real
+#42 := (= -1/3::real -1/3::real)
+#45 := (iff #42 true)
+#46 := [rewrite]: #45
+#43 := (iff #12 #42)
+#40 := (= #11 -1/3::real)
+#31 := -1::real
+#34 := (/ -1::real 3::real)
+#38 := (= #34 -1/3::real)
+#39 := [rewrite]: #38
+#35 := (= #11 #34)
+#32 := (= #9 -1::real)
+#33 := [rewrite]: #32
+#36 := [monotonicity #33]: #35
+#41 := [trans #36 #39]: #40
+#44 := [monotonicity #41 #41]: #43
+#48 := [trans #44 #46]: #47
+#51 := [monotonicity #48]: #50
+#55 := [trans #51 #53]: #54
+#30 := [asserted]: #13
+[mp #30 #55]: false
+unsat
+e4a7188c884c4cc69972b6b4c9f7bcbb1cd3ef93 43 0
+#2 := false
+#10 := 3::real
+#8 := 1::real
+#13 := (/ 1::real 3::real)
+#11 := (- 3::real)
+#9 := (- 1::real)
+#12 := (/ #9 #11)
+#14 := (= #12 #13)
+#15 := (not #14)
+#61 := (iff #15 false)
+#1 := true
+#56 := (not true)
+#59 := (iff #56 false)
+#60 := [rewrite]: #59
+#57 := (iff #15 #56)
+#54 := (iff #14 true)
+#42 := 1/3::real
+#49 := (= 1/3::real 1/3::real)
+#52 := (iff #49 true)
+#53 := [rewrite]: #52
+#50 := (iff #14 #49)
+#47 := (= #13 1/3::real)
+#48 := [rewrite]: #47
+#45 := (= #12 1/3::real)
+#36 := -3::real
+#33 := -1::real
+#39 := (/ -1::real -3::real)
+#43 := (= #39 1/3::real)
+#44 := [rewrite]: #43
+#40 := (= #12 #39)
+#37 := (= #11 -3::real)
+#38 := [rewrite]: #37
+#34 := (= #9 -1::real)
+#35 := [rewrite]: #34
+#41 := [monotonicity #35 #38]: #40
+#46 := [trans #41 #44]: #45
+#51 := [monotonicity #46 #48]: #50
+#55 := [trans #51 #53]: #54
+#58 := [monotonicity #55]: #57
+#62 := [trans #58 #60]: #61
+#32 := [asserted]: #15
+[mp #32 #62]: false
+unsat
+b7bd87b08d1cd38b7e0db00083a20eb0fa944982 26 0
+#2 := false
+decl f3 :: real
+#8 := f3
+#9 := 1::real
+#10 := (/ f3 1::real)
+#11 := (= #10 f3)
+#12 := (not #11)
+#43 := (iff #12 false)
+#1 := true
+#38 := (not true)
+#41 := (iff #38 false)
+#42 := [rewrite]: #41
+#39 := (iff #12 #38)
+#36 := (iff #11 true)
+#31 := (= f3 f3)
+#34 := (iff #31 true)
+#35 := [rewrite]: #34
+#32 := (iff #11 #31)
+#30 := [rewrite]: #11
+#33 := [monotonicity #30]: #32
+#37 := [trans #33 #35]: #36
+#40 := [monotonicity #37]: #39
+#44 := [trans #40 #42]: #43
+#29 := [asserted]: #12
+[mp #29 #44]: false
+unsat
+c3e8ed15657d66fdc6a77370b7e3994d0741653a 60 0
+#2 := false
+#8 := 0::real
+decl f3 :: real
+#9 := f3
+#65 := 2/3::real
+#66 := (* 2/3::real f3)
+#67 := (<= #66 0::real)
+#63 := (not #67)
+#52 := (<= f3 0::real)
+#69 := (or #52 #63)
+#72 := (not #69)
+#11 := 3::real
+#12 := (/ f3 3::real)
+#13 := (< #12 f3)
+#10 := (< 0::real f3)
+#14 := (implies #10 #13)
+#15 := (not #14)
+#75 := (iff #15 #72)
+#33 := 1/3::real
+#34 := (* 1/3::real f3)
+#37 := (< #34 f3)
+#43 := (not #10)
+#44 := (or #43 #37)
+#49 := (not #44)
+#73 := (iff #49 #72)
+#70 := (iff #44 #69)
+#64 := (iff #37 #63)
+#68 := [rewrite]: #64
+#61 := (iff #43 #52)
+#53 := (not #52)
+#56 := (not #53)
+#59 := (iff #56 #52)
+#60 := [rewrite]: #59
+#57 := (iff #43 #56)
+#54 := (iff #10 #53)
+#55 := [rewrite]: #54
+#58 := [monotonicity #55]: #57
+#62 := [trans #58 #60]: #61
+#71 := [monotonicity #62 #68]: #70
+#74 := [monotonicity #71]: #73
+#50 := (iff #15 #49)
+#47 := (iff #14 #44)
+#40 := (implies #10 #37)
+#45 := (iff #40 #44)
+#46 := [rewrite]: #45
+#41 := (iff #14 #40)
+#38 := (iff #13 #37)
+#35 := (= #12 #34)
+#36 := [rewrite]: #35
+#39 := [monotonicity #36]: #38
+#42 := [monotonicity #39]: #41
+#48 := [trans #42 #46]: #47
+#51 := [monotonicity #48]: #50
+#76 := [trans #51 #74]: #75
+#32 := [asserted]: #15
+#77 := [mp #32 #76]: #72
+#79 := [not-or-elim #77]: #67
+#78 := [not-or-elim #77]: #53
+[th-lemma #78 #79]: false
+unsat
+05dd4ddb8672ca6d3f31acc3ffa7f7986078b232 60 0
+#2 := false
+#9 := 0::real
+decl f3 :: real
+#8 := f3
+#67 := 2/3::real
+#68 := (* 2/3::real f3)
+#65 := (>= #68 0::real)
+#63 := (not #65)
+#54 := (>= f3 0::real)
+#69 := (or #54 #63)
+#72 := (not #69)
+#11 := 3::real
+#12 := (/ f3 3::real)
+#13 := (< f3 #12)
+#10 := (< f3 0::real)
+#14 := (implies #10 #13)
+#15 := (not #14)
+#75 := (iff #15 #72)
+#33 := 1/3::real
+#34 := (* 1/3::real f3)
+#37 := (< f3 #34)
+#43 := (not #10)
+#44 := (or #43 #37)
+#49 := (not #44)
+#73 := (iff #49 #72)
+#70 := (iff #44 #69)
+#64 := (iff #37 #63)
+#66 := [rewrite]: #64
+#61 := (iff #43 #54)
+#52 := (not #54)
+#56 := (not #52)
+#59 := (iff #56 #54)
+#60 := [rewrite]: #59
+#57 := (iff #43 #56)
+#53 := (iff #10 #52)
+#55 := [rewrite]: #53
+#58 := [monotonicity #55]: #57
+#62 := [trans #58 #60]: #61
+#71 := [monotonicity #62 #66]: #70
+#74 := [monotonicity #71]: #73
+#50 := (iff #15 #49)
+#47 := (iff #14 #44)
+#40 := (implies #10 #37)
+#45 := (iff #40 #44)
+#46 := [rewrite]: #45
+#41 := (iff #14 #40)
+#38 := (iff #13 #37)
+#35 := (= #12 #34)
+#36 := [rewrite]: #35
+#39 := [monotonicity #36]: #38
+#42 := [monotonicity #39]: #41
+#48 := [trans #42 #46]: #47
+#51 := [monotonicity #48]: #50
+#76 := [trans #51 #74]: #75
+#32 := [asserted]: #15
+#77 := [mp #32 #76]: #72
+#79 := [not-or-elim #77]: #65
+#78 := [not-or-elim #77]: #52
+[th-lemma #78 #79]: false
+unsat
+b7a066feb9f43bcc3b0373cd9c9287361b54eff6 36 0
+#2 := false
+decl f3 :: real
+#9 := f3
+#8 := 3::real
+#10 := (/ f3 3::real)
+#11 := (* 3::real #10)
+#12 := (= #11 f3)
+#13 := (not #12)
+#53 := (iff #13 false)
+#1 := true
+#48 := (not true)
+#51 := (iff #48 false)
+#52 := [rewrite]: #51
+#49 := (iff #13 #48)
+#46 := (iff #12 true)
+#41 := (= f3 f3)
+#44 := (iff #41 true)
+#45 := [rewrite]: #44
+#42 := (iff #12 #41)
+#31 := 1/3::real
+#32 := (* 1/3::real f3)
+#35 := (* 3::real #32)
+#38 := (= #35 f3)
+#39 := [rewrite]: #38
+#36 := (= #11 #35)
+#33 := (= #10 #32)
+#34 := [rewrite]: #33
+#37 := [monotonicity #34]: #36
+#40 := [trans #37 #39]: #12
+#43 := [monotonicity #40]: #42
+#47 := [trans #43 #45]: #46
+#50 := [monotonicity #47]: #49
+#54 := [trans #50 #52]: #53
+#30 := [asserted]: #13
+[mp #30 #54]: false
+unsat
+60b8185463f5b2ce142483f90c095e9accf8989c 35 0
+#2 := false
+decl f3 :: real
+#8 := f3
+#9 := 3::real
+#10 := (* f3 3::real)
+#11 := (/ #10 3::real)
+#12 := (= #11 f3)
+#13 := (not #12)
+#52 := (iff #13 false)
+#1 := true
+#47 := (not true)
+#50 := (iff #47 false)
+#51 := [rewrite]: #50
+#48 := (iff #13 #47)
+#45 := (iff #12 true)
+#40 := (= f3 f3)
+#43 := (iff #40 true)
+#44 := [rewrite]: #43
+#41 := (iff #12 #40)
+#31 := (* 3::real f3)
+#34 := (/ #31 3::real)
+#37 := (= #34 f3)
+#38 := [rewrite]: #37
+#35 := (= #11 #34)
+#32 := (= #10 #31)
+#33 := [rewrite]: #32
+#36 := [monotonicity #33]: #35
+#39 := [trans #36 #38]: #12
+#42 := [monotonicity #39]: #41
+#46 := [trans #42 #44]: #45
+#49 := [monotonicity #46]: #48
+#53 := [trans #49 #51]: #52
+#30 := [asserted]: #13
+[mp #30 #53]: false
+unsat
+8140105cc276b32acd77e4f2cd380fea1f24efd8 62 0
 #2 := false
 #8 := 0::real
 decl f3 :: real
@@ -47351,88 +47789,7 @@
 #80 := [not-or-elim #79]: #55
 [th-lemma #80 #81]: false
 unsat
-407d897cdd96a77ccd0a5e9af3ea8310eebe935c 36 0
-#2 := false
-#10 := 3::real
-#8 := 1::real
-#9 := (- 1::real)
-#11 := (/ #9 3::real)
-#12 := (= #11 #11)
-#13 := (not #12)
-#54 := (iff #13 false)
-#1 := true
-#49 := (not true)
-#52 := (iff #49 false)
-#53 := [rewrite]: #52
-#50 := (iff #13 #49)
-#47 := (iff #12 true)
-#37 := -1/3::real
-#42 := (= -1/3::real -1/3::real)
-#45 := (iff #42 true)
-#46 := [rewrite]: #45
-#43 := (iff #12 #42)
-#40 := (= #11 -1/3::real)
-#31 := -1::real
-#34 := (/ -1::real 3::real)
-#38 := (= #34 -1/3::real)
-#39 := [rewrite]: #38
-#35 := (= #11 #34)
-#32 := (= #9 -1::real)
-#33 := [rewrite]: #32
-#36 := [monotonicity #33]: #35
-#41 := [trans #36 #39]: #40
-#44 := [monotonicity #41 #41]: #43
-#48 := [trans #44 #46]: #47
-#51 := [monotonicity #48]: #50
-#55 := [trans #51 #53]: #54
-#30 := [asserted]: #13
-[mp #30 #55]: false
-unsat
-dae8653f4906be077917c7741a65becd7fc6baa9 43 0
-#2 := false
-#10 := 3::real
-#8 := 1::real
-#13 := (/ 1::real 3::real)
-#11 := (- 3::real)
-#9 := (- 1::real)
-#12 := (/ #9 #11)
-#14 := (= #12 #13)
-#15 := (not #14)
-#61 := (iff #15 false)
-#1 := true
-#56 := (not true)
-#59 := (iff #56 false)
-#60 := [rewrite]: #59
-#57 := (iff #15 #56)
-#54 := (iff #14 true)
-#42 := 1/3::real
-#49 := (= 1/3::real 1/3::real)
-#52 := (iff #49 true)
-#53 := [rewrite]: #52
-#50 := (iff #14 #49)
-#47 := (= #13 1/3::real)
-#48 := [rewrite]: #47
-#45 := (= #12 1/3::real)
-#36 := -3::real
-#33 := -1::real
-#39 := (/ -1::real -3::real)
-#43 := (= #39 1/3::real)
-#44 := [rewrite]: #43
-#40 := (= #12 #39)
-#37 := (= #11 -3::real)
-#38 := [rewrite]: #37
-#34 := (= #9 -1::real)
-#35 := [rewrite]: #34
-#41 := [monotonicity #35 #38]: #40
-#46 := [trans #41 #44]: #45
-#51 := [monotonicity #46 #48]: #50
-#55 := [trans #51 #53]: #54
-#58 := [monotonicity #55]: #57
-#62 := [trans #58 #60]: #61
-#32 := [asserted]: #15
-[mp #32 #62]: false
-unsat
-5f5bc75d0e15785a03976c15c11ed28c9527a9c1 62 0
+9a571f31a85693be816acc5c7b7aa7362cb820bb 62 0
 #2 := false
 #9 := 0::real
 decl f3 :: real
@@ -47495,34 +47852,7 @@
 #80 := [not-or-elim #79]: #54
 [th-lemma #80 #81]: false
 unsat
-760d3be0f602d4f12e9c082b34c9607286d80e81 26 0
-#2 := false
-decl f3 :: real
-#8 := f3
-#9 := 1::real
-#10 := (/ f3 1::real)
-#11 := (= #10 f3)
-#12 := (not #11)
-#43 := (iff #12 false)
-#1 := true
-#38 := (not true)
-#41 := (iff #38 false)
-#42 := [rewrite]: #41
-#39 := (iff #12 #38)
-#36 := (iff #11 true)
-#31 := (= f3 f3)
-#34 := (iff #31 true)
-#35 := [rewrite]: #34
-#32 := (iff #11 #31)
-#30 := [rewrite]: #11
-#33 := [monotonicity #30]: #32
-#37 := [trans #33 #35]: #36
-#40 := [monotonicity #37]: #39
-#44 := [trans #40 #42]: #43
-#29 := [asserted]: #12
-[mp #29 #44]: false
-unsat
-490baa52b7aa34ed9cb90263f1ecd055bf5a7ff6 75 0
+2f685a966903afb267e7499dd8ad6b2b43caca93 75 0
 #2 := false
 #8 := 0::real
 decl f3 :: real
@@ -47598,68 +47928,7 @@
 #96 := [unit-resolution #95 #92]: #90
 [th-lemma #89 #70 #96]: false
 unsat
-9c25469b6a76ae1f1d1ef57b133f6ba4622084dd 60 0
-#2 := false
-#8 := 0::real
-decl f3 :: real
-#9 := f3
-#65 := 2/3::real
-#66 := (* 2/3::real f3)
-#67 := (<= #66 0::real)
-#63 := (not #67)
-#52 := (<= f3 0::real)
-#69 := (or #52 #63)
-#72 := (not #69)
-#11 := 3::real
-#12 := (/ f3 3::real)
-#13 := (< #12 f3)
-#10 := (< 0::real f3)
-#14 := (implies #10 #13)
-#15 := (not #14)
-#75 := (iff #15 #72)
-#33 := 1/3::real
-#34 := (* 1/3::real f3)
-#37 := (< #34 f3)
-#43 := (not #10)
-#44 := (or #43 #37)
-#49 := (not #44)
-#73 := (iff #49 #72)
-#70 := (iff #44 #69)
-#64 := (iff #37 #63)
-#68 := [rewrite]: #64
-#61 := (iff #43 #52)
-#53 := (not #52)
-#56 := (not #53)
-#59 := (iff #56 #52)
-#60 := [rewrite]: #59
-#57 := (iff #43 #56)
-#54 := (iff #10 #53)
-#55 := [rewrite]: #54
-#58 := [monotonicity #55]: #57
-#62 := [trans #58 #60]: #61
-#71 := [monotonicity #62 #68]: #70
-#74 := [monotonicity #71]: #73
-#50 := (iff #15 #49)
-#47 := (iff #14 #44)
-#40 := (implies #10 #37)
-#45 := (iff #40 #44)
-#46 := [rewrite]: #45
-#41 := (iff #14 #40)
-#38 := (iff #13 #37)
-#35 := (= #12 #34)
-#36 := [rewrite]: #35
-#39 := [monotonicity #36]: #38
-#42 := [monotonicity #39]: #41
-#48 := [trans #42 #46]: #47
-#51 := [monotonicity #48]: #50
-#76 := [trans #51 #74]: #75
-#32 := [asserted]: #15
-#77 := [mp #32 #76]: #72
-#79 := [not-or-elim #77]: #67
-#78 := [not-or-elim #77]: #53
-[th-lemma #78 #79]: false
-unsat
-2dc606b7a7bd26f5089ecfacf1617fe8ce130325 132 0
+4a38ab99c62d0654d0c5bc149db3e769cf921224 132 0
 #2 := false
 #9 := 0::real
 decl f3 :: real
@@ -47792,68 +48061,7 @@
 #162 := [th-lemma]: #161
 [unit-resolution #162 #160 #143]: false
 unsat
-0eac487a9514449bb01231188977421b8bb134df 60 0
-#2 := false
-#9 := 0::real
-decl f3 :: real
-#8 := f3
-#67 := 2/3::real
-#68 := (* 2/3::real f3)
-#65 := (>= #68 0::real)
-#63 := (not #65)
-#54 := (>= f3 0::real)
-#69 := (or #54 #63)
-#72 := (not #69)
-#11 := 3::real
-#12 := (/ f3 3::real)
-#13 := (< f3 #12)
-#10 := (< f3 0::real)
-#14 := (implies #10 #13)
-#15 := (not #14)
-#75 := (iff #15 #72)
-#33 := 1/3::real
-#34 := (* 1/3::real f3)
-#37 := (< f3 #34)
-#43 := (not #10)
-#44 := (or #43 #37)
-#49 := (not #44)
-#73 := (iff #49 #72)
-#70 := (iff #44 #69)
-#64 := (iff #37 #63)
-#66 := [rewrite]: #64
-#61 := (iff #43 #54)
-#52 := (not #54)
-#56 := (not #52)
-#59 := (iff #56 #54)
-#60 := [rewrite]: #59
-#57 := (iff #43 #56)
-#53 := (iff #10 #52)
-#55 := [rewrite]: #53
-#58 := [monotonicity #55]: #57
-#62 := [trans #58 #60]: #61
-#71 := [monotonicity #62 #66]: #70
-#74 := [monotonicity #71]: #73
-#50 := (iff #15 #49)
-#47 := (iff #14 #44)
-#40 := (implies #10 #37)
-#45 := (iff #40 #44)
-#46 := [rewrite]: #45
-#41 := (iff #14 #40)
-#38 := (iff #13 #37)
-#35 := (= #12 #34)
-#36 := [rewrite]: #35
-#39 := [monotonicity #36]: #38
-#42 := [monotonicity #39]: #41
-#48 := [trans #42 #46]: #47
-#51 := [monotonicity #48]: #50
-#76 := [trans #51 #74]: #75
-#32 := [asserted]: #15
-#77 := [mp #32 #76]: #72
-#79 := [not-or-elim #77]: #65
-#78 := [not-or-elim #77]: #52
-[th-lemma #78 #79]: false
-unsat
-9db23a49c890fc2c240ef900b9554e7cff9b9382 103 0
+704113f68324742364e301291c46543b32d30adc 103 0
 #2 := false
 #8 := 0::real
 decl f3 :: real
@@ -47957,7 +48165,7 @@
 #124 := [unit-resolution #123 #120]: #115
 [th-lemma #110 #124]: false
 unsat
-949e408b2eea8c150227b26f7ecc3f591066ebd6 149 0
+099809f4b951e6c4b8dd87d068c12977d185ca9a 149 0
 #2 := false
 #9 := 0::real
 decl f3 :: real
@@ -48107,123 +48315,7 @@
 #170 := [unit-resolution #169 #155]: #97
 [th-lemma #159 #170 #167]: false
 unsat
-9e92ed579ec028b4688be273656268ff0a3d0850 57 0
-#2 := false
-#36 := 0::real
-decl f4 :: real
-#9 := f4
-decl f3 :: real
-#8 := f3
-#33 := -1::real
-#34 := (* -1::real f4)
-#35 := (+ f3 #34)
-#37 := (<= #35 0::real)
-#40 := (ite #37 f3 f4)
-#48 := (* -1::real #40)
-#49 := (+ f3 #48)
-#47 := (>= #49 0::real)
-#53 := (not #47)
-#10 := (<= f3 f4)
-#11 := (ite #10 f3 f4)
-#12 := (<= #11 f3)
-#13 := (not #12)
-#54 := (iff #13 #53)
-#51 := (iff #12 #47)
-#43 := (<= #40 f3)
-#46 := (iff #43 #47)
-#50 := [rewrite]: #46
-#44 := (iff #12 #43)
-#41 := (= #11 #40)
-#38 := (iff #10 #37)
-#39 := [rewrite]: #38
-#42 := [monotonicity #39]: #41
-#45 := [monotonicity #42]: #44
-#52 := [trans #45 #50]: #51
-#55 := [monotonicity #52]: #54
-#30 := [asserted]: #13
-#56 := [mp #30 #55]: #53
-#60 := (= f3 #40)
-#31 := (+ f4 #48)
-#65 := (>= #31 0::real)
-#61 := (= f4 #40)
-#62 := (not #37)
-#66 := [hypothesis]: #62
-#59 := (or #37 #61)
-#57 := [def-axiom]: #59
-#67 := [unit-resolution #57 #66]: #61
-#68 := (not #61)
-#69 := (or #68 #65)
-#70 := [th-lemma]: #69
-#71 := [unit-resolution #70 #67]: #65
-#72 := [th-lemma #56 #66 #71]: false
-#73 := [lemma #72]: #37
-#63 := (or #62 #60)
-#58 := [def-axiom]: #63
-#74 := [unit-resolution #58 #73]: #60
-#75 := (not #60)
-#76 := (or #75 #47)
-#77 := [th-lemma]: #76
-[unit-resolution #77 #74 #56]: false
-unsat
-49128a64794a78bd051fb157e2f6db1aa53acbe7 57 0
-#2 := false
-#36 := 0::real
-decl f4 :: real
-#9 := f4
-#33 := -1::real
-#34 := (* -1::real f4)
-decl f3 :: real
-#8 := f3
-#35 := (+ f3 #34)
-#37 := (<= #35 0::real)
-#40 := (ite #37 f3 f4)
-#61 := (= f4 #40)
-#65 := (not #61)
-#47 := (* -1::real #40)
-#48 := (+ f4 #47)
-#46 := (>= #48 0::real)
-#53 := (not #46)
-#10 := (<= f3 f4)
-#11 := (ite #10 f3 f4)
-#12 := (<= #11 f4)
-#13 := (not #12)
-#54 := (iff #13 #53)
-#51 := (iff #12 #46)
-#43 := (<= #40 f4)
-#49 := (iff #43 #46)
-#50 := [rewrite]: #49
-#44 := (iff #12 #43)
-#41 := (= #11 #40)
-#38 := (iff #10 #37)
-#39 := [rewrite]: #38
-#42 := [monotonicity #39]: #41
-#45 := [monotonicity #42]: #44
-#52 := [trans #45 #50]: #51
-#55 := [monotonicity #52]: #54
-#30 := [asserted]: #13
-#56 := [mp #30 #55]: #53
-#64 := [hypothesis]: #61
-#66 := (or #65 #46)
-#67 := [th-lemma]: #66
-#68 := [unit-resolution #67 #64 #56]: false
-#69 := [lemma #68]: #65
-#59 := (or #37 #61)
-#57 := [def-axiom]: #59
-#72 := [unit-resolution #57 #69]: #37
-#31 := (+ f3 #47)
-#71 := (>= #31 0::real)
-#60 := (= f3 #40)
-#62 := (not #37)
-#63 := (or #62 #60)
-#58 := [def-axiom]: #63
-#73 := [unit-resolution #58 #72]: #60
-#74 := (not #60)
-#75 := (or #74 #71)
-#76 := [th-lemma]: #75
-#77 := [unit-resolution #76 #73]: #71
-[th-lemma #77 #56 #72]: false
-unsat
-ba68972e04b6d39e23213630ff5fcde9dcaca230 114 0
+29b8934bcf2c51c8a551e0b17bf814bf8deedd61 114 0
 #2 := false
 #9 := 0::real
 decl f3 :: real
@@ -48338,7 +48430,123 @@
 #143 := [unit-resolution #142 #139]: #137
 [th-lemma #136 #134 #143]: false
 unsat
-a19015a56beb41efa267de75df62459c522bbc0a 103 0
+c7322eea663ce724ffdf3e5de8e20e490ab8caf6 57 0
+#2 := false
+#36 := 0::real
+decl f4 :: real
+#9 := f4
+decl f3 :: real
+#8 := f3
+#33 := -1::real
+#34 := (* -1::real f4)
+#35 := (+ f3 #34)
+#37 := (<= #35 0::real)
+#40 := (ite #37 f3 f4)
+#48 := (* -1::real #40)
+#49 := (+ f3 #48)
+#47 := (>= #49 0::real)
+#53 := (not #47)
+#10 := (<= f3 f4)
+#11 := (ite #10 f3 f4)
+#12 := (<= #11 f3)
+#13 := (not #12)
+#54 := (iff #13 #53)
+#51 := (iff #12 #47)
+#43 := (<= #40 f3)
+#46 := (iff #43 #47)
+#50 := [rewrite]: #46
+#44 := (iff #12 #43)
+#41 := (= #11 #40)
+#38 := (iff #10 #37)
+#39 := [rewrite]: #38
+#42 := [monotonicity #39]: #41
+#45 := [monotonicity #42]: #44
+#52 := [trans #45 #50]: #51
+#55 := [monotonicity #52]: #54
+#30 := [asserted]: #13
+#56 := [mp #30 #55]: #53
+#60 := (= f3 #40)
+#31 := (+ f4 #48)
+#65 := (>= #31 0::real)
+#61 := (= f4 #40)
+#62 := (not #37)
+#66 := [hypothesis]: #62
+#59 := (or #37 #61)
+#57 := [def-axiom]: #59
+#67 := [unit-resolution #57 #66]: #61
+#68 := (not #61)
+#69 := (or #68 #65)
+#70 := [th-lemma]: #69
+#71 := [unit-resolution #70 #67]: #65
+#72 := [th-lemma #56 #66 #71]: false
+#73 := [lemma #72]: #37
+#63 := (or #62 #60)
+#58 := [def-axiom]: #63
+#74 := [unit-resolution #58 #73]: #60
+#75 := (not #60)
+#76 := (or #75 #47)
+#77 := [th-lemma]: #76
+[unit-resolution #77 #74 #56]: false
+unsat
+f580fbf74c17d6b6fb5afde9912701d1c54bba4d 57 0
+#2 := false
+#36 := 0::real
+decl f4 :: real
+#9 := f4
+#33 := -1::real
+#34 := (* -1::real f4)
+decl f3 :: real
+#8 := f3
+#35 := (+ f3 #34)
+#37 := (<= #35 0::real)
+#40 := (ite #37 f3 f4)
+#61 := (= f4 #40)
+#65 := (not #61)
+#47 := (* -1::real #40)
+#48 := (+ f4 #47)
+#46 := (>= #48 0::real)
+#53 := (not #46)
+#10 := (<= f3 f4)
+#11 := (ite #10 f3 f4)
+#12 := (<= #11 f4)
+#13 := (not #12)
+#54 := (iff #13 #53)
+#51 := (iff #12 #46)
+#43 := (<= #40 f4)
+#49 := (iff #43 #46)
+#50 := [rewrite]: #49
+#44 := (iff #12 #43)
+#41 := (= #11 #40)
+#38 := (iff #10 #37)
+#39 := [rewrite]: #38
+#42 := [monotonicity #39]: #41
+#45 := [monotonicity #42]: #44
+#52 := [trans #45 #50]: #51
+#55 := [monotonicity #52]: #54
+#30 := [asserted]: #13
+#56 := [mp #30 #55]: #53
+#64 := [hypothesis]: #61
+#66 := (or #65 #46)
+#67 := [th-lemma]: #66
+#68 := [unit-resolution #67 #64 #56]: false
+#69 := [lemma #68]: #65
+#59 := (or #37 #61)
+#57 := [def-axiom]: #59
+#72 := [unit-resolution #57 #69]: #37
+#31 := (+ f3 #47)
+#71 := (>= #31 0::real)
+#60 := (= f3 #40)
+#62 := (not #37)
+#63 := (or #62 #60)
+#58 := [def-axiom]: #63
+#73 := [unit-resolution #58 #72]: #60
+#74 := (not #60)
+#75 := (or #74 #71)
+#76 := [th-lemma]: #75
+#77 := [unit-resolution #76 #73]: #71
+[th-lemma #77 #56 #72]: false
+unsat
+cc7a80da13c82cc207500cb8afbaae97318e63e2 103 0
 #2 := false
 #45 := 0::real
 decl f5 :: real
@@ -48442,123 +48650,7 @@
 #122 := [unit-resolution #121 #118]: #115
 [th-lemma #122 #92 #93]: false
 unsat
-385a845ab61bebe2a2caff1fe491e1a60a61d0b8 57 0
-#2 := false
-#36 := 0::real
-decl f4 :: real
-#9 := f4
-#33 := -1::real
-#34 := (* -1::real f4)
-decl f3 :: real
-#8 := f3
-#35 := (+ f3 #34)
-#37 := (<= #35 0::real)
-#40 := (ite #37 f4 f3)
-#61 := (= f3 #40)
-#65 := (not #61)
-#46 := (* -1::real #40)
-#47 := (+ f3 #46)
-#48 := (<= #47 0::real)
-#53 := (not #48)
-#10 := (<= f3 f4)
-#11 := (ite #10 f4 f3)
-#12 := (<= f3 #11)
-#13 := (not #12)
-#54 := (iff #13 #53)
-#51 := (iff #12 #48)
-#43 := (<= f3 #40)
-#49 := (iff #43 #48)
-#50 := [rewrite]: #49
-#44 := (iff #12 #43)
-#41 := (= #11 #40)
-#38 := (iff #10 #37)
-#39 := [rewrite]: #38
-#42 := [monotonicity #39]: #41
-#45 := [monotonicity #42]: #44
-#52 := [trans #45 #50]: #51
-#55 := [monotonicity #52]: #54
-#30 := [asserted]: #13
-#56 := [mp #30 #55]: #53
-#64 := [hypothesis]: #61
-#66 := (or #65 #48)
-#67 := [th-lemma]: #66
-#68 := [unit-resolution #67 #64 #56]: false
-#69 := [lemma #68]: #65
-#59 := (or #37 #61)
-#57 := [def-axiom]: #59
-#72 := [unit-resolution #57 #69]: #37
-#31 := (+ f4 #46)
-#70 := (<= #31 0::real)
-#60 := (= f4 #40)
-#62 := (not #37)
-#63 := (or #62 #60)
-#58 := [def-axiom]: #63
-#73 := [unit-resolution #58 #72]: #60
-#74 := (not #60)
-#75 := (or #74 #70)
-#76 := [th-lemma]: #75
-#77 := [unit-resolution #76 #73]: #70
-[th-lemma #77 #56 #72]: false
-unsat
-3bf0824ed026ba333e5bb1b1e54ccba54d70d87f 57 0
-#2 := false
-#35 := 0::real
-decl f4 :: real
-#9 := f4
-decl f3 :: real
-#8 := f3
-#33 := -1::real
-#37 := (* -1::real f4)
-#38 := (+ f3 #37)
-#36 := (>= #38 0::real)
-#40 := (ite #36 f3 f4)
-#46 := (* -1::real #40)
-#47 := (+ f3 #46)
-#48 := (<= #47 0::real)
-#53 := (not #48)
-#10 := (<= f4 f3)
-#11 := (ite #10 f3 f4)
-#12 := (<= f3 #11)
-#13 := (not #12)
-#54 := (iff #13 #53)
-#51 := (iff #12 #48)
-#43 := (<= f3 #40)
-#49 := (iff #43 #48)
-#50 := [rewrite]: #49
-#44 := (iff #12 #43)
-#41 := (= #11 #40)
-#34 := (iff #10 #36)
-#39 := [rewrite]: #34
-#42 := [monotonicity #39]: #41
-#45 := [monotonicity #42]: #44
-#52 := [trans #45 #50]: #51
-#55 := [monotonicity #52]: #54
-#30 := [asserted]: #13
-#56 := [mp #30 #55]: #53
-#60 := (= f3 #40)
-#31 := (+ f4 #46)
-#64 := (<= #31 0::real)
-#61 := (= f4 #40)
-#62 := (not #36)
-#66 := [hypothesis]: #62
-#59 := (or #36 #61)
-#57 := [def-axiom]: #59
-#67 := [unit-resolution #57 #66]: #61
-#68 := (not #61)
-#69 := (or #68 #64)
-#70 := [th-lemma]: #69
-#71 := [unit-resolution #70 #67]: #64
-#72 := [th-lemma #56 #66 #71]: false
-#73 := [lemma #72]: #36
-#63 := (or #62 #60)
-#58 := [def-axiom]: #63
-#74 := [unit-resolution #58 #73]: #60
-#75 := (not #60)
-#76 := (or #75 #48)
-#77 := [th-lemma]: #76
-[unit-resolution #77 #74 #56]: false
-unsat
-37f421aa1e81128329d966841b0fde5f3cd8d944 86 0
+23d9458c2d94638d741588b7f5f53da3b743fc4b 86 0
 #2 := false
 decl f3 :: real
 #8 := f3
@@ -48645,7 +48737,7 @@
 #105 := [trans #104 #83]: #51
 [unit-resolution #57 #105]: false
 unsat
-3b84d41e468d1d259c36f6a458b1bdd65a2c38ee 68 0
+ca3670773dec1aa185941c6109c12ef74c091595 68 0
 #2 := false
 #8 := 0::real
 decl f3 :: real
@@ -48714,111 +48806,7 @@
 #88 := [trans #87 #84]: #13
 [unit-resolution #71 #88]: false
 unsat
-56cfe43fc596233456f5d5ed3a41821e49f560df 103 0
-#2 := false
-#45 := 0::real
-decl f3 :: real
-#8 := f3
-decl f5 :: real
-#11 := f5
-#43 := -1::real
-#51 := (* -1::real f5)
-#63 := (+ f3 #51)
-#64 := (<= #63 0::real)
-#67 := (ite #64 f5 f3)
-#73 := (* -1::real #67)
-decl f4 :: real
-#9 := f4
-#74 := (+ f4 #73)
-#75 := (<= #74 0::real)
-#76 := (not #75)
-#52 := (+ f4 #51)
-#53 := (<= #52 0::real)
-#54 := (not #53)
-#47 := (* -1::real f4)
-#48 := (+ f3 #47)
-#46 := (>= #48 0::real)
-#44 := (not #46)
-#57 := (and #44 #54)
-#60 := (not #57)
-#81 := (or #60 #76)
-#84 := (not #81)
-#14 := (<= f3 f5)
-#15 := (ite #14 f5 f3)
-#16 := (< #15 f4)
-#12 := (< f5 f4)
-#10 := (< f3 f4)
-#13 := (and #10 #12)
-#17 := (implies #13 #16)
-#18 := (not #17)
-#87 := (iff #18 #84)
-#36 := (not #13)
-#37 := (or #36 #16)
-#40 := (not #37)
-#85 := (iff #40 #84)
-#82 := (iff #37 #81)
-#79 := (iff #16 #76)
-#70 := (< #67 f4)
-#77 := (iff #70 #76)
-#78 := [rewrite]: #77
-#71 := (iff #16 #70)
-#68 := (= #15 #67)
-#65 := (iff #14 #64)
-#66 := [rewrite]: #65
-#69 := [monotonicity #66]: #68
-#72 := [monotonicity #69]: #71
-#80 := [trans #72 #78]: #79
-#61 := (iff #36 #60)
-#58 := (iff #13 #57)
-#55 := (iff #12 #54)
-#56 := [rewrite]: #55
-#49 := (iff #10 #44)
-#50 := [rewrite]: #49
-#59 := [monotonicity #50 #56]: #58
-#62 := [monotonicity #59]: #61
-#83 := [monotonicity #62 #80]: #82
-#86 := [monotonicity #83]: #85
-#41 := (iff #18 #40)
-#38 := (iff #17 #37)
-#39 := [rewrite]: #38
-#42 := [monotonicity #39]: #41
-#88 := [trans #42 #86]: #87
-#35 := [asserted]: #18
-#89 := [mp #35 #88]: #84
-#93 := [not-or-elim #89]: #75
-#90 := [not-or-elim #89]: #57
-#91 := [and-elim #90]: #44
-#97 := (+ f3 #73)
-#116 := (>= #97 0::real)
-#104 := (= f3 #67)
-#105 := (not #64)
-#103 := (= f5 #67)
-#110 := (not #103)
-#100 := (+ f5 #73)
-#98 := (>= #100 0::real)
-#107 := (not #98)
-#92 := [and-elim #90]: #54
-#96 := [hypothesis]: #98
-#94 := [th-lemma #96 #92 #93]: false
-#108 := [lemma #94]: #107
-#109 := [hypothesis]: #103
-#111 := (or #110 #98)
-#112 := [th-lemma]: #111
-#113 := [unit-resolution #112 #109 #108]: false
-#114 := [lemma #113]: #110
-#106 := (or #105 #103)
-#101 := [def-axiom]: #106
-#117 := [unit-resolution #101 #114]: #105
-#102 := (or #64 #104)
-#99 := [def-axiom]: #102
-#118 := [unit-resolution #99 #117]: #104
-#119 := (not #104)
-#120 := (or #119 #116)
-#121 := [th-lemma]: #120
-#122 := [unit-resolution #121 #118]: #116
-[th-lemma #122 #91 #93]: false
-unsat
-7c5cf575f0888dae85c506d289c48f18d6429f5a 124 0
+a5005d5e2177d7e219c0762e500aaee9c6b18ed1 124 0
 #2 := false
 #13 := 0::real
 decl f4 :: real
@@ -48943,7 +48931,227 @@
 #145 := [unit-resolution #117 #144]: #106
 [th-lemma #131 #143 #141 #85 #145]: false
 unsat
-d61ccd78ced6caee45a422d1f20159081681ab1d 88 0
+dcaa54c51c4b138dac6c0178c67a0d52f05b3002 57 0
+#2 := false
+#36 := 0::real
+decl f4 :: real
+#9 := f4
+#33 := -1::real
+#34 := (* -1::real f4)
+decl f3 :: real
+#8 := f3
+#35 := (+ f3 #34)
+#37 := (<= #35 0::real)
+#40 := (ite #37 f4 f3)
+#61 := (= f3 #40)
+#65 := (not #61)
+#46 := (* -1::real #40)
+#47 := (+ f3 #46)
+#48 := (<= #47 0::real)
+#53 := (not #48)
+#10 := (<= f3 f4)
+#11 := (ite #10 f4 f3)
+#12 := (<= f3 #11)
+#13 := (not #12)
+#54 := (iff #13 #53)
+#51 := (iff #12 #48)
+#43 := (<= f3 #40)
+#49 := (iff #43 #48)
+#50 := [rewrite]: #49
+#44 := (iff #12 #43)
+#41 := (= #11 #40)
+#38 := (iff #10 #37)
+#39 := [rewrite]: #38
+#42 := [monotonicity #39]: #41
+#45 := [monotonicity #42]: #44
+#52 := [trans #45 #50]: #51
+#55 := [monotonicity #52]: #54
+#30 := [asserted]: #13
+#56 := [mp #30 #55]: #53
+#64 := [hypothesis]: #61
+#66 := (or #65 #48)
+#67 := [th-lemma]: #66
+#68 := [unit-resolution #67 #64 #56]: false
+#69 := [lemma #68]: #65
+#59 := (or #37 #61)
+#57 := [def-axiom]: #59
+#72 := [unit-resolution #57 #69]: #37
+#31 := (+ f4 #46)
+#70 := (<= #31 0::real)
+#60 := (= f4 #40)
+#62 := (not #37)
+#63 := (or #62 #60)
+#58 := [def-axiom]: #63
+#73 := [unit-resolution #58 #72]: #60
+#74 := (not #60)
+#75 := (or #74 #70)
+#76 := [th-lemma]: #75
+#77 := [unit-resolution #76 #73]: #70
+[th-lemma #77 #56 #72]: false
+unsat
+ef39e4f5f84c012d1a7afd1eb9ae7787355eb901 57 0
+#2 := false
+#35 := 0::real
+decl f4 :: real
+#9 := f4
+decl f3 :: real
+#8 := f3
+#33 := -1::real
+#37 := (* -1::real f4)
+#38 := (+ f3 #37)
+#36 := (>= #38 0::real)
+#40 := (ite #36 f3 f4)
+#46 := (* -1::real #40)
+#47 := (+ f3 #46)
+#48 := (<= #47 0::real)
+#53 := (not #48)
+#10 := (<= f4 f3)
+#11 := (ite #10 f3 f4)
+#12 := (<= f3 #11)
+#13 := (not #12)
+#54 := (iff #13 #53)
+#51 := (iff #12 #48)
+#43 := (<= f3 #40)
+#49 := (iff #43 #48)
+#50 := [rewrite]: #49
+#44 := (iff #12 #43)
+#41 := (= #11 #40)
+#34 := (iff #10 #36)
+#39 := [rewrite]: #34
+#42 := [monotonicity #39]: #41
+#45 := [monotonicity #42]: #44
+#52 := [trans #45 #50]: #51
+#55 := [monotonicity #52]: #54
+#30 := [asserted]: #13
+#56 := [mp #30 #55]: #53
+#60 := (= f3 #40)
+#31 := (+ f4 #46)
+#64 := (<= #31 0::real)
+#61 := (= f4 #40)
+#62 := (not #36)
+#66 := [hypothesis]: #62
+#59 := (or #36 #61)
+#57 := [def-axiom]: #59
+#67 := [unit-resolution #57 #66]: #61
+#68 := (not #61)
+#69 := (or #68 #64)
+#70 := [th-lemma]: #69
+#71 := [unit-resolution #70 #67]: #64
+#72 := [th-lemma #56 #66 #71]: false
+#73 := [lemma #72]: #36
+#63 := (or #62 #60)
+#58 := [def-axiom]: #63
+#74 := [unit-resolution #58 #73]: #60
+#75 := (not #60)
+#76 := (or #75 #48)
+#77 := [th-lemma]: #76
+[unit-resolution #77 #74 #56]: false
+unsat
+7eae59ca8b45a027fbc591665e655f191dd6bbc4 103 0
+#2 := false
+#45 := 0::real
+decl f3 :: real
+#8 := f3
+decl f5 :: real
+#11 := f5
+#43 := -1::real
+#51 := (* -1::real f5)
+#63 := (+ f3 #51)
+#64 := (<= #63 0::real)
+#67 := (ite #64 f5 f3)
+#73 := (* -1::real #67)
+decl f4 :: real
+#9 := f4
+#74 := (+ f4 #73)
+#75 := (<= #74 0::real)
+#76 := (not #75)
+#52 := (+ f4 #51)
+#53 := (<= #52 0::real)
+#54 := (not #53)
+#47 := (* -1::real f4)
+#48 := (+ f3 #47)
+#46 := (>= #48 0::real)
+#44 := (not #46)
+#57 := (and #44 #54)
+#60 := (not #57)
+#81 := (or #60 #76)
+#84 := (not #81)
+#14 := (<= f3 f5)
+#15 := (ite #14 f5 f3)
+#16 := (< #15 f4)
+#12 := (< f5 f4)
+#10 := (< f3 f4)
+#13 := (and #10 #12)
+#17 := (implies #13 #16)
+#18 := (not #17)
+#87 := (iff #18 #84)
+#36 := (not #13)
+#37 := (or #36 #16)
+#40 := (not #37)
+#85 := (iff #40 #84)
+#82 := (iff #37 #81)
+#79 := (iff #16 #76)
+#70 := (< #67 f4)
+#77 := (iff #70 #76)
+#78 := [rewrite]: #77
+#71 := (iff #16 #70)
+#68 := (= #15 #67)
+#65 := (iff #14 #64)
+#66 := [rewrite]: #65
+#69 := [monotonicity #66]: #68
+#72 := [monotonicity #69]: #71
+#80 := [trans #72 #78]: #79
+#61 := (iff #36 #60)
+#58 := (iff #13 #57)
+#55 := (iff #12 #54)
+#56 := [rewrite]: #55
+#49 := (iff #10 #44)
+#50 := [rewrite]: #49
+#59 := [monotonicity #50 #56]: #58
+#62 := [monotonicity #59]: #61
+#83 := [monotonicity #62 #80]: #82
+#86 := [monotonicity #83]: #85
+#41 := (iff #18 #40)
+#38 := (iff #17 #37)
+#39 := [rewrite]: #38
+#42 := [monotonicity #39]: #41
+#88 := [trans #42 #86]: #87
+#35 := [asserted]: #18
+#89 := [mp #35 #88]: #84
+#93 := [not-or-elim #89]: #75
+#90 := [not-or-elim #89]: #57
+#91 := [and-elim #90]: #44
+#97 := (+ f3 #73)
+#116 := (>= #97 0::real)
+#104 := (= f3 #67)
+#105 := (not #64)
+#103 := (= f5 #67)
+#110 := (not #103)
+#100 := (+ f5 #73)
+#98 := (>= #100 0::real)
+#107 := (not #98)
+#92 := [and-elim #90]: #54
+#96 := [hypothesis]: #98
+#94 := [th-lemma #96 #92 #93]: false
+#108 := [lemma #94]: #107
+#109 := [hypothesis]: #103
+#111 := (or #110 #98)
+#112 := [th-lemma]: #111
+#113 := [unit-resolution #112 #109 #108]: false
+#114 := [lemma #113]: #110
+#106 := (or #105 #103)
+#101 := [def-axiom]: #106
+#117 := [unit-resolution #101 #114]: #105
+#102 := (or #64 #104)
+#99 := [def-axiom]: #102
+#118 := [unit-resolution #99 #117]: #104
+#119 := (not #104)
+#120 := (or #119 #116)
+#121 := [th-lemma]: #120
+#122 := [unit-resolution #121 #118]: #116
+[th-lemma #122 #91 #93]: false
+unsat
+f8a9afea5f80c13265f7efc519efee9034d86aaa 88 0
 #2 := false
 decl f4 :: real
 #9 := f4
@@ -49032,7 +49240,7 @@
 #109 := [trans #108 #83]: #51
 [unit-resolution #57 #109]: false
 unsat
-ca09892ea7a85486f9b03ea9c10592c837c45810 66 0
+c008fb0a346a13df61c397707ef86eadaaaaf028 66 0
 #2 := false
 decl f3 :: real
 #9 := f3
@@ -49099,26 +49307,7 @@
 #86 := [trans #83 #85]: #33
 [unit-resolution #69 #86]: false
 unsat
-43eaef47fadbc62036dedee99a53f5627c2aea2d 18 0
-#2 := false
-decl f3 :: real
-#8 := f3
-#9 := (<= f3 f3)
-#10 := (not #9)
-#37 := (iff #10 false)
-#1 := true
-#32 := (not true)
-#35 := (iff #32 false)
-#36 := [rewrite]: #35
-#33 := (iff #10 #32)
-#30 := (iff #9 true)
-#31 := [rewrite]: #30
-#34 := [monotonicity #31]: #33
-#38 := [trans #34 #36]: #37
-#27 := [asserted]: #10
-[mp #27 #38]: false
-unsat
-2a53f1e9691b42c6b60d21a29f4be1df277b5c0e 228 0
+3ea0cce22deeb1cf0a91c70b01f38d938cbfca2b 228 0
 #2 := false
 #9 := 0::real
 decl f4 :: real
@@ -49347,7 +49536,26 @@
 #250 := [unit-resolution #196 #249]: #190
 [th-lemma #236 #250 #131 #234 #248 #238]: false
 unsat
-f4ae26a6b5fd11d6cc627389a122029146dcae87 50 0
+6fa8c48985ebc99327aefe25360343fbef04699f 18 0
+#2 := false
+decl f3 :: real
+#8 := f3
+#9 := (<= f3 f3)
+#10 := (not #9)
+#37 := (iff #10 false)
+#1 := true
+#32 := (not true)
+#35 := (iff #32 false)
+#36 := [rewrite]: #35
+#33 := (iff #10 #32)
+#30 := (iff #9 true)
+#31 := [rewrite]: #30
+#34 := [monotonicity #31]: #33
+#38 := [trans #34 #36]: #37
+#27 := [asserted]: #10
+[mp #27 #38]: false
+unsat
+6246909589ff44c53af5060ab09e0fb241168550 50 0
 #2 := false
 #44 := 0::real
 decl f4 :: real
@@ -49398,7 +49606,7 @@
 #68 := [not-or-elim #65]: #67
 [th-lemma #68 #66]: false
 unsat
-3c98043342d67c384e77c1b1c5755900a361dac4 55 0
+9a368b708e53fe0140c26e7ca790639aa465ac37 55 0
 #2 := false
 #43 := 0::real
 decl f4 :: real
@@ -49454,7 +49662,7 @@
 #72 := [not-or-elim #71]: #42
 [th-lemma #72 #73]: false
 unsat
-fe03b23c1657b4573ded6dbd4e82c4b99ddfa3be 52 0
+9d1b38a4c0ecf43766b852c7e4d288a3e46da901 52 0
 #2 := false
 #40 := 0::real
 decl f4 :: real
@@ -49507,7 +49715,417 @@
 #77 := [unit-resolution #76 #65]: #53
 [unit-resolution #77 #67]: false
 unsat
-d1bc3a4e60e3c1d280eca79821081b4e3318341f 43 0
+5631a67f2c650efd9761a68c96a2b28d9ed4cdd0 15 0
+#2 := false
+decl f3 :: real
+#8 := f3
+#9 := (< f3 f3)
+#10 := (not #9)
+#11 := (not #10)
+#33 := (iff #11 false)
+#31 := (iff #9 false)
+#32 := [rewrite]: #31
+#29 := (iff #11 #9)
+#30 := [rewrite]: #29
+#34 := [trans #30 #32]: #33
+#28 := [asserted]: #11
+[mp #28 #34]: false
+unsat
+c41bf12f912d9026a2697f2c615b98bf21c74fe7 75 0
+#2 := false
+#53 := 0::real
+decl f4 :: real
+#9 := f4
+#50 := -1::real
+#64 := (* -1::real f4)
+decl f3 :: real
+#8 := f3
+#65 := (+ f3 #64)
+#66 := (<= #65 0::real)
+#69 := (not #66)
+decl f5 :: real
+#11 := f5
+#51 := (* -1::real f5)
+#60 := (+ f3 #51)
+#61 := (<= #60 0::real)
+#52 := (+ f4 #51)
+#54 := (<= #52 0::real)
+#57 := (not #54)
+#78 := (or #57 #61 #69)
+#83 := (not #78)
+#13 := (<= f3 f5)
+#12 := (<= f4 f5)
+#14 := (implies #12 #13)
+#10 := (<= f3 f4)
+#15 := (implies #10 #14)
+#16 := (not #15)
+#86 := (iff #16 #83)
+#34 := (not #12)
+#35 := (or #34 #13)
+#41 := (not #10)
+#42 := (or #41 #35)
+#47 := (not #42)
+#84 := (iff #47 #83)
+#81 := (iff #42 #78)
+#72 := (or #57 #61)
+#75 := (or #69 #72)
+#79 := (iff #75 #78)
+#80 := [rewrite]: #79
+#76 := (iff #42 #75)
+#73 := (iff #35 #72)
+#62 := (iff #13 #61)
+#63 := [rewrite]: #62
+#58 := (iff #34 #57)
+#55 := (iff #12 #54)
+#56 := [rewrite]: #55
+#59 := [monotonicity #56]: #58
+#74 := [monotonicity #59 #63]: #73
+#70 := (iff #41 #69)
+#67 := (iff #10 #66)
+#68 := [rewrite]: #67
+#71 := [monotonicity #68]: #70
+#77 := [monotonicity #71 #74]: #76
+#82 := [trans #77 #80]: #81
+#85 := [monotonicity #82]: #84
+#48 := (iff #16 #47)
+#45 := (iff #15 #42)
+#38 := (implies #10 #35)
+#43 := (iff #38 #42)
+#44 := [rewrite]: #43
+#39 := (iff #15 #38)
+#36 := (iff #14 #35)
+#37 := [rewrite]: #36
+#40 := [monotonicity #37]: #39
+#46 := [trans #40 #44]: #45
+#49 := [monotonicity #46]: #48
+#87 := [trans #49 #85]: #86
+#33 := [asserted]: #16
+#88 := [mp #33 #87]: #83
+#92 := [not-or-elim #88]: #66
+#90 := (not #61)
+#91 := [not-or-elim #88]: #90
+#89 := [not-or-elim #88]: #54
+[th-lemma #89 #91 #92]: false
+unsat
+9f5fa7c760db14a67656e2a38b69825c2338f5ee 80 0
+#2 := false
+#53 := 0::real
+decl f4 :: real
+#9 := f4
+#50 := -1::real
+#66 := (* -1::real f4)
+decl f3 :: real
+#8 := f3
+#67 := (+ f3 #66)
+#65 := (>= #67 0::real)
+#64 := (not #65)
+decl f5 :: real
+#11 := f5
+#51 := (* -1::real f5)
+#60 := (+ f3 #51)
+#61 := (<= #60 0::real)
+#52 := (+ f4 #51)
+#54 := (<= #52 0::real)
+#57 := (not #54)
+#83 := (or #57 #61 #65)
+#88 := (not #83)
+#13 := (<= f3 f5)
+#12 := (<= f4 f5)
+#14 := (implies #12 #13)
+#10 := (< f3 f4)
+#15 := (implies #10 #14)
+#16 := (not #15)
+#91 := (iff #16 #88)
+#34 := (not #12)
+#35 := (or #34 #13)
+#41 := (not #10)
+#42 := (or #41 #35)
+#47 := (not #42)
+#89 := (iff #47 #88)
+#86 := (iff #42 #83)
+#77 := (or #57 #61)
+#80 := (or #65 #77)
+#84 := (iff #80 #83)
+#85 := [rewrite]: #84
+#81 := (iff #42 #80)
+#78 := (iff #35 #77)
+#62 := (iff #13 #61)
+#63 := [rewrite]: #62
+#58 := (iff #34 #57)
+#55 := (iff #12 #54)
+#56 := [rewrite]: #55
+#59 := [monotonicity #56]: #58
+#79 := [monotonicity #59 #63]: #78
+#75 := (iff #41 #65)
+#70 := (not #64)
+#73 := (iff #70 #65)
+#74 := [rewrite]: #73
+#71 := (iff #41 #70)
+#68 := (iff #10 #64)
+#69 := [rewrite]: #68
+#72 := [monotonicity #69]: #71
+#76 := [trans #72 #74]: #75
+#82 := [monotonicity #76 #79]: #81
+#87 := [trans #82 #85]: #86
+#90 := [monotonicity #87]: #89
+#48 := (iff #16 #47)
+#45 := (iff #15 #42)
+#38 := (implies #10 #35)
+#43 := (iff #38 #42)
+#44 := [rewrite]: #43
+#39 := (iff #15 #38)
+#36 := (iff #14 #35)
+#37 := [rewrite]: #36
+#40 := [monotonicity #37]: #39
+#46 := [trans #40 #44]: #45
+#49 := [monotonicity #46]: #48
+#92 := [trans #49 #90]: #91
+#33 := [asserted]: #16
+#93 := [mp #33 #92]: #88
+#97 := [not-or-elim #93]: #64
+#95 := (not #61)
+#96 := [not-or-elim #93]: #95
+#94 := [not-or-elim #93]: #54
+[th-lemma #94 #96 #97]: false
+unsat
+ef54c95730a86f29e142ca8c0703d6c3a6105c5e 80 0
+#2 := false
+#52 := 0::real
+decl f4 :: real
+#9 := f4
+#50 := -1::real
+#69 := (* -1::real f4)
+decl f3 :: real
+#8 := f3
+#70 := (+ f3 #69)
+#71 := (<= #70 0::real)
+#74 := (not #71)
+decl f5 :: real
+#11 := f5
+#54 := (* -1::real f5)
+#65 := (+ f3 #54)
+#66 := (<= #65 0::real)
+#55 := (+ f4 #54)
+#53 := (>= #55 0::real)
+#83 := (or #53 #66 #74)
+#88 := (not #83)
+#13 := (<= f3 f5)
+#12 := (< f4 f5)
+#14 := (implies #12 #13)
+#10 := (<= f3 f4)
+#15 := (implies #10 #14)
+#16 := (not #15)
+#91 := (iff #16 #88)
+#34 := (not #12)
+#35 := (or #34 #13)
+#41 := (not #10)
+#42 := (or #41 #35)
+#47 := (not #42)
+#89 := (iff #47 #88)
+#86 := (iff #42 #83)
+#77 := (or #53 #66)
+#80 := (or #74 #77)
+#84 := (iff #80 #83)
+#85 := [rewrite]: #84
+#81 := (iff #42 #80)
+#78 := (iff #35 #77)
+#67 := (iff #13 #66)
+#68 := [rewrite]: #67
+#63 := (iff #34 #53)
+#51 := (not #53)
+#58 := (not #51)
+#61 := (iff #58 #53)
+#62 := [rewrite]: #61
+#59 := (iff #34 #58)
+#56 := (iff #12 #51)
+#57 := [rewrite]: #56
+#60 := [monotonicity #57]: #59
+#64 := [trans #60 #62]: #63
+#79 := [monotonicity #64 #68]: #78
+#75 := (iff #41 #74)
+#72 := (iff #10 #71)
+#73 := [rewrite]: #72
+#76 := [monotonicity #73]: #75
+#82 := [monotonicity #76 #79]: #81
+#87 := [trans #82 #85]: #86
+#90 := [monotonicity #87]: #89
+#48 := (iff #16 #47)
+#45 := (iff #15 #42)
+#38 := (implies #10 #35)
+#43 := (iff #38 #42)
+#44 := [rewrite]: #43
+#39 := (iff #15 #38)
+#36 := (iff #14 #35)
+#37 := [rewrite]: #36
+#40 := [monotonicity #37]: #39
+#46 := [trans #40 #44]: #45
+#49 := [monotonicity #46]: #48
+#92 := [trans #49 #90]: #91
+#33 := [asserted]: #16
+#93 := [mp #33 #92]: #88
+#97 := [not-or-elim #93]: #71
+#94 := [not-or-elim #93]: #51
+#95 := (not #66)
+#96 := [not-or-elim #93]: #95
+[th-lemma #96 #94 #97]: false
+unsat
+0c5d7dbb8aa2814977e1a9556d08eabbc11dd00b 85 0
+#2 := false
+#52 := 0::real
+decl f4 :: real
+#9 := f4
+#50 := -1::real
+#72 := (* -1::real f4)
+decl f3 :: real
+#8 := f3
+#73 := (+ f3 #72)
+#71 := (>= #73 0::real)
+#70 := (not #71)
+decl f5 :: real
+#11 := f5
+#54 := (* -1::real f5)
+#67 := (+ f3 #54)
+#66 := (>= #67 0::real)
+#65 := (not #66)
+#55 := (+ f4 #54)
+#53 := (>= #55 0::real)
+#89 := (or #53 #65 #71)
+#94 := (not #89)
+#13 := (< f3 f5)
+#12 := (< f4 f5)
+#14 := (implies #12 #13)
+#10 := (< f3 f4)
+#15 := (implies #10 #14)
+#16 := (not #15)
+#97 := (iff #16 #94)
+#34 := (not #12)
+#35 := (or #34 #13)
+#41 := (not #10)
+#42 := (or #41 #35)
+#47 := (not #42)
+#95 := (iff #47 #94)
+#92 := (iff #42 #89)
+#83 := (or #53 #65)
+#86 := (or #71 #83)
+#90 := (iff #86 #89)
+#91 := [rewrite]: #90
+#87 := (iff #42 #86)
+#84 := (iff #35 #83)
+#68 := (iff #13 #65)
+#69 := [rewrite]: #68
+#63 := (iff #34 #53)
+#51 := (not #53)
+#58 := (not #51)
+#61 := (iff #58 #53)
+#62 := [rewrite]: #61
+#59 := (iff #34 #58)
+#56 := (iff #12 #51)
+#57 := [rewrite]: #56
+#60 := [monotonicity #57]: #59
+#64 := [trans #60 #62]: #63
+#85 := [monotonicity #64 #69]: #84
+#81 := (iff #41 #71)
+#76 := (not #70)
+#79 := (iff #76 #71)
+#80 := [rewrite]: #79
+#77 := (iff #41 #76)
+#74 := (iff #10 #70)
+#75 := [rewrite]: #74
+#78 := [monotonicity #75]: #77
+#82 := [trans #78 #80]: #81
+#88 := [monotonicity #82 #85]: #87
+#93 := [trans #88 #91]: #92
+#96 := [monotonicity #93]: #95
+#48 := (iff #16 #47)
+#45 := (iff #15 #42)
+#38 := (implies #10 #35)
+#43 := (iff #38 #42)
+#44 := [rewrite]: #43
+#39 := (iff #15 #38)
+#36 := (iff #14 #35)
+#37 := [rewrite]: #36
+#40 := [monotonicity #37]: #39
+#46 := [trans #40 #44]: #45
+#49 := [monotonicity #46]: #48
+#98 := [trans #49 #96]: #97
+#33 := [asserted]: #16
+#99 := [mp #33 #98]: #94
+#102 := [not-or-elim #99]: #70
+#100 := [not-or-elim #99]: #51
+#101 := [not-or-elim #99]: #66
+[th-lemma #101 #100 #102]: false
+unsat
+ad850a4d677f642c1e98f0e5f3b7d3f84b6c874a 69 0
+#2 := false
+#44 := 0::real
+decl f5 :: real
+#11 := f5
+#42 := -1::real
+#51 := (* -1::real f5)
+decl f3 :: real
+#8 := f3
+#62 := (+ f3 #51)
+#63 := (<= #62 0::real)
+#64 := (not #63)
+decl f4 :: real
+#9 := f4
+#52 := (+ f4 #51)
+#50 := (>= #52 0::real)
+#53 := (not #50)
+#46 := (* -1::real f4)
+#47 := (+ f3 #46)
+#45 := (>= #47 0::real)
+#43 := (not #45)
+#56 := (and #43 #53)
+#59 := (not #56)
+#74 := (or #59 #63)
+#77 := (not #74)
+#14 := (< f5 f3)
+#15 := (not #14)
+#12 := (< f4 f5)
+#10 := (< f3 f4)
+#13 := (and #10 #12)
+#16 := (implies #13 #15)
+#17 := (not #16)
+#80 := (iff #17 #77)
+#35 := (not #13)
+#36 := (or #35 #15)
+#39 := (not #36)
+#78 := (iff #39 #77)
+#75 := (iff #36 #74)
+#72 := (iff #15 #63)
+#67 := (not #64)
+#70 := (iff #67 #63)
+#71 := [rewrite]: #70
+#68 := (iff #15 #67)
+#65 := (iff #14 #64)
+#66 := [rewrite]: #65
+#69 := [monotonicity #66]: #68
+#73 := [trans #69 #71]: #72
+#60 := (iff #35 #59)
+#57 := (iff #13 #56)
+#54 := (iff #12 #53)
+#55 := [rewrite]: #54
+#48 := (iff #10 #43)
+#49 := [rewrite]: #48
+#58 := [monotonicity #49 #55]: #57
+#61 := [monotonicity #58]: #60
+#76 := [monotonicity #61 #73]: #75
+#79 := [monotonicity #76]: #78
+#40 := (iff #17 #39)
+#37 := (iff #16 #36)
+#38 := [rewrite]: #37
+#41 := [monotonicity #38]: #40
+#81 := [trans #41 #79]: #80
+#34 := [asserted]: #17
+#82 := [mp #34 #81]: #77
+#86 := [not-or-elim #82]: #64
+#83 := [not-or-elim #82]: #56
+#84 := [and-elim #83]: #43
+#85 := [and-elim #83]: #53
+[th-lemma #85 #84 #86]: false
+unsat
+7db3ee0924a4cb1b14a4cebbf1c598c0b8dbec57 43 0
 #2 := false
 decl f5 :: (-> S2 S2 S3)
 #15 := (:var 0 S2)
@@ -49551,7 +50169,7 @@
 #145 := [quant-inst]: #231
 [unit-resolution #145 #45 #567]: false
 unsat
-2fee6903df28c3fb9fdfe41b68f30e20ac97873f 43 0
+8502f9cde674da83d9cbc8a977644ab6e85265ed 43 0
 #2 := false
 decl f5 :: (-> S2 S2 S3)
 #15 := (:var 0 S2)
@@ -49595,23 +50213,7 @@
 #163 := [quant-inst]: #235
 [unit-resolution #163 #45 #573]: false
 unsat
-9e985a47bba5857a10e800553a31e2def569aae1 15 0
-#2 := false
-decl f3 :: real
-#8 := f3
-#9 := (< f3 f3)
-#10 := (not #9)
-#11 := (not #10)
-#33 := (iff #11 false)
-#31 := (iff #9 false)
-#32 := [rewrite]: #31
-#29 := (iff #11 #9)
-#30 := [rewrite]: #29
-#34 := [trans #30 #32]: #33
-#28 := [asserted]: #11
-[mp #28 #34]: false
-unsat
-27ecb5fef2d88b43d79e314c4934972296e27d35 85 0
+7028fd42bf273f5da876a633b9b2fac12fd44679 85 0
 #2 := false
 decl f5 :: S2
 #9 := f5
@@ -49697,83 +50299,7 @@
 #206 := [unit-resolution #170 #205]: #239
 [unit-resolution #206 #186]: false
 unsat
-06bb2cc4f48d1ee0c5893ec218e691f0f9b659a2 75 0
-#2 := false
-#53 := 0::real
-decl f4 :: real
-#9 := f4
-#50 := -1::real
-#64 := (* -1::real f4)
-decl f3 :: real
-#8 := f3
-#65 := (+ f3 #64)
-#66 := (<= #65 0::real)
-#69 := (not #66)
-decl f5 :: real
-#11 := f5
-#51 := (* -1::real f5)
-#60 := (+ f3 #51)
-#61 := (<= #60 0::real)
-#52 := (+ f4 #51)
-#54 := (<= #52 0::real)
-#57 := (not #54)
-#78 := (or #57 #61 #69)
-#83 := (not #78)
-#13 := (<= f3 f5)
-#12 := (<= f4 f5)
-#14 := (implies #12 #13)
-#10 := (<= f3 f4)
-#15 := (implies #10 #14)
-#16 := (not #15)
-#86 := (iff #16 #83)
-#34 := (not #12)
-#35 := (or #34 #13)
-#41 := (not #10)
-#42 := (or #41 #35)
-#47 := (not #42)
-#84 := (iff #47 #83)
-#81 := (iff #42 #78)
-#72 := (or #57 #61)
-#75 := (or #69 #72)
-#79 := (iff #75 #78)
-#80 := [rewrite]: #79
-#76 := (iff #42 #75)
-#73 := (iff #35 #72)
-#62 := (iff #13 #61)
-#63 := [rewrite]: #62
-#58 := (iff #34 #57)
-#55 := (iff #12 #54)
-#56 := [rewrite]: #55
-#59 := [monotonicity #56]: #58
-#74 := [monotonicity #59 #63]: #73
-#70 := (iff #41 #69)
-#67 := (iff #10 #66)
-#68 := [rewrite]: #67
-#71 := [monotonicity #68]: #70
-#77 := [monotonicity #71 #74]: #76
-#82 := [trans #77 #80]: #81
-#85 := [monotonicity #82]: #84
-#48 := (iff #16 #47)
-#45 := (iff #15 #42)
-#38 := (implies #10 #35)
-#43 := (iff #38 #42)
-#44 := [rewrite]: #43
-#39 := (iff #15 #38)
-#36 := (iff #14 #35)
-#37 := [rewrite]: #36
-#40 := [monotonicity #37]: #39
-#46 := [trans #40 #44]: #45
-#49 := [monotonicity #46]: #48
-#87 := [trans #49 #85]: #86
-#33 := [asserted]: #16
-#88 := [mp #33 #87]: #83
-#92 := [not-or-elim #88]: #66
-#90 := (not #61)
-#91 := [not-or-elim #88]: #90
-#89 := [not-or-elim #88]: #54
-[th-lemma #89 #91 #92]: false
-unsat
-8078fea9528cc0af35f2ba128aaa4ddff79359bb 155 0
+ab91d7002afe61891f92aa68c03a93dbeac150bb 155 0
 #2 := false
 decl f7 :: S2
 #12 := f7
@@ -49929,88 +50455,7 @@
 #263 := [def-axiom]: #176
 [unit-resolution #263 #289 #215 #569]: false
 unsat
-783cc208ed65b94bffd8eeb1bb636056967f88c2 80 0
-#2 := false
-#53 := 0::real
-decl f4 :: real
-#9 := f4
-#50 := -1::real
-#66 := (* -1::real f4)
-decl f3 :: real
-#8 := f3
-#67 := (+ f3 #66)
-#65 := (>= #67 0::real)
-#64 := (not #65)
-decl f5 :: real
-#11 := f5
-#51 := (* -1::real f5)
-#60 := (+ f3 #51)
-#61 := (<= #60 0::real)
-#52 := (+ f4 #51)
-#54 := (<= #52 0::real)
-#57 := (not #54)
-#83 := (or #57 #61 #65)
-#88 := (not #83)
-#13 := (<= f3 f5)
-#12 := (<= f4 f5)
-#14 := (implies #12 #13)
-#10 := (< f3 f4)
-#15 := (implies #10 #14)
-#16 := (not #15)
-#91 := (iff #16 #88)
-#34 := (not #12)
-#35 := (or #34 #13)
-#41 := (not #10)
-#42 := (or #41 #35)
-#47 := (not #42)
-#89 := (iff #47 #88)
-#86 := (iff #42 #83)
-#77 := (or #57 #61)
-#80 := (or #65 #77)
-#84 := (iff #80 #83)
-#85 := [rewrite]: #84
-#81 := (iff #42 #80)
-#78 := (iff #35 #77)
-#62 := (iff #13 #61)
-#63 := [rewrite]: #62
-#58 := (iff #34 #57)
-#55 := (iff #12 #54)
-#56 := [rewrite]: #55
-#59 := [monotonicity #56]: #58
-#79 := [monotonicity #59 #63]: #78
-#75 := (iff #41 #65)
-#70 := (not #64)
-#73 := (iff #70 #65)
-#74 := [rewrite]: #73
-#71 := (iff #41 #70)
-#68 := (iff #10 #64)
-#69 := [rewrite]: #68
-#72 := [monotonicity #69]: #71
-#76 := [trans #72 #74]: #75
-#82 := [monotonicity #76 #79]: #81
-#87 := [trans #82 #85]: #86
-#90 := [monotonicity #87]: #89
-#48 := (iff #16 #47)
-#45 := (iff #15 #42)
-#38 := (implies #10 #35)
-#43 := (iff #38 #42)
-#44 := [rewrite]: #43
-#39 := (iff #15 #38)
-#36 := (iff #14 #35)
-#37 := [rewrite]: #36
-#40 := [monotonicity #37]: #39
-#46 := [trans #40 #44]: #45
-#49 := [monotonicity #46]: #48
-#92 := [trans #49 #90]: #91
-#33 := [asserted]: #16
-#93 := [mp #33 #92]: #88
-#97 := [not-or-elim #93]: #64
-#95 := (not #61)
-#96 := [not-or-elim #93]: #95
-#94 := [not-or-elim #93]: #54
-[th-lemma #94 #96 #97]: false
-unsat
-94dd01d5ce8a66038a4fafc28cc8fc9a3d392fba 91 0
+c53da5ad16e14d4123258f05451b97645b05e47b 91 0
 #2 := false
 decl f9 :: S3
 #14 := f9
@@ -50102,88 +50547,7 @@
 #320 := [unit-resolution #230 #598]: #299
 [unit-resolution #320 #321]: false
 unsat
-8618512b77e4b49096ff35dd0be727f2675c21ad 80 0
-#2 := false
-#52 := 0::real
-decl f4 :: real
-#9 := f4
-#50 := -1::real
-#69 := (* -1::real f4)
-decl f3 :: real
-#8 := f3
-#70 := (+ f3 #69)
-#71 := (<= #70 0::real)
-#74 := (not #71)
-decl f5 :: real
-#11 := f5
-#54 := (* -1::real f5)
-#65 := (+ f3 #54)
-#66 := (<= #65 0::real)
-#55 := (+ f4 #54)
-#53 := (>= #55 0::real)
-#83 := (or #53 #66 #74)
-#88 := (not #83)
-#13 := (<= f3 f5)
-#12 := (< f4 f5)
-#14 := (implies #12 #13)
-#10 := (<= f3 f4)
-#15 := (implies #10 #14)
-#16 := (not #15)
-#91 := (iff #16 #88)
-#34 := (not #12)
-#35 := (or #34 #13)
-#41 := (not #10)
-#42 := (or #41 #35)
-#47 := (not #42)
-#89 := (iff #47 #88)
-#86 := (iff #42 #83)
-#77 := (or #53 #66)
-#80 := (or #74 #77)
-#84 := (iff #80 #83)
-#85 := [rewrite]: #84
-#81 := (iff #42 #80)
-#78 := (iff #35 #77)
-#67 := (iff #13 #66)
-#68 := [rewrite]: #67
-#63 := (iff #34 #53)
-#51 := (not #53)
-#58 := (not #51)
-#61 := (iff #58 #53)
-#62 := [rewrite]: #61
-#59 := (iff #34 #58)
-#56 := (iff #12 #51)
-#57 := [rewrite]: #56
-#60 := [monotonicity #57]: #59
-#64 := [trans #60 #62]: #63
-#79 := [monotonicity #64 #68]: #78
-#75 := (iff #41 #74)
-#72 := (iff #10 #71)
-#73 := [rewrite]: #72
-#76 := [monotonicity #73]: #75
-#82 := [monotonicity #76 #79]: #81
-#87 := [trans #82 #85]: #86
-#90 := [monotonicity #87]: #89
-#48 := (iff #16 #47)
-#45 := (iff #15 #42)
-#38 := (implies #10 #35)
-#43 := (iff #38 #42)
-#44 := [rewrite]: #43
-#39 := (iff #15 #38)
-#36 := (iff #14 #35)
-#37 := [rewrite]: #36
-#40 := [monotonicity #37]: #39
-#46 := [trans #40 #44]: #45
-#49 := [monotonicity #46]: #48
-#92 := [trans #49 #90]: #91
-#33 := [asserted]: #16
-#93 := [mp #33 #92]: #88
-#97 := [not-or-elim #93]: #71
-#94 := [not-or-elim #93]: #51
-#95 := (not #66)
-#96 := [not-or-elim #93]: #95
-[th-lemma #96 #94 #97]: false
-unsat
-f1064899dd3b2b344c829f77d7843d7c713a00e1 210 0
+4d6ae030a899d6a058af8429edba2402745a216b 210 0
 #2 := false
 decl f11 :: S5
 #16 := f11
@@ -50394,93 +50758,7 @@
 #321 := [def-axiom]: #234
 [unit-resolution #321 #597 #608 #332]: false
 unsat
-eb502ccd121bb8afc79076bd6ce9fee84004caac 85 0
-#2 := false
-#52 := 0::real
-decl f4 :: real
-#9 := f4
-#50 := -1::real
-#72 := (* -1::real f4)
-decl f3 :: real
-#8 := f3
-#73 := (+ f3 #72)
-#71 := (>= #73 0::real)
-#70 := (not #71)
-decl f5 :: real
-#11 := f5
-#54 := (* -1::real f5)
-#67 := (+ f3 #54)
-#66 := (>= #67 0::real)
-#65 := (not #66)
-#55 := (+ f4 #54)
-#53 := (>= #55 0::real)
-#89 := (or #53 #65 #71)
-#94 := (not #89)
-#13 := (< f3 f5)
-#12 := (< f4 f5)
-#14 := (implies #12 #13)
-#10 := (< f3 f4)
-#15 := (implies #10 #14)
-#16 := (not #15)
-#97 := (iff #16 #94)
-#34 := (not #12)
-#35 := (or #34 #13)
-#41 := (not #10)
-#42 := (or #41 #35)
-#47 := (not #42)
-#95 := (iff #47 #94)
-#92 := (iff #42 #89)
-#83 := (or #53 #65)
-#86 := (or #71 #83)
-#90 := (iff #86 #89)
-#91 := [rewrite]: #90
-#87 := (iff #42 #86)
-#84 := (iff #35 #83)
-#68 := (iff #13 #65)
-#69 := [rewrite]: #68
-#63 := (iff #34 #53)
-#51 := (not #53)
-#58 := (not #51)
-#61 := (iff #58 #53)
-#62 := [rewrite]: #61
-#59 := (iff #34 #58)
-#56 := (iff #12 #51)
-#57 := [rewrite]: #56
-#60 := [monotonicity #57]: #59
-#64 := [trans #60 #62]: #63
-#85 := [monotonicity #64 #69]: #84
-#81 := (iff #41 #71)
-#76 := (not #70)
-#79 := (iff #76 #71)
-#80 := [rewrite]: #79
-#77 := (iff #41 #76)
-#74 := (iff #10 #70)
-#75 := [rewrite]: #74
-#78 := [monotonicity #75]: #77
-#82 := [trans #78 #80]: #81
-#88 := [monotonicity #82 #85]: #87
-#93 := [trans #88 #91]: #92
-#96 := [monotonicity #93]: #95
-#48 := (iff #16 #47)
-#45 := (iff #15 #42)
-#38 := (implies #10 #35)
-#43 := (iff #38 #42)
-#44 := [rewrite]: #43
-#39 := (iff #15 #38)
-#36 := (iff #14 #35)
-#37 := [rewrite]: #36
-#40 := [monotonicity #37]: #39
-#46 := [trans #40 #44]: #45
-#49 := [monotonicity #46]: #48
-#98 := [trans #49 #96]: #97
-#33 := [asserted]: #16
-#99 := [mp #33 #98]: #94
-#102 := [not-or-elim #99]: #70
-#100 := [not-or-elim #99]: #51
-#101 := [not-or-elim #99]: #66
-[th-lemma #101 #100 #102]: false
-unsat
-53a01e2ecfe5fcdece4d167dc97ab57afbf5b665 144 0
+003bb45597da3e4b1800d61bab0b4e1d95dbdf99 144 0
 #2 := false
 decl f11 :: S3
 #16 := f11
@@ -50625,77 +50903,7 @@
 #432 := [unit-resolution #232 #431]: #301
 [unit-resolution #432 #433]: false
 unsat
-e52bb6d4fc18c5b1c7dd7b024810982ae6751f44 69 0
-#2 := false
-#44 := 0::real
-decl f5 :: real
-#11 := f5
-#42 := -1::real
-#51 := (* -1::real f5)
-decl f3 :: real
-#8 := f3
-#62 := (+ f3 #51)
-#63 := (<= #62 0::real)
-#64 := (not #63)
-decl f4 :: real
-#9 := f4
-#52 := (+ f4 #51)
-#50 := (>= #52 0::real)
-#53 := (not #50)
-#46 := (* -1::real f4)
-#47 := (+ f3 #46)
-#45 := (>= #47 0::real)
-#43 := (not #45)
-#56 := (and #43 #53)
-#59 := (not #56)
-#74 := (or #59 #63)
-#77 := (not #74)
-#14 := (< f5 f3)
-#15 := (not #14)
-#12 := (< f4 f5)
-#10 := (< f3 f4)
-#13 := (and #10 #12)
-#16 := (implies #13 #15)
-#17 := (not #16)
-#80 := (iff #17 #77)
-#35 := (not #13)
-#36 := (or #35 #15)
-#39 := (not #36)
-#78 := (iff #39 #77)
-#75 := (iff #36 #74)
-#72 := (iff #15 #63)
-#67 := (not #64)
-#70 := (iff #67 #63)
-#71 := [rewrite]: #70
-#68 := (iff #15 #67)
-#65 := (iff #14 #64)
-#66 := [rewrite]: #65
-#69 := [monotonicity #66]: #68
-#73 := [trans #69 #71]: #72
-#60 := (iff #35 #59)
-#57 := (iff #13 #56)
-#54 := (iff #12 #53)
-#55 := [rewrite]: #54
-#48 := (iff #10 #43)
-#49 := [rewrite]: #48
-#58 := [monotonicity #49 #55]: #57
-#61 := [monotonicity #58]: #60
-#76 := [monotonicity #61 #73]: #75
-#79 := [monotonicity #76]: #78
-#40 := (iff #17 #39)
-#37 := (iff #16 #36)
-#38 := [rewrite]: #37
-#41 := [monotonicity #38]: #40
-#81 := [trans #41 #79]: #80
-#34 := [asserted]: #17
-#82 := [mp #34 #81]: #77
-#86 := [not-or-elim #82]: #64
-#83 := [not-or-elim #82]: #56
-#84 := [and-elim #83]: #43
-#85 := [and-elim #83]: #53
-[th-lemma #85 #84 #86]: false
-unsat
-64328afb82eb17196ff1275de441cf37fad9bf4a 144 0
+31b5dadcaa7e7a3e69848f4809b04b844a6e8b56 144 0
 #2 := false
 decl f12 :: S3
 #17 := f12
@@ -50840,7 +51048,7 @@
 #432 := [unit-resolution #232 #431]: #301
 [unit-resolution #432 #433]: false
 unsat
-ca6daa73749cee4dc8b11e40a0201ba82325b005 103 0
+87752431ac232f724f4719a9dd746411dc8e0334 103 0
 #2 := false
 decl f6 :: S3
 #9 := f6
@@ -50944,7 +51152,7 @@
 #205 := [unit-resolution #171 #544]: #240
 [unit-resolution #205 #206]: false
 unsat
-0273cec9a3b2d721f13b81a5f7cb9a26c53610a3 111 0
+de0be68d62a4eef58cc08ca9b43494f9219ad474 111 0
 #2 := false
 decl f9 :: (-> S2 S3)
 decl f7 :: S2
@@ -51056,7 +51264,7 @@
 #66 := [not-or-elim #63]: #65
 [unit-resolution #66 #563]: false
 unsat
-59ef7a0dffeb5563b747aceef193a9e3d4f5d32d 117 0
+c992b1f07c8690f317843ec597ee745825b8d2b2 117 0
 #2 := false
 decl f6 :: (-> S3 S3 S2)
 decl f3 :: (-> S2 S3)
@@ -51174,7 +51382,7 @@
 #545 := [unit-resolution #170 #207]: #239
 [unit-resolution #545 #551]: false
 unsat
-13c013ce9b60f948f65ac05c9f64640c23a178bb 50 0
+6cc8ed53fe294f253a561b590f79ef0e0271d1c2 50 0
 #2 := false
 decl f3 :: (-> S2 S3 S4 S3 S4)
 #15 := (:var 1 S3)
@@ -51225,7 +51433,7 @@
 #155 := [quant-inst]: #241
 [unit-resolution #155 #55 #577]: false
 unsat
-1359b66e79f681fc73d3d69747fa44ddaf8fc00b 85 0
+d2a264ef3e43d9b7d009fa11ec172212f0bd6279 85 0
 #2 := false
 decl f8 :: (-> S3 S2 S4)
 decl f4 :: S2
@@ -51311,7 +51519,7 @@
 #241 := [quant-inst]: #250
 [unit-resolution #241 #592 #162]: false
 unsat
-85295dd87d7cebdd794ace783556d1b1d280619a 158 0
+13cd2a6082f8fe0947641619628aa7b793fa5eb0 158 0
 #2 := false
 decl f5 :: (-> S3 S2 S4)
 decl f3 :: S2
@@ -51470,7 +51678,7 @@
 #73 := [not-or-elim #72]: #71
 [unit-resolution #73 #578]: false
 unsat
-c0ecafa6d8b41ce5e1c9ab26d4838fdbd8c57d3d 72 0
+c3052f64b6a04dd1c8c5afaf04fa56dc411bc19c 72 0
 #2 := false
 decl f6 :: (-> S3 S2 S4 S3)
 #23 := (:var 0 S4)
@@ -51543,7 +51751,7 @@
 #265 := [quant-inst]: #264
 [unit-resolution #265 #73 #598]: false
 unsat
-dc41ee203ffc081fb445ff0000ddb4be9134fe60 81 0
+2dd7e2139c57ac13e3558d5a7c2ee7d6b5a8a254 81 0
 #2 := false
 decl f5 :: (-> S3 S2 S4)
 decl f3 :: S2
@@ -51625,7 +51833,7 @@
 #73 := [not-or-elim #72]: #71
 [unit-resolution #73 #230]: false
 unsat
-ae5be72196cf39e4c9692c5e427b5fb33965ed8f 187 0
+d43a3750252b86094159a0508ea79def7392e6da 187 0
 #2 := false
 decl f6 :: (-> S3 S2 S4)
 decl f5 :: S2
@@ -51813,7 +52021,7 @@
 #67 := [not-or-elim #66]: #65
 [unit-resolution #67 #496]: false
 unsat
-2f659b6a47ef547b90b00946538367bd9efdeee0 31 0
+d9c2fc1edc378ccb88037e9ae16a84f1542acb6f 31 0
 #2 := false
 decl f1 :: S1
 #4 := f1
@@ -51845,7 +52053,7 @@
 #62 := [asserted]: #12
 [mp #62 #79]: false
 unsat
-3329505aa680c82f608d61377279461a25778b94 58 0
+3df517f77a8847ec580730167b0937b5bdc9f32c 58 0
 #2 := false
 decl f3 :: (-> S2 S1)
 #13 := (:var 0 S2)
@@ -51904,7 +52112,7 @@
 #247 := [quant-inst]: #333
 [unit-resolution #247 #78 #667]: false
 unsat
-f68aac6d5243061ae9a95773075ebafc014b27eb 46 0
+1b877eab526e97351cb862fc0b9910541c0a7e4b 46 0
 #2 := false
 decl f3 :: (-> S2 S1)
 #17 := (:var 0 S2)
@@ -51951,7 +52159,7 @@
 #239 := [quant-inst]: #325
 [unit-resolution #239 #70 #666]: false
 unsat
-fca05a20824fb9ae1ae5b948aca1cdacb4ec3a3d 119 0
+9cf9419534c29ea2783ed62db9d7dac14034c845 119 0
 #2 := false
 decl f3 :: (-> S2 S2 S3 S1)
 decl f6 :: S3
@@ -52071,7 +52279,7 @@
 #320 := [unit-resolution #678 #329]: #319
 [unit-resolution #320 #318 #317]: false
 unsat
-3c528fd543bec1b0abe827803d678f1293dd7fb0 154 0
+db4d005483a981015636afecf4abdbe3eecdc540 154 0
 #2 := false
 decl f7 :: (-> S2 S3 S1)
 decl f6 :: S3
@@ -52226,7 +52434,7 @@
 #289 := [quant-inst]: #651
 [unit-resolution #289 #675 #294]: false
 unsat
-156b5bbccd88bf00c5d19d456c8d5c95867a3e79 128 0
+01e84163a16924ffd8e379badcb76f97c88eb8e1 128 0
 #2 := false
 decl f8 :: (-> S2 S3 S1)
 decl f6 :: S3
@@ -52355,7 +52563,7 @@
 #630 := [quant-inst]: #635
 [unit-resolution #630 #667 #637]: false
 unsat
-e9501ad699b29b1a0c65a378dbd0d68cdbadce66 146 0
+db25ac5ac3d60bebb553d3f6b343562b94a56253 146 0
 #2 := false
 decl f3 :: (-> S2 S2 S3 S1)
 decl f6 :: S3
@@ -52502,7 +52710,7 @@
 #632 := [unit-resolution #661 #642]: #635
 [unit-resolution #632 #631 #633]: false
 unsat
-60cdff9edbb9caaee078a8376311922be886cc87 121 0
+9033d495190da4d9088c43145881d91ed9687495 121 0
 #2 := false
 decl f3 :: (-> S2 S2 S3 S1)
 decl f5 :: S3
@@ -52624,7 +52832,7 @@
 #650 := [unit-resolution #323 #308]: #649
 [unit-resolution #650 #648 #647]: false
 unsat
-e7ec03d0541acce6c328f68f87047aa506800ef8 258 0
+d24feed52b94afbe0c3926c41e67eb725f57798d 258 0
 #2 := false
 decl f3 :: (-> S2 S3 S1)
 decl f8 :: S3
@@ -52883,7 +53091,7 @@
 #585 := [def-axiom]: #575
 [unit-resolution #585 #539 #546 #581]: false
 unsat
-0ff3bde207876a1e2bbf9f523fbb88a3af58b75b 153 0
+5b1f4025c8afe3b7c1a0fb50a6e0140ed51756ba 153 0
 #2 := false
 decl f7 :: (-> S2 S3 S1)
 decl f6 :: S3
@@ -53037,7 +53245,7 @@
 #339 := [unit-resolution #692 #689]: #338
 [unit-resolution #339 #337 #674]: false
 unsat
-af9d88d34e35638de5cad45a5fcb0943b421dc7f 141 0
+14e5476f4f3857f931521140926538df69a7900f 141 0
 #2 := false
 decl f7 :: (-> S2 S3 S1)
 decl f6 :: S3
@@ -53179,7 +53387,7 @@
 #312 := [quant-inst]: #652
 [unit-resolution #312 #668 #637]: false
 unsat
-1fe992aa332936da32d8835a880128a83f76f965 165 0
+17d0e6f19e122ea3db2aaeed6f9623f5ae4eb5b6 165 0
 #2 := false
 decl f7 :: (-> S2 S3 S1)
 decl f6 :: S3
@@ -53345,7 +53553,7 @@
 #665 := [def-axiom]: #664
 [unit-resolution #665 #655 #654]: false
 unsat
-71f085d2a9c00dbd0a999021926a678b17395642 164 0
+8c5227e6af822f9c7c2a140803c4af49bcf24bcb 164 0
 #2 := false
 decl f3 :: (-> S2 S2 S3 S1)
 decl f6 :: S3
@@ -53510,7 +53718,7 @@
 #626 := [unit-resolution #290 #354]: #625
 [unit-resolution #626 #636 #631]: false
 unsat
-4a2ec1009e57974d1e4c4125b313542f141f334f 142 0
+8156155a61a8ea554dc6bef7e96703a3d13ceb5f 142 0
 #2 := false
 decl f3 :: (-> S2 S2 S3 S1)
 decl f5 :: S3
@@ -53653,7 +53861,7 @@
 #655 := [unit-resolution #309 #645]: #295
 [unit-resolution #655 #294 #654]: false
 unsat
-9a719bb56314c3b4bab4bac5b0adde4b491a8497 279 0
+f0bdfff0ca207ce8dcf7aafa0f5fd49080a4003e 279 0
 #2 := false
 decl f3 :: (-> S2 S3 S1)
 decl f8 :: S3
@@ -53933,7 +54141,7 @@
 #580 := [def-axiom]: #579
 [unit-resolution #580 #526 #543 #572]: false
 unsat
-d1f4c7d9bd1f639df15a28175017a4537fb9637f 18 0
+0166ab8a0d2c326b9fbcb4a04b09f2f360525618 18 0
 #2 := false
 decl f3 :: S2
 #8 := f3
@@ -53952,211 +54160,3 @@
 #60 := [asserted]: #10
 [mp #60 #69]: false
 unsat
-c08976b5943421000363d731a8009a5d8459d100 93 0
-#2 := false
-decl f5 :: (-> S2 S1)
-decl f6 :: S2
-#16 := f6
-#20 := (f5 f6)
-decl f1 :: S1
-#4 := f1
-#65 := (= f1 #20)
-#84 := (not #65)
-decl f3 :: (-> S2 S3 S1)
-decl f4 :: S3
-#9 := f4
-#17 := (f3 f6 f4)
-#59 := (= f1 #17)
-#8 := (:var 0 S2)
-#12 := (f5 #8)
-#44 := (= f1 #12)
-#10 := (f3 #8 f4)
-#41 := (= f1 #10)
-#50 := (not #41)
-#51 := (or #50 #44)
-#56 := (forall (vars (?v0 S2)) #51)
-#62 := (and #56 #59)
-#71 := (not #62)
-#72 := (or #71 #65)
-#77 := (not #72)
-#21 := (= #20 f1)
-#18 := (= #17 f1)
-#13 := (= #12 f1)
-#11 := (= #10 f1)
-#14 := (implies #11 #13)
-#15 := (forall (vars (?v0 S2)) #14)
-#19 := (and #15 #18)
-#22 := (implies #19 #21)
-#23 := (not #22)
-#78 := (iff #23 #77)
-#75 := (iff #22 #72)
-#68 := (implies #62 #65)
-#73 := (iff #68 #72)
-#74 := [rewrite]: #73
-#69 := (iff #22 #68)
-#66 := (iff #21 #65)
-#67 := [rewrite]: #66
-#63 := (iff #19 #62)
-#60 := (iff #18 #59)
-#61 := [rewrite]: #60
-#57 := (iff #15 #56)
-#54 := (iff #14 #51)
-#47 := (implies #41 #44)
-#52 := (iff #47 #51)
-#53 := [rewrite]: #52
-#48 := (iff #14 #47)
-#45 := (iff #13 #44)
-#46 := [rewrite]: #45
-#42 := (iff #11 #41)
-#43 := [rewrite]: #42
-#49 := [monotonicity #43 #46]: #48
-#55 := [trans #49 #53]: #54
-#58 := [quant-intro #55]: #57
-#64 := [monotonicity #58 #61]: #63
-#70 := [monotonicity #64 #67]: #69
-#76 := [trans #70 #74]: #75
-#79 := [monotonicity #76]: #78
-#40 := [asserted]: #23
-#82 := [mp #40 #79]: #77
-#85 := [not-or-elim #82]: #84
-#80 := [not-or-elim #82]: #62
-#83 := [and-elim #80]: #59
-#572 := (pattern #12)
-#571 := (pattern #10)
-#573 := (forall (vars (?v0 S2)) (:pat #571 #572) #51)
-#576 := (iff #56 #573)
-#574 := (iff #51 #51)
-#575 := [refl]: #574
-#577 := [quant-intro #575]: #576
-#97 := (~ #56 #56)
-#95 := (~ #51 #51)
-#96 := [refl]: #95
-#98 := [nnf-pos #96]: #97
-#81 := [and-elim #80]: #56
-#87 := [mp~ #81 #98]: #56
-#578 := [mp #87 #577]: #573
-#155 := (not #59)
-#157 := (not #573)
-#244 := (or #157 #155 #65)
-#242 := (or #155 #65)
-#235 := (or #157 #242)
-#247 := (iff #235 #244)
-#175 := [rewrite]: #247
-#246 := [quant-inst]: #235
-#248 := [mp #246 #175]: #244
-[unit-resolution #248 #578 #83 #85]: false
-unsat
-9970f178c30353ca679caf02fa917fbf2a5e19ef 113 0
-#2 := false
-decl f3 :: (-> S2 S3 S1)
-decl f4 :: S3
-#9 := f4
-decl f6 :: S2
-#16 := f6
-#19 := (f3 f6 f4)
-decl f1 :: S1
-#4 := f1
-#57 := (= f1 #19)
-decl f5 :: (-> S2 S1)
-#17 := (f5 f6)
-#54 := (= f1 #17)
-#60 := (and #54 #57)
-#63 := (not #60)
-#8 := (:var 0 S2)
-#12 := (f5 #8)
-#45 := (= f1 #12)
-#10 := (f3 #8 f4)
-#42 := (= f1 #10)
-#48 := (and #42 #45)
-#51 := (exists (vars (?v0 S2)) #48)
-#66 := (or #51 #63)
-#69 := (not #66)
-#20 := (= #19 f1)
-#18 := (= #17 f1)
-#21 := (and #18 #20)
-#22 := (not #21)
-#13 := (= #12 f1)
-#11 := (= #10 f1)
-#14 := (and #11 #13)
-#15 := (exists (vars (?v0 S2)) #14)
-#23 := (or #15 #22)
-#24 := (not #23)
-#70 := (iff #24 #69)
-#67 := (iff #23 #66)
-#64 := (iff #22 #63)
-#61 := (iff #21 #60)
-#58 := (iff #20 #57)
-#59 := [rewrite]: #58
-#55 := (iff #18 #54)
-#56 := [rewrite]: #55
-#62 := [monotonicity #56 #59]: #61
-#65 := [monotonicity #62]: #64
-#52 := (iff #15 #51)
-#49 := (iff #14 #48)
-#46 := (iff #13 #45)
-#47 := [rewrite]: #46
-#43 := (iff #11 #42)
-#44 := [rewrite]: #43
-#50 := [monotonicity #44 #47]: #49
-#53 := [quant-intro #50]: #52
-#68 := [monotonicity #53 #65]: #67
-#71 := [monotonicity #68]: #70
-#41 := [asserted]: #24
-#74 := [mp #41 #71]: #69
-#75 := [not-or-elim #74]: #60
-#77 := [and-elim #75]: #57
-#76 := [and-elim #75]: #54
-#585 := (pattern #12)
-#584 := (pattern #10)
-#93 := (not #45)
-#92 := (not #42)
-#94 := (or #92 #93)
-#586 := (forall (vars (?v0 S2)) (:pat #584 #585) #94)
-#101 := (forall (vars (?v0 S2)) #94)
-#589 := (iff #101 #586)
-#587 := (iff #94 #94)
-#588 := [refl]: #587
-#590 := [quant-intro #588]: #589
-#87 := (not #48)
-#90 := (forall (vars (?v0 S2)) #87)
-#102 := (iff #90 #101)
-#99 := (iff #87 #94)
-#95 := (not #94)
-#83 := (not #95)
-#97 := (iff #83 #94)
-#98 := [rewrite]: #97
-#84 := (iff #87 #83)
-#85 := (iff #48 #95)
-#86 := [rewrite]: #85
-#96 := [monotonicity #86]: #84
-#100 := [trans #96 #98]: #99
-#103 := [quant-intro #100]: #102
-#72 := (not #51)
-#79 := (~ #72 #90)
-#88 := (~ #87 #87)
-#89 := [refl]: #88
-#80 := [nnf-neg #89]: #79
-#73 := [not-or-elim #74]: #72
-#91 := [mp~ #73 #80]: #90
-#104 := [mp #91 #103]: #101
-#591 := [mp #104 #590]: #586
-#255 := (not #57)
-#168 := (not #54)
-#248 := (not #586)
-#259 := (or #248 #168 #255)
-#169 := (or #255 #168)
-#260 := (or #248 #169)
-#578 := (iff #260 #259)
-#256 := (or #168 #255)
-#261 := (or #248 #256)
-#241 := (iff #261 #259)
-#576 := [rewrite]: #241
-#258 := (iff #260 #261)
-#170 := (iff #169 #256)
-#257 := [rewrite]: #170
-#262 := [monotonicity #257]: #258
-#235 := [trans #262 #576]: #578
-#188 := [quant-inst]: #260
-#365 := [mp #188 #235]: #259
-[unit-resolution #365 #591 #76 #77]: false
-unsat
--- a/src/HOL/SMT_Examples/SMT_Tests.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/SMT_Examples/SMT_Tests.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -8,7 +8,7 @@
 imports Complex_Main
 begin
 
-declare [[smt_solver=z3, z3_proofs=true]]
+declare [[smt_solver=z3, smt_oracle=false]]
 declare [[smt_certificates="~~/src/HOL/SMT_Examples/SMT_Tests.certs"]]
 declare [[smt_fixed=true]]
 
@@ -145,7 +145,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 [[z3_proofs=false, z3_options="AUTO_CONFIG=false SATURATE=true"]]
+  using [[smt_oracle=true, z3_options="AUTO_CONFIG=false SATURATE=true"]]
   by smt+
 
 lemma
@@ -165,7 +165,7 @@
   "(\<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)"
-  using [[z3_proofs=false]]
+  using [[smt_oracle=true]]
   by smt+
 
 lemma
@@ -622,6 +622,7 @@
   "p1 = (x, y) \<and> p2 = (y, x) \<longrightarrow> fst p1 = snd p2"
   "(fst (x, y) = snd (x, y)) = (x = y)"
   "(fst p = snd p) = (p = (snd p, fst p))"
+  using [[smt_trace=true]]
   by smt+
 
 
--- a/src/HOL/SMT_Examples/SMT_Word_Examples.certs	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/SMT_Examples/SMT_Word_Examples.certs	Wed Oct 27 11:11:35 2010 -0700
@@ -44,7 +44,7 @@
 unsat
 833f259a7dd52b3dddfc285b647a432c0daf13c1 1 0
 unsat
-234d4999f188127d090de762711164119fba49e9 1 0
+7d5a71ba893f865044699426792a024fe1612235 1 0
 unsat
 6adbd17723088c30f1d1f0f07e1e4df2076a51b2 1 0
 unsat
--- a/src/HOL/SMT_Examples/SMT_Word_Examples.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/SMT_Examples/SMT_Word_Examples.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -8,7 +8,7 @@
 imports Word
 begin
 
-declare [[smt_solver=z3, z3_proofs=false]]
+declare [[smt_solver=z3, smt_oracle=true]]
 declare [[smt_certificates="~~/src/HOL/SMT_Examples/SMT_Word_Examples.certs"]]
 declare [[smt_fixed=true]]
 
--- a/src/HOL/Sledgehammer.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Sledgehammer.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -7,7 +7,7 @@
 header {* Sledgehammer: Isabelle--ATP Linkup *}
 
 theory Sledgehammer
-imports ATP
+imports ATP SMT
 uses "Tools/Sledgehammer/sledgehammer_util.ML"
      "Tools/Sledgehammer/sledgehammer_filter.ML"
      "Tools/Sledgehammer/sledgehammer_atp_translate.ML"
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Function/partial_function.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -0,0 +1,238 @@
+(*  Title:      HOL/Tools/Function/partial_function.ML
+    Author:     Alexander Krauss, TU Muenchen
+
+Partial function definitions based on least fixed points in ccpos.
+*)
+
+signature PARTIAL_FUNCTION =
+sig
+  val setup: theory -> theory
+  val init: term -> term -> thm -> declaration
+
+  val add_partial_function: string -> (binding * typ option * mixfix) list ->
+    Attrib.binding * term -> local_theory -> local_theory
+
+  val add_partial_function_cmd: string -> (binding * string option * mixfix) list ->
+    Attrib.binding * string -> local_theory -> local_theory
+end;
+
+
+structure Partial_Function: PARTIAL_FUNCTION =
+struct
+
+(*** Context Data ***)
+
+structure Modes = Generic_Data
+(
+  type T = ((term * term) * thm) Symtab.table;
+  val empty = Symtab.empty;
+  val extend = I;
+  fun merge (a, b) = Symtab.merge (K true) (a, b);
+)
+
+fun init fixp mono fixp_eq phi =
+  let
+    val term = Morphism.term phi;
+    val data' = ((term fixp, term mono), Morphism.thm phi fixp_eq);
+    val mode = (* extract mode identifier from morphism prefix! *)
+      Binding.prefix_of (Morphism.binding phi (Binding.empty))
+      |> map fst |> space_implode ".";
+  in
+    if mode = "" then I
+    else Modes.map (Symtab.update (mode, data'))
+  end
+
+val known_modes = Symtab.keys o Modes.get o Context.Proof;
+val lookup_mode = Symtab.lookup o Modes.get o Context.Proof;
+
+
+structure Mono_Rules = Named_Thms
+(
+  val name = "partial_function_mono";
+  val description = "monotonicity rules for partial function definitions";
+);
+
+
+(*** Automated monotonicity proofs ***)
+
+fun strip_cases ctac = ctac #> Seq.map snd;
+
+(*rewrite conclusion with k-th assumtion*)
+fun rewrite_with_asm_tac ctxt k =
+  Subgoal.FOCUS (fn {context=ctxt', prems, ...} =>
+    Local_Defs.unfold_tac ctxt' [nth prems k]) ctxt;
+
+fun dest_case thy t =
+  case strip_comb t of
+    (Const (case_comb, _), args) =>
+      (case Datatype.info_of_case thy case_comb of
+         NONE => NONE
+       | SOME {case_rewrites, ...} =>
+           let
+             val lhs = prop_of (hd case_rewrites)
+               |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> fst;
+             val arity = length (snd (strip_comb lhs));
+             val conv = funpow (length args - arity) Conv.fun_conv
+               (Conv.rewrs_conv (map mk_meta_eq case_rewrites));
+           in
+             SOME (nth args (arity - 1), conv)
+           end)
+  | _ => NONE;
+
+(*split on case expressions*)
+val split_cases_tac = Subgoal.FOCUS_PARAMS (fn {context=ctxt, ...} =>
+  SUBGOAL (fn (t, i) => case t of
+    _ $ (_ $ Abs (_, _, body)) =>
+      (case dest_case (ProofContext.theory_of ctxt) body of
+         NONE => no_tac
+       | SOME (arg, conv) =>
+           let open Conv in
+              if not (null (loose_bnos arg)) then no_tac
+              else ((DETERM o strip_cases o Induct.cases_tac ctxt false [[SOME arg]] NONE [])
+                THEN_ALL_NEW (rewrite_with_asm_tac ctxt 0)
+                THEN_ALL_NEW etac @{thm thin_rl}
+                THEN_ALL_NEW (CONVERSION
+                  (params_conv ~1 (fn ctxt' =>
+                    arg_conv (arg_conv (abs_conv (K conv) ctxt'))) ctxt))) i
+           end)
+  | _ => no_tac) 1);
+
+(*monotonicity proof: apply rules + split case expressions*)
+fun mono_tac ctxt =
+  K (Local_Defs.unfold_tac ctxt [@{thm curry_def}])
+  THEN' (TRY o REPEAT_ALL_NEW
+   (resolve_tac (Mono_Rules.get ctxt)
+     ORELSE' split_cases_tac ctxt));
+
+
+(*** Auxiliary functions ***)
+
+(*positional instantiation with computed type substitution.
+  internal version of  attribute "[of s t u]".*)
+fun cterm_instantiate' cts thm =
+  let
+    val thy = Thm.theory_of_thm thm;
+    val vs = rev (Term.add_vars (prop_of thm) [])
+      |> map (Thm.cterm_of thy o Var);
+  in
+    cterm_instantiate (zip_options vs cts) thm
+  end;
+
+(*Returns t $ u, but instantiates the type of t to make the
+application type correct*)
+fun apply_inst ctxt t u =
+  let
+    val thy = ProofContext.theory_of ctxt;
+    val T = domain_type (fastype_of t);
+    val T' = fastype_of u;
+    val subst = Type.typ_match (Sign.tsig_of thy) (T, T') Vartab.empty
+      handle Type.TYPE_MATCH => raise TYPE ("apply_inst", [T, T'], [t, u])
+  in
+    map_types (Envir.norm_type subst) t $ u
+  end;
+
+fun head_conv cv ct =
+  if can Thm.dest_comb ct then Conv.fun_conv (head_conv cv) ct else cv ct;
+
+
+(*** currying transformation ***)
+
+fun curry_const (A, B, C) =
+  Const (@{const_name Product_Type.curry},
+    [HOLogic.mk_prodT (A, B) --> C, A, B] ---> C);
+
+fun mk_curry f =
+  case fastype_of f of
+    Type ("fun", [Type (_, [S, T]), U]) =>
+      curry_const (S, T, U) $ f
+  | T => raise TYPE ("mk_curry", [T], [f]);
+
+(* iterated versions. Nonstandard left-nested tuples arise naturally
+from "split o split o split"*)
+fun curry_n arity = funpow (arity - 1) mk_curry;
+fun uncurry_n arity = funpow (arity - 1) HOLogic.mk_split;
+
+val curry_uncurry_ss = HOL_basic_ss addsimps
+  [@{thm Product_Type.curry_split}, @{thm Product_Type.split_curry}]
+
+
+(*** partial_function definition ***)
+
+fun gen_add_partial_function prep mode fixes_raw eqn_raw lthy =
+  let
+    val ((fixp, mono), fixp_eq) = the (lookup_mode lthy mode)
+      handle Option.Option => error (cat_lines ["Unknown mode " ^ quote mode ^ ".",
+        "Known modes are " ^ commas_quote (known_modes lthy) ^ "."]);
+
+    val ((fixes, [(eq_abinding, eqn)]), _) = prep fixes_raw [eqn_raw] lthy;
+    val (_, _, plain_eqn) = Function_Lib.dest_all_all_ctx lthy eqn;
+
+    val ((f_binding, fT), mixfix) = the_single fixes;
+    val fname = Binding.name_of f_binding;
+
+    val cert = cterm_of (ProofContext.theory_of lthy);
+    val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop plain_eqn);
+    val (head, args) = strip_comb lhs;
+    val F = fold_rev lambda (head :: args) rhs;
+
+    val arity = length args;
+    val (aTs, bTs) = chop arity (binder_types fT);
+
+    val tupleT = foldl1 HOLogic.mk_prodT aTs;
+    val fT_uc = tupleT :: bTs ---> body_type fT;
+    val f_uc = Var ((fname, 0), fT_uc);
+    val x_uc = Var (("x", 0), tupleT);
+    val uncurry = lambda head (uncurry_n arity head);
+    val curry = lambda f_uc (curry_n arity f_uc);
+
+    val F_uc =
+      lambda f_uc (uncurry_n arity (F $ curry_n arity f_uc));
+
+    val mono_goal = apply_inst lthy mono (lambda f_uc (F_uc $ f_uc $ x_uc))
+      |> HOLogic.mk_Trueprop
+      |> Logic.all x_uc;
+
+    val mono_thm = Goal.prove_internal [] (cert mono_goal)
+        (K (mono_tac lthy 1))
+      |> Thm.forall_elim (cert x_uc);
+
+    val f_def_rhs = curry_n arity (apply_inst lthy fixp F_uc);
+    val f_def_binding = Binding.conceal (Binding.name (Thm.def_name fname));
+    val ((f, (_, f_def)), lthy') = Local_Theory.define
+      ((f_binding, mixfix), ((f_def_binding, []), f_def_rhs)) lthy;
+
+    val eqn = HOLogic.mk_eq (list_comb (f, args),
+        Term.betapplys (F, f :: args))
+      |> HOLogic.mk_Trueprop;
+
+    val unfold =
+      (cterm_instantiate' (map (SOME o cert) [uncurry, F, curry]) fixp_eq
+        OF [mono_thm, f_def])
+      |> Tactic.rule_by_tactic lthy (Simplifier.simp_tac curry_uncurry_ss 1);
+
+    val rec_rule = let open Conv in
+      Goal.prove lthy' (map (fst o dest_Free) args) [] eqn (fn _ =>
+        CONVERSION ((arg_conv o arg1_conv o head_conv o rewr_conv) (mk_meta_eq unfold)) 1
+        THEN rtac @{thm refl} 1) end;
+  in
+    lthy'
+    |> Local_Theory.note (eq_abinding, [rec_rule])
+    |-> (fn (_, rec') =>
+      Spec_Rules.add Spec_Rules.Equational ([f], rec')
+      #> Local_Theory.note ((Binding.qualify true fname (Binding.name "simps"), []), rec') #> snd)
+  end;
+
+val add_partial_function = gen_add_partial_function Specification.check_spec;
+val add_partial_function_cmd = gen_add_partial_function Specification.read_spec;
+
+val mode = Parse.$$$ "(" |-- Parse.xname --| Parse.$$$ ")";
+
+val _ = Outer_Syntax.local_theory
+  "partial_function" "define partial function" Keyword.thy_decl
+  ((mode -- (Parse.fixes -- (Parse.where_ |-- Parse_Spec.spec)))
+     >> (fn (mode, (fixes, spec)) => add_partial_function_cmd mode fixes spec));
+
+
+val setup = Mono_Rules.setup;
+
+end
--- a/src/HOL/Tools/Metis/metis_reconstruct.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/Metis/metis_reconstruct.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -754,8 +754,8 @@
         |> fold Int_Pair_Graph.add_deps_acyclic depss
         |> Int_Pair_Graph.topological_order
         handle Int_Pair_Graph.CYCLES _ =>
-               error "Cannot replay Metis proof in Isabelle without axiom of \
-                     \choice."
+               error "Cannot replay Metis proof in Isabelle without \
+                     \\"Hilbert_Choice\"."
       val params0 =
         [] |> fold (Term.add_vars o snd) (map_filter I axioms)
            |> map (`(Meson_Clausify.cluster_of_zapped_var_name o fst o fst))
@@ -796,6 +796,9 @@
                        THEN rotate_tac (rotation_for_subgoal i) i
 (*                       THEN PRIMITIVE (unify_first_prem_with_concl thy i) ###*)
                        THEN assume_tac i)))
+      handle ERROR _ =>
+             error ("Cannot replay Metis proof in Isabelle:\n\
+                    \Error when discharging Skolem assumptions.")
     end
 
 val setup = trace_setup
--- a/src/HOL/Tools/Metis/metis_tactics.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/Metis/metis_tactics.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -70,7 +70,7 @@
       val _ = trace_msg ctxt (fn () => "THEOREM CLAUSES")
       val _ = app (app (fn th => trace_msg ctxt (fn () => Display.string_of_thm ctxt th))) thss
       val (mode, {axioms, tfrees, old_skolems}) =
-        build_logic_map mode ctxt type_lits cls thss
+        prepare_metis_problem mode ctxt type_lits cls thss
       val _ = if null tfrees then ()
               else (trace_msg ctxt (fn () => "TFREE CLAUSES");
                     app (fn TyLitFree ((s, _), (s', _)) =>
--- a/src/HOL/Tools/Metis/metis_translate.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/Metis/metis_translate.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -31,7 +31,7 @@
   datatype fol_literal = FOLLiteral of bool * combterm
 
   datatype mode = FO | HO | FT
-  type logic_map =
+  type metis_problem =
     {axioms: (Metis_Thm.thm * thm) list,
      tfrees: type_literal list,
      old_skolems: (string * term) list}
@@ -68,15 +68,15 @@
   val combtyp_of : combterm -> combtyp
   val strip_combterm_comb : combterm -> combterm * combterm list
   val combterm_from_term :
-    theory -> int -> (string * typ) list -> term -> combterm * typ list
+    theory -> (string * typ) list -> term -> combterm * typ list
   val reveal_old_skolem_terms : (string * term) list -> term -> term
   val tfree_classes_of_terms : term list -> string list
   val tvar_classes_of_terms : term list -> string list
   val type_consts_of_terms : theory -> term list -> string list
   val string_of_mode : mode -> string
-  val build_logic_map :
+  val prepare_metis_problem :
     mode -> Proof.context -> bool -> thm list -> thm list list
-    -> mode * logic_map
+    -> mode * metis_problem
 end
 
 structure Metis_Translate : METIS_TRANSLATE =
@@ -215,10 +215,9 @@
     (s, Sign.the_const_type thy s) |> Sign.const_typargs thy |> length
 
 fun new_skolem_var_from_const s =
-  let
-    val ss = s |> space_explode Long_Name.separator
-    val n = length ss
-  in (nth ss (n - 2), nth ss (n - 3) |> Int.fromString |> the) end
+  let val ss = s |> space_explode Long_Name.separator in
+    (nth ss (length ss - 2), 0)
+  end
 
 
 (**** Definitions and functions for FOL clauses for TPTP format output ****)
@@ -396,17 +395,18 @@
   | simple_combtype_of (TVar (x, _)) =
     CombTVar (make_schematic_type_var x, string_of_indexname x)
 
-fun new_skolem_const_name th_no s num_T_args =
-  [new_skolem_const_prefix, string_of_int th_no, s, string_of_int num_T_args]
+fun new_skolem_const_name s num_T_args =
+  [new_skolem_const_prefix, s, string_of_int num_T_args]
   |> space_implode Long_Name.separator
 
-(* Converts a term (with combinators) into a combterm. Also accummulates sort
+(* Converts a term (with combinators) into a combterm. Also accumulates sort
    infomation. *)
-fun combterm_from_term thy th_no bs (P $ Q) =
-      let val (P', tsP) = combterm_from_term thy th_no bs P
-          val (Q', tsQ) = combterm_from_term thy th_no bs Q
-      in  (CombApp (P', Q'), union (op =) tsP tsQ)  end
-  | combterm_from_term thy _ _ (Const (c, T)) =
+fun combterm_from_term thy bs (P $ Q) =
+      let
+        val (P', tsP) = combterm_from_term thy bs P
+        val (Q', tsQ) = combterm_from_term thy bs Q
+      in (CombApp (P', Q'), union (op =) tsP tsQ) end
+  | combterm_from_term thy _ (Const (c, T)) =
       let
         val (tp, ts) = combtype_of T
         val tvar_list =
@@ -417,43 +417,42 @@
           |> map simple_combtype_of
         val c' = CombConst (`make_fixed_const c, tp, tvar_list)
       in  (c',ts)  end
-  | combterm_from_term _ _ _ (Free (v, T)) =
+  | combterm_from_term _ _ (Free (v, T)) =
       let val (tp, ts) = combtype_of T
           val v' = CombConst (`make_fixed_var v, tp, [])
       in  (v',ts)  end
-  | combterm_from_term _ th_no _ (Var (v as (s, _), T)) =
+  | combterm_from_term _ _ (Var (v as (s, _), T)) =
     let
       val (tp, ts) = combtype_of T
       val v' =
         if String.isPrefix Meson_Clausify.new_skolem_var_prefix s then
           let
             val tys = T |> strip_type |> swap |> op ::
-            val s' = new_skolem_const_name th_no s (length tys)
+            val s' = new_skolem_const_name s (length tys)
           in
             CombConst (`make_fixed_const s', tp, map simple_combtype_of tys)
           end
         else
           CombVar ((make_schematic_var v, string_of_indexname v), tp)
     in (v', ts) end
-  | combterm_from_term _ _ bs (Bound j) =
+  | combterm_from_term _ bs (Bound j) =
       let
         val (s, T) = nth bs j
         val (tp, ts) = combtype_of T
         val v' = CombConst (`make_bound_var s, tp, [])
       in (v', ts) end
-  | combterm_from_term _ _ _ (Abs _) = raise Fail "HOL clause: Abs"
+  | combterm_from_term _ _ (Abs _) = raise Fail "HOL clause: Abs"
 
-fun predicate_of thy th_no ((@{const Not} $ P), pos) =
-    predicate_of thy th_no (P, not pos)
-  | predicate_of thy th_no (t, pos) =
-    (combterm_from_term thy th_no [] (Envir.eta_contract t), pos)
+fun predicate_of thy ((@{const Not} $ P), pos) = predicate_of thy (P, not pos)
+  | predicate_of thy (t, pos) =
+    (combterm_from_term thy [] (Envir.eta_contract t), pos)
 
-fun literals_of_term1 args thy th_no (@{const Trueprop} $ P) =
-    literals_of_term1 args thy th_no P
-  | literals_of_term1 args thy th_no (@{const HOL.disj} $ P $ Q) =
-    literals_of_term1 (literals_of_term1 args thy th_no P) thy th_no Q
-  | literals_of_term1 (lits, ts) thy th_no P =
-    let val ((pred, ts'), pol) = predicate_of thy th_no (P, true) in
+fun literals_of_term1 args thy (@{const Trueprop} $ P) =
+    literals_of_term1 args thy P
+  | literals_of_term1 args thy (@{const HOL.disj} $ P $ Q) =
+    literals_of_term1 (literals_of_term1 args thy P) thy Q
+  | literals_of_term1 (lits, ts) thy P =
+    let val ((pred, ts'), pol) = predicate_of thy (P, true) in
       (FOLLiteral (pol, pred) :: lits, union (op =) ts ts')
     end
 val literals_of_term = literals_of_term1 ([], [])
@@ -607,9 +606,10 @@
             metis_lit pos "=" (map hol_term_to_fol_FT tms)
         | _ => metis_lit pos "{}" [hol_term_to_fol_FT tm])   (*hBOOL*);
 
-fun literals_of_hol_term thy th_no mode t =
-      let val (lits, types_sorts) = literals_of_term thy th_no t
-      in  (map (hol_literal_to_fol mode) lits, types_sorts) end;
+fun literals_of_hol_term thy mode t =
+  let val (lits, types_sorts) = literals_of_term thy t in
+    (map (hol_literal_to_fol mode) lits, types_sorts)
+  end
 
 (*Sign should be "true" for conjecture type constraints, "false" for type lits in clauses.*)
 fun metis_of_type_literals pos (TyLitVar ((s, _), (s', _))) =
@@ -623,13 +623,13 @@
 fun metis_of_tfree tf =
   Metis_Thm.axiom (Metis_LiteralSet.singleton (metis_of_type_literals true tf));
 
-fun hol_thm_to_fol is_conjecture th_no ctxt type_lits mode j old_skolems th =
+fun hol_thm_to_fol is_conjecture ctxt type_lits mode j old_skolems th =
   let
     val thy = ProofContext.theory_of ctxt
     val (old_skolems, (mlits, types_sorts)) =
      th |> prop_of |> Logic.strip_imp_concl
         |> conceal_old_skolem_terms j old_skolems
-        ||> (HOLogic.dest_Trueprop #> literals_of_hol_term thy th_no mode)
+        ||> (HOLogic.dest_Trueprop #> literals_of_hol_term thy mode)
   in
     if is_conjecture then
       (Metis_Thm.axiom (Metis_LiteralSet.fromList mlits),
@@ -662,7 +662,7 @@
 (* Logic maps manage the interface between HOL and first-order logic.        *)
 (* ------------------------------------------------------------------------- *)
 
-type logic_map =
+type metis_problem =
   {axioms: (Metis_Thm.thm * thm) list,
    tfrees: type_literal list,
    old_skolems: (string * term) list}
@@ -678,7 +678,7 @@
   end;
 
 (*Insert non-logical axioms corresponding to all accumulated TFrees*)
-fun add_tfrees {axioms, tfrees, old_skolems} : logic_map =
+fun add_tfrees {axioms, tfrees, old_skolems} : metis_problem =
      {axioms = map (rpair TrueI o metis_of_tfree) (distinct (op =) tfrees) @
                axioms,
       tfrees = tfrees, old_skolems = old_skolems}
@@ -722,7 +722,7 @@
   end;
 
 (* Function to generate metis clauses, including comb and type clauses *)
-fun build_logic_map mode0 ctxt type_lits cls thss =
+fun prepare_metis_problem mode0 ctxt type_lits cls thss =
   let val thy = ProofContext.theory_of ctxt
       (*The modes FO and FT are sticky. HO can be downgraded to FO.*)
       fun set_mode FO = FO
@@ -732,21 +732,20 @@
         | set_mode FT = FT
       val mode = set_mode mode0
       (*transform isabelle clause to metis clause *)
-      fun add_thm th_no is_conjecture (metis_ith, isa_ith)
-                  {axioms, tfrees, old_skolems} : logic_map =
+      fun add_thm is_conjecture (metis_ith, isa_ith)
+                  {axioms, tfrees, old_skolems} : metis_problem =
         let
           val (mth, tfree_lits, old_skolems) =
-            hol_thm_to_fol is_conjecture th_no ctxt type_lits mode (length axioms)
+            hol_thm_to_fol is_conjecture ctxt type_lits mode (length axioms)
                            old_skolems metis_ith
         in
            {axioms = (mth, Meson.make_meta_clause isa_ith) :: axioms,
             tfrees = union (op =) tfree_lits tfrees, old_skolems = old_skolems}
         end;
       val lmap = {axioms = [], tfrees = init_tfrees ctxt, old_skolems = []}
-                 |> fold (add_thm 0 true o `I) cls
+                 |> fold (add_thm true o `I) cls
                  |> add_tfrees
-                 |> fold (fn (th_no, ths) => fold (add_thm th_no false o `I) ths)
-                         (1 upto length thss ~~ thss)
+                 |> fold (fold (add_thm false o `I)) thss
       val clause_lists = map (Metis_Thm.clause o #1) (#axioms lmap)
       fun is_used c =
         exists (Metis_LiteralSet.exists (const_in_metis c o #2)) clause_lists
@@ -763,7 +762,7 @@
                                     []
                                   else
                                     thms)
-          in lmap |> fold (add_thm ~1 false) helper_ths end
+          in lmap |> fold (add_thm false) helper_ths end
   in
     (mode, add_type_thm (type_ext thy (maps (map prop_of) (cls :: thss))) lmap)
   end
--- a/src/HOL/Tools/Nitpick/nitpick.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/Nitpick/nitpick.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -216,7 +216,7 @@
         o Pretty.chunks o cons (Pretty.str "") o single
         o Pretty.mark Markup.hilite
       else
-        (fn s => (priority s; if debug then tracing s else ()))
+        (fn s => (Output.urgent_message s; if debug then tracing s else ()))
         o Pretty.string_of
     fun pprint_m f = () |> not auto ? pprint o f
     fun pprint_v f = () |> verbose ? pprint o f
@@ -989,7 +989,7 @@
            raise Interrupt
          else
            if passed_deadline deadline then
-             (priority "Nitpick ran out of time."; ("unknown", state))
+             (Output.urgent_message "Nitpick ran out of time."; ("unknown", state))
            else
              error "Nitpick was interrupted."
 
@@ -1040,7 +1040,7 @@
     val t = state |> Proof.raw_goal |> #goal |> prop_of
   in
     case Logic.count_prems t of
-      0 => (priority "No subgoal!"; ("none", state))
+      0 => (Output.urgent_message "No subgoal!"; ("none", state))
     | n =>
       let
         val t = Logic.goal_params t i |> fst
--- a/src/HOL/Tools/Nitpick/nitpick_hol.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -2054,7 +2054,7 @@
                     map (wf_constraint_for_triple rel) triples
                     |> foldr1 s_conj |> HOLogic.mk_Trueprop
          val _ = if debug then
-                   priority ("Wellfoundedness goal: " ^
+                   Output.urgent_message ("Wellfoundedness goal: " ^
                              Syntax.string_of_term ctxt prop ^ ".")
                  else
                    ()
--- a/src/HOL/Tools/Nitpick/nitpick_model.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/Nitpick/nitpick_model.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -1043,7 +1043,7 @@
                                      (fn (s, []) => TFree (s, HOLogic.typeS)
                                        | x => TFree x))
        val _ = if debug then
-                 priority ((if negate then "Genuineness" else "Spuriousness") ^
+                 Output.urgent_message ((if negate then "Genuineness" else "Spuriousness") ^
                            " goal: " ^ Syntax.string_of_term ctxt prop ^ ".")
                else
                  ()
--- a/src/HOL/Tools/Nitpick/nitpick_mono.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/Nitpick/nitpick_mono.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -505,29 +505,29 @@
   end
 
 fun solve max_var (lits, comps, sexps) =
-    let
-      fun do_assigns assigns =
-        SOME (literals_from_assignments max_var assigns lits
-              |> tap print_solution)
-      val _ = print_problem lits comps sexps
-      val prop = PropLogic.all (map prop_for_literal lits @
-                                map prop_for_comp comps @
-                                map prop_for_sign_expr sexps)
-      val default_val = bool_from_sign Minus
-    in
-      if PropLogic.eval (K default_val) prop then
-        do_assigns (K (SOME default_val))
-      else
-        let
-          (* use the first ML solver (to avoid startup overhead) *)
-          val solvers = !SatSolver.solvers
-                        |> filter (member (op =) ["dptsat", "dpll"] o fst)
-        in
-          case snd (hd solvers) prop of
-            SatSolver.SATISFIABLE assigns => do_assigns assigns
-          | _ => NONE
-        end
-    end
+  let
+    fun do_assigns assigns =
+      SOME (literals_from_assignments max_var assigns lits
+            |> tap print_solution)
+    val _ = print_problem lits comps sexps
+    val prop = PropLogic.all (map prop_for_literal lits @
+                              map prop_for_comp comps @
+                              map prop_for_sign_expr sexps)
+    val default_val = bool_from_sign Minus
+  in
+    if PropLogic.eval (K default_val) prop then
+      do_assigns (K (SOME default_val))
+    else
+      let
+        (* use the first ML solver (to avoid startup overhead) *)
+        val solvers = !SatSolver.solvers
+                      |> filter (member (op =) ["dptsat", "dpll"] o fst)
+      in
+        case snd (hd solvers) prop of
+          SatSolver.SATISFIABLE assigns => do_assigns assigns
+        | _ => NONE
+      end
+  end
 
 type mtype_schema = mtyp * constraint_set
 type mtype_context =
--- a/src/HOL/Tools/Predicate_Compile/core_data.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/Predicate_Compile/core_data.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -26,11 +26,6 @@
     needs_random : mode list
   };
 
-  (* general operations *)
-  val unify_consts : theory -> term list -> term list -> (term list * term list)
-  val mk_casesrule : Proof.context -> term -> thm list -> term
-  val preprocess_intro : theory -> thm -> thm
-
   structure PredData : THEORY_DATA
   
   (* queries *)
@@ -90,10 +85,6 @@
 
 open Predicate_Compile_Aux;
 
-(* FIXME: should be moved to Predicate_Compile_Aux *)
-val strip_intro_concl = (strip_comb o HOLogic.dest_Trueprop o Logic.strip_imp_concl o prop_of)
-
-
 (* book-keeping *)
 
 datatype predfun_data = PredfunData of {
@@ -195,7 +186,7 @@
 
 fun lookup_predfun_data ctxt name mode =
   Option.map rep_predfun_data
-    (AList.lookup (op =) (#predfun_data (the_pred_data ctxt name)) mode)
+    (AList.lookup eq_mode (#predfun_data (the_pred_data ctxt name)) mode)
 
 fun the_predfun_data ctxt name mode =
   case lookup_predfun_data ctxt name mode of
@@ -214,140 +205,6 @@
 val intros_graph_of =
   Graph.map (K (map snd o #intros o rep_pred_data)) o PredData.get o ProofContext.theory_of
 
-(** preprocessing rules **)
-
-fun Trueprop_conv cv ct =
-  case Thm.term_of ct of
-    Const (@{const_name Trueprop}, _) $ _ => Conv.arg_conv cv ct  
-  | _ => raise Fail "Trueprop_conv"
-
-fun preprocess_equality thy rule =
-  Conv.fconv_rule
-    (imp_prems_conv
-      (Trueprop_conv (Conv.try_conv (Conv.rewr_conv (Thm.symmetric @{thm Predicate.eq_is_eq})))))
-    (Thm.transfer thy rule)
-
-fun preprocess_intro thy = expand_tuples thy #> preprocess_equality thy
-
-(* importing introduction rules *)
-
-fun unify_consts thy cs intr_ts =
-  (let
-     val add_term_consts_2 = fold_aterms (fn Const c => insert (op =) c | _ => I);
-     fun varify (t, (i, ts)) =
-       let val t' = map_types (Logic.incr_tvar (i + 1)) (#2 (Type.varify_global [] t))
-       in (maxidx_of_term t', t'::ts) end;
-     val (i, cs') = List.foldr varify (~1, []) cs;
-     val (i', intr_ts') = List.foldr varify (i, []) intr_ts;
-     val rec_consts = fold add_term_consts_2 cs' [];
-     val intr_consts = fold add_term_consts_2 intr_ts' [];
-     fun unify (cname, cT) =
-       let val consts = map snd (filter (fn c => fst c = cname) intr_consts)
-       in fold (Sign.typ_unify thy) ((replicate (length consts) cT) ~~ consts) end;
-     val (env, _) = fold unify rec_consts (Vartab.empty, i');
-     val subst = map_types (Envir.norm_type env)
-   in (map subst cs', map subst intr_ts')
-   end) handle Type.TUNIFY =>
-     (warning "Occurrences of recursive constant have non-unifiable types"; (cs, intr_ts));
-
-fun import_intros inp_pred [] ctxt =
-  let
-    val ([outp_pred], ctxt') = Variable.import_terms true [inp_pred] ctxt
-    val T = fastype_of outp_pred
-    val paramTs = ho_argsT_of_typ (binder_types T)
-    val (param_names, ctxt'') = Variable.variant_fixes
-      (map (fn i => "p" ^ (string_of_int i)) (1 upto (length paramTs))) ctxt'
-    val params = map2 (curry Free) param_names paramTs
-  in
-    (((outp_pred, params), []), ctxt')
-  end
-  | import_intros inp_pred (th :: ths) ctxt =
-    let
-      val ((_, [th']), ctxt') = Variable.import true [th] ctxt
-      val thy = ProofContext.theory_of ctxt'
-      val (pred, args) = strip_intro_concl th'
-      val T = fastype_of pred
-      val ho_args = ho_args_of_typ T args
-      fun subst_of (pred', pred) =
-        let
-          val subst = Sign.typ_match thy (fastype_of pred', fastype_of pred) Vartab.empty
-            handle Type.TYPE_MATCH => error ("Type mismatch of predicate " ^ fst (dest_Const pred)
-            ^ " (trying to match " ^ Syntax.string_of_typ ctxt (fastype_of pred')
-            ^ " and " ^ Syntax.string_of_typ ctxt (fastype_of pred) ^ ")"
-            ^ " in " ^ Display.string_of_thm ctxt th)
-        in map (fn (indexname, (s, T)) => ((indexname, s), T)) (Vartab.dest subst) end
-      fun instantiate_typ th =
-        let
-          val (pred', _) = strip_intro_concl th
-          val _ = if not (fst (dest_Const pred) = fst (dest_Const pred')) then
-            raise Fail "Trying to instantiate another predicate" else ()
-        in Thm.certify_instantiate (subst_of (pred', pred), []) th end;
-      fun instantiate_ho_args th =
-        let
-          val (_, args') = (strip_comb o HOLogic.dest_Trueprop o Logic.strip_imp_concl o prop_of) th
-          val ho_args' = map dest_Var (ho_args_of_typ T args')
-        in Thm.certify_instantiate ([], ho_args' ~~ ho_args) th end
-      val outp_pred =
-        Term_Subst.instantiate (subst_of (inp_pred, pred), []) inp_pred
-      val ((_, ths'), ctxt1) =
-        Variable.import false (map (instantiate_typ #> instantiate_ho_args) ths) ctxt'
-    in
-      (((outp_pred, ho_args), th' :: ths'), ctxt1)
-    end
-
-(* generation of case rules from user-given introduction rules *)
-
-fun mk_args2 (Type (@{type_name Product_Type.prod}, [T1, T2])) st =
-    let
-      val (t1, st') = mk_args2 T1 st
-      val (t2, st'') = mk_args2 T2 st'
-    in
-      (HOLogic.mk_prod (t1, t2), st'')
-    end
-  (*| mk_args2 (T as Type ("fun", _)) (params, ctxt) = 
-    let
-      val (S, U) = strip_type T
-    in
-      if U = HOLogic.boolT then
-        (hd params, (tl params, ctxt))
-      else
-        let
-          val ([x], ctxt') = Variable.variant_fixes ["x"] ctxt
-        in
-          (Free (x, T), (params, ctxt'))
-        end
-    end*)
-  | mk_args2 T (params, ctxt) =
-    let
-      val ([x], ctxt') = Variable.variant_fixes ["x"] ctxt
-    in
-      (Free (x, T), (params, ctxt'))
-    end
-
-fun mk_casesrule ctxt pred introrules =
-  let
-    (* TODO: can be simplified if parameters are not treated specially ? *)
-    val (((pred, params), intros_th), ctxt1) = import_intros pred introrules ctxt
-    (* TODO: distinct required ? -- test case with more than one parameter! *)
-    val params = distinct (op aconv) params
-    val intros = map prop_of intros_th
-    val ([propname], ctxt2) = Variable.variant_fixes ["thesis"] ctxt1
-    val prop = HOLogic.mk_Trueprop (Free (propname, HOLogic.boolT))
-    val argsT = binder_types (fastype_of pred)
-    (* TODO: can be simplified if parameters are not treated specially ? <-- see uncommented code! *)
-    val (argvs, _) = fold_map mk_args2 argsT (params, ctxt2)
-    fun mk_case intro =
-      let
-        val (_, args) = (strip_comb o HOLogic.dest_Trueprop o Logic.strip_imp_concl) intro
-        val prems = Logic.strip_imp_prems intro
-        val eqprems =
-          map2 (HOLogic.mk_Trueprop oo (curry HOLogic.mk_eq)) argvs args
-        val frees = map Free (fold Term.add_frees (args @ prems) [])
-      in fold Logic.all frees (Logic.list_implies (eqprems @ prems, prop)) end
-    val assm = HOLogic.mk_Trueprop (list_comb (pred, argvs))
-    val cases = map mk_case intros
-  in Logic.list_implies (assm :: cases, prop) end;
-
 fun prove_casesrule ctxt (pred, (pre_cases_rule, nparams)) cases_rule =
   let
     val thy = ProofContext.theory_of ctxt
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -35,7 +35,7 @@
   val split_map_modeT : (mode -> typ -> typ option * typ option)
     -> mode -> typ list -> typ list * typ list
   val split_mode : mode -> term list -> term list * term list
-  val split_modeT' : mode -> typ list -> typ list * typ list
+  val split_modeT : mode -> typ list -> typ list * typ list
   val string_of_mode : mode -> string
   val ascii_string_of_mode : mode -> string
   (* premises *)
@@ -53,6 +53,7 @@
   val is_constr : Proof.context -> string -> bool
   val focus_ex : term -> Name.context -> ((string * typ) list * term) * Name.context
   val strip_all : term -> (string * typ) list * term
+  val strip_intro_concl : thm -> term * term list
   (* introduction rule combinators *)
   val map_atoms : (term -> term) -> term -> term
   val fold_atoms : (term -> 'a -> 'a) -> term -> 'a -> 'a
@@ -157,6 +158,11 @@
   val remove_equalities : theory -> thm -> thm
   val remove_pointless_clauses : thm -> thm list
   val peephole_optimisation : theory -> thm -> thm option
+  (* auxillary *)
+  val unify_consts : theory -> term list -> term list -> (term list * term list)
+  val mk_casesrule : Proof.context -> term -> thm list -> term
+  val preprocess_intro : theory -> thm -> thm
+  
   val define_quickcheck_predicate :
     term -> theory -> (((string * typ) * (string * typ) list) * thm) * theory
 end;
@@ -387,24 +393,22 @@
 fun map_filter_prod f (Const (@{const_name Pair}, _) $ t1 $ t2) =
   comb_option HOLogic.mk_prod (map_filter_prod f t1, map_filter_prod f t2)
   | map_filter_prod f t = f t
-
-(* obviously, split_mode' and split_modeT' do not match? where does that cause problems? *)
   
-fun split_modeT' mode Ts =
+fun split_modeT mode Ts =
   let
-    fun split_arg_mode' (Fun _) T = ([], [])
-      | split_arg_mode' (Pair (m1, m2)) (Type (@{type_name Product_Type.prod}, [T1, T2])) =
+    fun split_arg_mode (Fun _) T = ([], [])
+      | split_arg_mode (Pair (m1, m2)) (Type (@{type_name Product_Type.prod}, [T1, T2])) =
         let
-          val (i1, o1) = split_arg_mode' m1 T1
-          val (i2, o2) = split_arg_mode' m2 T2
+          val (i1, o1) = split_arg_mode m1 T1
+          val (i2, o2) = split_arg_mode m2 T2
         in
           (i1 @ i2, o1 @ o2)
         end
-      | split_arg_mode' Input T = ([T], [])
-      | split_arg_mode' Output T = ([], [T])
-      | split_arg_mode' _ _ = raise Fail "split_modeT': mode and type do not match"
+      | split_arg_mode Input T = ([T], [])
+      | split_arg_mode Output T = ([], [T])
+      | split_arg_mode _ _ = raise Fail "split_modeT: mode and type do not match"
   in
-    (pairself flat o split_list) (map2 split_arg_mode' (strip_fun_mode mode) Ts)
+    (pairself flat o split_list) (map2 split_arg_mode (strip_fun_mode mode) Ts)
   end
 
 fun string_of_mode mode =
@@ -546,6 +550,8 @@
     val t'' = Term.subst_bounds (rev vs, t');
   in ((ps', t''), nctxt') end;
 
+val strip_intro_concl = (strip_comb o HOLogic.dest_Trueprop o Logic.strip_imp_concl o prop_of)
+  
 (* introduction rule combinators *)
 
 fun map_atoms f intro = 
@@ -1048,6 +1054,144 @@
       (process_False (process_True (prop_of (process intro))))
   end
 
+
+(* importing introduction rules *)
+
+fun import_intros inp_pred [] ctxt =
+  let
+    val ([outp_pred], ctxt') = Variable.import_terms true [inp_pred] ctxt
+    val T = fastype_of outp_pred
+    val paramTs = ho_argsT_of_typ (binder_types T)
+    val (param_names, ctxt'') = Variable.variant_fixes
+      (map (fn i => "p" ^ (string_of_int i)) (1 upto (length paramTs))) ctxt'
+    val params = map2 (curry Free) param_names paramTs
+  in
+    (((outp_pred, params), []), ctxt')
+  end
+  | import_intros inp_pred (th :: ths) ctxt =
+    let
+      val ((_, [th']), ctxt') = Variable.import true [th] ctxt
+      val thy = ProofContext.theory_of ctxt'
+      val (pred, args) = strip_intro_concl th'
+      val T = fastype_of pred
+      val ho_args = ho_args_of_typ T args
+      fun subst_of (pred', pred) =
+        let
+          val subst = Sign.typ_match thy (fastype_of pred', fastype_of pred) Vartab.empty
+            handle Type.TYPE_MATCH => error ("Type mismatch of predicate " ^ fst (dest_Const pred)
+            ^ " (trying to match " ^ Syntax.string_of_typ ctxt (fastype_of pred')
+            ^ " and " ^ Syntax.string_of_typ ctxt (fastype_of pred) ^ ")"
+            ^ " in " ^ Display.string_of_thm ctxt th)
+        in map (fn (indexname, (s, T)) => ((indexname, s), T)) (Vartab.dest subst) end
+      fun instantiate_typ th =
+        let
+          val (pred', _) = strip_intro_concl th
+          val _ = if not (fst (dest_Const pred) = fst (dest_Const pred')) then
+            raise Fail "Trying to instantiate another predicate" else ()
+        in Thm.certify_instantiate (subst_of (pred', pred), []) th end;
+      fun instantiate_ho_args th =
+        let
+          val (_, args') = (strip_comb o HOLogic.dest_Trueprop o Logic.strip_imp_concl o prop_of) th
+          val ho_args' = map dest_Var (ho_args_of_typ T args')
+        in Thm.certify_instantiate ([], ho_args' ~~ ho_args) th end
+      val outp_pred =
+        Term_Subst.instantiate (subst_of (inp_pred, pred), []) inp_pred
+      val ((_, ths'), ctxt1) =
+        Variable.import false (map (instantiate_typ #> instantiate_ho_args) ths) ctxt'
+    in
+      (((outp_pred, ho_args), th' :: ths'), ctxt1)
+    end
+  
+(* generation of case rules from user-given introduction rules *)
+
+fun mk_args2 (Type (@{type_name Product_Type.prod}, [T1, T2])) st =
+    let
+      val (t1, st') = mk_args2 T1 st
+      val (t2, st'') = mk_args2 T2 st'
+    in
+      (HOLogic.mk_prod (t1, t2), st'')
+    end
+  (*| mk_args2 (T as Type ("fun", _)) (params, ctxt) = 
+    let
+      val (S, U) = strip_type T
+    in
+      if U = HOLogic.boolT then
+        (hd params, (tl params, ctxt))
+      else
+        let
+          val ([x], ctxt') = Variable.variant_fixes ["x"] ctxt
+        in
+          (Free (x, T), (params, ctxt'))
+        end
+    end*)
+  | mk_args2 T (params, ctxt) =
+    let
+      val ([x], ctxt') = Variable.variant_fixes ["x"] ctxt
+    in
+      (Free (x, T), (params, ctxt'))
+    end
+
+fun mk_casesrule ctxt pred introrules =
+  let
+    (* TODO: can be simplified if parameters are not treated specially ? *)
+    val (((pred, params), intros_th), ctxt1) = import_intros pred introrules ctxt
+    (* TODO: distinct required ? -- test case with more than one parameter! *)
+    val params = distinct (op aconv) params
+    val intros = map prop_of intros_th
+    val ([propname], ctxt2) = Variable.variant_fixes ["thesis"] ctxt1
+    val prop = HOLogic.mk_Trueprop (Free (propname, HOLogic.boolT))
+    val argsT = binder_types (fastype_of pred)
+    (* TODO: can be simplified if parameters are not treated specially ? <-- see uncommented code! *)
+    val (argvs, _) = fold_map mk_args2 argsT (params, ctxt2)
+    fun mk_case intro =
+      let
+        val (_, args) = (strip_comb o HOLogic.dest_Trueprop o Logic.strip_imp_concl) intro
+        val prems = Logic.strip_imp_prems intro
+        val eqprems =
+          map2 (HOLogic.mk_Trueprop oo (curry HOLogic.mk_eq)) argvs args
+        val frees = map Free (fold Term.add_frees (args @ prems) [])
+      in fold Logic.all frees (Logic.list_implies (eqprems @ prems, prop)) end
+    val assm = HOLogic.mk_Trueprop (list_comb (pred, argvs))
+    val cases = map mk_case intros
+  in Logic.list_implies (assm :: cases, prop) end;
+  
+
+(* unifying constants to have the same type variables *)
+
+fun unify_consts thy cs intr_ts =
+  (let
+     val add_term_consts_2 = fold_aterms (fn Const c => insert (op =) c | _ => I);
+     fun varify (t, (i, ts)) =
+       let val t' = map_types (Logic.incr_tvar (i + 1)) (#2 (Type.varify_global [] t))
+       in (maxidx_of_term t', t'::ts) end;
+     val (i, cs') = List.foldr varify (~1, []) cs;
+     val (i', intr_ts') = List.foldr varify (i, []) intr_ts;
+     val rec_consts = fold add_term_consts_2 cs' [];
+     val intr_consts = fold add_term_consts_2 intr_ts' [];
+     fun unify (cname, cT) =
+       let val consts = map snd (filter (fn c => fst c = cname) intr_consts)
+       in fold (Sign.typ_unify thy) ((replicate (length consts) cT) ~~ consts) end;
+     val (env, _) = fold unify rec_consts (Vartab.empty, i');
+     val subst = map_types (Envir.norm_type env)
+   in (map subst cs', map subst intr_ts')
+   end) handle Type.TUNIFY =>
+     (warning "Occurrences of recursive constant have non-unifiable types"; (cs, intr_ts));
+
+(* preprocessing rules *)
+
+fun Trueprop_conv cv ct =
+  case Thm.term_of ct of
+    Const (@{const_name Trueprop}, _) $ _ => Conv.arg_conv cv ct  
+  | _ => raise Fail "Trueprop_conv"
+
+fun preprocess_equality thy rule =
+  Conv.fconv_rule
+    (imp_prems_conv
+      (Trueprop_conv (Conv.try_conv (Conv.rewr_conv (Thm.symmetric @{thm Predicate.eq_is_eq})))))
+    (Thm.transfer thy rule)
+
+fun preprocess_intro thy = expand_tuples thy #> preprocess_equality thy
+
 (* defining a quickcheck predicate *)
 
 fun strip_imp_prems (Const(@{const_name HOL.implies}, _) $ A $ B) = A :: strip_imp_prems B
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -26,6 +26,8 @@
   val put_dseq_result : (unit -> term DSequence.dseq) -> Proof.context -> Proof.context
   val put_dseq_random_result : (unit -> int -> int -> int * int -> term DSequence.dseq * (int * int)) ->
     Proof.context -> Proof.context
+  val put_new_dseq_result : (unit -> int -> term Lazy_Sequence.lazy_sequence) ->
+    Proof.context -> Proof.context
   val put_lseq_random_result :
     (unit -> int -> int -> int * int -> int -> term Lazy_Sequence.lazy_sequence) ->
     Proof.context -> Proof.context
@@ -40,6 +42,7 @@
   val add_depth_limited_random_equations : options -> string list -> theory -> theory
   val add_random_dseq_equations : options -> string list -> theory -> theory
   val add_new_random_dseq_equations : options -> string list -> theory -> theory
+  val add_generator_dseq_equations : options -> string list -> theory -> theory
   val mk_tracing : string -> term -> term
   val prepare_intrs : options -> Proof.context -> string list -> thm list ->
     ((string * typ) list * string list * string list * (string * mode list) list *
@@ -54,7 +57,6 @@
   type moded_clause = term list * (Predicate_Compile_Aux.indprem * mode_derivation) list
   type 'a pred_mode_table = (string * ((bool * mode) * 'a) list) list
 
-
 end;
 
 structure Predicate_Compile_Core : PREDICATE_COMPILE_CORE =
@@ -165,7 +167,7 @@
 
 (* validity checks *)
 
-fun check_expected_modes preds options modes =
+fun check_expected_modes options preds modes =
   case expected_modes options of
     SOME (s, ms) => (case AList.lookup (op =) modes s of
       SOME modes =>
@@ -181,7 +183,7 @@
       | NONE => ())
   | NONE => ()
 
-fun check_proposed_modes preds options modes errors =
+fun check_proposed_modes options preds modes errors =
   map (fn (s, _) => case proposed_modes options s of
     SOME ms => (case AList.lookup (op =) modes s of
       SOME inferred_ms =>
@@ -317,7 +319,7 @@
     fn compfuns => fn s => fn T => fn mode => fn additional_arguments => fn compilation =>
     let
       val [depth] = additional_arguments
-      val (_, Ts) = split_modeT' mode (binder_types T)
+      val (_, Ts) = split_modeT mode (binder_types T)
       val T' = mk_predT compfuns (HOLogic.mk_tupleT Ts)
       val if_const = Const (@{const_name "If"}, @{typ bool} --> T' --> T' --> T')
     in
@@ -1377,8 +1379,8 @@
         options (lookup_mode, lookup_neg_mode, needs_random) ctxt preds all_modes param_vs clauses)
     val thy' = fold (fn (s, ms) => set_needs_random s ms) needs_random thy
     val modes = map (fn (p, mps) => (p, map fst mps)) moded_clauses
-    val _ = check_expected_modes preds options modes
-    val _ = check_proposed_modes preds options modes errors
+    val _ = check_expected_modes options preds modes
+    val _ = check_proposed_modes options preds modes errors
     val _ = print_modes options modes
     val _ = print_step options "Defining executable functions..."
     val thy'' =
@@ -1658,6 +1660,12 @@
 );
 val put_dseq_random_result = Dseq_Random_Result.put;
 
+structure New_Dseq_Result = Proof_Data (
+  type T = unit -> int -> term Lazy_Sequence.lazy_sequence
+  fun init _ () = error "New_Dseq_Random_Result"
+);
+val put_new_dseq_result = New_Dseq_Result.put;
+
 structure Lseq_Random_Result = Proof_Data (
   type T = unit -> int -> int -> int * int -> int -> term Lazy_Sequence.lazy_sequence
   fun init _ () = error "Lseq_Random_Result"
@@ -1671,8 +1679,10 @@
 val put_lseq_random_stats_result = Lseq_Random_Stats_Result.put;
 
 (*FIXME turn this into an LCF-guarded preprocessor for comprehensions*)
-fun analyze_compr ctxt compfuns param_user_modes (compilation, arguments) t_compr =
+fun analyze_compr ctxt (comp_modifiers, additional_arguments) param_user_modes
+  (compilation, arguments) t_compr =
   let
+    val compfuns = Comp_Mod.compfuns comp_modifiers
     val all_modes_of = all_modes_of compilation
     val split = case t_compr of (Const (@{const_name Collect}, _) $ t) => t
       | _ => error ("Not a set comprehension: " ^ Syntax.string_of_term ctxt t_compr);
@@ -1734,28 +1744,6 @@
           | d :: _ :: _ => (warning ("Multiple modes possible for comprehension "
                     ^ Syntax.string_of_term ctxt t_compr); d);
         val (_, outargs) = split_mode (head_mode_of deriv) all_args
-        val additional_arguments =
-          case compilation of
-            Pred => []
-          | Random => map (HOLogic.mk_number @{typ "code_numeral"}) arguments @
-            [@{term "(1, 1) :: code_numeral * code_numeral"}]
-          | Annotated => []
-          | Depth_Limited => [HOLogic.mk_number @{typ "code_numeral"} (hd arguments)]
-          | Depth_Limited_Random => map (HOLogic.mk_number @{typ "code_numeral"}) arguments @
-            [@{term "(1, 1) :: code_numeral * code_numeral"}]
-          | DSeq => []
-          | Pos_Random_DSeq => []
-          | New_Pos_Random_DSeq => []
-        val comp_modifiers =
-          case compilation of
-            Pred => predicate_comp_modifiers
-          | Random => random_comp_modifiers
-          | Depth_Limited => depth_limited_comp_modifiers
-          | Depth_Limited_Random => depth_limited_random_comp_modifiers
-          (*| Annotated => annotated_comp_modifiers*)
-          | DSeq => dseq_comp_modifiers
-          | Pos_Random_DSeq => pos_random_dseq_comp_modifiers
-          | New_Pos_Random_DSeq => new_pos_random_dseq_comp_modifiers
         val t_pred = compile_expr comp_modifiers ctxt
           (body, deriv) [] additional_arguments;
         val T_pred = dest_predT compfuns (fastype_of t_pred)
@@ -1775,14 +1763,32 @@
           | count' i (x' :: xs) = if x = x' then count' (i + 1) xs else count' i xs
       in count' 0 xs end
     fun accumulate xs = map (fn x => (x, count xs x)) (sort int_ord (distinct (op =) xs))
-    val compfuns =
+    val comp_modifiers =
       case compilation of
-        Random => PredicateCompFuns.compfuns
-      | DSeq => DSequence_CompFuns.compfuns
-      | Pos_Random_DSeq => Random_Sequence_CompFuns.compfuns
-      | New_Pos_Random_DSeq => New_Pos_Random_Sequence_CompFuns.depth_unlimited_compfuns
-      | _ => PredicateCompFuns.compfuns
-    val t = analyze_compr ctxt compfuns param_user_modes options t_compr;
+          Pred => predicate_comp_modifiers
+        | Random => random_comp_modifiers
+        | Depth_Limited => depth_limited_comp_modifiers
+        | Depth_Limited_Random => depth_limited_random_comp_modifiers
+        (*| Annotated => annotated_comp_modifiers*)
+        | DSeq => dseq_comp_modifiers
+        | Pos_Random_DSeq => pos_random_dseq_comp_modifiers
+        | New_Pos_Random_DSeq => new_pos_random_dseq_comp_modifiers
+        | Pos_Generator_DSeq => pos_generator_dseq_comp_modifiers
+    val compfuns = Comp_Mod.compfuns comp_modifiers
+    val additional_arguments =
+      case compilation of
+        Pred => []
+      | Random => map (HOLogic.mk_number @{typ "code_numeral"}) arguments @
+        [@{term "(1, 1) :: code_numeral * code_numeral"}]
+      | Annotated => []
+      | Depth_Limited => [HOLogic.mk_number @{typ "code_numeral"} (hd arguments)]
+      | Depth_Limited_Random => map (HOLogic.mk_number @{typ "code_numeral"}) arguments @
+        [@{term "(1, 1) :: code_numeral * code_numeral"}]
+      | DSeq => []
+      | Pos_Random_DSeq => []
+      | New_Pos_Random_DSeq => []
+      | Pos_Generator_DSeq => []
+    val t = analyze_compr ctxt (comp_modifiers, additional_arguments) param_user_modes options t_compr;
     val T = dest_predT compfuns (fastype_of t);
     val t' =
       if stats andalso compilation = New_Pos_Random_DSeq then
@@ -1793,7 +1799,7 @@
         mk_map compfuns T HOLogic.termT (HOLogic.term_of_const T) t
     val thy = ProofContext.theory_of ctxt
     val (ts, statistics) =
-      case compilation of
+      (case compilation of
        (* Random =>
           fst (Predicate.yieldn k
           (Code_Eval.eval NONE ("Predicate_Compile_Core.random_eval_ref", random_eval_ref)
@@ -1803,44 +1809,54 @@
           let
             val [nrandom, size, depth] = arguments
           in
-            rpair NONE (fst (DSequence.yieldn k
+            rpair NONE (TimeLimit.timeLimit (Time.fromSeconds 20) (fn () => fst (DSequence.yieldn k
               (Code_Runtime.dynamic_value_strict (Dseq_Random_Result.get, put_dseq_random_result, "Predicate_Compile_Core.put_dseq_random_result")
                 thy NONE (fn proc => fn g => fn nrandom => fn size => fn s => g nrandom size s |>> DSequence.map proc)
                   t' [] nrandom size
                 |> Random_Engine.run)
-              depth true))
+              depth true)) ())
           end
       | DSeq =>
-          rpair NONE (fst (DSequence.yieldn k
+          rpair NONE (TimeLimit.timeLimit (Time.fromSeconds 20) (fn () => fst (DSequence.yieldn k
             (Code_Runtime.dynamic_value_strict (Dseq_Result.get, put_dseq_result, "Predicate_Compile_Core.put_dseq_result")
-              thy NONE DSequence.map t' []) (the_single arguments) true))
+              thy NONE DSequence.map t' []) (the_single arguments) true)) ())
+      | Pos_Generator_DSeq =>
+          let
+            val depth = (the_single arguments)
+          in
+            rpair NONE (TimeLimit.timeLimit (Time.fromSeconds 20) (fn () => fst (Lazy_Sequence.yieldn k
+              (Code_Runtime.dynamic_value_strict (New_Dseq_Result.get, put_new_dseq_result, "Predicate_Compile_Core.put_new_dseq_result")
+              thy NONE (fn proc => fn g => fn depth => g depth |> Lazy_Sequence.mapa proc)
+              t' [] depth))) ())
+          end
       | New_Pos_Random_DSeq =>
           let
             val [nrandom, size, depth] = arguments
             val seed = Random_Engine.next_seed ()
           in
             if stats then
-              apsnd (SOME o accumulate) (split_list
-              (fst (Lazy_Sequence.yieldn k
+              apsnd (SOME o accumulate) (split_list (TimeLimit.timeLimit (Time.fromSeconds 20)
+              (fn () => fst (Lazy_Sequence.yieldn k
                 (Code_Runtime.dynamic_value_strict
                   (Lseq_Random_Stats_Result.get, put_lseq_random_stats_result, "Predicate_Compile_Core.put_lseq_random_stats_result")
                   thy NONE
                   (fn proc => fn g => fn nrandom => fn size => fn s => fn depth => g nrandom size s depth
                     |> Lazy_Sequence.mapa (apfst proc))
-                    t' [] nrandom size seed depth))))
+                    t' [] nrandom size seed depth))) ()))
             else rpair NONE
-              (fst (Lazy_Sequence.yieldn k
+              (TimeLimit.timeLimit (Time.fromSeconds 20) (fn () => fst (Lazy_Sequence.yieldn k
                 (Code_Runtime.dynamic_value_strict
                   (Lseq_Random_Result.get, put_lseq_random_result, "Predicate_Compile_Core.put_lseq_random_result")
                   thy NONE 
                   (fn proc => fn g => fn nrandom => fn size => fn s => fn depth => g nrandom size s depth
                     |> Lazy_Sequence.mapa proc)
-                    t' [] nrandom size seed depth)))
+                    t' [] nrandom size seed depth))) ())
           end
       | _ =>
-          rpair NONE (fst (Predicate.yieldn k
+          rpair NONE (TimeLimit.timeLimit (Time.fromSeconds 20) (fn () => fst (Predicate.yieldn k
             (Code_Runtime.dynamic_value_strict (Pred_Result.get, put_pred_result, "Predicate_Compile_Core.put_pred_result")
-              thy NONE Predicate.map t' [])))
+              thy NONE Predicate.map t' []))) ()))
+     handle TimeLimit.TimeOut => error "Reached timeout during execution of values"
   in ((T, ts), statistics) end;
 
 fun values ctxt param_user_modes ((raw_expected, stats), comp_options) k t_compr =
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_data.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -74,6 +74,11 @@
 *)
 val defining_term_of_introrule = defining_term_of_introrule_term o prop_of
 
+fun defining_const_of_introrule th =
+  case defining_term_of_introrule th
+   of Const (c, _) => c
+    | _ => raise TERM ("defining_const_of_introrule failed: Not a constant", [prop_of th])
+
 (*TODO*)
 fun is_introlike_term t = true
 
@@ -194,7 +199,7 @@
         defining_const_of_equation (normalize_equation thy th) = fst (dest_Const t) then
         SOME (normalize_equation thy th)
       else
-        if is_introlike th andalso defining_term_of_introrule th = t then
+        if is_introlike th andalso defining_const_of_introrule th = fst (dest_Const t) then
           SOME th
         else
           NONE
@@ -206,7 +211,8 @@
     | ths => rev ths
     val _ =
       if show_intermediate_results options then
-        Output.tracing (commas (map (Display.string_of_thm_global thy) spec))
+        Output.tracing ("Specification for " ^ (Syntax.string_of_term_global thy t) ^ ":\n" ^
+          commas (map (Display.string_of_thm_global thy) spec))
       else ()
   in
     spec
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_quickcheck.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -16,6 +16,8 @@
   val put_lseq_result :
     (unit -> int -> int -> int * int -> int -> term list Lazy_Sequence.lazy_sequence) ->
       Proof.context -> Proof.context;
+  val put_new_dseq_result : (unit -> int -> term list Lazy_Sequence.lazy_sequence) ->
+    Proof.context -> Proof.context
 
   val tracing : bool Unsynchronized.ref;
   val quiet : bool Unsynchronized.ref;
@@ -51,6 +53,12 @@
 );
 val put_lseq_result = Lseq_Result.put;
 
+structure New_Dseq_Result = Proof_Data (
+  type T = unit -> int -> term list Lazy_Sequence.lazy_sequence
+  fun init _ () = error "New_Dseq_Random_Result"
+);
+val put_new_dseq_result = New_Dseq_Result.put;
+
 val tracing = Unsynchronized.ref false;
 
 val quiet = Unsynchronized.ref true;
@@ -177,11 +185,19 @@
 val mk_bind = Predicate_Compile_Aux.mk_bind RandomPredCompFuns.compfuns
 
 val mk_new_randompredT =
-  Predicate_Compile_Aux.mk_predT New_Pos_Random_Sequence_CompFuns.depth_limited_compfuns
+  Predicate_Compile_Aux.mk_predT New_Pos_Random_Sequence_CompFuns.depth_unlimited_compfuns
 val mk_new_return =
-  Predicate_Compile_Aux.mk_single New_Pos_Random_Sequence_CompFuns.depth_limited_compfuns
+  Predicate_Compile_Aux.mk_single New_Pos_Random_Sequence_CompFuns.depth_unlimited_compfuns
 val mk_new_bind =
-  Predicate_Compile_Aux.mk_bind New_Pos_Random_Sequence_CompFuns.depth_limited_compfuns
+  Predicate_Compile_Aux.mk_bind New_Pos_Random_Sequence_CompFuns.depth_unlimited_compfuns
+
+val mk_new_dseqT =
+  Predicate_Compile_Aux.mk_predT New_Pos_DSequence_CompFuns.depth_unlimited_compfuns
+val mk_gen_return =
+  Predicate_Compile_Aux.mk_single New_Pos_DSequence_CompFuns.depth_unlimited_compfuns
+val mk_gen_bind =
+  Predicate_Compile_Aux.mk_bind New_Pos_DSequence_CompFuns.depth_unlimited_compfuns
+  
 
 val mk_split_lambda = HOLogic.tupled_lambda o HOLogic.mk_tuple
 
@@ -207,6 +223,8 @@
               Predicate_Compile_Core.add_random_dseq_equations options [full_constname] thy3
           | New_Pos_Random_DSeq =>
               Predicate_Compile_Core.add_new_random_dseq_equations options [full_constname] thy3
+          | Pos_Generator_DSeq =>
+              Predicate_Compile_Core.add_generator_dseq_equations options [full_constname] thy3
           (*| Depth_Limited_Random =>
               Predicate_Compile_Core.add_depth_limited_random_equations options [full_constname] thy3*))
     (*val _ = Predicate_Compile_Core.print_all_modes compilation thy4*)
@@ -223,6 +241,7 @@
             case compilation of
               Pos_Random_DSeq => mk_randompredT (HOLogic.mk_tupleT (map snd vs'))
             | New_Pos_Random_DSeq => mk_new_randompredT (HOLogic.mk_tupleT (map snd vs'))
+            | Pos_Generator_DSeq => mk_new_dseqT (HOLogic.mk_tupleT (map snd vs'))
             | Depth_Limited_Random =>
               [@{typ code_numeral}, @{typ code_numeral}, @{typ code_numeral},
               @{typ "code_numeral * code_numeral"}] ---> mk_predT (HOLogic.mk_tupleT (map snd vs'))
@@ -238,6 +257,9 @@
         | New_Pos_Random_DSeq => mk_new_bind (prog,
             mk_split_lambda (map Free vs') (mk_new_return (HOLogic.mk_list @{typ term}
             (map2 HOLogic.mk_term_of (map snd vs') (map Free vs')))))
+        | Pos_Generator_DSeq => mk_gen_bind (prog,
+            mk_split_lambda (map Free vs') (mk_gen_return (HOLogic.mk_list @{typ term}
+            (map2 HOLogic.mk_term_of (map snd vs') (map Free vs')))))
         | Depth_Limited_Random => fold_rev (curry absdummy)
             [@{typ code_numeral}, @{typ code_numeral}, @{typ code_numeral},
              @{typ "code_numeral * code_numeral"}]
@@ -274,7 +296,18 @@
                  val seed = Random_Engine.next_seed ()
                in compiled_term nrandom size seed depth end))
           end
-      | Depth_Limited_Random =>
+      | Pos_Generator_DSeq =>
+          let
+            val compiled_term =
+              Code_Runtime.dynamic_value_strict
+                (New_Dseq_Result.get, put_new_dseq_result, "Predicate_Compile_Quickcheck.put_new_dseq_result")
+                thy4 (SOME target)
+                (fn proc => fn g => fn depth => g depth |> (Lazy_Sequence.mapa o map) proc)
+                qc_term []
+          in
+            fn size => fn nrandom => fn depth => Option.map fst (Lazy_Sequence.yield (compiled_term depth))
+          end
+       | Depth_Limited_Random =>
           let
             val compiled_term = Code_Runtime.dynamic_value_strict
               (Pred_Result.get, put_pred_result, "Predicate_Compile_Quickcheck.put_pred_result")
@@ -294,7 +327,7 @@
     fun try' j =
       if j <= i then
         let
-          val _ = if quiet then () else priority ("Executing depth " ^ string_of_int j)
+          val _ = if quiet then () else Output.urgent_message ("Executing depth " ^ string_of_int j)
         in
           case f j handle Match => (if quiet then ()
              else warning "Exception Match raised during quickcheck"; NONE)
--- a/src/HOL/Tools/SMT/cvc3_solver.ML	Wed Oct 27 11:10:36 2010 -0700
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,48 +0,0 @@
-(*  Title:      HOL/Tools/SMT/cvc3_solver.ML
-    Author:     Sascha Boehme, TU Muenchen
-
-Interface of the SMT solver CVC3.
-*)
-
-signature CVC3_SOLVER =
-sig
-  val setup: theory -> theory
-end
-
-structure CVC3_Solver: CVC3_SOLVER =
-struct
-
-val solver_name = "cvc3"
-val env_var = "CVC3_SOLVER"
-
-val options = ["-lang", "smtlib", "-output-lang", "presentation"]
-
-val is_sat = String.isPrefix "Satisfiable."
-val is_unsat = String.isPrefix "Unsatisfiable."
-val is_unknown = String.isPrefix "Unknown."
-
-fun raise_cex real = raise SMT_Solver.SMT_COUNTEREXAMPLE (real, [])
-
-fun core_oracle (output, _) =
-  let
-    val empty_line = (fn "" => true | _ => false)
-    val split_first = (fn [] => ("", []) | l :: ls => (l, ls))
-    val (l, _) = split_first (snd (chop_while empty_line output))
-  in
-    if is_unsat l then @{cprop False}
-    else if is_sat l then raise_cex true
-    else if is_unknown l then raise_cex false
-    else raise SMT_Solver.SMT (solver_name ^ " failed")
-  end
-
-fun solver oracle _ = {
-  command = {env_var=env_var, remote_name=SOME solver_name},
-  arguments = options,
-  interface = SMTLIB_Interface.interface,
-  reconstruct = pair o oracle }
-
-val setup =
-  Thm.add_oracle (Binding.name solver_name, core_oracle) #-> (fn (_, oracle) =>
-  Context.theory_map (SMT_Solver.add_solver (solver_name, solver oracle)))
-
-end
--- a/src/HOL/Tools/SMT/smt_monomorph.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/SMT/smt_monomorph.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -6,7 +6,8 @@
 
 signature SMT_MONOMORPH =
 sig
-  val monomorph: thm list -> Proof.context -> thm list * Proof.context
+  val monomorph: (int * thm) list -> Proof.context ->
+    (int * thm) list * Proof.context
 end
 
 structure SMT_Monomorph: SMT_MONOMORPH =
@@ -33,9 +34,11 @@
 fun tvar_consts_of thm = collect_consts_if typ_has_tvars insert_const thm []
 
 
-fun incr_indexes thms =
-  let fun inc thm idx = (Thm.incr_indexes idx thm, Thm.maxidx_of thm + idx + 1)
-  in fst (fold_map inc thms 0) end
+fun incr_indexes irules =
+  let
+    fun inc (i, thm) idx =
+      ((i, Thm.incr_indexes idx thm), Thm.maxidx_of thm + idx + 1)
+  in fst (fold_map inc irules 0) end
 
 
 (* Compute all substitutions from the types "Ts" to all relevant
@@ -71,7 +74,7 @@
    (without schematic types) which do not occur in any of the
    previous rounds. Note that thus no schematic type variables are
    shared among theorems. *)
-fun specialize thy all_grounds new_grounds (thm, scs) =
+fun specialize thy all_grounds new_grounds (irule, scs) =
   let
     fun spec (subst, consts) next_grounds =
       [subst]
@@ -80,7 +83,7 @@
       |-> fold_map (apply_subst all_grounds consts)
   in
     fold_map spec scs #>> (fn scss =>
-    (thm, fold (fold (insert (eq_snd (op =)))) scss []))
+    (irule, fold (fold (insert (eq_snd (op =)))) scss []))
   end
 
 
@@ -89,16 +92,16 @@
    computation uses only the constants occurring with schematic type
    variables in the propositions. To ease comparisons, such sets of
    costants are always kept in their initial order. *)
-fun incremental_monomorph thy limit all_grounds new_grounds ths =
+fun incremental_monomorph thy limit all_grounds new_grounds irules =
   let
     val all_grounds' = Symtab.merge_list (op =) (all_grounds, new_grounds)
     val spec = specialize thy all_grounds' new_grounds
-    val (ths', new_grounds') = fold_map spec ths Symtab.empty
+    val (irs, new_grounds') = fold_map spec irules Symtab.empty
   in
-    if Symtab.is_empty new_grounds' then ths'
+    if Symtab.is_empty new_grounds' then irs
     else if limit > 0
-    then incremental_monomorph thy (limit-1) all_grounds' new_grounds' ths'
-    else (warning "SMT: monomorphization limit reached"; ths')
+    then incremental_monomorph thy (limit-1) all_grounds' new_grounds' irs
+    else (warning "SMT: monomorphization limit reached"; irs)
   end
 
 
@@ -137,9 +140,9 @@
 
     fun cert (ix, (S, T)) = pairself (Thm.ctyp_of thy) (TVar (ix, S), T)
 
-    fun inst thm subst =
+    fun inst (i, thm) subst =
       let val cTs = Vartab.fold (cons o cert) (fold complete Tenv subst) []
-      in Thm.instantiate (cTs, []) thm end
+      in (i, Thm.instantiate (cTs, []) thm) end
 
   in uncurry (map o inst) end
 
@@ -147,7 +150,7 @@
 fun mono_all ctxt _ [] monos = (monos, ctxt)
   | mono_all ctxt limit polys monos =
       let
-        fun invent_types thm ctxt =
+        fun invent_types (_, thm) ctxt =
           let val (vs, Ss) = split_list (Term.add_tvars (Thm.prop_of thm) [])
           in
             ctxt
@@ -159,7 +162,7 @@
         val thy = ProofContext.theory_of ctxt'
 
         val ths = polys
-          |> map (fn thm => (thm, [(Vartab.empty, tvar_consts_of thm)]))
+          |> map (fn (_, thm) => (thm, [(Vartab.empty, tvar_consts_of thm)]))
 
         (* all constant names occurring with schematic types *)
         val ns = fold (fold (fold (insert (op =) o fst) o snd) o snd) ths []
@@ -167,11 +170,11 @@
         (* all known instances with non-schematic types *)
         val grounds =
           Symtab.make (map (rpair []) ns)
-          |> fold (add_consts (K true)) monos
-          |> fold (add_consts (not o typ_has_tvars)) polys
+          |> fold (add_consts (K true) o snd) monos
+          |> fold (add_consts (not o typ_has_tvars) o snd) polys
       in
         polys
-        |> map (fn thm => (thm, [(Vartab.empty, tvar_consts_of thm)]))
+        |> map (fn (i, thm) => ((i, thm), [(Vartab.empty, tvar_consts_of thm)]))
         |> incremental_monomorph thy limit Symtab.empty grounds
         |> map (apsnd (filter_most_specific thy))
         |> flat o map2 (instantiate thy) Tenvs
@@ -192,9 +195,9 @@
      The initial set of theorems must not contain any schematic term
    variables, and the final list of theorems does not contain any
    schematic type variables anymore. *)
-fun monomorph thms ctxt =
-  thms
-  |> List.partition (Term.exists_type typ_has_tvars o Thm.prop_of)
+fun monomorph irules ctxt =
+  irules
+  |> List.partition (Term.exists_type typ_has_tvars o Thm.prop_of o snd)
   |>> incr_indexes
   |-> mono_all ctxt monomorph_limit
 
--- a/src/HOL/Tools/SMT/smt_normalize.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/SMT/smt_normalize.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -17,9 +17,10 @@
 
 signature SMT_NORMALIZE =
 sig
-  type extra_norm = thm list -> Proof.context -> thm list * Proof.context
-  val normalize: extra_norm -> bool -> thm list -> Proof.context ->
-    thm list * Proof.context
+  type extra_norm = bool -> (int * thm) list -> Proof.context ->
+    (int * thm) list * Proof.context
+  val normalize: extra_norm -> bool -> (int * thm) list -> Proof.context ->
+    (int * thm) list * Proof.context
   val atomize_conv: Proof.context -> conv
   val eta_expand_conv: (Proof.context -> conv) -> Proof.context -> conv
 end
@@ -52,8 +53,8 @@
     if_true_conv is_trivial_distinct (Conv.rewrs_conv thms)
 in
 fun trivial_distinct ctxt =
-  map ((Term.exists_subterm is_trivial_distinct o Thm.prop_of) ??
-    Conv.fconv_rule (Conv.top_conv distinct_conv ctxt))
+  map (apsnd ((Term.exists_subterm is_trivial_distinct o Thm.prop_of) ??
+    Conv.fconv_rule (Conv.top_conv distinct_conv ctxt)))
 end
 
 
@@ -72,8 +73,8 @@
   val unfold_conv = if_true_conv is_bool_case (Conv.rewrs_conv thms)
 in
 fun rewrite_bool_cases ctxt =
-  map ((Term.exists_subterm is_bool_case o Thm.prop_of) ??
-    Conv.fconv_rule (Conv.top_conv (K unfold_conv) ctxt))
+  map (apsnd ((Term.exists_subterm is_bool_case o Thm.prop_of) ??
+    Conv.fconv_rule (Conv.top_conv (K unfold_conv) ctxt)))
 end
 
 
@@ -108,8 +109,8 @@
     Conv.no_conv
 in
 fun normalize_numerals ctxt =
-  map ((Term.exists_subterm (is_strange_number ctxt) o Thm.prop_of) ??
-    Conv.fconv_rule (Conv.top_sweep_conv pos_conv ctxt))
+  map (apsnd ((Term.exists_subterm (is_strange_number ctxt) o Thm.prop_of) ??
+    Conv.fconv_rule (Conv.top_sweep_conv pos_conv ctxt)))
 end
 
 
@@ -117,7 +118,7 @@
 (* embedding of standard natural number operations into integer operations *)
 
 local
-  val nat_embedding = @{lemma
+  val nat_embedding = map (pair ~1) @{lemma
     "nat (int n) = n"
     "i >= 0 --> int (nat i) = i"
     "i < 0 --> int (nat i) = 0"
@@ -179,8 +180,8 @@
     Term.exists_subterm (member (op aconv) [@{term int}, @{term nat}])
 in
 fun nat_as_int ctxt =
-  map ((uses_nat_type o Thm.prop_of) ?? Conv.fconv_rule (conv ctxt)) #>
-  exists (uses_nat_int o Thm.prop_of) ?? append nat_embedding
+  map (apsnd ((uses_nat_type o Thm.prop_of) ?? Conv.fconv_rule (conv ctxt))) #>
+  exists (uses_nat_int o Thm.prop_of o snd) ?? append nat_embedding
 end
 
 
@@ -362,12 +363,13 @@
     if not (has_free_lambdas (Thm.prop_of thm)) then (thm, cx)
     else cx |> f (Thm.cprop_of thm) |>> (fn thm' => Thm.equal_elim thm' thm)
 in
-fun lift_lambdas thms ctxt =
+fun lift_lambdas irules ctxt =
   let
     val cx = (ctxt, Termtab.empty)
+    val (idxs, thms) = split_list irules
     val (thms', (ctxt', defs)) = fold_map (lift_lm (traverse [])) thms cx
     val eqs = Termtab.fold (cons o normalize_rule ctxt' o snd) defs []
-  in (eqs @ thms', ctxt') end
+  in (map (pair ~1) eqs @ (idxs ~~ thms'), ctxt') end
 end
 
 
@@ -384,8 +386,8 @@
     | (Free (n, _), ts) => add (free n) (length ts) #> fold traverse ts
     | (Abs (_, _, u), ts) => fold traverse (u :: ts)
     | (_, ts) => fold traverse ts)
-  val prune = (fn (n, (true, i)) => Symtab.update (n, i) | _ => I)
-  fun prune_tab tab = Symtab.fold prune tab Symtab.empty
+  fun prune tab = Symtab.fold (fn (n, (true, i)) =>
+    Symtab.update (n, i) | _ => I) tab Symtab.empty
 
   fun binop_conv cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2
   fun nary_conv conv1 conv2 ct =
@@ -395,7 +397,7 @@
     in conv (Symtab.update (free n, 0) tb) cx end)
   val fun_app_rule = @{lemma "f x == fun_app f x" by (simp add: fun_app_def)}
 in
-fun explicit_application ctxt thms =
+fun explicit_application ctxt irules =
   let
     fun sub_conv tb ctxt ct =
       (case Term.strip_comb (Thm.term_of ct) of
@@ -423,8 +425,8 @@
       if not (needs_exp_app tab (Thm.prop_of thm)) then thm
       else Conv.fconv_rule (sub_conv tab ctxt) thm
 
-    val tab = prune_tab (fold (traverse o Thm.prop_of) thms Symtab.empty)
-  in map (rewrite tab ctxt) thms end
+    val tab = prune (fold (traverse o Thm.prop_of o snd) irules Symtab.empty)
+  in map (apsnd (rewrite tab ctxt)) irules end
 end
 
 
@@ -465,11 +467,11 @@
         end)
 in
 
-fun datatype_selectors thms ctxt =
+fun datatype_selectors irules ctxt =
   let
-    val ns = Symtab.keys (fold (collect o Thm.prop_of) thms Symtab.empty)
+    val ns = Symtab.keys (fold (collect o Thm.prop_of o snd) irules Symtab.empty)
     val cs = fold (add_constructors (ProofContext.theory_of ctxt)) ns []
-  in (thms, fold add_selector cs ctxt) end
+  in (irules, fold add_selector cs ctxt) end
     (* FIXME: also generate hypothetical definitions for the selectors *)
 
 end
@@ -478,19 +480,20 @@
 
 (* combined normalization *)
 
-type extra_norm = thm list -> Proof.context -> thm list * Proof.context
+type extra_norm = bool -> (int * thm) list -> Proof.context ->
+  (int * thm) list * Proof.context
 
-fun with_context f thms ctxt = (f ctxt thms, ctxt)
+fun with_context f irules ctxt = (f ctxt irules, ctxt)
 
-fun normalize extra_norm with_datatypes thms ctxt =
-  thms
+fun normalize extra_norm with_datatypes irules ctxt =
+  irules
   |> trivial_distinct ctxt
   |> rewrite_bool_cases ctxt
   |> normalize_numerals ctxt
   |> nat_as_int ctxt
   |> rpair ctxt
-  |-> extra_norm
-  |-> with_context (fn cx => map (normalize_rule cx))
+  |-> extra_norm with_datatypes
+  |-> with_context (fn cx => map (apsnd (normalize_rule cx)))
   |-> SMT_Monomorph.monomorph
   |-> lift_lambdas
   |-> with_context explicit_application
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/SMT/smt_setup_solvers.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -0,0 +1,96 @@
+(*  Title:      HOL/Tools/SMT/smt_setup_solvers.ML
+    Author:     Sascha Boehme, TU Muenchen
+
+Setup SMT solvers.
+*)
+
+signature SMT_SETUP_SOLVERS =
+sig
+  val setup: theory -> theory
+end
+
+structure SMT_Setup_Solvers: SMT_SETUP_SOLVERS =
+struct
+
+fun outcome_of unsat sat unknown solver_name line =
+  if String.isPrefix unsat line then SMT_Solver.Unsat
+  else if String.isPrefix sat line then SMT_Solver.Sat
+  else if String.isPrefix unknown line then SMT_Solver.Unknown
+  else raise SMT_Solver.SMT (SMT_Solver.Other_Failure ("Solver " ^
+    quote solver_name ^ " failed, " ^ "see trace 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 (chop_while empty_line lines))
+  in (test_outcome solver_name l, ls) end
+
+
+(* CVC3 *)
+
+fun cvc3 () = {
+  name = "cvc3",
+  env_var = "CVC3_SOLVER",
+  is_remote = true,
+  options = K ["-lang", "smtlib", "-output-lang", "presentation"],
+  interface = SMTLIB_Interface.interface,
+  outcome =
+    on_first_line (outcome_of "Unsatisfiable." "Satisfiable." "Unknown."),
+  cex_parser = NONE,
+  reconstruct = NONE }
+
+
+(* Yices *)
+
+fun yices () = {
+  name = "yices",
+  env_var = "YICES_SOLVER",
+  is_remote = false,
+  options = K ["--smtlib"],
+  interface = SMTLIB_Interface.interface,
+  outcome = on_first_line (outcome_of "unsat" "sat" "unknown"),
+  cex_parser = NONE,
+  reconstruct = NONE }
+
+
+(* Z3 *)
+
+fun z3_options ctxt =
+  ["MODEL=true", "PRE_SIMPLIFY_EXPR=false", "CONTEXT_SIMPLIFIER=false", "-smt"]
+  |> not (Config.get ctxt SMT_Solver.oracle) ?
+       append ["DISPLAY_PROOF=true","PROOF_MODE=2"]
+
+fun z3_on_last_line solver_name lines =
+  let
+    fun junk l =
+      if String.isPrefix "WARNING: Out of allocated virtual memory" l
+      then raise SMT_Solver.SMT SMT_Solver.Out_Of_Memory
+      else
+        String.isPrefix "WARNING" l orelse
+        String.isPrefix "ERROR" l orelse
+        forall Symbol.is_ascii_blank (Symbol.explode l)
+  in
+    the_default ("", []) (try (swap o split_last) (filter_out junk lines))
+    |>> outcome_of "unsat" "sat" "unknown" solver_name
+  end
+
+fun z3 () = {
+  name = "z3",
+  env_var = "Z3_SOLVER",
+  is_remote = true,
+  options = z3_options,
+  interface = Z3_Interface.interface,
+  outcome = z3_on_last_line,
+  cex_parser = SOME Z3_Model.parse_counterex,
+  reconstruct = SOME Z3_Proof_Reconstruction.reconstruct }
+
+
+(* overall setup *)
+
+val setup =
+  SMT_Solver.add_solver (cvc3 ()) #>
+  SMT_Solver.add_solver (yices ()) #>
+  SMT_Solver.add_solver (z3 ())
+
+end
--- a/src/HOL/Tools/SMT/smt_solver.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/SMT/smt_solver.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -6,40 +6,60 @@
 
 signature SMT_SOLVER =
 sig
-  exception SMT of string
-  exception SMT_COUNTEREXAMPLE of bool * term list
+  datatype failure =
+    Counterexample of bool * term list |
+    Time_Out |
+    Out_Of_Memory |
+    Other_Failure of string
+  val string_of_failure: Proof.context -> string -> failure -> string
+  exception SMT of failure
 
   type interface = {
     extra_norm: SMT_Normalize.extra_norm,
     translate: SMT_Translate.config }
+  datatype outcome = Unsat | Sat | Unknown
   type solver_config = {
-    command: {env_var: string, remote_name: string option},
-    arguments: string list,
+    name: string,
+    env_var: string,
+    is_remote: bool,
+    options: Proof.context -> string list,
     interface: interface,
-    reconstruct: (string list * SMT_Translate.recon) -> Proof.context ->
-      thm * Proof.context }
+    outcome: string -> string list -> outcome * string list,
+    cex_parser: (Proof.context -> SMT_Translate.recon -> string list ->
+      term list) option,
+    reconstruct: (Proof.context -> SMT_Translate.recon -> string list ->
+      (int list * thm) * Proof.context) option }
 
   (*options*)
+  val oracle: bool Config.T
+  val filter_only: bool Config.T
+  val datatypes: bool Config.T
   val timeout: int Config.T
   val with_timeout: Proof.context -> ('a -> 'b) -> 'a -> 'b
   val trace: bool Config.T
   val trace_msg: Proof.context -> ('a -> string) -> 'a -> unit
+  val trace_used_facts: bool Config.T
 
   (*certificates*)
   val fixed_certificates: bool Config.T
   val select_certificates: string -> Context.generic -> Context.generic
 
   (*solvers*)
-  type solver = Proof.context -> thm list -> thm
-  type solver_info = Context.generic -> Pretty.T list
-  val add_solver: string * (Proof.context -> solver_config) ->
-    Context.generic -> Context.generic
+  type solver = bool option -> Proof.context -> (int * thm) list ->
+    int list * thm
+  val add_solver: solver_config -> theory -> theory
+  val set_solver_options: string -> string -> Context.generic ->
+    Context.generic
   val all_solver_names_of: Context.generic -> string list
-  val add_solver_info: string * solver_info -> Context.generic ->
-    Context.generic
   val solver_name_of: Context.generic -> string
   val select_solver: string -> Context.generic -> Context.generic
   val solver_of: Context.generic -> solver
+  val is_locally_installed: Proof.context -> bool
+
+  (*filter*)
+  val smt_filter: bool -> Time.time -> Proof.state -> ('a * thm) list -> int ->
+    {outcome: failure option, used_facts: 'a list,
+    run_time_in_msecs: int option}
 
   (*tactic*)
   val smt_tac': bool -> Proof.context -> thm list -> int -> Tactical.tactic
@@ -53,36 +73,68 @@
 structure SMT_Solver: SMT_SOLVER =
 struct
 
-exception SMT of string
-exception SMT_COUNTEREXAMPLE of bool * term list
+datatype failure =
+  Counterexample of bool * term list |
+  Time_Out |
+  Out_Of_Memory |
+  Other_Failure of string
 
+fun string_of_failure ctxt _ (Counterexample (real, ex)) =
+      let
+        val msg = (if real then "C" else "Potential c") ^ "ounterexample found"
+      in
+        if null ex then msg ^ "."
+        else Pretty.string_of (Pretty.big_list (msg ^ ":")
+          (map (Syntax.pretty_term ctxt) ex))
+      end
+  | string_of_failure _ name Time_Out = name ^ " timed out."
+  | string_of_failure _ name Out_Of_Memory = name ^ " ran out of memory."
+  | string_of_failure _ _ (Other_Failure msg) = msg
+
+exception SMT of failure
 
 type interface = {
   extra_norm: SMT_Normalize.extra_norm,
   translate: SMT_Translate.config }
 
+datatype outcome = Unsat | Sat | Unknown
+
 type solver_config = {
-  command: {env_var: string, remote_name: string option},
-  arguments: string list,
+  name: string,
+  env_var: string,
+  is_remote: bool,
+  options: Proof.context -> string list,
   interface: interface,
-  reconstruct: (string list * SMT_Translate.recon) -> Proof.context ->
-    thm * Proof.context }
+  outcome: string -> string list -> outcome * string list,
+  cex_parser: (Proof.context -> SMT_Translate.recon -> string list ->
+    term list) option,
+  reconstruct: (Proof.context -> SMT_Translate.recon -> string list ->
+    (int list * thm) * Proof.context) option }
 
 
 
 (* SMT options *)
 
+val (oracle, setup_oracle) = Attrib.config_bool "smt_oracle" (K true)
+
+val (filter_only, setup_filter_only) =
+  Attrib.config_bool "smt_filter_only" (K false)
+
+val (datatypes, setup_datatypes) = Attrib.config_bool "smt_datatypes" (K false)
+
 val (timeout, setup_timeout) = Attrib.config_int "smt_timeout" (K 30)
 
 fun with_timeout ctxt f x =
   TimeLimit.timeLimit (Time.fromSeconds (Config.get ctxt timeout)) f x
-  handle TimeLimit.TimeOut => raise SMT "timeout"
+  handle TimeLimit.TimeOut => raise SMT Time_Out
 
 val (trace, setup_trace) = Attrib.config_bool "smt_trace" (K false)
 
 fun trace_msg ctxt f x =
   if Config.get ctxt trace then tracing (f x) else ()
 
+val (trace_used_facts, setup_trace_used_facts) =
+  Attrib.config_bool "smt_trace_used_facts" (K false)
 
 
 (* SMT certificates *)
@@ -109,24 +161,34 @@
 
 (* interface to external solvers *)
 
+fun get_local_solver env_var =
+  let val local_solver = getenv env_var
+  in if local_solver <> "" then SOME local_solver else NONE end
+
 local
 
-fun choose {env_var, remote_name} =
+fun choose (rm, env_var, is_remote, name) =
   let
-    val local_solver = getenv env_var
-    val remote_solver = the_default "" remote_name
+    val force_local = (case rm of SOME false => true | _ => false)
+    val force_remote = (case rm of SOME true => true | _ => false)
+    val lsolver = get_local_solver env_var
     val remote_url = getenv "REMOTE_SMT_URL"
+    val trace = if is_some rm then K () else tracing
   in
-    if local_solver <> ""
+    if not force_remote andalso is_some lsolver
     then 
-     (tracing ("Invoking local SMT solver " ^ quote local_solver ^ " ...");
-      [local_solver])
-    else if remote_solver <> ""
+     (trace ("Invoking local SMT solver " ^ quote (the lsolver) ^ " ...");
+      [the lsolver])
+    else if not force_local andalso is_remote
     then
-     (tracing ("Invoking remote SMT solver " ^ quote remote_solver ^ " at " ^
+     (trace ("Invoking remote SMT solver " ^ quote name ^ " at " ^
         quote remote_url ^ " ...");
-      [getenv "REMOTE_SMT", remote_solver])
-    else error ("Undefined Isabelle environment variable: " ^ quote env_var)
+      [getenv "REMOTE_SMT", name])
+    else if force_remote
+    then error ("The SMT solver " ^ quote name ^ " is not remotely available.")
+    else error ("The SMT solver " ^ quote name ^ " has not been found " ^
+      "on this computer. Please set the Isabelle environment variable " ^
+      quote env_var ^ ".")
   end
 
 fun make_cmd solver args problem_path proof_path = space_implode " " (
@@ -166,6 +228,9 @@
 
 end
 
+fun trace_assms ctxt = trace_msg ctxt (Pretty.string_of o
+  Pretty.big_list "SMT assertions:" o map (Display.pretty_thm ctxt o snd))
+
 fun trace_recon_data ctxt ({typs, terms, ...} : SMT_Translate.recon) =
   let
     fun pretty_eq n p = Pretty.block [Pretty.str n, Pretty.str " = ", p]
@@ -177,46 +242,86 @@
       Pretty.big_list "functions:" (map pretty_term (Symtab.dest terms))])) ()
   end
 
-fun invoke translate_config comments command arguments thms ctxt =
-  thms
-  |> SMT_Translate.translate translate_config ctxt comments
-  ||> tap (trace_recon_data ctxt)
-  |>> run_solver ctxt command arguments
-  |> rpair ctxt
+fun invoke translate_config name cmd more_opts options irules ctxt =
+  let
+    val args = more_opts @ options ctxt
+    val comments = ("solver: " ^ name) ::
+      ("timeout: " ^ string_of_int (Config.get ctxt timeout)) ::
+      "arguments:" :: args
+  in
+    irules
+    |> tap (trace_assms ctxt)
+    |> SMT_Translate.translate translate_config ctxt comments
+    ||> tap (trace_recon_data ctxt)
+    |>> run_solver ctxt cmd args
+    |> rpair ctxt
+  end
 
 fun discharge_definitions thm =
   if Thm.nprems_of thm = 0 then thm
   else discharge_definitions (@{thm reflexive} RS thm)
 
-fun gen_solver name solver ctxt prems =
+fun set_has_datatypes with_datatypes translate =
+  let
+    val {prefixes, header, strict, builtins, serialize} = translate
+    val {builtin_typ, builtin_num, builtin_fun, has_datatypes} = builtins
+    val with_datatypes' = has_datatypes andalso with_datatypes
+    val builtins' = {builtin_typ=builtin_typ, builtin_num=builtin_num,
+      builtin_fun=builtin_fun, has_datatypes=with_datatypes}
+    val translate' = {prefixes=prefixes, header=header, strict=strict,
+      builtins=builtins', serialize=serialize}
+  in (with_datatypes', translate') end
+
+fun trace_assumptions ctxt irules idxs =
   let
-    val {command, arguments, interface, reconstruct} = solver ctxt
-    val comments = ("solver: " ^ name) ::
-      ("timeout: " ^ string_of_int (Config.get ctxt timeout)) ::
-      "arguments:" :: arguments
+    val thms = filter (fn i => i >= 0) idxs
+      |> map_filter (AList.lookup (op =) irules)
+  in
+    if Config.get ctxt trace_used_facts andalso length thms > 0
+    then
+      tracing (Pretty.string_of (Pretty.big_list "SMT used facts:"
+        (map (Display.pretty_thm ctxt) thms)))
+    else ()
+  end
+
+fun gen_solver name info rm ctxt irules =
+  let
+    val {env_var, is_remote, options, more_options, interface, reconstruct} =
+      info
     val {extra_norm, translate} = interface
-    val {builtins, ...} = translate
-    val {has_datatypes, ...} = builtins
+    val (with_datatypes, translate') =
+      set_has_datatypes (Config.get ctxt datatypes) translate
+    val cmd = (rm, env_var, is_remote, name)
   in
-    (prems, ctxt)
-    |-> SMT_Normalize.normalize extra_norm has_datatypes
-    |-> invoke translate comments command arguments
+    (irules, ctxt)
+    |-> SMT_Normalize.normalize extra_norm with_datatypes
+    |-> invoke translate' name cmd more_options options
     |-> reconstruct
-    |-> (fn thm => fn ctxt' => thm
+    |-> (fn (idxs, thm) => fn ctxt' => thm
     |> singleton (ProofContext.export ctxt' ctxt)
-    |> discharge_definitions)
+    |> discharge_definitions
+    |> tap (fn _ => trace_assumptions ctxt irules idxs)
+    |> pair idxs)
   end
 
 
 
 (* solver store *)
 
-type solver = Proof.context -> thm list -> thm
-type solver_info = Context.generic -> Pretty.T list
+type solver = bool option -> Proof.context -> (int * thm) list -> int list * thm
+
+type solver_info = {
+  env_var: string,
+  is_remote: bool,
+  options: Proof.context -> string list,
+  more_options: string list,
+  interface: interface,
+  reconstruct: string list * SMT_Translate.recon -> Proof.context ->
+    (int list * thm) * Proof.context }
 
 structure Solvers = Generic_Data
 (
-  type T = ((Proof.context -> solver_config) * solver_info) Symtab.table
+  type T = solver_info Symtab.table
   val empty = Symtab.empty
   val extend = I
   fun merge data = Symtab.merge (K true) data
@@ -224,10 +329,48 @@
 )
 
 val no_solver = "(none)"
-val add_solver = Solvers.map o Symtab.update_new o apsnd (rpair (K []))
+
+fun set_solver_options name opts = Solvers.map (Symtab.map_entry name (fn
+  {env_var, is_remote, options, interface, reconstruct, ...} =>
+  {env_var=env_var, is_remote=is_remote, options=options,
+   more_options=String.tokens (Symbol.is_ascii_blank o str) opts,
+   interface=interface, reconstruct=reconstruct}))
+
+local
+  fun finish outcome cex_parser reconstruct ocl (output, recon) ctxt =
+    (case outcome output of
+      (Unsat, ls) =>
+        if not (Config.get ctxt oracle) andalso is_some reconstruct
+        then the reconstruct ctxt recon ls
+        else (([], ocl ()), ctxt)
+    | (result, ls) =>
+        let val ts = (case cex_parser of SOME f => f ctxt recon ls | _ => [])
+        in raise SMT (Counterexample (result = Sat, ts)) end)
+in
+
+fun add_solver cfg =
+  let
+    val {name, env_var, is_remote, options, interface, outcome, cex_parser,
+      reconstruct} = cfg
+
+    fun core_oracle () = @{cprop False}
+
+    fun solver ocl = { env_var=env_var, is_remote=is_remote, options=options,
+      more_options=[], interface=interface,
+      reconstruct=finish (outcome name) cex_parser reconstruct ocl }
+  in
+    Thm.add_oracle (Binding.name name, core_oracle) #-> (fn (_, ocl) =>
+    Context.theory_map (Solvers.map (Symtab.update_new (name, solver ocl)))) #>
+    Attrib.setup (Binding.name (name ^ "_options"))
+      (Scan.lift (Parse.$$$ "=" |-- Args.name) >>
+        (Thm.declaration_attribute o K o set_solver_options name))
+      ("Additional command line options for SMT solver " ^ quote name)
+  end
+
+end
+
 val all_solver_names_of = Symtab.keys o Solvers.get
 val lookup_solver = Symtab.lookup o Solvers.get
-fun add_solver_info (n, i) = Solvers.map (Symtab.map_entry n (apsnd (K i)))
 
 
 
@@ -251,56 +394,75 @@
 fun raw_solver_of context name =
   (case lookup_solver context name of
     NONE => error "No SMT solver selected"
-  | SOME (s, _) => s)
+  | SOME s => s)
 
 fun solver_of context =
   let val name = solver_name_of context
   in gen_solver name (raw_solver_of context name) end
 
+fun is_locally_installed ctxt =
+  let
+    val name = solver_name_of (Context.Proof ctxt)
+    val {env_var, ...} = raw_solver_of (Context.Proof ctxt) name
+  in is_some (get_local_solver env_var) end
+
+
+
+(* SMT filter *)
+
+val has_topsort = Term.exists_type (Term.exists_subtype (fn
+    TFree (_, []) => true
+  | TVar (_, []) => true
+  | _ => false))
+
+fun smt_solver rm ctxt irules =
+  (* without this test, we would run into problems when atomizing the rules: *)
+  if exists (has_topsort o Thm.prop_of o snd) irules
+  then raise SMT (Other_Failure "proof state contains the universal sort {}")
+  else solver_of (Context.Proof ctxt) rm ctxt irules
+
+fun smt_filter run_remote time_limit st xrules i =
+  let
+    val {facts, goal, ...} = Proof.goal st
+    val ctxt =
+      Proof.context_of st
+      |> Config.put timeout (Time.toSeconds time_limit)
+      |> Config.put oracle false
+      |> Config.put filter_only true
+    val cprop =
+      Thm.cprem_of goal i
+      |> Thm.rhs_of o SMT_Normalize.atomize_conv ctxt
+      |> Thm.capply @{cterm Trueprop} o Thm.capply @{cterm Not} o Thm.dest_arg
+    val irs = map (pair ~1) (Thm.assume cprop :: facts)
+    val rm = SOME run_remote
+  in
+    split_list xrules
+    ||> distinct (op =) o fst o smt_solver rm ctxt o append irs o map_index I
+    |-> map_filter o try o nth
+    |> (fn xs => {outcome=NONE, used_facts=xs, run_time_in_msecs=NONE})
+  end
+  handle SMT fail => {outcome=SOME fail, used_facts=[], run_time_in_msecs=NONE}
+  (* FIXME: measure runtime *)
+
 
 
 (* SMT tactic *)
 
-local
-  fun pretty_cex ctxt (real, ex) =
-    let
-      val msg = if real then "SMT: counterexample found"
-        else "SMT: potential counterexample found"
-    in
-      if null ex then msg ^ "."
-      else Pretty.string_of (Pretty.big_list (msg ^ ":")
-        (map (Syntax.pretty_term ctxt) ex))
-    end
-
-  fun fail_tac f msg st = (f msg; Tactical.no_tac st)
-
-  fun SAFE pass_exns tac ctxt i st =
-    if pass_exns then tac ctxt i st
-    else (tac ctxt i st
-      handle SMT msg => fail_tac (trace_msg ctxt (prefix "SMT: ")) msg st
-           | SMT_COUNTEREXAMPLE ce => fail_tac tracing (pretty_cex ctxt ce) st)
-
-  fun smt_solver rules ctxt = solver_of (Context.Proof ctxt) ctxt rules
-
-  val has_topsort = Term.exists_type (Term.exists_subtype (fn
-      TFree (_, []) => true
-    | TVar (_, []) => true
-    | _ => false))
-in
 fun smt_tac' pass_exns ctxt rules =
   CONVERSION (SMT_Normalize.atomize_conv ctxt)
   THEN' Tactic.rtac @{thm ccontr}
-  THEN' SUBPROOF (fn {context, prems, ...} =>
-    let val thms = rules @ prems
+  THEN' SUBPROOF (fn {context=cx, prems, ...} =>
+    let
+      fun solve () = snd (smt_solver NONE cx (map (pair ~1) (rules @ prems)))
+      val name = solver_name_of (Context.Proof cx)
+      val trace = trace_msg cx (prefix "SMT: " o string_of_failure cx name)
     in
-      if exists (has_topsort o Thm.prop_of) thms
-      then fail_tac (trace_msg context I)
-        "SMT: proof state contains the universal sort {}"
-      else SAFE pass_exns (Tactic.rtac o smt_solver thms) context 1
+      (if pass_exns then SOME (solve ())
+       else (SOME (solve ()) handle SMT fail => (trace fail; NONE)))
+      |> (fn SOME thm => Tactic.rtac thm 1 | _ => Tactical.no_tac)
     end) ctxt
 
 val smt_tac = smt_tac' false
-end
 
 val smt_method =
   Scan.optional Attrib.thms [] >>
@@ -316,8 +478,12 @@
     (Scan.lift (Parse.$$$ "=" |-- Args.name) >>
       (Thm.declaration_attribute o K o select_solver))
     "SMT solver configuration" #>
+  setup_oracle #>
+  setup_filter_only #>
+  setup_datatypes #>
   setup_timeout #>
   setup_trace #>
+  setup_trace_used_facts #>
   setup_fixed_certificates #>
   Attrib.setup @{binding smt_certificates}
     (Scan.lift (Parse.$$$ "=" |-- Args.name) >>
@@ -332,13 +498,12 @@
     val t = string_of_int (Config.get_generic context timeout)
     val names = sort_strings (all_solver_names_of context)
     val ns = if null names then [no_solver] else names
-    val take_info = (fn (_, []) => NONE | info => SOME info)
-    val infos =
-      Solvers.get context
-      |> Symtab.dest
-      |> map_filter (fn (n, (_, info)) => take_info (n, info context))
-      |> sort (prod_ord string_ord (K EQUAL))
-      |> map (fn (n, ps) => Pretty.big_list (n ^ ":") ps)
+    val n = solver_name_of context
+    val opts =
+      (case Symtab.lookup (Solvers.get context) n of
+        NONE => []
+      | SOME {more_options, options, ...} =>
+          more_options @ options (Context.proof_of context))
     val certs_filename =
       (case get_certificates_path context of
         SOME path => Path.implode path
@@ -347,17 +512,20 @@
       else "false"
   in
     Pretty.writeln (Pretty.big_list "SMT setup:" [
-      Pretty.str ("Current SMT solver: " ^ solver_name_of context),
+      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: " ^
+        (if Config.get_generic context oracle then "false" else "true")),
       Pretty.str ("Certificates cache: " ^ certs_filename),
-      Pretty.str ("Fixed certificates: " ^ fixed),
-      Pretty.big_list "Solver-specific settings:" infos])
+      Pretty.str ("Fixed certificates: " ^ fixed)])
   end
 
 val _ =
   Outer_Syntax.improper_command "smt_status"
-    "show the available SMT solvers and the currently selected solver" Keyword.diag
+    "show the available SMT solvers and the currently selected solver"
+    Keyword.diag
     (Scan.succeed (Toplevel.no_timing o Toplevel.keep (fn state =>
       print_setup (Context.Proof (Toplevel.context_of state)))))
 
--- a/src/HOL/Tools/SMT/smt_translate.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/SMT/smt_translate.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -43,9 +43,9 @@
     typs: typ Symtab.table,
     terms: term Symtab.table,
     unfolds: thm list,
-    assms: thm list }
+    assms: (int * thm) list }
 
-  val translate: config -> Proof.context -> string list -> thm list ->
+  val translate: config -> Proof.context -> string list -> (int * thm) list ->
     string * recon
 end
 
@@ -101,7 +101,7 @@
   typs: typ Symtab.table,
   terms: term Symtab.table,
   unfolds: thm list,
-  assms: thm list }
+  assms: (int * thm) list }
 
 
 
@@ -244,8 +244,9 @@
           else as_term (in_term t)
       | _ => as_term (in_term t))
   in
-    map (normalize ctxt) #> (fn thms => ((unfold_rules, term_bool' :: thms),
-    map (in_form o prop_of) (term_bool :: thms)))
+    map (apsnd (normalize ctxt)) #> (fn irules =>
+    ((unfold_rules, (~1, term_bool') :: irules),
+     map (in_form o prop_of o snd) ((~1, term_bool) :: irules)))
   end
 
 
@@ -318,7 +319,7 @@
              (nth Tds i, map (mk_constructor ctxt (Tds, Tfs) (nth Tds i)) cs))
          end)
 
-fun relaxed thms = (([], thms), map prop_of thms)
+fun relaxed irules = (([], irules), map (prop_of o snd) irules)
 
 fun with_context header f (ths, ts) =
   let val (us, context) = fold_map f ts empty_context
--- a/src/HOL/Tools/SMT/smtlib_interface.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/SMT/smtlib_interface.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -15,7 +15,7 @@
   val add_builtins: builtins -> Context.generic -> Context.generic
   val add_logic: (term list -> string option) -> Context.generic ->
     Context.generic
-  val extra_norm: bool -> SMT_Normalize.extra_norm
+  val extra_norm: SMT_Normalize.extra_norm
   val interface: SMT_Solver.interface
 end
 
@@ -30,7 +30,7 @@
 (** facts about uninterpreted constants **)
 
 infix 2 ??
-fun (ex ?? f) thms = if exists (ex o Thm.prop_of) thms then f thms else thms
+fun (ex ?? f) irules = irules |> exists (ex o Thm.prop_of o snd) irules ? f
 
 
 (* pairs *)
@@ -40,7 +40,7 @@
 val pair_type = (fn Type (@{type_name Product_Type.prod}, _) => true | _ => false)
 val exists_pair_type = Term.exists_type (Term.exists_subtype pair_type)
 
-val add_pair_rules = exists_pair_type ?? append pair_rules
+val add_pair_rules = exists_pair_type ?? append (map (pair ~1) pair_rules)
 
 
 (* function update *)
@@ -50,7 +50,7 @@
 val is_fun_upd = (fn Const (@{const_name fun_upd}, _) => true | _ => false)
 val exists_fun_upd = Term.exists_subterm is_fun_upd
 
-val add_fun_upd_rules = exists_fun_upd ?? append fun_upd_rules
+val add_fun_upd_rules = exists_fun_upd ?? append (map (pair ~1) fun_upd_rules)
 
 
 (* abs/min/max *)
@@ -88,11 +88,11 @@
 
 (* include additional facts *)
 
-fun extra_norm has_datatypes thms ctxt =
-  thms
+fun extra_norm has_datatypes irules ctxt =
+  irules
   |> not has_datatypes ? add_pair_rules
   |> add_fun_upd_rules
-  |> map (unfold_abs_min_max_defs ctxt)
+  |> map (apsnd (unfold_abs_min_max_defs ctxt))
   |> rpair ctxt
 
 
@@ -283,7 +283,7 @@
 (** interfaces **)
 
 val interface = {
-  extra_norm = extra_norm false,
+  extra_norm = extra_norm,
   translate = {
     prefixes = {
       sort_prefix = "S",
--- a/src/HOL/Tools/SMT/yices_solver.ML	Wed Oct 27 11:10:36 2010 -0700
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,44 +0,0 @@
-(*  Title:      HOL/Tools/SMT/yices_solver.ML
-    Author:     Sascha Boehme, TU Muenchen
-
-Interface of the SMT solver Yices.
-*)
-
-signature YICES_SOLVER =
-sig
-  val setup: theory -> theory
-end
-
-structure Yices_Solver: YICES_SOLVER =
-struct
-
-val solver_name = "yices"
-val env_var = "YICES_SOLVER"
-
-val options = ["--smtlib"]
-
-fun raise_cex real = raise SMT_Solver.SMT_COUNTEREXAMPLE (real, [])
-
-fun core_oracle (output, _) =
-  let
-    val empty_line = (fn "" => true | _ => false)
-    val split_first = (fn [] => ("", []) | l :: ls => (l, ls))
-    val (l, _) = split_first (snd (chop_while empty_line output))
-  in
-    if String.isPrefix "unsat" l then @{cprop False}
-    else if String.isPrefix "sat" l then raise_cex true
-    else if String.isPrefix "unknown" l then raise_cex false
-    else raise SMT_Solver.SMT (solver_name ^ " failed")
-  end
-
-fun solver oracle _ = {
-  command = {env_var=env_var, remote_name=NONE},
-  arguments = options,
-  interface = SMTLIB_Interface.interface,
-  reconstruct = pair o oracle }
-
-val setup =
-  Thm.add_oracle (Binding.name solver_name, core_oracle) #-> (fn (_, oracle) =>
-  Context.theory_map (SMT_Solver.add_solver (solver_name, solver oracle)))
-
-end
--- a/src/HOL/Tools/SMT/z3_interface.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/SMT/z3_interface.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -8,7 +8,7 @@
 sig
   type builtin_fun = string * typ -> term list -> (string * term list) option
   val add_builtin_funs: builtin_fun -> Context.generic -> Context.generic
-  val interface: bool -> SMT_Solver.interface
+  val interface: SMT_Solver.interface
 
   datatype sym = Sym of string * sym list
   type mk_builtins = {
@@ -74,13 +74,13 @@
     | is_int_div_mod @{term "op mod :: int => _"} = true
     | is_int_div_mod _ = false
 
-  fun add_div_mod thms =
-    if exists (Term.exists_subterm is_int_div_mod o Thm.prop_of) thms
-    then [@{thm div_by_z3div}, @{thm mod_by_z3mod}] @ thms
-    else thms
+  fun add_div_mod irules =
+    if exists (Term.exists_subterm is_int_div_mod o Thm.prop_of o snd) irules
+    then [(~1, @{thm div_by_z3div}), (~1, @{thm mod_by_z3mod})] @ irules
+    else irules
 
-  fun extra_norm' has_datatypes thms =
-    SMTLIB_Interface.extra_norm has_datatypes (add_div_mod thms)
+  fun extra_norm' has_datatypes =
+    SMTLIB_Interface.extra_norm has_datatypes o add_div_mod
 
   fun z3_builtin_fun' _ (@{const_name z3div}, _) ts = SOME ("div", ts)
     | z3_builtin_fun' _ (@{const_name z3mod}, _) ts = SOME ("mod", ts)
@@ -95,8 +95,8 @@
   is_some (z3_builtin_fun' ctxt c ts) orelse 
   is_builtin_pred ctxt (n, Term.strip_type T ||> as_propT |> (op --->))
 
-fun interface has_datatypes = {
-  extra_norm = extra_norm' has_datatypes,
+val interface = {
+  extra_norm = extra_norm',
   translate = {
     prefixes = prefixes,
     strict = strict,
@@ -105,7 +105,7 @@
       builtin_typ = builtin_typ,
       builtin_num = builtin_num,
       builtin_fun = z3_builtin_fun',
-      has_datatypes = has_datatypes},
+      has_datatypes = true},
     serialize = serialize}}
 
 end
--- a/src/HOL/Tools/SMT/z3_proof_parser.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/SMT/z3_proof_parser.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -240,8 +240,8 @@
 
 (* core parser *)
 
-fun parse_exn line_no msg = raise SMT_Solver.SMT ("Z3 proof parser (line " ^
-  string_of_int line_no ^ "): " ^ msg)
+fun parse_exn line_no msg = raise SMT_Solver.SMT (SMT_Solver.Other_Failure
+  ("Z3 proof parser (line " ^ string_of_int line_no ^ "): " ^ msg))
 
 fun scan_exn msg ((line_no, _), _) = parse_exn line_no msg
 
--- a/src/HOL/Tools/SMT/z3_proof_reconstruction.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/SMT/z3_proof_reconstruction.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -7,9 +7,8 @@
 signature Z3_PROOF_RECONSTRUCTION =
 sig
   val add_z3_rule: thm -> Context.generic -> Context.generic
-  val trace_assms: bool Config.T
-  val reconstruct: string list * SMT_Translate.recon -> Proof.context ->
-    thm * Proof.context
+  val reconstruct: Proof.context -> SMT_Translate.recon -> string list ->
+    (int list * thm) * Proof.context
   val setup: theory -> theory
 end
 
@@ -20,7 +19,8 @@
 structure T = Z3_Proof_Tools
 structure L = Z3_Proof_Literals
 
-fun z3_exn msg = raise SMT_Solver.SMT ("Z3 proof reconstruction: " ^ msg)
+fun z3_exn msg = raise SMT_Solver.SMT (SMT_Solver.Other_Failure
+  ("Z3 proof reconstruction: " ^ msg))
 
 
 
@@ -137,9 +137,6 @@
 
 (* assumption *)
 
-val (trace_assms, trace_assms_setup) =
-  Attrib.config_bool "z3_trace_assms" (K false)
-
 local
   val remove_trigger = @{lemma "trigger t p == p"
     by (rule eq_reflection, rule trigger_def)}
@@ -149,28 +146,28 @@
   fun rewrite_conv ctxt eqs = Simplifier.full_rewrite
     (Simplifier.context ctxt Simplifier.empty_ss addsimps eqs)
 
-  fun rewrites ctxt eqs = map (Conv.fconv_rule (rewrite_conv ctxt eqs))
+  fun rewrites f ctxt eqs = map (f (Conv.fconv_rule (rewrite_conv ctxt eqs)))
 
-  fun trace ctxt thm =
-    if Config.get ctxt trace_assms
-    then tracing (Display.string_of_thm ctxt thm)
-    else ()
-
+  fun burrow_snd_option f (i, thm) = Option.map (pair i) (f thm)
   fun lookup_assm ctxt assms ct =
-    (case T.net_instance assms ct of
-      SOME thm => (trace ctxt thm; thm)
+    (case T.net_instance' burrow_snd_option assms ct of
+      SOME ithm => ithm
     | _ => z3_exn ("not asserted: " ^
         quote (Syntax.string_of_term ctxt (Thm.term_of ct))))
 in
 fun prepare_assms ctxt unfolds assms =
   let
-    val unfolds' = rewrites ctxt [L.rewrite_true] unfolds
-    val assms' = rewrites ctxt (union Thm.eq_thm unfolds' prep_rules) assms
-  in (unfolds', T.thm_net_of assms') end
+    val unfolds' = rewrites I ctxt [L.rewrite_true] unfolds
+    val assms' =
+      rewrites apsnd ctxt (union Thm.eq_thm unfolds' prep_rules) assms
+  in (unfolds', T.thm_net_of snd assms') end
 
 fun asserted ctxt (unfolds, assms) ct =
   let val revert_conv = rewrite_conv ctxt unfolds
-  in Thm (T.with_conv revert_conv (lookup_assm ctxt assms) ct) end
+  in Thm (T.with_conv revert_conv (snd o lookup_assm ctxt assms) ct) end
+
+fun find_assm ctxt (unfolds, assms) ct =
+  fst (lookup_assm ctxt assms (Thm.rhs_of (rewrite_conv ctxt unfolds ct)))
 end
 
 
@@ -745,91 +742,115 @@
 
 (* overall reconstruction procedure *)
 
-fun not_supported r =
-  raise Fail ("Z3: proof rule not implemented: " ^ quote (P.string_of_rule r))
-
-fun prove ctxt unfolds assms vars =
-  let
-    val assms' = prepare_assms ctxt unfolds assms
-    val simpset = T.make_simpset ctxt (Z3_Simps.get ctxt)
+local
+  fun not_supported r = raise Fail ("Z3: proof rule not implemented: " ^
+    quote (P.string_of_rule r))
 
-    fun step r ps ct (cxp as (cx, ptab)) =
-      (case (r, ps) of
-        (* core rules *)
-        (P.TrueAxiom, _) => (Thm L.true_thm, cxp)
-      | (P.Asserted, _) => (asserted cx assms' ct, cxp)
-      | (P.Goal, _) => (asserted cx assms' ct, cxp)
-      | (P.ModusPonens, [(p, _), (q, _)]) => (mp q (thm_of p), cxp)
-      | (P.ModusPonensOeq, [(p, _), (q, _)]) => (mp q (thm_of p), cxp)
-      | (P.AndElim, [(p, i)]) => and_elim (p, i) ct ptab ||> pair cx
-      | (P.NotOrElim, [(p, i)]) => not_or_elim (p, i) ct ptab ||> pair cx
-      | (P.Hypothesis, _) => (Thm (Thm.assume ct), cxp)
-      | (P.Lemma, [(p, _)]) => (lemma (thm_of p) ct, cxp)
-      | (P.UnitResolution, (p, _) :: ps) =>
-          (unit_resolution (thm_of p) (map (thm_of o fst) ps) ct, cxp)
-      | (P.IffTrue, [(p, _)]) => (iff_true (thm_of p), cxp)
-      | (P.IffFalse, [(p, _)]) => (iff_false (thm_of p), cxp)
-      | (P.Distributivity, _) => (distributivity cx ct, cxp)
-      | (P.DefAxiom, _) => (def_axiom cx ct, cxp)
-      | (P.IntroDef, _) => intro_def ct cx ||> rpair ptab
-      | (P.ApplyDef, [(p, _)]) => (apply_def (thm_of p), cxp)
-      | (P.IffOeq, [(p, _)]) => (p, cxp)
-      | (P.NnfPos, _) => (nnf cx vars (map fst ps) ct, cxp)
-      | (P.NnfNeg, _) => (nnf cx vars (map fst ps) ct, cxp)
+  fun step assms simpset vars r ps ct (cxp as (cx, ptab)) =
+    (case (r, ps) of
+      (* core rules *)
+      (P.TrueAxiom, _) => (Thm L.true_thm, cxp)
+    | (P.Asserted, _) => (asserted cx assms ct, cxp)
+    | (P.Goal, _) => (asserted cx assms ct, cxp)
+    | (P.ModusPonens, [(p, _), (q, _)]) => (mp q (thm_of p), cxp)
+    | (P.ModusPonensOeq, [(p, _), (q, _)]) => (mp q (thm_of p), cxp)
+    | (P.AndElim, [(p, i)]) => and_elim (p, i) ct ptab ||> pair cx
+    | (P.NotOrElim, [(p, i)]) => not_or_elim (p, i) ct ptab ||> pair cx
+    | (P.Hypothesis, _) => (Thm (Thm.assume ct), cxp)
+    | (P.Lemma, [(p, _)]) => (lemma (thm_of p) ct, cxp)
+    | (P.UnitResolution, (p, _) :: ps) =>
+        (unit_resolution (thm_of p) (map (thm_of o fst) ps) ct, cxp)
+    | (P.IffTrue, [(p, _)]) => (iff_true (thm_of p), cxp)
+    | (P.IffFalse, [(p, _)]) => (iff_false (thm_of p), cxp)
+    | (P.Distributivity, _) => (distributivity cx ct, cxp)
+    | (P.DefAxiom, _) => (def_axiom cx ct, cxp)
+    | (P.IntroDef, _) => intro_def ct cx ||> rpair ptab
+    | (P.ApplyDef, [(p, _)]) => (apply_def (thm_of p), cxp)
+    | (P.IffOeq, [(p, _)]) => (p, cxp)
+    | (P.NnfPos, _) => (nnf cx vars (map fst ps) ct, cxp)
+    | (P.NnfNeg, _) => (nnf cx vars (map fst ps) ct, cxp)
 
-        (* equality rules *)
-      | (P.Reflexivity, _) => (refl ct, cxp)
-      | (P.Symmetry, [(p, _)]) => (symm p, cxp)
-      | (P.Transitivity, [(p, _), (q, _)]) => (trans p q, cxp)
-      | (P.Monotonicity, _) => (monotonicity (map fst ps) ct, cxp)
-      | (P.Commutativity, _) => (commutativity ct, cxp)
+      (* equality rules *)
+    | (P.Reflexivity, _) => (refl ct, cxp)
+    | (P.Symmetry, [(p, _)]) => (symm p, cxp)
+    | (P.Transitivity, [(p, _), (q, _)]) => (trans p q, cxp)
+    | (P.Monotonicity, _) => (monotonicity (map fst ps) ct, cxp)
+    | (P.Commutativity, _) => (commutativity ct, cxp)
+
+      (* quantifier rules *)
+    | (P.QuantIntro, [(p, _)]) => (quant_intro vars p ct, cxp)
+    | (P.PullQuant, _) => (pull_quant cx ct, cxp)
+    | (P.PushQuant, _) => (push_quant cx ct, cxp)
+    | (P.ElimUnusedVars, _) => (elim_unused_vars cx ct, cxp)
+    | (P.DestEqRes, _) => (dest_eq_res cx ct, cxp)
+    | (P.QuantInst, _) => (quant_inst ct, cxp)
+    | (P.Skolemize, _) => skolemize ct cx ||> rpair ptab
+
+      (* theory rules *)
+    | (P.ThLemma, _) =>
+        (th_lemma cx simpset (map (thm_of o fst) ps) ct, cxp)
+    | (P.Rewrite, _) => (rewrite cx simpset [] ct, cxp)
+    | (P.RewriteStar, ps) =>
+        (rewrite cx simpset (map fst ps) ct, cxp)
 
-        (* quantifier rules *)
-      | (P.QuantIntro, [(p, _)]) => (quant_intro vars p ct, cxp)
-      | (P.PullQuant, _) => (pull_quant cx ct, cxp)
-      | (P.PushQuant, _) => (push_quant cx ct, cxp)
-      | (P.ElimUnusedVars, _) => (elim_unused_vars cx ct, cxp)
-      | (P.DestEqRes, _) => (dest_eq_res cx ct, cxp)
-      | (P.QuantInst, _) => (quant_inst ct, cxp)
-      | (P.Skolemize, _) => skolemize ct cx ||> rpair ptab
+    | (P.NnfStar, _) => not_supported r
+    | (P.CnfStar, _) => not_supported r
+    | (P.TransitivityStar, _) => not_supported r
+    | (P.PullQuantStar, _) => not_supported r
 
-        (* theory rules *)
-      | (P.ThLemma, _) =>
-          (th_lemma cx simpset (map (thm_of o fst) ps) ct, cxp)
-      | (P.Rewrite, _) => (rewrite cx simpset [] ct, cxp)
-      | (P.RewriteStar, ps) =>
-          (rewrite cx simpset (map fst ps) ct, cxp)
+    | _ => raise Fail ("Z3: proof rule " ^ quote (P.string_of_rule r) ^
+       " has an unexpected number of arguments."))
 
-      | (P.NnfStar, _) => not_supported r
-      | (P.CnfStar, _) => not_supported r
-      | (P.TransitivityStar, _) => not_supported r
-      | (P.PullQuantStar, _) => not_supported r
-
-      | _ => raise Fail ("Z3: proof rule " ^ quote (P.string_of_rule r) ^
-         " has an unexpected number of arguments."))
-
-    fun conclude idx rule prop (ps, cxp) =
-      trace_rule idx step rule ps prop cxp
-      |-> (fn p => apsnd (Inttab.update (idx, Proved p)) #> pair p)
+  fun prove ctxt assms vars =
+    let
+      val simpset = T.make_simpset ctxt (Z3_Simps.get ctxt)
+ 
+      fun conclude idx rule prop (ps, cxp) =
+        trace_rule idx (step assms simpset vars) rule ps prop cxp
+        |-> (fn p => apsnd (Inttab.update (idx, Proved p)) #> pair p)
+ 
+      fun lookup idx (cxp as (_, ptab)) =
+        (case Inttab.lookup ptab idx of
+          SOME (Unproved (P.Proof_Step {rule, prems, prop})) =>
+            fold_map lookup prems cxp
+            |>> map2 rpair prems
+            |> conclude idx rule prop
+        | SOME (Proved p) => (p, cxp)
+        | NONE => z3_exn ("unknown proof id: " ^ quote (string_of_int idx)))
+ 
+      fun result (p, (cx, _)) = (thm_of p, cx)
+    in
+      (fn idx => result o lookup idx o pair ctxt o Inttab.map (K Unproved))
+    end
 
-    fun lookup idx (cxp as (_, ptab)) =
-      (case Inttab.lookup ptab idx of
-        SOME (Unproved (P.Proof_Step {rule, prems, prop})) =>
-          fold_map lookup prems cxp
-          |>> map2 rpair prems
-          |> conclude idx rule prop
-      | SOME (Proved p) => (p, cxp)
-      | NONE => z3_exn ("unknown proof id: " ^ quote (string_of_int idx)))
+  fun filter_assms ctxt assms ptab =
+    let
+      fun step r ct =
+        (case r of
+          P.Asserted => insert (op =) (find_assm ctxt assms ct)
+        | P.Goal => insert (op =) (find_assm ctxt assms ct)
+        | _ => I)
 
-    fun result (p, (cx, _)) = (thm_of p, cx)
+      fun lookup idx =
+        (case Inttab.lookup ptab idx of
+          SOME (P.Proof_Step {rule, prems, prop}) =>
+            fold lookup prems #> step rule prop
+        | NONE => z3_exn ("unknown proof id: " ^ quote (string_of_int idx)))
+    in lookup end
+in
+
+fun reconstruct ctxt {typs, terms, unfolds, assms} output =
+  let
+    val (idx, (ptab, vars, cx)) = P.parse ctxt typs terms output
+    val assms' = prepare_assms cx unfolds assms
   in
-    (fn (idx, ptab) => result (lookup idx (ctxt, Inttab.map (K Unproved) ptab)))
+    if Config.get cx SMT_Solver.filter_only
+    then ((filter_assms cx assms' ptab idx [], @{thm TrueI}), cx)
+    else apfst (pair []) (prove cx assms' vars idx ptab)
   end
 
-fun reconstruct (output, {typs, terms, unfolds, assms}) ctxt =
-  P.parse ctxt typs terms output
-  |> (fn (idx, (ptab, vars, cx)) => prove cx unfolds assms vars (idx, ptab))
+end
 
-val setup = trace_assms_setup #> z3_rules_setup #> Z3_Simps.setup
+val setup = z3_rules_setup #> Z3_Simps.setup
 
 end
--- a/src/HOL/Tools/SMT/z3_proof_tools.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/SMT/z3_proof_tools.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -13,7 +13,9 @@
   val as_meta_eq: cterm -> cterm
 
   (* theorem nets *)
-  val thm_net_of: thm list -> thm Net.net
+  val thm_net_of: ('a -> thm) -> 'a list -> 'a Net.net
+  val net_instance': ((thm -> thm option) -> 'a -> 'a option) -> 'a Net.net ->
+    cterm -> 'a option
   val net_instance: thm Net.net -> cterm -> thm option
 
   (* proof combinators *)
@@ -67,16 +69,18 @@
 
 (* theorem nets *)
 
-fun thm_net_of thms =
-  let fun insert thm = Net.insert_term (K false) (Thm.prop_of thm, thm)
-  in fold insert thms Net.empty end
+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)
 
-fun first_of thms ct = get_first (maybe_instantiate ct) thms
-fun net_instance net ct = first_of (Net.match_term net (Thm.term_of ct)) ct
+fun net_instance' f net ct =
+  let fun first_of f xthms ct = get_first (f (maybe_instantiate ct)) xthms 
+  in first_of f (Net.match_term net (Thm.term_of ct)) ct end
+val net_instance = net_instance' I
 
 
 
--- a/src/HOL/Tools/SMT/z3_solver.ML	Wed Oct 27 11:10:36 2010 -0700
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,86 +0,0 @@
-(*  Title:      HOL/Tools/SMT/z3_solver.ML
-    Author:     Sascha Boehme, TU Muenchen
-
-Interface of the SMT solver Z3.
-*)
-
-signature Z3_SOLVER =
-sig
-  val proofs: bool Config.T
-  val options: string Config.T
-  val datatypes: bool Config.T
-  val setup: theory -> theory
-end
-
-structure Z3_Solver: Z3_SOLVER =
-struct
-
-val solver_name = "z3"
-val env_var = "Z3_SOLVER"
-
-val (proofs, proofs_setup) = Attrib.config_bool "z3_proofs" (K false)
-val (options, options_setup) = Attrib.config_string "z3_options" (K "")
-val (datatypes, datatypes_setup) = Attrib.config_bool "z3_datatypes" (K false)
-
-fun add xs ys = ys @ xs
-
-fun explode_options s = String.tokens (Symbol.is_ascii_blank o str) s
-
-fun get_options ctxt =
-  ["MODEL=true", "PRE_SIMPLIFY_EXPR=false", "CONTEXT_SIMPLIFIER=false"]
-  |> Config.get ctxt proofs ? add ["DISPLAY_PROOF=true", "PROOF_MODE=2"]
-  |> add (explode_options (Config.get ctxt options))
-
-fun pretty_config context = [
-  Pretty.str ("With proofs: " ^
-    (if Config.get_generic context proofs then "true" else "false")),
-  Pretty.str ("Options: " ^
-    space_implode " " (get_options (Context.proof_of context))) ]
-
-fun cmdline_options ctxt =
-  get_options ctxt
-  |> add ["-smt"]
-
-fun raise_cex ctxt real recon ls =
-  let val cex = Z3_Model.parse_counterex ctxt recon ls
-  in raise SMT_Solver.SMT_COUNTEREXAMPLE (real, cex) end
-
-fun if_unsat f (output, recon) ctxt =
-  let
-    fun jnk l =
-      String.isPrefix "WARNING" l orelse
-      String.isPrefix "ERROR" l orelse
-      forall Symbol.is_ascii_blank (Symbol.explode l)
-    val (ls, l) = the_default ([], "") (try split_last (filter_out jnk output))
-  in
-    if String.isPrefix "unsat" l then f (ls, recon) ctxt
-    else if String.isPrefix "sat" l then raise_cex ctxt true recon ls
-    else if String.isPrefix "unknown" l then raise_cex ctxt false recon ls
-    else raise SMT_Solver.SMT (solver_name ^ " failed")
-  end
-
-val core_oracle = uncurry (if_unsat (K (K @{cprop False})))
-
-val prover = if_unsat Z3_Proof_Reconstruction.reconstruct
-
-fun solver oracle ctxt =
-  let
-    val with_datatypes = Config.get ctxt datatypes
-    val with_proof = not with_datatypes andalso Config.get ctxt proofs
-                     (* FIXME: implement proof reconstruction for datatypes *)
-  in
-   {command = {env_var=env_var, remote_name=SOME solver_name},
-    arguments = cmdline_options ctxt,
-    interface = Z3_Interface.interface with_datatypes,
-    reconstruct = if with_proof then prover else (fn r => `(oracle o pair r))}
-  end
-
-val setup =
-  proofs_setup #>
-  options_setup #>
-  datatypes_setup #>
-  Thm.add_oracle (Binding.name solver_name, core_oracle) #-> (fn (_, oracle) =>
-  Context.theory_map (SMT_Solver.add_solver (solver_name, solver oracle))) #>
-  Context.theory_map (SMT_Solver.add_solver_info (solver_name, pretty_config))
-
-end
--- a/src/HOL/Tools/Sledgehammer/sledgehammer.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -8,11 +8,11 @@
 
 signature SLEDGEHAMMER =
 sig
-  type failure = ATP_Systems.failure
+  type failure = ATP_Proof.failure
   type locality = Sledgehammer_Filter.locality
   type relevance_fudge = Sledgehammer_Filter.relevance_fudge
   type relevance_override = Sledgehammer_Filter.relevance_override
-  type prepared_formula = Sledgehammer_ATP_Translate.prepared_formula
+  type translated_formula = Sledgehammer_ATP_Translate.translated_formula
   type minimize_command = Sledgehammer_ATP_Reconstruct.minimize_command
 
   type params =
@@ -30,21 +30,21 @@
      timeout: Time.time,
      expect: string}
 
-  datatype axiom =
-    Unprepared of (string * locality) * thm |
-    Prepared of term * ((string * locality) * prepared_formula) option
+  datatype fact =
+    Untranslated of (string * locality) * thm |
+    Translated of term * ((string * locality) * translated_formula) option
 
   type prover_problem =
     {state: Proof.state,
      goal: thm,
      subgoal: int,
      subgoal_count: int,
-     axioms: axiom list,
+     facts: fact list,
      only: bool}
 
   type prover_result =
     {outcome: failure option,
-     used_axioms: (string * locality) list,
+     used_facts: (string * locality) list,
      run_time_in_msecs: int option,
      message: string}
 
@@ -99,7 +99,7 @@
 
 fun is_prover_installed ctxt name =
   let val thy = ProofContext.theory_of ctxt in
-    if is_smt_prover name then true (* FIXME *)
+    if is_smt_prover name then SMT_Solver.is_locally_installed ctxt
     else is_atp_installed thy name
   end
 
@@ -115,7 +115,27 @@
 val atp_irrelevant_consts =
   [@{const_name False}, @{const_name True}, @{const_name If}, @{const_name Let},
    @{const_name HOL.eq}]
-val smt_irrelevant_consts = atp_irrelevant_consts (* FIXME *)
+
+(* FIXME: check types *)
+val smt_irrelevant_consts =
+  atp_irrelevant_consts @
+  [@{const_name distinct},
+   (* numeral-related: *)
+   @{const_name number_of}, @{const_name Int.Pls}, @{const_name Int.Min},
+   @{const_name Int.Bit0}, @{const_name Int.Bit1}, @{const_name Suc},
+   (* FIXME: @{const_name Numeral0}, @{const_name Numeral1}, *)
+   (* int => nat *)
+   @{const_name nat},
+   (* nat => int *)
+   (* FIXME: @{const_name int}, *)
+   (* for "nat", "int", and "real": *)
+   @{const_name plus}, @{const_name uminus}, @{const_name minus},
+   @{const_name times}, @{const_name less}, @{const_name less_eq},
+   @{const_name abs}, @{const_name min}, @{const_name max},
+   (* for "nat" and "div": *)
+   @{const_name div}, @{const_name mod} (* , *)
+   (* for real: *)
+   (* FIXME: @{const_name "op /"} *)]
 
 fun irrelevant_consts_for_prover name =
   if is_smt_prover name then smt_irrelevant_consts else atp_irrelevant_consts
@@ -169,8 +189,8 @@
       sort_strings (available_atps thy) @ smt_prover_names
       |> List.partition (String.isPrefix remote_prefix)
   in
-    priority ("Available provers: " ^ commas (local_provers @ remote_provers) ^
-              ".")
+    Output.urgent_message ("Available provers: " ^
+                           commas (local_provers @ remote_provers) ^ ".")
   end
 
 fun kill_provers () = Async_Manager.kill_threads das_Tool "provers"
@@ -194,22 +214,22 @@
    timeout: Time.time,
    expect: string}
 
-datatype axiom =
-  Unprepared of (string * locality) * thm |
-  Prepared of term * ((string * locality) * prepared_formula) option
+datatype fact =
+  Untranslated of (string * locality) * thm |
+  Translated of term * ((string * locality) * translated_formula) option
 
 type prover_problem =
   {state: Proof.state,
    goal: thm,
    subgoal: int,
    subgoal_count: int,
-   axioms: axiom list,
+   facts: fact list,
    only: bool}
 
 type prover_result =
   {outcome: failure option,
    message: string,
-   used_axioms: (string * locality) list,
+   used_facts: (string * locality) list,
    run_time_in_msecs: int option}
 
 type prover = params -> minimize_command -> prover_problem -> prover_result
@@ -232,11 +252,11 @@
   |> Exn.release
   |> tap (after path)
 
-fun prover_description ctxt ({blocking, verbose, ...} : params) name num_axioms
-                       i n goal =
+fun prover_description ctxt ({blocking, verbose, ...} : params) name num_facts i
+                       n goal =
   quote name ^
   (if verbose then
-     " with " ^ string_of_int num_axioms ^ " fact" ^ plural_s num_axioms
+     " with " ^ string_of_int num_facts ^ " fact" ^ plural_s num_facts
    else
      "") ^
   " on " ^ (if n = 1 then "goal" else "subgoal " ^ string_of_int i) ^ ":" ^
@@ -250,10 +270,12 @@
 
 (* generic TPTP-based ATPs *)
 
-fun dest_Unprepared (Unprepared p) = p
-  | dest_Unprepared (Prepared _) = raise Fail "dest_Unprepared"
-fun prepared_axiom ctxt (Unprepared p) = prepare_axiom ctxt p
-  | prepared_axiom _ (Prepared p) = p
+fun dest_Untranslated (Untranslated p) = p
+  | dest_Untranslated (Translated _) = raise Fail "dest_Untranslated"
+fun translated_fact ctxt (Untranslated p) = translate_fact ctxt p
+  | translated_fact _ (Translated p) = p
+fun fact_name (Untranslated ((name, _), _)) = SOME name
+  | fact_name (Translated (_, p)) = Option.map (fst o fst) p
 
 fun int_option_add (SOME m) (SOME n) = SOME (m + n)
   | int_option_add _ _ = NONE
@@ -269,13 +291,13 @@
         ({debug, verbose, overlord, full_types, explicit_apply,
           max_relevant, isar_proof, isar_shrink_factor, timeout, ...} : params)
         minimize_command
-        ({state, goal, subgoal, axioms, only, ...} : prover_problem) =
+        ({state, goal, subgoal, facts, only, ...} : prover_problem) =
   let
     val ctxt = Proof.context_of state
     val (_, hyp_ts, concl_t) = strip_subgoal goal subgoal
-    val axioms =
-      axioms |> not only ? take (the_default default_max_relevant max_relevant)
-             |> map (prepared_axiom ctxt)
+    val facts =
+      facts |> not only ? take (the_default default_max_relevant max_relevant)
+            |> map (translated_fact ctxt)
     val dest_dir = if overlord then getenv "ISABELLE_HOME_USER"
                    else Config.get ctxt dest_dir
     val problem_prefix = Config.get ctxt problem_prefix
@@ -334,9 +356,9 @@
                       proof_delims known_failures output
               in (output, msecs, tstplike_proof, outcome) end
             val readable_names = debug andalso overlord
-            val (atp_problem, pool, conjecture_offset, axiom_names) =
+            val (atp_problem, pool, conjecture_offset, fact_names) =
               prepare_atp_problem ctxt readable_names explicit_forall full_types
-                                  explicit_apply hyp_ts concl_t axioms
+                                  explicit_apply hyp_ts concl_t facts
             val ss = tptp_strings_for_atp_problem use_conjecture_for_hypotheses
                                                   atp_problem
             val _ = File.write_list probfile ss
@@ -358,7 +380,7 @@
                               (output, int_option_add msecs0 msecs,
                                tstplike_proof, outcome))
                      | result => result)
-          in ((pool, conjecture_shape, axiom_names), result) end
+          in ((pool, conjecture_shape, fact_names), result) end
         else
           error ("Bad executable: " ^ Path.implode command ^ ".")
 
@@ -371,24 +393,23 @@
         ()
       else
         File.write (Path.explode (Path.implode probfile ^ "_proof")) output
-    val ((pool, conjecture_shape, axiom_names),
+    val ((pool, conjecture_shape, fact_names),
          (output, msecs, tstplike_proof, outcome)) =
       with_path cleanup export run_on problem_path_name
-    val (conjecture_shape, axiom_names) =
-      repair_conjecture_shape_and_axiom_names output conjecture_shape
-                                              axiom_names
+    val (conjecture_shape, fact_names) =
+      repair_conjecture_shape_and_fact_names output conjecture_shape fact_names
     val important_message =
       if random () <= important_message_keep_factor then
         extract_important_message output
       else
         ""
-    val (message, used_axioms) =
+    val (message, used_facts) =
       case outcome of
         NONE =>
         proof_text isar_proof
             (pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
             (proof_banner auto, full_types, minimize_command, tstplike_proof,
-             axiom_names, goal, subgoal)
+             fact_names, goal, subgoal)
         |>> (fn message =>
                 message ^ (if verbose then
                              "\nATP real CPU time: " ^
@@ -402,43 +423,47 @@
                    ""))
       | SOME failure => (string_for_failure failure, [])
   in
-    {outcome = outcome, message = message, used_axioms = used_axioms,
+    {outcome = outcome, message = message, used_facts = used_facts,
      run_time_in_msecs = msecs}
   end
 
-(* FIXME: dummy *)
-fun saschas_run_smt_solver remote timeout state axioms i =
-  (OS.Process.sleep (Time.fromMilliseconds 1500);
-   {outcome = NONE, used_axioms = axioms |> take 5 |> map fst,
-    run_time_in_msecs = NONE})
+fun failure_from_smt_failure (SMT_Solver.Counterexample _) = Unprovable
+  | failure_from_smt_failure SMT_Solver.Time_Out = TimedOut
+  | failure_from_smt_failure SMT_Solver.Out_Of_Memory = OutOfResources
+  | failure_from_smt_failure _ = UnknownError
 
 fun run_smt_solver remote ({timeout, ...} : params) minimize_command
-                   ({state, subgoal, subgoal_count, axioms, ...}
+                   ({state, subgoal, subgoal_count, facts, ...}
                     : prover_problem) =
   let
-    val {outcome, used_axioms, run_time_in_msecs} =
-      saschas_run_smt_solver remote timeout state
-                             (map_filter (try dest_Unprepared) axioms) subgoal
+    val ctxt = Proof.context_of state
+    val {outcome, used_facts, run_time_in_msecs} =
+      SMT_Solver.smt_filter remote timeout state
+                            (map_filter (try dest_Untranslated) facts) subgoal
     val message =
-      if outcome = NONE then
+      case outcome of
+        NONE =>
         try_command_line (proof_banner false)
                          (apply_on_subgoal subgoal subgoal_count ^
-                          command_call smtN (map fst used_axioms)) ^
-        minimize_line minimize_command (map fst used_axioms)
-      else
-        ""
+                          command_call smtN (map fst used_facts)) ^
+        minimize_line minimize_command (map fst used_facts)
+      | SOME SMT_Solver.Time_Out => "Timed out."
+      | SOME (SMT_Solver.Counterexample _) => "The SMT problem is unprovable."
+      | SOME failure =>
+        SMT_Solver.string_of_failure ctxt "The SMT solver" failure
+    val outcome = outcome |> Option.map failure_from_smt_failure (* FIXME *)
   in
-    {outcome = outcome, used_axioms = used_axioms,
+    {outcome = outcome, used_facts = used_facts,
      run_time_in_msecs = run_time_in_msecs, message = message}
   end
 
 fun get_prover thy auto name =
-  if is_smt_prover name then run_smt_solver (String.isPrefix remote_prefix)
+  if is_smt_prover name then run_smt_solver (String.isPrefix remote_prefix name)
   else run_atp auto name (get_atp thy name)
 
 fun run_prover (params as {blocking, debug, max_relevant, timeout, expect, ...})
                auto minimize_command
-               (problem as {state, goal, subgoal, subgoal_count, axioms, ...})
+               (problem as {state, goal, subgoal, subgoal_count, facts, ...})
                name =
   let
     val thy = Proof.theory_of state
@@ -447,9 +472,9 @@
     val death_time = Time.+ (birth_time, timeout)
     val max_relevant =
       the_default (default_max_relevant_for_prover thy name) max_relevant
-    val num_axioms = Int.min (length axioms, max_relevant)
+    val num_facts = Int.min (length facts, max_relevant)
     val desc =
-      prover_description ctxt params name num_axioms subgoal subgoal_count goal
+      prover_description ctxt params name num_facts subgoal subgoal_count goal
     fun go () =
       let
         fun really_go () =
@@ -481,7 +506,7 @@
       end
     else if blocking then
       let val (success, message) = TimeLimit.timeLimit timeout go () in
-        List.app priority
+        List.app Output.urgent_message
                  (Async_Manager.break_into_chunks [desc ^ "\n" ^ message]);
         (success, state)
       end
@@ -490,25 +515,26 @@
        (false, state))
   end
 
-fun run_sledgehammer (params as {blocking, provers, full_types,
+fun run_sledgehammer (params as {blocking, debug, provers, full_types,
                                  relevance_thresholds, max_relevant, ...})
                      auto i (relevance_override as {only, ...}) minimize_command
                      state =
   if null provers then
     error "No prover is set."
   else case subgoal_count state of
-    0 => (priority "No subgoal!"; (false, state))
+    0 => (Output.urgent_message "No subgoal!"; (false, state))
   | n =>
     let
       val _ = Proof.assert_backward state
       val thy = Proof.theory_of state
-      val {context = ctxt, facts = chained_ths, goal} = Proof.goal state
+      val ctxt = Proof.context_of state
+      val {facts = chained_ths, goal, ...} = Proof.goal state
       val (_, hyp_ts, concl_t) = strip_subgoal goal i
       val _ = () |> not blocking ? kill_provers
-      val _ = if auto then () else priority "Sledgehammering..."
+      val _ = if auto then () else Output.urgent_message "Sledgehammering..."
       val (smts, atps) = provers |> List.partition is_smt_prover
-      fun run_provers full_types irrelevant_consts relevance_fudge maybe_prepare
-                      provers (res as (success, state)) =
+      fun run_provers label full_types irrelevant_consts relevance_fudge
+                      maybe_translate provers (res as (success, state)) =
         if success orelse null provers then
           res
         else
@@ -520,16 +546,27 @@
                 0 |> fold (Integer.max o default_max_relevant_for_prover thy)
                           provers
                   |> auto ? (fn n => n div auto_max_relevant_divisor)
-            val axioms =
+            val facts =
               relevant_facts ctxt full_types relevance_thresholds
                              max_max_relevant irrelevant_consts relevance_fudge
                              relevance_override chained_ths hyp_ts concl_t
-              |> map maybe_prepare
+              |> map maybe_translate
             val problem =
               {state = state, goal = goal, subgoal = i, subgoal_count = n,
-               axioms = axioms, only = only}
+               facts = facts, only = only}
             val run_prover = run_prover params auto minimize_command
           in
+            if debug then
+              Output.urgent_message (label ^ ": " ^
+                  (if null facts then
+                     "Found no relevant facts."
+                   else
+                     "Including (up to) " ^ string_of_int (length facts) ^
+                     " relevant fact" ^ plural_s (length facts) ^ ":\n" ^
+                     (facts |> map_filter fact_name
+                            |> space_implode " ") ^ "."))
+            else
+              ();
             if auto then
               fold (fn prover => fn (true, state) => (true, state)
                                   | (false, _) => run_prover problem prover)
@@ -540,11 +577,11 @@
                       |> exists fst |> rpair state
           end
       val run_atps =
-        run_provers full_types atp_irrelevant_consts atp_relevance_fudge
-                    (Prepared o prepare_axiom ctxt) atps
+        run_provers "ATPs" full_types atp_irrelevant_consts atp_relevance_fudge
+                    (Translated o translate_fact ctxt) atps
       val run_smts =
-        run_provers true smt_irrelevant_consts smt_relevance_fudge Unprepared
-                    smts
+        run_provers "SMT" true smt_irrelevant_consts smt_relevance_fudge
+                    Untranslated smts
       fun run_atps_and_smt_solvers () =
         [run_atps, run_smts] |> Par_List.map (fn f => f (false, state) |> K ())
     in
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_atp_reconstruct.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_atp_reconstruct.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -1,4 +1,4 @@
-(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_reconstruct.ML
+(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_atp_reconstruct.ML
     Author:     Lawrence C. Paulson, Cambridge University Computer Laboratory
     Author:     Claire Quigley, Cambridge University Computer Laboratory
     Author:     Jasmin Blanchette, TU Muenchen
@@ -17,7 +17,7 @@
     string Symtab.table * bool * int * Proof.context * int list list
   type text_result = string * (string * locality) list
 
-  val repair_conjecture_shape_and_axiom_names :
+  val repair_conjecture_shape_and_fact_names :
     string -> int list list -> (string * locality) list vector
     -> int list list * (string * locality) list vector
   val apply_on_subgoal : int -> int -> string
@@ -57,7 +57,7 @@
 
 (** SPASS's Flotter hack **)
 
-(* This is a hack required for keeping track of axioms after they have been
+(* This is a hack required for keeping track of facts after they have been
    clausified by SPASS's Flotter tool. The "ATP/scripts/spass" script is also
    part of this hack. *)
 
@@ -84,8 +84,7 @@
   #> explode #> filter_out Symbol.is_blank #> parse_clause_formula_relation
   #> fst
 
-fun repair_conjecture_shape_and_axiom_names output conjecture_shape
-                                            axiom_names =
+fun repair_conjecture_shape_and_fact_names output conjecture_shape fact_names =
   if String.isSubstring set_ClauseFormulaRelationN output then
     let
       val j0 = hd (hd conjecture_shape)
@@ -97,10 +96,9 @@
         |> map (fn s => find_index (curry (op =) s) seq + 1)
       fun names_for_number j =
         j |> AList.lookup (op =) name_map |> these
-          |> map_filter (try (unprefix axiom_prefix)) |> map unascii_of
+          |> map_filter (try (unprefix fact_prefix)) |> map unascii_of
           |> map (fn name =>
-                     (name, name |> find_first_in_list_vector axiom_names
-                                 |> the)
+                     (name, name |> find_first_in_list_vector fact_names |> the)
                      handle Option.Option =>
                             error ("No such fact: " ^ quote name ^ "."))
     in
@@ -108,7 +106,7 @@
        seq |> map names_for_number |> Vector.fromList)
     end
   else
-    (conjecture_shape, axiom_names)
+    (conjecture_shape, fact_names)
 
 
 (** Soft-core proof reconstruction: Metis one-liner **)
@@ -139,38 +137,37 @@
       "\nTo minimize the number of lemmas, try this: " ^
       Markup.markup Markup.sendback command ^ "."
 
-fun resolve_axiom axiom_names ((_, SOME s)) =
-    (case strip_prefix_and_unascii axiom_prefix s of
-       SOME s' => (case find_first_in_list_vector axiom_names s' of
+fun resolve_fact fact_names ((_, SOME s)) =
+    (case strip_prefix_and_unascii fact_prefix s of
+       SOME s' => (case find_first_in_list_vector fact_names s' of
                      SOME x => [(s', x)]
                    | NONE => [])
      | NONE => [])
-  | resolve_axiom axiom_names (num, NONE) =
+  | resolve_fact fact_names (num, NONE) =
     case Int.fromString num of
       SOME j =>
-      if j > 0 andalso j <= Vector.length axiom_names then
-        Vector.sub (axiom_names, j - 1)
+      if j > 0 andalso j <= Vector.length fact_names then
+        Vector.sub (fact_names, j - 1)
       else
         []
     | NONE => []
 
-fun add_fact axiom_names (Inference (name, _, [])) =
-    append (resolve_axiom axiom_names name)
+fun add_fact fact_names (Inference (name, _, [])) =
+    append (resolve_fact fact_names name)
   | add_fact _ _ = I
 
-fun used_facts_in_tstplike_proof axiom_names =
-  atp_proof_from_tstplike_string #> rpair [] #-> fold (add_fact axiom_names)
+fun used_facts_in_tstplike_proof fact_names =
+  atp_proof_from_tstplike_string #> rpair [] #-> fold (add_fact fact_names)
 
-fun used_facts axiom_names =
-  used_facts_in_tstplike_proof axiom_names
+fun used_facts fact_names =
+  used_facts_in_tstplike_proof fact_names
   #> List.partition (curry (op =) Chained o snd)
   #> pairself (sort_distinct (string_ord o pairself fst))
 
 fun metis_proof_text (banner, full_types, minimize_command,
-                      tstplike_proof, axiom_names, goal, i) =
+                      tstplike_proof, fact_names, goal, i) =
   let
-    val (chained_lemmas, other_lemmas) =
-      used_facts axiom_names tstplike_proof
+    val (chained_lemmas, other_lemmas) = used_facts fact_names tstplike_proof
     val n = Logic.count_prems (prop_of goal)
   in
     (metis_line banner full_types i n (map fst other_lemmas) ^
@@ -233,7 +230,7 @@
                       | NONE => ~1
   in if k >= 0 then [k] else [] end
 
-fun is_axiom conjecture_shape = not o null o resolve_axiom conjecture_shape
+fun is_fact conjecture_shape = not o null o resolve_fact conjecture_shape
 fun is_conjecture conjecture_shape = not o null o resolve_conjecture conjecture_shape
 
 fun raw_label_for_name conjecture_shape name =
@@ -491,14 +488,14 @@
   | replace_dependencies_in_line p (Inference (name, t, deps)) =
     Inference (name, t, fold (union (op =) o replace_one_dependency p) deps [])
 
-(* Discard axioms; consolidate adjacent lines that prove the same formula, since
+(* Discard facts; consolidate adjacent lines that prove the same formula, since
    they differ only in type information.*)
 fun add_line _ _ (line as Definition _) lines = line :: lines
-  | add_line conjecture_shape axiom_names (Inference (name, t, [])) lines =
-    (* No dependencies: axiom, conjecture, or (for Vampire) internal axioms or
+  | add_line conjecture_shape fact_names (Inference (name, t, [])) lines =
+    (* No dependencies: fact, conjecture, or (for Vampire) internal facts or
        definitions. *)
-    if is_axiom axiom_names name then
-      (* Axioms are not proof lines. *)
+    if is_fact fact_names name then
+      (* Facts are not proof lines. *)
       if is_only_type_information t then
         map (replace_dependencies_in_line (name, [])) lines
       (* Is there a repetition? If so, replace later line by earlier one. *)
@@ -541,10 +538,10 @@
 
 fun add_desired_line _ _ _ _ (line as Definition (name, _, _)) (j, lines) =
     (j, line :: map (replace_dependencies_in_line (name, [])) lines)
-  | add_desired_line isar_shrink_factor conjecture_shape axiom_names frees
+  | add_desired_line isar_shrink_factor conjecture_shape fact_names frees
                      (Inference (name, t, deps)) (j, lines) =
     (j + 1,
-     if is_axiom axiom_names name orelse
+     if is_fact fact_names name orelse
         is_conjecture conjecture_shape name orelse
         (* the last line must be kept *)
         j = 0 orelse
@@ -580,20 +577,20 @@
 fun smart_case_split [] facts = ByMetis facts
   | smart_case_split proofs facts = CaseSplit (proofs, facts)
 
-fun add_fact_from_dependency conjecture_shape axiom_names name =
-  if is_axiom axiom_names name then
-    apsnd (union (op =) (map fst (resolve_axiom axiom_names name)))
+fun add_fact_from_dependency conjecture_shape fact_names name =
+  if is_fact fact_names name then
+    apsnd (union (op =) (map fst (resolve_fact fact_names name)))
   else
     apfst (insert (op =) (raw_label_for_name conjecture_shape name))
 
 fun step_for_line _ _ _ (Definition (_, t1, t2)) = Let (t1, t2)
   | step_for_line conjecture_shape _ _ (Inference (name, t, [])) =
     Assume (raw_label_for_name conjecture_shape name, t)
-  | step_for_line conjecture_shape axiom_names j (Inference (name, t, deps)) =
+  | step_for_line conjecture_shape fact_names j (Inference (name, t, deps)) =
     Have (if j = 1 then [Show] else [],
           raw_label_for_name conjecture_shape name,
           fold_rev forall_of (map Var (Term.add_vars t [])) t,
-          ByMetis (fold (add_fact_from_dependency conjecture_shape axiom_names)
+          ByMetis (fold (add_fact_from_dependency conjecture_shape fact_names)
                         deps ([], [])))
 
 fun repair_name "$true" = "c_True"
@@ -606,8 +603,9 @@
     else
       s
 
-fun isar_proof_from_tstplike_proof pool ctxt full_types tfrees isar_shrink_factor
-        tstplike_proof conjecture_shape axiom_names params frees =
+fun isar_proof_from_tstplike_proof pool ctxt full_types tfrees
+        isar_shrink_factor tstplike_proof conjecture_shape fact_names params
+        frees =
   let
     val lines =
       tstplike_proof
@@ -615,14 +613,14 @@
       |> nasty_atp_proof pool
       |> map_term_names_in_atp_proof repair_name
       |> decode_lines ctxt full_types tfrees
-      |> rpair [] |-> fold_rev (add_line conjecture_shape axiom_names)
+      |> rpair [] |-> fold_rev (add_line conjecture_shape fact_names)
       |> rpair [] |-> fold_rev add_nontrivial_line
       |> rpair (0, []) |-> fold_rev (add_desired_line isar_shrink_factor
-                                             conjecture_shape axiom_names frees)
+                                             conjecture_shape fact_names frees)
       |> snd
   in
     (if null params then [] else [Fix params]) @
-    map2 (step_for_line conjecture_shape axiom_names) (length lines downto 1)
+    map2 (step_for_line conjecture_shape fact_names) (length lines downto 1)
          lines
   end
 
@@ -915,7 +913,7 @@
 
 fun isar_proof_text (pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
                     (other_params as (_, full_types, _, tstplike_proof,
-                                      axiom_names, goal, i)) =
+                                      fact_names, goal, i)) =
   let
     val (params, hyp_ts, concl_t) = strip_subgoal goal i
     val frees = fold Term.add_frees (concl_t :: hyp_ts) []
@@ -924,7 +922,7 @@
     val (one_line_proof, lemma_names) = metis_proof_text other_params
     fun isar_proof_for () =
       case isar_proof_from_tstplike_proof pool ctxt full_types tfrees
-               isar_shrink_factor tstplike_proof conjecture_shape axiom_names
+               isar_shrink_factor tstplike_proof conjecture_shape fact_names
                params frees
            |> redirect_proof hyp_ts concl_t
            |> kill_duplicate_assumptions_in_proof
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_atp_translate.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_atp_translate.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -1,4 +1,4 @@
-(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_translate.ML
+(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_atp_translate.ML
     Author:     Fabian Immler, TU Muenchen
     Author:     Makarius
     Author:     Jasmin Blanchette, TU Muenchen
@@ -9,16 +9,16 @@
 signature SLEDGEHAMMER_ATP_TRANSLATE =
 sig
   type 'a problem = 'a ATP_Problem.problem
-  type prepared_formula
+  type translated_formula
 
-  val axiom_prefix : string
+  val fact_prefix : string
   val conjecture_prefix : string
-  val prepare_axiom :
+  val translate_fact :
     Proof.context -> (string * 'a) * thm
-    -> term * ((string * 'a) * prepared_formula) option
+    -> term * ((string * 'a) * translated_formula) option
   val prepare_atp_problem :
     Proof.context -> bool -> bool -> bool -> bool -> term list -> term
-    -> (term * ((string * 'a) * prepared_formula) option) list
+    -> (term * ((string * 'a) * translated_formula) option) list
     -> string problem * string Symtab.table * int * (string * 'a) list vector
 end;
 
@@ -29,7 +29,7 @@
 open Metis_Translate
 open Sledgehammer_Util
 
-val axiom_prefix = "ax_"
+val fact_prefix = "fact_"
 val conjecture_prefix = "conj_"
 val helper_prefix = "help_"
 val class_rel_clause_prefix = "clrel_";
@@ -39,7 +39,7 @@
 (* Freshness almost guaranteed! *)
 val sledgehammer_weak_prefix = "Sledgehammer:"
 
-type prepared_formula =
+type translated_formula =
   {name: string,
    kind: kind,
    combformula: (name, combterm) formula,
@@ -53,7 +53,7 @@
 
 fun combformula_for_prop thy =
   let
-    val do_term = combterm_from_term thy ~1
+    val do_term = combterm_from_term thy
     fun do_quant bs q s T t' =
       let val s = Name.variant (map fst bs) s in
         do_formula ((s, T) :: bs) t'
@@ -174,7 +174,7 @@
       | aux _ = raise Fail "aux"
   in perhaps (try aux) end
 
-(* making axiom and conjecture formulas *)
+(* making fact and conjecture formulas *)
 fun make_formula ctxt presimp name kind t =
   let
     val thy = ProofContext.theory_of ctxt
@@ -194,7 +194,7 @@
      ctypes_sorts = ctypes_sorts}
   end
 
-fun make_axiom ctxt presimp ((name, loc), th) =
+fun make_fact ctxt presimp ((name, loc), th) =
   case make_formula ctxt presimp name Axiom (prop_of th) of
     {combformula = AAtom (CombConst (("c_True", _), _, _)), ...} => NONE
   | formula => SOME ((name, loc), formula)
@@ -214,7 +214,7 @@
 fun count_combformula (AQuant (_, _, phi)) = count_combformula phi
   | count_combformula (AConn (_, phis)) = fold count_combformula phis
   | count_combformula (AAtom tm) = count_combterm tm
-fun count_prepared_formula ({combformula, ...} : prepared_formula) =
+fun count_translated_formula ({combformula, ...} : translated_formula) =
   count_combformula combformula
 
 val optional_helpers =
@@ -232,10 +232,10 @@
   [optional_helpers, optional_typed_helpers] |> maps (maps fst)
   |> sort_distinct string_ord |> map (rpair 0) |> Symtab.make
 
-fun get_helper_facts ctxt is_FO full_types conjectures axioms =
+fun get_helper_facts ctxt is_FO full_types conjectures facts =
   let
     val ct =
-      fold (fold count_prepared_formula) [conjectures, axioms] init_counters
+      fold (fold count_translated_formula) [conjectures, facts] init_counters
     fun is_needed c = the (Symtab.lookup ct c) > 0
     fun baptize th = ((Thm.get_name_hint th, false), th)
   in
@@ -244,32 +244,32 @@
      |> maps (fn (ss, ths) =>
                  if exists is_needed ss then map baptize ths else [])) @
     (if is_FO then [] else map baptize mandatory_helpers)
-    |> map_filter (Option.map snd o make_axiom ctxt false)
+    |> map_filter (Option.map snd o make_fact ctxt false)
   end
 
-fun prepare_axiom ctxt (ax as (_, th)) = (prop_of th, make_axiom ctxt true ax)
+fun translate_fact ctxt (ax as (_, th)) = (prop_of th, make_fact ctxt true ax)
 
-fun prepare_formulas ctxt full_types hyp_ts concl_t axioms =
+fun translate_formulas ctxt full_types hyp_ts concl_t facts =
   let
     val thy = ProofContext.theory_of ctxt
-    val (axiom_ts, prepared_axioms) = ListPair.unzip axioms
-    (* Remove existing axioms from the conjecture, as this can dramatically
+    val (fact_ts, translated_facts) = ListPair.unzip facts
+    (* Remove existing facts from the conjecture, as this can dramatically
        boost an ATP's performance (for some reason). *)
-    val hyp_ts = hyp_ts |> filter_out (member (op aconv) axiom_ts)
+    val hyp_ts = hyp_ts |> filter_out (member (op aconv) fact_ts)
     val goal_t = Logic.list_implies (hyp_ts, concl_t)
     val is_FO = Meson.is_fol_term thy goal_t
     val subs = tfree_classes_of_terms [goal_t]
-    val supers = tvar_classes_of_terms axiom_ts
-    val tycons = type_consts_of_terms thy (goal_t :: axiom_ts)
-    (* TFrees in the conjecture; TVars in the axioms *)
+    val supers = tvar_classes_of_terms fact_ts
+    val tycons = type_consts_of_terms thy (goal_t :: fact_ts)
+    (* TFrees in the conjecture; TVars in the facts *)
     val conjectures = make_conjecture ctxt (hyp_ts @ [concl_t])
-    val (axiom_names, axioms) = ListPair.unzip (map_filter I prepared_axioms)
-    val helper_facts = get_helper_facts ctxt is_FO full_types conjectures axioms
+    val (fact_names, facts) = ListPair.unzip (map_filter I translated_facts)
+    val helper_facts = get_helper_facts ctxt is_FO full_types conjectures facts
     val (supers', arity_clauses) = make_arity_clauses thy tycons supers
     val class_rel_clauses = make_class_rel_clauses thy subs supers'
   in
-    (axiom_names |> map single |> Vector.fromList,
-     (conjectures, axioms, helper_facts, class_rel_clauses, arity_clauses))
+    (fact_names |> map single |> Vector.fromList,
+     (conjectures, facts, helper_facts, class_rel_clauses, arity_clauses))
   end
 
 fun wrap_type ty t = ATerm ((type_wrapper_name, type_wrapper_name), [ty, t])
@@ -322,14 +322,14 @@
       | aux (AAtom tm) = AAtom (fo_term_for_combterm full_types tm)
   in aux end
 
-fun formula_for_axiom full_types
-                      ({combformula, ctypes_sorts, ...} : prepared_formula) =
+fun formula_for_fact full_types
+                     ({combformula, ctypes_sorts, ...} : translated_formula) =
   mk_ahorn (map (formula_for_fo_literal o fo_literal_for_type_literal)
                 (type_literals_for_types ctypes_sorts))
            (formula_for_combformula full_types combformula)
 
 fun problem_line_for_fact prefix full_types (formula as {name, kind, ...}) =
-  Fof (prefix ^ ascii_of name, kind, formula_for_axiom full_types formula)
+  Fof (prefix ^ ascii_of name, kind, formula_for_fact full_types formula)
 
 fun problem_line_for_class_rel_clause (ClassRelClause {name, subclass,
                                                        superclass, ...}) =
@@ -353,11 +353,12 @@
                      (fo_literal_for_arity_literal conclLit)))
 
 fun problem_line_for_conjecture full_types
-                           ({name, kind, combformula, ...} : prepared_formula) =
+        ({name, kind, combformula, ...} : translated_formula) =
   Fof (conjecture_prefix ^ name, kind,
        formula_for_combformula full_types combformula)
 
-fun free_type_literals_for_conjecture ({ctypes_sorts, ...} : prepared_formula) =
+fun free_type_literals_for_conjecture
+        ({ctypes_sorts, ...} : translated_formula) =
   map fo_literal_for_type_literal (type_literals_for_types ctypes_sorts)
 
 fun problem_line_for_free_type j lit =
@@ -496,13 +497,13 @@
       (const_table_for_problem explicit_apply problem) problem
 
 fun prepare_atp_problem ctxt readable_names explicit_forall full_types
-                        explicit_apply hyp_ts concl_t axioms =
+                        explicit_apply hyp_ts concl_t facts =
   let
     val thy = ProofContext.theory_of ctxt
-    val (axiom_names, (conjectures, axioms, helper_facts, class_rel_clauses,
-                       arity_clauses)) =
-      prepare_formulas ctxt full_types hyp_ts concl_t axioms
-    val axiom_lines = map (problem_line_for_fact axiom_prefix full_types) axioms
+    val (fact_names, (conjectures, facts, helper_facts, class_rel_clauses,
+                      arity_clauses)) =
+      translate_formulas ctxt full_types hyp_ts concl_t facts
+    val fact_lines = map (problem_line_for_fact fact_prefix full_types) facts
     val helper_lines =
       map (problem_line_for_fact helper_prefix full_types) helper_facts
     val conjecture_lines =
@@ -514,7 +515,7 @@
     (* Reordering these might or might not confuse the proof reconstruction
        code or the SPASS Flotter hack. *)
     val problem =
-      [("Relevant facts", axiom_lines),
+      [("Relevant facts", fact_lines),
        ("Class relationships", class_rel_lines),
        ("Arity declarations", arity_lines),
        ("Helper facts", helper_lines),
@@ -523,12 +524,12 @@
       |> repair_problem thy explicit_forall full_types explicit_apply
     val (problem, pool) = nice_atp_problem readable_names problem
     val conjecture_offset =
-      length axiom_lines + length class_rel_lines + length arity_lines
+      length fact_lines + length class_rel_lines + length arity_lines
       + length helper_lines
   in
     (problem,
      case pool of SOME the_pool => snd the_pool | NONE => Symtab.empty,
-     conjecture_offset, axiom_names)
+     conjecture_offset, fact_names)
   end
 
 end;
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_filter.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_filter.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -34,7 +34,7 @@
      only : bool}
 
   val trace : bool Unsynchronized.ref
-  val name_thm_pairs_from_ref :
+  val fact_from_ref :
     Proof.context -> unit Symtab.table -> thm list
     -> Facts.ref * Attrib.src list -> ((string * locality) * thm) list
   val relevant_facts :
@@ -90,7 +90,7 @@
   (name |> Symtab.defined reserved name ? quote) ^
   (if multi then "(" ^ Int.toString j ^ ")" else "")
 
-fun name_thm_pairs_from_ref ctxt reserved chained_ths (xthm as (xref, args)) =
+fun fact_from_ref ctxt reserved chained_ths (xthm as (xref, args)) =
   let
     val ths = Attrib.eval_thms ctxt [xthm]
     val bracket =
@@ -283,7 +283,7 @@
 
 structure PType_Tab = Table(type key = ptype val ord = ptype_ord)
 
-fun count_axiom_consts thy fudge =
+fun count_fact_consts thy fudge =
   let
     fun do_const const (s, T) ts =
       (* Two-dimensional table update. Constant maps to types maps to count. *)
@@ -357,9 +357,9 @@
   | locality_bonus {assum_bonus, ...} Assum = assum_bonus
   | locality_bonus {chained_bonus, ...} Chained = chained_bonus
 
-fun axiom_weight fudge loc const_tab relevant_consts axiom_consts =
-  case axiom_consts |> List.partition (pconst_hyper_mem I relevant_consts)
-                    ||> filter_out (pconst_hyper_mem swap relevant_consts) of
+fun fact_weight fudge loc const_tab relevant_consts fact_consts =
+  case fact_consts |> List.partition (pconst_hyper_mem I relevant_consts)
+                   ||> filter_out (pconst_hyper_mem swap relevant_consts) of
     ([], _) => 0.0
   | (rel, irrel) =>
     let
@@ -372,33 +372,14 @@
       val res = rel_weight / (rel_weight + irrel_weight)
     in if Real.isFinite res then res else 0.0 end
 
-(* FIXME: experiment
-fun debug_axiom_weight fudge loc const_tab relevant_consts axiom_consts =
-  case axiom_consts |> List.partition (pconst_hyper_mem I relevant_consts)
-                    ||> filter_out (pconst_hyper_mem swap relevant_consts) of
-    ([], _) => 0.0
-  | (rel, irrel) =>
-    let
-      val irrel = irrel |> filter_out (pconst_mem swap rel)
-      val rels_weight =
-        0.0 |> fold (curry (op +) o rel_pconst_weight const_tab) rel
-      val irrels_weight =
-        ~ (locality_bonus fudge loc)
-        |> fold (curry (op +) o irrel_pconst_weight fudge const_tab) irrel
-val _ = tracing (PolyML.makestring ("REL: ", map (`(rel_pconst_weight const_tab)) rel))
-val _ = tracing (PolyML.makestring ("IRREL: ", map (`(irrel_pconst_weight fudge const_tab)) irrel))
-      val res = rels_weight / (rels_weight + irrels_weight)
-    in if Real.isFinite res then res else 0.0 end
-*)
-
-fun pconsts_in_axiom thy irrelevant_consts t =
+fun pconsts_in_fact thy irrelevant_consts t =
   Symtab.fold (fn (s, pss) => fold (cons o pair s) pss)
               (pconsts_in_terms thy irrelevant_consts true (SOME true) [t]) []
-fun pair_consts_axiom thy irrelevant_consts fudge axiom =
-  case axiom |> snd |> theory_const_prop_of fudge
-             |> pconsts_in_axiom thy irrelevant_consts of
+fun pair_consts_fact thy irrelevant_consts fudge fact =
+  case fact |> snd |> theory_const_prop_of fudge
+            |> pconsts_in_fact thy irrelevant_consts of
     [] => NONE
-  | consts => SOME ((axiom, consts), NONE)
+  | consts => SOME ((fact, consts), NONE)
 
 type annotated_thm =
   (((unit -> string) * locality) * thm) * (string * ptype) list
@@ -426,44 +407,37 @@
     (accepts, more_rejects @ rejects)
   end
 
-fun if_empty_replace_with_locality thy irrelevant_consts axioms loc tab =
+fun if_empty_replace_with_locality thy irrelevant_consts facts loc tab =
   if Symtab.is_empty tab then
     pconsts_in_terms thy irrelevant_consts false (SOME false)
         (map_filter (fn ((_, loc'), th) =>
-                        if loc' = loc then SOME (prop_of th) else NONE) axioms)
+                        if loc' = loc then SOME (prop_of th) else NONE) facts)
   else
     tab
 
 fun relevance_filter ctxt threshold0 decay max_relevant irrelevant_consts
         (fudge as {threshold_divisor, ridiculous_threshold, ...})
-        ({add, del, ...} : relevance_override) axioms goal_ts =
+        ({add, del, ...} : relevance_override) facts goal_ts =
   let
     val thy = ProofContext.theory_of ctxt
-    val const_tab = fold (count_axiom_consts thy fudge) axioms Symtab.empty
+    val const_tab = fold (count_fact_consts thy fudge) facts Symtab.empty
     val goal_const_tab =
       pconsts_in_terms thy irrelevant_consts false (SOME false) goal_ts
-      |> fold (if_empty_replace_with_locality thy irrelevant_consts axioms)
+      |> fold (if_empty_replace_with_locality thy irrelevant_consts facts)
               [Chained, Assum, Local]
     val add_ths = Attrib.eval_thms ctxt add
     val del_ths = Attrib.eval_thms ctxt del
+    val facts = facts |> filter_out (member Thm.eq_thm del_ths o snd)
     fun iter j remaining_max threshold rel_const_tab hopeless hopeful =
       let
-        fun game_over rejects =
-          (* Add "add:" facts. *)
-          if null add_ths then
-            []
-          else
-            map_filter (fn ((p as (_, th), _), _) =>
-                           if member Thm.eq_thm add_ths th then SOME p
-                           else NONE) rejects
-        fun relevant [] rejects [] =
+        fun relevant [] _ [] =
             (* Nothing has been added this iteration. *)
             if j = 0 andalso threshold >= ridiculous_threshold then
               (* First iteration? Try again. *)
               iter 0 max_relevant (threshold / threshold_divisor) rel_const_tab
                    hopeless hopeful
             else
-              game_over (rejects @ hopeless)
+              []
           | relevant candidates rejects [] =
             let
               val (accepts, more_rejects) =
@@ -496,27 +470,20 @@
                           |> map string_for_hyper_pconst));
               map (fst o fst) accepts @
               (if remaining_max = 0 then
-                 game_over (hopeful_rejects @ map (apsnd SOME) hopeless_rejects)
+                 []
                else
                  iter (j + 1) remaining_max threshold rel_const_tab'
                       hopeless_rejects hopeful_rejects)
             end
           | relevant candidates rejects
-                     (((ax as (((_, loc), _), axiom_consts)), cached_weight)
+                     (((ax as (((_, loc), _), fact_consts)), cached_weight)
                       :: hopeful) =
             let
               val weight =
                 case cached_weight of
                   SOME w => w
-                | NONE => axiom_weight fudge loc const_tab rel_const_tab
-                                       axiom_consts
-(* FIXME: experiment
-val name = fst (fst (fst ax)) ()
-val _ = if String.isSubstring "positive_minus" name orelse String.isSubstring "not_exp_le_zero" name then
-tracing ("*** " ^ name ^ PolyML.makestring (debug_axiom_weight fudge loc const_tab rel_const_tab axiom_consts))
-else
-()
-*)
+                | NONE => fact_weight fudge loc const_tab rel_const_tab
+                                      fact_consts
             in
               if weight >= threshold then
                 relevant ((ax, weight) :: candidates) rejects hopeful
@@ -532,12 +499,17 @@
                       |> map string_for_hyper_pconst));
           relevant [] [] hopeful
         end
+    fun add_add_ths accepts =
+      (facts |> filter ((member Thm.eq_thm add_ths
+                         andf (not o member (Thm.eq_thm o apsnd snd) accepts))
+                        o snd))
+      @ accepts
   in
-    axioms |> filter_out (member Thm.eq_thm del_ths o snd)
-           |> map_filter (pair_consts_axiom thy irrelevant_consts fudge)
-           |> iter 0 max_relevant threshold0 goal_const_tab []
-           |> tap (fn res => trace_msg (fn () =>
-                                "Total relevant: " ^ Int.toString (length res)))
+    facts |> map_filter (pair_consts_fact thy irrelevant_consts fudge)
+          |> iter 0 max_relevant threshold0 goal_const_tab []
+          |> not (null add_ths) ? add_add_ths
+          |> tap (fn res => trace_msg (fn () =>
+                               "Total relevant: " ^ Int.toString (length res)))
   end
 
 
@@ -550,12 +522,10 @@
 (*Reject theorems with names like "List.filter.filter_list_def" or
   "Accessible_Part.acc.defs", as these are definitions arising from packages.*)
 fun is_package_def a =
-  let val names = Long_Name.explode a
-  in
-     length names > 2 andalso
-     not (hd names = "local") andalso
-     String.isSuffix "_def" a  orelse  String.isSuffix "_defs" a
-  end;
+  let val names = Long_Name.explode a in
+    (length names > 2 andalso not (hd names = "local") andalso
+     String.isSuffix "_def" a) orelse String.isSuffix "_defs" a
+  end
 
 fun mk_fact_table f xs =
   fold (Termtab.update o `(prop_of o f)) xs Termtab.empty
@@ -718,9 +688,9 @@
           (Term.add_vars t [] |> sort_wrt (fst o fst))
   |> fst
 
-fun all_name_thms_pairs ctxt reserved full_types
-        ({intro_bonus, elim_bonus, simp_bonus, ...} : relevance_fudge) add_ths
-        chained_ths =
+fun all_facts ctxt reserved full_types
+              ({intro_bonus, elim_bonus, simp_bonus, ...} : relevance_fudge)
+              add_ths chained_ths =
   let
     val thy = ProofContext.theory_of ctxt
     val global_facts = Global_Theory.facts_of thy
@@ -805,7 +775,7 @@
 
 (* The single-name theorems go after the multiple-name ones, so that single
    names are preferred when both are available. *)
-fun name_thm_pairs ctxt respect_no_atp =
+fun rearrange_facts ctxt respect_no_atp =
   List.partition (fst o snd) #> op @ #> map (apsnd snd)
   #> respect_no_atp ? filter_out (No_ATPs.member ctxt o snd)
 
@@ -821,25 +791,25 @@
                           1.0 / Real.fromInt (max_relevant + 1))
     val add_ths = Attrib.eval_thms ctxt add
     val reserved = reserved_isar_keyword_table ()
-    val axioms =
+    val facts =
       (if only then
          maps (map (fn ((name, loc), th) => ((K name, loc), (true, th)))
-               o name_thm_pairs_from_ref ctxt reserved chained_ths) add
+               o fact_from_ref ctxt reserved chained_ths) add
        else
-         all_name_thms_pairs ctxt reserved full_types fudge add_ths chained_ths)
-      |> name_thm_pairs ctxt (respect_no_atp andalso not only)
+         all_facts ctxt reserved full_types fudge add_ths chained_ths)
+      |> rearrange_facts ctxt (respect_no_atp andalso not only)
       |> uniquify
   in
-    trace_msg (fn () => "Considering " ^ Int.toString (length axioms) ^
-                        " theorems");
+    trace_msg (fn () => "Considering " ^ Int.toString (length facts) ^
+                        " facts");
     (if only orelse threshold1 < 0.0 then
-       axioms
+       facts
      else if threshold0 > 1.0 orelse threshold0 > threshold1 orelse
              max_relevant = 0 then
        []
      else
        relevance_filter ctxt threshold0 decay max_relevant irrelevant_consts
-                        fudge override axioms (concl_t :: hyp_ts))
+                        fudge override facts (concl_t :: hyp_ts))
     |> map (apfst (apfst (fn f => f ())))
   end
 
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_isar.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -130,9 +130,6 @@
   val extend = I
   fun merge p : T = AList.merge (op =) (K true) p)
 
-(* FIXME: dummy *)
-fun is_smt_solver_installed ctxt = true
-
 fun remotify_prover_if_available_and_not_already_remote thy name =
   if String.isPrefix remote_prefix name then
     SOME name
@@ -157,13 +154,15 @@
       prover :: avoid_too_many_local_threads thy n provers
     end
 
+fun num_processors () = try Thread.numProcessors () |> the_default 1
+
 (* The first ATP of the list is used by Auto Sledgehammer. Because of the low
    timeout, it makes sense to put SPASS first. *)
 fun default_provers_param_value ctxt =
   let val thy = ProofContext.theory_of ctxt in
     [spassN, eN, vampireN, sine_eN (* FIXME: , smtN *)]
     |> map_filter (remotify_prover_if_not_installed ctxt)
-    |> avoid_too_many_local_threads thy (Thread.numProcessors ())
+    |> avoid_too_many_local_threads thy (num_processors ())
     |> space_implode " "
   end
 
@@ -360,6 +359,6 @@
                        (minimize_command [] 1) state
     end
 
-val setup = Auto_Tools.register_tool ("sledgehammer", auto_sledgehammer)
+val setup = Auto_Tools.register_tool (sledgehammerN, auto_sledgehammer)
 
 end;
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -44,42 +44,42 @@
 
 fun test_facts ({debug, verbose, overlord, provers, full_types, isar_proof,
                  isar_shrink_factor, ...} : params) (prover : prover)
-               explicit_apply timeout i n state axioms =
+               explicit_apply timeout i n state facts =
   let
     val _ =
-      priority ("Testing " ^ n_facts (map fst axioms) ^ "...")
+      Output.urgent_message ("Testing " ^ n_facts (map fst facts) ^ "...")
     val params =
       {blocking = true, debug = debug, verbose = verbose, overlord = overlord,
        provers = provers, full_types = full_types,
        explicit_apply = explicit_apply, relevance_thresholds = (1.01, 1.01),
        max_relevant = NONE, isar_proof = isar_proof,
        isar_shrink_factor = isar_shrink_factor, timeout = timeout, expect = ""}
-    val axioms =
-      axioms |> maps (fn (n, ths) => ths |> map (Unprepared o pair n))
+    val facts =
+      facts |> maps (fn (n, ths) => ths |> map (Untranslated o pair n))
     val {goal, ...} = Proof.goal state
     val problem =
       {state = state, goal = goal, subgoal = i, subgoal_count = n,
-       axioms = axioms, only = true}
-    val result as {outcome, used_axioms, ...} = prover params (K "") problem
+       facts = facts, only = true}
+    val result as {outcome, used_facts, ...} = prover params (K "") problem
   in
-    priority (case outcome of
+    Output.urgent_message (case outcome of
                 SOME failure => string_for_failure failure
               | NONE =>
-                if length used_axioms = length axioms then "Found proof."
-                else "Found proof with " ^ n_facts used_axioms ^ ".");
+                if length used_facts = length facts then "Found proof."
+                else "Found proof with " ^ n_facts used_facts ^ ".");
     result
   end
 
-(* minimalization of thms *)
+(* minimalization of facts *)
 
-fun filter_used_axioms used = filter (member (op =) used o fst)
+fun filter_used_facts used = filter (member (op =) used o fst)
 
 fun sublinear_minimize _ [] p = p
   | sublinear_minimize test (x :: xs) (seen, result) =
     case test (xs @ seen) of
-      result as {outcome = NONE, used_axioms, ...} : prover_result =>
-      sublinear_minimize test (filter_used_axioms used_axioms xs)
-                         (filter_used_axioms used_axioms seen, result)
+      result as {outcome = NONE, used_facts, ...} : prover_result =>
+      sublinear_minimize test (filter_used_facts used_facts xs)
+                         (filter_used_facts used_facts seen, result)
     | _ => sublinear_minimize test xs (x :: seen, result)
 
 (* Give the ATP some slack. The ATP gets further slack because the Sledgehammer
@@ -89,23 +89,23 @@
 
 fun minimize_facts {provers = [], ...} _ _ _ _ = error "No prover is set."
   | minimize_facts (params as {provers = prover_name :: _, timeout, ...}) i n
-                   state axioms =
+                   state facts =
   let
     val thy = Proof.theory_of state
     val prover = get_prover thy false prover_name
     val msecs = Time.toMilliseconds timeout
-    val _ = priority ("Sledgehammer minimize: " ^ quote prover_name ^ ".")
+    val _ = Output.urgent_message ("Sledgehammer minimize: " ^ quote prover_name ^ ".")
     val {goal, ...} = Proof.goal state
     val (_, hyp_ts, concl_t) = strip_subgoal goal i
     val explicit_apply =
       not (forall (Meson.is_fol_term thy)
-                  (concl_t :: hyp_ts @ maps (map prop_of o snd) axioms))
+                  (concl_t :: hyp_ts @ maps (map prop_of o snd) facts))
     fun do_test timeout =
       test_facts params prover explicit_apply timeout i n state
     val timer = Timer.startRealTimer ()
   in
-    (case do_test timeout axioms of
-       result as {outcome = NONE, used_axioms, ...} =>
+    (case do_test timeout facts of
+       result as {outcome = NONE, used_facts, ...} =>
        let
          val time = Timer.checkRealTimer timer
          val new_timeout =
@@ -113,9 +113,9 @@
            |> Time.fromMilliseconds
          val (min_thms, {message, ...}) =
            sublinear_minimize (do_test new_timeout)
-               (filter_used_axioms used_axioms axioms) ([], result)
+               (filter_used_facts used_facts facts) ([], result)
          val n = length min_thms
-         val _ = priority (cat_lines
+         val _ = Output.urgent_message (cat_lines
            ["Minimized: " ^ string_of_int n ^ " fact" ^ plural_s n] ^
             (case length (filter (curry (op =) Chained o snd o fst) min_thms) of
                0 => ""
@@ -140,15 +140,14 @@
     val ctxt = Proof.context_of state
     val reserved = reserved_isar_keyword_table ()
     val chained_ths = #facts (Proof.goal state)
-    val axioms =
-      maps (map (apsnd single)
-            o name_thm_pairs_from_ref ctxt reserved chained_ths) refs
+    val facts =
+      maps (map (apsnd single) o fact_from_ref ctxt reserved chained_ths) refs
   in
     case subgoal_count state of
-      0 => priority "No subgoal!"
+      0 => Output.urgent_message "No subgoal!"
     | n =>
       (kill_provers ();
-       priority (#2 (minimize_facts params i n state axioms)))
+       Output.urgent_message (#2 (minimize_facts params i n state facts)))
   end
 
 end;
--- a/src/HOL/Tools/async_manager.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/async_manager.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -105,7 +105,7 @@
     [] => ()
   | msgs as (tool, _) :: _ =>
     let val ss = break_into_chunks (map snd msgs) in
-      List.app priority (tool ^ ": " ^ hd ss :: tl ss)
+      List.app Output.urgent_message (tool ^ ": " ^ hd ss :: tl ss)
     end
 
 fun check_thread_manager () = Synchronized.change global_state
@@ -185,7 +185,7 @@
                        if tool' = tool then SOME (th, (tool', Time.now (), desc))
                        else NONE) active
       val state' = make_state manager timeout_heap [] (killing @ canceling) messages store;
-      val _ = if null killing then () else priority ("Killed active " ^ workers ^ ".")
+      val _ = if null killing then () else Output.urgent_message ("Killed active " ^ workers ^ ".")
     in state' end);
 
 
@@ -218,7 +218,7 @@
       case map_filter canceling_info canceling of
         [] => []
       | ss => "Trying to interrupt the following " ^ workers ^ ":" :: ss
-  in priority (space_implode "\n\n" (running @ interrupting)) end
+  in Output.urgent_message (space_implode "\n\n" (running @ interrupting)) end
 
 fun thread_messages tool worker opt_limit =
   let
@@ -230,7 +230,7 @@
         (if length tool_store <= limit then ":"
          else " (" ^ string_of_int limit ^ " displayed):");
   in
-    List.app priority (header :: break_into_chunks
+    List.app Output.urgent_message (header :: break_into_chunks
                                      (map snd (#1 (chop limit tool_store))))
   end
 
--- a/src/HOL/Tools/inductive_codegen.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/inductive_codegen.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -844,7 +844,7 @@
     val test_fn' = !test_fn;
     val dummy_report = ([], false)
     fun test k = (deepen bound (fn i =>
-      (priority ("Search depth: " ^ string_of_int i);
+      (Output.urgent_message ("Search depth: " ^ string_of_int i);
        test_fn' (i, values, k+offset))) start, dummy_report);
   in test end;
 
--- a/src/HOL/Tools/refute.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/refute.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -1133,31 +1133,31 @@
                             ". Available solvers: " ^
                             commas (map (quote o fst) (!SatSolver.solvers)) ^ ".")
           in
-            priority "Invoking SAT solver...";
+            Output.urgent_message "Invoking SAT solver...";
             (case solver fm of
               SatSolver.SATISFIABLE assignment =>
-                (priority ("*** Model found: ***\n" ^ print_model ctxt model
+                (Output.urgent_message ("*** Model found: ***\n" ^ print_model ctxt model
                   (fn i => case assignment i of SOME b => b | NONE => true));
                  if maybe_spurious then "potential" else "genuine")
             | SatSolver.UNSATISFIABLE _ =>
-                (priority "No model exists.";
+                (Output.urgent_message "No model exists.";
                 case next_universe universe sizes minsize maxsize of
                   SOME universe' => find_model_loop universe'
-                | NONE => (priority
-                  "Search terminated, no larger universe within the given limits.";
-                  "none"))
+                | NONE => (Output.urgent_message
+                    "Search terminated, no larger universe within the given limits.";
+                    "none"))
             | SatSolver.UNKNOWN =>
-                (priority "No model found.";
+                (Output.urgent_message "No model found.";
                 case next_universe universe sizes minsize maxsize of
                   SOME universe' => find_model_loop universe'
-                | NONE           => (priority
+                | NONE => (Output.urgent_message
                   "Search terminated, no larger universe within the given limits.";
                   "unknown"))) handle SatSolver.NOT_CONFIGURED =>
               (error ("SAT solver " ^ quote satsolver ^ " is not configured.");
                "unknown")
           end
           handle MAXVARS_EXCEEDED =>
-            (priority ("Search terminated, number of Boolean variables ("
+            (Output.urgent_message ("Search terminated, number of Boolean variables ("
               ^ string_of_int maxvars ^ " allowed) exceeded.");
               "unknown")
 
@@ -1179,14 +1179,14 @@
     maxtime>=0 orelse
       error ("\"maxtime\" is " ^ string_of_int maxtime ^ ", must be at least 0");
     (* enter loop with or without time limit *)
-    priority ("Trying to find a model that "
+    Output.urgent_message ("Trying to find a model that "
       ^ (if negate then "refutes" else "satisfies") ^ ": "
       ^ Syntax.string_of_term ctxt t);
     if maxtime > 0 then (
       TimeLimit.timeLimit (Time.fromSeconds maxtime)
         wrapper ()
       handle TimeLimit.TimeOut =>
-        (priority ("Search terminated, time limit (" ^
+        (Output.urgent_message ("Search terminated, time limit (" ^
             string_of_int maxtime
             ^ (if maxtime=1 then " second" else " seconds") ^ ") exceeded.");
          check_expect "unknown")
@@ -1273,7 +1273,7 @@
     val t = th |> prop_of
   in
     if Logic.count_prems t = 0 then
-      priority "No subgoal!"
+      Output.urgent_message "No subgoal!"
     else
       let
         val assms = map term_of (Assumption.all_assms_of ctxt)
--- a/src/HOL/Tools/try.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/Tools/try.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -40,6 +40,7 @@
       SOME (x, _) => nprems_of (post x) < nprems_of goal
     | NONE => false
   end
+  handle TimeLimit.TimeOut => false
 
 fun do_generic timeout_opt command pre post apply st =
   let val timer = Timer.startRealTimer () in
@@ -52,15 +53,6 @@
 fun named_method thy name =
   Method.method thy (Args.src ((name, []), Position.none))
 
-fun apply_named_method name ctxt =
-  let val thy = ProofContext.theory_of ctxt in
-    Method.apply (named_method thy name) ctxt []
-  end
-
-fun do_named_method name timeout_opt st =
-  do_generic timeout_opt name (#goal o Proof.goal) snd
-             (apply_named_method name (Proof.context_of st)) st
-
 fun apply_named_method_on_first_goal name ctxt =
   let val thy = ProofContext.theory_of ctxt in
     Proof.refine (Method.SelectGoals (1, Method.Basic (named_method thy name)))
@@ -77,7 +69,8 @@
 
 val named_methods =
   [("simp", false), ("auto", true), ("fast", false), ("fastsimp", false),
-   ("force", false), ("blast", false), ("arith", false)]
+   ("force", false), ("blast", false), ("metis", false), ("linarith", false),
+   ("presburger", false)]
 val do_methods = map do_named_method named_methods
 
 fun time_string (s, ms) = s ^ ": " ^ string_of_int ms ^ " ms"
@@ -103,7 +96,7 @@
                                     Pretty.markup Markup.hilite
                                                   [Pretty.str message]])
                     else
-                      tap (fn _ => priority message)))
+                      tap (fn _ => Output.urgent_message message)))
     end
 
 val invoke_try = fst oo do_try false
--- a/src/HOL/ex/Fundefs.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/HOL/ex/Fundefs.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -5,7 +5,7 @@
 header {* Examples of function definitions *}
 
 theory Fundefs 
-imports Main
+imports Parity Monad_Syntax
 begin
 
 subsection {* Very basic *}
@@ -208,6 +208,31 @@
 thm my_monoid.foldL.simps
 thm my_monoid.foldR_foldL
 
+
+subsection {* Partial Function Definitions *}
+
+text {* Partial functions in the option monad: *}
+
+partial_function (option)
+  collatz :: "nat \<Rightarrow> nat list option"
+where
+  "collatz n =
+  (if n \<le> 1 then Some [n]
+   else if even n 
+     then do { ns \<leftarrow> collatz (n div 2); Some (n # ns) }
+     else do { ns \<leftarrow> collatz (3 * n + 1);  Some (n # ns)})"
+
+declare collatz.simps[code]
+value "collatz 23"
+
+
+text {* Tail-recursive functions: *}
+
+partial_function (tailrec) fixpoint :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a"
+where
+  "fixpoint f x = (if f x = x then x else fixpoint f (f x))"
+
+
 subsection {* Regression tests *}
 
 text {* The following examples mainly serve as tests for the 
--- a/src/Pure/General/markup.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/Pure/General/markup.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -120,11 +120,11 @@
   val reportN: string val report: T
   val no_reportN: string val no_report: T
   val badN: string val bad: T
-  val no_output: output * output
-  val default_output: T -> output * output
-  val add_mode: string -> (T -> output * output) -> unit
-  val output: T -> output * output
-  val enclose: T -> output -> output
+  val no_output: Output.output * Output.output
+  val default_output: T -> Output.output * Output.output
+  val add_mode: string -> (T -> Output.output * Output.output) -> unit
+  val output: T -> Output.output * Output.output
+  val enclose: T -> Output.output -> Output.output
   val markup: T -> string -> string
 end;
 
--- a/src/Pure/General/output.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/Pure/General/output.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -6,9 +6,7 @@
 
 signature BASIC_OUTPUT =
 sig
-  type output = string
   val writeln: string -> unit
-  val priority: string -> unit
   val tracing: string -> unit
   val warning: string -> unit
   val legacy_feature: string -> unit
@@ -22,6 +20,7 @@
 signature OUTPUT =
 sig
   include BASIC_OUTPUT
+  type output = string
   val default_output: string -> output * int
   val default_escape: output -> string
   val add_mode: string -> (string -> output * int) -> (output -> string) -> unit
@@ -31,21 +30,22 @@
   val raw_stdout: output -> unit
   val raw_stderr: output -> unit
   val raw_writeln: output -> unit
-  val writeln_fn: (output -> unit) Unsynchronized.ref
-  val priority_fn: (output -> unit) Unsynchronized.ref
-  val tracing_fn: (output -> unit) Unsynchronized.ref
-  val warning_fn: (output -> unit) Unsynchronized.ref
-  val error_fn: (output -> unit) Unsynchronized.ref
-  val debug_fn: (output -> unit) Unsynchronized.ref
-  val prompt_fn: (output -> unit) Unsynchronized.ref
-  val status_fn: (output -> unit) Unsynchronized.ref
-  val report_fn: (output -> unit) Unsynchronized.ref
+  structure Private_Hooks:
+  sig
+    val writeln_fn: (output -> unit) Unsynchronized.ref
+    val urgent_message_fn: (output -> unit) Unsynchronized.ref
+    val tracing_fn: (output -> unit) Unsynchronized.ref
+    val warning_fn: (output -> unit) Unsynchronized.ref
+    val error_fn: (output -> unit) Unsynchronized.ref
+    val prompt_fn: (output -> unit) Unsynchronized.ref
+    val status_fn: (output -> unit) Unsynchronized.ref
+    val report_fn: (output -> unit) Unsynchronized.ref
+  end
+  val urgent_message: string -> unit
   val error_msg: string -> unit
   val prompt: string -> unit
   val status: string -> unit
   val report: string -> unit
-  val debugging: bool Unsynchronized.ref
-  val debug: (unit -> string) -> unit
 end;
 
 structure Output: OUTPUT =
@@ -88,30 +88,29 @@
 
 (* Isabelle output channels *)
 
-val writeln_fn = Unsynchronized.ref raw_writeln;
-val priority_fn = Unsynchronized.ref (fn s => ! writeln_fn s);
-val tracing_fn = Unsynchronized.ref (fn s => ! writeln_fn s);
-val warning_fn = Unsynchronized.ref (raw_stdout o suffix "\n" o prefix_lines "### ");
-val error_fn = Unsynchronized.ref (raw_stdout o suffix "\n" o prefix_lines "*** ");
-val debug_fn = Unsynchronized.ref (raw_stdout o suffix "\n" o prefix_lines "::: ");
-val prompt_fn = Unsynchronized.ref raw_stdout;
-val status_fn = Unsynchronized.ref (fn _: string => ());
-val report_fn = Unsynchronized.ref (fn _: string => ());
+structure Private_Hooks =
+struct
+  val writeln_fn = Unsynchronized.ref raw_writeln;
+  val urgent_message_fn = Unsynchronized.ref (fn s => ! writeln_fn s);
+  val tracing_fn = Unsynchronized.ref (fn s => ! writeln_fn s);
+  val warning_fn = Unsynchronized.ref (raw_stdout o suffix "\n" o prefix_lines "### ");
+  val error_fn = Unsynchronized.ref (raw_stdout o suffix "\n" o prefix_lines "*** ");
+  val prompt_fn = Unsynchronized.ref raw_stdout;
+  val status_fn = Unsynchronized.ref (fn _: string => ());
+  val report_fn = Unsynchronized.ref (fn _: string => ());
+end;
 
-fun writeln s = ! writeln_fn (output s);
-fun priority s = ! priority_fn (output s);
-fun tracing s = ! tracing_fn (output s);
-fun warning s = ! warning_fn (output s);
-fun error_msg s = ! error_fn (output s);
-fun prompt s = ! prompt_fn (output s);
-fun status s = ! status_fn (output s);
-fun report s = ! report_fn (output s);
+fun writeln s = ! Private_Hooks.writeln_fn (output s);
+fun urgent_message s = ! Private_Hooks.urgent_message_fn (output s);
+fun tracing s = ! Private_Hooks.tracing_fn (output s);
+fun warning s = ! Private_Hooks.warning_fn (output s);
+fun error_msg s = ! Private_Hooks.error_fn (output s);
+fun prompt s = ! Private_Hooks.prompt_fn (output s);
+fun status s = ! Private_Hooks.status_fn (output s);
+fun report s = ! Private_Hooks.report_fn (output s);
 
 fun legacy_feature s = warning ("Legacy feature! " ^ s);
 
-val debugging = Unsynchronized.ref false;
-fun debug s = if ! debugging then ! debug_fn (output (s ())) else ()
-
 
 
 (** timing **)
--- a/src/Pure/General/pretty.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/Pure/General/pretty.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -21,8 +21,8 @@
 
 signature PRETTY =
 sig
-  val default_indent: string -> int -> output
-  val add_mode: string -> (string -> int -> output) -> unit
+  val default_indent: string -> int -> Output.output
+  val add_mode: string -> (string -> int -> Output.output) -> unit
   type T
   val str: string -> T
   val brk: int -> T
@@ -55,7 +55,7 @@
   val margin_default: int Unsynchronized.ref
   val symbolicN: string
   val output_buffer: int option -> T -> Buffer.T
-  val output: int option -> T -> output
+  val output: int option -> T -> Output.output
   val string_of_margin: int -> T -> string
   val string_of: T -> string
   val str_of: T -> string
@@ -103,9 +103,10 @@
 (** printing items: compound phrases, strings, and breaks **)
 
 abstype T =
-  Block of (output * output) * T list * int * int |  (*markup output, body, indentation, length*)
-  String of output * int |                           (*text, length*)
-  Break of bool * int                                (*mandatory flag, width if not taken*)
+    Block of (Output.output * Output.output) * T list * int * int
+      (*markup output, body, indentation, length*)
+  | String of Output.output * int  (*text, length*)
+  | Break of bool * int  (*mandatory flag, width if not taken*)
 with
 
 fun length (Block (_, _, _, len)) = len
--- a/src/Pure/General/symbol.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/Pure/General/symbol.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -66,7 +66,7 @@
   val bump_string: string -> string
   val length: symbol list -> int
   val xsymbolsN: string
-  val output: string -> output * int
+  val output: string -> Output.output * int
 end;
 
 structure Symbol: SYMBOL =
--- a/src/Pure/General/xml.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/Pure/General/xml.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -16,7 +16,7 @@
   val header: string
   val text: string -> string
   val element: string -> attributes -> string list -> string
-  val output_markup: Markup.T -> output * output
+  val output_markup: Markup.T -> Output.output * Output.output
   val string_of: tree -> string
   val output: tree -> TextIO.outstream -> unit
   val parse_comments: string list -> unit * string list
--- a/src/Pure/Isar/class_declaration.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/Pure/Isar/class_declaration.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -129,9 +129,15 @@
                   ^ Syntax.string_of_sort_global thy a_sort)
             | _ => error "Multiple type variables in class specification";
       in (map o map_atyps) (K (TFree (Name.aT, fixate_sort))) Ts end;
-    val after_infer_fixate = (map o map_atyps)
-      (fn T as TFree _ => T | T as TVar (vi, sort) =>
-        if Type_Infer.is_param vi then Type_Infer.param 0 (Name.aT, sort) else T);
+    fun after_infer_fixate Ts =
+      let
+        val sort' = (fold o fold_atyps) (fn T as TFree _ => I | T as TVar (vi, sort) =>
+          if Type_Infer.is_param vi then inter_sort sort else I) Ts [];
+      in
+        (map o map_atyps)
+          (fn T as TFree _ => T | T as TVar (vi, _) =>
+            if Type_Infer.is_param vi then Type_Infer.param 0 (Name.aT, sort') else T) Ts
+      end;
     fun add_typ_check level name f = Context.proof_map
       (Syntax.add_typ_check level name (fn Ts => fn ctxt =>
         let val Ts' = f Ts in if eq_list (op =) (Ts, Ts') then NONE else SOME (Ts', ctxt) end));
--- a/src/Pure/Isar/code.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/Pure/Isar/code.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -542,7 +542,7 @@
     val (rep, lhs) = dest_comb full_lhs
       handle TERM _ => bad "Not an abstract equation";
     val (rep_const, ty) = dest_Const rep;
-    val (tyco, sorts) = ((apsnd o map) (snd o dest_TVar) o dest_Type o domain_type) ty
+    val (tyco, Ts) = (dest_Type o domain_type) ty
       handle TERM _ => bad "Not an abstract equation"
            | TYPE _ => bad "Not an abstract equation";
     val _ = case some_tyco of SOME tyco' => if tyco = tyco' then ()
@@ -553,8 +553,8 @@
       else bad ("Projection mismatch: " ^ quote rep_const ^ " vs. " ^ quote rep');
     val _ = check_eqn thy { allow_nonlinear = false,
       allow_consts = false, allow_pats = false } thm (lhs, rhs);
-    val _ = if forall (Sign.subsort thy) (sorts ~~ map snd  vs') then ()
-      else error ("Sort constraints on type arguments are weaker than in abstype certificate.")
+    val _ = if forall (Sign.of_sort thy) (Ts ~~ map snd vs') then ()
+      else error ("Type arguments do not satisfy sort constraints of abstype certificate.");
   in (thm, tyco) end;
 
 fun assert_eqn thy = error_thm (gen_assert_eqn thy true);
--- a/src/Pure/Isar/proof.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/Pure/Isar/proof.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -972,7 +972,7 @@
       if ! testing then () else Proof_Display.print_results int ctxt res;
     fun print_rule ctxt th =
       if ! testing then rule := SOME th
-      else if int then priority (Proof_Display.string_of_rule ctxt "Successful" th)
+      else if int then Output.urgent_message (Proof_Display.string_of_rule ctxt "Successful" th)
       else ();
     val test_proof =
       try (local_skip_proof true)
--- a/src/Pure/Isar/toplevel.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/Pure/Isar/toplevel.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -231,7 +231,8 @@
     (((fn () => body () handle exn => if Exn.is_interrupt exn then () else reraise exn)
         |> Runtime.debugging
         |> Runtime.toplevel_error
-          (fn exn => priority ("## INTERNAL ERROR ##\n" ^ ML_Compiler.exn_message exn))),
+          (fn exn =>
+            Output.urgent_message ("## INTERNAL ERROR ##\n" ^ ML_Compiler.exn_message exn))),
       Simple_Thread.attributes interrupts);
 
 
--- a/src/Pure/ML/ml_antiquote.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/Pure/ML/ml_antiquote.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -57,6 +57,9 @@
 
 (** misc antiquotations **)
 
+val _ = inline "assert"
+  (Scan.succeed "(fn b => if b then () else raise General.Fail \"Assertion failed\")");
+
 val _ = inline "make_string" (Scan.succeed ml_make_string);
 
 val _ = value "binding"
--- a/src/Pure/Proof/extraction.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/Pure/Proof/extraction.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -119,7 +119,7 @@
 fun extr_name s vs = Long_Name.append "extr" (space_implode "_" (s :: vs));
 fun corr_name s vs = extr_name s vs ^ "_correctness";
 
-fun msg d s = priority (Symbol.spaces d ^ s);
+fun msg d s = Output.urgent_message (Symbol.spaces d ^ s);
 
 fun vars_of t = map Var (rev (Term.add_vars t []));
 fun frees_of t = map Free (rev (Term.add_frees t []));
--- a/src/Pure/ProofGeneral/proof_general_emacs.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/Pure/ProofGeneral/proof_general_emacs.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -74,14 +74,14 @@
   | s => Output.raw_writeln (enclose bg en (prefix_lines prfx s)));
 
 fun setup_messages () =
- (Output.writeln_fn := message "" "" "";
-  Output.status_fn := (fn _ => ());
-  Output.report_fn := (fn _ => ());
-  Output.priority_fn := message (special "I") (special "J") "";
-  Output.tracing_fn := message (special "I" ^ special "V") (special "J") "";
-  Output.warning_fn := message (special "K") (special "L") "### ";
-  Output.error_fn := message (special "M") (special "N") "*** ";
-  Output.prompt_fn := (fn s => Output.raw_stdout (render s ^ special "S")));
+ (Output.Private_Hooks.writeln_fn := message "" "" "";
+  Output.Private_Hooks.status_fn := (fn _ => ());
+  Output.Private_Hooks.report_fn := (fn _ => ());
+  Output.Private_Hooks.urgent_message_fn := message (special "I") (special "J") "";
+  Output.Private_Hooks.tracing_fn := message (special "I" ^ special "V") (special "J") "";
+  Output.Private_Hooks.warning_fn := message (special "K") (special "L") "### ";
+  Output.Private_Hooks.error_fn := message (special "M") (special "N") "*** ";
+  Output.Private_Hooks.prompt_fn := (fn s => Output.raw_stdout (render s ^ special "S")));
 
 fun panic s =
   (message (special "M") (special "N") "!!! " ("## SYSTEM EXIT ##\n" ^ s); exit 1);
@@ -146,7 +146,7 @@
 
 (* restart top-level loop (keeps most state information) *)
 
-val welcome = priority o Session.welcome;
+val welcome = Output.urgent_message o Session.welcome;
 
 fun restart () =
  (sync_thy_loader ();
@@ -227,7 +227,7 @@
           Output.default_output Output.default_escape;
          Markup.add_mode ProofGeneralPgip.proof_general_emacsN YXML.output_markup;
          setup_messages ();
-         ProofGeneralPgip.pgip_channel_emacs (! Output.priority_fn);
+         ProofGeneralPgip.pgip_channel_emacs (! Output.Private_Hooks.urgent_message_fn);
          setup_thy_loader ();
          setup_present_hook ();
          initialized := true);
--- a/src/Pure/ProofGeneral/proof_general_pgip.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/Pure/ProofGeneral/proof_general_pgip.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -160,13 +160,13 @@
    add new lines explicitly in PGIP: they are left implicit.  It means that PGIP messages
    can't be written without newlines. *)
 fun setup_messages () =
- (Output.writeln_fn := (fn s => normalmsg Message s);
-  Output.status_fn := (fn _ => ());
-  Output.report_fn := (fn _ => ());
-  Output.priority_fn := (fn s => normalmsg Status s);
-  Output.tracing_fn := (fn s => normalmsg Tracing s);
-  Output.warning_fn := (fn s => errormsg Message Warning s);
-  Output.error_fn := (fn s => errormsg Message Fatal s));
+ (Output.Private_Hooks.writeln_fn := (fn s => normalmsg Message s);
+  Output.Private_Hooks.status_fn := (fn _ => ());
+  Output.Private_Hooks.report_fn := (fn _ => ());
+  Output.Private_Hooks.urgent_message_fn := (fn s => normalmsg Status s);
+  Output.Private_Hooks.tracing_fn := (fn s => normalmsg Tracing s);
+  Output.Private_Hooks.warning_fn := (fn s => errormsg Message Warning s);
+  Output.Private_Hooks.error_fn := (fn s => errormsg Message Fatal s));
 
 
 (* immediate messages *)
@@ -231,7 +231,7 @@
 
 (* restart top-level loop (keeps most state information) *)
 
-val welcome = priority o Session.welcome;
+val welcome = Output.urgent_message o Session.welcome;
 
 fun restart () =
     (sync_thy_loader ();
@@ -807,14 +807,15 @@
                                 PgipTypes.string_of_pgipurl url)
         | NONE => (openfile_retract filepath;
                    changecwd_dir filedir;
-                   priority ("Working in file: " ^ PgipTypes.string_of_pgipurl url);
+                   Output.urgent_message ("Working in file: " ^ PgipTypes.string_of_pgipurl url);
                    currently_open_file := SOME url)
   end
 
 fun closefile _ =
     case !currently_open_file of
         SOME f => (inform_file_processed f (Isar.state ());
-                   priority ("Finished working in file: " ^ PgipTypes.string_of_pgipurl f);
+                   Output.urgent_message
+                    ("Finished working in file: " ^ PgipTypes.string_of_pgipurl f);
                    currently_open_file := NONE)
       | NONE => raise PGIP ("<closefile> when no file is open!")
 
@@ -835,7 +836,7 @@
 fun abortfile _ =
     case !currently_open_file of
         SOME f => (isarcmd "init_toplevel";
-                   priority ("Aborted working in file: " ^
+                   Output.urgent_message ("Aborted working in file: " ^
                              PgipTypes.string_of_pgipurl f);
                    currently_open_file := NONE)
       | NONE => raise PGIP ("<abortfile> when no file is open!")
@@ -846,7 +847,7 @@
     in
         case !currently_open_file of
             SOME f => raise PGIP ("<retractfile> when a file is open!")
-          | NONE => (priority ("Retracting file: " ^ PgipTypes.string_of_pgipurl url);
+          | NONE => (Output.urgent_message ("Retracting file: " ^ PgipTypes.string_of_pgipurl url);
                      (* TODO: next should be in thy loader, here just for testing *)
                      let
                          val name = thy_name url
--- a/src/Pure/System/isabelle_process.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/Pure/System/isabelle_process.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -57,42 +57,33 @@
 end;
 
 
-(* message markup *)
+(* message channels *)
 
 local
 
-fun chunk s = string_of_int (size s) ^ "\n" ^ s;
+fun chunk s = [string_of_int (size s), "\n", s];
 
 fun message _ _ _ "" = ()
-  | message out_stream ch raw_props body =
+  | message mbox ch raw_props body =
       let
         val robust_props = map (pairself YXML.escape_controls) raw_props;
         val header = YXML.string_of (XML.Elem ((ch, robust_props), []));
-      in TextIO.output (out_stream, chunk header ^ chunk body) (*atomic output!*) end;
+      in Mailbox.send mbox (chunk header @ chunk body) end;
 
-in
-
-fun standard_message out_stream with_serial ch body =
-  message out_stream ch
+fun standard_message mbox with_serial ch body =
+  message mbox ch
     ((if with_serial then cons (Markup.serialN, serial_string ()) else I)
       (Position.properties_of (Position.thread_data ()))) body;
 
-fun init_message out_stream =
-  message out_stream "A" [] (Session.welcome ());
-
-end;
-
-
-(* channels *)
-
-local
-
-fun auto_flush stream =
+fun message_output mbox out_stream =
   let
-    val _ = TextIO.StreamIO.setBufferMode (TextIO.getOutstream stream, IO.BLOCK_BUF);
-    fun loop () =
-      (OS.Process.sleep (Time.fromMilliseconds 20); try TextIO.flushOut stream; loop ());
-  in loop end;
+    fun loop receive =
+      (case receive mbox of
+        SOME msg =>
+         (List.app (fn s => TextIO.output (out_stream, s)) msg;
+          loop (Mailbox.receive_timeout (Time.fromMilliseconds 20)))
+      | NONE => (try TextIO.flushOut out_stream; loop (SOME o Mailbox.receive)));
+  in fn () => loop (SOME o Mailbox.receive) end;
 
 in
 
@@ -102,19 +93,22 @@
     val in_stream = TextIO.openIn in_fifo;
     val out_stream = TextIO.openOut out_fifo;
 
-    val _ = Simple_Thread.fork false (auto_flush out_stream);
-    val _ = Simple_Thread.fork false (auto_flush TextIO.stdOut);
-    val _ = Simple_Thread.fork false (auto_flush TextIO.stdErr);
+    val _ = TextIO.StreamIO.setBufferMode (TextIO.getOutstream TextIO.stdOut, IO.LINE_BUF);
+    val _ = TextIO.StreamIO.setBufferMode (TextIO.getOutstream TextIO.stdErr, IO.LINE_BUF);
+
+    val mbox = Mailbox.create () : string list Mailbox.T;
+    val _ = Simple_Thread.fork false (message_output mbox out_stream);
   in
-    Output.status_fn   := standard_message out_stream false "B";
-    Output.report_fn   := standard_message out_stream false "C";
-    Output.writeln_fn  := standard_message out_stream true "D";
-    Output.tracing_fn  := standard_message out_stream true "E";
-    Output.warning_fn  := standard_message out_stream true "F";
-    Output.error_fn    := standard_message out_stream true "G";
-    Output.priority_fn := ! Output.writeln_fn;
-    Output.prompt_fn   := ignore;
-    (in_stream, out_stream)
+    Output.Private_Hooks.status_fn := standard_message mbox false "B";
+    Output.Private_Hooks.report_fn := standard_message mbox false "C";
+    Output.Private_Hooks.writeln_fn := standard_message mbox true "D";
+    Output.Private_Hooks.tracing_fn := standard_message mbox true "E";
+    Output.Private_Hooks.warning_fn := standard_message mbox true "F";
+    Output.Private_Hooks.error_fn := standard_message mbox true "G";
+    Output.Private_Hooks.urgent_message_fn := ! Output.Private_Hooks.writeln_fn;
+    Output.Private_Hooks.prompt_fn := ignore;
+    message mbox "A" [] (Session.welcome ());
+    in_stream
   end;
 
 end;
@@ -179,8 +173,7 @@
     val _ = Unsynchronized.change print_mode
       (fold (update op =) [isabelle_processN, Keyword.keyword_statusN, Pretty.symbolicN]);
 
-    val (in_stream, out_stream) = setup_channels in_fifo out_fifo;
-    val _ = init_message out_stream;
+    val in_stream = setup_channels in_fifo out_fifo;
     val _ = Keyword.status ();
     val _ = Output.status (Markup.markup Markup.ready "process ready");
   in loop in_stream end));
--- a/src/Pure/Thy/thy_info.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/Pure/Thy/thy_info.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -141,7 +141,7 @@
   else
     let
       val succs = thy_graph Graph.all_succs [name];
-      val _ = priority (loader_msg "removing" succs);
+      val _ = Output.urgent_message (loader_msg "removing" succs);
       val _ = List.app (perform Remove) succs;
       val _ = change_thys (Graph.del_nodes succs);
     in () end);
@@ -222,7 +222,7 @@
 fun load_thy initiators update_time (deps: deps) text name parent_thys =
   let
     val _ = kill_thy name;
-    val _ = priority ("Loading theory " ^ quote name ^ required_by " " initiators);
+    val _ = Output.urgent_message ("Loading theory " ^ quote name ^ required_by " " initiators);
 
     val {master = (thy_path, _), ...} = deps;
     val dir = Path.dir thy_path;
@@ -337,7 +337,7 @@
   in
     NAMED_CRITICAL "Thy_Info" (fn () =>
      (kill_thy name;
-      priority ("Registering theory " ^ quote name);
+      Output.urgent_message ("Registering theory " ^ quote name);
       map get_theory parents;
       change_thys (new_entry name parents (SOME deps, SOME theory));
       perform Update name))
--- a/src/Pure/Thy/thy_syntax.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/Pure/Thy/thy_syntax.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -10,7 +10,7 @@
     (Token.T, (Symbol_Pos.T, Position.T * (Symbol.symbol, (string, 'a)
       Source.source) Source.source) Source.source) Source.source
   val parse_tokens: Scan.lexicon * Scan.lexicon -> Position.T -> string -> Token.T list
-  val present_token: Token.T -> output
+  val present_token: Token.T -> Output.output
   val report_token: Token.T -> unit
   datatype span_kind = Command of string | Ignored | Malformed
   type span
@@ -19,7 +19,7 @@
   val span_range: span -> Position.range
   val span_source: (Token.T, 'a) Source.source -> (span, (Token.T, 'a) Source.source) Source.source
   val parse_spans: Scan.lexicon * Scan.lexicon -> Position.T -> string -> span list
-  val present_span: span -> output
+  val present_span: span -> Output.output
   val report_span: span -> unit
   val atom_source: (span, 'a) Source.source ->
     (span * span list * bool, (span, 'a) Source.source) Source.source
--- a/src/Pure/thm.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/Pure/thm.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -335,14 +335,14 @@
 (*** Derivations and Theorems ***)
 
 abstype thm = Thm of
- deriv *                                        (*derivation*)
- {thy_ref: theory_ref,                          (*dynamic reference to theory*)
-  tags: Properties.T,                           (*additional annotations/comments*)
-  maxidx: int,                                  (*maximum index of any Var or TVar*)
-  shyps: sort Ord_List.T,                        (*sort hypotheses*)
-  hyps: term Ord_List.T,                         (*hypotheses*)
-  tpairs: (term * term) list,                   (*flex-flex pairs*)
-  prop: term}                                   (*conclusion*)
+ deriv *                        (*derivation*)
+ {thy_ref: theory_ref,          (*dynamic reference to theory*)
+  tags: Properties.T,           (*additional annotations/comments*)
+  maxidx: int,                  (*maximum index of any Var or TVar*)
+  shyps: sort Ord_List.T,       (*sort hypotheses*)
+  hyps: term Ord_List.T,        (*hypotheses*)
+  tpairs: (term * term) list,   (*flex-flex pairs*)
+  prop: term}                   (*conclusion*)
 and deriv = Deriv of
  {promises: (serial * thm future) Ord_List.T,
   body: Proofterm.proof_body}
--- a/src/Pure/variable.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/Pure/variable.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -77,7 +77,7 @@
   binds: (typ * term) Vartab.table,     (*term bindings*)
   type_occs: string list Symtab.table,  (*type variables -- possibly within term variables*)
   maxidx: int,                          (*maximum var index*)
-  sorts: sort Ord_List.T,                (*declared sort occurrences*)
+  sorts: sort Ord_List.T,               (*declared sort occurrences*)
   constraints:
     typ Vartab.table *                  (*type constraints*)
     sort Vartab.table};                 (*default sorts*)
--- a/src/Tools/Code/code_preproc.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/Tools/Code/code_preproc.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -480,9 +480,9 @@
     val consts = fold_aterms
       (fn Const (c, _) => insert (op =) c | _ => I) t' [];
     val (algebra', eqngr') = obtain false thy consts [t'];
-    val result = evaluator algebra' eqngr' vs' t';
   in
-    evaluator algebra' eqngr' vs' t'
+    t'
+    |> evaluator algebra' eqngr' vs'
     |> postproc (postprocess_term thy o resubst)
   end;
 
--- a/src/Tools/Code/code_runtime.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/Tools/Code/code_runtime.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -29,6 +29,7 @@
     -> theory -> theory
   datatype truth = Holds
   val put_truth: (unit -> truth) -> Proof.context -> Proof.context
+  val trace: bool Unsynchronized.ref
   val polyml_as_definition: (binding * typ) list -> Path.T list -> theory -> theory
 end;
 
@@ -59,8 +60,11 @@
   #> fold (Code_Target.add_reserved target) ["oo", "ooo", "oooo", "upto", "downto", "orf", "andf"]));
        (*avoid further pervasive infix names*)
 
+val trace = Unsynchronized.ref false;
+
 fun exec verbose code =
-  ML_Context.exec (fn () => Secure.use_text ML_Env.local_context (0, "generated code") false code);
+  (if ! trace then tracing code else ();
+  ML_Context.exec (fn () => Secure.use_text ML_Env.local_context (0, "generated code") false code));
 
 fun value ctxt (get, put, put_ml) (prelude, value) =
   let
--- a/src/Tools/Code_Generator.thy	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/Tools/Code_Generator.thy	Wed Oct 27 11:11:35 2010 -0700
@@ -9,7 +9,7 @@
 uses
   "~~/src/Tools/cache_io.ML"
   "~~/src/Tools/auto_tools.ML"
-  "~~/src/Tools/auto_solve.ML"
+  "~~/src/Tools/solve_direct.ML"
   "~~/src/Tools/quickcheck.ML"
   "~~/src/Tools/value.ML"
   "~~/src/Tools/Code/code_preproc.ML" 
@@ -26,7 +26,7 @@
 begin
 
 setup {*
-  Auto_Solve.setup
+  Solve_Direct.setup
   #> Code_Preproc.setup
   #> Code_Simp.setup
   #> Code_ML.setup
--- a/src/Tools/auto_solve.ML	Wed Oct 27 11:10:36 2010 -0700
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,80 +0,0 @@
-(*  Title:      Tools/auto_solve.ML
-    Author:     Timothy Bourke and Gerwin Klein, NICTA
-
-Check whether a newly stated theorem can be solved directly by an
-existing theorem.  Duplicate lemmas can be detected in this way.
-
-The implementation relies critically on the Find_Theorems solves
-feature.
-*)
-
-signature AUTO_SOLVE =
-sig
-  val auto : bool Unsynchronized.ref
-  val max_solutions : int Unsynchronized.ref
-  val setup : theory -> theory
-end;
-
-structure Auto_Solve : AUTO_SOLVE =
-struct
-
-(* preferences *)
-
-val auto = Unsynchronized.ref false;
-val max_solutions = Unsynchronized.ref 5;
-
-val _ =
-  ProofGeneralPgip.add_preference Preferences.category_tracing
-  (Unsynchronized.setmp auto true (fn () =>
-    Preferences.bool_pref auto
-      "auto-solve" "Try to solve conjectures with existing theorems.") ());
-
-
-(* hook *)
-
-fun auto_solve state =
-  let
-    val ctxt = Proof.context_of state;
-
-    val crits = [(true, Find_Theorems.Solves)];
-    fun find g = snd (Find_Theorems.find_theorems ctxt (SOME g) (SOME (! max_solutions)) false crits);
-
-    fun prt_result (goal, results) =
-      let
-        val msg =
-          (case goal of
-            NONE => "The current goal"
-          | SOME sg => Syntax.string_of_term ctxt (Thm.term_of sg));
-      in
-        Pretty.big_list
-          (msg ^ " could be solved directly with:")
-          (map (Find_Theorems.pretty_thm ctxt) results)
-      end;
-
-    fun seek_against_goal () =
-      (case try Proof.simple_goal state of
-        NONE => NONE
-      | SOME {goal = st, ...} =>
-          let
-            val subgoals = Drule.strip_imp_prems (Thm.cprop_of st);
-            val rs =
-              if length subgoals = 1
-              then [(NONE, find st)]
-              else map (fn sg => (SOME sg, find (Goal.init sg))) subgoals;
-            val results = filter_out (null o snd) rs;
-          in if null results then NONE else SOME results end);
-
-    fun go () =
-      (case try seek_against_goal () of
-        SOME (SOME results) =>
-          (true, state |> Proof.goal_message
-                  (fn () => Pretty.chunks
-                    [Pretty.str "",
-                      Pretty.markup Markup.hilite
-                        (separate (Pretty.brk 0) (map prt_result results))]))
-      | _ => (false, state));
-  in if not (!auto) then (false, state) else go () end;
-
-val setup = Auto_Tools.register_tool ("solve", auto_solve)
-
-end;
--- a/src/Tools/quickcheck.ML	Wed Oct 27 11:10:36 2010 -0700
+++ b/src/Tools/quickcheck.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -188,16 +188,20 @@
           case iterate (fn () => tester (k - 1)) i empty_report
            of (NONE, report) => apsnd (cons report) (with_testers k testers)
             | (SOME q, report) => (SOME q, [report]);
-    fun with_size k reports = if k > size then (NONE, reports)
-      else (if quiet then () else priority ("Test data size: " ^ string_of_int k);
+    fun with_size k reports =
+      if k > size then (NONE, reports)
+      else
+       (if quiet then () else Output.urgent_message ("Test data size: " ^ string_of_int k);
         let
           val (result, new_report) = with_testers k testers
           val reports = ((k, new_report) :: reports)
         in case result of NONE => with_size (k + 1) reports | SOME q => (SOME q, reports) end);
-    val ((result, reports), exec_time) = cpu_time "quickcheck execution"
+    val ((result, reports), exec_time) =
+      TimeLimit.timeLimit (Time.fromSeconds 20) (fn () => cpu_time "quickcheck execution"
       (fn () => apfst
          (fn result => case result of NONE => NONE
-        | SOME ts => SOME (names ~~ ts)) (with_size 1 []))
+        | SOME ts => SOME (names ~~ ts)) (with_size 1 []))) ()
+      handle TimeLimit.TimeOut => error "Reached timeout during Quickcheck"
   in
     (result, ([exec_time, comp_time], if report ctxt then SOME reports else NONE))
   end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/solve_direct.ML	Wed Oct 27 11:11:35 2010 -0700
@@ -0,0 +1,105 @@
+(*  Title:      Tools/solve_direct.ML
+    Author:     Timothy Bourke and Gerwin Klein, NICTA
+
+Check whether a newly stated theorem can be solved directly by an
+existing theorem.  Duplicate lemmas can be detected in this way.
+
+The implementation relies critically on the Find_Theorems solves
+feature.
+*)
+
+signature SOLVE_DIRECT =
+sig
+  val auto: bool Unsynchronized.ref
+  val max_solutions: int Unsynchronized.ref
+  val solve_direct: bool -> Proof.state -> bool * Proof.state
+  val setup: theory -> theory
+end;
+
+structure Solve_Direct: SOLVE_DIRECT =
+struct
+
+val solve_directN = "solve_direct";
+
+
+(* preferences *)
+
+val auto = Unsynchronized.ref false;
+val max_solutions = Unsynchronized.ref 5;
+
+val _ =
+  ProofGeneralPgip.add_preference Preferences.category_tracing
+  (Unsynchronized.setmp auto true (fn () =>
+    Preferences.bool_pref auto
+      "auto-solve-direct"
+      ("Run " ^ quote solve_directN ^ " automatically.")) ());
+
+
+(* solve_direct command *)
+
+fun solve_direct auto state =
+  let
+    val ctxt = Proof.context_of state;
+
+    val crits = [(true, Find_Theorems.Solves)];
+    fun find g =
+      snd (Find_Theorems.find_theorems ctxt (SOME g) (SOME (! max_solutions)) false crits);
+
+    fun prt_result (goal, results) =
+      let
+        val msg =
+          (if auto then "Auto " else "") ^ solve_directN ^ ": " ^
+          (case goal of
+            NONE => "The current goal"
+          | SOME sg => Syntax.string_of_term ctxt (Thm.term_of sg));
+      in
+        Pretty.big_list
+          (msg ^ " can be solved directly with")
+          (map (Find_Theorems.pretty_thm ctxt) results)
+      end;
+
+    fun seek_against_goal () =
+      (case try Proof.simple_goal state of
+        NONE => NONE
+      | SOME {goal = st, ...} =>
+          let
+            val subgoals = Drule.strip_imp_prems (Thm.cprop_of st);
+            val rs =
+              if length subgoals = 1
+              then [(NONE, find st)]
+              else map (fn sg => (SOME sg, find (Goal.init sg))) subgoals;
+            val results = filter_out (null o snd) rs;
+          in if null results then NONE else SOME results end);
+    fun message results = separate (Pretty.brk 0) (map prt_result results);
+  in
+    (case try seek_against_goal () of
+      SOME (SOME results) =>
+        (true,
+          state |>
+           (if auto then
+             Proof.goal_message
+               (fn () =>
+                 Pretty.chunks
+                  [Pretty.str "",
+                   Pretty.markup Markup.hilite (message results)])
+            else
+              tap (fn _ =>
+                Output.urgent_message (Pretty.string_of (Pretty.chunks (message results))))))
+    | _ => (false, state))
+  end;
+
+val _ =
+  Outer_Syntax.improper_command solve_directN
+    "try to solve conjectures directly with existing theorems" Keyword.diag
+    (Scan.succeed
+      (Toplevel.keep (fn state => ignore (solve_direct false (Toplevel.proof_of state)))));
+
+
+(* hook *)
+
+fun auto_solve_direct state =
+  if not (! auto) then (false, state) else solve_direct true state;
+
+val setup = Auto_Tools.register_tool (solve_directN, auto_solve_direct);
+
+end;