merged
authorhuffman
Sat, 30 Oct 2010 12:25:18 -0700
changeset 40328 ae8d187600e7
parent 40327 1dfdbd66093a (current diff)
parent 40291 012ed4426fda (diff)
child 40329 73f2b99b549d
merged
src/FOL/ex/Iff_Oracle.thy
--- a/Admin/churn	Fri Oct 29 17:15:28 2010 -0700
+++ b/Admin/churn	Sat Oct 30 12:25:18 2010 -0700
@@ -2,4 +2,4 @@
 
 ADMIN="$(cd "$(dirname "$0")"; cd "$(pwd -P)"; pwd)"
 cd "$ADMIN"
-hg churn --aliases user-aliases --progress
+hg churn --aliases user-aliases
--- a/Admin/user-aliases	Fri Oct 29 17:15:28 2010 -0700
+++ b/Admin/user-aliases	Sat Oct 30 12:25:18 2010 -0700
@@ -7,3 +7,8 @@
 immler@in.tum.de immler
 tsewell@rubicon.NSW.bigpond.net.au tsewell
 tsewell@nicta.com.au tsewell
+kaliszyk@in.tum.de kaliszyk
+Philipp\ Meyer meyerp
+Timothy\ Bourke tbourke
+noschinl@in.tum.de noschinl
+brianh@cs.pdx.edu huffman
--- a/CONTRIBUTORS	Fri Oct 29 17:15:28 2010 -0700
+++ b/CONTRIBUTORS	Sat Oct 30 12:25:18 2010 -0700
@@ -6,6 +6,9 @@
 Contributions to this Isabelle version
 --------------------------------------
 
+* October 2010: Dmitriy Traytel, TUM
+  Coercive subtyping via subtype constraints.
+
 * September 2010: Florian Haftmann, TUM
   Refined concepts for evaluation, i.e. normalisation of terms using different techniques.
 
--- a/NEWS	Fri Oct 29 17:15:28 2010 -0700
+++ b/NEWS	Sat Oct 30 12:25:18 2010 -0700
@@ -56,6 +56,10 @@
 
 *** Pure ***
 
+* Support for real valued configuration options, using simplistic
+floating-point notation that coincides with the inner syntax for
+float_token.
+
 * Interpretation command 'interpret' accepts a list of equations like
 'interpretation' does.
 
@@ -76,6 +80,11 @@
 
 *** HOL ***
 
+* Quickcheck now has a configurable time limit which is set to 30 seconds
+by default. This can be changed by adding [timeout = n] to the quickcheck
+command. The time limit for auto quickcheck is still set independently,
+by default to 5 seconds.
+
 * 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
@@ -339,6 +348,10 @@
 Fail if the argument is false.  Due to inlining the source position of
 failed assertions is included in the error output.
 
+* Discontinued antiquotation @{theory_ref}, which is obsolete since ML
+text is in practice always evaluated with a stable theory checkpoint.
+Minor INCOMPATIBILITY, use (Theory.check_thy @{theory}) instead.
+
 * Renamed setmp_noncritical to Unsynchronized.setmp to emphasize its
 meaning.
 
--- a/doc-src/IsarImplementation/Thy/Logic.thy	Fri Oct 29 17:15:28 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/Logic.thy	Sat Oct 30 12:25:18 2010 -0700
@@ -201,7 +201,7 @@
   ;
   'sort' sort
   ;
-  ('type\_name' | 'type\_abbrev' | 'nonterminal') nameref
+  ('type_name' | 'type_abbrev' | 'nonterminal') nameref
   ;
   'typ' type
   ;
@@ -437,7 +437,7 @@
   \end{matharray}
 
   \begin{rail}
-  ('const\_name' | 'const\_abbrev') nameref
+  ('const_name' | 'const_abbrev') nameref
   ;
   'const' ('(' (type + ',') ')')?
   ;
--- a/doc-src/IsarImplementation/Thy/ML.thy	Fri Oct 29 17:15:28 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/ML.thy	Sat Oct 30 12:25:18 2010 -0700
@@ -463,7 +463,7 @@
   \end{enumerate}
 
   Such weakly structured layout should be use with great care.  Here
-  is a counter-example involving @{ML_text let} expressions:
+  are some counter-examples involving @{ML_text let} expressions:
 
   \begin{verbatim}
   (* WRONG *)
@@ -472,6 +472,10 @@
       val y = ...
     in ... end
 
+  fun foo x = let
+    val y = ...
+  in ... end
+
   fun foo x =
   let
     val y = ...
@@ -746,7 +750,7 @@
 
   This can be avoided by \emph{canonical argument order}, which
   observes certain standard patterns and minimizes adhoc permutations
-  in their application.  In Isabelle/ML, large portions text can be
+  in their application.  In Isabelle/ML, large portions of 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.
@@ -783,7 +787,7 @@
   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
   @{text "insert c \<circ> insert b \<circ> insert a"}.  The slightly awkward
-  inversion if the composition order is due to conventional
+  inversion of the composition order is due to conventional
   mathematical notation, which can be easily amended as explained
   below.
 *}
@@ -1134,9 +1138,9 @@
   language definition.  It was excluded from the SML97 version to
   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 command transactions), but requires user code to be
-  strictly transparent wrt.\ interrupts.
+  interruptibility (which is an indispensable tool to implement
+  managed 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,
--- a/doc-src/IsarImplementation/Thy/Prelim.thy	Fri Oct 29 17:15:28 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/Prelim.thy	Sat Oct 30 12:25:18 2010 -0700
@@ -226,11 +226,10 @@
 text %mlantiq {*
   \begin{matharray}{rcl}
   @{ML_antiquotation_def "theory"} & : & @{text ML_antiquotation} \\
-  @{ML_antiquotation_def "theory_ref"} & : & @{text ML_antiquotation} \\
   \end{matharray}
 
   \begin{rail}
-  ('theory' | 'theory\_ref') nameref?
+  'theory' nameref?
   ;
   \end{rail}
 
@@ -243,10 +242,6 @@
   theory @{text "A"} of the background theory of the current context
   --- as abstract value.
 
-  \item @{text "@{theory_ref}"} is similar to @{text "@{theory}"}, but
-  produces a @{ML_type theory_ref} via @{ML "Theory.check_thy"} as
-  explained above.
-
   \end{description}
 *}
 
@@ -519,7 +514,7 @@
   of this module do not care about the declaration order, since that
   data structure forces its own arrangement of elements.
 
-  Observe how the @{verbatim merge} operation joins the data slots of
+  Observe how the @{ML_text merge} operation joins the data slots of
   the two constituents: @{ML Ord_List.union} prevents duplication of
   common data from different branches, thus avoiding the danger of
   exponential blowup.  Plain list append etc.\ must never be used for
@@ -602,6 +597,8 @@
   bool Config.T * (theory -> theory)"} \\
   @{index_ML Attrib.config_int: "string -> (Context.generic -> int) ->
   int Config.T * (theory -> theory)"} \\
+  @{index_ML Attrib.config_real: "string -> (Context.generic -> real) ->
+  real Config.T * (theory -> theory)"} \\
   @{index_ML Attrib.config_string: "string -> (Context.generic -> string) ->
   string Config.T * (theory -> theory)"} \\
   \end{mldecls}
@@ -622,9 +619,9 @@
   needs to be applied to the theory initially, in order to make
   concrete declaration syntax available to the user.
 
-  \item @{ML Attrib.config_int} and @{ML Attrib.config_string} work
-  like @{ML Attrib.config_bool}, but for types @{ML_type int} and
-  @{ML_type string}, respectively.
+  \item @{ML Attrib.config_int}, @{ML Attrib.config_real}, and @{ML
+  Attrib.config_string} work like @{ML Attrib.config_bool}, but for
+  types @{ML_type int} and @{ML_type string}, respectively.
 
   \end{description}
 *}
@@ -657,6 +654,17 @@
   ML_val {* @{assert} (Config.get @{context} my_flag = true) *}
 qed
 
+text {* Here is another example involving ML type @{ML_type real}
+  (floating-point numbers). *}
+
+ML {*
+  val (airspeed_velocity, airspeed_velocity_setup) =
+    Attrib.config_real "airspeed_velocity" (K 0.0)
+*}
+setup airspeed_velocity_setup
+
+declare [[airspeed_velocity = 9.9]]
+
 
 section {* Names \label{sec:names} *}
 
--- a/doc-src/IsarImplementation/Thy/Proof.thy	Fri Oct 29 17:15:28 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/Proof.thy	Sat Oct 30 12:25:18 2010 -0700
@@ -182,7 +182,7 @@
 
 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
+  literally: @{text "x"} is mapped again to @{text "x"}, and
   attempting to fix it again in the subsequent context is an error.
   Alternatively, fixed parameters can be renamed explicitly as
   follows: *}
@@ -194,7 +194,7 @@
 *}
 
 text {* The following ML code can now work with the invented names of
-  @{verbatim x1}, @{verbatim x2}, @{verbatim x3}, without depending on
+  @{text x1}, @{text x2}, @{text x3}, without depending on
   the details on the system policy for introducing these variants.
   Recall that within a proof body the system always invents fresh
   ``skolem constants'', e.g.\ as follows: *}
--- a/doc-src/IsarImplementation/Thy/document/Logic.tex	Fri Oct 29 17:15:28 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/document/Logic.tex	Sat Oct 30 12:25:18 2010 -0700
@@ -218,7 +218,7 @@
   ;
   'sort' sort
   ;
-  ('type\_name' | 'type\_abbrev' | 'nonterminal') nameref
+  ('type_name' | 'type_abbrev' | 'nonterminal') nameref
   ;
   'typ' type
   ;
@@ -464,7 +464,7 @@
   \end{matharray}
 
   \begin{rail}
-  ('const\_name' | 'const\_abbrev') nameref
+  ('const_name' | 'const_abbrev') nameref
   ;
   'const' ('(' (type + ',') ')')?
   ;
--- a/doc-src/IsarImplementation/Thy/document/ML.tex	Fri Oct 29 17:15:28 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/document/ML.tex	Sat Oct 30 12:25:18 2010 -0700
@@ -484,7 +484,7 @@
   \end{enumerate}
 
   Such weakly structured layout should be use with great care.  Here
-  is a counter-example involving \verb|let| expressions:
+  are some counter-examples involving \verb|let| expressions:
 
   \begin{verbatim}
   (* WRONG *)
@@ -493,6 +493,10 @@
       val y = ...
     in ... end
 
+  fun foo x = let
+    val y = ...
+  in ... end
+
   fun foo x =
   let
     val y = ...
@@ -944,7 +948,7 @@
 
   This can be avoided by \emph{canonical argument order}, which
   observes certain standard patterns and minimizes adhoc permutations
-  in their application.  In Isabelle/ML, large portions text can be
+  in their application.  In Isabelle/ML, large portions of 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.
@@ -979,7 +983,7 @@
   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
   \isa{insert\ c\ {\isasymcirc}\ insert\ b\ {\isasymcirc}\ insert\ a}.  The slightly awkward
-  inversion if the composition order is due to conventional
+  inversion of the composition order is due to conventional
   mathematical notation, which can be easily amended as explained
   below.%
 \end{isamarkuptext}%
@@ -1446,9 +1450,9 @@
   language definition.  It was excluded from the SML97 version to
   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 command transactions), but requires user code to be
-  strictly transparent wrt.\ interrupts.
+  interruptibility (which is an indispensable tool to implement
+  managed 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,
--- a/doc-src/IsarImplementation/Thy/document/Prelim.tex	Fri Oct 29 17:15:28 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/document/Prelim.tex	Sat Oct 30 12:25:18 2010 -0700
@@ -262,11 +262,10 @@
 \begin{isamarkuptext}%
 \begin{matharray}{rcl}
   \indexdef{}{ML antiquotation}{theory}\hypertarget{ML antiquotation.theory}{\hyperlink{ML antiquotation.theory}{\mbox{\isa{theory}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
-  \indexdef{}{ML antiquotation}{theory\_ref}\hypertarget{ML antiquotation.theory-ref}{\hyperlink{ML antiquotation.theory-ref}{\mbox{\isa{theory{\isacharunderscore}ref}}}} & : & \isa{ML{\isacharunderscore}antiquotation} \\
   \end{matharray}
 
   \begin{rail}
-  ('theory' | 'theory\_ref') nameref?
+  'theory' nameref?
   ;
   \end{rail}
 
@@ -279,10 +278,6 @@
   theory \isa{A} of the background theory of the current context
   --- as abstract value.
 
-  \item \isa{{\isacharat}{\isacharbraceleft}theory{\isacharunderscore}ref{\isacharbraceright}} is similar to \isa{{\isacharat}{\isacharbraceleft}theory{\isacharbraceright}}, but
-  produces a \verb|theory_ref| via \verb|Theory.check_thy| as
-  explained above.
-
   \end{description}%
 \end{isamarkuptext}%
 \isamarkuptrue%
@@ -799,6 +794,8 @@
 \verb|  bool Config.T * (theory -> theory)| \\
   \indexdef{}{ML}{Attrib.config\_int}\verb|Attrib.config_int: string -> (Context.generic -> int) ->|\isasep\isanewline%
 \verb|  int Config.T * (theory -> theory)| \\
+  \indexdef{}{ML}{Attrib.config\_real}\verb|Attrib.config_real: string -> (Context.generic -> real) ->|\isasep\isanewline%
+\verb|  real Config.T * (theory -> theory)| \\
   \indexdef{}{ML}{Attrib.config\_string}\verb|Attrib.config_string: string -> (Context.generic -> string) ->|\isasep\isanewline%
 \verb|  string Config.T * (theory -> theory)| \\
   \end{mldecls}
@@ -818,9 +815,8 @@
   needs to be applied to the theory initially, in order to make
   concrete declaration syntax available to the user.
 
-  \item \verb|Attrib.config_int| and \verb|Attrib.config_string| work
-  like \verb|Attrib.config_bool|, but for types \verb|int| and
-  \verb|string|, respectively.
+  \item \verb|Attrib.config_int|, \verb|Attrib.config_real|, and \verb|Attrib.config_string| work like \verb|Attrib.config_bool|, but for
+  types \verb|int| and \verb|string|, respectively.
 
   \end{description}%
 \end{isamarkuptext}%
@@ -1024,6 +1020,34 @@
 %
 \endisadelimproof
 %
+\begin{isamarkuptext}%
+Here is another example involving ML type \verb|real|
+  (floating-point numbers).%
+\end{isamarkuptext}%
+\isamarkuptrue%
+%
+\isadelimML
+%
+\endisadelimML
+%
+\isatagML
+\isacommand{ML}\isamarkupfalse%
+\ {\isacharverbatimopen}\isanewline
+\ \ val\ {\isacharparenleft}airspeed{\isacharunderscore}velocity{\isacharcomma}\ airspeed{\isacharunderscore}velocity{\isacharunderscore}setup{\isacharparenright}\ {\isacharequal}\isanewline
+\ \ \ \ Attrib{\isachardot}config{\isacharunderscore}real\ {\isachardoublequote}airspeed{\isacharunderscore}velocity{\isachardoublequote}\ {\isacharparenleft}K\ {\isadigit{0}}{\isachardot}{\isadigit{0}}{\isacharparenright}\isanewline
+{\isacharverbatimclose}\isanewline
+\isacommand{setup}\isamarkupfalse%
+\ airspeed{\isacharunderscore}velocity{\isacharunderscore}setup%
+\endisatagML
+{\isafoldML}%
+%
+\isadelimML
+%
+\endisadelimML
+\isanewline
+\isanewline
+\isacommand{declare}\isamarkupfalse%
+\ {\isacharbrackleft}{\isacharbrackleft}airspeed{\isacharunderscore}velocity\ {\isacharequal}\ {\isadigit{9}}{\isachardot}{\isadigit{9}}{\isacharbrackright}{\isacharbrackright}%
 \isamarkupsection{Names \label{sec:names}%
 }
 \isamarkuptrue%
--- a/doc-src/IsarImplementation/Thy/document/Proof.tex	Fri Oct 29 17:15:28 2010 -0700
+++ b/doc-src/IsarImplementation/Thy/document/Proof.tex	Sat Oct 30 12:25:18 2010 -0700
@@ -301,7 +301,7 @@
 \begin{isamarkuptext}%
 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
+  literally: \isa{x} is mapped again to \isa{x}, and
   attempting to fix it again in the subsequent context is an error.
   Alternatively, fixed parameters can be renamed explicitly as
   follows:%
@@ -332,7 +332,7 @@
 %
 \begin{isamarkuptext}%
 The following ML code can now work with the invented names of
-  \verb|x1|, \verb|x2|, \verb|x3|, without depending on
+  \isa{x{\isadigit{1}}}, \isa{x{\isadigit{2}}}, \isa{x{\isadigit{3}}}, without depending on
   the details on the system policy for introducing these variants.
   Recall that within a proof body the system always invents fresh
   ``skolem constants'', e.g.\ as follows:%
--- a/doc-src/IsarRef/Thy/Document_Preparation.thy	Fri Oct 29 17:15:28 2010 -0700
+++ b/doc-src/IsarRef/Thy/Document_Preparation.thy	Sat Oct 30 12:25:18 2010 -0700
@@ -84,7 +84,7 @@
   \begin{rail}
     ('chapter' | 'section' | 'subsection' | 'subsubsection' | 'text') target? text
     ;
-    ('header' | 'text\_raw' | 'sect' | 'subsect' | 'subsubsect' | 'txt' | 'txt\_raw') text
+    ('header' | 'text_raw' | 'sect' | 'subsect' | 'subsubsect' | 'txt' | 'txt_raw') text
     ;
   \end{rail}
 
@@ -198,10 +198,10 @@
       'goals' options |
       'subgoals' options |
       'prf' options thmrefs |
-      'full\_prf' options thmrefs |
+      'full_prf' options thmrefs |
       'ML' options name |
-      'ML\_type' options name |
-      'ML\_struct' options name
+      'ML_type' options name |
+      'ML_struct' options name
     ;
     options: '[' (option * ',') ']'
     ;
@@ -468,7 +468,7 @@
   \end{matharray}
 
   \begin{rail}
-    ('display\_drafts' | 'print\_drafts') (name +)
+    ('display_drafts' | 'print_drafts') (name +)
     ;
   \end{rail}
 
--- a/doc-src/IsarRef/Thy/Generic.thy	Fri Oct 29 17:15:28 2010 -0700
+++ b/doc-src/IsarRef/Thy/Generic.thy	Sat Oct 30 12:25:18 2010 -0700
@@ -6,16 +6,15 @@
 
 section {* Configuration options *}
 
-text {*
-  Isabelle/Pure maintains a record of named configuration options
-  within the theory or proof context, with values of type @{ML_type
-  bool}, @{ML_type int}, or @{ML_type string}.  Tools may declare
-  options in ML, and then refer to these values (relative to the
-  context).  Thus global reference variables are easily avoided.  The
-  user may change the value of a configuration option by means of an
-  associated attribute of the same name.  This form of context
-  declaration works particularly well with commands such as @{command
-  "declare"} or @{command "using"}.
+text {* Isabelle/Pure maintains a record of named configuration
+  options within the theory or proof context, with values of type
+  @{ML_type bool}, @{ML_type int}, @{ML_type real}, or @{ML_type
+  string}.  Tools may declare options in ML, and then refer to these
+  values (relative to the context).  Thus global reference variables
+  are easily avoided.  The user may change the value of a
+  configuration option by means of an associated attribute of the same
+  name.  This form of context declaration works particularly well with
+  commands such as @{command "declare"} or @{command "using"}.
 
   For historical reasons, some tools cannot take the full proof
   context into account and merely refer to the background theory.
@@ -27,7 +26,7 @@
   \end{matharray}
 
   \begin{rail}
-    name ('=' ('true' | 'false' | int | name))?
+    name ('=' ('true' | 'false' | int | float | name))?
   \end{rail}
 
   \begin{description}
@@ -284,14 +283,14 @@
   \end{matharray}
 
   \begin{rail}
-    ( 'rule\_tac' | 'erule\_tac' | 'drule\_tac' | 'frule\_tac' | 'cut\_tac' | 'thin\_tac' ) goalspec?
+    ( 'rule_tac' | 'erule_tac' | 'drule_tac' | 'frule_tac' | 'cut_tac' | 'thin_tac' ) goalspec?
     ( insts thmref | thmrefs )
     ;
-    'subgoal\_tac' goalspec? (prop +)
+    'subgoal_tac' goalspec? (prop +)
     ;
-    'rename\_tac' goalspec? (name +)
+    'rename_tac' goalspec? (name +)
     ;
-    'rotate\_tac' goalspec? int?
+    'rotate_tac' goalspec? int?
     ;
     ('tactic' | 'raw_tactic') text
     ;
@@ -364,10 +363,10 @@
 
   \indexouternonterm{simpmod}
   \begin{rail}
-    ('simp' | 'simp\_all') opt? (simpmod *)
+    ('simp' | 'simp_all') opt? (simpmod *)
     ;
 
-    opt: '(' ('no\_asm' | 'no\_asm\_simp' | 'no\_asm\_use' | 'asm\_lr' ) ')'
+    opt: '(' ('no_asm' | 'no_asm_simp' | 'no_asm_use' | 'asm_lr' ) ')'
     ;
     simpmod: ('add' | 'del' | 'only' | 'cong' (() | 'add' | 'del') |
       'split' (() | 'add' | 'del')) ':' thmrefs
@@ -471,7 +470,7 @@
   \end{matharray}
 
   \begin{rail}
-    'simproc\_setup' name '(' (term + '|') ')' '=' text \\ ('identifier' (nameref+))?
+    'simproc_setup' name '(' (term + '|') ')' '=' text \\ ('identifier' (nameref+))?
     ;
 
     'simproc' (('add' ':')? | 'del' ':') (name+)
@@ -519,7 +518,7 @@
     'simplified' opt? thmrefs?
     ;
 
-    opt: '(' ('no\_asm' | 'no\_asm\_simp' | 'no\_asm\_use') ')'
+    opt: '(' ('no_asm' | 'no_asm_simp' | 'no_asm_use') ')'
     ;
   \end{rail}
 
@@ -785,7 +784,7 @@
     ;
     'atomize' ('(' 'full' ')')?
     ;
-    'rule\_format' ('(' 'noasm' ')')?
+    'rule_format' ('(' 'noasm' ')')?
     ;
   \end{rail}
 
--- a/doc-src/IsarRef/Thy/HOL_Specific.thy	Fri Oct 29 17:15:28 2010 -0700
+++ b/doc-src/IsarRef/Thy/HOL_Specific.thy	Sat Oct 30 12:25:18 2010 -0700
@@ -72,7 +72,7 @@
   \end{matharray}
 
   \begin{rail}
-    'split\_format' ((( name * ) + 'and') | ('(' 'complete' ')'))
+    'split_format' ((( name * ) + 'and') | ('(' 'complete' ')'))
     ;
   \end{rail}
 
@@ -355,7 +355,7 @@
   \begin{rail}
     'datatype' (dtspec + 'and')
     ;
-    'rep\_datatype' ('(' (name +) ')')? (term +)
+    'rep_datatype' ('(' (name +) ')')? (term +)
     ;
 
     dtspec: parname? typespec mixfix? '=' (cons + '|')
@@ -501,9 +501,9 @@
   \begin{rail}
     'relation' term
     ;
-    'lexicographic\_order' ( clasimpmod * )
+    'lexicographic_order' ( clasimpmod * )
     ;
-    'size\_change' ( orders ( clasimpmod * ) )
+    'size_change' ( orders ( clasimpmod * ) )
     ;
     orders: ( 'max' | 'min' | 'ms' ) *
   \end{rail}
@@ -633,7 +633,7 @@
     ;
     hints: '(' 'hints' ( recdefmod * ) ')'
     ;
-    recdefmod: (('recdef\_simp' | 'recdef\_cong' | 'recdef\_wf') (() | 'add' | 'del') ':' thmrefs) | clasimpmod
+    recdefmod: (('recdef_simp' | 'recdef_cong' | 'recdef_wf') (() | 'add' | 'del') ':' thmrefs) | clasimpmod
     ;
     tc: nameref ('(' nat ')')?
     ;
@@ -672,7 +672,7 @@
   \end{matharray}
 
   \begin{rail}
-    ('recdef\_simp' | 'recdef\_cong' | 'recdef\_wf') (() | 'add' | 'del')
+    ('recdef_simp' | 'recdef_cong' | 'recdef_wf') (() | 'add' | 'del')
     ;
   \end{rail}
 *}
@@ -712,7 +712,7 @@
   \end{matharray}
 
   \begin{rail}
-    ('inductive' | 'inductive\_set' | 'coinductive' | 'coinductive\_set') target? fixes ('for' fixes)? \\
+    ('inductive' | 'inductive_set' | 'coinductive' | 'coinductive_set') target? fixes ('for' fixes)? \\
     ('where' clauses)? ('monos' thmrefs)?
     ;
     clauses: (thmdecl? prop + '|')
@@ -941,14 +941,29 @@
 
     \begin{description}
 
-      \item[size] specifies the maximum size of the search space for
-        assignment values.
+    \item[@{text size}] specifies the maximum size of the search space
+    for assignment values.
+
+    \item[@{text iterations}] sets how many sets of assignments are
+    generated for each particular size.
+
+    \item[@{text no_assms}] specifies whether assumptions in
+    structured proofs should be ignored.
+
+    \item[@{text timeout}] sets the time limit in seconds.
 
-      \item[iterations] sets how many sets of assignments are
-        generated for each particular size.
+    \item[@{text default_type}] sets the type(s) generally used to
+    instantiate type variables.
+
+    \item[@{text report}] if set quickcheck reports how many tests
+    fulfilled the preconditions.
 
-      \item[no\_assms] specifies whether assumptions in
-        structured proofs should be ignored.
+    \item[@{text quiet}] if not set quickcheck informs about the
+    current size for assignment values.
+
+    \item[@{text expect}] can be used to check if the user's
+    expectation was met (@{text no_expectation}, @{text
+    no_counterexample}, or @{text counterexample}).
 
     \end{description}
 
@@ -976,13 +991,13 @@
   \end{matharray}
 
   \begin{rail}
-    'case\_tac' goalspec? term rule?
+    'case_tac' goalspec? term rule?
     ;
-    'induct\_tac' goalspec? (insts * 'and') rule?
+    'induct_tac' goalspec? (insts * 'and') rule?
     ;
-    'ind\_cases' (prop +) ('for' (name +)) ?
+    'ind_cases' (prop +) ('for' (name +)) ?
     ;
-    'inductive\_cases' (thmdecl? (prop +) + 'and')
+    'inductive_cases' (thmdecl? (prop +) + 'and')
     ;
 
     rule: ('rule' ':' thmref)
@@ -1063,8 +1078,8 @@
   \end{matharray}
 
   \begin{rail}
-     'export\_code' ( constexpr + ) \\
-       ( ( 'in' target ( 'module\_name' string ) ? \\
+     'export_code' ( constexpr + ) \\
+       ( ( 'in' target ( 'module_name' string ) ? \\
         ( 'file' ( string | '-' ) ) ? ( '(' args ')' ) ?) + ) ?
     ;
 
@@ -1086,10 +1101,10 @@
     'code' ( 'del' | 'abstype' | 'abstract' ) ?
     ;
 
-    'code\_abort' ( const + )
+    'code_abort' ( const + )
     ;
 
-    'code\_datatype' ( const + )
+    'code_datatype' ( const + )
     ;
 
     'code_inline' ( 'del' ) ?
@@ -1098,41 +1113,41 @@
     'code_post' ( 'del' ) ?
     ;
 
-    'code\_thms' ( constexpr + ) ?
+    'code_thms' ( constexpr + ) ?
     ;
 
-    'code\_deps' ( constexpr + ) ?
+    'code_deps' ( constexpr + ) ?
     ;
 
-    'code\_const' (const + 'and') \\
+    'code_const' (const + 'and') \\
       ( ( '(' target ( syntax ? + 'and' ) ')' ) + )
     ;
 
-    'code\_type' (typeconstructor + 'and') \\
+    'code_type' (typeconstructor + 'and') \\
       ( ( '(' target ( syntax ? + 'and' ) ')' ) + )
     ;
 
-    'code\_class' (class + 'and') \\
+    'code_class' (class + 'and') \\
       ( ( '(' target \\ ( string ? + 'and' ) ')' ) + )
     ;
 
-    'code\_instance' (( typeconstructor '::' class ) + 'and') \\
+    'code_instance' (( typeconstructor '::' class ) + 'and') \\
       ( ( '(' target ( '-' ? + 'and' ) ')' ) + )
     ;
 
-    'code\_reserved' target ( string + )
+    'code_reserved' target ( string + )
     ;
 
-    'code\_monad' const const target
+    'code_monad' const const target
     ;
 
-    'code\_include' target ( string ( string | '-') )
+    'code_include' target ( string ( string | '-') )
     ;
 
-    'code\_modulename' target ( ( string string ) + )
+    'code_modulename' target ( ( string string ) + )
     ;
 
-    'code\_reflect' string ( 'datatypes' ( string '=' ( string + '|' ) + 'and' ) ) ? \\
+    'code_reflect' string ( 'datatypes' ( string '=' ( string + '|' ) + 'and' ) ) ? \\
       ( 'functions' ( string + ) ) ? ( 'file' string ) ?
     ;
 
@@ -1266,7 +1281,7 @@
   \end{matharray}
 
   \begin{rail}
-  ( 'code\_module' | 'code\_library' ) modespec ? name ? \\
+  ( 'code_module' | 'code_library' ) modespec ? name ? \\
     ( 'file' name ) ? ( 'imports' ( name + ) ) ? \\
     'contains' ( ( name '=' term ) + | term + )
   ;
@@ -1274,13 +1289,13 @@
   modespec: '(' ( name * ) ')'
   ;
 
-  'consts\_code' (codespec +)
+  'consts_code' (codespec +)
   ;
 
   codespec: const template attachment ?
   ;
 
-  'types\_code' (tycodespec +)
+  'types_code' (tycodespec +)
   ;
 
   tycodespec: name template attachment ?
@@ -1311,7 +1326,7 @@
   \end{matharray}
 
   \begin{rail}
-  ('specification' | 'ax\_specification') '(' (decl +) ')' \\ (thmdecl? prop +)
+  ('specification' | 'ax_specification') '(' (decl +) ')' \\ (thmdecl? prop +)
   ;
   decl: ((name ':')? term '(' 'overloaded' ')'?)
   \end{rail}
--- a/doc-src/IsarRef/Thy/Inner_Syntax.thy	Fri Oct 29 17:15:28 2010 -0700
+++ b/doc-src/IsarRef/Thy/Inner_Syntax.thy	Sat Oct 30 12:25:18 2010 -0700
@@ -31,7 +31,7 @@
     ;
     'thm' modes? thmrefs
     ;
-    ( 'prf' | 'full\_prf' ) modes? thmrefs?
+    ( 'prf' | 'full_prf' ) modes? thmrefs?
     ;
     'pr' modes? nat?
     ;
@@ -364,9 +364,9 @@
   \end{matharray}
 
   \begin{rail}
-    ('type\_notation' | 'no\_type\_notation') target? mode? \\ (nameref mixfix + 'and')
+    ('type_notation' | 'no_type_notation') target? mode? \\ (nameref mixfix + 'and')
     ;
-    ('notation' | 'no\_notation') target? mode? \\ (nameref structmixfix + 'and')
+    ('notation' | 'no_notation') target? mode? \\ (nameref structmixfix + 'and')
     ;
     'write' mode? (nameref structmixfix + 'and')
     ;
@@ -730,9 +730,9 @@
   \begin{rail}
     'nonterminals' (name +)
     ;
-    ('syntax' | 'no\_syntax') mode? (constdecl +)
+    ('syntax' | 'no_syntax') mode? (constdecl +)
     ;
-    ('translations' | 'no\_translations') (transpat ('==' | '=>' | '<=' | rightleftharpoons | rightharpoonup | leftharpoondown) transpat +)
+    ('translations' | 'no_translations') (transpat ('==' | '=>' | '<=' | rightleftharpoons | rightharpoonup | leftharpoondown) transpat +)
     ;
 
     mode: ('(' ( name | 'output' | name 'output' ) ')')
@@ -786,8 +786,8 @@
   \end{matharray}
 
   \begin{rail}
-  ( 'parse\_ast\_translation' | 'parse\_translation' | 'print\_translation' |
-    'typed\_print\_translation' | 'print\_ast\_translation' ) ('(advanced)')? text
+  ( 'parse_ast_translation' | 'parse_translation' | 'print_translation' |
+    'typed_print_translation' | 'print_ast_translation' ) ('(advanced)')? text
   ;
   \end{rail}
 
--- a/doc-src/IsarRef/Thy/Misc.thy	Fri Oct 29 17:15:28 2010 -0700
+++ b/doc-src/IsarRef/Thy/Misc.thy	Sat Oct 30 12:25:18 2010 -0700
@@ -21,19 +21,19 @@
   \end{matharray}
 
   \begin{rail}
-    ('print\_theory' | 'print\_theorems') ('!'?)
+    ('print_theory' | 'print_theorems') ('!'?)
     ;
 
-    'find\_theorems' (('(' (nat)? ('with\_dups')? ')')?) (thmcriterion *)
+    'find_theorems' (('(' (nat)? ('with_dups')? ')')?) (thmcriterion *)
     ;
     thmcriterion: ('-'?) ('name' ':' nameref | 'intro' | 'elim' | 'dest' |
       'solves' | 'simp' ':' term | term)
     ;
-    'find\_consts' (constcriterion *)
+    'find_consts' (constcriterion *)
     ;
     constcriterion: ('-'?) ('name' ':' nameref | 'strict' ':' type | type)
     ;
-    'thm\_deps' thmrefs
+    'thm_deps' thmrefs
     ;
   \end{rail}
 
@@ -143,7 +143,7 @@
   \end{matharray}
 
   \begin{rail}
-    ('cd' | 'use\_thy') name
+    ('cd' | 'use_thy') name
     ;
   \end{rail}
 
--- a/doc-src/IsarRef/Thy/Outer_Syntax.thy	Fri Oct 29 17:15:28 2010 -0700
+++ b/doc-src/IsarRef/Thy/Outer_Syntax.thy	Sat Oct 30 12:25:18 2010 -0700
@@ -108,6 +108,7 @@
     @{syntax_def longident} & = & @{text "ident("}@{verbatim "."}@{text "ident)\<^sup>+"} \\
     @{syntax_def symident} & = & @{text "sym\<^sup>+  |  "}@{verbatim "\\"}@{verbatim "<"}@{text ident}@{verbatim ">"} \\
     @{syntax_def nat} & = & @{text "digit\<^sup>+"} \\
+    @{syntax_def float} & = & @{syntax_ref nat}@{verbatim "."}@{syntax_ref nat}@{text "  |  "}@{verbatim "-"}@{syntax_ref nat}@{verbatim "."}@{syntax_ref nat} \\
     @{syntax_def var} & = & @{verbatim "?"}@{text "ident  |  "}@{verbatim "?"}@{text ident}@{verbatim "."}@{text nat} \\
     @{syntax_def typefree} & = & @{verbatim "'"}@{text ident} \\
     @{syntax_def typevar} & = & @{verbatim "?"}@{text "typefree  |  "}@{verbatim "?"}@{text typefree}@{verbatim "."}@{text nat} \\
@@ -365,7 +366,7 @@
 
   \indexoutertoken{atom}\indexouternonterm{args}\indexouternonterm{attributes}
   \begin{rail}
-    atom: nameref | typefree | typevar | var | nat | keyword
+    atom: nameref | typefree | typevar | var | nat | float | keyword
     ;
     arg: atom | '(' args ')' | '[' args ']'
     ;
--- a/doc-src/IsarRef/Thy/Proof.thy	Fri Oct 29 17:15:28 2010 -0700
+++ b/doc-src/IsarRef/Thy/Proof.thy	Sat Oct 30 12:25:18 2010 -0700
@@ -432,11 +432,11 @@
 
   \begin{rail}
     ('lemma' | 'theorem' | 'corollary' |
-     'schematic\_lemma' | 'schematic\_theorem' | 'schematic\_corollary') target? (goal | longgoal)
+     'schematic_lemma' | 'schematic_theorem' | 'schematic_corollary') target? (goal | longgoal)
     ;
     ('have' | 'show' | 'hence' | 'thus') goal
     ;
-    'print\_statement' modes? thmrefs
+    'print_statement' modes? thmrefs
     ;
   
     goal: (props + 'and')
@@ -834,7 +834,7 @@
   \end{matharray}
 
   \begin{rail}
-    ( 'apply' | 'apply\_end' ) method
+    ( 'apply' | 'apply_end' ) method
     ;
     'defer' nat?
     ;
@@ -896,7 +896,7 @@
   \end{matharray}
 
   \begin{rail}
-    'method\_setup' name '=' text text
+    'method_setup' name '=' text text
     ;
   \end{rail}
 
@@ -1190,9 +1190,9 @@
     caseref: nameref attributes?
     ;
 
-    'case\_names' (name +)
+    'case_names' (name +)
     ;
-    'case\_conclusion' name (name *)
+    'case_conclusion' name (name *)
     ;
     'params' ((name *) + 'and')
     ;
--- a/doc-src/IsarRef/Thy/Spec.thy	Fri Oct 29 17:15:28 2010 -0700
+++ b/doc-src/IsarRef/Thy/Spec.thy	Sat Oct 30 12:25:18 2010 -0700
@@ -159,7 +159,7 @@
 
 text {*
   \begin{matharray}{rcll}
-    @{command_def "axiomatization"} & : & @{text "theory \<rightarrow> theory"} & (axiomatic!)\\
+    @{command_def "axiomatization"} & : & @{text "theory \<rightarrow> theory"} & (axiomatic!) \\
     @{command_def "definition"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
     @{attribute_def "defn"} & : & @{text attribute} \\
     @{command_def "abbreviation"} & : & @{text "local_theory \<rightarrow> local_theory"} \\
@@ -322,8 +322,8 @@
   A locale instance consists of a reference to a locale and either
   positional or named parameter instantiations.  Identical
   instantiations (that is, those that instante a parameter by itself)
-  may be omitted.  The notation `\_' enables to omit the instantiation
-  for a parameter inside a positional instantiation.
+  may be omitted.  The notation `@{text "_"}' enables to omit the
+  instantiation for a parameter inside a positional instantiation.
 
   Terms in instantiations are from the context the locale expressions
   is declared in.  Local names may be added to this context with the
@@ -360,7 +360,7 @@
   \begin{rail}
     'locale' name ('=' locale)? 'begin'?
     ;
-    'print\_locale' '!'? nameref
+    'print_locale' '!'? nameref
     ;
     locale: contextelem+ | localeexpr ('+' (contextelem+))?
     ;
@@ -503,7 +503,7 @@
     ;
     'interpret' localeexpr equations?
     ;
-    'print\_interps' nameref
+    'print_interps' nameref
     ;
     equations: 'where' (thmdecl? prop + 'and')
     ;
@@ -630,9 +630,9 @@
     ;
     'instance' nameref ('<' | subseteq) nameref
     ;
-    'print\_classes'
+    'print_classes'
     ;
-    'class\_deps'
+    'class_deps'
     ;
 
     superclassexpr: nameref | (nameref '+' superclassexpr)
@@ -840,9 +840,9 @@
   \begin{rail}
     'use' name
     ;
-    ('ML' | 'ML\_prf' | 'ML\_val' | 'ML\_command' | 'setup' | 'local\_setup') text
+    ('ML' | 'ML_prf' | 'ML_val' | 'ML_command' | 'setup' | 'local_setup') text
     ;
-    'attribute\_setup' name '=' text text
+    'attribute_setup' name '=' text text
     ;
   \end{rail}
 
@@ -930,7 +930,7 @@
     ;
     'classrel' (nameref ('<' | subseteq) nameref + 'and')
     ;
-    'default\_sort' sort
+    'default_sort' sort
     ;
   \end{rail}
 
@@ -1139,8 +1139,8 @@
   asserted, and records within the internal derivation object how
   presumed theorems depend on unproven suppositions.
 
-  \begin{matharray}{rcl}
-    @{command_def "oracle"} & : & @{text "theory \<rightarrow> theory"} \\
+  \begin{matharray}{rcll}
+    @{command_def "oracle"} & : & @{text "theory \<rightarrow> theory"} & (axiomatic!) \\
   \end{matharray}
 
   \begin{rail}
@@ -1159,7 +1159,7 @@
 
   \end{description}
 
-  See @{"file" "~~/src/FOL/ex/Iff_Oracle.thy"} for a worked example of
+  See @{"file" "~~/src/HOL/ex/Iff_Oracle.thy"} for a worked example of
   defining a new primitive rule as oracle, and turning it into a proof
   method.
 *}
@@ -1176,7 +1176,7 @@
   \end{matharray}
 
   \begin{rail}
-    ( 'hide\_class' | 'hide\_type' | 'hide\_const' | 'hide\_fact' ) ('(open)')? (nameref + )
+    ( 'hide_class' | 'hide_type' | 'hide_const' | 'hide_fact' ) ('(open)')? (nameref + )
     ;
   \end{rail}
 
--- a/doc-src/IsarRef/Thy/ZF_Specific.thy	Fri Oct 29 17:15:28 2010 -0700
+++ b/doc-src/IsarRef/Thy/ZF_Specific.thy	Sat Oct 30 12:25:18 2010 -0700
@@ -67,11 +67,11 @@
     ;
     monos: ('monos' thmrefs)?
     ;
-    condefs: ('con\_defs' thmrefs)?
+    condefs: ('con_defs' thmrefs)?
     ;
-    typeintros: ('type\_intros' thmrefs)?
+    typeintros: ('type_intros' thmrefs)?
     ;
-    typeelims: ('type\_elims' thmrefs)?
+    typeelims: ('type_elims' thmrefs)?
     ;
   \end{rail}
 
@@ -126,7 +126,7 @@
   \end{matharray}
 
   \begin{rail}
-    ('case\_tac' | 'induct\_tac') goalspec? name
+    ('case_tac' | 'induct_tac') goalspec? name
     ;
     indcases (prop +)
     ;
--- a/doc-src/IsarRef/Thy/document/Document_Preparation.tex	Fri Oct 29 17:15:28 2010 -0700
+++ b/doc-src/IsarRef/Thy/document/Document_Preparation.tex	Sat Oct 30 12:25:18 2010 -0700
@@ -105,7 +105,7 @@
   \begin{rail}
     ('chapter' | 'section' | 'subsection' | 'subsubsection' | 'text') target? text
     ;
-    ('header' | 'text\_raw' | 'sect' | 'subsect' | 'subsubsect' | 'txt' | 'txt\_raw') text
+    ('header' | 'text_raw' | 'sect' | 'subsect' | 'subsubsect' | 'txt' | 'txt_raw') text
     ;
   \end{rail}
 
@@ -216,10 +216,10 @@
       'goals' options |
       'subgoals' options |
       'prf' options thmrefs |
-      'full\_prf' options thmrefs |
+      'full_prf' options thmrefs |
       'ML' options name |
-      'ML\_type' options name |
-      'ML\_struct' options name
+      'ML_type' options name |
+      'ML_struct' options name
     ;
     options: '[' (option * ',') ']'
     ;
@@ -487,7 +487,7 @@
   \end{matharray}
 
   \begin{rail}
-    ('display\_drafts' | 'print\_drafts') (name +)
+    ('display_drafts' | 'print_drafts') (name +)
     ;
   \end{rail}
 
--- a/doc-src/IsarRef/Thy/document/Generic.tex	Fri Oct 29 17:15:28 2010 -0700
+++ b/doc-src/IsarRef/Thy/document/Generic.tex	Sat Oct 30 12:25:18 2010 -0700
@@ -27,13 +27,14 @@
 \isamarkuptrue%
 %
 \begin{isamarkuptext}%
-Isabelle/Pure maintains a record of named configuration options
-  within the theory or proof context, with values of type \verb|bool|, \verb|int|, or \verb|string|.  Tools may declare
-  options in ML, and then refer to these values (relative to the
-  context).  Thus global reference variables are easily avoided.  The
-  user may change the value of a configuration option by means of an
-  associated attribute of the same name.  This form of context
-  declaration works particularly well with commands such as \hyperlink{command.declare}{\mbox{\isa{\isacommand{declare}}}} or \hyperlink{command.using}{\mbox{\isa{\isacommand{using}}}}.
+Isabelle/Pure maintains a record of named configuration
+  options within the theory or proof context, with values of type
+  \verb|bool|, \verb|int|, \verb|real|, or \verb|string|.  Tools may declare options in ML, and then refer to these
+  values (relative to the context).  Thus global reference variables
+  are easily avoided.  The user may change the value of a
+  configuration option by means of an associated attribute of the same
+  name.  This form of context declaration works particularly well with
+  commands such as \hyperlink{command.declare}{\mbox{\isa{\isacommand{declare}}}} or \hyperlink{command.using}{\mbox{\isa{\isacommand{using}}}}.
 
   For historical reasons, some tools cannot take the full proof
   context into account and merely refer to the background theory.
@@ -45,7 +46,7 @@
   \end{matharray}
 
   \begin{rail}
-    name ('=' ('true' | 'false' | int | name))?
+    name ('=' ('true' | 'false' | int | float | name))?
   \end{rail}
 
   \begin{description}
@@ -302,14 +303,14 @@
   \end{matharray}
 
   \begin{rail}
-    ( 'rule\_tac' | 'erule\_tac' | 'drule\_tac' | 'frule\_tac' | 'cut\_tac' | 'thin\_tac' ) goalspec?
+    ( 'rule_tac' | 'erule_tac' | 'drule_tac' | 'frule_tac' | 'cut_tac' | 'thin_tac' ) goalspec?
     ( insts thmref | thmrefs )
     ;
-    'subgoal\_tac' goalspec? (prop +)
+    'subgoal_tac' goalspec? (prop +)
     ;
-    'rename\_tac' goalspec? (name +)
+    'rename_tac' goalspec? (name +)
     ;
-    'rotate\_tac' goalspec? int?
+    'rotate_tac' goalspec? int?
     ;
     ('tactic' | 'raw_tactic') text
     ;
@@ -384,10 +385,10 @@
 
   \indexouternonterm{simpmod}
   \begin{rail}
-    ('simp' | 'simp\_all') opt? (simpmod *)
+    ('simp' | 'simp_all') opt? (simpmod *)
     ;
 
-    opt: '(' ('no\_asm' | 'no\_asm\_simp' | 'no\_asm\_use' | 'asm\_lr' ) ')'
+    opt: '(' ('no_asm' | 'no_asm_simp' | 'no_asm_use' | 'asm_lr' ) ')'
     ;
     simpmod: ('add' | 'del' | 'only' | 'cong' (() | 'add' | 'del') |
       'split' (() | 'add' | 'del')) ':' thmrefs
@@ -490,7 +491,7 @@
   \end{matharray}
 
   \begin{rail}
-    'simproc\_setup' name '(' (term + '|') ')' '=' text \\ ('identifier' (nameref+))?
+    'simproc_setup' name '(' (term + '|') ')' '=' text \\ ('identifier' (nameref+))?
     ;
 
     'simproc' (('add' ':')? | 'del' ':') (name+)
@@ -538,7 +539,7 @@
     'simplified' opt? thmrefs?
     ;
 
-    opt: '(' ('no\_asm' | 'no\_asm\_simp' | 'no\_asm\_use') ')'
+    opt: '(' ('no_asm' | 'no_asm_simp' | 'no_asm_use') ')'
     ;
   \end{rail}
 
@@ -809,7 +810,7 @@
     ;
     'atomize' ('(' 'full' ')')?
     ;
-    'rule\_format' ('(' 'noasm' ')')?
+    'rule_format' ('(' 'noasm' ')')?
     ;
   \end{rail}
 
--- a/doc-src/IsarRef/Thy/document/HOL_Specific.tex	Fri Oct 29 17:15:28 2010 -0700
+++ b/doc-src/IsarRef/Thy/document/HOL_Specific.tex	Sat Oct 30 12:25:18 2010 -0700
@@ -93,7 +93,7 @@
   \end{matharray}
 
   \begin{rail}
-    'split\_format' ((( name * ) + 'and') | ('(' 'complete' ')'))
+    'split_format' ((( name * ) + 'and') | ('(' 'complete' ')'))
     ;
   \end{rail}
 
@@ -364,7 +364,7 @@
   \begin{rail}
     'datatype' (dtspec + 'and')
     ;
-    'rep\_datatype' ('(' (name +) ')')? (term +)
+    'rep_datatype' ('(' (name +) ')')? (term +)
     ;
 
     dtspec: parname? typespec mixfix? '=' (cons + '|')
@@ -511,9 +511,9 @@
   \begin{rail}
     'relation' term
     ;
-    'lexicographic\_order' ( clasimpmod * )
+    'lexicographic_order' ( clasimpmod * )
     ;
-    'size\_change' ( orders ( clasimpmod * ) )
+    'size_change' ( orders ( clasimpmod * ) )
     ;
     orders: ( 'max' | 'min' | 'ms' ) *
   \end{rail}
@@ -642,7 +642,7 @@
     ;
     hints: '(' 'hints' ( recdefmod * ) ')'
     ;
-    recdefmod: (('recdef\_simp' | 'recdef\_cong' | 'recdef\_wf') (() | 'add' | 'del') ':' thmrefs) | clasimpmod
+    recdefmod: (('recdef_simp' | 'recdef_cong' | 'recdef_wf') (() | 'add' | 'del') ':' thmrefs) | clasimpmod
     ;
     tc: nameref ('(' nat ')')?
     ;
@@ -680,7 +680,7 @@
   \end{matharray}
 
   \begin{rail}
-    ('recdef\_simp' | 'recdef\_cong' | 'recdef\_wf') (() | 'add' | 'del')
+    ('recdef_simp' | 'recdef_cong' | 'recdef_wf') (() | 'add' | 'del')
     ;
   \end{rail}%
 \end{isamarkuptext}%
@@ -722,7 +722,7 @@
   \end{matharray}
 
   \begin{rail}
-    ('inductive' | 'inductive\_set' | 'coinductive' | 'coinductive\_set') target? fixes ('for' fixes)? \\
+    ('inductive' | 'inductive_set' | 'coinductive' | 'coinductive_set') target? fixes ('for' fixes)? \\
     ('where' clauses)? ('monos' thmrefs)?
     ;
     clauses: (thmdecl? prop + '|')
@@ -959,14 +959,28 @@
 
     \begin{description}
 
-      \item[size] specifies the maximum size of the search space for
-        assignment values.
+    \item[\isa{size}] specifies the maximum size of the search space
+    for assignment values.
+
+    \item[\isa{iterations}] sets how many sets of assignments are
+    generated for each particular size.
+
+    \item[\isa{no{\isacharunderscore}assms}] specifies whether assumptions in
+    structured proofs should be ignored.
+
+    \item[\isa{timeout}] sets the time limit in seconds.
 
-      \item[iterations] sets how many sets of assignments are
-        generated for each particular size.
+    \item[\isa{default{\isacharunderscore}type}] sets the type(s) generally used to
+    instantiate type variables.
+
+    \item[\isa{report}] if set quickcheck reports how many tests
+    fulfilled the preconditions.
 
-      \item[no\_assms] specifies whether assumptions in
-        structured proofs should be ignored.
+    \item[\isa{quiet}] if not set quickcheck informs about the
+    current size for assignment values.
+
+    \item[\isa{expect}] can be used to check if the user's
+    expectation was met (\isa{no{\isacharunderscore}expectation}, \isa{no{\isacharunderscore}counterexample}, or \isa{counterexample}).
 
     \end{description}
 
@@ -996,13 +1010,13 @@
   \end{matharray}
 
   \begin{rail}
-    'case\_tac' goalspec? term rule?
+    'case_tac' goalspec? term rule?
     ;
-    'induct\_tac' goalspec? (insts * 'and') rule?
+    'induct_tac' goalspec? (insts * 'and') rule?
     ;
-    'ind\_cases' (prop +) ('for' (name +)) ?
+    'ind_cases' (prop +) ('for' (name +)) ?
     ;
-    'inductive\_cases' (thmdecl? (prop +) + 'and')
+    'inductive_cases' (thmdecl? (prop +) + 'and')
     ;
 
     rule: ('rule' ':' thmref)
@@ -1079,8 +1093,8 @@
   \end{matharray}
 
   \begin{rail}
-     'export\_code' ( constexpr + ) \\
-       ( ( 'in' target ( 'module\_name' string ) ? \\
+     'export_code' ( constexpr + ) \\
+       ( ( 'in' target ( 'module_name' string ) ? \\
         ( 'file' ( string | '-' ) ) ? ( '(' args ')' ) ?) + ) ?
     ;
 
@@ -1102,10 +1116,10 @@
     'code' ( 'del' | 'abstype' | 'abstract' ) ?
     ;
 
-    'code\_abort' ( const + )
+    'code_abort' ( const + )
     ;
 
-    'code\_datatype' ( const + )
+    'code_datatype' ( const + )
     ;
 
     'code_inline' ( 'del' ) ?
@@ -1114,41 +1128,41 @@
     'code_post' ( 'del' ) ?
     ;
 
-    'code\_thms' ( constexpr + ) ?
+    'code_thms' ( constexpr + ) ?
     ;
 
-    'code\_deps' ( constexpr + ) ?
+    'code_deps' ( constexpr + ) ?
     ;
 
-    'code\_const' (const + 'and') \\
+    'code_const' (const + 'and') \\
       ( ( '(' target ( syntax ? + 'and' ) ')' ) + )
     ;
 
-    'code\_type' (typeconstructor + 'and') \\
+    'code_type' (typeconstructor + 'and') \\
       ( ( '(' target ( syntax ? + 'and' ) ')' ) + )
     ;
 
-    'code\_class' (class + 'and') \\
+    'code_class' (class + 'and') \\
       ( ( '(' target \\ ( string ? + 'and' ) ')' ) + )
     ;
 
-    'code\_instance' (( typeconstructor '::' class ) + 'and') \\
+    'code_instance' (( typeconstructor '::' class ) + 'and') \\
       ( ( '(' target ( '-' ? + 'and' ) ')' ) + )
     ;
 
-    'code\_reserved' target ( string + )
+    'code_reserved' target ( string + )
     ;
 
-    'code\_monad' const const target
+    'code_monad' const const target
     ;
 
-    'code\_include' target ( string ( string | '-') )
+    'code_include' target ( string ( string | '-') )
     ;
 
-    'code\_modulename' target ( ( string string ) + )
+    'code_modulename' target ( ( string string ) + )
     ;
 
-    'code\_reflect' string ( 'datatypes' ( string '=' ( string + '|' ) + 'and' ) ) ? \\
+    'code_reflect' string ( 'datatypes' ( string '=' ( string + '|' ) + 'and' ) ) ? \\
       ( 'functions' ( string + ) ) ? ( 'file' string ) ?
     ;
 
@@ -1279,7 +1293,7 @@
   \end{matharray}
 
   \begin{rail}
-  ( 'code\_module' | 'code\_library' ) modespec ? name ? \\
+  ( 'code_module' | 'code_library' ) modespec ? name ? \\
     ( 'file' name ) ? ( 'imports' ( name + ) ) ? \\
     'contains' ( ( name '=' term ) + | term + )
   ;
@@ -1287,13 +1301,13 @@
   modespec: '(' ( name * ) ')'
   ;
 
-  'consts\_code' (codespec +)
+  'consts_code' (codespec +)
   ;
 
   codespec: const template attachment ?
   ;
 
-  'types\_code' (tycodespec +)
+  'types_code' (tycodespec +)
   ;
 
   tycodespec: name template attachment ?
@@ -1325,7 +1339,7 @@
   \end{matharray}
 
   \begin{rail}
-  ('specification' | 'ax\_specification') '(' (decl +) ')' \\ (thmdecl? prop +)
+  ('specification' | 'ax_specification') '(' (decl +) ')' \\ (thmdecl? prop +)
   ;
   decl: ((name ':')? term '(' 'overloaded' ')'?)
   \end{rail}
--- a/doc-src/IsarRef/Thy/document/Inner_Syntax.tex	Fri Oct 29 17:15:28 2010 -0700
+++ b/doc-src/IsarRef/Thy/document/Inner_Syntax.tex	Sat Oct 30 12:25:18 2010 -0700
@@ -53,7 +53,7 @@
     ;
     'thm' modes? thmrefs
     ;
-    ( 'prf' | 'full\_prf' ) modes? thmrefs?
+    ( 'prf' | 'full_prf' ) modes? thmrefs?
     ;
     'pr' modes? nat?
     ;
@@ -387,9 +387,9 @@
   \end{matharray}
 
   \begin{rail}
-    ('type\_notation' | 'no\_type\_notation') target? mode? \\ (nameref mixfix + 'and')
+    ('type_notation' | 'no_type_notation') target? mode? \\ (nameref mixfix + 'and')
     ;
-    ('notation' | 'no\_notation') target? mode? \\ (nameref structmixfix + 'and')
+    ('notation' | 'no_notation') target? mode? \\ (nameref structmixfix + 'and')
     ;
     'write' mode? (nameref structmixfix + 'and')
     ;
@@ -749,9 +749,9 @@
   \begin{rail}
     'nonterminals' (name +)
     ;
-    ('syntax' | 'no\_syntax') mode? (constdecl +)
+    ('syntax' | 'no_syntax') mode? (constdecl +)
     ;
-    ('translations' | 'no\_translations') (transpat ('==' | '=>' | '<=' | rightleftharpoons | rightharpoonup | leftharpoondown) transpat +)
+    ('translations' | 'no_translations') (transpat ('==' | '=>' | '<=' | rightleftharpoons | rightharpoonup | leftharpoondown) transpat +)
     ;
 
     mode: ('(' ( name | 'output' | name 'output' ) ')')
@@ -806,8 +806,8 @@
   \end{matharray}
 
   \begin{rail}
-  ( 'parse\_ast\_translation' | 'parse\_translation' | 'print\_translation' |
-    'typed\_print\_translation' | 'print\_ast\_translation' ) ('(advanced)')? text
+  ( 'parse_ast_translation' | 'parse_translation' | 'print_translation' |
+    'typed_print_translation' | 'print_ast_translation' ) ('(advanced)')? text
   ;
   \end{rail}
 
--- a/doc-src/IsarRef/Thy/document/Misc.tex	Fri Oct 29 17:15:28 2010 -0700
+++ b/doc-src/IsarRef/Thy/document/Misc.tex	Sat Oct 30 12:25:18 2010 -0700
@@ -41,19 +41,19 @@
   \end{matharray}
 
   \begin{rail}
-    ('print\_theory' | 'print\_theorems') ('!'?)
+    ('print_theory' | 'print_theorems') ('!'?)
     ;
 
-    'find\_theorems' (('(' (nat)? ('with\_dups')? ')')?) (thmcriterion *)
+    'find_theorems' (('(' (nat)? ('with_dups')? ')')?) (thmcriterion *)
     ;
     thmcriterion: ('-'?) ('name' ':' nameref | 'intro' | 'elim' | 'dest' |
       'solves' | 'simp' ':' term | term)
     ;
-    'find\_consts' (constcriterion *)
+    'find_consts' (constcriterion *)
     ;
     constcriterion: ('-'?) ('name' ':' nameref | 'strict' ':' type | type)
     ;
-    'thm\_deps' thmrefs
+    'thm_deps' thmrefs
     ;
   \end{rail}
 
@@ -164,7 +164,7 @@
   \end{matharray}
 
   \begin{rail}
-    ('cd' | 'use\_thy') name
+    ('cd' | 'use_thy') name
     ;
   \end{rail}
 
--- a/doc-src/IsarRef/Thy/document/Outer_Syntax.tex	Fri Oct 29 17:15:28 2010 -0700
+++ b/doc-src/IsarRef/Thy/document/Outer_Syntax.tex	Sat Oct 30 12:25:18 2010 -0700
@@ -127,6 +127,7 @@
     \indexdef{}{syntax}{longident}\hypertarget{syntax.longident}{\hyperlink{syntax.longident}{\mbox{\isa{longident}}}} & = & \isa{{\isachardoublequote}ident{\isacharparenleft}{\isachardoublequote}}\verb|.|\isa{{\isachardoublequote}ident{\isacharparenright}\isactrlsup {\isacharplus}{\isachardoublequote}} \\
     \indexdef{}{syntax}{symident}\hypertarget{syntax.symident}{\hyperlink{syntax.symident}{\mbox{\isa{symident}}}} & = & \isa{{\isachardoublequote}sym\isactrlsup {\isacharplus}\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|\|\verb|<|\isa{ident}\verb|>| \\
     \indexdef{}{syntax}{nat}\hypertarget{syntax.nat}{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}} & = & \isa{{\isachardoublequote}digit\isactrlsup {\isacharplus}{\isachardoublequote}} \\
+    \indexdef{}{syntax}{float}\hypertarget{syntax.float}{\hyperlink{syntax.float}{\mbox{\isa{float}}}} & = & \indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}}\verb|.|\indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}}\isa{{\isachardoublequote}\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|-|\indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}}\verb|.|\indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}} \\
     \indexdef{}{syntax}{var}\hypertarget{syntax.var}{\hyperlink{syntax.var}{\mbox{\isa{var}}}} & = & \verb|?|\isa{{\isachardoublequote}ident\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|?|\isa{ident}\verb|.|\isa{nat} \\
     \indexdef{}{syntax}{typefree}\hypertarget{syntax.typefree}{\hyperlink{syntax.typefree}{\mbox{\isa{typefree}}}} & = & \verb|'|\isa{ident} \\
     \indexdef{}{syntax}{typevar}\hypertarget{syntax.typevar}{\hyperlink{syntax.typevar}{\mbox{\isa{typevar}}}} & = & \verb|?|\isa{{\isachardoublequote}typefree\ \ {\isacharbar}\ \ {\isachardoublequote}}\verb|?|\isa{typefree}\verb|.|\isa{nat} \\
@@ -391,7 +392,7 @@
 
   \indexoutertoken{atom}\indexouternonterm{args}\indexouternonterm{attributes}
   \begin{rail}
-    atom: nameref | typefree | typevar | var | nat | keyword
+    atom: nameref | typefree | typevar | var | nat | float | keyword
     ;
     arg: atom | '(' args ')' | '[' args ']'
     ;
--- a/doc-src/IsarRef/Thy/document/Proof.tex	Fri Oct 29 17:15:28 2010 -0700
+++ b/doc-src/IsarRef/Thy/document/Proof.tex	Sat Oct 30 12:25:18 2010 -0700
@@ -453,11 +453,11 @@
 
   \begin{rail}
     ('lemma' | 'theorem' | 'corollary' |
-     'schematic\_lemma' | 'schematic\_theorem' | 'schematic\_corollary') target? (goal | longgoal)
+     'schematic_lemma' | 'schematic_theorem' | 'schematic_corollary') target? (goal | longgoal)
     ;
     ('have' | 'show' | 'hence' | 'thus') goal
     ;
-    'print\_statement' modes? thmrefs
+    'print_statement' modes? thmrefs
     ;
   
     goal: (props + 'and')
@@ -843,7 +843,7 @@
   \end{matharray}
 
   \begin{rail}
-    ( 'apply' | 'apply\_end' ) method
+    ( 'apply' | 'apply_end' ) method
     ;
     'defer' nat?
     ;
@@ -903,7 +903,7 @@
   \end{matharray}
 
   \begin{rail}
-    'method\_setup' name '=' text text
+    'method_setup' name '=' text text
     ;
   \end{rail}
 
@@ -1192,9 +1192,9 @@
     caseref: nameref attributes?
     ;
 
-    'case\_names' (name +)
+    'case_names' (name +)
     ;
-    'case\_conclusion' name (name *)
+    'case_conclusion' name (name *)
     ;
     'params' ((name *) + 'and')
     ;
--- a/doc-src/IsarRef/Thy/document/Spec.tex	Fri Oct 29 17:15:28 2010 -0700
+++ b/doc-src/IsarRef/Thy/document/Spec.tex	Sat Oct 30 12:25:18 2010 -0700
@@ -179,7 +179,7 @@
 %
 \begin{isamarkuptext}%
 \begin{matharray}{rcll}
-    \indexdef{}{command}{axiomatization}\hypertarget{command.axiomatization}{\hyperlink{command.axiomatization}{\mbox{\isa{\isacommand{axiomatization}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} & (axiomatic!)\\
+    \indexdef{}{command}{axiomatization}\hypertarget{command.axiomatization}{\hyperlink{command.axiomatization}{\mbox{\isa{\isacommand{axiomatization}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} & (axiomatic!) \\
     \indexdef{}{command}{definition}\hypertarget{command.definition}{\hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}} & : & \isa{{\isachardoublequote}local{\isacharunderscore}theory\ {\isasymrightarrow}\ local{\isacharunderscore}theory{\isachardoublequote}} \\
     \indexdef{}{attribute}{defn}\hypertarget{attribute.defn}{\hyperlink{attribute.defn}{\mbox{\isa{defn}}}} & : & \isa{attribute} \\
     \indexdef{}{command}{abbreviation}\hypertarget{command.abbreviation}{\hyperlink{command.abbreviation}{\mbox{\isa{\isacommand{abbreviation}}}}} & : & \isa{{\isachardoublequote}local{\isacharunderscore}theory\ {\isasymrightarrow}\ local{\isacharunderscore}theory{\isachardoublequote}} \\
@@ -346,8 +346,8 @@
   A locale instance consists of a reference to a locale and either
   positional or named parameter instantiations.  Identical
   instantiations (that is, those that instante a parameter by itself)
-  may be omitted.  The notation `\_' enables to omit the instantiation
-  for a parameter inside a positional instantiation.
+  may be omitted.  The notation `\isa{{\isachardoublequote}{\isacharunderscore}{\isachardoublequote}}' enables to omit the
+  instantiation for a parameter inside a positional instantiation.
 
   Terms in instantiations are from the context the locale expressions
   is declared in.  Local names may be added to this context with the
@@ -385,7 +385,7 @@
   \begin{rail}
     'locale' name ('=' locale)? 'begin'?
     ;
-    'print\_locale' '!'? nameref
+    'print_locale' '!'? nameref
     ;
     locale: contextelem+ | localeexpr ('+' (contextelem+))?
     ;
@@ -525,7 +525,7 @@
     ;
     'interpret' localeexpr equations?
     ;
-    'print\_interps' nameref
+    'print_interps' nameref
     ;
     equations: 'where' (thmdecl? prop + 'and')
     ;
@@ -653,9 +653,9 @@
     ;
     'instance' nameref ('<' | subseteq) nameref
     ;
-    'print\_classes'
+    'print_classes'
     ;
-    'class\_deps'
+    'class_deps'
     ;
 
     superclassexpr: nameref | (nameref '+' superclassexpr)
@@ -862,9 +862,9 @@
   \begin{rail}
     'use' name
     ;
-    ('ML' | 'ML\_prf' | 'ML\_val' | 'ML\_command' | 'setup' | 'local\_setup') text
+    ('ML' | 'ML_prf' | 'ML_val' | 'ML_command' | 'setup' | 'local_setup') text
     ;
-    'attribute\_setup' name '=' text text
+    'attribute_setup' name '=' text text
     ;
   \end{rail}
 
@@ -966,7 +966,7 @@
     ;
     'classrel' (nameref ('<' | subseteq) nameref + 'and')
     ;
-    'default\_sort' sort
+    'default_sort' sort
     ;
   \end{rail}
 
@@ -1179,8 +1179,8 @@
   asserted, and records within the internal derivation object how
   presumed theorems depend on unproven suppositions.
 
-  \begin{matharray}{rcl}
-    \indexdef{}{command}{oracle}\hypertarget{command.oracle}{\hyperlink{command.oracle}{\mbox{\isa{\isacommand{oracle}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} \\
+  \begin{matharray}{rcll}
+    \indexdef{}{command}{oracle}\hypertarget{command.oracle}{\hyperlink{command.oracle}{\mbox{\isa{\isacommand{oracle}}}}} & : & \isa{{\isachardoublequote}theory\ {\isasymrightarrow}\ theory{\isachardoublequote}} & (axiomatic!) \\
   \end{matharray}
 
   \begin{rail}
@@ -1199,7 +1199,7 @@
 
   \end{description}
 
-  See \hyperlink{file.~~/src/FOL/ex/Iff-Oracle.thy}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}FOL{\isacharslash}ex{\isacharslash}Iff{\isacharunderscore}Oracle{\isachardot}thy}}}} for a worked example of
+  See \hyperlink{file.~~/src/HOL/ex/Iff-Oracle.thy}{\mbox{\isa{\isatt{{\isachartilde}{\isachartilde}{\isacharslash}src{\isacharslash}HOL{\isacharslash}ex{\isacharslash}Iff{\isacharunderscore}Oracle{\isachardot}thy}}}} for a worked example of
   defining a new primitive rule as oracle, and turning it into a proof
   method.%
 \end{isamarkuptext}%
@@ -1218,7 +1218,7 @@
   \end{matharray}
 
   \begin{rail}
-    ( 'hide\_class' | 'hide\_type' | 'hide\_const' | 'hide\_fact' ) ('(open)')? (nameref + )
+    ( 'hide_class' | 'hide_type' | 'hide_const' | 'hide_fact' ) ('(open)')? (nameref + )
     ;
   \end{rail}
 
--- a/doc-src/IsarRef/Thy/document/ZF_Specific.tex	Fri Oct 29 17:15:28 2010 -0700
+++ b/doc-src/IsarRef/Thy/document/ZF_Specific.tex	Sat Oct 30 12:25:18 2010 -0700
@@ -91,11 +91,11 @@
     ;
     monos: ('monos' thmrefs)?
     ;
-    condefs: ('con\_defs' thmrefs)?
+    condefs: ('con_defs' thmrefs)?
     ;
-    typeintros: ('type\_intros' thmrefs)?
+    typeintros: ('type_intros' thmrefs)?
     ;
-    typeelims: ('type\_elims' thmrefs)?
+    typeelims: ('type_elims' thmrefs)?
     ;
   \end{rail}
 
@@ -153,7 +153,7 @@
   \end{matharray}
 
   \begin{rail}
-    ('case\_tac' | 'induct\_tac') goalspec? name
+    ('case_tac' | 'induct_tac') goalspec? name
     ;
     indcases (prop +)
     ;
--- a/doc-src/Main/Docs/Main_Doc.thy	Fri Oct 29 17:15:28 2010 -0700
+++ b/doc-src/Main/Docs/Main_Doc.thy	Sat Oct 30 12:25:18 2010 -0700
@@ -501,6 +501,7 @@
 @{const List.lexn} & @{term_type_only List.lexn "('a*'a)set\<Rightarrow>nat\<Rightarrow>('a list * 'a list)set"}\\
 @{const List.lexord} & @{term_type_only List.lexord "('a*'a)set\<Rightarrow>('a list * 'a list)set"}\\
 @{const List.listrel} & @{term_type_only List.listrel "('a*'a)set\<Rightarrow>('a list * 'a list)set"}\\
+@{const List.listrel1} & @{term_type_only List.listrel1 "('a*'a)set\<Rightarrow>('a list * 'a list)set"}\\
 @{const List.lists} & @{term_type_only List.lists "'a set\<Rightarrow>'a list set"}\\
 @{const List.listset} & @{term_type_only List.listset "'a set list \<Rightarrow> 'a list set"}\\
 @{const List.listsum} & @{typeof List.listsum}\\
--- a/doc-src/Main/Docs/document/Main_Doc.tex	Fri Oct 29 17:15:28 2010 -0700
+++ b/doc-src/Main/Docs/document/Main_Doc.tex	Sat Oct 30 12:25:18 2010 -0700
@@ -511,6 +511,7 @@
 \isa{lexn} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ set\ {\isasymRightarrow}\ nat\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ list\ {\isasymtimes}\ {\isacharprime}a\ list{\isacharparenright}\ set}\\
 \isa{lexord} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ set\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ list\ {\isasymtimes}\ {\isacharprime}a\ list{\isacharparenright}\ set}\\
 \isa{listrel} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ set\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ list\ {\isasymtimes}\ {\isacharprime}a\ list{\isacharparenright}\ set}\\
+\isa{listrel{\isadigit{1}}} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}a{\isacharparenright}\ set\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ list\ {\isasymtimes}\ {\isacharprime}a\ list{\isacharparenright}\ set}\\
 \isa{lists} & \isa{{\isacharprime}a\ set\ {\isasymRightarrow}\ {\isacharprime}a\ list\ set}\\
 \isa{listset} & \isa{{\isacharprime}a\ set\ list\ {\isasymRightarrow}\ {\isacharprime}a\ list\ set}\\
 \isa{listsum} & \isa{{\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}a}\\
@@ -565,15 +566,15 @@
 \bigskip
 
 \begin{supertabular}{@ {} l @ {~::~} l @ {}}
-\isa{Map{\isachardot}empty} & \isa{{\isacharprime}a\ {\isasymrightharpoonup}\ {\isacharprime}b}\\
-\isa{op\ {\isacharplus}{\isacharplus}} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymrightharpoonup}\ {\isacharprime}b{\isacharparenright}\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymrightharpoonup}\ {\isacharprime}b{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymrightharpoonup}\ {\isacharprime}b}\\
-\isa{op\ {\isasymcirc}\isactrlsub m} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymrightharpoonup}\ {\isacharprime}b{\isacharparenright}\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}c\ {\isasymrightharpoonup}\ {\isacharprime}a{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}c\ {\isasymrightharpoonup}\ {\isacharprime}b}\\
-\isa{op\ {\isacharbar}{\isacharbackquote}} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymrightharpoonup}\ {\isacharprime}b{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}a\ set\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymrightharpoonup}\ {\isacharprime}b}\\
-\isa{dom} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymrightharpoonup}\ {\isacharprime}b{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}a\ set}\\
-\isa{ran} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymrightharpoonup}\ {\isacharprime}b{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}b\ set}\\
-\isa{op\ {\isasymsubseteq}\isactrlsub m} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymrightharpoonup}\ {\isacharprime}b{\isacharparenright}\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymrightharpoonup}\ {\isacharprime}b{\isacharparenright}\ {\isasymRightarrow}\ bool}\\
-\isa{map{\isacharunderscore}of} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}b{\isacharparenright}\ list\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymrightharpoonup}\ {\isacharprime}b}\\
-\isa{map{\isacharunderscore}upds} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymrightharpoonup}\ {\isacharprime}b{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}b\ list\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymrightharpoonup}\ {\isacharprime}b}\\
+\isa{Map{\isachardot}empty} & \isa{{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b\ option}\\
+\isa{op\ {\isacharplus}{\isacharplus}} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b\ option{\isacharparenright}\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b\ option{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b\ option}\\
+\isa{op\ {\isasymcirc}\isactrlsub m} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b\ option{\isacharparenright}\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}c\ {\isasymRightarrow}\ {\isacharprime}a\ option{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}c\ {\isasymRightarrow}\ {\isacharprime}b\ option}\\
+\isa{op\ {\isacharbar}{\isacharbackquote}} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b\ option{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}a\ set\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b\ option}\\
+\isa{dom} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b\ option{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}a\ set}\\
+\isa{ran} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b\ option{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}b\ set}\\
+\isa{op\ {\isasymsubseteq}\isactrlsub m} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b\ option{\isacharparenright}\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b\ option{\isacharparenright}\ {\isasymRightarrow}\ bool}\\
+\isa{map{\isacharunderscore}of} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymtimes}\ {\isacharprime}b{\isacharparenright}\ list\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b\ option}\\
+\isa{map{\isacharunderscore}upds} & \isa{{\isacharparenleft}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b\ option{\isacharparenright}\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}b\ list\ {\isasymRightarrow}\ {\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}b\ option}\\
 \end{supertabular}
 
 \subsubsection*{Syntax}
--- a/src/FOL/IsaMakefile	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/FOL/IsaMakefile	Sat Oct 30 12:25:18 2010 -0700
@@ -45,14 +45,13 @@
 FOL-ex: FOL $(LOG)/FOL-ex.gz
 
 $(LOG)/FOL-ex.gz: $(OUT)/FOL ex/First_Order_Logic.thy ex/If.thy		\
-  ex/Iff_Oracle.thy ex/Nat.thy ex/Nat_Class.thy ex/Natural_Numbers.thy	\
+  ex/Nat.thy ex/Nat_Class.thy ex/Natural_Numbers.thy			\
   ex/Locale_Test/Locale_Test.thy ex/Locale_Test/Locale_Test1.thy	\
   ex/Locale_Test/Locale_Test2.thy ex/Locale_Test/Locale_Test3.thy	\
-  ex/Miniscope.thy ex/Prolog.thy ex/ROOT.ML				\
-  ex/Classical.thy ex/document/root.tex ex/Foundation.thy		\
-  ex/Intuitionistic.thy ex/Intro.thy ex/Propositional_Int.thy		\
-  ex/Propositional_Cla.thy ex/Quantifiers_Int.thy			\
-  ex/Quantifiers_Cla.thy
+  ex/Miniscope.thy ex/Prolog.thy ex/ROOT.ML ex/Classical.thy		\
+  ex/document/root.tex ex/Foundation.thy ex/Intuitionistic.thy		\
+  ex/Intro.thy ex/Propositional_Int.thy ex/Propositional_Cla.thy	\
+  ex/Quantifiers_Int.thy ex/Quantifiers_Cla.thy
 	@$(ISABELLE_TOOL) usedir $(OUT)/FOL ex
 
 
--- a/src/FOL/ex/Iff_Oracle.thy	Fri Oct 29 17:15:28 2010 -0700
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,76 +0,0 @@
-(*  Title:      FOL/ex/Iff_Oracle.thy
-    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
-    Copyright   1996  University of Cambridge
-*)
-
-header {* Example of Declaring an Oracle *}
-
-theory Iff_Oracle
-imports FOL
-begin
-
-subsection {* Oracle declaration *}
-
-text {*
-  This oracle makes tautologies of the form @{text "P <-> P <-> P <-> P"}.
-  The length is specified by an integer, which is checked to be even
-  and positive.
-*}
-
-oracle iff_oracle = {*
-  let
-    fun mk_iff 1 = Var (("P", 0), @{typ o})
-      | mk_iff n = FOLogic.iff $ Var (("P", 0), @{typ o}) $ mk_iff (n - 1);
-  in
-    fn (thy, n) =>
-      if n > 0 andalso n mod 2 = 0
-      then Thm.cterm_of thy (FOLogic.mk_Trueprop (mk_iff n))
-      else raise Fail ("iff_oracle: " ^ string_of_int n)
-  end
-*}
-
-
-subsection {* Oracle as low-level rule *}
-
-ML {* iff_oracle (@{theory}, 2) *}
-ML {* iff_oracle (@{theory}, 10) *}
-ML {* Thm.proof_body_of (iff_oracle (@{theory}, 10)) *}
-
-text {* These oracle calls had better fail. *}
-
-ML {*
-  (iff_oracle (@{theory}, 5); error "?")
-    handle Fail _ => warning "Oracle failed, as expected"
-*}
-
-ML {*
-  (iff_oracle (@{theory}, 1); error "?")
-    handle Fail _ => warning "Oracle failed, as expected"
-*}
-
-
-subsection {* Oracle as proof method *}
-
-method_setup iff = {*
-  Scan.lift Parse.nat >> (fn n => fn ctxt =>
-    SIMPLE_METHOD
-      (HEADGOAL (Tactic.rtac (iff_oracle (ProofContext.theory_of ctxt, n)))
-        handle Fail _ => no_tac))
-*} "iff oracle"
-
-
-lemma "A <-> A"
-  by (iff 2)
-
-lemma "A <-> A <-> A <-> A <-> A <-> A <-> A <-> A <-> A <-> A"
-  by (iff 10)
-
-lemma "A <-> A <-> A <-> A <-> A"
-  apply (iff 5)?
-  oops
-
-lemma A
-  apply (iff 1)?
-  oops
-
-end
--- a/src/FOL/ex/ROOT.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/FOL/ex/ROOT.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -18,8 +18,7 @@
   "Propositional_Cla",
   "Quantifiers_Cla",
   "Miniscope",
-  "If",
-  "Iff_Oracle"
+  "If"
 ];
 
 (*regression test for locales -- sets several global flags!*)
--- a/src/HOL/Algebra/AbelCoset.thy	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Algebra/AbelCoset.thy	Sat Oct 30 12:25:18 2010 -0700
@@ -13,7 +13,7 @@
 text {* Hiding @{text "<+>"} from @{theory Sum_Type} until I come
   up with better syntax here *}
 
-no_notation Plus (infixr "<+>" 65)
+no_notation Sum_Type.Plus (infixr "<+>" 65)
 
 definition
   a_r_coset    :: "[_, 'a set, 'a] \<Rightarrow> 'a set"    (infixl "+>\<index>" 60)
--- a/src/HOL/Boogie/Tools/boogie_loader.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Boogie/Tools/boogie_loader.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -191,7 +191,7 @@
             | NONE => NONE)
         | is_unique _ = NONE
       fun mk_unique_axiom T ts =
-        Const (@{const_name distinct}, HOLogic.listT T --> @{typ bool}) $
+        Const (@{const_name SMT.distinct}, HOLogic.listT T --> @{typ bool}) $
           HOLogic.mk_list T ts
     in
       map_filter is_unique fns
@@ -383,7 +383,7 @@
   fun mk_list T = HOLogic.mk_list T
 
   fun mk_distinct T ts =
-    Const (@{const_name distinct}, HOLogic.listT T --> @{typ bool}) $
+    Const (@{const_name SMT.distinct}, HOLogic.listT T --> @{typ bool}) $
       mk_list T ts
 
   fun quant name f = scan_line name (num -- num -- num) >> pair f
--- a/src/HOL/Complete_Partial_Order.thy	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Complete_Partial_Order.thy	Sat Oct 30 12:25:18 2010 -0700
@@ -258,6 +258,6 @@
 
 end
 
-
+hide_const (open) lub iterates fixp admissible
 
 end
--- a/src/HOL/HOL.thy	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/HOL.thy	Sat Oct 30 12:25:18 2010 -0700
@@ -1978,11 +1978,9 @@
 *} "solve goal by normalization"
 
 
-(*
 subsection {* Try *}
 
 setup {* Try.setup *}
-*)
 
 subsection {* Counterexample Search Units *}
 
--- a/src/HOL/Imperative_HOL/Heap_Monad.thy	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Imperative_HOL/Heap_Monad.thy	Sat Oct 30 12:25:18 2010 -0700
@@ -16,6 +16,10 @@
   and transform the heap, or fail *}
 datatype 'a Heap = Heap "heap \<Rightarrow> ('a \<times> heap) option"
 
+lemma [code, code del]:
+  "(Code_Evaluation.term_of :: 'a::typerep Heap \<Rightarrow> Code_Evaluation.term) = Code_Evaluation.term_of"
+  ..
+
 primrec execute :: "'a Heap \<Rightarrow> heap \<Rightarrow> ('a \<times> heap) option" where
   [code del]: "execute (Heap f) = f"
 
@@ -398,6 +402,82 @@
   ultimately show ?case by (simp, simp only: execute_bind(1), simp add: execute_simps)
 qed
 
+
+subsection {* Partial function definition setup *}
+
+definition Heap_ord :: "'a Heap \<Rightarrow> 'a Heap \<Rightarrow> bool" where
+  "Heap_ord = img_ord execute (fun_ord option_ord)"
+
+definition Heap_lub :: "('a Heap \<Rightarrow> bool) \<Rightarrow> 'a Heap" where
+  "Heap_lub = img_lub execute Heap (fun_lub (flat_lub None))"
+
+interpretation heap!: partial_function_definitions Heap_ord Heap_lub
+proof -
+  have "partial_function_definitions (fun_ord option_ord) (fun_lub (flat_lub None))"
+    by (rule partial_function_lift) (rule flat_interpretation)
+  then have "partial_function_definitions (img_ord execute (fun_ord option_ord))
+      (img_lub execute Heap (fun_lub (flat_lub None)))"
+    by (rule partial_function_image) (auto intro: Heap_eqI)
+  then show "partial_function_definitions Heap_ord Heap_lub"
+    by (simp only: Heap_ord_def Heap_lub_def)
+qed
+
+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
+
+
 subsection {* Code generator setup *}
 
 subsubsection {* Logical intermediate layer *}
@@ -593,76 +673,6 @@
 
 *}
 
-
-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/IsaMakefile	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/IsaMakefile	Sat Oct 30 12:25:18 2010 -0700
@@ -337,6 +337,7 @@
   Tools/Sledgehammer/sledgehammer_atp_translate.ML \
   Tools/Sledgehammer/sledgehammer_util.ML \
   Tools/SMT/smtlib_interface.ML \
+  Tools/SMT/smt_builtin.ML \
   Tools/SMT/smt_monomorph.ML \
   Tools/SMT/smt_normalize.ML \
   Tools/SMT/smt_setup_solvers.ML \
@@ -1011,25 +1012,24 @@
   Number_Theory/Primes.thy ex/Abstract_NAT.thy ex/Antiquote.thy		\
   ex/Arith_Examples.thy ex/Arithmetic_Series_Complex.thy ex/BT.thy	\
   ex/BinEx.thy ex/Binary.thy ex/CTL.thy ex/Chinese.thy			\
-  ex/Classical.thy ex/CodegenSML_Test.thy ex/Coherent.thy		\
-  ex/Dedekind_Real.thy ex/Efficient_Nat_examples.thy			\
+  ex/Classical.thy ex/CodegenSML_Test.thy ex/Coercion_Examples.thy	\
+  ex/Coherent.thy ex/Dedekind_Real.thy ex/Efficient_Nat_examples.thy	\
   ex/Eval_Examples.thy ex/Fundefs.thy ex/Gauge_Integration.thy		\
   ex/Groebner_Examples.thy ex/Guess.thy ex/HarmonicSeries.thy		\
   ex/Hebrew.thy ex/Hex_Bin_Examples.thy ex/Higher_Order_Logic.thy	\
-  ex/Induction_Schema.thy ex/InductiveInvariant.thy			\
+  ex/Iff_Oracle.thy ex/Induction_Schema.thy ex/InductiveInvariant.thy	\
   ex/InductiveInvariant_examples.thy ex/Intuitionistic.thy		\
   ex/Lagrange.thy ex/LocaleTest2.thy ex/MT.thy ex/MergeSort.thy		\
   ex/Meson_Test.thy ex/MonoidGroup.thy ex/Multiquote.thy ex/NatSum.thy	\
-  ex/Normalization_by_Evaluation.thy					\
-  ex/Numeral.thy ex/PER.thy ex/PresburgerEx.thy ex/Primrec.thy		\
-  ex/Quickcheck_Examples.thy ex/Quickcheck_Lattice_Examples.thy		\
-  ex/ROOT.ML ex/Recdefs.thy ex/Records.thy ex/ReflectionEx.thy		\
-  ex/Refute_Examples.thy ex/SAT_Examples.thy ex/SVC_Oracle.thy		\
-  ex/Serbian.thy ex/Sqrt.thy ex/Sqrt_Script.thy ex/Sudoku.thy		\
-  ex/Tarski.thy ex/Termination.thy ex/Transfer_Ex.thy ex/Tree23.thy	\
-  ex/Unification.thy ex/While_Combinator_Example.thy			\
-  ex/document/root.bib ex/document/root.tex ex/set.thy ex/svc_funcs.ML	\
-  ex/svc_test.thy
+  ex/Normalization_by_Evaluation.thy ex/Numeral.thy ex/PER.thy		\
+  ex/PresburgerEx.thy ex/Primrec.thy ex/Quickcheck_Examples.thy		\
+  ex/Quickcheck_Lattice_Examples.thy ex/ROOT.ML ex/Recdefs.thy		\
+  ex/Records.thy ex/ReflectionEx.thy ex/Refute_Examples.thy		\
+  ex/SAT_Examples.thy ex/SVC_Oracle.thy ex/Serbian.thy ex/Sqrt.thy	\
+  ex/Sqrt_Script.thy ex/Sudoku.thy ex/Tarski.thy ex/Termination.thy	\
+  ex/Transfer_Ex.thy ex/Tree23.thy ex/Unification.thy			\
+  ex/While_Combinator_Example.thy ex/document/root.bib			\
+  ex/document/root.tex ex/set.thy ex/svc_funcs.ML ex/svc_test.thy
 	@$(ISABELLE_TOOL) usedir $(OUT)/HOL ex
 
 
--- a/src/HOL/Library/Multiset.thy	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Library/Multiset.thy	Sat Oct 30 12:25:18 2010 -0700
@@ -1279,7 +1279,7 @@
 subsubsection {* Monotonicity of multiset union *}
 
 lemma mult1_union:
-  "(B, D) \<in> mult1 r ==> trans r ==> (C + B, C + D) \<in> mult1 r"
+  "(B, D) \<in> mult1 r ==> (C + B, C + D) \<in> mult1 r"
 apply (unfold mult1_def)
 apply auto
 apply (rule_tac x = a in exI)
@@ -1290,8 +1290,8 @@
 lemma union_less_mono2: "B \<subset># D ==> C + B \<subset># C + (D::'a::order multiset)"
 apply (unfold less_multiset_def mult_def)
 apply (erule trancl_induct)
- apply (blast intro: mult1_union transI order_less_trans r_into_trancl)
-apply (blast intro: mult1_union transI order_less_trans r_into_trancl trancl_trans)
+ apply (blast intro: mult1_union)
+apply (blast intro: mult1_union trancl_trans)
 done
 
 lemma union_less_mono1: "B \<subset># D ==> B + C \<subset># D + (C::'a::order multiset)"
--- a/src/HOL/List.thy	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/List.thy	Sat Oct 30 12:25:18 2010 -0700
@@ -1576,6 +1576,10 @@
   "map f (butlast xs) = butlast (map f xs)"
   by (induct xs) simp_all
 
+lemma snoc_eq_iff_butlast:
+  "xs @ [x] = ys \<longleftrightarrow> (ys \<noteq> [] & butlast ys = xs & last ys = x)"
+by (metis append_butlast_last_id append_is_Nil_conv butlast_snoc last_snoc not_Cons_self)
+
 
 subsubsection {* @{text take} and @{text drop} *}
 
@@ -2288,6 +2292,10 @@
   "xs = ys \<longleftrightarrow> list_all2 (op =) xs ys"
 by (induct xs ys rule: list_induct2') auto
 
+lemma list_eq_iff_zip_eq:
+  "xs = ys \<longleftrightarrow> length xs = length ys \<and> (\<forall>(x,y) \<in> set (zip xs ys). x = y)"
+by(auto simp add: set_zip list_all2_eq list_all2_conv_all_nth cong: conj_cong)
+
 
 subsubsection {* @{text foldl} and @{text foldr} *}
 
@@ -4457,7 +4465,7 @@
   by blast
 
 
-subsection {* Lexicographic combination of measure functions *}
+subsubsection {* Lexicographic combination of measure functions *}
 
 text {* These are useful for termination proofs *}
 
@@ -4482,7 +4490,133 @@
 by auto
 
 
-subsubsection {* Lifting a Relation on List Elements to the Lists *}
+subsubsection {* Lifting Relations to Lists: one element *}
+
+definition listrel1 :: "('a \<times> 'a) set \<Rightarrow> ('a list \<times> 'a list) set" where
+"listrel1 r = {(xs,ys).
+   \<exists>us z z' vs. xs = us @ z # vs \<and> (z,z') \<in> r \<and> ys = us @ z' # vs}"
+
+lemma listrel1I:
+  "\<lbrakk> (x, y) \<in> r;  xs = us @ x # vs;  ys = us @ y # vs \<rbrakk> \<Longrightarrow>
+  (xs, ys) \<in> listrel1 r"
+unfolding listrel1_def by auto
+
+lemma listrel1E:
+  "\<lbrakk> (xs, ys) \<in> listrel1 r;
+     !!x y us vs. \<lbrakk> (x, y) \<in> r;  xs = us @ x # vs;  ys = us @ y # vs \<rbrakk> \<Longrightarrow> P
+   \<rbrakk> \<Longrightarrow> P"
+unfolding listrel1_def by auto
+
+lemma not_Nil_listrel1 [iff]: "([], xs) \<notin> listrel1 r"
+unfolding listrel1_def by blast
+
+lemma not_listrel1_Nil [iff]: "(xs, []) \<notin> listrel1 r"
+unfolding listrel1_def by blast
+
+lemma Cons_listrel1_Cons [iff]:
+  "(x # xs, y # ys) \<in> listrel1 r \<longleftrightarrow>
+   (x,y) \<in> r \<and> xs = ys \<or> x = y \<and> (xs, ys) \<in> listrel1 r"
+by (simp add: listrel1_def Cons_eq_append_conv) (blast)
+
+lemma listrel1I1: "(x,y) \<in> r \<Longrightarrow> (x # xs, y # xs) \<in> listrel1 r"
+by (metis Cons_listrel1_Cons)
+
+lemma listrel1I2: "(xs, ys) \<in> listrel1 r \<Longrightarrow> (x # xs, x # ys) \<in> listrel1 r"
+by (metis Cons_listrel1_Cons)
+
+lemma append_listrel1I:
+  "(xs, ys) \<in> listrel1 r \<and> us = vs \<or> xs = ys \<and> (us, vs) \<in> listrel1 r
+    \<Longrightarrow> (xs @ us, ys @ vs) \<in> listrel1 r"
+unfolding listrel1_def
+by auto (blast intro: append_eq_appendI)+
+
+lemma Cons_listrel1E1[elim!]:
+  assumes "(x # xs, ys) \<in> listrel1 r"
+    and "\<And>y. ys = y # xs \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> R"
+    and "\<And>zs. ys = x # zs \<Longrightarrow> (xs, zs) \<in> listrel1 r \<Longrightarrow> R"
+  shows R
+using assms by (cases ys) blast+
+
+lemma Cons_listrel1E2[elim!]:
+  assumes "(xs, y # ys) \<in> listrel1 r"
+    and "\<And>x. xs = x # ys \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> R"
+    and "\<And>zs. xs = y # zs \<Longrightarrow> (zs, ys) \<in> listrel1 r \<Longrightarrow> R"
+  shows R
+using assms by (cases xs) blast+
+
+lemma snoc_listrel1_snoc_iff:
+  "(xs @ [x], ys @ [y]) \<in> listrel1 r
+    \<longleftrightarrow> (xs, ys) \<in> listrel1 r \<and> x = y \<or> xs = ys \<and> (x,y) \<in> r" (is "?L \<longleftrightarrow> ?R")
+proof
+  assume ?L thus ?R
+    by (fastsimp simp: listrel1_def snoc_eq_iff_butlast butlast_append)
+next
+  assume ?R then show ?L unfolding listrel1_def by force
+qed
+
+lemma listrel1_eq_len: "(xs,ys) \<in> listrel1 r \<Longrightarrow> length xs = length ys"
+unfolding listrel1_def by auto
+
+lemma listrel1_mono:
+  "r \<subseteq> s \<Longrightarrow> listrel1 r \<subseteq> listrel1 s"
+unfolding listrel1_def by blast
+
+
+lemma listrel1_converse: "listrel1 (r^-1) = (listrel1 r)^-1"
+unfolding listrel1_def by blast
+
+lemma in_listrel1_converse:
+  "(x,y) : listrel1 (r^-1) \<longleftrightarrow> (x,y) : (listrel1 r)^-1"
+unfolding listrel1_def by blast
+
+lemma listrel1_iff_update:
+  "(xs,ys) \<in> (listrel1 r)
+   \<longleftrightarrow> (\<exists>y n. (xs ! n, y) \<in> r \<and> n < length xs \<and> ys = xs[n:=y])" (is "?L \<longleftrightarrow> ?R")
+proof
+  assume "?L"
+  then obtain x y u v where "xs = u @ x # v"  "ys = u @ y # v"  "(x,y) \<in> r"
+    unfolding listrel1_def by auto
+  then have "ys = xs[length u := y]" and "length u < length xs"
+    and "(xs ! length u, y) \<in> r" by auto
+  then show "?R" by auto
+next
+  assume "?R"
+  then obtain x y n where "(xs!n, y) \<in> r" "n < size xs" "ys = xs[n:=y]" "x = xs!n"
+    by auto
+  then obtain u v where "xs = u @ x # v" and "ys = u @ y # v" and "(x, y) \<in> r"
+    by (auto intro: upd_conv_take_nth_drop id_take_nth_drop)
+  then show "?L" by (auto simp: listrel1_def)
+qed
+
+
+text{* Accessible part of @{term listrel1} relations: *}
+
+lemma Cons_acc_listrel1I [intro!]:
+  "x \<in> acc r \<Longrightarrow> xs \<in> acc (listrel1 r) \<Longrightarrow> (x # xs) \<in> acc (listrel1 r)"
+apply (induct arbitrary: xs set: acc)
+apply (erule thin_rl)
+apply (erule acc_induct)
+apply (rule accI)
+apply (blast)
+done
+
+lemma lists_accD: "xs \<in> lists (acc r) \<Longrightarrow> xs \<in> acc (listrel1 r)"
+apply (induct set: lists)
+ apply (rule accI)
+ apply simp
+apply (rule accI)
+apply (fast dest: acc_downward)
+done
+
+lemma lists_accI: "xs \<in> acc (listrel1 r) \<Longrightarrow> xs \<in> lists (acc r)"
+apply (induct set: acc)
+apply clarify
+apply (rule accI)
+apply (fastsimp dest!: in_set_conv_decomp[THEN iffD1] simp: listrel1_def)
+done
+
+
+subsubsection {* Lifting Relations to Lists: all elements *}
 
 inductive_set
   listrel :: "('a * 'a)set => ('a list * 'a list)set"
@@ -4497,6 +4631,24 @@
 inductive_cases listrel_Cons2 [elim!]: "(xs,y#ys) \<in> listrel r"
 
 
+lemma listrel_eq_len:  "(xs, ys) \<in> listrel r \<Longrightarrow> length xs = length ys"
+by(induct rule: listrel.induct) auto
+
+lemma listrel_iff_zip: "(xs,ys) : listrel r \<longleftrightarrow>
+  length xs = length ys & (\<forall>(x,y) \<in> set(zip xs ys). (x,y) \<in> r)" (is "?L \<longleftrightarrow> ?R")
+proof
+  assume ?L thus ?R by induct (auto intro: listrel_eq_len)
+next
+  assume ?R thus ?L
+    apply (clarify)
+    by (induct rule: list_induct2) (auto intro: listrel.intros)
+qed
+
+lemma listrel_iff_nth: "(xs,ys) : listrel r \<longleftrightarrow>
+  length xs = length ys & (\<forall>n < length xs. (xs!n, ys!n) \<in> r)" (is "?L \<longleftrightarrow> ?R")
+by (auto simp add: all_set_conv_all_nth listrel_iff_zip)
+
+
 lemma listrel_mono: "r \<subseteq> s \<Longrightarrow> listrel r \<subseteq> listrel s"
 apply clarify  
 apply (erule listrel.induct)
@@ -4532,6 +4684,16 @@
 theorem equiv_listrel: "equiv A r \<Longrightarrow> equiv (lists A) (listrel r)"
 by (simp add: equiv_def listrel_refl_on listrel_sym listrel_trans) 
 
+lemma listrel_rtrancl_refl[iff]: "(xs,xs) : listrel(r^*)"
+using listrel_refl_on[of UNIV, OF refl_rtrancl]
+by(auto simp: refl_on_def)
+
+lemma listrel_rtrancl_trans:
+  "\<lbrakk> (xs,ys) : listrel(r^*);  (ys,zs) : listrel(r^*) \<rbrakk>
+  \<Longrightarrow> (xs,zs) : listrel(r^*)"
+by (metis listrel_trans trans_def trans_rtrancl)
+
+
 lemma listrel_Nil [simp]: "listrel r `` {[]} = {[]}"
 by (blast intro: listrel.intros)
 
@@ -4539,6 +4701,74 @@
      "listrel r `` {x#xs} = set_Cons (r``{x}) (listrel r `` {xs})"
 by (auto simp add: set_Cons_def intro: listrel.intros)
 
+text {* Relating @{term listrel1}, @{term listrel} and closures: *}
+
+lemma listrel1_rtrancl_subset_rtrancl_listrel1:
+  "listrel1 (r^*) \<subseteq> (listrel1 r)^*"
+proof (rule subrelI)
+  fix xs ys assume 1: "(xs,ys) \<in> listrel1 (r^*)"
+  { fix x y us vs
+    have "(x,y) : r^* \<Longrightarrow> (us @ x # vs, us @ y # vs) : (listrel1 r)^*"
+    proof(induct rule: rtrancl.induct)
+      case rtrancl_refl show ?case by simp
+    next
+      case rtrancl_into_rtrancl thus ?case
+        by (metis listrel1I rtrancl.rtrancl_into_rtrancl)
+    qed }
+  thus "(xs,ys) \<in> (listrel1 r)^*" using 1 by(blast elim: listrel1E)
+qed
+
+lemma rtrancl_listrel1_eq_len: "(x,y) \<in> (listrel1 r)^* \<Longrightarrow> length x = length y"
+by (induct rule: rtrancl.induct) (auto intro: listrel1_eq_len)
+
+lemma rtrancl_listrel1_ConsI1:
+  "(xs,ys) : (listrel1 r)^* \<Longrightarrow> (x#xs,x#ys) : (listrel1 r)^*"
+apply(induct rule: rtrancl.induct)
+ apply simp
+by (metis listrel1I2 rtrancl.rtrancl_into_rtrancl)
+
+lemma rtrancl_listrel1_ConsI2:
+  "(x,y) \<in> r^* \<Longrightarrow> (xs, ys) \<in> (listrel1 r)^*
+  \<Longrightarrow> (x # xs, y # ys) \<in> (listrel1 r)^*"
+  by (blast intro: rtrancl_trans rtrancl_listrel1_ConsI1 
+    subsetD[OF listrel1_rtrancl_subset_rtrancl_listrel1 listrel1I1])
+
+lemma listrel1_subset_listrel:
+  "r \<subseteq> r' \<Longrightarrow> refl r' \<Longrightarrow> listrel1 r \<subseteq> listrel(r')"
+by(auto elim!: listrel1E simp add: listrel_iff_zip set_zip refl_on_def)
+
+lemma listrel_reflcl_if_listrel1:
+  "(xs,ys) : listrel1 r \<Longrightarrow> (xs,ys) : listrel(r^*)"
+by(erule listrel1E)(auto simp add: listrel_iff_zip set_zip)
+
+lemma listrel_rtrancl_eq_rtrancl_listrel1: "listrel (r^*) = (listrel1 r)^*"
+proof
+  { fix x y assume "(x,y) \<in> listrel (r^*)"
+    then have "(x,y) \<in> (listrel1 r)^*"
+    by induct (auto intro: rtrancl_listrel1_ConsI2) }
+  then show "listrel (r^*) \<subseteq> (listrel1 r)^*"
+    by (rule subrelI)
+next
+  show "listrel (r^*) \<supseteq> (listrel1 r)^*"
+  proof(rule subrelI)
+    fix xs ys assume "(xs,ys) \<in> (listrel1 r)^*"
+    then show "(xs,ys) \<in> listrel (r^*)"
+    proof induct
+      case base show ?case by(auto simp add: listrel_iff_zip set_zip)
+    next
+      case (step ys zs)
+      thus ?case  by (metis listrel_reflcl_if_listrel1 listrel_rtrancl_trans)
+    qed
+  qed
+qed
+
+lemma rtrancl_listrel1_if_listrel:
+  "(xs,ys) : listrel r \<Longrightarrow> (xs,ys) : (listrel1 r)^*"
+by(metis listrel_rtrancl_eq_rtrancl_listrel1 subsetD[OF listrel_mono] r_into_rtrancl subsetI)
+
+lemma listrel_subset_rtrancl_listrel1: "listrel r \<subseteq> (listrel1 r)^*"
+by(fast intro:rtrancl_listrel1_if_listrel)
+
 
 subsection {* Size function *}
 
--- a/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer_filter.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer_filter.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -95,8 +95,7 @@
   else
     ()
 
-val default_max_relevant = 300
-val prover_name = ATP_Systems.eN (* arbitrary ATP *)
+val default_prover = ATP_Systems.eN (* arbitrary ATP *)
 
 fun with_index (i, s) = s ^ "@" ^ string_of_int i
 
@@ -107,11 +106,16 @@
        SOME proofs =>
        let
          val {context = ctxt, facts, goal} = Proof.goal pre
+         val thy = ProofContext.theory_of ctxt
+         val prover = AList.lookup (op =) args "prover"
+                      |> the_default default_prover
+         val default_max_relevant =
+           Sledgehammer.default_max_relevant_for_prover thy prover
          val irrelevant_consts =
-           Sledgehammer.irrelevant_consts_for_prover prover_name
+           Sledgehammer.irrelevant_consts_for_prover prover
          val relevance_fudge =
            extract_relevance_fudge args
-               (Sledgehammer.relevance_fudge_for_prover prover_name)
+               (Sledgehammer.relevance_fudge_for_prover prover)
          val relevance_override = {add = [], del = [], only = false}
          val {relevance_thresholds, full_types, max_relevant, ...} =
            Sledgehammer_Isar.default_params ctxt args
--- a/src/HOL/Mutabelle/mutabelle.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Mutabelle/mutabelle.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -497,15 +497,19 @@
  (map_types (inst_type insts) (freeze t));
 
 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));
-  Output.urgent_message "executable"; true) handle ERROR s =>
-    (Output.urgent_message ("not executable: " ^ s); false);
+  ((Quickcheck.test_term
+      (Context.proof_map (Quickcheck.map_test_params (apfst (K (1, 1)))) (ProofContext.init_global thy))
+      (SOME (!testgen_name)) (preprocess thy insts (prop_of th));
+    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 
-     (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))
+     (Output.urgent_message ("qc_recursive: " ^ string_of_int (length xs));
+     ((x, pretty (the_default [] (fst (Quickcheck.test_term
+       (Context.proof_map (Quickcheck.map_test_params (apfst (K (sz, qciter))))
+         (ProofContext.init_global usedthy))
+       (SOME (!testgen_name)) (preprocess usedthy insts x))))) :: acc))
           handle ERROR msg => (Output.urgent_message msg; qc_recursive usedthy xs insts sz qciter acc);
 
 
--- a/src/HOL/Mutabelle/mutabelle_extra.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Mutabelle/mutabelle_extra.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -53,8 +53,6 @@
 
 (* quickcheck options *)
 (*val quickcheck_generator = "SML"*)
-val iterations = 10
-val size = 5
 
 exception RANDOM;
 
@@ -96,8 +94,8 @@
 fun invoke_quickcheck quickcheck_generator thy t =
   TimeLimit.timeLimit (Time.fromSeconds (!Auto_Tools.time_limit))
       (fn _ =>
-          case Quickcheck.gen_test_term (ProofContext.init_global thy) true (SOME quickcheck_generator)
-                                    size iterations (preprocess thy [] t) of
+          case Quickcheck.test_term (ProofContext.init_global thy)
+              (SOME quickcheck_generator) (preprocess thy [] t) of
             (NONE, (time_report, report)) => (NoCex, (time_report, report))
           | (SOME _, (time_report, report)) => (GenuineCex, (time_report, report))) ()
   handle TimeLimit.TimeOut => (Timeout, ([("timelimit", !Auto_Tools.time_limit)], NONE))
@@ -249,8 +247,12 @@
     exists (member (op =) forbidden_mutant_constnames) consts
   end
 
-fun is_executable_term thy t = can (TimeLimit.timeLimit (Time.fromMilliseconds 2000) (Quickcheck.test_term
- (ProofContext.init_global thy) false (SOME "SML") 1 0)) (preprocess thy [] t)
+fun is_executable_term thy t =
+  can (TimeLimit.timeLimit (Time.fromMilliseconds 2000)
+    (Quickcheck.test_term
+      (Context.proof_map (Quickcheck.map_test_params (apfst (K (1, 0)))) (ProofContext.init_global thy))
+      (SOME "SML"))) (preprocess thy [] t)
+
 fun is_executable_thm thy th = is_executable_term thy (prop_of th)
 
 val freezeT =
@@ -422,6 +424,7 @@
 fun mutate_theorems_and_write_report thy mtds thms file_name =
   let
     val _ = Output.urgent_message "Starting Mutabelle..."
+    val Quickcheck.Test_Params test_params = Quickcheck.test_params_of (ProofContext.init_global thy)
     val path = Path.explode file_name
     (* for normal report: *)
     (*val (gen_create_entry, gen_string_for_entry) = (create_entry, string_for_entry)*)
@@ -437,8 +440,8 @@
       "; num_mutations: " ^ string_of_int num_mutations ^ "\n" ^
     "QC options = " ^
       (*"quickcheck_generator: " ^ quickcheck_generator ^ ";*)
-      "size: " ^ string_of_int size ^
-      "; iterations: " ^ string_of_int iterations ^ "\n");
+      "size: " ^ string_of_int (#size test_params) ^
+      "; iterations: " ^ string_of_int (#iterations test_params) ^ "\n");
     map (File.append path o gen_string_for_entry o mutate_theorem gen_create_entry thy mtds) thms;
     ()
   end
--- a/src/HOL/Partial_Function.thy	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Partial_Function.thy	Sat Oct 30 12:25:18 2010 -0700
@@ -26,6 +26,11 @@
 lemma call_mono[partial_function_mono]: "monotone (fun_ord ord) ord (\<lambda>f. f t)"
 by (rule monotoneI) (auto simp: fun_ord_def)
 
+lemma let_mono[partial_function_mono]:
+  "(\<And>x. monotone orda ordb (\<lambda>f. b f x))
+  \<Longrightarrow> monotone orda ordb (\<lambda>f. Let t (b f))"
+by (simp add: Let_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)"
@@ -242,6 +247,7 @@
   show "option_ord (Option.bind (B f) (\<lambda>y. C y f)) (Option.bind (B g) (\<lambda>y'. C y' g))" .
 qed
 
+hide_const (open) chain
 
 end
 
--- a/src/HOL/Predicate_Compile_Examples/Code_Prolog_Examples.thy	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Predicate_Compile_Examples/Code_Prolog_Examples.thy	Sat Oct 30 12:25:18 2010 -0700
@@ -105,7 +105,7 @@
 
 section {* Example symbolic derivation *}
 
-hide_const Plus Pow
+hide_const Pow
 
 datatype expr = Log expr | Mult expr expr | Div expr expr | x | Num int | Plus expr expr
   | Minus expr expr | Uminus expr | Pow expr int | Exp expr
--- a/src/HOL/Predicate_Compile_Examples/ROOT.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Predicate_Compile_Examples/ROOT.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -12,11 +12,12 @@
 if getenv "EXEC_SWIPL" = "" andalso getenv "EXEC_YAP" = "" then
   (warning "No prolog system found - skipping some example theories"; ())
 else
-  (use_thys [
-    "Code_Prolog_Examples",
-    "Context_Free_Grammar_Example",
-    "Hotel_Example_Prolog",
-    "Lambda_Example",
-    "List_Examples"];
-   Unsynchronized.setmp quick_and_dirty true use_thys ["Reg_Exp_Example"]);
-
+  if not (getenv "EXEC_SWIPL" = "") orelse (getenv "SWIPL_VERSION" = "5.10.1") then
+     (use_thys [
+        "Code_Prolog_Examples",
+        "Context_Free_Grammar_Example",
+        "Hotel_Example_Prolog",
+        "Lambda_Example",
+        "List_Examples"];
+      Unsynchronized.setmp quick_and_dirty true use_thys ["Reg_Exp_Example"])
+    else (warning "Unsupported SWI-Prolog version - skipping some example theories"; ());
--- a/src/HOL/SMT.thy	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/SMT.thy	Sat Oct 30 12:25:18 2010 -0700
@@ -9,6 +9,7 @@
 uses
   "Tools/Datatype/datatype_selectors.ML"
   ("Tools/SMT/smt_monomorph.ML")
+  ("Tools/SMT/smt_builtin.ML")
   ("Tools/SMT/smt_normalize.ML")
   ("Tools/SMT/smt_translate.ML")
   ("Tools/SMT/smt_solver.ML")
@@ -49,6 +50,20 @@
 
 
 
+subsection {* Distinctness *}
+
+text {*
+As an abbreviation for a quadratic number of inequalities, SMT solvers
+provide a built-in @{text distinct}.  To avoid confusion with the
+already defined (and more general) @{term List.distinct}, a separate
+constant is defined.
+*}
+
+definition distinct :: "'a list \<Rightarrow> bool"
+where "distinct xs = List.distinct xs"
+
+
+
 subsection {* Higher-order encoding *}
 
 text {*
@@ -112,6 +127,7 @@
 subsection {* Setup *}
 
 use "Tools/SMT/smt_monomorph.ML"
+use "Tools/SMT/smt_builtin.ML"
 use "Tools/SMT/smt_normalize.ML"
 use "Tools/SMT/smt_translate.ML"
 use "Tools/SMT/smt_solver.ML"
@@ -314,7 +330,7 @@
 
 hide_type (open) pattern
 hide_const Pattern term_eq
-hide_const (open) trigger pat nopat fun_app z3div z3mod
+hide_const (open) trigger pat nopat distinct fun_app z3div z3mod
 
 
 
--- a/src/HOL/SMT_Examples/SMT_Examples.thy	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/SMT_Examples/SMT_Examples.thy	Sat Oct 30 12:25:18 2010 -0700
@@ -293,11 +293,11 @@
 
 lemma "(0 \<le> y + -1 * x \<or> \<not> 0 \<le> x \<or> 0 \<le> (x::int)) = (\<not> False)" by smt
 
-lemma "distinct [x < (3::int), 3 \<le> x]" by smt
+lemma "SMT.distinct [x < (3::int), 3 \<le> x]" by smt
 
 lemma
   assumes "a > (0::int)"
-  shows "distinct [a, a * 2, a - a]"
+  shows "SMT.distinct [a, a * 2, a - a]"
   using assms by smt
 
 lemma "
@@ -430,7 +430,7 @@
    False \<or> P = (x - 1 = y) \<or> (\<not>P \<longrightarrow> False)"
   by smt
 
-lemma "distinct [a + (1::nat), a * 2 + 3, a - a]" by smt
+lemma "SMT.distinct [a + (1::nat), a * 2 + 3, a - a]" by smt
 
 lemma "int (nat \<bar>x::int\<bar>) = \<bar>x\<bar>" by smt
 
--- a/src/HOL/SMT_Examples/SMT_Tests.certs	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/SMT_Examples/SMT_Tests.certs	Sat Oct 30 12:25:18 2010 -0700
@@ -7,96 +7,6 @@
 #25 := [asserted]: #8
 [mp #25 #27]: false
 unsat
-ac5b1925e47f9c59b180fdccd321c68a4a01a710 16 0
-#2 := false
-#8 := (not false)
-#9 := (not #8)
-#34 := (iff #9 false)
-#1 := true
-#29 := (not true)
-#32 := (iff #29 false)
-#33 := [rewrite]: #32
-#30 := (iff #9 #29)
-#27 := (iff #8 true)
-#28 := [rewrite]: #27
-#31 := [monotonicity #28]: #30
-#35 := [trans #31 #33]: #34
-#26 := [asserted]: #9
-[mp #26 #35]: false
-unsat
-b13645ce0b47b432b924b38e0bebde6cbcb236f9 21 0
-#2 := false
-#1 := true
-#8 := (not true)
-#9 := (not #8)
-#10 := (not #9)
-#39 := (iff #10 false)
-#28 := (iff #8 false)
-#29 := [rewrite]: #28
-#37 := (iff #10 #8)
-#35 := (iff #9 true)
-#30 := (not false)
-#33 := (iff #30 true)
-#34 := [rewrite]: #33
-#31 := (iff #9 #30)
-#32 := [monotonicity #29]: #31
-#36 := [trans #32 #34]: #35
-#38 := [monotonicity #36]: #37
-#40 := [trans #38 #29]: #39
-#27 := [asserted]: #10
-[mp #27 #40]: false
-unsat
-8ec60d0baade8bddd39b6e00ec153f6bd8949184 16 0
-#2 := false
-#1 := true
-#8 := (and true true)
-#9 := (not #8)
-#34 := (iff #9 false)
-#29 := (not true)
-#32 := (iff #29 false)
-#33 := [rewrite]: #32
-#30 := (iff #9 #29)
-#27 := (iff #8 true)
-#28 := [rewrite]: #27
-#31 := [monotonicity #28]: #30
-#35 := [trans #31 #33]: #34
-#26 := [asserted]: #9
-[mp #26 #35]: false
-unsat
-b448617db1567a507fc596f5b21ade645ca6e0f9 16 0
-#2 := false
-#1 := true
-#8 := (or true false)
-#9 := (not #8)
-#34 := (iff #9 false)
-#29 := (not true)
-#32 := (iff #29 false)
-#33 := [rewrite]: #32
-#30 := (iff #9 #29)
-#27 := (iff #8 true)
-#28 := [rewrite]: #27
-#31 := [monotonicity #28]: #30
-#35 := [trans #31 #33]: #34
-#26 := [asserted]: #9
-[mp #26 #35]: false
-unsat
-bf545ecaa58049cbf454eac70ee3077f64564cca 16 0
-#2 := false
-#1 := true
-#8 := (implies false true)
-#9 := (not #8)
-#34 := (iff #9 false)
-#29 := (not true)
-#32 := (iff #29 false)
-#33 := [rewrite]: #32
-#30 := (iff #9 #29)
-#27 := (iff #8 true)
-#28 := [rewrite]: #27
-#31 := [monotonicity #28]: #30
-#35 := [trans #31 #33]: #34
-#26 := [asserted]: #9
-[mp #26 #35]: false
-unsat
 a38979487be0aafcf2d387180e92dfd3bd43e483 33 0
 #2 := false
 decl f1 :: S1
@@ -131,6 +41,23 @@
 #29 := [asserted]: #12
 [mp #29 #49]: false
 unsat
+ac5b1925e47f9c59b180fdccd321c68a4a01a710 16 0
+#2 := false
+#8 := (not false)
+#9 := (not #8)
+#34 := (iff #9 false)
+#1 := true
+#29 := (not true)
+#32 := (iff #29 false)
+#33 := [rewrite]: #32
+#30 := (iff #9 #29)
+#27 := (iff #8 true)
+#28 := [rewrite]: #27
+#31 := [monotonicity #28]: #30
+#35 := [trans #31 #33]: #34
+#26 := [asserted]: #9
+[mp #26 #35]: false
+unsat
 b2faf72eaa234b8ac1acedfef5768f99ca3da86b 41 0
 #2 := false
 decl f1 :: S1
@@ -173,6 +100,28 @@
 #30 := [asserted]: #13
 [mp #30 #57]: false
 unsat
+b13645ce0b47b432b924b38e0bebde6cbcb236f9 21 0
+#2 := false
+#1 := true
+#8 := (not true)
+#9 := (not #8)
+#10 := (not #9)
+#39 := (iff #10 false)
+#28 := (iff #8 false)
+#29 := [rewrite]: #28
+#37 := (iff #10 #8)
+#35 := (iff #9 true)
+#30 := (not false)
+#33 := (iff #30 true)
+#34 := [rewrite]: #33
+#31 := (iff #9 #30)
+#32 := [monotonicity #29]: #31
+#36 := [trans #32 #34]: #35
+#38 := [monotonicity #36]: #37
+#40 := [trans #38 #29]: #39
+#27 := [asserted]: #10
+[mp #27 #40]: false
+unsat
 fa3a2c67c2833da196c4009cb283e94a9f22829b 65 0
 #2 := false
 decl f1 :: S1
@@ -239,6 +188,23 @@
 #33 := [asserted]: #16
 [mp #33 #81]: false
 unsat
+8ec60d0baade8bddd39b6e00ec153f6bd8949184 16 0
+#2 := false
+#1 := true
+#8 := (and true true)
+#9 := (not #8)
+#34 := (iff #9 false)
+#29 := (not true)
+#32 := (iff #29 false)
+#33 := [rewrite]: #32
+#30 := (iff #9 #29)
+#27 := (iff #8 true)
+#28 := [rewrite]: #27
+#31 := [monotonicity #28]: #30
+#35 := [trans #31 #33]: #34
+#26 := [asserted]: #9
+[mp #26 #35]: false
+unsat
 85a9b542a9c6d64e778f2449f892d6d55e153689 29 0
 #2 := false
 decl f1 :: S1
@@ -269,6 +235,23 @@
 #28 := [asserted]: #11
 [mp #28 #45]: false
 unsat
+b448617db1567a507fc596f5b21ade645ca6e0f9 16 0
+#2 := false
+#1 := true
+#8 := (or true false)
+#9 := (not #8)
+#34 := (iff #9 false)
+#29 := (not true)
+#32 := (iff #29 false)
+#33 := [rewrite]: #32
+#30 := (iff #9 #29)
+#27 := (iff #8 true)
+#28 := [rewrite]: #27
+#31 := [monotonicity #28]: #30
+#35 := [trans #31 #33]: #34
+#26 := [asserted]: #9
+[mp #26 #35]: false
+unsat
 a21434fdffa6243cc7410eb0aff669b84c5be474 41 0
 #2 := false
 decl f1 :: S1
@@ -311,6 +294,23 @@
 #30 := [asserted]: #13
 [mp #30 #57]: false
 unsat
+bf545ecaa58049cbf454eac70ee3077f64564cca 16 0
+#2 := false
+#1 := true
+#8 := (implies false true)
+#9 := (not #8)
+#34 := (iff #9 false)
+#29 := (not true)
+#32 := (iff #29 false)
+#33 := [rewrite]: #32
+#30 := (iff #9 #29)
+#27 := (iff #8 true)
+#28 := [rewrite]: #27
+#31 := [monotonicity #28]: #30
+#35 := [trans #31 #33]: #34
+#26 := [asserted]: #9
+[mp #26 #35]: false
+unsat
 a6a664cc4c3ff45fb3211b72c0f520208b902c9b 47 0
 #2 := false
 decl f1 :: S1
@@ -359,6 +359,156 @@
 #32 := [asserted]: #15
 [mp #32 #62]: 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
 d99e41f56c960ba88eb8ca95c5a49bd2181a7626 47 0
 #2 := false
 decl f1 :: S1
@@ -407,6 +557,73 @@
 #32 := [asserted]: #15
 [mp #32 #62]: 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
 a8e5086b6e21c7bf961a22426cb7ae14847089d7 55 0
 #2 := false
 decl f3 :: S1
@@ -463,6 +680,74 @@
 #71 := [and-elim #70]: #33
 [mp #71 #77]: 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
 e2d0843914ae918fb9bd06a2d6949232327217fa 55 0
 #2 := false
 decl f1 :: S1
@@ -519,6 +804,74 @@
 #33 := [asserted]: #16
 [mp #33 #70]: 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
 00d9c16458cf2825e776c32505a513a412ede70f 55 0
 #2 := false
 decl f1 :: S1
@@ -575,6 +928,193 @@
 #33 := [asserted]: #16
 [mp #33 #70]: 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
 bca31cbd8bb6114896c32a11dfae7025f2fc07b1 49 0
 #2 := false
 decl f3 :: S1
@@ -680,6 +1220,40 @@
 #68 := [and-elim #66]: #34
 [mp #68 #74]: 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
 38d7d5f12a8c933e456bba4a04c7645f4c87ea49 60 0
 #2 := false
 decl f4 :: S1
@@ -741,6 +1315,40 @@
 #70 := [mp #35 #67]: #64
 [mp #70 #68]: 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
 27121de8ceb5985c14cca2c4563e9cad4d2ebd3c 79 0
 #2 := false
 decl f1 :: S1
@@ -821,6 +1429,25 @@
 #36 := [asserted]: #19
 [mp #36 #93]: 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
 dc14c19019afcaa63507e4b503184496ed6bfd2f 94 0
 #2 := false
 decl f5 :: S1
@@ -916,6 +1543,36 @@
 #96 := [not-or-elim #89]: #70
 [mp #96 #112]: 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
 682b764e75beeea86518e9d265c22d014c6c1f1d 114 0
 #2 := false
 decl f5 :: S1
@@ -1031,6 +1688,41 @@
 #136 := [unit-resolution #109 #135]: #87
 [unit-resolution #93 #136 #134]: 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
 c134dad603b5004748b4c36796dfbdf41bc232cc 121 0
 #2 := false
 decl f5 :: S1
@@ -1153,6 +1845,34 @@
 #99 := [not-or-elim #92]: #98
 [mp #99 #138]: 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
 7afff0259006c1c4b6a802d510f1b96695b3e9ad 110 0
 #2 := false
 decl f4 :: S1
@@ -1264,6 +1984,38 @@
 #84 := [and-elim #83]: #44
 [mp #84 #129]: 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
 83b90287eef3783de4d8fa4aad70d03dd0abddbf 127 0
 #2 := false
 decl f5 :: S1
@@ -1392,6 +2144,70 @@
 #149 := [unit-resolution #131 #147]: #97
 [unit-resolution #108 #149 #148]: 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
 a7f2c916e2192b153aaac29ea6f75d0842fd84d2 66 0
 #2 := false
 decl f3 :: S1
@@ -1459,6 +2275,202 @@
 #74 := [not-or-elim #84]: #33
 [unit-resolution #69 #74]: 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
 d398c106ccd0941d46bab8e5feda32347dd02740 156 0
 #2 := false
 decl f5 :: S1
@@ -1616,6 +2628,43 @@
 #114 := [def-axiom]: #113
 [unit-resolution #114 #172 #170]: 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
 78ae843b278e5f0cc65999ee60776dec7c9f683f 114 0
 #2 := false
 decl f5 :: S1
@@ -1731,6 +2780,84 @@
 #101 := [mp #36 #100]: #94
 [mp #101 #130]: 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
 016fee2f7f490b7727a4579bedbeaa0a03271f5c 198 0
 #2 := false
 decl f5 :: S1
@@ -1930,6 +3057,35 @@
 #141 := [not-or-elim #136]: #93
 [mp #141 #215]: 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
 3e0305075f83d65498b4a11b8ca72526b8e2d46d 134 0
 #2 := false
 decl f5 :: S1
@@ -2065,6 +3221,37 @@
 #121 := [not-or-elim #116]: #81
 [mp #121 #151]: 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
 78ecef8a6265ee63380698aee66f7437c6eb9791 162 0
 #2 := false
 decl f5 :: S1
@@ -2228,6 +3415,47 @@
 #178 := [unit-resolution #128 #177]: #70
 [unit-resolution #136 #178 #176 #168]: 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
 cbabdc011b8a302b61906c6191f0d7af27e25599 54 0
 #2 := false
 decl f1 :: S1
@@ -2283,6 +3511,40 @@
 #31 := [asserted]: #14
 [mp #31 #69]: 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
 ae8651ca265765f3bb029935ca54bf350a439257 144 0
 #2 := false
 decl f5 :: S1
@@ -2428,6 +3690,98 @@
 #126 := [not-or-elim #121]: #102
 [mp #126 #161]: 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
 f037e85f433fe4b6db0edebe38d7599d0d5f9a56 121 0
 #2 := false
 decl f5 :: S1
@@ -2735,6 +4089,32 @@
 #34 := [asserted]: #17
 [mp #34 #75]: 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
 cb74096d43b27b6df4bed4b142ce429518455d83 121 0
 #2 := false
 decl f5 :: S1
@@ -2857,6 +4237,41 @@
 #115 := [not-or-elim #109]: #90
 [mp #115 #138]: 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
 f67715569c1aa586d3c8e9a94e103b17f1f34614 84 0
 #2 := false
 decl f4 :: S1
@@ -2942,1645 +4357,6 @@
 #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)
@@ -4677,6 +4453,54 @@
 #102 := [and-elim #101]: #76
 [unit-resolution #102 #115]: 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
 48a51f989350a0e584db03a2950a272906be29b0 300 0
 #2 := false
 decl f4 :: (-> S2 S1)
@@ -4978,6 +4802,48 @@
 #609 := [quant-inst]: #615
 [unit-resolution #609 #257 #603]: 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
 8d1c135b088bf498c079d2958e8764a8092679cf 201 0
 #2 := false
 decl f3 :: (-> S2 S1)
@@ -5180,6 +5046,67 @@
 #211 := [quant-inst]: #320
 [unit-resolution #211 #595 #319]: 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
 c2e19ee7461e4fbfe65c34d98111e3914f852baf 244 0
 #2 := false
 decl f3 :: (-> S2 S1)
@@ -5425,6 +5352,79 @@
 #616 := [mp #343 #237]: #621
 [unit-resolution #616 #629 #330]: 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
 35b6826e84c97e2db1e3cd91ad69abb16f131b93 49 0
 #2 := false
 decl f1 :: S1
@@ -5475,6 +5475,32 @@
 #35 := [asserted]: #18
 [mp #35 #65]: false
 unsat
+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
 2e6f238fe51369f58bc6454a136532e4c6f3bd82 136 0
 #2 := false
 decl f3 :: (-> S2 S2 S1)
@@ -5612,6 +5638,41 @@
 #593 := [quant-inst]: #592
 [unit-resolution #593 #610 #258]: false
 unsat
+b28b1c7eb24edf6acf5466ddd1304a6b881a60f9 34 0
+#2 := false
+decl f1 :: S1
+#4 := f1
+decl f3 :: (-> S2 S1)
+#8 := (:var 0 S2)
+#9 := (f3 #8)
+#10 := (= #9 f1)
+#11 := (exists (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 := (exists (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
 a15d00efdb475e30bdd25ef3ed2227a8ea9c519f 124 0
 #2 := false
 decl f3 :: (-> S2 S1)
@@ -5737,383 +5798,6 @@
 #207 := [mp #217 #206]: #569
 [unit-resolution #207 #576 #88 #209]: false
 unsat
-0e7fa24f774b5d4840baee92f17adcdc6e36469d 64 0
-#2 := false
-decl f3 :: (-> S2 S2 S3)
-decl f4 :: S2
-#14 := f4
-decl f5 :: S2
-#16 := f5
-#18 := (f3 f5 f4)
-#17 := (f3 f4 f5)
-#19 := (= #17 #18)
-#62 := (not #19)
-#8 := (:var 1 S2)
-#9 := (:var 0 S2)
-#11 := (f3 #9 #8)
-#10 := (f3 #8 #9)
-#12 := (= #10 #11)
-#13 := (forall (vars (?v0 S2) (?v1 S2)) #12)
-#52 := (not #13)
-#53 := (or #52 #19)
-#58 := (not #53)
-#15 := (= f4 f4)
-#20 := (and #15 #19)
-#21 := (implies #13 #20)
-#22 := (not #21)
-#59 := (iff #22 #58)
-#56 := (iff #21 #53)
-#49 := (implies #13 #19)
-#54 := (iff #49 #53)
-#55 := [rewrite]: #54
-#50 := (iff #21 #49)
-#47 := (iff #20 #19)
-#1 := true
-#42 := (and true #19)
-#45 := (iff #42 #19)
-#46 := [rewrite]: #45
-#43 := (iff #20 #42)
-#40 := (iff #15 true)
-#41 := [rewrite]: #40
-#44 := [monotonicity #41]: #43
-#48 := [trans #44 #46]: #47
-#51 := [monotonicity #48]: #50
-#57 := [trans #51 #55]: #56
-#60 := [monotonicity #57]: #59
-#39 := [asserted]: #22
-#63 := [mp #39 #60]: #58
-#64 := [not-or-elim #63]: #62
-#547 := (pattern #11)
-#546 := (pattern #10)
-#548 := (forall (vars (?v0 S2) (?v1 S2)) (:pat #546 #547) #12)
-#551 := (iff #13 #548)
-#549 := (iff #12 #12)
-#550 := [refl]: #549
-#552 := [quant-intro #550]: #551
-#74 := (~ #13 #13)
-#72 := (~ #12 #12)
-#73 := [refl]: #72
-#75 := [nnf-pos #73]: #74
-#61 := [not-or-elim #63]: #13
-#67 := [mp~ #61 #75]: #13
-#553 := [mp #67 #552]: #548
-#129 := (not #548)
-#216 := (or #129 #19)
-#130 := [quant-inst]: #216
-[unit-resolution #130 #553 #64]: false
-unsat
-e97ea15e571d05d4542f0baceec935a45a20adf5 250 0
-#2 := false
-decl f3 :: (-> S2 S1)
-decl f7 :: S2
-#22 := f7
-#25 := (f3 f7)
-decl f1 :: S1
-#4 := f1
-#94 := (= f1 #25)
-#283 := (not #94)
-decl f4 :: (-> S2 S1)
-#8 := (:var 0 S2)
-#11 := (f4 #8)
-#702 := (pattern #11)
-#9 := (f3 #8)
-#701 := (pattern #9)
-#57 := (= f1 #11)
-#54 := (= f1 #9)
-#63 := (not #54)
-#64 := (or #63 #57)
-#703 := (forall (vars (?v0 S2)) (:pat #701 #702) #64)
-#69 := (forall (vars (?v0 S2)) #64)
-#706 := (iff #69 #703)
-#704 := (iff #64 #64)
-#705 := [refl]: #704
-#707 := [quant-intro #705]: #706
-#187 := (~ #69 #69)
-#201 := (~ #64 #64)
-#202 := [refl]: #201
-#188 := [nnf-pos #202]: #187
-decl f6 :: (-> S2 S1)
-#23 := (f6 f7)
-#91 := (= f1 #23)
-#78 := (not #57)
-#86 := (and #63 #78)
-#103 := (not #86)
-#136 := (or #103 #91 #94)
-#139 := (forall (vars (?v0 S2)) #136)
-decl f5 :: S2
-#15 := f5
-#16 := (f4 f5)
-#72 := (= f1 #16)
-#75 := (not #72)
-#148 := (and #69 #75 #139)
-#153 := (not #148)
-#32 := (f4 f7)
-#118 := (= f1 #32)
-#159 := (or #91 #118 #153)
-#164 := (not #159)
-#33 := (= #32 f1)
-#24 := (= #23 f1)
-#34 := (or #24 #33)
-#26 := (= #25 f1)
-#27 := (or #24 #26)
-#10 := (= #9 f1)
-#20 := (not #10)
-#12 := (= #11 f1)
-#19 := (not #12)
-#21 := (and #19 #20)
-#28 := (implies #21 #27)
-#29 := (forall (vars (?v0 S2)) #28)
-#17 := (= #16 f1)
-#18 := (not #17)
-#30 := (and #18 #29)
-#13 := (implies #10 #12)
-#14 := (forall (vars (?v0 S2)) #13)
-#31 := (and #14 #30)
-#35 := (implies #31 #34)
-#36 := (not #35)
-#167 := (iff #36 #164)
-#121 := (or #91 #118)
-#97 := (or #91 #94)
-#104 := (or #103 #97)
-#109 := (forall (vars (?v0 S2)) #104)
-#112 := (and #75 #109)
-#115 := (and #69 #112)
-#127 := (not #115)
-#128 := (or #127 #121)
-#133 := (not #128)
-#165 := (iff #133 #164)
-#162 := (iff #128 #159)
-#156 := (or #153 #121)
-#160 := (iff #156 #159)
-#161 := [rewrite]: #160
-#157 := (iff #128 #156)
-#154 := (iff #127 #153)
-#151 := (iff #115 #148)
-#142 := (and #75 #139)
-#145 := (and #69 #142)
-#149 := (iff #145 #148)
-#150 := [rewrite]: #149
-#146 := (iff #115 #145)
-#143 := (iff #112 #142)
-#140 := (iff #109 #139)
-#137 := (iff #104 #136)
-#138 := [rewrite]: #137
-#141 := [quant-intro #138]: #140
-#144 := [monotonicity #141]: #143
-#147 := [monotonicity #144]: #146
-#152 := [trans #147 #150]: #151
-#155 := [monotonicity #152]: #154
-#158 := [monotonicity #155]: #157
-#163 := [trans #158 #161]: #162
-#166 := [monotonicity #163]: #165
-#134 := (iff #36 #133)
-#131 := (iff #35 #128)
-#124 := (implies #115 #121)
-#129 := (iff #124 #128)
-#130 := [rewrite]: #129
-#125 := (iff #35 #124)
-#122 := (iff #34 #121)
-#119 := (iff #33 #118)
-#120 := [rewrite]: #119
-#92 := (iff #24 #91)
-#93 := [rewrite]: #92
-#123 := [monotonicity #93 #120]: #122
-#116 := (iff #31 #115)
-#113 := (iff #30 #112)
-#110 := (iff #29 #109)
-#107 := (iff #28 #104)
-#100 := (implies #86 #97)
-#105 := (iff #100 #104)
-#106 := [rewrite]: #105
-#101 := (iff #28 #100)
-#98 := (iff #27 #97)
-#95 := (iff #26 #94)
-#96 := [rewrite]: #95
-#99 := [monotonicity #93 #96]: #98
-#89 := (iff #21 #86)
-#83 := (and #78 #63)
-#87 := (iff #83 #86)
-#88 := [rewrite]: #87
-#84 := (iff #21 #83)
-#81 := (iff #20 #63)
-#55 := (iff #10 #54)
-#56 := [rewrite]: #55
-#82 := [monotonicity #56]: #81
-#79 := (iff #19 #78)
-#58 := (iff #12 #57)
-#59 := [rewrite]: #58
-#80 := [monotonicity #59]: #79
-#85 := [monotonicity #80 #82]: #84
-#90 := [trans #85 #88]: #89
-#102 := [monotonicity #90 #99]: #101
-#108 := [trans #102 #106]: #107
-#111 := [quant-intro #108]: #110
-#76 := (iff #18 #75)
-#73 := (iff #17 #72)
-#74 := [rewrite]: #73
-#77 := [monotonicity #74]: #76
-#114 := [monotonicity #77 #111]: #113
-#70 := (iff #14 #69)
-#67 := (iff #13 #64)
-#60 := (implies #54 #57)
-#65 := (iff #60 #64)
-#66 := [rewrite]: #65
-#61 := (iff #13 #60)
-#62 := [monotonicity #56 #59]: #61
-#68 := [trans #62 #66]: #67
-#71 := [quant-intro #68]: #70
-#117 := [monotonicity #71 #114]: #116
-#126 := [monotonicity #117 #123]: #125
-#132 := [trans #126 #130]: #131
-#135 := [monotonicity #132]: #134
-#168 := [trans #135 #166]: #167
-#53 := [asserted]: #36
-#169 := [mp #53 #168]: #164
-#174 := [not-or-elim #169]: #148
-#175 := [and-elim #174]: #69
-#185 := [mp~ #175 #188]: #69
-#708 := [mp #185 #707]: #703
-#172 := (not #118)
-#173 := [not-or-elim #169]: #172
-#285 := (not #703)
-#372 := (or #285 #283 #118)
-#370 := (or #283 #118)
-#363 := (or #285 #370)
-#375 := (iff #363 #372)
-#303 := [rewrite]: #375
-#374 := [quant-inst]: #363
-#376 := [mp #374 #303]: #372
-#398 := [unit-resolution #376 #173 #708]: #283
-#217 := (or #54 #57 #94)
-#709 := (forall (vars (?v0 S2)) (:pat #701 #702) #217)
-#222 := (forall (vars (?v0 S2)) #217)
-#712 := (iff #222 #709)
-#710 := (iff #217 #217)
-#711 := [refl]: #710
-#713 := [quant-intro #711]: #712
-#192 := (or #103 #94)
-#197 := (forall (vars (?v0 S2)) #192)
-#223 := (iff #197 #222)
-#220 := (iff #192 #217)
-#203 := (or #54 #57)
-#214 := (or #203 #94)
-#218 := (iff #214 #217)
-#219 := [rewrite]: #218
-#215 := (iff #192 #214)
-#212 := (iff #103 #203)
-#204 := (not #203)
-#207 := (not #204)
-#210 := (iff #207 #203)
-#211 := [rewrite]: #210
-#208 := (iff #103 #207)
-#205 := (iff #86 #204)
-#206 := [rewrite]: #205
-#209 := [monotonicity #206]: #208
-#213 := [trans #209 #211]: #212
-#216 := [monotonicity #213]: #215
-#221 := [trans #216 #219]: #220
-#224 := [quant-intro #221]: #223
-#184 := (~ #197 #197)
-#186 := (~ #192 #192)
-#183 := [refl]: #186
-#180 := [nnf-pos #183]: #184
-#198 := (iff #139 #197)
-#195 := (iff #136 #192)
-#189 := (or #103 false #94)
-#193 := (iff #189 #192)
-#194 := [rewrite]: #193
-#190 := (iff #136 #189)
-#181 := (iff #91 false)
-#170 := (not #91)
-#171 := [not-or-elim #169]: #170
-#182 := [iff-false #171]: #181
-#191 := [monotonicity #182]: #190
-#196 := [trans #191 #194]: #195
-#199 := [quant-intro #196]: #198
-#177 := [and-elim #174]: #139
-#200 := [mp #177 #199]: #197
-#178 := [mp~ #200 #180]: #197
-#225 := [mp #178 #224]: #222
-#714 := [mp #225 #713]: #709
-#356 := (not #709)
-#693 := (or #356 #94 #118)
-#284 := (or #94 #118 #94)
-#695 := (or #356 #284)
-#697 := (iff #695 #693)
-#371 := (or #94 #118)
-#482 := (or #356 #371)
-#362 := (iff #482 #693)
-#696 := [rewrite]: #362
-#689 := (iff #695 #482)
-#373 := (iff #284 #371)
-#377 := [rewrite]: #373
-#361 := [monotonicity #377]: #689
-#698 := [trans #361 #696]: #697
-#350 := [quant-inst]: #695
-#699 := [mp #350 #698]: #693
-[unit-resolution #699 #173 #714 #398]: false
-unsat
-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
-decl f3 :: (-> S2 S1)
-#8 := (:var 0 S2)
-#9 := (f3 #8)
-#10 := (= #9 f1)
-#11 := (exists (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 := (exists (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
 043713d1f8d615cef2b57b8a8ae1162f9c51de61 258 0
 #2 := false
 decl f4 :: (-> S2 S1)
@@ -6373,6 +6057,71 @@
 #579 := [quant-inst]: #578
 [unit-resolution #579 #432 #573]: false
 unsat
+0e7fa24f774b5d4840baee92f17adcdc6e36469d 64 0
+#2 := false
+decl f3 :: (-> S2 S2 S3)
+decl f4 :: S2
+#14 := f4
+decl f5 :: S2
+#16 := f5
+#18 := (f3 f5 f4)
+#17 := (f3 f4 f5)
+#19 := (= #17 #18)
+#62 := (not #19)
+#8 := (:var 1 S2)
+#9 := (:var 0 S2)
+#11 := (f3 #9 #8)
+#10 := (f3 #8 #9)
+#12 := (= #10 #11)
+#13 := (forall (vars (?v0 S2) (?v1 S2)) #12)
+#52 := (not #13)
+#53 := (or #52 #19)
+#58 := (not #53)
+#15 := (= f4 f4)
+#20 := (and #15 #19)
+#21 := (implies #13 #20)
+#22 := (not #21)
+#59 := (iff #22 #58)
+#56 := (iff #21 #53)
+#49 := (implies #13 #19)
+#54 := (iff #49 #53)
+#55 := [rewrite]: #54
+#50 := (iff #21 #49)
+#47 := (iff #20 #19)
+#1 := true
+#42 := (and true #19)
+#45 := (iff #42 #19)
+#46 := [rewrite]: #45
+#43 := (iff #20 #42)
+#40 := (iff #15 true)
+#41 := [rewrite]: #40
+#44 := [monotonicity #41]: #43
+#48 := [trans #44 #46]: #47
+#51 := [monotonicity #48]: #50
+#57 := [trans #51 #55]: #56
+#60 := [monotonicity #57]: #59
+#39 := [asserted]: #22
+#63 := [mp #39 #60]: #58
+#64 := [not-or-elim #63]: #62
+#547 := (pattern #11)
+#546 := (pattern #10)
+#548 := (forall (vars (?v0 S2) (?v1 S2)) (:pat #546 #547) #12)
+#551 := (iff #13 #548)
+#549 := (iff #12 #12)
+#550 := [refl]: #549
+#552 := [quant-intro #550]: #551
+#74 := (~ #13 #13)
+#72 := (~ #12 #12)
+#73 := [refl]: #72
+#75 := [nnf-pos #73]: #74
+#61 := [not-or-elim #63]: #13
+#67 := [mp~ #61 #75]: #13
+#553 := [mp #67 #552]: #548
+#129 := (not #548)
+#216 := (or #129 #19)
+#130 := [quant-inst]: #216
+[unit-resolution #130 #553 #64]: false
+unsat
 7fcf7ce1521f6f1599b2fe6f93ba5dee1cbd7b76 222 0
 #2 := false
 decl f3 :: (-> S2 S1)
@@ -6596,6 +6345,257 @@
 #610 := [quant-inst]: #609
 [unit-resolution #610 #339 #320]: false
 unsat
+e97ea15e571d05d4542f0baceec935a45a20adf5 250 0
+#2 := false
+decl f3 :: (-> S2 S1)
+decl f7 :: S2
+#22 := f7
+#25 := (f3 f7)
+decl f1 :: S1
+#4 := f1
+#94 := (= f1 #25)
+#283 := (not #94)
+decl f4 :: (-> S2 S1)
+#8 := (:var 0 S2)
+#11 := (f4 #8)
+#702 := (pattern #11)
+#9 := (f3 #8)
+#701 := (pattern #9)
+#57 := (= f1 #11)
+#54 := (= f1 #9)
+#63 := (not #54)
+#64 := (or #63 #57)
+#703 := (forall (vars (?v0 S2)) (:pat #701 #702) #64)
+#69 := (forall (vars (?v0 S2)) #64)
+#706 := (iff #69 #703)
+#704 := (iff #64 #64)
+#705 := [refl]: #704
+#707 := [quant-intro #705]: #706
+#187 := (~ #69 #69)
+#201 := (~ #64 #64)
+#202 := [refl]: #201
+#188 := [nnf-pos #202]: #187
+decl f6 :: (-> S2 S1)
+#23 := (f6 f7)
+#91 := (= f1 #23)
+#78 := (not #57)
+#86 := (and #63 #78)
+#103 := (not #86)
+#136 := (or #103 #91 #94)
+#139 := (forall (vars (?v0 S2)) #136)
+decl f5 :: S2
+#15 := f5
+#16 := (f4 f5)
+#72 := (= f1 #16)
+#75 := (not #72)
+#148 := (and #69 #75 #139)
+#153 := (not #148)
+#32 := (f4 f7)
+#118 := (= f1 #32)
+#159 := (or #91 #118 #153)
+#164 := (not #159)
+#33 := (= #32 f1)
+#24 := (= #23 f1)
+#34 := (or #24 #33)
+#26 := (= #25 f1)
+#27 := (or #24 #26)
+#10 := (= #9 f1)
+#20 := (not #10)
+#12 := (= #11 f1)
+#19 := (not #12)
+#21 := (and #19 #20)
+#28 := (implies #21 #27)
+#29 := (forall (vars (?v0 S2)) #28)
+#17 := (= #16 f1)
+#18 := (not #17)
+#30 := (and #18 #29)
+#13 := (implies #10 #12)
+#14 := (forall (vars (?v0 S2)) #13)
+#31 := (and #14 #30)
+#35 := (implies #31 #34)
+#36 := (not #35)
+#167 := (iff #36 #164)
+#121 := (or #91 #118)
+#97 := (or #91 #94)
+#104 := (or #103 #97)
+#109 := (forall (vars (?v0 S2)) #104)
+#112 := (and #75 #109)
+#115 := (and #69 #112)
+#127 := (not #115)
+#128 := (or #127 #121)
+#133 := (not #128)
+#165 := (iff #133 #164)
+#162 := (iff #128 #159)
+#156 := (or #153 #121)
+#160 := (iff #156 #159)
+#161 := [rewrite]: #160
+#157 := (iff #128 #156)
+#154 := (iff #127 #153)
+#151 := (iff #115 #148)
+#142 := (and #75 #139)
+#145 := (and #69 #142)
+#149 := (iff #145 #148)
+#150 := [rewrite]: #149
+#146 := (iff #115 #145)
+#143 := (iff #112 #142)
+#140 := (iff #109 #139)
+#137 := (iff #104 #136)
+#138 := [rewrite]: #137
+#141 := [quant-intro #138]: #140
+#144 := [monotonicity #141]: #143
+#147 := [monotonicity #144]: #146
+#152 := [trans #147 #150]: #151
+#155 := [monotonicity #152]: #154
+#158 := [monotonicity #155]: #157
+#163 := [trans #158 #161]: #162
+#166 := [monotonicity #163]: #165
+#134 := (iff #36 #133)
+#131 := (iff #35 #128)
+#124 := (implies #115 #121)
+#129 := (iff #124 #128)
+#130 := [rewrite]: #129
+#125 := (iff #35 #124)
+#122 := (iff #34 #121)
+#119 := (iff #33 #118)
+#120 := [rewrite]: #119
+#92 := (iff #24 #91)
+#93 := [rewrite]: #92
+#123 := [monotonicity #93 #120]: #122
+#116 := (iff #31 #115)
+#113 := (iff #30 #112)
+#110 := (iff #29 #109)
+#107 := (iff #28 #104)
+#100 := (implies #86 #97)
+#105 := (iff #100 #104)
+#106 := [rewrite]: #105
+#101 := (iff #28 #100)
+#98 := (iff #27 #97)
+#95 := (iff #26 #94)
+#96 := [rewrite]: #95
+#99 := [monotonicity #93 #96]: #98
+#89 := (iff #21 #86)
+#83 := (and #78 #63)
+#87 := (iff #83 #86)
+#88 := [rewrite]: #87
+#84 := (iff #21 #83)
+#81 := (iff #20 #63)
+#55 := (iff #10 #54)
+#56 := [rewrite]: #55
+#82 := [monotonicity #56]: #81
+#79 := (iff #19 #78)
+#58 := (iff #12 #57)
+#59 := [rewrite]: #58
+#80 := [monotonicity #59]: #79
+#85 := [monotonicity #80 #82]: #84
+#90 := [trans #85 #88]: #89
+#102 := [monotonicity #90 #99]: #101
+#108 := [trans #102 #106]: #107
+#111 := [quant-intro #108]: #110
+#76 := (iff #18 #75)
+#73 := (iff #17 #72)
+#74 := [rewrite]: #73
+#77 := [monotonicity #74]: #76
+#114 := [monotonicity #77 #111]: #113
+#70 := (iff #14 #69)
+#67 := (iff #13 #64)
+#60 := (implies #54 #57)
+#65 := (iff #60 #64)
+#66 := [rewrite]: #65
+#61 := (iff #13 #60)
+#62 := [monotonicity #56 #59]: #61
+#68 := [trans #62 #66]: #67
+#71 := [quant-intro #68]: #70
+#117 := [monotonicity #71 #114]: #116
+#126 := [monotonicity #117 #123]: #125
+#132 := [trans #126 #130]: #131
+#135 := [monotonicity #132]: #134
+#168 := [trans #135 #166]: #167
+#53 := [asserted]: #36
+#169 := [mp #53 #168]: #164
+#174 := [not-or-elim #169]: #148
+#175 := [and-elim #174]: #69
+#185 := [mp~ #175 #188]: #69
+#708 := [mp #185 #707]: #703
+#172 := (not #118)
+#173 := [not-or-elim #169]: #172
+#285 := (not #703)
+#372 := (or #285 #283 #118)
+#370 := (or #283 #118)
+#363 := (or #285 #370)
+#375 := (iff #363 #372)
+#303 := [rewrite]: #375
+#374 := [quant-inst]: #363
+#376 := [mp #374 #303]: #372
+#398 := [unit-resolution #376 #173 #708]: #283
+#217 := (or #54 #57 #94)
+#709 := (forall (vars (?v0 S2)) (:pat #701 #702) #217)
+#222 := (forall (vars (?v0 S2)) #217)
+#712 := (iff #222 #709)
+#710 := (iff #217 #217)
+#711 := [refl]: #710
+#713 := [quant-intro #711]: #712
+#192 := (or #103 #94)
+#197 := (forall (vars (?v0 S2)) #192)
+#223 := (iff #197 #222)
+#220 := (iff #192 #217)
+#203 := (or #54 #57)
+#214 := (or #203 #94)
+#218 := (iff #214 #217)
+#219 := [rewrite]: #218
+#215 := (iff #192 #214)
+#212 := (iff #103 #203)
+#204 := (not #203)
+#207 := (not #204)
+#210 := (iff #207 #203)
+#211 := [rewrite]: #210
+#208 := (iff #103 #207)
+#205 := (iff #86 #204)
+#206 := [rewrite]: #205
+#209 := [monotonicity #206]: #208
+#213 := [trans #209 #211]: #212
+#216 := [monotonicity #213]: #215
+#221 := [trans #216 #219]: #220
+#224 := [quant-intro #221]: #223
+#184 := (~ #197 #197)
+#186 := (~ #192 #192)
+#183 := [refl]: #186
+#180 := [nnf-pos #183]: #184
+#198 := (iff #139 #197)
+#195 := (iff #136 #192)
+#189 := (or #103 false #94)
+#193 := (iff #189 #192)
+#194 := [rewrite]: #193
+#190 := (iff #136 #189)
+#181 := (iff #91 false)
+#170 := (not #91)
+#171 := [not-or-elim #169]: #170
+#182 := [iff-false #171]: #181
+#191 := [monotonicity #182]: #190
+#196 := [trans #191 #194]: #195
+#199 := [quant-intro #196]: #198
+#177 := [and-elim #174]: #139
+#200 := [mp #177 #199]: #197
+#178 := [mp~ #200 #180]: #197
+#225 := [mp #178 #224]: #222
+#714 := [mp #225 #713]: #709
+#356 := (not #709)
+#693 := (or #356 #94 #118)
+#284 := (or #94 #118 #94)
+#695 := (or #356 #284)
+#697 := (iff #695 #693)
+#371 := (or #94 #118)
+#482 := (or #356 #371)
+#362 := (iff #482 #693)
+#696 := [rewrite]: #362
+#689 := (iff #695 #482)
+#373 := (iff #284 #371)
+#377 := [rewrite]: #373
+#361 := [monotonicity #377]: #689
+#698 := [trans #361 #696]: #697
+#350 := [quant-inst]: #695
+#699 := [mp #350 #698]: #693
+[unit-resolution #699 #173 #714 #398]: false
+unsat
 8031eb3bcd3fb97da7b874e36b0d4d2ae215ec7b 49 0
 #2 := false
 decl f1 :: S1
@@ -6646,6 +6646,10 @@
 #35 := [asserted]: #18
 [mp #35 #65]: false
 unsat
+2317912b27971033dc831a6ce5ef74d0f20fbca2 1 2
+unsat
+WARNING: failed to find a pattern for quantifier (quantifier id: k!9)
+
 301a246cbbd414b8b75c6b801fee2f734cf2af47 122 0
 #2 := false
 decl f3 :: (-> S2 S1)
@@ -6769,20 +6773,8 @@
 #287 := [quant-inst]: #200
 [unit-resolution #287 #620 #138]: false
 unsat
-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)
@@ -6889,6 +6881,8 @@
 #261 := [quant-inst]: #188
 [unit-resolution #261 #590 #259]: false
 unsat
+f7daa562701d4999ed8a3f5ec133f08651bf6190 1 0
+unsat
 ff2e1b7168f418022909b6d3ec628836562df57d 230 0
 #2 := false
 decl f3 :: (-> S1 S1)
@@ -7120,6 +7114,8 @@
 #309 := [quant-inst]: #308
 [unit-resolution #309 #590 #312]: false
 unsat
+fdbe39cdd485517d8ec1964f2c4fb7b6646fe608 1 0
+unsat
 258fe7f9c05eff8e107778930f55bcfa407e2fde 156 0
 #2 := false
 decl f3 :: (-> S1 S2 S1)
@@ -7277,6 +7273,8 @@
 #269 := [quant-inst]: #264
 [unit-resolution #269 #592 #292]: false
 unsat
+779c5fc38fd08b8484ec1da52f6a17e6b849eac4 1 0
+unsat
 6c0f54222b59d5e628ff948749ac67e96364c6a1 74 0
 #2 := false
 decl f3 :: (-> S2 S3 S4)
@@ -7352,6 +7350,8 @@
 #548 := [mp #55 #547]: #543
 [unit-resolution #548 #537]: false
 unsat
+b720bb30cc87882c3f660a5bf9c8737dde4c14c4 1 0
+unsat
 cba176f179e597e7b48393532028be8634e0c4f7 108 0
 #2 := false
 decl f3 :: (-> S1 S1)
@@ -7661,8 +7661,6 @@
 unsat
 ce6d32e932c7cece567323ec34f7a5540258fa7f 1 0
 unsat
-22c7d1d1136fafc0a09e243a12ede1ee4fee416f 1 0
-unsat
 cfd9d2f647046e282a1080aae51b4b88be291157 107 0
 #2 := false
 decl f3 :: (-> S2 S1)
@@ -7771,6 +7769,8 @@
 #279 := [quant-inst]: #193
 [unit-resolution #279 #102 #618]: false
 unsat
+22c7d1d1136fafc0a09e243a12ede1ee4fee416f 1 0
+unsat
 9ea0d22c6270219f7cef6ee3d8f429562349b2a3 75 0
 #2 := false
 decl f1 :: S1
@@ -8055,6 +8055,100 @@
 #622 := [mp #616 #257]: #277
 [unit-resolution #622 #645 #591 #596]: false
 unsat
+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
 628763241ae198f2ea4a54c641ca6e8df6d851ca 250 0
 #2 := false
 decl f3 :: (-> S2 S1)
@@ -8306,6 +8400,120 @@
 #616 := [mp #628 #615]: #622
 [unit-resolution #616 #128 #676 #607 #606]: 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
 64bd43a9d49de65761e5220e25dbbe3dc269bf0e 238 0
 #2 := false
 decl f3 :: (-> S2 S1)
@@ -8545,214 +8753,6 @@
 #338 := [mp #643 #632]: #276
 [unit-resolution #338 #664 #617 #616]: false
 unsat
-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
@@ -9063,6 +9063,10 @@
 #29 := [asserted]: #12
 [mp #29 #46]: false
 unsat
+46acb4cd4b5353c32c529ffe509fc6f26ac2f2a1 3 0
+#2 := false
+[asserted]: false
+unsat
 ac00789473bf074e49a8a6419c1f2a1042912d6a 48 0
 #2 := false
 decl f1 :: S1
@@ -9112,6 +9116,25 @@
 #31 := [asserted]: #14
 [mp #31 #63]: false
 unsat
+040123bf68e09d0b72cf990436bb29a623c102b8 18 0
+#2 := false
+#9 := 1::int
+#8 := 0::int
+#10 := (< 0::int 1::int)
+#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
 75052914158f379cd343f693f354895fa579a89c 53 0
 #2 := false
 decl f1 :: S1
@@ -9166,6 +9189,25 @@
 #31 := [asserted]: #14
 [mp #31 #68]: false
 unsat
+be1600749b493592f603ae3a7f9dcbc68d688187 18 0
+#2 := false
+#9 := 1::int
+#8 := 0::int
+#10 := (<= 0::int 1::int)
+#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
 4bcd26d70bcaf24bb295031b732101a3dec2c643 56 0
 #2 := false
 decl f1 :: S1
@@ -9223,6 +9265,25 @@
 #32 := [asserted]: #15
 [mp #32 #71]: false
 unsat
+d6985ded37fcab252cac3fbb09a30e977c68973d 18 0
+#2 := false
+#9 := 2345678901::int
+#8 := 123456789::int
+#10 := (< 123456789::int 2345678901::int)
+#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
 7eeb56537bd13a6f6bc35096a4ddfb33dceb295f 89 0
 #2 := false
 decl f4 :: S1
@@ -9371,6 +9432,37 @@
 #59 := [not-or-elim #57]: #58
 [mp #59 #74]: false
 unsat
+12f34f8245ed278a095c475d68eb5a5711d23981 30 0
+#2 := false
+decl f3 :: (-> int S2)
+#25 := 1::int
+#28 := (f3 1::int)
+#13 := 0::int
+#26 := (+ 0::int 1::int)
+#27 := (f3 #26)
+#29 := (= #27 #28)
+#30 := (not #29)
+#148 := (iff #30 false)
+#1 := true
+#143 := (not true)
+#146 := (iff #143 false)
+#147 := [rewrite]: #146
+#144 := (iff #30 #143)
+#141 := (iff #29 true)
+#136 := (= #28 #28)
+#139 := (iff #136 true)
+#140 := [rewrite]: #139
+#137 := (iff #29 #136)
+#133 := (= #26 1::int)
+#134 := [rewrite]: #133
+#135 := [monotonicity #134]: #29
+#138 := [monotonicity #135]: #137
+#142 := [trans #138 #140]: #141
+#145 := [monotonicity #142]: #144
+#149 := [trans #145 #147]: #148
+#132 := [asserted]: #30
+[mp #132 #149]: false
+unsat
 8e8baf35daf48781c204954ba8d7df5254aac6a8 50 0
 #2 := false
 decl f5 :: S2
@@ -9450,6 +9542,40 @@
 #43 := [not-or-elim #42]: #41
 [unit-resolution #43 #51]: false
 unsat
+fca9345b5fb82879eab30d8a35a71a4cfb012222 33 0
+#2 := false
+#27 := 1::int
+decl f4 :: (-> S2 int)
+decl f5 :: S2
+#25 := f5
+#26 := (f4 f5)
+#28 := (+ #26 1::int)
+#29 := (< #26 #28)
+#30 := (not #29)
+#151 := (iff #30 false)
+#133 := (+ 1::int #26)
+#136 := (< #26 #133)
+#139 := (not #136)
+#149 := (iff #139 false)
+#1 := true
+#144 := (not true)
+#147 := (iff #144 false)
+#148 := [rewrite]: #147
+#145 := (iff #139 #144)
+#142 := (iff #136 true)
+#143 := [rewrite]: #142
+#146 := [monotonicity #143]: #145
+#150 := [trans #146 #148]: #149
+#140 := (iff #30 #139)
+#137 := (iff #29 #136)
+#134 := (= #28 #133)
+#135 := [rewrite]: #134
+#138 := [monotonicity #135]: #137
+#141 := [monotonicity #138]: #140
+#152 := [trans #141 #150]: #151
+#132 := [asserted]: #30
+[mp #132 #152]: false
+unsat
 a8014c2ba4ebe428717a3516d1900ca3b7591555 73 0
 #2 := false
 decl f3 :: (-> S1 S1)
@@ -9524,558 +9650,6 @@
 #146 := [quant-inst]: #232
 [unit-resolution #146 #566 #74]: false
 unsat
-c0906c23e4b0650863a2ef0d68bafa4e06a9dd19 106 0
-#2 := false
-decl f3 :: (-> S2 S2 S1)
-decl ?v0!0 :: S2
-#72 := ?v0!0
-#83 := (f3 ?v0!0 ?v0!0)
-decl f1 :: S1
-#4 := f1
-#75 := (= f1 #83)
-#76 := (not #75)
-#9 := (:var 0 S2)
-#16 := (f3 #9 #9)
-#50 := (= f1 #16)
-#53 := (forall (vars (?v0 S2)) #50)
-#69 := (not #53)
-#84 := (~ #69 #76)
-#85 := [sk]: #84
-#8 := (:var 1 S2)
-#12 := (f3 #9 #8)
-#41 := (= f1 #12)
-#10 := (f3 #8 #9)
-#38 := (= f1 #10)
-#44 := (and #38 #41)
-#47 := (forall (vars (?v0 S2) (?v1 S2)) #44)
-#59 := (not #47)
-#60 := (or #59 #53)
-#65 := (not #60)
-#17 := (= #16 f1)
-#18 := (forall (vars (?v0 S2)) #17)
-#13 := (= #12 f1)
-#11 := (= #10 f1)
-#14 := (and #11 #13)
-#15 := (forall (vars (?v0 S2) (?v1 S2)) #14)
-#19 := (implies #15 #18)
-#20 := (not #19)
-#66 := (iff #20 #65)
-#63 := (iff #19 #60)
-#56 := (implies #47 #53)
-#61 := (iff #56 #60)
-#62 := [rewrite]: #61
-#57 := (iff #19 #56)
-#54 := (iff #18 #53)
-#51 := (iff #17 #50)
-#52 := [rewrite]: #51
-#55 := [quant-intro #52]: #54
-#48 := (iff #15 #47)
-#45 := (iff #14 #44)
-#42 := (iff #13 #41)
-#43 := [rewrite]: #42
-#39 := (iff #11 #38)
-#40 := [rewrite]: #39
-#46 := [monotonicity #40 #43]: #45
-#49 := [quant-intro #46]: #48
-#58 := [monotonicity #49 #55]: #57
-#64 := [trans #58 #62]: #63
-#67 := [monotonicity #64]: #66
-#37 := [asserted]: #20
-#70 := [mp #37 #67]: #65
-#71 := [not-or-elim #70]: #69
-#88 := [mp~ #71 #85]: #76
-#577 := (pattern #12)
-#576 := (pattern #10)
-#87 := (not #41)
-#86 := (not #38)
-#89 := (or #86 #87)
-#90 := (not #89)
-#578 := (forall (vars (?v0 S2) (?v1 S2)) (:pat #576 #577) #90)
-#93 := (forall (vars (?v0 S2) (?v1 S2)) #90)
-#581 := (iff #93 #578)
-#579 := (iff #90 #90)
-#580 := [refl]: #579
-#582 := [quant-intro #580]: #581
-#94 := (iff #47 #93)
-#91 := (iff #44 #90)
-#92 := [rewrite]: #91
-#95 := [quant-intro #92]: #94
-#81 := (~ #47 #47)
-#79 := (~ #44 #44)
-#80 := [refl]: #79
-#82 := [nnf-pos #80]: #81
-#68 := [not-or-elim #70]: #47
-#74 := [mp~ #68 #82]: #47
-#96 := [mp #74 #95]: #93
-#583 := [mp #96 #582]: #578
-#250 := (not #578)
-#254 := (or #250 #75)
-#160 := (or #76 #76)
-#247 := (not #160)
-#233 := (or #250 #247)
-#570 := (iff #233 #254)
-#357 := (iff #254 #254)
-#564 := [rewrite]: #357
-#180 := (iff #247 #75)
-#162 := (not #76)
-#251 := (iff #162 #75)
-#252 := [rewrite]: #251
-#249 := (iff #247 #162)
-#161 := (iff #160 #76)
-#248 := [rewrite]: #161
-#240 := [monotonicity #248]: #249
-#253 := [trans #240 #252]: #180
-#227 := [monotonicity #253]: #570
-#238 := [trans #227 #564]: #570
-#568 := [quant-inst]: #233
-#239 := [mp #568 #238]: #254
-[unit-resolution #239 #583 #88]: false
-unsat
-58aaefc3b25d8181c01a2139430686aeca861bf3 66 0
-#2 := false
-decl f4 :: S1
-#10 := f4
-decl f1 :: S1
-#4 := f1
-#37 := (= f1 f4)
-decl f3 :: S1
-#8 := f3
-#34 := (= f1 f3)
-#40 := (or #34 #37)
-#84 := (iff #40 false)
-#79 := (or false false)
-#82 := (iff #79 false)
-#83 := [rewrite]: #82
-#80 := (iff #40 #79)
-#75 := (iff #37 false)
-#66 := (not #37)
-#43 := (not #34)
-#49 := (and #43 #40)
-#57 := (not #49)
-#58 := (or #37 #57)
-#63 := (not #58)
-#11 := (= f4 f1)
-#9 := (= f3 f1)
-#13 := (not #9)
-#12 := (or #9 #11)
-#14 := (and #12 #13)
-#15 := (implies #14 #11)
-#16 := (not #15)
-#64 := (iff #16 #63)
-#61 := (iff #15 #58)
-#54 := (implies #49 #37)
-#59 := (iff #54 #58)
-#60 := [rewrite]: #59
-#55 := (iff #15 #54)
-#38 := (iff #11 #37)
-#39 := [rewrite]: #38
-#52 := (iff #14 #49)
-#46 := (and #40 #43)
-#50 := (iff #46 #49)
-#51 := [rewrite]: #50
-#47 := (iff #14 #46)
-#44 := (iff #13 #43)
-#35 := (iff #9 #34)
-#36 := [rewrite]: #35
-#45 := [monotonicity #36]: #44
-#41 := (iff #12 #40)
-#42 := [monotonicity #36 #39]: #41
-#48 := [monotonicity #42 #45]: #47
-#53 := [trans #48 #51]: #52
-#56 := [monotonicity #53 #39]: #55
-#62 := [trans #56 #60]: #61
-#65 := [monotonicity #62]: #64
-#33 := [asserted]: #16
-#68 := [mp #33 #65]: #63
-#67 := [not-or-elim #68]: #66
-#76 := [iff-false #67]: #75
-#77 := (iff #34 false)
-#69 := [not-or-elim #68]: #49
-#70 := [and-elim #69]: #43
-#78 := [iff-false #70]: #77
-#81 := [monotonicity #78 #76]: #80
-#85 := [trans #81 #83]: #84
-#71 := [and-elim #69]: #40
-[mp #71 #85]: false
-unsat
-d80d7bb97197c962c0b81d44c7a8cfd617b6ca33 59 0
-#2 := false
-decl f1 :: S1
-#4 := f1
-decl f6 :: S1
-#15 := f6
-#16 := (= f6 f1)
-decl f5 :: S1
-#13 := f5
-#14 := (= f5 f1)
-#17 := (and #14 #16)
-decl f4 :: S1
-#10 := f4
-#11 := (= f4 f1)
-decl f3 :: S1
-#8 := f3
-#9 := (= f3 f1)
-#12 := (and #9 #11)
-#18 := (or #12 #17)
-#19 := (implies #18 #18)
-#20 := (not #19)
-#71 := (iff #20 false)
-#1 := true
-#66 := (not true)
-#69 := (iff #66 false)
-#70 := [rewrite]: #69
-#67 := (iff #20 #66)
-#64 := (iff #19 true)
-#50 := (= f1 f6)
-#47 := (= f1 f5)
-#53 := (and #47 #50)
-#41 := (= f1 f4)
-#38 := (= f1 f3)
-#44 := (and #38 #41)
-#56 := (or #44 #53)
-#59 := (implies #56 #56)
-#62 := (iff #59 true)
-#63 := [rewrite]: #62
-#60 := (iff #19 #59)
-#57 := (iff #18 #56)
-#54 := (iff #17 #53)
-#51 := (iff #16 #50)
-#52 := [rewrite]: #51
-#48 := (iff #14 #47)
-#49 := [rewrite]: #48
-#55 := [monotonicity #49 #52]: #54
-#45 := (iff #12 #44)
-#42 := (iff #11 #41)
-#43 := [rewrite]: #42
-#39 := (iff #9 #38)
-#40 := [rewrite]: #39
-#46 := [monotonicity #40 #43]: #45
-#58 := [monotonicity #46 #55]: #57
-#61 := [monotonicity #58 #58]: #60
-#65 := [trans #61 #63]: #64
-#68 := [monotonicity #65]: #67
-#72 := [trans #68 #70]: #71
-#37 := [asserted]: #20
-[mp #37 #72]: false
-unsat
-46acb4cd4b5353c32c529ffe509fc6f26ac2f2a1 3 0
-#2 := false
-[asserted]: false
-unsat
-040123bf68e09d0b72cf990436bb29a623c102b8 18 0
-#2 := false
-#9 := 1::int
-#8 := 0::int
-#10 := (< 0::int 1::int)
-#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
-be1600749b493592f603ae3a7f9dcbc68d688187 18 0
-#2 := false
-#9 := 1::int
-#8 := 0::int
-#10 := (<= 0::int 1::int)
-#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
-d6985ded37fcab252cac3fbb09a30e977c68973d 18 0
-#2 := false
-#9 := 2345678901::int
-#8 := 123456789::int
-#10 := (< 123456789::int 2345678901::int)
-#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
-12f34f8245ed278a095c475d68eb5a5711d23981 30 0
-#2 := false
-decl f3 :: (-> int S2)
-#25 := 1::int
-#28 := (f3 1::int)
-#13 := 0::int
-#26 := (+ 0::int 1::int)
-#27 := (f3 #26)
-#29 := (= #27 #28)
-#30 := (not #29)
-#148 := (iff #30 false)
-#1 := true
-#143 := (not true)
-#146 := (iff #143 false)
-#147 := [rewrite]: #146
-#144 := (iff #30 #143)
-#141 := (iff #29 true)
-#136 := (= #28 #28)
-#139 := (iff #136 true)
-#140 := [rewrite]: #139
-#137 := (iff #29 #136)
-#133 := (= #26 1::int)
-#134 := [rewrite]: #133
-#135 := [monotonicity #134]: #29
-#138 := [monotonicity #135]: #137
-#142 := [trans #138 #140]: #141
-#145 := [monotonicity #142]: #144
-#149 := [trans #145 #147]: #148
-#132 := [asserted]: #30
-[mp #132 #149]: false
-unsat
-fb0035d341cb2e8ac6449dbe61b4ac5d9e12a151 225 0
-#2 := false
-#13 := 0::int
-decl f4 :: (-> S2 int)
-decl f3 :: (-> int S2)
-decl f5 :: S2
-#25 := f5
-#26 := (f4 f5)
-#27 := 1::int
-#135 := (+ 1::int #26)
-#138 := (f3 #135)
-#141 := (f4 #138)
-#150 := -1::int
-#153 := (* -1::int #141)
-#154 := (+ #26 #153)
-#152 := (>= #154 0::int)
-#28 := (+ #26 1::int)
-#29 := (f3 #28)
-#30 := (f4 #29)
-#31 := (< #26 #30)
-#32 := (not #31)
-#164 := (iff #32 #152)
-#144 := (< #26 #141)
-#147 := (not #144)
-#162 := (iff #147 #152)
-#151 := (not #152)
-#157 := (not #151)
-#160 := (iff #157 #152)
-#161 := [rewrite]: #160
-#158 := (iff #147 #157)
-#155 := (iff #144 #151)
-#156 := [rewrite]: #155
-#159 := [monotonicity #156]: #158
-#163 := [trans #159 #161]: #162
-#148 := (iff #32 #147)
-#145 := (iff #31 #144)
-#142 := (= #30 #141)
-#139 := (= #29 #138)
-#136 := (= #28 #135)
-#137 := [rewrite]: #136
-#140 := [monotonicity #137]: #139
-#143 := [monotonicity #140]: #142
-#146 := [monotonicity #143]: #145
-#149 := [monotonicity #146]: #148
-#165 := [trans #149 #163]: #164
-#134 := [asserted]: #32
-#166 := [mp #134 #165]: #152
-#347 := (>= #141 0::int)
-#274 := (= #141 0::int)
-#642 := (>= #26 -1::int)
-#649 := (not #642)
-#648 := (= #154 -1::int)
-#336 := (not #648)
-#630 := (<= #154 -1::int)
-#621 := (not #630)
-#625 := (or #621 #151)
-#622 := [th-lemma]: #625
-#626 := [unit-resolution #622 #166]: #621
-#337 := (or #336 #630)
-#338 := [th-lemma]: #337
-#339 := [unit-resolution #338 #626]: #336
-#14 := (:var 0 int)
-#16 := (f3 #14)
-#662 := (pattern #16)
-#75 := (>= #14 0::int)
-#76 := (not #75)
-#17 := (f4 #16)
-#57 := (= #14 #17)
-#82 := (or #57 #76)
-#663 := (forall (vars (?v0 int)) (:pat #662) #82)
-#87 := (forall (vars (?v0 int)) #82)
-#666 := (iff #87 #663)
-#664 := (iff #82 #82)
-#665 := [refl]: #664
-#667 := [quant-intro #665]: #666
-#170 := (~ #87 #87)
-#167 := (~ #82 #82)
-#182 := [refl]: #167
-#171 := [nnf-pos #182]: #170
-#18 := (= #17 #14)
-#15 := (<= 0::int #14)
-#19 := (implies #15 #18)
-#20 := (forall (vars (?v0 int)) #19)
-#90 := (iff #20 #87)
-#64 := (not #15)
-#65 := (or #64 #57)
-#70 := (forall (vars (?v0 int)) #65)
-#88 := (iff #70 #87)
-#85 := (iff #65 #82)
-#79 := (or #76 #57)
-#83 := (iff #79 #82)
-#84 := [rewrite]: #83
-#80 := (iff #65 #79)
-#77 := (iff #64 #76)
-#73 := (iff #15 #75)
-#74 := [rewrite]: #73
-#78 := [monotonicity #74]: #77
-#81 := [monotonicity #78]: #80
-#86 := [trans #81 #84]: #85
-#89 := [quant-intro #86]: #88
-#71 := (iff #20 #70)
-#68 := (iff #19 #65)
-#61 := (implies #15 #57)
-#66 := (iff #61 #65)
-#67 := [rewrite]: #66
-#62 := (iff #19 #61)
-#59 := (iff #18 #57)
-#60 := [rewrite]: #59
-#63 := [monotonicity #60]: #62
-#69 := [trans #63 #67]: #68
-#72 := [quant-intro #69]: #71
-#91 := [trans #72 #89]: #90
-#56 := [asserted]: #20
-#92 := [mp #56 #91]: #87
-#183 := [mp~ #92 #171]: #87
-#668 := [mp #183 #667]: #663
-#645 := (not #663)
-#288 := (or #645 #649 #648)
-#330 := (>= #135 0::int)
-#331 := (not #330)
-#311 := (= #135 #141)
-#646 := (or #311 #331)
-#629 := (or #645 #646)
-#633 := (iff #629 #288)
-#643 := (or #649 #648)
-#293 := (or #645 #643)
-#631 := (iff #293 #288)
-#632 := [rewrite]: #631
-#294 := (iff #629 #293)
-#644 := (iff #646 #643)
-#652 := (or #648 #649)
-#303 := (iff #652 #643)
-#308 := [rewrite]: #303
-#647 := (iff #646 #652)
-#650 := (iff #331 #649)
-#316 := (iff #330 #642)
-#317 := [rewrite]: #316
-#651 := [monotonicity #317]: #650
-#305 := (iff #311 #648)
-#435 := [rewrite]: #305
-#653 := [monotonicity #435 #651]: #647
-#304 := [trans #653 #308]: #644
-#295 := [monotonicity #304]: #294
-#634 := [trans #295 #632]: #633
-#292 := [quant-inst]: #629
-#635 := [mp #292 #634]: #288
-#617 := [unit-resolution #635 #668 #339]: #649
-#22 := (= #17 0::int)
-#123 := (or #22 #75)
-#669 := (forall (vars (?v0 int)) (:pat #662) #123)
-#128 := (forall (vars (?v0 int)) #123)
-#672 := (iff #128 #669)
-#670 := (iff #123 #123)
-#671 := [refl]: #670
-#673 := [quant-intro #671]: #672
-#172 := (~ #128 #128)
-#184 := (~ #123 #123)
-#185 := [refl]: #184
-#173 := [nnf-pos #185]: #172
-#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)
-#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
-#186 := [mp~ #133 #173]: #128
-#674 := [mp #186 #673]: #669
-#640 := (not #669)
-#638 := (or #640 #274 #642)
-#637 := (or #274 #330)
-#641 := (or #640 #637)
-#247 := (iff #641 #638)
-#639 := (or #274 #642)
-#628 := (or #640 #639)
-#352 := (iff #628 #638)
-#353 := [rewrite]: #352
-#350 := (iff #641 #628)
-#279 := (iff #637 #639)
-#280 := [monotonicity #317]: #279
-#351 := [monotonicity #280]: #350
-#623 := [trans #351 #353]: #247
-#627 := [quant-inst]: #641
-#624 := [mp #627 #623]: #638
-#618 := [unit-resolution #624 #674 #617]: #274
-#333 := (not #274)
-#615 := (or #333 #347)
-#619 := [th-lemma]: #615
-#616 := [unit-resolution #619 #618]: #347
-[th-lemma #617 #616 #166]: false
-unsat
 f572500dcd49c30edfbceabc272b009fb11814d2 428 0
 #2 := false
 decl f4 :: (-> S2 int)
@@ -10505,9 +10079,243 @@
 #393 := [unit-resolution #403 #411 #414]: #443
 [unit-resolution #393 #424]: false
 unsat
-1b539d34eedd00cd7cacdb9d650b014baeb31cd1 409 0
-#2 := false
-#181 := -1::int
+c0906c23e4b0650863a2ef0d68bafa4e06a9dd19 106 0
+#2 := false
+decl f3 :: (-> S2 S2 S1)
+decl ?v0!0 :: S2
+#72 := ?v0!0
+#83 := (f3 ?v0!0 ?v0!0)
+decl f1 :: S1
+#4 := f1
+#75 := (= f1 #83)
+#76 := (not #75)
+#9 := (:var 0 S2)
+#16 := (f3 #9 #9)
+#50 := (= f1 #16)
+#53 := (forall (vars (?v0 S2)) #50)
+#69 := (not #53)
+#84 := (~ #69 #76)
+#85 := [sk]: #84
+#8 := (:var 1 S2)
+#12 := (f3 #9 #8)
+#41 := (= f1 #12)
+#10 := (f3 #8 #9)
+#38 := (= f1 #10)
+#44 := (and #38 #41)
+#47 := (forall (vars (?v0 S2) (?v1 S2)) #44)
+#59 := (not #47)
+#60 := (or #59 #53)
+#65 := (not #60)
+#17 := (= #16 f1)
+#18 := (forall (vars (?v0 S2)) #17)
+#13 := (= #12 f1)
+#11 := (= #10 f1)
+#14 := (and #11 #13)
+#15 := (forall (vars (?v0 S2) (?v1 S2)) #14)
+#19 := (implies #15 #18)
+#20 := (not #19)
+#66 := (iff #20 #65)
+#63 := (iff #19 #60)
+#56 := (implies #47 #53)
+#61 := (iff #56 #60)
+#62 := [rewrite]: #61
+#57 := (iff #19 #56)
+#54 := (iff #18 #53)
+#51 := (iff #17 #50)
+#52 := [rewrite]: #51
+#55 := [quant-intro #52]: #54
+#48 := (iff #15 #47)
+#45 := (iff #14 #44)
+#42 := (iff #13 #41)
+#43 := [rewrite]: #42
+#39 := (iff #11 #38)
+#40 := [rewrite]: #39
+#46 := [monotonicity #40 #43]: #45
+#49 := [quant-intro #46]: #48
+#58 := [monotonicity #49 #55]: #57
+#64 := [trans #58 #62]: #63
+#67 := [monotonicity #64]: #66
+#37 := [asserted]: #20
+#70 := [mp #37 #67]: #65
+#71 := [not-or-elim #70]: #69
+#88 := [mp~ #71 #85]: #76
+#577 := (pattern #12)
+#576 := (pattern #10)
+#87 := (not #41)
+#86 := (not #38)
+#89 := (or #86 #87)
+#90 := (not #89)
+#578 := (forall (vars (?v0 S2) (?v1 S2)) (:pat #576 #577) #90)
+#93 := (forall (vars (?v0 S2) (?v1 S2)) #90)
+#581 := (iff #93 #578)
+#579 := (iff #90 #90)
+#580 := [refl]: #579
+#582 := [quant-intro #580]: #581
+#94 := (iff #47 #93)
+#91 := (iff #44 #90)
+#92 := [rewrite]: #91
+#95 := [quant-intro #92]: #94
+#81 := (~ #47 #47)
+#79 := (~ #44 #44)
+#80 := [refl]: #79
+#82 := [nnf-pos #80]: #81
+#68 := [not-or-elim #70]: #47
+#74 := [mp~ #68 #82]: #47
+#96 := [mp #74 #95]: #93
+#583 := [mp #96 #582]: #578
+#250 := (not #578)
+#254 := (or #250 #75)
+#160 := (or #76 #76)
+#247 := (not #160)
+#233 := (or #250 #247)
+#570 := (iff #233 #254)
+#357 := (iff #254 #254)
+#564 := [rewrite]: #357
+#180 := (iff #247 #75)
+#162 := (not #76)
+#251 := (iff #162 #75)
+#252 := [rewrite]: #251
+#249 := (iff #247 #162)
+#161 := (iff #160 #76)
+#248 := [rewrite]: #161
+#240 := [monotonicity #248]: #249
+#253 := [trans #240 #252]: #180
+#227 := [monotonicity #253]: #570
+#238 := [trans #227 #564]: #570
+#568 := [quant-inst]: #233
+#239 := [mp #568 #238]: #254
+[unit-resolution #239 #583 #88]: false
+unsat
+58aaefc3b25d8181c01a2139430686aeca861bf3 66 0
+#2 := false
+decl f4 :: S1
+#10 := f4
+decl f1 :: S1
+#4 := f1
+#37 := (= f1 f4)
+decl f3 :: S1
+#8 := f3
+#34 := (= f1 f3)
+#40 := (or #34 #37)
+#84 := (iff #40 false)
+#79 := (or false false)
+#82 := (iff #79 false)
+#83 := [rewrite]: #82
+#80 := (iff #40 #79)
+#75 := (iff #37 false)
+#66 := (not #37)
+#43 := (not #34)
+#49 := (and #43 #40)
+#57 := (not #49)
+#58 := (or #37 #57)
+#63 := (not #58)
+#11 := (= f4 f1)
+#9 := (= f3 f1)
+#13 := (not #9)
+#12 := (or #9 #11)
+#14 := (and #12 #13)
+#15 := (implies #14 #11)
+#16 := (not #15)
+#64 := (iff #16 #63)
+#61 := (iff #15 #58)
+#54 := (implies #49 #37)
+#59 := (iff #54 #58)
+#60 := [rewrite]: #59
+#55 := (iff #15 #54)
+#38 := (iff #11 #37)
+#39 := [rewrite]: #38
+#52 := (iff #14 #49)
+#46 := (and #40 #43)
+#50 := (iff #46 #49)
+#51 := [rewrite]: #50
+#47 := (iff #14 #46)
+#44 := (iff #13 #43)
+#35 := (iff #9 #34)
+#36 := [rewrite]: #35
+#45 := [monotonicity #36]: #44
+#41 := (iff #12 #40)
+#42 := [monotonicity #36 #39]: #41
+#48 := [monotonicity #42 #45]: #47
+#53 := [trans #48 #51]: #52
+#56 := [monotonicity #53 #39]: #55
+#62 := [trans #56 #60]: #61
+#65 := [monotonicity #62]: #64
+#33 := [asserted]: #16
+#68 := [mp #33 #65]: #63
+#67 := [not-or-elim #68]: #66
+#76 := [iff-false #67]: #75
+#77 := (iff #34 false)
+#69 := [not-or-elim #68]: #49
+#70 := [and-elim #69]: #43
+#78 := [iff-false #70]: #77
+#81 := [monotonicity #78 #76]: #80
+#85 := [trans #81 #83]: #84
+#71 := [and-elim #69]: #40
+[mp #71 #85]: false
+unsat
+d80d7bb97197c962c0b81d44c7a8cfd617b6ca33 59 0
+#2 := false
+decl f1 :: S1
+#4 := f1
+decl f6 :: S1
+#15 := f6
+#16 := (= f6 f1)
+decl f5 :: S1
+#13 := f5
+#14 := (= f5 f1)
+#17 := (and #14 #16)
+decl f4 :: S1
+#10 := f4
+#11 := (= f4 f1)
+decl f3 :: S1
+#8 := f3
+#9 := (= f3 f1)
+#12 := (and #9 #11)
+#18 := (or #12 #17)
+#19 := (implies #18 #18)
+#20 := (not #19)
+#71 := (iff #20 false)
+#1 := true
+#66 := (not true)
+#69 := (iff #66 false)
+#70 := [rewrite]: #69
+#67 := (iff #20 #66)
+#64 := (iff #19 true)
+#50 := (= f1 f6)
+#47 := (= f1 f5)
+#53 := (and #47 #50)
+#41 := (= f1 f4)
+#38 := (= f1 f3)
+#44 := (and #38 #41)
+#56 := (or #44 #53)
+#59 := (implies #56 #56)
+#62 := (iff #59 true)
+#63 := [rewrite]: #62
+#60 := (iff #19 #59)
+#57 := (iff #18 #56)
+#54 := (iff #17 #53)
+#51 := (iff #16 #50)
+#52 := [rewrite]: #51
+#48 := (iff #14 #47)
+#49 := [rewrite]: #48
+#55 := [monotonicity #49 #52]: #54
+#45 := (iff #12 #44)
+#42 := (iff #11 #41)
+#43 := [rewrite]: #42
+#39 := (iff #9 #38)
+#40 := [rewrite]: #39
+#46 := [monotonicity #40 #43]: #45
+#58 := [monotonicity #46 #55]: #57
+#61 := [monotonicity #58 #58]: #60
+#65 := [trans #61 #63]: #64
+#68 := [monotonicity #65]: #67
+#72 := [trans #68 #70]: #71
+#37 := [asserted]: #20
+[mp #37 #72]: false
+unsat
+5682ca831360f99dba925cc3023ce41a2bdd0f63 364 0
+#2 := false
+#13 := 0::int
 decl f4 :: (-> S2 int)
 decl f3 :: (-> int S2)
 decl f6 :: S2
@@ -10517,403 +10325,358 @@
 #25 := f5
 #26 := (f4 f5)
 #29 := (+ #26 #28)
+#161 := 2::int
+#162 := (+ 2::int #29)
+#167 := (f3 #162)
+#170 := (f4 #167)
+#179 := -1::int
+#182 := (* -1::int #170)
 #30 := 1::int
-#145 := (+ 1::int #29)
-#148 := (f3 #145)
-#151 := (f4 #148)
-#676 := (* -1::int #151)
-#336 := (+ #28 #676)
-#341 := (+ #26 #336)
-#280 := (>= #341 -1::int)
-#426 := (not #280)
-#163 := (+ 1::int #28)
-#166 := (f3 #163)
-#169 := (f4 #166)
-#185 := (* -1::int #169)
-#602 := (+ #28 #185)
-#574 := (<= #602 -1::int)
-#612 := (= #602 -1::int)
-#610 := (>= #28 -1::int)
-#13 := 0::int
-#554 := (>= #28 0::int)
-#479 := (= #28 0::int)
-#274 := (f3 #28)
-#557 := (f4 #274)
-#535 := (= #557 0::int)
-#462 := (not #554)
-#458 := [hypothesis]: #462
-#536 := (or #535 #554)
+#143 := (+ 1::int #29)
+#146 := (f3 #143)
+#149 := (f4 #146)
+#183 := (+ #149 #182)
+#181 := (>= #183 0::int)
+#35 := (+ #28 1::int)
+#34 := (+ #26 1::int)
+#36 := (+ #34 #35)
+#37 := (f3 #36)
+#38 := (f4 #37)
+#31 := (+ #29 1::int)
+#32 := (f3 #31)
+#33 := (f4 #32)
+#39 := (< #33 #38)
+#40 := (not #39)
+#193 := (iff #40 #181)
+#173 := (< #149 #170)
+#176 := (not #173)
+#191 := (iff #176 #181)
+#180 := (not #181)
+#186 := (not #180)
+#189 := (iff #186 #181)
+#190 := [rewrite]: #189
+#187 := (iff #176 #186)
+#184 := (iff #173 #180)
+#185 := [rewrite]: #184
+#188 := [monotonicity #185]: #187
+#192 := [trans #188 #190]: #191
+#177 := (iff #40 #176)
+#174 := (iff #39 #173)
+#171 := (= #38 #170)
+#168 := (= #37 #167)
+#165 := (= #36 #162)
+#155 := (+ 1::int #28)
+#152 := (+ 1::int #26)
+#158 := (+ #152 #155)
+#163 := (= #158 #162)
+#164 := [rewrite]: #163
+#159 := (= #36 #158)
+#156 := (= #35 #155)
+#157 := [rewrite]: #156
+#153 := (= #34 #152)
+#154 := [rewrite]: #153
+#160 := [monotonicity #154 #157]: #159
+#166 := [trans #160 #164]: #165
+#169 := [monotonicity #166]: #168
+#172 := [monotonicity #169]: #171
+#150 := (= #33 #149)
+#147 := (= #32 #146)
+#144 := (= #31 #143)
+#145 := [rewrite]: #144
+#148 := [monotonicity #145]: #147
+#151 := [monotonicity #148]: #150
+#175 := [monotonicity #151 #172]: #174
+#178 := [monotonicity #175]: #177
+#194 := [trans #178 #192]: #193
+#142 := [asserted]: #40
+#195 := [mp #142 #194]: #181
+#637 := -2::int
+#641 := (+ #28 #182)
+#642 := (+ #26 #641)
+#517 := (<= #642 -2::int)
+#643 := (= #642 -2::int)
+#486 := (>= #29 -2::int)
+#337 := (>= #29 -1::int)
+#606 := (>= #26 0::int)
+#577 := (= #26 0::int)
+#269 := (f3 #26)
+#611 := (f4 #269)
+#600 := (= #611 0::int)
+#562 := (not #606)
+#563 := [hypothesis]: #562
+#601 := (or #600 #606)
 #14 := (:var 0 int)
 #16 := (f3 #14)
-#695 := (pattern #16)
-#85 := (>= #14 0::int)
+#691 := (pattern #16)
+#83 := (>= #14 0::int)
 #17 := (f4 #16)
 #22 := (= #17 0::int)
-#133 := (or #22 #85)
-#702 := (forall (vars (?v0 int)) (:pat #695) #133)
-#138 := (forall (vars (?v0 int)) #133)
-#705 := (iff #138 #702)
-#703 := (iff #133 #133)
-#704 := [refl]: #703
-#706 := [quant-intro #704]: #705
-#205 := (~ #138 #138)
-#217 := (~ #133 #133)
-#218 := [refl]: #217
-#206 := [nnf-pos #218]: #205
+#131 := (or #22 #83)
+#698 := (forall (vars (?v0 int)) (:pat #691) #131)
+#136 := (forall (vars (?v0 int)) #131)
+#701 := (iff #136 #698)
+#699 := (iff #131 #131)
+#700 := [refl]: #699
+#702 := [quant-intro #700]: #701
+#201 := (~ #136 #136)
+#213 := (~ #131 #131)
+#214 := [refl]: #213
+#202 := [nnf-pos #214]: #201
 #21 := (< #14 0::int)
 #23 := (implies #21 #22)
 #24 := (forall (vars (?v0 int)) #23)
-#141 := (iff #24 #138)
-#104 := (= 0::int #17)
-#110 := (not #21)
-#111 := (or #110 #104)
-#116 := (forall (vars (?v0 int)) #111)
-#139 := (iff #116 #138)
-#136 := (iff #111 #133)
-#130 := (or #85 #22)
-#134 := (iff #130 #133)
-#135 := [rewrite]: #134
-#131 := (iff #111 #130)
-#128 := (iff #104 #22)
-#129 := [rewrite]: #128
-#126 := (iff #110 #85)
-#86 := (not #85)
-#121 := (not #86)
-#124 := (iff #121 #85)
-#125 := [rewrite]: #124
-#122 := (iff #110 #121)
-#119 := (iff #21 #86)
-#120 := [rewrite]: #119
-#123 := [monotonicity #120]: #122
-#127 := [trans #123 #125]: #126
-#132 := [monotonicity #127 #129]: #131
-#137 := [trans #132 #135]: #136
-#140 := [quant-intro #137]: #139
-#117 := (iff #24 #116)
-#114 := (iff #23 #111)
-#107 := (implies #21 #104)
-#112 := (iff #107 #111)
-#113 := [rewrite]: #112
-#108 := (iff #23 #107)
-#105 := (iff #22 #104)
-#106 := [rewrite]: #105
-#109 := [monotonicity #106]: #108
-#115 := [trans #109 #113]: #114
-#118 := [quant-intro #115]: #117
-#142 := [trans #118 #140]: #141
-#103 := [asserted]: #24
-#143 := [mp #103 #142]: #138
-#219 := [mp~ #143 #206]: #138
-#707 := [mp #219 #706]: #702
-#371 := (not #702)
-#541 := (or #371 #535 #554)
-#542 := (or #371 #536)
-#543 := (iff #542 #541)
-#544 := [rewrite]: #543
-#547 := [quant-inst]: #542
-#545 := [mp #547 #544]: #541
-#463 := [unit-resolution #545 #707]: #536
-#442 := [unit-resolution #463 #458]: #535
-#449 := (= #28 #557)
-#359 := (= f6 #274)
+#139 := (iff #24 #136)
+#102 := (= 0::int #17)
+#108 := (not #21)
+#109 := (or #108 #102)
+#114 := (forall (vars (?v0 int)) #109)
+#137 := (iff #114 #136)
+#134 := (iff #109 #131)
+#128 := (or #83 #22)
+#132 := (iff #128 #131)
+#133 := [rewrite]: #132
+#129 := (iff #109 #128)
+#126 := (iff #102 #22)
+#127 := [rewrite]: #126
+#124 := (iff #108 #83)
+#84 := (not #83)
+#119 := (not #84)
+#122 := (iff #119 #83)
+#123 := [rewrite]: #122
+#120 := (iff #108 #119)
+#117 := (iff #21 #84)
+#118 := [rewrite]: #117
+#121 := [monotonicity #118]: #120
+#125 := [trans #121 #123]: #124
+#130 := [monotonicity #125 #127]: #129
+#135 := [trans #130 #133]: #134
+#138 := [quant-intro #135]: #137
+#115 := (iff #24 #114)
+#112 := (iff #23 #109)
+#105 := (implies #21 #102)
+#110 := (iff #105 #109)
+#111 := [rewrite]: #110
+#106 := (iff #23 #105)
+#103 := (iff #22 #102)
+#104 := [rewrite]: #103
+#107 := [monotonicity #104]: #106
+#113 := [trans #107 #111]: #112
+#116 := [quant-intro #113]: #115
+#140 := [trans #116 #138]: #139
+#101 := [asserted]: #24
+#141 := [mp #101 #140]: #136
+#215 := [mp~ #141 #202]: #136
+#703 := [mp #215 #702]: #698
+#651 := (not #698)
+#597 := (or #651 #600 #606)
+#604 := (or #651 #601)
+#589 := (iff #604 #597)
+#591 := [rewrite]: #589
+#588 := [quant-inst]: #604
+#592 := [mp #588 #591]: #597
+#564 := [unit-resolution #592 #703]: #601
+#565 := [unit-resolution #564 #563]: #600
+#566 := (= #26 #611)
+#354 := (= f5 #269)
 #8 := (:var 0 S2)
 #9 := (f4 #8)
-#687 := (pattern #9)
+#683 := (pattern #9)
 #10 := (f3 #9)
-#60 := (= #8 #10)
-#688 := (forall (vars (?v0 S2)) (:pat #687) #60)
-#63 := (forall (vars (?v0 S2)) #60)
-#689 := (iff #63 #688)
-#691 := (iff #688 #688)
-#692 := [rewrite]: #691
-#690 := [rewrite]: #689
-#693 := [trans #690 #692]: #689
-#213 := (~ #63 #63)
-#211 := (~ #60 #60)
-#212 := [refl]: #211
-#214 := [nnf-pos #212]: #213
+#58 := (= #8 #10)
+#684 := (forall (vars (?v0 S2)) (:pat #683) #58)
+#61 := (forall (vars (?v0 S2)) #58)
+#685 := (iff #61 #684)
+#687 := (iff #684 #684)
+#688 := [rewrite]: #687
+#686 := [rewrite]: #685
+#689 := [trans #686 #688]: #685
+#209 := (~ #61 #61)
+#207 := (~ #58 #58)
+#208 := [refl]: #207
+#210 := [nnf-pos #208]: #209
 #11 := (= #10 #8)
 #12 := (forall (vars (?v0 S2)) #11)
-#64 := (iff #12 #63)
-#61 := (iff #11 #60)
-#62 := [rewrite]: #61
-#65 := [quant-intro #62]: #64
-#59 := [asserted]: #12
-#68 := [mp #59 #65]: #63
-#202 := [mp~ #68 #214]: #63
-#694 := [mp #202 #693]: #688
-#360 := (not #688)
-#361 := (or #360 #359)
-#365 := [quant-inst]: #361
-#445 := [unit-resolution #365 #694]: #359
-#451 := [monotonicity #445]: #449
-#452 := [trans #451 #442]: #479
-#450 := (not #479)
-#453 := (or #450 #554)
-#441 := [th-lemma]: #453
-#444 := [unit-resolution #441 #458 #452]: false
-#446 := [lemma #444]: #554
-#304 := (or #462 #610)
-#447 := [th-lemma]: #304
-#443 := [unit-resolution #447 #446]: #610
-#605 := (not #610)
-#593 := (or #605 #612)
-#67 := (= #14 #17)
-#92 := (or #67 #86)
-#696 := (forall (vars (?v0 int)) (:pat #695) #92)
-#97 := (forall (vars (?v0 int)) #92)
-#699 := (iff #97 #696)
-#697 := (iff #92 #92)
-#698 := [refl]: #697
-#700 := [quant-intro #698]: #699
-#203 := (~ #97 #97)
-#200 := (~ #92 #92)
-#215 := [refl]: #200
-#204 := [nnf-pos #215]: #203
+#62 := (iff #12 #61)
+#59 := (iff #11 #58)
+#60 := [rewrite]: #59
+#63 := [quant-intro #60]: #62
+#57 := [asserted]: #12
+#66 := [mp #57 #63]: #61
+#198 := [mp~ #66 #210]: #61
+#690 := [mp #198 #689]: #684
+#356 := (not #684)
+#347 := (or #356 #354)
+#358 := [quant-inst]: #347
+#560 := [unit-resolution #358 #690]: #354
+#552 := [monotonicity #560]: #566
+#548 := [trans #552 #565]: #577
+#549 := (not #577)
+#550 := (or #549 #606)
+#553 := [th-lemma]: #550
+#554 := [unit-resolution #553 #563 #548]: false
+#555 := [lemma #554]: #606
+#602 := (>= #28 0::int)
+#578 := (= #28 0::int)
+#270 := (f3 #28)
+#603 := (f4 #270)
+#593 := (= #603 0::int)
+#551 := (not #602)
+#557 := [hypothesis]: #551
+#594 := (or #593 #602)
+#581 := (or #651 #593 #602)
+#582 := (or #651 #594)
+#585 := (iff #582 #581)
+#586 := [rewrite]: #585
+#584 := [quant-inst]: #582
+#583 := [mp #584 #586]: #581
+#530 := [unit-resolution #583 #703]: #594
+#531 := [unit-resolution #530 #557]: #593
+#425 := (= #28 #603)
+#355 := (= f6 #270)
+#357 := (or #356 #355)
+#361 := [quant-inst]: #357
+#532 := [unit-resolution #361 #690]: #355
+#536 := [monotonicity #532]: #425
+#537 := [trans #536 #531]: #578
+#538 := (not #578)
+#543 := (or #538 #602)
+#539 := [th-lemma]: #543
+#540 := [unit-resolution #539 #557 #537]: false
+#541 := [lemma #540]: #602
+#556 := (or #551 #562 #337)
+#544 := [th-lemma]: #556
+#545 := [unit-resolution #544 #541 #555]: #337
+#674 := (not #337)
+#546 := (or #674 #486)
+#542 := [th-lemma]: #546
+#547 := [unit-resolution #542 #545]: #486
+#630 := (not #486)
+#627 := (or #630 #643)
+#65 := (= #14 #17)
+#90 := (or #65 #84)
+#692 := (forall (vars (?v0 int)) (:pat #691) #90)
+#95 := (forall (vars (?v0 int)) #90)
+#695 := (iff #95 #692)
+#693 := (iff #90 #90)
+#694 := [refl]: #693
+#696 := [quant-intro #694]: #695
+#199 := (~ #95 #95)
+#196 := (~ #90 #90)
+#211 := [refl]: #196
+#200 := [nnf-pos #211]: #199
 #18 := (= #17 #14)
 #15 := (<= 0::int #14)
 #19 := (implies #15 #18)
 #20 := (forall (vars (?v0 int)) #19)
-#100 := (iff #20 #97)
-#74 := (not #15)
-#75 := (or #74 #67)
-#80 := (forall (vars (?v0 int)) #75)
-#98 := (iff #80 #97)
-#95 := (iff #75 #92)
-#89 := (or #86 #67)
-#93 := (iff #89 #92)
-#94 := [rewrite]: #93
-#90 := (iff #75 #89)
-#87 := (iff #74 #86)
-#83 := (iff #15 #85)
-#84 := [rewrite]: #83
-#88 := [monotonicity #84]: #87
-#91 := [monotonicity #88]: #90
-#96 := [trans #91 #94]: #95
-#99 := [quant-intro #96]: #98
-#81 := (iff #20 #80)
-#78 := (iff #19 #75)
-#71 := (implies #15 #67)
-#76 := (iff #71 #75)
-#77 := [rewrite]: #76
-#72 := (iff #19 #71)
-#69 := (iff #18 #67)
-#70 := [rewrite]: #69
-#73 := [monotonicity #70]: #72
-#79 := [trans #73 #77]: #78
-#82 := [quant-intro #79]: #81
-#101 := [trans #82 #99]: #100
-#66 := [asserted]: #20
-#102 := [mp #66 #101]: #97
-#216 := [mp~ #102 #204]: #97
-#701 := [mp #216 #700]: #696
-#670 := (not #696)
-#594 := (or #670 #605 #612)
-#609 := (>= #163 0::int)
-#611 := (not #609)
-#603 := (= #163 #169)
-#600 := (or #603 #611)
-#599 := (or #670 #600)
-#591 := (iff #599 #594)
-#586 := (or #670 #593)
-#590 := (iff #586 #594)
-#587 := [rewrite]: #590
-#588 := (iff #599 #586)
-#597 := (iff #600 #593)
-#601 := (or #612 #605)
-#595 := (iff #601 #593)
-#596 := [rewrite]: #595
-#608 := (iff #600 #601)
-#606 := (iff #611 #605)
-#615 := (iff #609 #610)
-#604 := [rewrite]: #615
-#607 := [monotonicity #604]: #606
-#613 := (iff #603 #612)
-#614 := [rewrite]: #613
-#592 := [monotonicity #614 #607]: #608
-#598 := [trans #592 #596]: #597
-#589 := [monotonicity #598]: #588
-#571 := [trans #589 #587]: #591
-#585 := [quant-inst]: #599
-#572 := [mp #585 #571]: #594
-#448 := [unit-resolution #572 #701]: #593
-#438 := [unit-resolution #448 #443]: #612
-#428 := (not #612)
-#430 := (or #428 #574)
-#431 := [th-lemma]: #430
-#434 := [unit-resolution #431 #438]: #574
-#154 := (+ 1::int #26)
-#157 := (f3 #154)
-#160 := (f4 #157)
-#184 := (* -1::int #160)
-#488 := (+ #26 #184)
-#512 := (<= #488 -1::int)
-#489 := (= #488 -1::int)
-#633 := (>= #26 -1::int)
-#570 := (>= #26 0::int)
-#478 := (= #26 0::int)
-#273 := (f3 #26)
-#556 := (f4 #273)
-#552 := (= #556 0::int)
-#480 := (not #570)
-#481 := [hypothesis]: #480
-#553 := (or #552 #570)
-#558 := (or #371 #552 #570)
-#559 := (or #371 #553)
-#555 := (iff #559 #558)
-#561 := [rewrite]: #555
-#560 := [quant-inst]: #559
-#534 := [mp #560 #561]: #558
-#482 := [unit-resolution #534 #707]: #553
-#483 := [unit-resolution #482 #481]: #552
-#484 := (= #26 #556)
-#358 := (= f5 #273)
-#351 := (or #360 #358)
-#362 := [quant-inst]: #351
-#466 := [unit-resolution #362 #694]: #358
-#454 := [monotonicity #466]: #484
-#455 := [trans #454 #483]: #478
-#456 := (not #478)
-#457 := (or #456 #570)
-#459 := [th-lemma]: #457
-#460 := [unit-resolution #459 #481 #455]: false
-#461 := [lemma #460]: #570
-#435 := (or #480 #633)
-#439 := [th-lemma]: #435
-#432 := [unit-resolution #439 #461]: #633
-#629 := (not #633)
-#637 := (or #489 #629)
-#467 := (or #670 #489 #629)
-#645 := (>= #154 0::int)
-#646 := (not #645)
-#641 := (= #154 #160)
-#647 := (or #641 #646)
-#469 := (or #670 #647)
-#628 := (iff #469 #467)
-#623 := (or #670 #637)
-#627 := (iff #623 #467)
-#625 := [rewrite]: #627
-#624 := (iff #469 #623)
-#631 := (iff #647 #637)
-#630 := (iff #646 #629)
-#634 := (iff #645 #633)
-#635 := [rewrite]: #634
-#636 := [monotonicity #635]: #630
-#490 := (iff #641 #489)
-#632 := [rewrite]: #490
-#638 := [monotonicity #632 #636]: #631
-#626 := [monotonicity #638]: #624
-#510 := [trans #626 #625]: #628
-#470 := [quant-inst]: #469
-#511 := [mp #470 #510]: #467
-#440 := [unit-resolution #511 #701]: #637
-#433 := [unit-resolution #440 #432]: #489
-#436 := (not #489)
-#437 := (or #436 #512)
-#423 := [th-lemma]: #437
-#425 := [unit-resolution #423 #433]: #512
-#186 := (+ #184 #185)
-#187 := (+ #151 #186)
-#183 := (>= #187 0::int)
-#37 := (+ #28 1::int)
-#38 := (f3 #37)
-#39 := (f4 #38)
-#34 := (+ #26 1::int)
-#35 := (f3 #34)
-#36 := (f4 #35)
-#40 := (+ #36 #39)
-#31 := (+ #29 1::int)
-#32 := (f3 #31)
-#33 := (f4 #32)
-#41 := (< #33 #40)
-#42 := (not #41)
-#197 := (iff #42 #183)
-#172 := (+ #160 #169)
-#175 := (< #151 #172)
-#178 := (not #175)
-#195 := (iff #178 #183)
-#182 := (not #183)
-#190 := (not #182)
-#193 := (iff #190 #183)
-#194 := [rewrite]: #193
-#191 := (iff #178 #190)
-#188 := (iff #175 #182)
-#189 := [rewrite]: #188
-#192 := [monotonicity #189]: #191
-#196 := [trans #192 #194]: #195
-#179 := (iff #42 #178)
-#176 := (iff #41 #175)
-#173 := (= #40 #172)
-#170 := (= #39 #169)
-#167 := (= #38 #166)
-#164 := (= #37 #163)
-#165 := [rewrite]: #164
-#168 := [monotonicity #165]: #167
-#171 := [monotonicity #168]: #170
-#161 := (= #36 #160)
-#158 := (= #35 #157)
-#155 := (= #34 #154)
-#156 := [rewrite]: #155
-#159 := [monotonicity #156]: #158
-#162 := [monotonicity #159]: #161
-#174 := [monotonicity #162 #171]: #173
-#152 := (= #33 #151)
-#149 := (= #32 #148)
-#146 := (= #31 #145)
-#147 := [rewrite]: #146
-#150 := [monotonicity #147]: #149
-#153 := [monotonicity #150]: #152
-#177 := [monotonicity #153 #174]: #176
-#180 := [monotonicity #177]: #179
-#198 := [trans #180 #196]: #197
-#144 := [asserted]: #42
-#199 := [mp #144 #198]: #183
-#427 := (not #574)
-#424 := (not #512)
-#409 := (or #426 #182 #424 #427)
-#411 := [th-lemma]: #409
-#412 := [unit-resolution #411 #199 #425 #434]: #426
-#677 := (= #341 -1::int)
-#321 := (>= #29 -1::int)
-#413 := (or #321 #629 #462)
-#414 := [th-lemma]: #413
-#415 := [unit-resolution #414 #432 #446]: #321
-#326 := (not #321)
-#667 := (or #326 #677)
-#672 := (or #670 #326 #677)
-#682 := (>= #145 0::int)
-#683 := (not #682)
-#680 := (= #145 #151)
-#686 := (or #680 #683)
-#312 := (or #670 #686)
-#383 := (iff #312 #672)
-#673 := (or #670 #667)
-#660 := (iff #673 #672)
+#98 := (iff #20 #95)
+#72 := (not #15)
+#73 := (or #72 #65)
+#78 := (forall (vars (?v0 int)) #73)
+#96 := (iff #78 #95)
+#93 := (iff #73 #90)
+#87 := (or #84 #65)
+#91 := (iff #87 #90)
+#92 := [rewrite]: #91
+#88 := (iff #73 #87)
+#85 := (iff #72 #84)
+#81 := (iff #15 #83)
+#82 := [rewrite]: #81
+#86 := [monotonicity #82]: #85
+#89 := [monotonicity #86]: #88
+#94 := [trans #89 #92]: #93
+#97 := [quant-intro #94]: #96
+#79 := (iff #20 #78)
+#76 := (iff #19 #73)
+#69 := (implies #15 #65)
+#74 := (iff #69 #73)
+#75 := [rewrite]: #74
+#70 := (iff #19 #69)
+#67 := (iff #18 #65)
+#68 := [rewrite]: #67
+#71 := [monotonicity #68]: #70
+#77 := [trans #71 #75]: #76
+#80 := [quant-intro #77]: #79
+#99 := [trans #80 #97]: #98
+#64 := [asserted]: #20
+#100 := [mp #64 #99]: #95
+#212 := [mp~ #100 #200]: #95
+#697 := [mp #212 #696]: #692
+#664 := (not #692)
+#619 := (or #664 #630 #643)
+#639 := (>= #162 0::int)
+#640 := (not #639)
+#635 := (= #162 #170)
+#636 := (or #635 #640)
+#620 := (or #664 #636)
+#508 := (iff #620 #619)
+#623 := (or #664 #627)
+#506 := (iff #623 #619)
+#507 := [rewrite]: #506
+#621 := (iff #620 #623)
+#465 := (iff #636 #627)
+#626 := (or #643 #630)
+#634 := (iff #626 #627)
+#463 := [rewrite]: #634
+#632 := (iff #636 #626)
+#631 := (iff #640 #630)
+#628 := (iff #639 #486)
+#629 := [rewrite]: #628
+#625 := [monotonicity #629]: #631
+#484 := (iff #635 #643)
+#485 := [rewrite]: #484
+#633 := [monotonicity #485 #625]: #632
+#466 := [trans #633 #463]: #465
+#624 := [monotonicity #466]: #621
+#467 := [trans #624 #507]: #508
+#622 := [quant-inst]: #620
+#615 := [mp #622 #467]: #619
+#527 := [unit-resolution #615 #697]: #627
+#473 := [unit-resolution #527 #547]: #643
+#528 := (not #643)
+#509 := (or #528 #517)
+#533 := [th-lemma]: #509
+#534 := [unit-resolution #533 #473]: #517
+#680 := (* -1::int #149)
+#681 := (+ #28 #680)
+#676 := (+ #26 #681)
+#379 := (>= #676 -1::int)
+#682 := (= #676 -1::int)
+#324 := (or #674 #682)
+#659 := (or #664 #674 #682)
+#464 := (>= #143 0::int)
+#671 := (not #464)
+#678 := (= #143 #149)
+#679 := (or #678 #671)
+#665 := (or #664 #679)
+#667 := (iff #665 #659)
+#666 := (or #664 #324)
+#309 := (iff #666 #659)
+#669 := [rewrite]: #309
+#668 := (iff #665 #666)
+#662 := (iff #679 #324)
+#321 := (or #682 #674)
+#660 := (iff #321 #324)
 #661 := [rewrite]: #660
-#671 := (iff #312 #673)
-#669 := (iff #686 #667)
-#664 := (or #677 #326)
-#668 := (iff #664 #667)
-#663 := [rewrite]: #668
-#665 := (iff #686 #664)
-#327 := (iff #683 #326)
-#662 := (iff #682 #321)
-#325 := [rewrite]: #662
-#328 := [monotonicity #325]: #327
-#337 := (iff #680 #677)
-#678 := [rewrite]: #337
-#666 := [monotonicity #678 #328]: #665
-#307 := [trans #666 #663]: #669
-#674 := [monotonicity #307]: #671
-#384 := [trans #674 #661]: #383
-#313 := [quant-inst]: #312
-#385 := [mp #313 #384]: #672
-#416 := [unit-resolution #385 #701]: #667
-#417 := [unit-resolution #416 #415]: #677
-#418 := (not #677)
-#419 := (or #418 #280)
-#420 := [th-lemma]: #419
-[unit-resolution #420 #417 #412]: false
+#322 := (iff #679 #321)
+#317 := (iff #671 #674)
+#673 := (iff #464 #337)
+#333 := [rewrite]: #673
+#658 := [monotonicity #333]: #317
+#672 := (iff #678 #682)
+#332 := [rewrite]: #672
+#323 := [monotonicity #332 #658]: #322
+#663 := [trans #323 #661]: #662
+#308 := [monotonicity #663]: #668
+#670 := [trans #308 #669]: #667
+#303 := [quant-inst]: #665
+#656 := [mp #303 #670]: #659
+#529 := [unit-resolution #656 #697]: #324
+#535 := [unit-resolution #529 #545]: #682
+#503 := (not #682)
+#510 := (or #503 #379)
+#469 := [th-lemma]: #510
+#500 := [unit-resolution #469 #535]: #379
+[th-lemma #500 #534 #195]: false
 unsat
 b265c055659e0ea1b48e0304606789db36e05540 60 0
 #2 := false
@@ -11075,6 +10838,67 @@
 #136 := [asserted]: #34
 [mp #136 #154]: false
 unsat
+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
 b5b340366dfb8289e173cd594b95953e97e705ac 42 0
 #2 := false
 decl f3 :: (-> int S2)
@@ -11118,6 +10942,186 @@
 #140 := [asserted]: #38
 [mp #140 #158]: false
 unsat
+2a686ca258d75e408aba2310eab7bfde718856c5 179 0
+#2 := false
+decl f3 :: (-> int S2)
+decl f4 :: (-> S2 int)
+decl f6 :: S2
+#27 := f6
+#28 := (f4 f6)
+#138 := -1::int
+#139 := (* -1::int #28)
+decl f5 :: S2
+#25 := f5
+#26 := (f4 f5)
+#140 := (+ #26 #139)
+#143 := (f3 #140)
+#13 := 0::int
+#32 := (f3 0::int)
+#149 := (= #32 #143)
+#649 := (f4 #143)
+#597 := (f3 #649)
+#596 := (= #597 #143)
+#499 := (= #143 #597)
+#8 := (:var 0 S2)
+#9 := (f4 #8)
+#665 := (pattern #9)
+#10 := (f3 #9)
+#53 := (= #8 #10)
+#666 := (forall (vars (?v0 S2)) (:pat #665) #53)
+#56 := (forall (vars (?v0 S2)) #53)
+#667 := (iff #56 #666)
+#669 := (iff #666 #666)
+#670 := [rewrite]: #669
+#668 := [rewrite]: #667
+#671 := [trans #668 #670]: #667
+#192 := (~ #56 #56)
+#190 := (~ #53 #53)
+#191 := [refl]: #190
+#193 := [nnf-pos #191]: #192
+#11 := (= #10 #8)
+#12 := (forall (vars (?v0 S2)) #11)
+#57 := (iff #12 #56)
+#54 := (iff #11 #53)
+#55 := [rewrite]: #54
+#58 := [quant-intro #55]: #57
+#52 := [asserted]: #12
+#61 := [mp #52 #58]: #56
+#179 := [mp~ #61 #193]: #56
+#672 := [mp #179 #671]: #666
+#338 := (not #666)
+#502 := (or #338 #499)
+#503 := [quant-inst]: #502
+#501 := [unit-resolution #503 #672]: #499
+#600 := [symm #501]: #596
+#506 := (= #32 #597)
+#504 := (= 0::int #649)
+#622 := (= #649 0::int)
+#161 := (>= #140 0::int)
+#160 := (not #161)
+#154 := (not #149)
+#167 := (and #154 #160)
+#30 := (- #26 #28)
+#31 := (f3 #30)
+#33 := (= #31 #32)
+#34 := (not #33)
+#29 := (< #26 #28)
+#35 := (and #29 #34)
+#172 := (iff #35 #167)
+#157 := (and #29 #154)
+#170 := (iff #157 #167)
+#164 := (and #160 #154)
+#168 := (iff #164 #167)
+#169 := [rewrite]: #168
+#165 := (iff #157 #164)
+#162 := (iff #29 #160)
+#163 := [rewrite]: #162
+#166 := [monotonicity #163]: #165
+#171 := [trans #166 #169]: #170
+#158 := (iff #35 #157)
+#155 := (iff #34 #154)
+#152 := (iff #33 #149)
+#146 := (= #143 #32)
+#150 := (iff #146 #149)
+#151 := [rewrite]: #150
+#147 := (iff #33 #146)
+#144 := (= #31 #143)
+#141 := (= #30 #140)
+#142 := [rewrite]: #141
+#145 := [monotonicity #142]: #144
+#148 := [monotonicity #145]: #147
+#153 := [trans #148 #151]: #152
+#156 := [monotonicity #153]: #155
+#159 := [monotonicity #156]: #158
+#173 := [trans #159 #171]: #172
+#137 := [asserted]: #35
+#174 := [mp #137 #173]: #167
+#176 := [and-elim #174]: #160
+#14 := (:var 0 int)
+#16 := (f3 #14)
+#673 := (pattern #16)
+#78 := (>= #14 0::int)
+#17 := (f4 #16)
+#22 := (= #17 0::int)
+#126 := (or #22 #78)
+#680 := (forall (vars (?v0 int)) (:pat #673) #126)
+#131 := (forall (vars (?v0 int)) #126)
+#683 := (iff #131 #680)
+#681 := (iff #126 #126)
+#682 := [refl]: #681
+#684 := [quant-intro #682]: #683
+#182 := (~ #131 #131)
+#196 := (~ #126 #126)
+#197 := [refl]: #196
+#183 := [nnf-pos #197]: #182
+#21 := (< #14 0::int)
+#23 := (implies #21 #22)
+#24 := (forall (vars (?v0 int)) #23)
+#134 := (iff #24 #131)
+#97 := (= 0::int #17)
+#103 := (not #21)
+#104 := (or #103 #97)
+#109 := (forall (vars (?v0 int)) #104)
+#132 := (iff #109 #131)
+#129 := (iff #104 #126)
+#123 := (or #78 #22)
+#127 := (iff #123 #126)
+#128 := [rewrite]: #127
+#124 := (iff #104 #123)
+#121 := (iff #97 #22)
+#122 := [rewrite]: #121
+#119 := (iff #103 #78)
+#79 := (not #78)
+#114 := (not #79)
+#117 := (iff #114 #78)
+#118 := [rewrite]: #117
+#115 := (iff #103 #114)
+#112 := (iff #21 #79)
+#113 := [rewrite]: #112
+#116 := [monotonicity #113]: #115
+#120 := [trans #116 #118]: #119
+#125 := [monotonicity #120 #122]: #124
+#130 := [trans #125 #128]: #129
+#133 := [quant-intro #130]: #132
+#110 := (iff #24 #109)
+#107 := (iff #23 #104)
+#100 := (implies #21 #97)
+#105 := (iff #100 #104)
+#106 := [rewrite]: #105
+#101 := (iff #23 #100)
+#98 := (iff #22 #97)
+#99 := [rewrite]: #98
+#102 := [monotonicity #99]: #101
+#108 := [trans #102 #106]: #107
+#111 := [quant-intro #108]: #110
+#135 := [trans #111 #133]: #134
+#96 := [asserted]: #24
+#136 := [mp #96 #135]: #131
+#198 := [mp~ #136 #183]: #131
+#685 := [mp #198 #684]: #680
+#619 := (not #680)
+#625 := (or #619 #161 #622)
+#617 := (or #622 #161)
+#466 := (or #619 #617)
+#607 := (iff #466 #625)
+#618 := (or #161 #622)
+#468 := (or #619 #618)
+#612 := (iff #468 #625)
+#613 := [rewrite]: #612
+#610 := (iff #466 #468)
+#623 := (iff #617 #618)
+#624 := [rewrite]: #623
+#611 := [monotonicity #624]: #610
+#608 := [trans #611 #613]: #607
+#467 := [quant-inst]: #466
+#614 := [mp #467 #608]: #625
+#494 := [unit-resolution #614 #685 #176]: #622
+#505 := [symm #494]: #504
+#599 := [monotonicity #505]: #506
+#587 := [trans #599 #600]: #149
+#175 := [and-elim #174]: #154
+[unit-resolution #175 #587]: false
+unsat
 3c1f7d30ea379c9ce2207a011fc3d11958eb42f1 438 0
 #2 := false
 #13 := 0::int
@@ -11557,247 +11561,6 @@
 #404 := [unit-resolution #379 #434]: #485
 [unit-resolution #404 #396]: false
 unsat
-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)
-decl f6 :: S2
-#27 := f6
-#28 := (f4 f6)
-#138 := -1::int
-#139 := (* -1::int #28)
-decl f5 :: S2
-#25 := f5
-#26 := (f4 f5)
-#140 := (+ #26 #139)
-#143 := (f3 #140)
-#13 := 0::int
-#32 := (f3 0::int)
-#149 := (= #32 #143)
-#649 := (f4 #143)
-#597 := (f3 #649)
-#596 := (= #597 #143)
-#499 := (= #143 #597)
-#8 := (:var 0 S2)
-#9 := (f4 #8)
-#665 := (pattern #9)
-#10 := (f3 #9)
-#53 := (= #8 #10)
-#666 := (forall (vars (?v0 S2)) (:pat #665) #53)
-#56 := (forall (vars (?v0 S2)) #53)
-#667 := (iff #56 #666)
-#669 := (iff #666 #666)
-#670 := [rewrite]: #669
-#668 := [rewrite]: #667
-#671 := [trans #668 #670]: #667
-#192 := (~ #56 #56)
-#190 := (~ #53 #53)
-#191 := [refl]: #190
-#193 := [nnf-pos #191]: #192
-#11 := (= #10 #8)
-#12 := (forall (vars (?v0 S2)) #11)
-#57 := (iff #12 #56)
-#54 := (iff #11 #53)
-#55 := [rewrite]: #54
-#58 := [quant-intro #55]: #57
-#52 := [asserted]: #12
-#61 := [mp #52 #58]: #56
-#179 := [mp~ #61 #193]: #56
-#672 := [mp #179 #671]: #666
-#338 := (not #666)
-#502 := (or #338 #499)
-#503 := [quant-inst]: #502
-#501 := [unit-resolution #503 #672]: #499
-#600 := [symm #501]: #596
-#506 := (= #32 #597)
-#504 := (= 0::int #649)
-#622 := (= #649 0::int)
-#161 := (>= #140 0::int)
-#160 := (not #161)
-#154 := (not #149)
-#167 := (and #154 #160)
-#30 := (- #26 #28)
-#31 := (f3 #30)
-#33 := (= #31 #32)
-#34 := (not #33)
-#29 := (< #26 #28)
-#35 := (and #29 #34)
-#172 := (iff #35 #167)
-#157 := (and #29 #154)
-#170 := (iff #157 #167)
-#164 := (and #160 #154)
-#168 := (iff #164 #167)
-#169 := [rewrite]: #168
-#165 := (iff #157 #164)
-#162 := (iff #29 #160)
-#163 := [rewrite]: #162
-#166 := [monotonicity #163]: #165
-#171 := [trans #166 #169]: #170
-#158 := (iff #35 #157)
-#155 := (iff #34 #154)
-#152 := (iff #33 #149)
-#146 := (= #143 #32)
-#150 := (iff #146 #149)
-#151 := [rewrite]: #150
-#147 := (iff #33 #146)
-#144 := (= #31 #143)
-#141 := (= #30 #140)
-#142 := [rewrite]: #141
-#145 := [monotonicity #142]: #144
-#148 := [monotonicity #145]: #147
-#153 := [trans #148 #151]: #152
-#156 := [monotonicity #153]: #155
-#159 := [monotonicity #156]: #158
-#173 := [trans #159 #171]: #172
-#137 := [asserted]: #35
-#174 := [mp #137 #173]: #167
-#176 := [and-elim #174]: #160
-#14 := (:var 0 int)
-#16 := (f3 #14)
-#673 := (pattern #16)
-#78 := (>= #14 0::int)
-#17 := (f4 #16)
-#22 := (= #17 0::int)
-#126 := (or #22 #78)
-#680 := (forall (vars (?v0 int)) (:pat #673) #126)
-#131 := (forall (vars (?v0 int)) #126)
-#683 := (iff #131 #680)
-#681 := (iff #126 #126)
-#682 := [refl]: #681
-#684 := [quant-intro #682]: #683
-#182 := (~ #131 #131)
-#196 := (~ #126 #126)
-#197 := [refl]: #196
-#183 := [nnf-pos #197]: #182
-#21 := (< #14 0::int)
-#23 := (implies #21 #22)
-#24 := (forall (vars (?v0 int)) #23)
-#134 := (iff #24 #131)
-#97 := (= 0::int #17)
-#103 := (not #21)
-#104 := (or #103 #97)
-#109 := (forall (vars (?v0 int)) #104)
-#132 := (iff #109 #131)
-#129 := (iff #104 #126)
-#123 := (or #78 #22)
-#127 := (iff #123 #126)
-#128 := [rewrite]: #127
-#124 := (iff #104 #123)
-#121 := (iff #97 #22)
-#122 := [rewrite]: #121
-#119 := (iff #103 #78)
-#79 := (not #78)
-#114 := (not #79)
-#117 := (iff #114 #78)
-#118 := [rewrite]: #117
-#115 := (iff #103 #114)
-#112 := (iff #21 #79)
-#113 := [rewrite]: #112
-#116 := [monotonicity #113]: #115
-#120 := [trans #116 #118]: #119
-#125 := [monotonicity #120 #122]: #124
-#130 := [trans #125 #128]: #129
-#133 := [quant-intro #130]: #132
-#110 := (iff #24 #109)
-#107 := (iff #23 #104)
-#100 := (implies #21 #97)
-#105 := (iff #100 #104)
-#106 := [rewrite]: #105
-#101 := (iff #23 #100)
-#98 := (iff #22 #97)
-#99 := [rewrite]: #98
-#102 := [monotonicity #99]: #101
-#108 := [trans #102 #106]: #107
-#111 := [quant-intro #108]: #110
-#135 := [trans #111 #133]: #134
-#96 := [asserted]: #24
-#136 := [mp #96 #135]: #131
-#198 := [mp~ #136 #183]: #131
-#685 := [mp #198 #684]: #680
-#619 := (not #680)
-#625 := (or #619 #161 #622)
-#617 := (or #622 #161)
-#466 := (or #619 #617)
-#607 := (iff #466 #625)
-#618 := (or #161 #622)
-#468 := (or #619 #618)
-#612 := (iff #468 #625)
-#613 := [rewrite]: #612
-#610 := (iff #466 #468)
-#623 := (iff #617 #618)
-#624 := [rewrite]: #623
-#611 := [monotonicity #624]: #610
-#608 := [trans #611 #613]: #607
-#467 := [quant-inst]: #466
-#614 := [mp #467 #608]: #625
-#494 := [unit-resolution #614 #685 #176]: #622
-#505 := [symm #494]: #504
-#599 := [monotonicity #505]: #506
-#587 := [trans #599 #600]: #149
-#175 := [and-elim #174]: #154
-[unit-resolution #175 #587]: false
-unsat
 d2811f187a26b77cca78876fe4e7b2248df3c75e 321 0
 #2 := false
 decl f3 :: (-> int S2)
@@ -12450,6 +12213,40 @@
 #500 := [trans #499 #542]: #188
 [unit-resolution #198 #500]: false
 unsat
+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
 38a43c830e9f2f628495aa4330122f2c2da981cd 517 0
 #2 := false
 #13 := 0::int
@@ -12968,40 +12765,6 @@
 #282 := [unit-resolution #419 #286]: #319
 [unit-resolution #424 #282 #300]: false
 unsat
-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)
@@ -13711,6 +13474,281 @@
 #271 := [asserted]: #59
 [unit-resolution #271 #738]: false
 unsat
+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
 426e56a2fa6cae689a233b52c380ff3e1f4c5d0f 287 0
 #2 := false
 decl f5 :: (-> int S2)
@@ -13999,6 +14037,328 @@
 #270 := [asserted]: #58
 [unit-resolution #270 #691]: 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
 f595f2976198bb59087cfc3afdc24567f743886b 287 0
 #2 := false
 decl f5 :: (-> int S2)
@@ -14586,6 +14946,311 @@
 #271 := [asserted]: #59
 [unit-resolution #271 #706]: 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
 4183c3ff4a9f3923626099be345775dc19db2d14 360 0
 #2 := false
 decl f5 :: (-> int S2)
@@ -14947,6 +15612,309 @@
 #557 := [trans #620 #572]: #272
 [unit-resolution #280 #557]: 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
 f2f039ccfb9f8bb6046ed67b2d5f9e4965b93c1c 288 0
 #2 := false
 decl f5 :: (-> int S2)
@@ -15236,6 +16204,320 @@
 #270 := [asserted]: #58
 [unit-resolution #270 #692]: 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
 ed1dcbda8b48a8192bb34e0311b0082775d33ed7 302 0
 #2 := false
 decl f5 :: (-> int S2)
@@ -15539,6 +16821,454 @@
 #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
 8f1427496ff61632aac173e792df93e6d2f8fec2 292 0
 #2 := false
 decl f5 :: (-> int S2)
@@ -16371,6 +18101,312 @@
 #593 := [trans #591 #608]: #781
 [unit-resolution #661 #593]: 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
 b68469389563a7c2adfa58d6c3a299893ac9030d 605 0
 #2 := false
 decl f5 :: (-> int S2)
@@ -16977,2279 +19013,6 @@
 #603 := [unit-resolution #535 #671]: #61
 [unit-resolution #603 #596]: false
 unsat
-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)
@@ -19579,1516 +19342,6 @@
 #271 := [asserted]: #59
 [unit-resolution #271 #708]: false
 unsat
-30266f09c986583d8c5407f4e044da694063d38f 303 0
-#2 := false
-decl f5 :: (-> int S2)
-#11 := 0::int
-#56 := (f5 0::int)
-decl f4 :: (-> int int int)
-#53 := 3::int
-#54 := (f4 3::int 3::int)
-#55 := (f5 #54)
-#57 := (= #55 #56)
-#832 := (= #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
-#777 := (not #852)
-#662 := (or #777 #832)
-#430 := (* -1::int 3::int)
-#515 := (mod #430 #430)
-#516 := (+ #54 #515)
-#517 := (= #516 0::int)
-#508 := (mod 3::int 3::int)
-#519 := (* -1::int #508)
-#520 := (+ #54 #519)
-#448 := (= #520 0::int)
-#521 := (<= 3::int 0::int)
-#518 := (or #521 #521)
-#522 := (not #518)
-#501 := (>= 3::int 0::int)
-#836 := (or #521 #501)
-#838 := (not #836)
-#495 := (or #838 #522)
-#625 := (ite #495 #448 #517)
-#506 := (= 3::int 0::int)
-#507 := (ite #506 #832 #625)
-#839 := (= 3::int #54)
-#840 := (ite #506 #839 #507)
-#677 := (or #777 #840)
-#680 := (iff #677 #662)
-#681 := (iff #662 #662)
-#682 := [rewrite]: #681
-#773 := (iff #840 #832)
-#843 := (= #54 3::int)
-#668 := (ite false #843 #832)
-#776 := (iff #668 #832)
-#678 := [rewrite]: #776
-#669 := (iff #840 #668)
-#785 := (iff #507 #832)
-#780 := (ite false #832 #832)
-#784 := (iff #780 #832)
-#782 := [rewrite]: #784
-#781 := (iff #507 #780)
-#626 := (iff #625 #832)
-#1 := true
-#793 := (ite true #832 #832)
-#795 := (iff #793 #832)
-#624 := [rewrite]: #795
-#794 := (iff #625 #793)
-#786 := (iff #517 #832)
-#791 := (= #516 #54)
-#807 := (+ #54 0::int)
-#805 := (= #807 #54)
-#809 := [rewrite]: #805
-#789 := (= #516 #807)
-#646 := (= #515 0::int)
-#801 := -3::int
-#802 := (mod -3::int -3::int)
-#804 := (= #802 0::int)
-#645 := [rewrite]: #804
-#803 := (= #515 #802)
-#796 := (= #430 -3::int)
-#797 := [rewrite]: #796
-#798 := [monotonicity #797 #797]: #803
-#647 := [trans #798 #645]: #646
-#790 := [monotonicity #647]: #789
-#792 := [trans #790 #809]: #791
-#787 := [monotonicity #792]: #786
-#799 := (iff #448 #832)
-#806 := (= #520 #54)
-#808 := (= #520 #807)
-#528 := (= #519 0::int)
-#815 := (* -1::int 0::int)
-#526 := (= #815 0::int)
-#527 := [rewrite]: #526
-#812 := (= #519 #815)
-#537 := (= #508 0::int)
-#811 := [rewrite]: #537
-#816 := [monotonicity #811]: #812
-#529 := [trans #816 #527]: #528
-#523 := [monotonicity #529]: #808
-#810 := [trans #523 #809]: #806
-#800 := [monotonicity #810]: #799
-#814 := (iff #495 true)
-#835 := (or false true)
-#482 := (iff #835 true)
-#483 := [rewrite]: #482
-#436 := (iff #495 #835)
-#542 := (iff #522 true)
-#831 := (not false)
-#540 := (iff #831 true)
-#541 := [rewrite]: #540
-#817 := (iff #522 #831)
-#830 := (iff #518 false)
-#464 := (or false false)
-#469 := (iff #464 false)
-#470 := [rewrite]: #469
-#827 := (iff #518 #464)
-#493 := (iff #521 false)
-#498 := [rewrite]: #493
-#829 := [monotonicity #498 #498]: #827
-#828 := [trans #829 #470]: #830
-#818 := [monotonicity #828]: #817
-#543 := [trans #818 #541]: #542
-#820 := (iff #838 false)
-#821 := (not true)
-#824 := (iff #821 false)
-#825 := [rewrite]: #824
-#822 := (iff #838 #821)
-#484 := (iff #836 true)
-#478 := (iff #836 #835)
-#834 := (iff #501 true)
-#494 := [rewrite]: #834
-#819 := [monotonicity #498 #494]: #478
-#485 := [trans #819 #483]: #484
-#823 := [monotonicity #485]: #822
-#826 := [trans #823 #825]: #820
-#813 := [monotonicity #826 #543]: #436
-#536 := [trans #813 #483]: #814
-#788 := [monotonicity #536 #800 #787]: #794
-#627 := [trans #788 #624]: #626
-#841 := (iff #506 false)
-#842 := [rewrite]: #841
-#783 := [monotonicity #842 #627]: #781
-#667 := [trans #783 #782]: #785
-#837 := (iff #839 #843)
-#833 := [rewrite]: #837
-#628 := [monotonicity #842 #833 #667]: #669
-#774 := [trans #628 #678]: #773
-#673 := [monotonicity #774]: #680
-#683 := [trans #673 #682]: #680
-#679 := [quant-inst]: #677
-#684 := [mp #679 #683]: #662
-#704 := [unit-resolution #684 #857]: #832
-#700 := [monotonicity #704]: #57
-#58 := (not #57)
-#270 := [asserted]: #58
-[unit-resolution #270 #700]: false
-unsat
-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)
@@ -21938,6 +20191,844 @@
 #1103 := [trans #935 #1102]: #300
 [unit-resolution #308 #1103]: false
 unsat
+30266f09c986583d8c5407f4e044da694063d38f 303 0
+#2 := false
+decl f5 :: (-> int S2)
+#11 := 0::int
+#56 := (f5 0::int)
+decl f4 :: (-> int int int)
+#53 := 3::int
+#54 := (f4 3::int 3::int)
+#55 := (f5 #54)
+#57 := (= #55 #56)
+#832 := (= #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
+#777 := (not #852)
+#662 := (or #777 #832)
+#430 := (* -1::int 3::int)
+#515 := (mod #430 #430)
+#516 := (+ #54 #515)
+#517 := (= #516 0::int)
+#508 := (mod 3::int 3::int)
+#519 := (* -1::int #508)
+#520 := (+ #54 #519)
+#448 := (= #520 0::int)
+#521 := (<= 3::int 0::int)
+#518 := (or #521 #521)
+#522 := (not #518)
+#501 := (>= 3::int 0::int)
+#836 := (or #521 #501)
+#838 := (not #836)
+#495 := (or #838 #522)
+#625 := (ite #495 #448 #517)
+#506 := (= 3::int 0::int)
+#507 := (ite #506 #832 #625)
+#839 := (= 3::int #54)
+#840 := (ite #506 #839 #507)
+#677 := (or #777 #840)
+#680 := (iff #677 #662)
+#681 := (iff #662 #662)
+#682 := [rewrite]: #681
+#773 := (iff #840 #832)
+#843 := (= #54 3::int)
+#668 := (ite false #843 #832)
+#776 := (iff #668 #832)
+#678 := [rewrite]: #776
+#669 := (iff #840 #668)
+#785 := (iff #507 #832)
+#780 := (ite false #832 #832)
+#784 := (iff #780 #832)
+#782 := [rewrite]: #784
+#781 := (iff #507 #780)
+#626 := (iff #625 #832)
+#1 := true
+#793 := (ite true #832 #832)
+#795 := (iff #793 #832)
+#624 := [rewrite]: #795
+#794 := (iff #625 #793)
+#786 := (iff #517 #832)
+#791 := (= #516 #54)
+#807 := (+ #54 0::int)
+#805 := (= #807 #54)
+#809 := [rewrite]: #805
+#789 := (= #516 #807)
+#646 := (= #515 0::int)
+#801 := -3::int
+#802 := (mod -3::int -3::int)
+#804 := (= #802 0::int)
+#645 := [rewrite]: #804
+#803 := (= #515 #802)
+#796 := (= #430 -3::int)
+#797 := [rewrite]: #796
+#798 := [monotonicity #797 #797]: #803
+#647 := [trans #798 #645]: #646
+#790 := [monotonicity #647]: #789
+#792 := [trans #790 #809]: #791
+#787 := [monotonicity #792]: #786
+#799 := (iff #448 #832)
+#806 := (= #520 #54)
+#808 := (= #520 #807)
+#528 := (= #519 0::int)
+#815 := (* -1::int 0::int)
+#526 := (= #815 0::int)
+#527 := [rewrite]: #526
+#812 := (= #519 #815)
+#537 := (= #508 0::int)
+#811 := [rewrite]: #537
+#816 := [monotonicity #811]: #812
+#529 := [trans #816 #527]: #528
+#523 := [monotonicity #529]: #808
+#810 := [trans #523 #809]: #806
+#800 := [monotonicity #810]: #799
+#814 := (iff #495 true)
+#835 := (or false true)
+#482 := (iff #835 true)
+#483 := [rewrite]: #482
+#436 := (iff #495 #835)
+#542 := (iff #522 true)
+#831 := (not false)
+#540 := (iff #831 true)
+#541 := [rewrite]: #540
+#817 := (iff #522 #831)
+#830 := (iff #518 false)
+#464 := (or false false)
+#469 := (iff #464 false)
+#470 := [rewrite]: #469
+#827 := (iff #518 #464)
+#493 := (iff #521 false)
+#498 := [rewrite]: #493
+#829 := [monotonicity #498 #498]: #827
+#828 := [trans #829 #470]: #830
+#818 := [monotonicity #828]: #817
+#543 := [trans #818 #541]: #542
+#820 := (iff #838 false)
+#821 := (not true)
+#824 := (iff #821 false)
+#825 := [rewrite]: #824
+#822 := (iff #838 #821)
+#484 := (iff #836 true)
+#478 := (iff #836 #835)
+#834 := (iff #501 true)
+#494 := [rewrite]: #834
+#819 := [monotonicity #498 #494]: #478
+#485 := [trans #819 #483]: #484
+#823 := [monotonicity #485]: #822
+#826 := [trans #823 #825]: #820
+#813 := [monotonicity #826 #543]: #436
+#536 := [trans #813 #483]: #814
+#788 := [monotonicity #536 #800 #787]: #794
+#627 := [trans #788 #624]: #626
+#841 := (iff #506 false)
+#842 := [rewrite]: #841
+#783 := [monotonicity #842 #627]: #781
+#667 := [trans #783 #782]: #785
+#837 := (iff #839 #843)
+#833 := [rewrite]: #837
+#628 := [monotonicity #842 #833 #667]: #669
+#774 := [trans #628 #678]: #773
+#673 := [monotonicity #774]: #680
+#683 := [trans #673 #682]: #680
+#679 := [quant-inst]: #677
+#684 := [mp #679 #683]: #662
+#704 := [unit-resolution #684 #857]: #832
+#700 := [monotonicity #704]: #57
+#58 := (not #57)
+#270 := [asserted]: #58
+[unit-resolution #270 #700]: false
+unsat
+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
 7c27e22543c792c36da02bc4168bcd3f7ce7432d 961 0
 #2 := false
 decl f5 :: (-> int S2)
@@ -22900,6 +21991,678 @@
 #1225 := [trans #1224 #1218]: #301
 [unit-resolution #309 #1225]: 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
 85ac6a29deaf91400dbfbdf638baba5914d49e18 60 0
 #2 := false
 #13 := 0::int
@@ -22961,6 +22724,67 @@
 #630 := [unit-resolution #632 #631]: #266
 [th-lemma #159 #629 #630]: 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
 ec431e8216f1d8b2f670cdb0885c1e07fa83e950 60 0
 #2 := false
 #13 := 0::int
@@ -23022,6 +22846,67 @@
 #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
 db10b3baed96f5ad894a70d823fdba86b4614481 253 0
 #2 := false
 #13 := 0::int
@@ -23276,507 +23161,6 @@
 #473 := [lemma #472]: #233
 [unit-resolution #631 #473 #487]: false
 unsat
-89bc3a6333da51700f7daac14eb84da9abc113a1 103 0
-#2 := false
-#13 := 0::int
-decl f4 :: (-> S2 int)
-decl f7 :: S2
-#30 := f7
-#31 := (f4 f7)
-decl f6 :: S2
-#27 := f6
-#28 := (f4 f6)
-#143 := -1::int
-#146 := (* -1::int #31)
-#150 := (+ #28 #146)
-#151 := (<= #150 0::int)
-#154 := (ite #151 #28 #31)
-#162 := (* -1::int #154)
-decl f5 :: S2
-#25 := f5
-#26 := (f4 f5)
-#163 := (+ #26 #162)
-#161 := (>= #163 0::int)
-#177 := (* -1::int #28)
-#178 := (+ #26 #177)
-#176 := (>= #178 0::int)
-#175 := (not #176)
-#147 := (+ #26 #146)
-#145 := (>= #147 0::int)
-#144 := (not #145)
-#187 := (and #144 #161 #175)
-#33 := (<= #28 #31)
-#34 := (ite #33 #28 #31)
-#35 := (< #26 #34)
-#36 := (not #35)
-#32 := (< #26 #31)
-#37 := (and #32 #36)
-#29 := (< #26 #28)
-#38 := (and #29 #37)
-#190 := (iff #38 #187)
-#181 := (and #144 #161)
-#184 := (and #175 #181)
-#188 := (iff #184 #187)
-#189 := [rewrite]: #188
-#185 := (iff #38 #184)
-#182 := (iff #37 #181)
-#173 := (iff #36 #161)
-#160 := (not #161)
-#168 := (not #160)
-#171 := (iff #168 #161)
-#172 := [rewrite]: #171
-#169 := (iff #36 #168)
-#166 := (iff #35 #160)
-#157 := (< #26 #154)
-#164 := (iff #157 #160)
-#165 := [rewrite]: #164
-#158 := (iff #35 #157)
-#155 := (= #34 #154)
-#152 := (iff #33 #151)
-#153 := [rewrite]: #152
-#156 := [monotonicity #153]: #155
-#159 := [monotonicity #156]: #158
-#167 := [trans #159 #165]: #166
-#170 := [monotonicity #167]: #169
-#174 := [trans #170 #172]: #173
-#148 := (iff #32 #144)
-#149 := [rewrite]: #148
-#183 := [monotonicity #149 #174]: #182
-#179 := (iff #29 #175)
-#180 := [rewrite]: #179
-#186 := [monotonicity #180 #183]: #185
-#191 := [trans #186 #189]: #190
-#140 := [asserted]: #38
-#192 := [mp #140 #191]: #187
-#194 := [and-elim #192]: #161
-#193 := [and-elim #192]: #144
-#365 := (+ #31 #162)
-#638 := (<= #365 0::int)
-#353 := (= #31 #154)
-#268 := (not #151)
-#267 := (= #28 #154)
-#643 := (not #267)
-#652 := (+ #28 #162)
-#374 := (<= #652 0::int)
-#645 := (not #374)
-#195 := [and-elim #192]: #175
-#366 := [hypothesis]: #374
-#367 := [th-lemma #366 #195 #194]: false
-#646 := [lemma #367]: #645
-#361 := [hypothesis]: #267
-#647 := (or #643 #374)
-#644 := [th-lemma]: #647
-#648 := [unit-resolution #644 #361 #646]: false
-#637 := [lemma #648]: #643
-#354 := (or #268 #267)
-#355 := [def-axiom]: #354
-#634 := [unit-resolution #355 #637]: #268
-#346 := (or #151 #353)
-#357 := [def-axiom]: #346
-#635 := [unit-resolution #357 #634]: #353
-#640 := (not #353)
-#641 := (or #640 #638)
-#636 := [th-lemma]: #641
-#642 := [unit-resolution #636 #635]: #638
-[th-lemma #642 #193 #194]: false
-unsat
-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
@@ -24215,6 +23599,228 @@
 #462 := [unit-resolution #492 #479]: #669
 [th-lemma #462 #528 #194 #475 #461]: false
 unsat
+89bc3a6333da51700f7daac14eb84da9abc113a1 103 0
+#2 := false
+#13 := 0::int
+decl f4 :: (-> S2 int)
+decl f7 :: S2
+#30 := f7
+#31 := (f4 f7)
+decl f6 :: S2
+#27 := f6
+#28 := (f4 f6)
+#143 := -1::int
+#146 := (* -1::int #31)
+#150 := (+ #28 #146)
+#151 := (<= #150 0::int)
+#154 := (ite #151 #28 #31)
+#162 := (* -1::int #154)
+decl f5 :: S2
+#25 := f5
+#26 := (f4 f5)
+#163 := (+ #26 #162)
+#161 := (>= #163 0::int)
+#177 := (* -1::int #28)
+#178 := (+ #26 #177)
+#176 := (>= #178 0::int)
+#175 := (not #176)
+#147 := (+ #26 #146)
+#145 := (>= #147 0::int)
+#144 := (not #145)
+#187 := (and #144 #161 #175)
+#33 := (<= #28 #31)
+#34 := (ite #33 #28 #31)
+#35 := (< #26 #34)
+#36 := (not #35)
+#32 := (< #26 #31)
+#37 := (and #32 #36)
+#29 := (< #26 #28)
+#38 := (and #29 #37)
+#190 := (iff #38 #187)
+#181 := (and #144 #161)
+#184 := (and #175 #181)
+#188 := (iff #184 #187)
+#189 := [rewrite]: #188
+#185 := (iff #38 #184)
+#182 := (iff #37 #181)
+#173 := (iff #36 #161)
+#160 := (not #161)
+#168 := (not #160)
+#171 := (iff #168 #161)
+#172 := [rewrite]: #171
+#169 := (iff #36 #168)
+#166 := (iff #35 #160)
+#157 := (< #26 #154)
+#164 := (iff #157 #160)
+#165 := [rewrite]: #164
+#158 := (iff #35 #157)
+#155 := (= #34 #154)
+#152 := (iff #33 #151)
+#153 := [rewrite]: #152
+#156 := [monotonicity #153]: #155
+#159 := [monotonicity #156]: #158
+#167 := [trans #159 #165]: #166
+#170 := [monotonicity #167]: #169
+#174 := [trans #170 #172]: #173
+#148 := (iff #32 #144)
+#149 := [rewrite]: #148
+#183 := [monotonicity #149 #174]: #182
+#179 := (iff #29 #175)
+#180 := [rewrite]: #179
+#186 := [monotonicity #180 #183]: #185
+#191 := [trans #186 #189]: #190
+#140 := [asserted]: #38
+#192 := [mp #140 #191]: #187
+#194 := [and-elim #192]: #161
+#193 := [and-elim #192]: #144
+#365 := (+ #31 #162)
+#638 := (<= #365 0::int)
+#353 := (= #31 #154)
+#268 := (not #151)
+#267 := (= #28 #154)
+#643 := (not #267)
+#652 := (+ #28 #162)
+#374 := (<= #652 0::int)
+#645 := (not #374)
+#195 := [and-elim #192]: #175
+#366 := [hypothesis]: #374
+#367 := [th-lemma #366 #195 #194]: false
+#646 := [lemma #367]: #645
+#361 := [hypothesis]: #267
+#647 := (or #643 #374)
+#644 := [th-lemma]: #647
+#648 := [unit-resolution #644 #361 #646]: false
+#637 := [lemma #648]: #643
+#354 := (or #268 #267)
+#355 := [def-axiom]: #354
+#634 := [unit-resolution #355 #637]: #268
+#346 := (or #151 #353)
+#357 := [def-axiom]: #346
+#635 := [unit-resolution #357 #634]: #353
+#640 := (not #353)
+#641 := (or #640 #638)
+#636 := [th-lemma]: #641
+#642 := [unit-resolution #636 #635]: #638
+[th-lemma #642 #193 #194]: false
+unsat
+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
 452992a55139d640bc8969517cd3166d0f13bb37 103 0
 #2 := false
 #13 := 0::int
@@ -24319,6 +23925,163 @@
 #642 := [unit-resolution #636 #635]: #639
 [th-lemma #642 #195 #194]: 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
 c2cc039657273f4c31549e108b87c8f33092d45f 119 0
 #2 := false
 decl f3 :: (-> int S2)
@@ -24744,6 +24507,24 @@
 #260 := [th-lemma]: #618
 [unit-resolution #260 #139 #617]: false
 unsat
+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
 3413e63e105cab337755fe6950808ba664d2597a 75 0
 #2 := false
 decl f3 :: (-> int S2)
@@ -24820,6 +24601,32 @@
 #153 := [and-elim #151]: #32
 [unit-resolution #153 #339]: 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
 7c99e4c706a5bfe7c4bac45e274e4bb5f5ba171b 20 0
 #2 := false
 decl f4 :: (-> S2 int)
@@ -24841,6 +24648,24 @@
 #130 := [asserted]: #28
 [mp #130 #141]: 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
 2c1b55c23f4e7de8da8d5f6f9bfa461e094d5fdc 319 0
 #2 := false
 #13 := 0::int
@@ -25161,6 +24986,82 @@
 #507 := [unit-resolution #506 #504]: #623
 [th-lemma #507 #516 #163 #162]: 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
 b1d5ad1c199444dbf6fe73759c2c90d7c60b146e 325 0
 #2 := false
 #13 := 0::int
@@ -25487,6 +25388,37 @@
 #515 := [th-lemma]: #519
 [unit-resolution #515 #510 #517]: 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
 a489de1c5d2012bec455302dfb4702c1a1df77d2 38 0
 #2 := false
 #13 := 0::int
@@ -25526,298 +25458,160 @@
 #228 := [unit-resolution #313 #154]: #144
 [unit-resolution #228 #155]: false
 unsat
-08e6db4ea57a4ed26d47431f93a880d0e622c6f9 292 0
-#2 := false
-#13 := 0::int
+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
+957a81798a07625d9f9bb6cae6421523a9942b09 70 0
+#2 := false
+#149 := -1::int
 decl f4 :: (-> S2 int)
-decl f3 :: (-> int S2)
-decl f5 :: S2
-#25 := f5
-#26 := (f4 f5)
-#30 := 1::int
-#139 := (+ 1::int #26)
-#142 := (f3 #139)
-#145 := (f4 #142)
-#157 := -1::int
-#165 := (* -1::int #145)
 decl f6 :: S2
 #27 := f6
 #28 := (f4 f6)
-#166 := (+ #28 #165)
-#164 := (>= #166 0::int)
-#160 := (* -1::int #28)
-#161 := (+ #26 #160)
-#159 := (>= #161 0::int)
-#319 := (+ #26 #165)
-#643 := (>= #319 -1::int)
-#324 := (= #319 -1::int)
-#661 := (>= #26 -1::int)
-#634 := (>= #26 0::int)
-#613 := (= #26 0::int)
-#344 := (f3 #26)
-#349 := (f4 #344)
-#631 := (= #349 0::int)
-#606 := (not #634)
-#607 := [hypothesis]: #606
-#635 := (or #631 #634)
-#14 := (:var 0 int)
-#16 := (f3 #14)
-#678 := (pattern #16)
-#79 := (>= #14 0::int)
-#17 := (f4 #16)
-#22 := (= #17 0::int)
-#127 := (or #22 #79)
-#685 := (forall (vars (?v0 int)) (:pat #678) #127)
-#132 := (forall (vars (?v0 int)) #127)
-#688 := (iff #132 #685)
-#686 := (iff #127 #127)
-#687 := [refl]: #686
-#689 := [quant-intro #687]: #688
-#188 := (~ #132 #132)
-#200 := (~ #127 #127)
-#201 := [refl]: #200
-#189 := [nnf-pos #201]: #188
-#21 := (< #14 0::int)
-#23 := (implies #21 #22)
-#24 := (forall (vars (?v0 int)) #23)
-#135 := (iff #24 #132)
-#98 := (= 0::int #17)
-#104 := (not #21)
-#105 := (or #104 #98)
-#110 := (forall (vars (?v0 int)) #105)
-#133 := (iff #110 #132)
-#130 := (iff #105 #127)
-#124 := (or #79 #22)
-#128 := (iff #124 #127)
-#129 := [rewrite]: #128
-#125 := (iff #105 #124)
-#122 := (iff #98 #22)
-#123 := [rewrite]: #122
-#120 := (iff #104 #79)
-#80 := (not #79)
-#115 := (not #80)
-#118 := (iff #115 #79)
-#119 := [rewrite]: #118
-#116 := (iff #104 #115)
-#113 := (iff #21 #80)
-#114 := [rewrite]: #113
-#117 := [monotonicity #114]: #116
-#121 := [trans #117 #119]: #120
-#126 := [monotonicity #121 #123]: #125
-#131 := [trans #126 #129]: #130
-#134 := [quant-intro #131]: #133
-#111 := (iff #24 #110)
-#108 := (iff #23 #105)
-#101 := (implies #21 #98)
-#106 := (iff #101 #105)
-#107 := [rewrite]: #106
-#102 := (iff #23 #101)
-#99 := (iff #22 #98)
-#100 := [rewrite]: #99
-#103 := [monotonicity #100]: #102
-#109 := [trans #103 #107]: #108
-#112 := [quant-intro #109]: #111
-#136 := [trans #112 #134]: #135
-#97 := [asserted]: #24
-#137 := [mp #97 #136]: #132
-#202 := [mp~ #137 #189]: #132
-#690 := [mp #202 #689]: #685
-#263 := (not #685)
-#625 := (or #263 #631 #634)
-#626 := (or #263 #635)
-#622 := (iff #626 #625)
-#623 := [rewrite]: #622
-#627 := [quant-inst]: #626
-#628 := [mp #627 #623]: #625
-#609 := [unit-resolution #628 #690]: #635
-#610 := [unit-resolution #609 #607]: #631
-#611 := (= #26 #349)
-#348 := (= f5 #344)
-#8 := (:var 0 S2)
-#9 := (f4 #8)
-#670 := (pattern #9)
-#10 := (f3 #9)
-#54 := (= #8 #10)
-#671 := (forall (vars (?v0 S2)) (:pat #670) #54)
-#57 := (forall (vars (?v0 S2)) #54)
-#672 := (iff #57 #671)
-#674 := (iff #671 #671)
-#675 := [rewrite]: #674
-#673 := [rewrite]: #672
-#676 := [trans #673 #675]: #672
-#196 := (~ #57 #57)
-#194 := (~ #54 #54)
-#195 := [refl]: #194
-#197 := [nnf-pos #195]: #196
-#11 := (= #10 #8)
-#12 := (forall (vars (?v0 S2)) #11)
-#58 := (iff #12 #57)
-#55 := (iff #11 #54)
-#56 := [rewrite]: #55
-#59 := [quant-intro #56]: #58
-#53 := [asserted]: #12
-#62 := [mp #53 #59]: #57
-#185 := [mp~ #62 #197]: #57
-#677 := [mp #185 #676]: #671
-#664 := (not #671)
-#321 := (or #664 #348)
-#451 := [quant-inst]: #321
-#608 := [unit-resolution #451 #677]: #348
-#493 := [monotonicity #608]: #611
-#494 := [trans #493 #610]: #613
-#495 := (not #613)
-#454 := (or #495 #634)
-#602 := [th-lemma]: #454
-#504 := [unit-resolution #602 #607 #494]: false
-#599 := [lemma #504]: #634
-#619 := (or #606 #661)
-#600 := [th-lemma]: #619
-#603 := [unit-resolution #600 #599]: #661
-#308 := (not #661)
-#311 := (or #324 #308)
-#61 := (= #14 #17)
-#86 := (or #61 #80)
-#679 := (forall (vars (?v0 int)) (:pat #678) #86)
-#91 := (forall (vars (?v0 int)) #86)
-#682 := (iff #91 #679)
-#680 := (iff #86 #86)
-#681 := [refl]: #680
-#683 := [quant-intro #681]: #682
-#186 := (~ #91 #91)
-#183 := (~ #86 #86)
-#198 := [refl]: #183
-#187 := [nnf-pos #198]: #186
-#18 := (= #17 #14)
-#15 := (<= 0::int #14)
-#19 := (implies #15 #18)
-#20 := (forall (vars (?v0 int)) #19)
-#94 := (iff #20 #91)
-#68 := (not #15)
-#69 := (or #68 #61)
-#74 := (forall (vars (?v0 int)) #69)
-#92 := (iff #74 #91)
-#89 := (iff #69 #86)
-#83 := (or #80 #61)
-#87 := (iff #83 #86)
-#88 := [rewrite]: #87
-#84 := (iff #69 #83)
-#81 := (iff #68 #80)
-#77 := (iff #15 #79)
-#78 := [rewrite]: #77
-#82 := [monotonicity #78]: #81
-#85 := [monotonicity #82]: #84
-#90 := [trans #85 #88]: #89
-#93 := [quant-intro #90]: #92
-#75 := (iff #20 #74)
-#72 := (iff #19 #69)
-#65 := (implies #15 #61)
-#70 := (iff #65 #69)
-#71 := [rewrite]: #70
-#66 := (iff #19 #65)
-#63 := (iff #18 #61)
-#64 := [rewrite]: #63
-#67 := [monotonicity #64]: #66
-#73 := [trans #67 #71]: #72
-#76 := [quant-intro #73]: #75
-#95 := [trans #76 #93]: #94
-#60 := [asserted]: #20
-#96 := [mp #60 #95]: #91
-#199 := [mp~ #96 #187]: #91
-#684 := [mp #199 #683]: #679
-#649 := (not #679)
-#650 := (or #649 #324 #308)
-#666 := (>= #139 0::int)
-#667 := (not #666)
-#669 := (= #139 #145)
-#659 := (or #669 #667)
-#651 := (or #649 #659)
-#296 := (iff #651 #650)
-#652 := (or #649 #311)
-#655 := (iff #652 #650)
-#295 := [rewrite]: #655
-#290 := (iff #651 #652)
-#647 := (iff #659 #311)
-#309 := (iff #667 #308)
-#304 := (iff #666 #661)
-#645 := [rewrite]: #304
-#310 := [monotonicity #645]: #309
-#660 := (iff #669 #324)
-#320 := [rewrite]: #660
-#648 := [monotonicity #320 #310]: #647
-#653 := [monotonicity #648]: #290
-#656 := [trans #653 #295]: #296
-#646 := [quant-inst]: #651
-#654 := [mp #646 #656]: #650
-#488 := [unit-resolution #654 #684]: #311
-#503 := [unit-resolution #488 #603]: #324
-#505 := (not #324)
-#506 := (or #505 #643)
-#499 := [th-lemma]: #506
-#507 := [unit-resolution #499 #503]: #643
-#158 := (not #159)
-#508 := [hypothesis]: #158
-#169 := (not #164)
-#342 := (or #159 #169)
-#175 := (iff #159 #164)
+#152 := (* -1::int #28)
+decl f5 :: S2
+#25 := f5
+#26 := (f4 f5)
+#153 := (+ #26 #152)
+#156 := (<= #153 -1::int)
+#13 := 0::int
+#151 := (>= #153 0::int)
+#150 := (not #151)
+#159 := (not #156)
+#357 := [hypothesis]: #159
+#336 := (or #150 #156)
+#165 := (iff #151 #156)
+#30 := 1::int
 #31 := (+ #26 1::int)
-#32 := (f3 #31)
-#33 := (f4 #32)
-#34 := (<= #33 #28)
-#35 := (not #34)
+#32 := (<= #31 #28)
+#33 := (not #32)
 #29 := (< #26 #28)
-#36 := (iff #29 #35)
-#180 := (iff #36 #175)
-#148 := (<= #145 #28)
-#151 := (not #148)
-#154 := (iff #29 #151)
-#178 := (iff #154 #175)
-#172 := (iff #158 #169)
-#176 := (iff #172 #175)
-#177 := [rewrite]: #176
-#173 := (iff #154 #172)
-#170 := (iff #151 #169)
-#167 := (iff #148 #164)
-#168 := [rewrite]: #167
-#171 := [monotonicity #168]: #170
-#162 := (iff #29 #158)
-#163 := [rewrite]: #162
-#174 := [monotonicity #163 #171]: #173
-#179 := [trans #174 #177]: #178
-#155 := (iff #36 #154)
-#152 := (iff #35 #151)
-#149 := (iff #34 #148)
-#146 := (= #33 #145)
-#143 := (= #32 #142)
-#140 := (= #31 #139)
-#141 := [rewrite]: #140
-#144 := [monotonicity #141]: #143
-#147 := [monotonicity #144]: #146
-#150 := [monotonicity #147]: #149
-#153 := [monotonicity #150]: #152
-#156 := [monotonicity #153]: #155
-#181 := [trans #156 #179]: #180
-#138 := [asserted]: #36
-#182 := [mp #138 #181]: #175
-#256 := (not #175)
-#341 := (or #159 #169 #256)
-#257 := [def-axiom]: #341
-#343 := [unit-resolution #257 #182]: #342
-#509 := [unit-resolution #343 #508]: #169
-#510 := (not #643)
-#511 := (or #510 #159 #164)
-#604 := [th-lemma]: #511
-#601 := [unit-resolution #604 #509 #508 #507]: false
-#605 := [lemma #601]: #159
-#346 := (or #158 #164)
-#334 := (or #158 #164 #256)
-#345 := [def-axiom]: #334
-#347 := [unit-resolution #345 #182]: #346
-#592 := [unit-resolution #347 #605]: #164
-#657 := (<= #319 -1::int)
-#594 := (or #505 #657)
-#586 := [th-lemma]: #594
-#583 := [unit-resolution #586 #503]: #657
-[th-lemma #583 #605 #592]: false
+#34 := (iff #29 #33)
+#170 := (iff #34 #165)
+#137 := (+ 1::int #26)
+#140 := (<= #137 #28)
+#143 := (not #140)
+#146 := (iff #29 #143)
+#168 := (iff #146 #165)
+#162 := (iff #150 #159)
+#166 := (iff #162 #165)
+#167 := [rewrite]: #166
+#163 := (iff #146 #162)
+#160 := (iff #143 #159)
+#157 := (iff #140 #156)
+#158 := [rewrite]: #157
+#161 := [monotonicity #158]: #160
+#154 := (iff #29 #150)
+#155 := [rewrite]: #154
+#164 := [monotonicity #155 #161]: #163
+#169 := [trans #164 #167]: #168
+#147 := (iff #34 #146)
+#144 := (iff #33 #143)
+#141 := (iff #32 #140)
+#138 := (= #31 #137)
+#139 := [rewrite]: #138
+#142 := [monotonicity #139]: #141
+#145 := [monotonicity #142]: #144
+#148 := [monotonicity #145]: #147
+#171 := [trans #148 #169]: #170
+#136 := [asserted]: #34
+#172 := [mp #136 #171]: #165
+#246 := (not #165)
+#324 := (or #150 #156 #246)
+#335 := [def-axiom]: #324
+#337 := [unit-resolution #335 #172]: #336
+#358 := [unit-resolution #337 #357]: #150
+#359 := [th-lemma #357 #358]: false
+#253 := [lemma #359]: #156
+#629 := (or #159 #150)
+#630 := [th-lemma]: #629
+#352 := [unit-resolution #630 #253]: #150
+#332 := (or #151 #159)
+#331 := (or #151 #159 #246)
+#247 := [def-axiom]: #331
+#333 := [unit-resolution #247 #172]: #332
+[unit-resolution #333 #352 #253]: 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
 7a80c4532f5787abddc50ad2e7bc9838390acb73 11 0
 #2 := false
@@ -25831,6 +25625,40 @@
 #129 := [asserted]: #27
 [mp #129 #133]: 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
 a5e0b28483c053128bc0be113d20a82250433b74 54 0
 #2 := false
 #13 := 0::int
@@ -25886,6 +25714,52 @@
 #171 := [and-elim #170]: #144
 [th-lemma #171 #172 #173]: 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
 4d82a9a048129e417a9e0a1825ce1ac4532cabd0 55 0
 #2 := false
 #13 := 0::int
@@ -25998,6 +25872,33 @@
 #173 := [and-elim #171]: #152
 [th-lemma #173 #172 #174]: 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
 2fa07586f302aa9750784686f2cfa4df7ec4584e 61 0
 #2 := false
 #13 := 0::int
@@ -26060,6 +25961,33 @@
 #179 := [and-elim #177]: #149
 [th-lemma #179 #178 #180]: 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
 fe16712a1fa51f4a62675a12f177489494d4ff85 53 0
 #2 := false
 #13 := 0::int
@@ -26114,393 +26042,6 @@
 #171 := [and-elim #169]: #149
 [th-lemma #171 #170 #172]: false
 unsat
-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
@@ -26622,6 +26163,33 @@
 #31 := [asserted]: #14
 [mp #31 #79]: false
 unsat
+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
 1c28009871674ce12da75bf1dced4812d4d59e07 61 0
 #2 := false
 decl f3 :: int
@@ -26684,33 +26252,6 @@
 #31 := [asserted]: #14
 [mp #31 #79]: false
 unsat
-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
@@ -26744,6 +26285,33 @@
 #30 := [asserted]: #13
 [mp #30 #50]: false
 unsat
+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
 31fc2c351446233f2061fde4a4592b401c54f862 64 0
 #2 := false
 #12 := 0::int
@@ -26809,6 +26377,33 @@
 #32 := [asserted]: #15
 [mp #32 #81]: false
 unsat
+592736b964503a2e2a8b3d310866adf35fdeb70d 26 0
+#2 := false
+#8 := 0::int
+decl f3 :: int
+#9 := f3
+#10 := (* 0::int f3)
+#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
 7e5077e88974e2fa3da11f2fc124b2fc85a4f520 44 0
 #2 := false
 decl f3 :: int
@@ -26854,152 +26449,6 @@
 #31 := [asserted]: #14
 [mp #31 #61]: false
 unsat
-891b65b932309e97c69155ee809ef85f6496b3c5 42 0
-#2 := false
-decl f3 :: int
-#8 := f3
-decl f4 :: int
-#9 := f4
-#11 := (- f4)
-#12 := (+ #11 f3)
-#10 := (- f3 f4)
-#13 := (= #10 #12)
-#14 := (not #13)
-#58 := (iff #14 false)
-#1 := true
-#53 := (not true)
-#56 := (iff #53 false)
-#57 := [rewrite]: #56
-#54 := (iff #14 #53)
-#51 := (iff #13 true)
-#32 := -1::int
-#33 := (* -1::int f4)
-#34 := (+ f3 #33)
-#46 := (= #34 #34)
-#49 := (iff #46 true)
-#50 := [rewrite]: #49
-#47 := (iff #13 #46)
-#44 := (= #12 #34)
-#39 := (+ #33 f3)
-#42 := (= #39 #34)
-#43 := [rewrite]: #42
-#40 := (= #12 #39)
-#37 := (= #11 #33)
-#38 := [rewrite]: #37
-#41 := [monotonicity #38]: #40
-#45 := [trans #41 #43]: #44
-#35 := (= #10 #34)
-#36 := [rewrite]: #35
-#48 := [monotonicity #36 #45]: #47
-#52 := [trans #48 #50]: #51
-#55 := [monotonicity #52]: #54
-#59 := [trans #55 #57]: #58
-#31 := [asserted]: #14
-[mp #31 #59]: false
-unsat
-e90209aecaafc9e5a8d65cc3bfe9923d664d1180 48 0
-#2 := false
-decl f5 :: int
-#11 := f5
-decl f4 :: int
-#9 := f4
-#13 := (+ f4 f5)
-decl f3 :: int
-#8 := f3
-#14 := (- f3 #13)
-#10 := (- f3 f4)
-#12 := (- #10 f5)
-#15 := (= #12 #14)
-#16 := (not #15)
-#63 := (iff #16 false)
-#1 := true
-#58 := (not true)
-#61 := (iff #58 false)
-#62 := [rewrite]: #61
-#59 := (iff #16 #58)
-#56 := (iff #15 true)
-#34 := -1::int
-#42 := (* -1::int f5)
-#35 := (* -1::int f4)
-#43 := (+ #35 #42)
-#44 := (+ f3 #43)
-#51 := (= #44 #44)
-#54 := (iff #51 true)
-#55 := [rewrite]: #54
-#52 := (iff #15 #51)
-#49 := (= #14 #44)
-#50 := [rewrite]: #49
-#47 := (= #12 #44)
-#36 := (+ f3 #35)
-#39 := (- #36 f5)
-#45 := (= #39 #44)
-#46 := [rewrite]: #45
-#40 := (= #12 #39)
-#37 := (= #10 #36)
-#38 := [rewrite]: #37
-#41 := [monotonicity #38]: #40
-#48 := [trans #41 #46]: #47
-#53 := [monotonicity #48 #50]: #52
-#57 := [trans #53 #55]: #56
-#60 := [monotonicity #57]: #59
-#64 := [trans #60 #62]: #63
-#33 := [asserted]: #16
-[mp #33 #64]: false
-unsat
-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
-#9 := f3
-#10 := (* 0::int f3)
-#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
 89e44e496d1c9e19d85ee56ab3fb92079351f1b0 26 0
 #2 := false
 decl f3 :: int
@@ -27054,6 +26503,49 @@
 #29 := [asserted]: #12
 [mp #29 #44]: false
 unsat
+891b65b932309e97c69155ee809ef85f6496b3c5 42 0
+#2 := false
+decl f3 :: int
+#8 := f3
+decl f4 :: int
+#9 := f4
+#11 := (- f4)
+#12 := (+ #11 f3)
+#10 := (- f3 f4)
+#13 := (= #10 #12)
+#14 := (not #13)
+#58 := (iff #14 false)
+#1 := true
+#53 := (not true)
+#56 := (iff #53 false)
+#57 := [rewrite]: #56
+#54 := (iff #14 #53)
+#51 := (iff #13 true)
+#32 := -1::int
+#33 := (* -1::int f4)
+#34 := (+ f3 #33)
+#46 := (= #34 #34)
+#49 := (iff #46 true)
+#50 := [rewrite]: #49
+#47 := (iff #13 #46)
+#44 := (= #12 #34)
+#39 := (+ #33 f3)
+#42 := (= #39 #34)
+#43 := [rewrite]: #42
+#40 := (= #12 #39)
+#37 := (= #11 #33)
+#38 := [rewrite]: #37
+#41 := [monotonicity #38]: #40
+#45 := [trans #41 #43]: #44
+#35 := (= #10 #34)
+#36 := [rewrite]: #35
+#48 := [monotonicity #36 #45]: #47
+#52 := [trans #48 #50]: #51
+#55 := [monotonicity #52]: #54
+#59 := [trans #55 #57]: #58
+#31 := [asserted]: #14
+[mp #31 #59]: false
+unsat
 f1f03cdf8e0332171bfb331e70236cd578ef1b1a 40 0
 #2 := false
 decl f3 :: int
@@ -27095,6 +26587,55 @@
 #31 := [asserted]: #14
 [mp #31 #58]: false
 unsat
+e90209aecaafc9e5a8d65cc3bfe9923d664d1180 48 0
+#2 := false
+decl f5 :: int
+#11 := f5
+decl f4 :: int
+#9 := f4
+#13 := (+ f4 f5)
+decl f3 :: int
+#8 := f3
+#14 := (- f3 #13)
+#10 := (- f3 f4)
+#12 := (- #10 f5)
+#15 := (= #12 #14)
+#16 := (not #15)
+#63 := (iff #16 false)
+#1 := true
+#58 := (not true)
+#61 := (iff #58 false)
+#62 := [rewrite]: #61
+#59 := (iff #16 #58)
+#56 := (iff #15 true)
+#34 := -1::int
+#42 := (* -1::int f5)
+#35 := (* -1::int f4)
+#43 := (+ #35 #42)
+#44 := (+ f3 #43)
+#51 := (= #44 #44)
+#54 := (iff #51 true)
+#55 := [rewrite]: #54
+#52 := (iff #15 #51)
+#49 := (= #14 #44)
+#50 := [rewrite]: #49
+#47 := (= #12 #44)
+#36 := (+ f3 #35)
+#39 := (- #36 f5)
+#45 := (= #39 #44)
+#46 := [rewrite]: #45
+#40 := (= #12 #39)
+#37 := (= #10 #36)
+#38 := [rewrite]: #37
+#41 := [monotonicity #38]: #40
+#48 := [trans #41 #46]: #47
+#53 := [monotonicity #48 #50]: #52
+#57 := [trans #53 #55]: #56
+#60 := [monotonicity #57]: #59
+#64 := [trans #60 #62]: #63
+#33 := [asserted]: #16
+[mp #33 #64]: false
+unsat
 dcb91b564372ef4df7bba10a66d7a6eecfdd1891 35 0
 #2 := false
 decl f3 :: int
@@ -27131,35 +26672,6 @@
 #31 := [asserted]: #14
 [mp #31 #53]: false
 unsat
-e58e9441ca33f62ea126f129a1302457840ee037 28 0
-#2 := false
-#8 := 3::int
-decl f3 :: int
-#9 := f3
-#11 := (* f3 3::int)
-#10 := (* 3::int f3)
-#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
 dc9c4895f383c2bc2209a55e4f8478193dad358d 244 0
 #2 := false
 #11 := 0::int
@@ -27405,6 +26917,304 @@
 #683 := [mp #696 #682]: #695
 [unit-resolution #683 #742 #167]: false
 unsat
+e58e9441ca33f62ea126f129a1302457840ee037 28 0
+#2 := false
+#8 := 3::int
+decl f3 :: int
+#9 := f3
+#11 := (* f3 3::int)
+#10 := (* 3::int f3)
+#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
+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
 6b27a316bc3cb1671937528bbb2aa4c5f16dd00c 257 0
 #2 := false
 #11 := 0::int
@@ -27663,6 +27473,283 @@
 #673 := [mp #687 #520]: #680
 [unit-resolution #673 #743 #168]: 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
 ac9468ec7d34e56bdf0ee161b89fe3a4afadc438 281 0
 #2 := false
 #11 := 0::int
@@ -28227,6 +28314,305 @@
 #667 := [mp #573 #670]: #572
 [unit-resolution #667 #742 #168]: 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
 892e25c8bdf423f86a8231d77f83b5c173fd91c0 292 0
 #2 := false
 #38 := 3::int
@@ -28520,6 +28906,303 @@
 #660 := [mp #653 #664]: #659
 [unit-resolution #660 #743 #169]: 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
 dcc2c3c66de3e673b0cfeb392eb7d2e8ec04caf3 335 0
 #2 := false
 decl f3 :: (-> int int int)
@@ -28856,6 +29539,314 @@
 #500 := [mp #530 #497]: #170
 [unit-resolution #178 #500]: 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
 2a60c54232ee9900c0a7254451730bed52200c08 306 0
 #2 := false
 #11 := 0::int
@@ -29163,6 +30154,315 @@
 #682 := [mp #588 #685]: #587
 [unit-resolution #682 #757 #193]: 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
 923c4bec65f1ce0be6fa5d4ec7cf2b004a307f06 315 0
 #2 := false
 #66 := -1::int
@@ -29791,6 +31091,324 @@
 #661 := [mp #659 #672]: #660
 [unit-resolution #661 #750 #188]: 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
 ef319beaaae1961d7e0fb146b0bd21608ccf741d 362 0
 #2 := false
 #11 := 0::int
@@ -30154,6 +31772,334 @@
 #563 := [th-lemma]: #562
 [unit-resolution #563 #561 #633]: false
 unsat
+07f15f3f61db6e3805237fa8c50e95e2181bdd46 327 0
+#2 := false
+#11 := 0::int
+decl f4 :: (-> int int int)
+#66 := -1::int
+#38 := 1::int
+#172 := (f4 1::int -1::int)
+#175 := (= #172 0::int)
+#188 := (not #175)
+#39 := (- 1::int)
+#40 := (f4 1::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
+#580 := (not #759)
+#588 := (or #580 #175)
+#338 := (* -1::int -1::int)
+#422 := (* -1::int 1::int)
+#423 := (mod #422 #338)
+#424 := (+ #172 #423)
+#415 := (= #424 0::int)
+#426 := (mod 1::int -1::int)
+#427 := (* -1::int #426)
+#428 := (+ #172 #427)
+#425 := (= #428 0::int)
+#429 := (<= -1::int 0::int)
+#408 := (<= 1::int 0::int)
+#743 := (or #408 #429)
+#745 := (not #743)
+#402 := (>= 1::int 0::int)
+#532 := (or #429 #402)
+#739 := (not #532)
+#413 := (or #739 #745)
+#414 := (ite #413 #425 #415)
+#746 := (= 1::int 0::int)
+#747 := (ite #746 #175 #414)
+#748 := (= 1::int #172)
+#749 := (= -1::int 0::int)
+#744 := (ite #749 #748 #747)
+#589 := (or #580 #744)
+#591 := (iff #589 #588)
+#685 := (iff #588 #588)
+#682 := [rewrite]: #685
+#586 := (iff #744 #175)
+#405 := (= #172 1::int)
+#680 := (ite false #405 #175)
+#569 := (iff #680 #175)
+#584 := [rewrite]: #569
+#681 := (iff #744 #680)
+#683 := (iff #747 #175)
+#688 := (ite false #175 #175)
+#689 := (iff #688 #175)
+#692 := [rewrite]: #689
+#576 := (iff #747 #688)
+#574 := (iff #414 #175)
+#690 := (iff #414 #688)
+#534 := (iff #415 #175)
+#531 := (= #424 #172)
+#706 := (+ #172 0::int)
+#703 := (= #706 #172)
+#704 := [rewrite]: #703
+#695 := (= #424 #706)
+#700 := (= #423 0::int)
+#697 := (mod -1::int 1::int)
+#693 := (= #697 0::int)
+#694 := [rewrite]: #693
+#698 := (= #423 #697)
+#554 := (= #338 1::int)
+#696 := [rewrite]: #554
+#552 := (= #422 -1::int)
+#553 := [rewrite]: #552
+#699 := [monotonicity #553 #696]: #698
+#701 := [trans #699 #694]: #700
+#702 := [monotonicity #701]: #695
+#533 := [trans #702 #704]: #531
+#687 := [monotonicity #533]: #534
+#705 := (iff #425 #175)
+#709 := (= #428 #172)
+#707 := (= #428 #706)
+#713 := (= #427 0::int)
+#714 := (* -1::int 0::int)
+#712 := (= #714 0::int)
+#716 := [rewrite]: #712
+#715 := (= #427 #714)
+#435 := (= #426 0::int)
+#436 := [rewrite]: #435
+#430 := [monotonicity #436]: #715
+#717 := [trans #430 #716]: #713
+#708 := [monotonicity #717]: #707
+#710 := [trans #708 #704]: #709
+#711 := [monotonicity #710]: #705
+#433 := (iff #413 false)
+#444 := (or false false)
+#719 := (iff #444 false)
+#723 := [rewrite]: #719
+#718 := (iff #413 #444)
+#721 := (iff #745 false)
+#1 := true
+#727 := (not true)
+#734 := (iff #727 false)
+#736 := [rewrite]: #734
+#344 := (iff #745 #727)
+#449 := (iff #743 true)
+#738 := (or false true)
+#447 := (iff #738 true)
+#448 := [rewrite]: #447
+#724 := (iff #743 #738)
+#385 := (iff #429 true)
+#726 := [rewrite]: #385
+#737 := (iff #408 false)
+#735 := [rewrite]: #737
+#725 := [monotonicity #735 #726]: #724
+#450 := [trans #725 #448]: #449
+#720 := [monotonicity #450]: #344
+#443 := [trans #720 #736]: #721
+#376 := (iff #739 false)
+#733 := (iff #739 #727)
+#731 := (iff #532 true)
+#391 := (or true true)
+#729 := (iff #391 true)
+#730 := [rewrite]: #729
+#392 := (iff #532 #391)
+#389 := (iff #402 true)
+#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
+#691 := [monotonicity #434 #711 #687]: #690
+#575 := [trans #691 #692]: #574
+#401 := (iff #746 false)
+#742 := [rewrite]: #401
+#535 := [monotonicity #742 #575]: #576
+#585 := [trans #535 #692]: #683
+#400 := (iff #748 #405)
+#741 := [rewrite]: #400
+#750 := (iff #749 false)
+#740 := [rewrite]: #750
+#684 := [monotonicity #740 #741 #585]: #681
+#587 := [trans #684 #584]: #586
+#592 := [monotonicity #587]: #591
+#686 := [trans #592 #682]: #591
+#590 := [quant-inst]: #589
+#673 := [mp #590 #686]: #588
+[unit-resolution #673 #764 #193]: false
+unsat
 24bbfda15529f2b078998e7a57a330dfd637cef3 282 0
 #2 := false
 #11 := 0::int
@@ -30437,6 +32383,336 @@
 #668 := [mp #574 #671]: #573
 [unit-resolution #668 #743 #168]: false
 unsat
+8144b6d093845ae8dee3931af66c9361d99b4453 329 0
+#2 := false
+#11 := 0::int
+decl f4 :: (-> int int int)
+#67 := -1::int
+#38 := 3::int
+#173 := (f4 3::int -1::int)
+#176 := (= #173 0::int)
+#189 := (not #176)
+#39 := 1::int
+#40 := (- 1::int)
+#41 := (f4 3::int #40)
+#42 := (= #41 0::int)
+#43 := (not #42)
+#192 := (iff #43 #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 #43 #184)
+#182 := (iff #42 #179)
+#180 := (iff #176 #179)
+#181 := [rewrite]: #180
+#177 := (iff #42 #176)
+#174 := (= #41 #173)
+#171 := (= #40 -1::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
+#170 := [asserted]: #43
+#194 := [mp #170 #193]: #189
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#29 := (f4 #8 #9)
+#759 := (pattern #29)
+#71 := (* -1::int #9)
+#68 := (* -1::int #8)
+#126 := (mod #68 #71)
+#271 := (+ #29 #126)
+#272 := (= #271 0::int)
+#30 := (mod #8 #9)
+#268 := (* -1::int #30)
+#269 := (+ #29 #268)
+#270 := (= #269 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)
+#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)
+#760 := (forall (vars (?v0 int) (?v1 int)) (:pat #759) #275)
+#278 := (forall (vars (?v0 int) (?v1 int)) #275)
+#763 := (iff #278 #760)
+#761 := (iff #275 #275)
+#762 := [refl]: #761
+#764 := [quant-intro #762]: #763
+#132 := (* -1::int #126)
+#235 := (ite #217 #30 #132)
+#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
+#101 := (not #100)
+#94 := (not #93)
+#104 := (and #94 #101)
+#90 := (not #89)
+#97 := (and #90 #94)
+#107 := (or #97 #104)
+#152 := (ite #107 #30 #132)
+#155 := (ite #12 0::int #152)
+#158 := (ite #13 #8 #155)
+#161 := (= #29 #158)
+#164 := (forall (vars (?v0 int) (?v1 int)) #161)
+#248 := (iff #164 #247)
+#245 := (iff #161 #244)
+#242 := (= #158 #241)
+#239 := (= #155 #238)
+#236 := (= #152 #235)
+#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
+#237 := [monotonicity #221]: #236
+#240 := [monotonicity #237]: #239
+#243 := [monotonicity #240]: #242
+#246 := [monotonicity #243]: #245
+#249 := [quant-intro #246]: #248
+#198 := (~ #164 #164)
+#195 := (~ #161 #161)
+#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)
+#167 := (iff #37 #164)
+#61 := (and #16 #18)
+#64 := (or #17 #61)
+#137 := (ite #64 #30 #132)
+#140 := (ite #12 0::int #137)
+#143 := (ite #13 #8 #140)
+#146 := (= #29 #143)
+#149 := (forall (vars (?v0 int) (?v1 int)) #146)
+#165 := (iff #149 #164)
+#162 := (iff #146 #161)
+#159 := (= #143 #158)
+#156 := (= #140 #155)
+#153 := (= #137 #152)
+#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
+#154 := [monotonicity #109]: #153
+#157 := [monotonicity #154]: #156
+#160 := [monotonicity #157]: #159
+#163 := [monotonicity #160]: #162
+#166 := [quant-intro #163]: #165
+#150 := (iff #37 #149)
+#147 := (iff #36 #146)
+#144 := (= #35 #143)
+#141 := (= #34 #140)
+#138 := (= #33 #137)
+#135 := (= #32 #132)
+#129 := (- #126)
+#133 := (= #129 #132)
+#134 := [rewrite]: #133
+#130 := (= #32 #129)
+#127 := (= #31 #126)
+#72 := (= #23 #71)
+#73 := [rewrite]: #72
+#69 := (= #22 #68)
+#70 := [rewrite]: #69
+#128 := [monotonicity #70 #73]: #127
+#131 := [monotonicity #128]: #130
+#136 := [trans #131 #134]: #135
+#65 := (iff #20 #64)
+#62 := (iff #19 #61)
+#63 := [rewrite]: #62
+#66 := [monotonicity #63]: #65
+#139 := [monotonicity #66 #136]: #138
+#142 := [monotonicity #139]: #141
+#145 := [monotonicity #142]: #144
+#148 := [monotonicity #145]: #147
+#151 := [quant-intro #148]: #150
+#168 := [trans #151 #166]: #167
+#125 := [asserted]: #37
+#169 := [mp #125 #168]: #164
+#209 := [mp~ #169 #199]: #164
+#250 := [mp #209 #249]: #247
+#281 := [mp #250 #280]: #278
+#765 := [mp #281 #764]: #760
+#589 := (not #760)
+#590 := (or #589 #176)
+#339 := (* -1::int -1::int)
+#423 := (* -1::int 3::int)
+#424 := (mod #423 #339)
+#425 := (+ #173 #424)
+#416 := (= #425 0::int)
+#427 := (mod 3::int -1::int)
+#428 := (* -1::int #427)
+#429 := (+ #173 #428)
+#426 := (= #429 0::int)
+#430 := (<= -1::int 0::int)
+#409 := (<= 3::int 0::int)
+#744 := (or #409 #430)
+#746 := (not #744)
+#403 := (>= 3::int 0::int)
+#533 := (or #430 #403)
+#740 := (not #533)
+#414 := (or #740 #746)
+#415 := (ite #414 #426 #416)
+#747 := (= 3::int 0::int)
+#748 := (ite #747 #176 #415)
+#749 := (= 3::int #173)
+#750 := (= -1::int 0::int)
+#745 := (ite #750 #749 #748)
+#591 := (or #589 #745)
+#593 := (iff #591 #590)
+#683 := (iff #590 #590)
+#687 := [rewrite]: #683
+#588 := (iff #745 #176)
+#406 := (= #173 3::int)
+#682 := (ite false #406 #176)
+#585 := (iff #682 #176)
+#587 := [rewrite]: #585
+#685 := (iff #745 #682)
+#586 := (iff #748 #176)
+#691 := (ite false #176 #176)
+#693 := (iff #691 #176)
+#575 := [rewrite]: #693
+#536 := (iff #748 #691)
+#576 := (iff #415 #176)
+#692 := (iff #415 #691)
+#688 := (iff #416 #176)
+#534 := (= #425 #173)
+#707 := (+ #173 0::int)
+#704 := (= #707 #173)
+#705 := [rewrite]: #704
+#703 := (= #425 #707)
+#702 := (= #424 0::int)
+#553 := -3::int
+#699 := (mod -3::int 1::int)
+#695 := (= #699 0::int)
+#701 := [rewrite]: #695
+#700 := (= #424 #699)
+#697 := (= #339 1::int)
+#698 := [rewrite]: #697
+#554 := (= #423 -3::int)
+#555 := [rewrite]: #554
+#694 := [monotonicity #555 #698]: #700
+#696 := [trans #694 #701]: #702
+#532 := [monotonicity #696]: #703
+#535 := [trans #532 #705]: #534
+#689 := [monotonicity #535]: #688
+#706 := (iff #426 #176)
+#710 := (= #429 #173)
+#708 := (= #429 #707)
+#714 := (= #428 0::int)
+#715 := (* -1::int 0::int)
+#713 := (= #715 0::int)
+#717 := [rewrite]: #713
+#716 := (= #428 #715)
+#436 := (= #427 0::int)
+#437 := [rewrite]: #436
+#431 := [monotonicity #437]: #716
+#718 := [trans #431 #717]: #714
+#709 := [monotonicity #718]: #708
+#711 := [trans #709 #705]: #710
+#712 := [monotonicity #711]: #706
+#434 := (iff #414 false)
+#445 := (or false false)
+#720 := (iff #445 false)
+#724 := [rewrite]: #720
+#719 := (iff #414 #445)
+#722 := (iff #746 false)
+#1 := true
+#728 := (not true)
+#735 := (iff #728 false)
+#737 := [rewrite]: #735
+#345 := (iff #746 #728)
+#450 := (iff #744 true)
+#739 := (or false true)
+#448 := (iff #739 true)
+#449 := [rewrite]: #448
+#725 := (iff #744 #739)
+#386 := (iff #430 true)
+#727 := [rewrite]: #386
+#738 := (iff #409 false)
+#736 := [rewrite]: #738
+#726 := [monotonicity #736 #727]: #725
+#451 := [trans #726 #449]: #450
+#721 := [monotonicity #451]: #345
+#444 := [trans #721 #737]: #722
+#377 := (iff #740 false)
+#734 := (iff #740 #728)
+#732 := (iff #533 true)
+#392 := (or true true)
+#730 := (iff #392 true)
+#731 := [rewrite]: #730
+#393 := (iff #533 #392)
+#390 := (iff #403 true)
+#391 := [rewrite]: #390
+#729 := [monotonicity #727 #391]: #393
+#733 := [trans #729 #731]: #732
+#372 := [monotonicity #733]: #734
+#378 := [trans #372 #737]: #377
+#723 := [monotonicity #378 #444]: #719
+#435 := [trans #723 #724]: #434
+#690 := [monotonicity #435 #712 #689]: #692
+#577 := [trans #690 #575]: #576
+#402 := (iff #747 false)
+#743 := [rewrite]: #402
+#684 := [monotonicity #743 #577]: #536
+#681 := [trans #684 #575]: #586
+#401 := (iff #749 #406)
+#742 := [rewrite]: #401
+#751 := (iff #750 false)
+#741 := [rewrite]: #751
+#570 := [monotonicity #741 #742 #681]: #685
+#581 := [trans #570 #587]: #588
+#686 := [monotonicity #581]: #593
+#674 := [trans #686 #687]: #593
+#592 := [quant-inst]: #591
+#676 := [mp #592 #674]: #590
+[unit-resolution #676 #765 #194]: false
+unsat
 12ef6b4ae9675610f423327a4c22119f401927ad 307 0
 #2 := false
 #11 := 0::int
@@ -30745,6 +33021,313 @@
 #687 := [mp #582 #594]: #588
 [unit-resolution #687 #759 #194]: false
 unsat
+a20f91759967cf78b23640f832ca57a185975634 306 0
+#2 := false
+#11 := 0::int
+decl f4 :: (-> int int int)
+#67 := -1::int
+decl f5 :: int
+#38 := f5
+#173 := (f4 f5 -1::int)
+#176 := (= #173 0::int)
+#189 := (not #176)
+#39 := 1::int
+#40 := (- 1::int)
+#41 := (f4 f5 #40)
+#42 := (= #41 0::int)
+#43 := (not #42)
+#192 := (iff #43 #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 #43 #184)
+#182 := (iff #42 #179)
+#180 := (iff #176 #179)
+#181 := [rewrite]: #180
+#177 := (iff #42 #176)
+#174 := (= #41 #173)
+#171 := (= #40 -1::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
+#170 := [asserted]: #43
+#194 := [mp #170 #193]: #189
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#29 := (f4 #8 #9)
+#759 := (pattern #29)
+#71 := (* -1::int #9)
+#68 := (* -1::int #8)
+#126 := (mod #68 #71)
+#271 := (+ #29 #126)
+#272 := (= #271 0::int)
+#30 := (mod #8 #9)
+#268 := (* -1::int #30)
+#269 := (+ #29 #268)
+#270 := (= #269 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)
+#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)
+#760 := (forall (vars (?v0 int) (?v1 int)) (:pat #759) #275)
+#278 := (forall (vars (?v0 int) (?v1 int)) #275)
+#763 := (iff #278 #760)
+#761 := (iff #275 #275)
+#762 := [refl]: #761
+#764 := [quant-intro #762]: #763
+#132 := (* -1::int #126)
+#235 := (ite #217 #30 #132)
+#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
+#101 := (not #100)
+#94 := (not #93)
+#104 := (and #94 #101)
+#90 := (not #89)
+#97 := (and #90 #94)
+#107 := (or #97 #104)
+#152 := (ite #107 #30 #132)
+#155 := (ite #12 0::int #152)
+#158 := (ite #13 #8 #155)
+#161 := (= #29 #158)
+#164 := (forall (vars (?v0 int) (?v1 int)) #161)
+#248 := (iff #164 #247)
+#245 := (iff #161 #244)
+#242 := (= #158 #241)
+#239 := (= #155 #238)
+#236 := (= #152 #235)
+#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
+#237 := [monotonicity #221]: #236
+#240 := [monotonicity #237]: #239
+#243 := [monotonicity #240]: #242
+#246 := [monotonicity #243]: #245
+#249 := [quant-intro #246]: #248
+#198 := (~ #164 #164)
+#195 := (~ #161 #161)
+#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)
+#167 := (iff #37 #164)
+#61 := (and #16 #18)
+#64 := (or #17 #61)
+#137 := (ite #64 #30 #132)
+#140 := (ite #12 0::int #137)
+#143 := (ite #13 #8 #140)
+#146 := (= #29 #143)
+#149 := (forall (vars (?v0 int) (?v1 int)) #146)
+#165 := (iff #149 #164)
+#162 := (iff #146 #161)
+#159 := (= #143 #158)
+#156 := (= #140 #155)
+#153 := (= #137 #152)
+#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
+#154 := [monotonicity #109]: #153
+#157 := [monotonicity #154]: #156
+#160 := [monotonicity #157]: #159
+#163 := [monotonicity #160]: #162
+#166 := [quant-intro #163]: #165
+#150 := (iff #37 #149)
+#147 := (iff #36 #146)
+#144 := (= #35 #143)
+#141 := (= #34 #140)
+#138 := (= #33 #137)
+#135 := (= #32 #132)
+#129 := (- #126)
+#133 := (= #129 #132)
+#134 := [rewrite]: #133
+#130 := (= #32 #129)
+#127 := (= #31 #126)
+#72 := (= #23 #71)
+#73 := [rewrite]: #72
+#69 := (= #22 #68)
+#70 := [rewrite]: #69
+#128 := [monotonicity #70 #73]: #127
+#131 := [monotonicity #128]: #130
+#136 := [trans #131 #134]: #135
+#65 := (iff #20 #64)
+#62 := (iff #19 #61)
+#63 := [rewrite]: #62
+#66 := [monotonicity #63]: #65
+#139 := [monotonicity #66 #136]: #138
+#142 := [monotonicity #139]: #141
+#145 := [monotonicity #142]: #144
+#148 := [monotonicity #145]: #147
+#151 := [quant-intro #148]: #150
+#168 := [trans #151 #166]: #167
+#125 := [asserted]: #37
+#169 := [mp #125 #168]: #164
+#209 := [mp~ #169 #199]: #164
+#250 := [mp #209 #249]: #247
+#281 := [mp #250 #280]: #278
+#765 := [mp #281 #764]: #760
+#703 := (not #760)
+#532 := (or #703 #176)
+#339 := (* -1::int -1::int)
+#423 := (* -1::int f5)
+#424 := (mod #423 #339)
+#425 := (+ #173 #424)
+#416 := (= #425 0::int)
+#427 := (mod f5 -1::int)
+#428 := (* -1::int #427)
+#429 := (+ #173 #428)
+#426 := (= #429 0::int)
+#430 := (<= -1::int 0::int)
+#409 := (<= f5 0::int)
+#744 := (or #409 #430)
+#746 := (not #744)
+#403 := (>= f5 0::int)
+#533 := (or #430 #403)
+#740 := (not #533)
+#414 := (or #740 #746)
+#415 := (ite #414 #426 #416)
+#747 := (= f5 0::int)
+#748 := (ite #747 #176 #415)
+#749 := (= f5 #173)
+#750 := (= -1::int 0::int)
+#745 := (ite #750 #749 #748)
+#534 := (or #703 #745)
+#688 := (iff #534 #532)
+#691 := (iff #532 #532)
+#692 := [rewrite]: #691
+#702 := (iff #745 #176)
+#699 := (ite false #749 #176)
+#695 := (iff #699 #176)
+#701 := [rewrite]: #695
+#700 := (iff #745 #699)
+#697 := (iff #748 #176)
+#706 := (ite #747 #176 #176)
+#554 := (iff #706 #176)
+#555 := [rewrite]: #554
+#712 := (iff #748 #706)
+#710 := (iff #415 #176)
+#707 := (ite false #426 #176)
+#704 := (iff #707 #176)
+#705 := [rewrite]: #704
+#708 := (iff #415 #707)
+#714 := (iff #416 #176)
+#713 := (= #425 #173)
+#436 := (+ #173 0::int)
+#716 := (= #436 #173)
+#431 := [rewrite]: #716
+#437 := (= #425 #436)
+#434 := (= #424 0::int)
+#445 := (mod #423 1::int)
+#720 := (= #445 0::int)
+#724 := [rewrite]: #720
+#719 := (= #424 #445)
+#722 := (= #339 1::int)
+#444 := [rewrite]: #722
+#723 := [monotonicity #444]: #719
+#435 := [trans #723 #724]: #434
+#715 := [monotonicity #435]: #437
+#717 := [trans #715 #431]: #713
+#718 := [monotonicity #717]: #714
+#345 := (iff #414 false)
+#726 := (or false false)
+#450 := (iff #726 false)
+#451 := [rewrite]: #450
+#448 := (iff #414 #726)
+#739 := (iff #746 false)
+#1 := true
+#392 := (not true)
+#730 := (iff #392 false)
+#731 := [rewrite]: #730
+#738 := (iff #746 #392)
+#377 := (iff #744 true)
+#728 := (or #409 true)
+#735 := (iff #728 true)
+#737 := [rewrite]: #735
+#734 := (iff #744 #728)
+#401 := (iff #430 true)
+#406 := [rewrite]: #401
+#372 := [monotonicity #406]: #734
+#378 := [trans #372 #737]: #377
+#736 := [monotonicity #378]: #738
+#725 := [trans #736 #731]: #739
+#732 := (iff #740 false)
+#393 := (iff #740 #392)
+#390 := (iff #533 true)
+#742 := (or true #403)
+#386 := (iff #742 true)
+#727 := [rewrite]: #386
+#402 := (iff #533 #742)
+#743 := [monotonicity #406]: #402
+#391 := [trans #743 #727]: #390
+#729 := [monotonicity #391]: #393
+#733 := [trans #729 #731]: #732
+#449 := [monotonicity #733 #725]: #448
+#721 := [trans #449 #451]: #345
+#709 := [monotonicity #721 #718]: #708
+#711 := [trans #709 #705]: #710
+#553 := [monotonicity #711]: #712
+#698 := [trans #553 #555]: #697
+#751 := (iff #750 false)
+#741 := [rewrite]: #751
+#694 := [monotonicity #741 #698]: #700
+#696 := [trans #694 #701]: #702
+#689 := [monotonicity #696]: #688
+#690 := [trans #689 #692]: #688
+#535 := [quant-inst]: #534
+#693 := [mp #535 #690]: #532
+[unit-resolution #693 #765 #194]: false
+unsat
 35df6ca3b9d8297d2ae93fc18a5a2fbecd67e27f 296 0
 #2 := false
 #11 := 0::int
@@ -31329,6 +33912,306 @@
 #643 := [mp #568 #651]: #566
 [unit-resolution #643 #733 #169]: false
 unsat
+450a206c661160bbd83de0094374a847857dd804 299 0
+#2 := false
+#11 := 0::int
+decl f4 :: (-> int int int)
+#38 := 3::int
+#39 := (f4 0::int 3::int)
+#40 := (= #39 0::int)
+#41 := (not #40)
+#168 := [asserted]: #41
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#29 := (f4 #8 #9)
+#744 := (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)
+#745 := (forall (vars (?v0 int) (?v1 int)) (:pat #744) #259)
+#262 := (forall (vars (?v0 int) (?v1 int)) #259)
+#748 := (iff #262 #745)
+#746 := (iff #259 #259)
+#747 := [refl]: #746
+#749 := [quant-intro #747]: #748
+#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
+#750 := [mp #265 #749]: #745
+#666 := (not #745)
+#667 := (or #666 #40)
+#323 := (* -1::int 3::int)
+#408 := (* -1::int 0::int)
+#409 := (mod #408 #323)
+#410 := (+ #39 #409)
+#401 := (= #410 0::int)
+#412 := (mod 0::int 3::int)
+#413 := (* -1::int #412)
+#341 := (+ #39 #413)
+#414 := (= #341 0::int)
+#411 := (<= 3::int 0::int)
+#415 := (<= 0::int 0::int)
+#394 := (or #415 #411)
+#729 := (not #394)
+#731 := (>= 0::int 0::int)
+#388 := (or #411 #731)
+#518 := (not #388)
+#725 := (or #518 #729)
+#399 := (ite #725 #414 #401)
+#400 := (= 0::int 0::int)
+#732 := (ite #400 #40 #399)
+#169 := (= 0::int #39)
+#733 := (= 3::int 0::int)
+#734 := (ite #733 #169 #732)
+#670 := (or #666 #734)
+#570 := (iff #670 #667)
+#573 := (iff #667 #667)
+#566 := [rewrite]: #573
+#669 := (iff #734 #40)
+#687 := (ite false #40 #40)
+#517 := (iff #687 #40)
+#519 := [rewrite]: #517
+#562 := (iff #734 #687)
+#560 := (iff #732 #40)
+#1 := true
+#674 := (ite true #40 #40)
+#675 := (iff #674 #40)
+#678 := [rewrite]: #675
+#676 := (iff #732 #674)
+#520 := (iff #399 #40)
+#681 := (iff #399 #687)
+#680 := (iff #401 #40)
+#685 := (= #410 #39)
+#701 := (+ #39 0::int)
+#702 := (= #701 #39)
+#699 := [rewrite]: #702
+#683 := (= #410 #701)
+#540 := (= #409 0::int)
+#689 := -3::int
+#696 := (mod 0::int -3::int)
+#538 := (= #696 0::int)
+#539 := [rewrite]: #538
+#691 := (= #409 #696)
+#690 := (= #323 -3::int)
+#695 := [rewrite]: #690
+#420 := (= #408 0::int)
+#421 := [rewrite]: #420
+#697 := [monotonicity #421 #695]: #691
+#682 := [trans #697 #539]: #540
+#684 := [monotonicity #682]: #683
+#679 := [trans #684 #699]: #685
+#686 := [monotonicity #679]: #680
+#693 := (iff #414 #40)
+#703 := (= #341 #39)
+#416 := (= #341 #701)
+#422 := (= #413 0::int)
+#709 := (= #413 #408)
+#708 := (= #412 0::int)
+#705 := [rewrite]: #708
+#419 := [monotonicity #705]: #709
+#700 := [trans #419 #421]: #422
+#698 := [monotonicity #700]: #416
+#692 := [trans #698 #699]: #703
+#694 := [monotonicity #692]: #693
+#430 := (iff #725 false)
+#436 := (or false false)
+#707 := (iff #436 false)
+#429 := [rewrite]: #707
+#329 := (iff #725 #436)
+#434 := (iff #729 false)
+#714 := (not true)
+#717 := (iff #714 false)
+#718 := [rewrite]: #717
+#711 := (iff #729 #714)
+#724 := (iff #394 true)
+#722 := (or true false)
+#723 := (iff #722 true)
+#721 := [rewrite]: #723
+#362 := (iff #394 #722)
+#386 := (iff #411 false)
+#391 := [rewrite]: #386
+#357 := (iff #415 true)
+#720 := [rewrite]: #357
+#363 := [monotonicity #720 #391]: #362
+#710 := [trans #363 #721]: #724
+#433 := [monotonicity #710]: #711
+#435 := [trans #433 #718]: #434
+#713 := (iff #518 false)
+#715 := (iff #518 #714)
+#377 := (iff #388 true)
+#728 := (or false true)
+#375 := (iff #728 true)
+#376 := [rewrite]: #375
+#371 := (iff #388 #728)
+#727 := (iff #731 true)
+#387 := [rewrite]: #727
+#712 := [monotonicity #391 #387]: #371
+#378 := [trans #712 #376]: #377
+#716 := [monotonicity #378]: #715
+#719 := [trans #716 #718]: #713
+#706 := [monotonicity #719 #435]: #329
+#704 := [trans #706 #429]: #430
+#688 := [monotonicity #704 #694 #686]: #681
+#673 := [trans #688 #519]: #520
+#736 := (iff #400 true)
+#726 := [rewrite]: #736
+#677 := [monotonicity #726 #673]: #676
+#561 := [trans #677 #678]: #560
+#175 := (iff #169 #40)
+#176 := [rewrite]: #175
+#735 := (iff #733 false)
+#730 := [rewrite]: #735
+#521 := [monotonicity #730 #176 #561]: #562
+#571 := [trans #521 #519]: #669
+#572 := [monotonicity #571]: #570
+#574 := [trans #572 #566]: #570
+#555 := [quant-inst]: #670
+#575 := [mp #555 #574]: #667
+[unit-resolution #575 #750 #168]: false
+unsat
 fc9ffeceefa2f1854d4c8adb404766a307c30ad9 310 0
 #2 := false
 #41 := 1::int
@@ -31640,6 +34523,325 @@
 #627 := [mp #637 #626]: #634
 [unit-resolution #627 #734 #170]: false
 unsat
+b8be49aa798ad99c4b186f0fa5e2759eac880cd8 318 0
+#2 := false
+#11 := 0::int
+decl f4 :: (-> int int int)
+#170 := -3::int
+#173 := (f4 0::int -3::int)
+#176 := (= #173 0::int)
+#189 := (not #176)
+#38 := 3::int
+#39 := (- 3::int)
+#40 := (f4 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)
+#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
+#578 := (not #761)
+#537 := (or #578 #176)
+#339 := (* -1::int -3::int)
+#424 := (* -1::int 0::int)
+#425 := (mod #424 #339)
+#426 := (+ #173 #425)
+#417 := (= #426 0::int)
+#428 := (mod 0::int -3::int)
+#429 := (* -1::int #428)
+#357 := (+ #173 #429)
+#430 := (= #357 0::int)
+#427 := (<= -3::int 0::int)
+#431 := (<= 0::int 0::int)
+#410 := (or #431 #427)
+#745 := (not #410)
+#747 := (>= 0::int 0::int)
+#404 := (or #427 #747)
+#534 := (not #404)
+#741 := (or #534 #745)
+#415 := (ite #741 #430 #417)
+#416 := (= 0::int 0::int)
+#748 := (ite #416 #176 #415)
+#749 := (= -3::int 0::int)
+#750 := (ite #749 #179 #748)
+#685 := (or #578 #750)
+#682 := (iff #685 #537)
+#686 := (iff #537 #537)
+#571 := [rewrite]: #686
+#576 := (iff #750 #176)
+#701 := (ite false #176 #176)
+#702 := (iff #701 #176)
+#703 := [rewrite]: #702
+#691 := (iff #750 #701)
+#692 := (iff #748 #176)
+#1 := true
+#533 := (ite true #176 #176)
+#689 := (iff #533 #176)
+#690 := [rewrite]: #689
+#535 := (iff #748 #533)
+#697 := (iff #415 #176)
+#695 := (iff #415 #701)
+#699 := (iff #417 #176)
+#556 := (= #426 #173)
+#437 := (+ #173 0::int)
+#717 := (= #437 #173)
+#432 := [rewrite]: #717
+#554 := (= #426 #437)
+#707 := (= #425 0::int)
+#710 := (mod 0::int 3::int)
+#711 := (= #710 0::int)
+#712 := [rewrite]: #711
+#705 := (= #425 #710)
+#708 := (= #339 3::int)
+#709 := [rewrite]: #708
+#721 := (= #424 0::int)
+#725 := [rewrite]: #721
+#706 := [monotonicity #725 #709]: #705
+#713 := [trans #706 #712]: #707
+#555 := [monotonicity #713]: #554
+#698 := [trans #555 #432]: #556
+#700 := [monotonicity #698]: #699
+#715 := (iff #430 #176)
+#714 := (= #357 #173)
+#438 := (= #357 #437)
+#435 := (= #429 0::int)
+#720 := (= #429 #424)
+#445 := (= #428 0::int)
+#446 := [rewrite]: #445
+#724 := [monotonicity #446]: #720
+#436 := [trans #724 #725]: #435
+#716 := [monotonicity #436]: #438
+#718 := [trans #716 #432]: #714
+#719 := [monotonicity #718]: #715
+#722 := (iff #741 false)
+#449 := (or false false)
+#452 := (iff #449 false)
+#345 := [rewrite]: #452
+#450 := (iff #741 #449)
+#726 := (iff #745 false)
+#730 := (not true)
+#733 := (iff #730 false)
+#734 := [rewrite]: #733
+#737 := (iff #745 #730)
+#379 := (iff #410 true)
+#744 := (or true true)
+#391 := (iff #744 true)
+#392 := [rewrite]: #391
+#738 := (iff #410 #744)
+#402 := (iff #427 true)
+#407 := [rewrite]: #402
+#373 := (iff #431 true)
+#736 := [rewrite]: #373
+#378 := [monotonicity #736 #407]: #738
+#739 := [trans #378 #392]: #379
+#740 := [monotonicity #739]: #737
+#727 := [trans #740 #734]: #726
+#729 := (iff #534 false)
+#731 := (iff #534 #730)
+#393 := (iff #404 true)
+#387 := (iff #404 #744)
+#743 := (iff #747 true)
+#403 := [rewrite]: #743
+#728 := [monotonicity #407 #403]: #387
+#394 := [trans #728 #392]: #393
+#732 := [monotonicity #394]: #731
+#735 := [trans #732 #734]: #729
+#451 := [monotonicity #735 #727]: #450
+#723 := [trans #451 #345]: #722
+#696 := [monotonicity #723 #719 #700]: #695
+#704 := [trans #696 #703]: #697
+#752 := (iff #416 true)
+#742 := [rewrite]: #752
+#536 := [monotonicity #742 #704]: #535
+#693 := [trans #536 #690]: #692
+#751 := (iff #749 false)
+#746 := [rewrite]: #751
+#694 := [monotonicity #746 #188 #693]: #691
+#577 := [trans #694 #703]: #576
+#683 := [monotonicity #577]: #682
+#586 := [trans #683 #571]: #682
+#587 := [quant-inst]: #685
+#588 := [mp #587 #586]: #537
+[unit-resolution #588 #766 #194]: false
+unsat
 334231bac1c10cb79d865cb2c86779cf7dd19631 324 0
 #2 := false
 #68 := -1::int
@@ -31965,6 +35167,329 @@
 #674 := [mp #681 #673]: #669
 [unit-resolution #674 #762 #198]: false
 unsat
+ed5e3568452fe6600500ab5071290e80a7175922 322 0
+#2 := false
+#38 := 1::int
+decl f4 :: (-> int int int)
+#39 := 3::int
+#40 := (f4 1::int 3::int)
+#41 := (= #40 1::int)
+#42 := (not #41)
+#169 := [asserted]: #42
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#29 := (f4 #8 #9)
+#744 := (pattern #29)
+#11 := 0::int
+#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
+#664 := (not #745)
+#660 := (or #664 #41)
+#324 := (* -1::int 3::int)
+#408 := (* -1::int 1::int)
+#409 := (mod #408 #324)
+#410 := (+ #40 #409)
+#401 := (= #410 0::int)
+#412 := (mod 1::int 3::int)
+#413 := (* -1::int #412)
+#414 := (+ #40 #413)
+#411 := (= #414 0::int)
+#415 := (<= 3::int 0::int)
+#394 := (<= 1::int 0::int)
+#729 := (or #394 #415)
+#731 := (not #729)
+#388 := (>= 1::int 0::int)
+#518 := (or #415 #388)
+#725 := (not #518)
+#399 := (or #725 #731)
+#400 := (ite #399 #411 #401)
+#732 := (= #40 0::int)
+#733 := (= 1::int 0::int)
+#734 := (ite #733 #732 #400)
+#170 := (= 1::int #40)
+#735 := (= 3::int 0::int)
+#730 := (ite #735 #170 #734)
+#665 := (or #664 #730)
+#655 := (iff #665 #660)
+#657 := (iff #660 #660)
+#651 := [rewrite]: #657
+#662 := (iff #730 #41)
+#659 := (ite false #41 #41)
+#650 := (iff #659 #41)
+#652 := [rewrite]: #650
+#661 := (iff #730 #659)
+#668 := (iff #734 #41)
+#575 := (ite false #732 #41)
+#578 := (iff #575 #41)
+#671 := [rewrite]: #578
+#576 := (iff #734 #575)
+#566 := (iff #400 #41)
+#562 := -2::int
+#521 := (= #40 -2::int)
+#1 := true
+#670 := (ite true #41 #521)
+#572 := (iff #670 #41)
+#573 := [rewrite]: #572
+#555 := (iff #400 #670)
+#666 := (iff #401 #521)
+#680 := 2::int
+#673 := (+ 2::int #40)
+#678 := (= #673 0::int)
+#669 := (iff #678 #521)
+#571 := [rewrite]: #669
+#560 := (iff #401 #678)
+#677 := (= #410 #673)
+#517 := (+ #40 2::int)
+#674 := (= #517 #673)
+#676 := [rewrite]: #674
+#519 := (= #410 #517)
+#681 := (= #409 2::int)
+#540 := -3::int
+#684 := (mod -1::int -3::int)
+#686 := (= #684 2::int)
+#687 := [rewrite]: #686
+#685 := (= #409 #684)
+#682 := (= #324 -3::int)
+#683 := [rewrite]: #682
+#422 := (= #408 -1::int)
+#700 := [rewrite]: #422
+#679 := [monotonicity #700 #683]: #685
+#688 := [trans #679 #687]: #681
+#520 := [monotonicity #688]: #519
+#675 := [trans #520 #676]: #677
+#561 := [monotonicity #675]: #560
+#667 := [trans #561 #571]: #666
+#538 := (iff #411 #41)
+#703 := (+ -1::int #40)
+#690 := (= #703 0::int)
+#691 := (iff #690 #41)
+#697 := [rewrite]: #691
+#695 := (iff #411 #690)
+#694 := (= #414 #703)
+#698 := (+ #40 -1::int)
+#692 := (= #698 #703)
+#693 := [rewrite]: #692
+#702 := (= #414 #698)
+#701 := (= #413 -1::int)
+#420 := (= #413 #408)
+#709 := (= #412 1::int)
+#419 := [rewrite]: #709
+#421 := [monotonicity #419]: #420
+#416 := [trans #421 #700]: #701
+#699 := [monotonicity #416]: #702
+#689 := [trans #699 #693]: #694
+#696 := [monotonicity #689]: #695
+#539 := [trans #696 #697]: #538
+#708 := (iff #399 true)
+#712 := (or false true)
+#377 := (iff #712 true)
+#378 := [rewrite]: #377
+#430 := (iff #399 #712)
+#707 := (iff #731 true)
+#434 := (not false)
+#330 := (iff #434 true)
+#706 := [rewrite]: #330
+#435 := (iff #731 #434)
+#711 := (iff #729 false)
+#363 := (or false false)
+#724 := (iff #363 false)
+#710 := [rewrite]: #724
+#723 := (iff #729 #363)
+#727 := (iff #415 false)
+#387 := [rewrite]: #727
+#722 := (iff #394 false)
+#362 := [rewrite]: #722
+#721 := [monotonicity #362 #387]: #723
+#433 := [trans #721 #710]: #711
+#436 := [monotonicity #433]: #435
+#429 := [trans #436 #706]: #707
+#357 := (iff #725 false)
+#716 := (not true)
+#713 := (iff #716 false)
+#719 := [rewrite]: #713
+#717 := (iff #725 #716)
+#714 := (iff #518 true)
+#375 := (iff #518 #712)
+#728 := (iff #388 true)
+#371 := [rewrite]: #728
+#376 := [monotonicity #387 #371]: #375
+#715 := [trans #376 #378]: #714
+#718 := [monotonicity #715]: #717
+#720 := [trans #718 #719]: #357
+#704 := [monotonicity #720 #429]: #430
+#705 := [trans #704 #378]: #708
+#570 := [monotonicity #705 #539 #667]: #555
+#574 := [trans #570 #573]: #566
+#386 := (iff #733 false)
+#391 := [rewrite]: #386
+#577 := [monotonicity #391 #574]: #576
+#672 := [trans #577 #671]: #668
+#176 := (iff #170 #41)
+#177 := [rewrite]: #176
+#736 := (iff #735 false)
+#726 := [rewrite]: #736
+#653 := [monotonicity #726 #177 #672]: #661
+#663 := [trans #653 #652]: #662
+#656 := [monotonicity #663]: #655
+#658 := [trans #656 #651]: #655
+#654 := [quant-inst]: #665
+#642 := [mp #654 #658]: #660
+[unit-resolution #642 #750 #169]: false
+unsat
 60a3e615568433502bc15c7d472c0ece644ee329 323 0
 #2 := false
 #68 := -1::int
@@ -32289,6 +35814,304 @@
 #673 := [mp #671 #684]: #672
 [unit-resolution #673 #762 #198]: false
 unsat
+1897b65589b01a36e8c901ddcdce3eb69e66fd4d 297 0
+#2 := false
+#11 := 0::int
+decl f4 :: (-> int int int)
+#38 := 3::int
+#39 := (f4 3::int 3::int)
+#40 := (= #39 0::int)
+#41 := (not #40)
+#168 := [asserted]: #41
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#29 := (f4 #8 #9)
+#744 := (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)
+#745 := (forall (vars (?v0 int) (?v1 int)) (:pat #744) #259)
+#262 := (forall (vars (?v0 int) (?v1 int)) #259)
+#748 := (iff #262 #745)
+#746 := (iff #259 #259)
+#747 := [refl]: #746
+#749 := [quant-intro #747]: #748
+#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
+#750 := [mp #265 #749]: #745
+#667 := (not #745)
+#670 := (or #667 #40)
+#323 := (* -1::int 3::int)
+#408 := (mod #323 #323)
+#409 := (+ #39 #408)
+#410 := (= #409 0::int)
+#401 := (mod 3::int 3::int)
+#412 := (* -1::int #401)
+#413 := (+ #39 #412)
+#341 := (= #413 0::int)
+#414 := (<= 3::int 0::int)
+#411 := (or #414 #414)
+#415 := (not #411)
+#394 := (>= 3::int 0::int)
+#729 := (or #414 #394)
+#731 := (not #729)
+#388 := (or #731 #415)
+#518 := (ite #388 #341 #410)
+#725 := (= 3::int 0::int)
+#399 := (ite #725 #40 #518)
+#400 := (= 3::int #39)
+#732 := (ite #725 #400 #399)
+#555 := (or #667 #732)
+#572 := (iff #555 #670)
+#566 := (iff #670 #670)
+#574 := [rewrite]: #566
+#571 := (iff #732 #40)
+#730 := (= #39 3::int)
+#560 := (ite false #730 #40)
+#521 := (iff #560 #40)
+#669 := [rewrite]: #521
+#561 := (iff #732 #560)
+#675 := (iff #399 #40)
+#520 := (ite false #40 #40)
+#676 := (iff #520 #40)
+#677 := [rewrite]: #676
+#673 := (iff #399 #520)
+#517 := (iff #518 #40)
+#1 := true
+#680 := (ite true #40 #40)
+#681 := (iff #680 #40)
+#688 := [rewrite]: #681
+#686 := (iff #518 #680)
+#685 := (iff #410 #40)
+#683 := (= #409 #39)
+#422 := (+ #39 0::int)
+#416 := (= #422 #39)
+#698 := [rewrite]: #416
+#540 := (= #409 #422)
+#538 := (= #408 0::int)
+#693 := -3::int
+#690 := (mod -3::int -3::int)
+#691 := (= #690 0::int)
+#697 := [rewrite]: #691
+#695 := (= #408 #690)
+#694 := (= #323 -3::int)
+#689 := [rewrite]: #694
+#696 := [monotonicity #689 #689]: #695
+#539 := [trans #696 #697]: #538
+#682 := [monotonicity #539]: #540
+#684 := [trans #682 #698]: #683
+#679 := [monotonicity #684]: #685
+#703 := (iff #341 #40)
+#702 := (= #413 #39)
+#700 := (= #413 #422)
+#420 := (= #412 0::int)
+#704 := (* -1::int 0::int)
+#709 := (= #704 0::int)
+#419 := [rewrite]: #709
+#708 := (= #412 #704)
+#429 := (= #401 0::int)
+#430 := [rewrite]: #429
+#705 := [monotonicity #430]: #708
+#421 := [trans #705 #419]: #420
+#701 := [monotonicity #421]: #700
+#699 := [trans #701 #698]: #702
+#692 := [monotonicity #699]: #703
+#706 := (iff #388 true)
+#387 := (or false true)
+#712 := (iff #387 true)
+#375 := [rewrite]: #712
+#436 := (iff #388 #387)
+#434 := (iff #415 true)
+#721 := (not false)
+#711 := (iff #721 true)
+#433 := [rewrite]: #711
+#724 := (iff #415 #721)
+#363 := (iff #411 false)
+#719 := (or false false)
+#722 := (iff #719 false)
+#362 := [rewrite]: #722
+#357 := (iff #411 #719)
+#726 := (iff #414 false)
+#386 := [rewrite]: #726
+#720 := [monotonicity #386 #386]: #357
+#723 := [trans #720 #362]: #363
+#710 := [monotonicity #723]: #724
+#435 := [trans #710 #433]: #434
+#718 := (iff #731 false)
+#378 := (not true)
+#716 := (iff #378 false)
+#717 := [rewrite]: #716
+#714 := (iff #731 #378)
+#376 := (iff #729 true)
+#728 := (iff #729 #387)
+#391 := (iff #394 true)
+#727 := [rewrite]: #391
+#371 := [monotonicity #386 #727]: #728
+#377 := [trans #371 #375]: #376
+#715 := [monotonicity #377]: #714
+#713 := [trans #715 #717]: #718
+#329 := [monotonicity #713 #435]: #436
+#707 := [trans #329 #375]: #706
+#687 := [monotonicity #707 #692 #679]: #686
+#519 := [trans #687 #688]: #517
+#733 := (iff #725 false)
+#734 := [rewrite]: #733
+#674 := [monotonicity #734 #519]: #673
+#678 := [trans #674 #677]: #675
+#735 := (iff #400 #730)
+#736 := [rewrite]: #735
+#562 := [monotonicity #734 #736 #678]: #561
+#666 := [trans #562 #669]: #571
+#573 := [monotonicity #666]: #572
+#575 := [trans #573 #574]: #572
+#570 := [quant-inst]: #555
+#576 := [mp #570 #575]: #670
+[unit-resolution #576 #750 #168]: false
+unsat
 0217fb3975a8ed7626e13b93029371f1151b43e4 326 0
 #2 := false
 #179 := -2::int
@@ -32616,6 +36439,335 @@
 #654 := [mp #652 #658]: #661
 [unit-resolution #654 #753 #190]: false
 unsat
+0950f059c8c59b021ddf8b08664ab965652bd6ab 328 0
+#2 := false
+#41 := 2::int
+decl f4 :: (-> int int int)
+#39 := 3::int
+#38 := 5::int
+#40 := (f4 5::int 3::int)
+#42 := (= #40 2::int)
+#43 := (not #42)
+#170 := [asserted]: #43
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#29 := (f4 #8 #9)
+#736 := (pattern #29)
+#11 := 0::int
+#67 := -1::int
+#71 := (* -1::int #9)
+#68 := (* -1::int #8)
+#126 := (mod #68 #71)
+#247 := (+ #29 #126)
+#248 := (= #247 0::int)
+#30 := (mod #8 #9)
+#244 := (* -1::int #30)
+#245 := (+ #29 #244)
+#246 := (= #245 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)
+#249 := (ite #193 #246 #248)
+#243 := (= #29 0::int)
+#12 := (= #8 0::int)
+#250 := (ite #12 #243 #249)
+#242 := (= #8 #29)
+#13 := (= #9 0::int)
+#251 := (ite #13 #242 #250)
+#737 := (forall (vars (?v0 int) (?v1 int)) (:pat #736) #251)
+#254 := (forall (vars (?v0 int) (?v1 int)) #251)
+#740 := (iff #254 #737)
+#738 := (iff #251 #251)
+#739 := [refl]: #738
+#741 := [quant-intro #739]: #740
+#132 := (* -1::int #126)
+#211 := (ite #193 #30 #132)
+#214 := (ite #12 0::int #211)
+#217 := (ite #13 #8 #214)
+#220 := (= #29 #217)
+#223 := (forall (vars (?v0 int) (?v1 int)) #220)
+#255 := (iff #223 #254)
+#252 := (iff #220 #251)
+#253 := [rewrite]: #252
+#256 := [quant-intro #253]: #255
+#101 := (not #100)
+#94 := (not #93)
+#104 := (and #94 #101)
+#90 := (not #89)
+#97 := (and #90 #94)
+#107 := (or #97 #104)
+#152 := (ite #107 #30 #132)
+#155 := (ite #12 0::int #152)
+#158 := (ite #13 #8 #155)
+#161 := (= #29 #158)
+#164 := (forall (vars (?v0 int) (?v1 int)) #161)
+#224 := (iff #164 #223)
+#221 := (iff #161 #220)
+#218 := (= #158 #217)
+#215 := (= #155 #214)
+#212 := (= #152 #211)
+#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
+#213 := [monotonicity #197]: #212
+#216 := [monotonicity #213]: #215
+#219 := [monotonicity #216]: #218
+#222 := [monotonicity #219]: #221
+#225 := [quant-intro #222]: #224
+#174 := (~ #164 #164)
+#171 := (~ #161 #161)
+#184 := [refl]: #171
+#175 := [nnf-pos #184]: #174
+#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)
+#167 := (iff #37 #164)
+#61 := (and #16 #18)
+#64 := (or #17 #61)
+#137 := (ite #64 #30 #132)
+#140 := (ite #12 0::int #137)
+#143 := (ite #13 #8 #140)
+#146 := (= #29 #143)
+#149 := (forall (vars (?v0 int) (?v1 int)) #146)
+#165 := (iff #149 #164)
+#162 := (iff #146 #161)
+#159 := (= #143 #158)
+#156 := (= #140 #155)
+#153 := (= #137 #152)
+#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
+#154 := [monotonicity #109]: #153
+#157 := [monotonicity #154]: #156
+#160 := [monotonicity #157]: #159
+#163 := [monotonicity #160]: #162
+#166 := [quant-intro #163]: #165
+#150 := (iff #37 #149)
+#147 := (iff #36 #146)
+#144 := (= #35 #143)
+#141 := (= #34 #140)
+#138 := (= #33 #137)
+#135 := (= #32 #132)
+#129 := (- #126)
+#133 := (= #129 #132)
+#134 := [rewrite]: #133
+#130 := (= #32 #129)
+#127 := (= #31 #126)
+#72 := (= #23 #71)
+#73 := [rewrite]: #72
+#69 := (= #22 #68)
+#70 := [rewrite]: #69
+#128 := [monotonicity #70 #73]: #127
+#131 := [monotonicity #128]: #130
+#136 := [trans #131 #134]: #135
+#65 := (iff #20 #64)
+#62 := (iff #19 #61)
+#63 := [rewrite]: #62
+#66 := [monotonicity #63]: #65
+#139 := [monotonicity #66 #136]: #138
+#142 := [monotonicity #139]: #141
+#145 := [monotonicity #142]: #144
+#148 := [monotonicity #145]: #147
+#151 := [quant-intro #148]: #150
+#168 := [trans #151 #166]: #167
+#125 := [asserted]: #37
+#169 := [mp #125 #168]: #164
+#185 := [mp~ #169 #175]: #164
+#226 := [mp #185 #225]: #223
+#257 := [mp #226 #256]: #254
+#742 := [mp #257 #741]: #737
+#643 := (not #737)
+#650 := (or #643 #42)
+#315 := (* -1::int 3::int)
+#400 := (* -1::int 5::int)
+#401 := (mod #400 #315)
+#402 := (+ #40 #401)
+#393 := (= #402 0::int)
+#404 := (mod 5::int 3::int)
+#405 := (* -1::int #404)
+#333 := (+ #40 #405)
+#406 := (= #333 0::int)
+#403 := (<= 3::int 0::int)
+#407 := (<= 5::int 0::int)
+#386 := (or #407 #403)
+#721 := (not #386)
+#723 := (>= 5::int 0::int)
+#380 := (or #403 #723)
+#510 := (not #380)
+#717 := (or #510 #721)
+#391 := (ite #717 #406 #393)
+#392 := (= #40 0::int)
+#724 := (= 5::int 0::int)
+#725 := (ite #724 #392 #391)
+#726 := (= 5::int #40)
+#727 := (= 3::int 0::int)
+#722 := (ite #727 #726 #725)
+#634 := (or #643 #722)
+#637 := (iff #634 #650)
+#639 := (iff #650 #650)
+#640 := [rewrite]: #639
+#648 := (iff #722 #42)
+#383 := (= #40 5::int)
+#656 := (ite false #383 #42)
+#646 := (iff #656 #42)
+#647 := [rewrite]: #646
+#652 := (iff #722 #656)
+#654 := (iff #725 #42)
+#651 := (ite false #392 #42)
+#642 := (iff #651 #42)
+#644 := [rewrite]: #642
+#653 := (iff #725 #651)
+#660 := (iff #391 #42)
+#562 := (= #40 -1::int)
+#1 := true
+#567 := (ite true #42 #562)
+#570 := (iff #567 #42)
+#663 := [rewrite]: #570
+#568 := (iff #391 #567)
+#558 := (iff #393 #562)
+#665 := 1::int
+#554 := (+ 1::int #40)
+#659 := (= #554 0::int)
+#564 := (iff #659 #562)
+#565 := [rewrite]: #564
+#662 := (iff #393 #659)
+#563 := (= #402 #554)
+#670 := (+ #40 1::int)
+#513 := (= #670 #554)
+#661 := [rewrite]: #513
+#552 := (= #402 #670)
+#669 := (= #401 1::int)
+#679 := -3::int
+#671 := -5::int
+#509 := (mod -5::int -3::int)
+#666 := (= #509 1::int)
+#668 := [rewrite]: #666
+#511 := (= #401 #509)
+#673 := (= #315 -3::int)
+#680 := [rewrite]: #673
+#672 := (= #400 -5::int)
+#678 := [rewrite]: #672
+#512 := [monotonicity #678 #680]: #511
+#667 := [trans #512 #668]: #669
+#553 := [monotonicity #667]: #552
+#658 := [trans #553 #661]: #563
+#547 := [monotonicity #658]: #662
+#566 := [trans #547 #565]: #558
+#676 := (iff #406 #42)
+#690 := -2::int
+#682 := (+ -2::int #40)
+#530 := (= #682 0::int)
+#674 := (iff #530 #42)
+#675 := [rewrite]: #674
+#531 := (iff #406 #530)
+#683 := (= #333 #682)
+#685 := (+ #40 -2::int)
+#687 := (= #685 #682)
+#688 := [rewrite]: #687
+#686 := (= #333 #685)
+#695 := (= #405 -2::int)
+#692 := (* -1::int 2::int)
+#694 := (= #692 -2::int)
+#691 := [rewrite]: #694
+#693 := (= #405 #692)
+#413 := (= #404 2::int)
+#414 := [rewrite]: #413
+#408 := [monotonicity #414]: #693
+#684 := [trans #408 #691]: #695
+#681 := [monotonicity #684]: #686
+#689 := [trans #681 #688]: #683
+#532 := [monotonicity #689]: #531
+#677 := [trans #532 #675]: #676
+#411 := (iff #717 true)
+#369 := (or false true)
+#707 := (iff #369 true)
+#708 := [rewrite]: #707
+#697 := (iff #717 #369)
+#696 := (iff #721 true)
+#321 := (not false)
+#421 := (iff #321 true)
+#422 := [rewrite]: #421
+#698 := (iff #721 #321)
+#427 := (iff #386 false)
+#716 := (or false false)
+#425 := (iff #716 false)
+#426 := [rewrite]: #425
+#702 := (iff #386 #716)
+#363 := (iff #403 false)
+#704 := [rewrite]: #363
+#715 := (iff #407 false)
+#713 := [rewrite]: #715
+#703 := [monotonicity #713 #704]: #702
+#428 := [trans #703 #426]: #427
+#699 := [monotonicity #428]: #698
+#700 := [trans #699 #422]: #696
+#354 := (iff #510 false)
+#705 := (not true)
+#712 := (iff #705 false)
+#714 := [rewrite]: #712
+#711 := (iff #510 #705)
+#709 := (iff #380 true)
+#370 := (iff #380 #369)
+#367 := (iff #723 true)
+#368 := [rewrite]: #367
+#706 := [monotonicity #704 #368]: #370
+#710 := [trans #706 #708]: #709
+#349 := [monotonicity #710]: #711
+#355 := [trans #349 #714]: #354
+#701 := [monotonicity #355 #700]: #697
+#412 := [trans #701 #708]: #411
+#569 := [monotonicity #412 #677 #566]: #568
+#664 := [trans #569 #663]: #660
+#379 := (iff #724 false)
+#720 := [rewrite]: #379
+#645 := [monotonicity #720 #664]: #653
+#655 := [trans #645 #644]: #654
+#378 := (iff #726 #383)
+#719 := [rewrite]: #378
+#728 := (iff #727 false)
+#718 := [rewrite]: #728
+#657 := [monotonicity #718 #719 #655]: #652
+#649 := [trans #657 #647]: #648
+#638 := [monotonicity #649]: #637
+#636 := [trans #638 #640]: #637
+#635 := [quant-inst]: #634
+#641 := [mp #635 #636]: #650
+[unit-resolution #641 #742 #170]: false
+unsat
 a5e7ca3a1a908b3cfa96261b62a5e5f564ff08b5 321 0
 #2 := false
 #67 := -1::int
@@ -33262,5923 +37414,6 @@
 #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)
-#66 := -1::int
-#38 := 1::int
-#172 := (f4 1::int -1::int)
-#175 := (= #172 0::int)
-#188 := (not #175)
-#39 := (- 1::int)
-#40 := (f4 1::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
-#580 := (not #759)
-#588 := (or #580 #175)
-#338 := (* -1::int -1::int)
-#422 := (* -1::int 1::int)
-#423 := (mod #422 #338)
-#424 := (+ #172 #423)
-#415 := (= #424 0::int)
-#426 := (mod 1::int -1::int)
-#427 := (* -1::int #426)
-#428 := (+ #172 #427)
-#425 := (= #428 0::int)
-#429 := (<= -1::int 0::int)
-#408 := (<= 1::int 0::int)
-#743 := (or #408 #429)
-#745 := (not #743)
-#402 := (>= 1::int 0::int)
-#532 := (or #429 #402)
-#739 := (not #532)
-#413 := (or #739 #745)
-#414 := (ite #413 #425 #415)
-#746 := (= 1::int 0::int)
-#747 := (ite #746 #175 #414)
-#748 := (= 1::int #172)
-#749 := (= -1::int 0::int)
-#744 := (ite #749 #748 #747)
-#589 := (or #580 #744)
-#591 := (iff #589 #588)
-#685 := (iff #588 #588)
-#682 := [rewrite]: #685
-#586 := (iff #744 #175)
-#405 := (= #172 1::int)
-#680 := (ite false #405 #175)
-#569 := (iff #680 #175)
-#584 := [rewrite]: #569
-#681 := (iff #744 #680)
-#683 := (iff #747 #175)
-#688 := (ite false #175 #175)
-#689 := (iff #688 #175)
-#692 := [rewrite]: #689
-#576 := (iff #747 #688)
-#574 := (iff #414 #175)
-#690 := (iff #414 #688)
-#534 := (iff #415 #175)
-#531 := (= #424 #172)
-#706 := (+ #172 0::int)
-#703 := (= #706 #172)
-#704 := [rewrite]: #703
-#695 := (= #424 #706)
-#700 := (= #423 0::int)
-#697 := (mod -1::int 1::int)
-#693 := (= #697 0::int)
-#694 := [rewrite]: #693
-#698 := (= #423 #697)
-#554 := (= #338 1::int)
-#696 := [rewrite]: #554
-#552 := (= #422 -1::int)
-#553 := [rewrite]: #552
-#699 := [monotonicity #553 #696]: #698
-#701 := [trans #699 #694]: #700
-#702 := [monotonicity #701]: #695
-#533 := [trans #702 #704]: #531
-#687 := [monotonicity #533]: #534
-#705 := (iff #425 #175)
-#709 := (= #428 #172)
-#707 := (= #428 #706)
-#713 := (= #427 0::int)
-#714 := (* -1::int 0::int)
-#712 := (= #714 0::int)
-#716 := [rewrite]: #712
-#715 := (= #427 #714)
-#435 := (= #426 0::int)
-#436 := [rewrite]: #435
-#430 := [monotonicity #436]: #715
-#717 := [trans #430 #716]: #713
-#708 := [monotonicity #717]: #707
-#710 := [trans #708 #704]: #709
-#711 := [monotonicity #710]: #705
-#433 := (iff #413 false)
-#444 := (or false false)
-#719 := (iff #444 false)
-#723 := [rewrite]: #719
-#718 := (iff #413 #444)
-#721 := (iff #745 false)
-#1 := true
-#727 := (not true)
-#734 := (iff #727 false)
-#736 := [rewrite]: #734
-#344 := (iff #745 #727)
-#449 := (iff #743 true)
-#738 := (or false true)
-#447 := (iff #738 true)
-#448 := [rewrite]: #447
-#724 := (iff #743 #738)
-#385 := (iff #429 true)
-#726 := [rewrite]: #385
-#737 := (iff #408 false)
-#735 := [rewrite]: #737
-#725 := [monotonicity #735 #726]: #724
-#450 := [trans #725 #448]: #449
-#720 := [monotonicity #450]: #344
-#443 := [trans #720 #736]: #721
-#376 := (iff #739 false)
-#733 := (iff #739 #727)
-#731 := (iff #532 true)
-#391 := (or true true)
-#729 := (iff #391 true)
-#730 := [rewrite]: #729
-#392 := (iff #532 #391)
-#389 := (iff #402 true)
-#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
-#691 := [monotonicity #434 #711 #687]: #690
-#575 := [trans #691 #692]: #574
-#401 := (iff #746 false)
-#742 := [rewrite]: #401
-#535 := [monotonicity #742 #575]: #576
-#585 := [trans #535 #692]: #683
-#400 := (iff #748 #405)
-#741 := [rewrite]: #400
-#750 := (iff #749 false)
-#740 := [rewrite]: #750
-#684 := [monotonicity #740 #741 #585]: #681
-#587 := [trans #684 #584]: #586
-#592 := [monotonicity #587]: #591
-#686 := [trans #592 #682]: #591
-#590 := [quant-inst]: #589
-#673 := [mp #590 #686]: #588
-[unit-resolution #673 #764 #193]: false
-unsat
-8144b6d093845ae8dee3931af66c9361d99b4453 329 0
-#2 := false
-#11 := 0::int
-decl f4 :: (-> int int int)
-#67 := -1::int
-#38 := 3::int
-#173 := (f4 3::int -1::int)
-#176 := (= #173 0::int)
-#189 := (not #176)
-#39 := 1::int
-#40 := (- 1::int)
-#41 := (f4 3::int #40)
-#42 := (= #41 0::int)
-#43 := (not #42)
-#192 := (iff #43 #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 #43 #184)
-#182 := (iff #42 #179)
-#180 := (iff #176 #179)
-#181 := [rewrite]: #180
-#177 := (iff #42 #176)
-#174 := (= #41 #173)
-#171 := (= #40 -1::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
-#170 := [asserted]: #43
-#194 := [mp #170 #193]: #189
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#29 := (f4 #8 #9)
-#759 := (pattern #29)
-#71 := (* -1::int #9)
-#68 := (* -1::int #8)
-#126 := (mod #68 #71)
-#271 := (+ #29 #126)
-#272 := (= #271 0::int)
-#30 := (mod #8 #9)
-#268 := (* -1::int #30)
-#269 := (+ #29 #268)
-#270 := (= #269 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)
-#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)
-#760 := (forall (vars (?v0 int) (?v1 int)) (:pat #759) #275)
-#278 := (forall (vars (?v0 int) (?v1 int)) #275)
-#763 := (iff #278 #760)
-#761 := (iff #275 #275)
-#762 := [refl]: #761
-#764 := [quant-intro #762]: #763
-#132 := (* -1::int #126)
-#235 := (ite #217 #30 #132)
-#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
-#101 := (not #100)
-#94 := (not #93)
-#104 := (and #94 #101)
-#90 := (not #89)
-#97 := (and #90 #94)
-#107 := (or #97 #104)
-#152 := (ite #107 #30 #132)
-#155 := (ite #12 0::int #152)
-#158 := (ite #13 #8 #155)
-#161 := (= #29 #158)
-#164 := (forall (vars (?v0 int) (?v1 int)) #161)
-#248 := (iff #164 #247)
-#245 := (iff #161 #244)
-#242 := (= #158 #241)
-#239 := (= #155 #238)
-#236 := (= #152 #235)
-#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
-#237 := [monotonicity #221]: #236
-#240 := [monotonicity #237]: #239
-#243 := [monotonicity #240]: #242
-#246 := [monotonicity #243]: #245
-#249 := [quant-intro #246]: #248
-#198 := (~ #164 #164)
-#195 := (~ #161 #161)
-#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)
-#167 := (iff #37 #164)
-#61 := (and #16 #18)
-#64 := (or #17 #61)
-#137 := (ite #64 #30 #132)
-#140 := (ite #12 0::int #137)
-#143 := (ite #13 #8 #140)
-#146 := (= #29 #143)
-#149 := (forall (vars (?v0 int) (?v1 int)) #146)
-#165 := (iff #149 #164)
-#162 := (iff #146 #161)
-#159 := (= #143 #158)
-#156 := (= #140 #155)
-#153 := (= #137 #152)
-#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
-#154 := [monotonicity #109]: #153
-#157 := [monotonicity #154]: #156
-#160 := [monotonicity #157]: #159
-#163 := [monotonicity #160]: #162
-#166 := [quant-intro #163]: #165
-#150 := (iff #37 #149)
-#147 := (iff #36 #146)
-#144 := (= #35 #143)
-#141 := (= #34 #140)
-#138 := (= #33 #137)
-#135 := (= #32 #132)
-#129 := (- #126)
-#133 := (= #129 #132)
-#134 := [rewrite]: #133
-#130 := (= #32 #129)
-#127 := (= #31 #126)
-#72 := (= #23 #71)
-#73 := [rewrite]: #72
-#69 := (= #22 #68)
-#70 := [rewrite]: #69
-#128 := [monotonicity #70 #73]: #127
-#131 := [monotonicity #128]: #130
-#136 := [trans #131 #134]: #135
-#65 := (iff #20 #64)
-#62 := (iff #19 #61)
-#63 := [rewrite]: #62
-#66 := [monotonicity #63]: #65
-#139 := [monotonicity #66 #136]: #138
-#142 := [monotonicity #139]: #141
-#145 := [monotonicity #142]: #144
-#148 := [monotonicity #145]: #147
-#151 := [quant-intro #148]: #150
-#168 := [trans #151 #166]: #167
-#125 := [asserted]: #37
-#169 := [mp #125 #168]: #164
-#209 := [mp~ #169 #199]: #164
-#250 := [mp #209 #249]: #247
-#281 := [mp #250 #280]: #278
-#765 := [mp #281 #764]: #760
-#589 := (not #760)
-#590 := (or #589 #176)
-#339 := (* -1::int -1::int)
-#423 := (* -1::int 3::int)
-#424 := (mod #423 #339)
-#425 := (+ #173 #424)
-#416 := (= #425 0::int)
-#427 := (mod 3::int -1::int)
-#428 := (* -1::int #427)
-#429 := (+ #173 #428)
-#426 := (= #429 0::int)
-#430 := (<= -1::int 0::int)
-#409 := (<= 3::int 0::int)
-#744 := (or #409 #430)
-#746 := (not #744)
-#403 := (>= 3::int 0::int)
-#533 := (or #430 #403)
-#740 := (not #533)
-#414 := (or #740 #746)
-#415 := (ite #414 #426 #416)
-#747 := (= 3::int 0::int)
-#748 := (ite #747 #176 #415)
-#749 := (= 3::int #173)
-#750 := (= -1::int 0::int)
-#745 := (ite #750 #749 #748)
-#591 := (or #589 #745)
-#593 := (iff #591 #590)
-#683 := (iff #590 #590)
-#687 := [rewrite]: #683
-#588 := (iff #745 #176)
-#406 := (= #173 3::int)
-#682 := (ite false #406 #176)
-#585 := (iff #682 #176)
-#587 := [rewrite]: #585
-#685 := (iff #745 #682)
-#586 := (iff #748 #176)
-#691 := (ite false #176 #176)
-#693 := (iff #691 #176)
-#575 := [rewrite]: #693
-#536 := (iff #748 #691)
-#576 := (iff #415 #176)
-#692 := (iff #415 #691)
-#688 := (iff #416 #176)
-#534 := (= #425 #173)
-#707 := (+ #173 0::int)
-#704 := (= #707 #173)
-#705 := [rewrite]: #704
-#703 := (= #425 #707)
-#702 := (= #424 0::int)
-#553 := -3::int
-#699 := (mod -3::int 1::int)
-#695 := (= #699 0::int)
-#701 := [rewrite]: #695
-#700 := (= #424 #699)
-#697 := (= #339 1::int)
-#698 := [rewrite]: #697
-#554 := (= #423 -3::int)
-#555 := [rewrite]: #554
-#694 := [monotonicity #555 #698]: #700
-#696 := [trans #694 #701]: #702
-#532 := [monotonicity #696]: #703
-#535 := [trans #532 #705]: #534
-#689 := [monotonicity #535]: #688
-#706 := (iff #426 #176)
-#710 := (= #429 #173)
-#708 := (= #429 #707)
-#714 := (= #428 0::int)
-#715 := (* -1::int 0::int)
-#713 := (= #715 0::int)
-#717 := [rewrite]: #713
-#716 := (= #428 #715)
-#436 := (= #427 0::int)
-#437 := [rewrite]: #436
-#431 := [monotonicity #437]: #716
-#718 := [trans #431 #717]: #714
-#709 := [monotonicity #718]: #708
-#711 := [trans #709 #705]: #710
-#712 := [monotonicity #711]: #706
-#434 := (iff #414 false)
-#445 := (or false false)
-#720 := (iff #445 false)
-#724 := [rewrite]: #720
-#719 := (iff #414 #445)
-#722 := (iff #746 false)
-#1 := true
-#728 := (not true)
-#735 := (iff #728 false)
-#737 := [rewrite]: #735
-#345 := (iff #746 #728)
-#450 := (iff #744 true)
-#739 := (or false true)
-#448 := (iff #739 true)
-#449 := [rewrite]: #448
-#725 := (iff #744 #739)
-#386 := (iff #430 true)
-#727 := [rewrite]: #386
-#738 := (iff #409 false)
-#736 := [rewrite]: #738
-#726 := [monotonicity #736 #727]: #725
-#451 := [trans #726 #449]: #450
-#721 := [monotonicity #451]: #345
-#444 := [trans #721 #737]: #722
-#377 := (iff #740 false)
-#734 := (iff #740 #728)
-#732 := (iff #533 true)
-#392 := (or true true)
-#730 := (iff #392 true)
-#731 := [rewrite]: #730
-#393 := (iff #533 #392)
-#390 := (iff #403 true)
-#391 := [rewrite]: #390
-#729 := [monotonicity #727 #391]: #393
-#733 := [trans #729 #731]: #732
-#372 := [monotonicity #733]: #734
-#378 := [trans #372 #737]: #377
-#723 := [monotonicity #378 #444]: #719
-#435 := [trans #723 #724]: #434
-#690 := [monotonicity #435 #712 #689]: #692
-#577 := [trans #690 #575]: #576
-#402 := (iff #747 false)
-#743 := [rewrite]: #402
-#684 := [monotonicity #743 #577]: #536
-#681 := [trans #684 #575]: #586
-#401 := (iff #749 #406)
-#742 := [rewrite]: #401
-#751 := (iff #750 false)
-#741 := [rewrite]: #751
-#570 := [monotonicity #741 #742 #681]: #685
-#581 := [trans #570 #587]: #588
-#686 := [monotonicity #581]: #593
-#674 := [trans #686 #687]: #593
-#592 := [quant-inst]: #591
-#676 := [mp #592 #674]: #590
-[unit-resolution #676 #765 #194]: false
-unsat
-a20f91759967cf78b23640f832ca57a185975634 306 0
-#2 := false
-#11 := 0::int
-decl f4 :: (-> int int int)
-#67 := -1::int
-decl f5 :: int
-#38 := f5
-#173 := (f4 f5 -1::int)
-#176 := (= #173 0::int)
-#189 := (not #176)
-#39 := 1::int
-#40 := (- 1::int)
-#41 := (f4 f5 #40)
-#42 := (= #41 0::int)
-#43 := (not #42)
-#192 := (iff #43 #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 #43 #184)
-#182 := (iff #42 #179)
-#180 := (iff #176 #179)
-#181 := [rewrite]: #180
-#177 := (iff #42 #176)
-#174 := (= #41 #173)
-#171 := (= #40 -1::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
-#170 := [asserted]: #43
-#194 := [mp #170 #193]: #189
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#29 := (f4 #8 #9)
-#759 := (pattern #29)
-#71 := (* -1::int #9)
-#68 := (* -1::int #8)
-#126 := (mod #68 #71)
-#271 := (+ #29 #126)
-#272 := (= #271 0::int)
-#30 := (mod #8 #9)
-#268 := (* -1::int #30)
-#269 := (+ #29 #268)
-#270 := (= #269 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)
-#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)
-#760 := (forall (vars (?v0 int) (?v1 int)) (:pat #759) #275)
-#278 := (forall (vars (?v0 int) (?v1 int)) #275)
-#763 := (iff #278 #760)
-#761 := (iff #275 #275)
-#762 := [refl]: #761
-#764 := [quant-intro #762]: #763
-#132 := (* -1::int #126)
-#235 := (ite #217 #30 #132)
-#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
-#101 := (not #100)
-#94 := (not #93)
-#104 := (and #94 #101)
-#90 := (not #89)
-#97 := (and #90 #94)
-#107 := (or #97 #104)
-#152 := (ite #107 #30 #132)
-#155 := (ite #12 0::int #152)
-#158 := (ite #13 #8 #155)
-#161 := (= #29 #158)
-#164 := (forall (vars (?v0 int) (?v1 int)) #161)
-#248 := (iff #164 #247)
-#245 := (iff #161 #244)
-#242 := (= #158 #241)
-#239 := (= #155 #238)
-#236 := (= #152 #235)
-#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
-#237 := [monotonicity #221]: #236
-#240 := [monotonicity #237]: #239
-#243 := [monotonicity #240]: #242
-#246 := [monotonicity #243]: #245
-#249 := [quant-intro #246]: #248
-#198 := (~ #164 #164)
-#195 := (~ #161 #161)
-#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)
-#167 := (iff #37 #164)
-#61 := (and #16 #18)
-#64 := (or #17 #61)
-#137 := (ite #64 #30 #132)
-#140 := (ite #12 0::int #137)
-#143 := (ite #13 #8 #140)
-#146 := (= #29 #143)
-#149 := (forall (vars (?v0 int) (?v1 int)) #146)
-#165 := (iff #149 #164)
-#162 := (iff #146 #161)
-#159 := (= #143 #158)
-#156 := (= #140 #155)
-#153 := (= #137 #152)
-#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
-#154 := [monotonicity #109]: #153
-#157 := [monotonicity #154]: #156
-#160 := [monotonicity #157]: #159
-#163 := [monotonicity #160]: #162
-#166 := [quant-intro #163]: #165
-#150 := (iff #37 #149)
-#147 := (iff #36 #146)
-#144 := (= #35 #143)
-#141 := (= #34 #140)
-#138 := (= #33 #137)
-#135 := (= #32 #132)
-#129 := (- #126)
-#133 := (= #129 #132)
-#134 := [rewrite]: #133
-#130 := (= #32 #129)
-#127 := (= #31 #126)
-#72 := (= #23 #71)
-#73 := [rewrite]: #72
-#69 := (= #22 #68)
-#70 := [rewrite]: #69
-#128 := [monotonicity #70 #73]: #127
-#131 := [monotonicity #128]: #130
-#136 := [trans #131 #134]: #135
-#65 := (iff #20 #64)
-#62 := (iff #19 #61)
-#63 := [rewrite]: #62
-#66 := [monotonicity #63]: #65
-#139 := [monotonicity #66 #136]: #138
-#142 := [monotonicity #139]: #141
-#145 := [monotonicity #142]: #144
-#148 := [monotonicity #145]: #147
-#151 := [quant-intro #148]: #150
-#168 := [trans #151 #166]: #167
-#125 := [asserted]: #37
-#169 := [mp #125 #168]: #164
-#209 := [mp~ #169 #199]: #164
-#250 := [mp #209 #249]: #247
-#281 := [mp #250 #280]: #278
-#765 := [mp #281 #764]: #760
-#703 := (not #760)
-#532 := (or #703 #176)
-#339 := (* -1::int -1::int)
-#423 := (* -1::int f5)
-#424 := (mod #423 #339)
-#425 := (+ #173 #424)
-#416 := (= #425 0::int)
-#427 := (mod f5 -1::int)
-#428 := (* -1::int #427)
-#429 := (+ #173 #428)
-#426 := (= #429 0::int)
-#430 := (<= -1::int 0::int)
-#409 := (<= f5 0::int)
-#744 := (or #409 #430)
-#746 := (not #744)
-#403 := (>= f5 0::int)
-#533 := (or #430 #403)
-#740 := (not #533)
-#414 := (or #740 #746)
-#415 := (ite #414 #426 #416)
-#747 := (= f5 0::int)
-#748 := (ite #747 #176 #415)
-#749 := (= f5 #173)
-#750 := (= -1::int 0::int)
-#745 := (ite #750 #749 #748)
-#534 := (or #703 #745)
-#688 := (iff #534 #532)
-#691 := (iff #532 #532)
-#692 := [rewrite]: #691
-#702 := (iff #745 #176)
-#699 := (ite false #749 #176)
-#695 := (iff #699 #176)
-#701 := [rewrite]: #695
-#700 := (iff #745 #699)
-#697 := (iff #748 #176)
-#706 := (ite #747 #176 #176)
-#554 := (iff #706 #176)
-#555 := [rewrite]: #554
-#712 := (iff #748 #706)
-#710 := (iff #415 #176)
-#707 := (ite false #426 #176)
-#704 := (iff #707 #176)
-#705 := [rewrite]: #704
-#708 := (iff #415 #707)
-#714 := (iff #416 #176)
-#713 := (= #425 #173)
-#436 := (+ #173 0::int)
-#716 := (= #436 #173)
-#431 := [rewrite]: #716
-#437 := (= #425 #436)
-#434 := (= #424 0::int)
-#445 := (mod #423 1::int)
-#720 := (= #445 0::int)
-#724 := [rewrite]: #720
-#719 := (= #424 #445)
-#722 := (= #339 1::int)
-#444 := [rewrite]: #722
-#723 := [monotonicity #444]: #719
-#435 := [trans #723 #724]: #434
-#715 := [monotonicity #435]: #437
-#717 := [trans #715 #431]: #713
-#718 := [monotonicity #717]: #714
-#345 := (iff #414 false)
-#726 := (or false false)
-#450 := (iff #726 false)
-#451 := [rewrite]: #450
-#448 := (iff #414 #726)
-#739 := (iff #746 false)
-#1 := true
-#392 := (not true)
-#730 := (iff #392 false)
-#731 := [rewrite]: #730
-#738 := (iff #746 #392)
-#377 := (iff #744 true)
-#728 := (or #409 true)
-#735 := (iff #728 true)
-#737 := [rewrite]: #735
-#734 := (iff #744 #728)
-#401 := (iff #430 true)
-#406 := [rewrite]: #401
-#372 := [monotonicity #406]: #734
-#378 := [trans #372 #737]: #377
-#736 := [monotonicity #378]: #738
-#725 := [trans #736 #731]: #739
-#732 := (iff #740 false)
-#393 := (iff #740 #392)
-#390 := (iff #533 true)
-#742 := (or true #403)
-#386 := (iff #742 true)
-#727 := [rewrite]: #386
-#402 := (iff #533 #742)
-#743 := [monotonicity #406]: #402
-#391 := [trans #743 #727]: #390
-#729 := [monotonicity #391]: #393
-#733 := [trans #729 #731]: #732
-#449 := [monotonicity #733 #725]: #448
-#721 := [trans #449 #451]: #345
-#709 := [monotonicity #721 #718]: #708
-#711 := [trans #709 #705]: #710
-#553 := [monotonicity #711]: #712
-#698 := [trans #553 #555]: #697
-#751 := (iff #750 false)
-#741 := [rewrite]: #751
-#694 := [monotonicity #741 #698]: #700
-#696 := [trans #694 #701]: #702
-#689 := [monotonicity #696]: #688
-#690 := [trans #689 #692]: #688
-#535 := [quant-inst]: #534
-#693 := [mp #535 #690]: #532
-[unit-resolution #693 #765 #194]: false
-unsat
-450a206c661160bbd83de0094374a847857dd804 299 0
-#2 := false
-#11 := 0::int
-decl f4 :: (-> int int int)
-#38 := 3::int
-#39 := (f4 0::int 3::int)
-#40 := (= #39 0::int)
-#41 := (not #40)
-#168 := [asserted]: #41
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#29 := (f4 #8 #9)
-#744 := (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)
-#745 := (forall (vars (?v0 int) (?v1 int)) (:pat #744) #259)
-#262 := (forall (vars (?v0 int) (?v1 int)) #259)
-#748 := (iff #262 #745)
-#746 := (iff #259 #259)
-#747 := [refl]: #746
-#749 := [quant-intro #747]: #748
-#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
-#750 := [mp #265 #749]: #745
-#666 := (not #745)
-#667 := (or #666 #40)
-#323 := (* -1::int 3::int)
-#408 := (* -1::int 0::int)
-#409 := (mod #408 #323)
-#410 := (+ #39 #409)
-#401 := (= #410 0::int)
-#412 := (mod 0::int 3::int)
-#413 := (* -1::int #412)
-#341 := (+ #39 #413)
-#414 := (= #341 0::int)
-#411 := (<= 3::int 0::int)
-#415 := (<= 0::int 0::int)
-#394 := (or #415 #411)
-#729 := (not #394)
-#731 := (>= 0::int 0::int)
-#388 := (or #411 #731)
-#518 := (not #388)
-#725 := (or #518 #729)
-#399 := (ite #725 #414 #401)
-#400 := (= 0::int 0::int)
-#732 := (ite #400 #40 #399)
-#169 := (= 0::int #39)
-#733 := (= 3::int 0::int)
-#734 := (ite #733 #169 #732)
-#670 := (or #666 #734)
-#570 := (iff #670 #667)
-#573 := (iff #667 #667)
-#566 := [rewrite]: #573
-#669 := (iff #734 #40)
-#687 := (ite false #40 #40)
-#517 := (iff #687 #40)
-#519 := [rewrite]: #517
-#562 := (iff #734 #687)
-#560 := (iff #732 #40)
-#1 := true
-#674 := (ite true #40 #40)
-#675 := (iff #674 #40)
-#678 := [rewrite]: #675
-#676 := (iff #732 #674)
-#520 := (iff #399 #40)
-#681 := (iff #399 #687)
-#680 := (iff #401 #40)
-#685 := (= #410 #39)
-#701 := (+ #39 0::int)
-#702 := (= #701 #39)
-#699 := [rewrite]: #702
-#683 := (= #410 #701)
-#540 := (= #409 0::int)
-#689 := -3::int
-#696 := (mod 0::int -3::int)
-#538 := (= #696 0::int)
-#539 := [rewrite]: #538
-#691 := (= #409 #696)
-#690 := (= #323 -3::int)
-#695 := [rewrite]: #690
-#420 := (= #408 0::int)
-#421 := [rewrite]: #420
-#697 := [monotonicity #421 #695]: #691
-#682 := [trans #697 #539]: #540
-#684 := [monotonicity #682]: #683
-#679 := [trans #684 #699]: #685
-#686 := [monotonicity #679]: #680
-#693 := (iff #414 #40)
-#703 := (= #341 #39)
-#416 := (= #341 #701)
-#422 := (= #413 0::int)
-#709 := (= #413 #408)
-#708 := (= #412 0::int)
-#705 := [rewrite]: #708
-#419 := [monotonicity #705]: #709
-#700 := [trans #419 #421]: #422
-#698 := [monotonicity #700]: #416
-#692 := [trans #698 #699]: #703
-#694 := [monotonicity #692]: #693
-#430 := (iff #725 false)
-#436 := (or false false)
-#707 := (iff #436 false)
-#429 := [rewrite]: #707
-#329 := (iff #725 #436)
-#434 := (iff #729 false)
-#714 := (not true)
-#717 := (iff #714 false)
-#718 := [rewrite]: #717
-#711 := (iff #729 #714)
-#724 := (iff #394 true)
-#722 := (or true false)
-#723 := (iff #722 true)
-#721 := [rewrite]: #723
-#362 := (iff #394 #722)
-#386 := (iff #411 false)
-#391 := [rewrite]: #386
-#357 := (iff #415 true)
-#720 := [rewrite]: #357
-#363 := [monotonicity #720 #391]: #362
-#710 := [trans #363 #721]: #724
-#433 := [monotonicity #710]: #711
-#435 := [trans #433 #718]: #434
-#713 := (iff #518 false)
-#715 := (iff #518 #714)
-#377 := (iff #388 true)
-#728 := (or false true)
-#375 := (iff #728 true)
-#376 := [rewrite]: #375
-#371 := (iff #388 #728)
-#727 := (iff #731 true)
-#387 := [rewrite]: #727
-#712 := [monotonicity #391 #387]: #371
-#378 := [trans #712 #376]: #377
-#716 := [monotonicity #378]: #715
-#719 := [trans #716 #718]: #713
-#706 := [monotonicity #719 #435]: #329
-#704 := [trans #706 #429]: #430
-#688 := [monotonicity #704 #694 #686]: #681
-#673 := [trans #688 #519]: #520
-#736 := (iff #400 true)
-#726 := [rewrite]: #736
-#677 := [monotonicity #726 #673]: #676
-#561 := [trans #677 #678]: #560
-#175 := (iff #169 #40)
-#176 := [rewrite]: #175
-#735 := (iff #733 false)
-#730 := [rewrite]: #735
-#521 := [monotonicity #730 #176 #561]: #562
-#571 := [trans #521 #519]: #669
-#572 := [monotonicity #571]: #570
-#574 := [trans #572 #566]: #570
-#555 := [quant-inst]: #670
-#575 := [mp #555 #574]: #667
-[unit-resolution #575 #750 #168]: false
-unsat
-b8be49aa798ad99c4b186f0fa5e2759eac880cd8 318 0
-#2 := false
-#11 := 0::int
-decl f4 :: (-> int int int)
-#170 := -3::int
-#173 := (f4 0::int -3::int)
-#176 := (= #173 0::int)
-#189 := (not #176)
-#38 := 3::int
-#39 := (- 3::int)
-#40 := (f4 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)
-#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
-#578 := (not #761)
-#537 := (or #578 #176)
-#339 := (* -1::int -3::int)
-#424 := (* -1::int 0::int)
-#425 := (mod #424 #339)
-#426 := (+ #173 #425)
-#417 := (= #426 0::int)
-#428 := (mod 0::int -3::int)
-#429 := (* -1::int #428)
-#357 := (+ #173 #429)
-#430 := (= #357 0::int)
-#427 := (<= -3::int 0::int)
-#431 := (<= 0::int 0::int)
-#410 := (or #431 #427)
-#745 := (not #410)
-#747 := (>= 0::int 0::int)
-#404 := (or #427 #747)
-#534 := (not #404)
-#741 := (or #534 #745)
-#415 := (ite #741 #430 #417)
-#416 := (= 0::int 0::int)
-#748 := (ite #416 #176 #415)
-#749 := (= -3::int 0::int)
-#750 := (ite #749 #179 #748)
-#685 := (or #578 #750)
-#682 := (iff #685 #537)
-#686 := (iff #537 #537)
-#571 := [rewrite]: #686
-#576 := (iff #750 #176)
-#701 := (ite false #176 #176)
-#702 := (iff #701 #176)
-#703 := [rewrite]: #702
-#691 := (iff #750 #701)
-#692 := (iff #748 #176)
-#1 := true
-#533 := (ite true #176 #176)
-#689 := (iff #533 #176)
-#690 := [rewrite]: #689
-#535 := (iff #748 #533)
-#697 := (iff #415 #176)
-#695 := (iff #415 #701)
-#699 := (iff #417 #176)
-#556 := (= #426 #173)
-#437 := (+ #173 0::int)
-#717 := (= #437 #173)
-#432 := [rewrite]: #717
-#554 := (= #426 #437)
-#707 := (= #425 0::int)
-#710 := (mod 0::int 3::int)
-#711 := (= #710 0::int)
-#712 := [rewrite]: #711
-#705 := (= #425 #710)
-#708 := (= #339 3::int)
-#709 := [rewrite]: #708
-#721 := (= #424 0::int)
-#725 := [rewrite]: #721
-#706 := [monotonicity #725 #709]: #705
-#713 := [trans #706 #712]: #707
-#555 := [monotonicity #713]: #554
-#698 := [trans #555 #432]: #556
-#700 := [monotonicity #698]: #699
-#715 := (iff #430 #176)
-#714 := (= #357 #173)
-#438 := (= #357 #437)
-#435 := (= #429 0::int)
-#720 := (= #429 #424)
-#445 := (= #428 0::int)
-#446 := [rewrite]: #445
-#724 := [monotonicity #446]: #720
-#436 := [trans #724 #725]: #435
-#716 := [monotonicity #436]: #438
-#718 := [trans #716 #432]: #714
-#719 := [monotonicity #718]: #715
-#722 := (iff #741 false)
-#449 := (or false false)
-#452 := (iff #449 false)
-#345 := [rewrite]: #452
-#450 := (iff #741 #449)
-#726 := (iff #745 false)
-#730 := (not true)
-#733 := (iff #730 false)
-#734 := [rewrite]: #733
-#737 := (iff #745 #730)
-#379 := (iff #410 true)
-#744 := (or true true)
-#391 := (iff #744 true)
-#392 := [rewrite]: #391
-#738 := (iff #410 #744)
-#402 := (iff #427 true)
-#407 := [rewrite]: #402
-#373 := (iff #431 true)
-#736 := [rewrite]: #373
-#378 := [monotonicity #736 #407]: #738
-#739 := [trans #378 #392]: #379
-#740 := [monotonicity #739]: #737
-#727 := [trans #740 #734]: #726
-#729 := (iff #534 false)
-#731 := (iff #534 #730)
-#393 := (iff #404 true)
-#387 := (iff #404 #744)
-#743 := (iff #747 true)
-#403 := [rewrite]: #743
-#728 := [monotonicity #407 #403]: #387
-#394 := [trans #728 #392]: #393
-#732 := [monotonicity #394]: #731
-#735 := [trans #732 #734]: #729
-#451 := [monotonicity #735 #727]: #450
-#723 := [trans #451 #345]: #722
-#696 := [monotonicity #723 #719 #700]: #695
-#704 := [trans #696 #703]: #697
-#752 := (iff #416 true)
-#742 := [rewrite]: #752
-#536 := [monotonicity #742 #704]: #535
-#693 := [trans #536 #690]: #692
-#751 := (iff #749 false)
-#746 := [rewrite]: #751
-#694 := [monotonicity #746 #188 #693]: #691
-#577 := [trans #694 #703]: #576
-#683 := [monotonicity #577]: #682
-#586 := [trans #683 #571]: #682
-#587 := [quant-inst]: #685
-#588 := [mp #587 #586]: #537
-[unit-resolution #588 #766 #194]: false
-unsat
-ed5e3568452fe6600500ab5071290e80a7175922 322 0
-#2 := false
-#38 := 1::int
-decl f4 :: (-> int int int)
-#39 := 3::int
-#40 := (f4 1::int 3::int)
-#41 := (= #40 1::int)
-#42 := (not #41)
-#169 := [asserted]: #42
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#29 := (f4 #8 #9)
-#744 := (pattern #29)
-#11 := 0::int
-#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
-#664 := (not #745)
-#660 := (or #664 #41)
-#324 := (* -1::int 3::int)
-#408 := (* -1::int 1::int)
-#409 := (mod #408 #324)
-#410 := (+ #40 #409)
-#401 := (= #410 0::int)
-#412 := (mod 1::int 3::int)
-#413 := (* -1::int #412)
-#414 := (+ #40 #413)
-#411 := (= #414 0::int)
-#415 := (<= 3::int 0::int)
-#394 := (<= 1::int 0::int)
-#729 := (or #394 #415)
-#731 := (not #729)
-#388 := (>= 1::int 0::int)
-#518 := (or #415 #388)
-#725 := (not #518)
-#399 := (or #725 #731)
-#400 := (ite #399 #411 #401)
-#732 := (= #40 0::int)
-#733 := (= 1::int 0::int)
-#734 := (ite #733 #732 #400)
-#170 := (= 1::int #40)
-#735 := (= 3::int 0::int)
-#730 := (ite #735 #170 #734)
-#665 := (or #664 #730)
-#655 := (iff #665 #660)
-#657 := (iff #660 #660)
-#651 := [rewrite]: #657
-#662 := (iff #730 #41)
-#659 := (ite false #41 #41)
-#650 := (iff #659 #41)
-#652 := [rewrite]: #650
-#661 := (iff #730 #659)
-#668 := (iff #734 #41)
-#575 := (ite false #732 #41)
-#578 := (iff #575 #41)
-#671 := [rewrite]: #578
-#576 := (iff #734 #575)
-#566 := (iff #400 #41)
-#562 := -2::int
-#521 := (= #40 -2::int)
-#1 := true
-#670 := (ite true #41 #521)
-#572 := (iff #670 #41)
-#573 := [rewrite]: #572
-#555 := (iff #400 #670)
-#666 := (iff #401 #521)
-#680 := 2::int
-#673 := (+ 2::int #40)
-#678 := (= #673 0::int)
-#669 := (iff #678 #521)
-#571 := [rewrite]: #669
-#560 := (iff #401 #678)
-#677 := (= #410 #673)
-#517 := (+ #40 2::int)
-#674 := (= #517 #673)
-#676 := [rewrite]: #674
-#519 := (= #410 #517)
-#681 := (= #409 2::int)
-#540 := -3::int
-#684 := (mod -1::int -3::int)
-#686 := (= #684 2::int)
-#687 := [rewrite]: #686
-#685 := (= #409 #684)
-#682 := (= #324 -3::int)
-#683 := [rewrite]: #682
-#422 := (= #408 -1::int)
-#700 := [rewrite]: #422
-#679 := [monotonicity #700 #683]: #685
-#688 := [trans #679 #687]: #681
-#520 := [monotonicity #688]: #519
-#675 := [trans #520 #676]: #677
-#561 := [monotonicity #675]: #560
-#667 := [trans #561 #571]: #666
-#538 := (iff #411 #41)
-#703 := (+ -1::int #40)
-#690 := (= #703 0::int)
-#691 := (iff #690 #41)
-#697 := [rewrite]: #691
-#695 := (iff #411 #690)
-#694 := (= #414 #703)
-#698 := (+ #40 -1::int)
-#692 := (= #698 #703)
-#693 := [rewrite]: #692
-#702 := (= #414 #698)
-#701 := (= #413 -1::int)
-#420 := (= #413 #408)
-#709 := (= #412 1::int)
-#419 := [rewrite]: #709
-#421 := [monotonicity #419]: #420
-#416 := [trans #421 #700]: #701
-#699 := [monotonicity #416]: #702
-#689 := [trans #699 #693]: #694
-#696 := [monotonicity #689]: #695
-#539 := [trans #696 #697]: #538
-#708 := (iff #399 true)
-#712 := (or false true)
-#377 := (iff #712 true)
-#378 := [rewrite]: #377
-#430 := (iff #399 #712)
-#707 := (iff #731 true)
-#434 := (not false)
-#330 := (iff #434 true)
-#706 := [rewrite]: #330
-#435 := (iff #731 #434)
-#711 := (iff #729 false)
-#363 := (or false false)
-#724 := (iff #363 false)
-#710 := [rewrite]: #724
-#723 := (iff #729 #363)
-#727 := (iff #415 false)
-#387 := [rewrite]: #727
-#722 := (iff #394 false)
-#362 := [rewrite]: #722
-#721 := [monotonicity #362 #387]: #723
-#433 := [trans #721 #710]: #711
-#436 := [monotonicity #433]: #435
-#429 := [trans #436 #706]: #707
-#357 := (iff #725 false)
-#716 := (not true)
-#713 := (iff #716 false)
-#719 := [rewrite]: #713
-#717 := (iff #725 #716)
-#714 := (iff #518 true)
-#375 := (iff #518 #712)
-#728 := (iff #388 true)
-#371 := [rewrite]: #728
-#376 := [monotonicity #387 #371]: #375
-#715 := [trans #376 #378]: #714
-#718 := [monotonicity #715]: #717
-#720 := [trans #718 #719]: #357
-#704 := [monotonicity #720 #429]: #430
-#705 := [trans #704 #378]: #708
-#570 := [monotonicity #705 #539 #667]: #555
-#574 := [trans #570 #573]: #566
-#386 := (iff #733 false)
-#391 := [rewrite]: #386
-#577 := [monotonicity #391 #574]: #576
-#672 := [trans #577 #671]: #668
-#176 := (iff #170 #41)
-#177 := [rewrite]: #176
-#736 := (iff #735 false)
-#726 := [rewrite]: #736
-#653 := [monotonicity #726 #177 #672]: #661
-#663 := [trans #653 #652]: #662
-#656 := [monotonicity #663]: #655
-#658 := [trans #656 #651]: #655
-#654 := [quant-inst]: #665
-#642 := [mp #654 #658]: #660
-[unit-resolution #642 #750 #169]: false
-unsat
-1897b65589b01a36e8c901ddcdce3eb69e66fd4d 297 0
-#2 := false
-#11 := 0::int
-decl f4 :: (-> int int int)
-#38 := 3::int
-#39 := (f4 3::int 3::int)
-#40 := (= #39 0::int)
-#41 := (not #40)
-#168 := [asserted]: #41
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#29 := (f4 #8 #9)
-#744 := (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)
-#745 := (forall (vars (?v0 int) (?v1 int)) (:pat #744) #259)
-#262 := (forall (vars (?v0 int) (?v1 int)) #259)
-#748 := (iff #262 #745)
-#746 := (iff #259 #259)
-#747 := [refl]: #746
-#749 := [quant-intro #747]: #748
-#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
-#750 := [mp #265 #749]: #745
-#667 := (not #745)
-#670 := (or #667 #40)
-#323 := (* -1::int 3::int)
-#408 := (mod #323 #323)
-#409 := (+ #39 #408)
-#410 := (= #409 0::int)
-#401 := (mod 3::int 3::int)
-#412 := (* -1::int #401)
-#413 := (+ #39 #412)
-#341 := (= #413 0::int)
-#414 := (<= 3::int 0::int)
-#411 := (or #414 #414)
-#415 := (not #411)
-#394 := (>= 3::int 0::int)
-#729 := (or #414 #394)
-#731 := (not #729)
-#388 := (or #731 #415)
-#518 := (ite #388 #341 #410)
-#725 := (= 3::int 0::int)
-#399 := (ite #725 #40 #518)
-#400 := (= 3::int #39)
-#732 := (ite #725 #400 #399)
-#555 := (or #667 #732)
-#572 := (iff #555 #670)
-#566 := (iff #670 #670)
-#574 := [rewrite]: #566
-#571 := (iff #732 #40)
-#730 := (= #39 3::int)
-#560 := (ite false #730 #40)
-#521 := (iff #560 #40)
-#669 := [rewrite]: #521
-#561 := (iff #732 #560)
-#675 := (iff #399 #40)
-#520 := (ite false #40 #40)
-#676 := (iff #520 #40)
-#677 := [rewrite]: #676
-#673 := (iff #399 #520)
-#517 := (iff #518 #40)
-#1 := true
-#680 := (ite true #40 #40)
-#681 := (iff #680 #40)
-#688 := [rewrite]: #681
-#686 := (iff #518 #680)
-#685 := (iff #410 #40)
-#683 := (= #409 #39)
-#422 := (+ #39 0::int)
-#416 := (= #422 #39)
-#698 := [rewrite]: #416
-#540 := (= #409 #422)
-#538 := (= #408 0::int)
-#693 := -3::int
-#690 := (mod -3::int -3::int)
-#691 := (= #690 0::int)
-#697 := [rewrite]: #691
-#695 := (= #408 #690)
-#694 := (= #323 -3::int)
-#689 := [rewrite]: #694
-#696 := [monotonicity #689 #689]: #695
-#539 := [trans #696 #697]: #538
-#682 := [monotonicity #539]: #540
-#684 := [trans #682 #698]: #683
-#679 := [monotonicity #684]: #685
-#703 := (iff #341 #40)
-#702 := (= #413 #39)
-#700 := (= #413 #422)
-#420 := (= #412 0::int)
-#704 := (* -1::int 0::int)
-#709 := (= #704 0::int)
-#419 := [rewrite]: #709
-#708 := (= #412 #704)
-#429 := (= #401 0::int)
-#430 := [rewrite]: #429
-#705 := [monotonicity #430]: #708
-#421 := [trans #705 #419]: #420
-#701 := [monotonicity #421]: #700
-#699 := [trans #701 #698]: #702
-#692 := [monotonicity #699]: #703
-#706 := (iff #388 true)
-#387 := (or false true)
-#712 := (iff #387 true)
-#375 := [rewrite]: #712
-#436 := (iff #388 #387)
-#434 := (iff #415 true)
-#721 := (not false)
-#711 := (iff #721 true)
-#433 := [rewrite]: #711
-#724 := (iff #415 #721)
-#363 := (iff #411 false)
-#719 := (or false false)
-#722 := (iff #719 false)
-#362 := [rewrite]: #722
-#357 := (iff #411 #719)
-#726 := (iff #414 false)
-#386 := [rewrite]: #726
-#720 := [monotonicity #386 #386]: #357
-#723 := [trans #720 #362]: #363
-#710 := [monotonicity #723]: #724
-#435 := [trans #710 #433]: #434
-#718 := (iff #731 false)
-#378 := (not true)
-#716 := (iff #378 false)
-#717 := [rewrite]: #716
-#714 := (iff #731 #378)
-#376 := (iff #729 true)
-#728 := (iff #729 #387)
-#391 := (iff #394 true)
-#727 := [rewrite]: #391
-#371 := [monotonicity #386 #727]: #728
-#377 := [trans #371 #375]: #376
-#715 := [monotonicity #377]: #714
-#713 := [trans #715 #717]: #718
-#329 := [monotonicity #713 #435]: #436
-#707 := [trans #329 #375]: #706
-#687 := [monotonicity #707 #692 #679]: #686
-#519 := [trans #687 #688]: #517
-#733 := (iff #725 false)
-#734 := [rewrite]: #733
-#674 := [monotonicity #734 #519]: #673
-#678 := [trans #674 #677]: #675
-#735 := (iff #400 #730)
-#736 := [rewrite]: #735
-#562 := [monotonicity #734 #736 #678]: #561
-#666 := [trans #562 #669]: #571
-#573 := [monotonicity #666]: #572
-#575 := [trans #573 #574]: #572
-#570 := [quant-inst]: #555
-#576 := [mp #570 #575]: #670
-[unit-resolution #576 #750 #168]: false
-unsat
-0950f059c8c59b021ddf8b08664ab965652bd6ab 328 0
-#2 := false
-#41 := 2::int
-decl f4 :: (-> int int int)
-#39 := 3::int
-#38 := 5::int
-#40 := (f4 5::int 3::int)
-#42 := (= #40 2::int)
-#43 := (not #42)
-#170 := [asserted]: #43
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#29 := (f4 #8 #9)
-#736 := (pattern #29)
-#11 := 0::int
-#67 := -1::int
-#71 := (* -1::int #9)
-#68 := (* -1::int #8)
-#126 := (mod #68 #71)
-#247 := (+ #29 #126)
-#248 := (= #247 0::int)
-#30 := (mod #8 #9)
-#244 := (* -1::int #30)
-#245 := (+ #29 #244)
-#246 := (= #245 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)
-#249 := (ite #193 #246 #248)
-#243 := (= #29 0::int)
-#12 := (= #8 0::int)
-#250 := (ite #12 #243 #249)
-#242 := (= #8 #29)
-#13 := (= #9 0::int)
-#251 := (ite #13 #242 #250)
-#737 := (forall (vars (?v0 int) (?v1 int)) (:pat #736) #251)
-#254 := (forall (vars (?v0 int) (?v1 int)) #251)
-#740 := (iff #254 #737)
-#738 := (iff #251 #251)
-#739 := [refl]: #738
-#741 := [quant-intro #739]: #740
-#132 := (* -1::int #126)
-#211 := (ite #193 #30 #132)
-#214 := (ite #12 0::int #211)
-#217 := (ite #13 #8 #214)
-#220 := (= #29 #217)
-#223 := (forall (vars (?v0 int) (?v1 int)) #220)
-#255 := (iff #223 #254)
-#252 := (iff #220 #251)
-#253 := [rewrite]: #252
-#256 := [quant-intro #253]: #255
-#101 := (not #100)
-#94 := (not #93)
-#104 := (and #94 #101)
-#90 := (not #89)
-#97 := (and #90 #94)
-#107 := (or #97 #104)
-#152 := (ite #107 #30 #132)
-#155 := (ite #12 0::int #152)
-#158 := (ite #13 #8 #155)
-#161 := (= #29 #158)
-#164 := (forall (vars (?v0 int) (?v1 int)) #161)
-#224 := (iff #164 #223)
-#221 := (iff #161 #220)
-#218 := (= #158 #217)
-#215 := (= #155 #214)
-#212 := (= #152 #211)
-#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
-#213 := [monotonicity #197]: #212
-#216 := [monotonicity #213]: #215
-#219 := [monotonicity #216]: #218
-#222 := [monotonicity #219]: #221
-#225 := [quant-intro #222]: #224
-#174 := (~ #164 #164)
-#171 := (~ #161 #161)
-#184 := [refl]: #171
-#175 := [nnf-pos #184]: #174
-#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)
-#167 := (iff #37 #164)
-#61 := (and #16 #18)
-#64 := (or #17 #61)
-#137 := (ite #64 #30 #132)
-#140 := (ite #12 0::int #137)
-#143 := (ite #13 #8 #140)
-#146 := (= #29 #143)
-#149 := (forall (vars (?v0 int) (?v1 int)) #146)
-#165 := (iff #149 #164)
-#162 := (iff #146 #161)
-#159 := (= #143 #158)
-#156 := (= #140 #155)
-#153 := (= #137 #152)
-#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
-#154 := [monotonicity #109]: #153
-#157 := [monotonicity #154]: #156
-#160 := [monotonicity #157]: #159
-#163 := [monotonicity #160]: #162
-#166 := [quant-intro #163]: #165
-#150 := (iff #37 #149)
-#147 := (iff #36 #146)
-#144 := (= #35 #143)
-#141 := (= #34 #140)
-#138 := (= #33 #137)
-#135 := (= #32 #132)
-#129 := (- #126)
-#133 := (= #129 #132)
-#134 := [rewrite]: #133
-#130 := (= #32 #129)
-#127 := (= #31 #126)
-#72 := (= #23 #71)
-#73 := [rewrite]: #72
-#69 := (= #22 #68)
-#70 := [rewrite]: #69
-#128 := [monotonicity #70 #73]: #127
-#131 := [monotonicity #128]: #130
-#136 := [trans #131 #134]: #135
-#65 := (iff #20 #64)
-#62 := (iff #19 #61)
-#63 := [rewrite]: #62
-#66 := [monotonicity #63]: #65
-#139 := [monotonicity #66 #136]: #138
-#142 := [monotonicity #139]: #141
-#145 := [monotonicity #142]: #144
-#148 := [monotonicity #145]: #147
-#151 := [quant-intro #148]: #150
-#168 := [trans #151 #166]: #167
-#125 := [asserted]: #37
-#169 := [mp #125 #168]: #164
-#185 := [mp~ #169 #175]: #164
-#226 := [mp #185 #225]: #223
-#257 := [mp #226 #256]: #254
-#742 := [mp #257 #741]: #737
-#643 := (not #737)
-#650 := (or #643 #42)
-#315 := (* -1::int 3::int)
-#400 := (* -1::int 5::int)
-#401 := (mod #400 #315)
-#402 := (+ #40 #401)
-#393 := (= #402 0::int)
-#404 := (mod 5::int 3::int)
-#405 := (* -1::int #404)
-#333 := (+ #40 #405)
-#406 := (= #333 0::int)
-#403 := (<= 3::int 0::int)
-#407 := (<= 5::int 0::int)
-#386 := (or #407 #403)
-#721 := (not #386)
-#723 := (>= 5::int 0::int)
-#380 := (or #403 #723)
-#510 := (not #380)
-#717 := (or #510 #721)
-#391 := (ite #717 #406 #393)
-#392 := (= #40 0::int)
-#724 := (= 5::int 0::int)
-#725 := (ite #724 #392 #391)
-#726 := (= 5::int #40)
-#727 := (= 3::int 0::int)
-#722 := (ite #727 #726 #725)
-#634 := (or #643 #722)
-#637 := (iff #634 #650)
-#639 := (iff #650 #650)
-#640 := [rewrite]: #639
-#648 := (iff #722 #42)
-#383 := (= #40 5::int)
-#656 := (ite false #383 #42)
-#646 := (iff #656 #42)
-#647 := [rewrite]: #646
-#652 := (iff #722 #656)
-#654 := (iff #725 #42)
-#651 := (ite false #392 #42)
-#642 := (iff #651 #42)
-#644 := [rewrite]: #642
-#653 := (iff #725 #651)
-#660 := (iff #391 #42)
-#562 := (= #40 -1::int)
-#1 := true
-#567 := (ite true #42 #562)
-#570 := (iff #567 #42)
-#663 := [rewrite]: #570
-#568 := (iff #391 #567)
-#558 := (iff #393 #562)
-#665 := 1::int
-#554 := (+ 1::int #40)
-#659 := (= #554 0::int)
-#564 := (iff #659 #562)
-#565 := [rewrite]: #564
-#662 := (iff #393 #659)
-#563 := (= #402 #554)
-#670 := (+ #40 1::int)
-#513 := (= #670 #554)
-#661 := [rewrite]: #513
-#552 := (= #402 #670)
-#669 := (= #401 1::int)
-#679 := -3::int
-#671 := -5::int
-#509 := (mod -5::int -3::int)
-#666 := (= #509 1::int)
-#668 := [rewrite]: #666
-#511 := (= #401 #509)
-#673 := (= #315 -3::int)
-#680 := [rewrite]: #673
-#672 := (= #400 -5::int)
-#678 := [rewrite]: #672
-#512 := [monotonicity #678 #680]: #511
-#667 := [trans #512 #668]: #669
-#553 := [monotonicity #667]: #552
-#658 := [trans #553 #661]: #563
-#547 := [monotonicity #658]: #662
-#566 := [trans #547 #565]: #558
-#676 := (iff #406 #42)
-#690 := -2::int
-#682 := (+ -2::int #40)
-#530 := (= #682 0::int)
-#674 := (iff #530 #42)
-#675 := [rewrite]: #674
-#531 := (iff #406 #530)
-#683 := (= #333 #682)
-#685 := (+ #40 -2::int)
-#687 := (= #685 #682)
-#688 := [rewrite]: #687
-#686 := (= #333 #685)
-#695 := (= #405 -2::int)
-#692 := (* -1::int 2::int)
-#694 := (= #692 -2::int)
-#691 := [rewrite]: #694
-#693 := (= #405 #692)
-#413 := (= #404 2::int)
-#414 := [rewrite]: #413
-#408 := [monotonicity #414]: #693
-#684 := [trans #408 #691]: #695
-#681 := [monotonicity #684]: #686
-#689 := [trans #681 #688]: #683
-#532 := [monotonicity #689]: #531
-#677 := [trans #532 #675]: #676
-#411 := (iff #717 true)
-#369 := (or false true)
-#707 := (iff #369 true)
-#708 := [rewrite]: #707
-#697 := (iff #717 #369)
-#696 := (iff #721 true)
-#321 := (not false)
-#421 := (iff #321 true)
-#422 := [rewrite]: #421
-#698 := (iff #721 #321)
-#427 := (iff #386 false)
-#716 := (or false false)
-#425 := (iff #716 false)
-#426 := [rewrite]: #425
-#702 := (iff #386 #716)
-#363 := (iff #403 false)
-#704 := [rewrite]: #363
-#715 := (iff #407 false)
-#713 := [rewrite]: #715
-#703 := [monotonicity #713 #704]: #702
-#428 := [trans #703 #426]: #427
-#699 := [monotonicity #428]: #698
-#700 := [trans #699 #422]: #696
-#354 := (iff #510 false)
-#705 := (not true)
-#712 := (iff #705 false)
-#714 := [rewrite]: #712
-#711 := (iff #510 #705)
-#709 := (iff #380 true)
-#370 := (iff #380 #369)
-#367 := (iff #723 true)
-#368 := [rewrite]: #367
-#706 := [monotonicity #704 #368]: #370
-#710 := [trans #706 #708]: #709
-#349 := [monotonicity #710]: #711
-#355 := [trans #349 #714]: #354
-#701 := [monotonicity #355 #700]: #697
-#412 := [trans #701 #708]: #411
-#569 := [monotonicity #412 #677 #566]: #568
-#664 := [trans #569 #663]: #660
-#379 := (iff #724 false)
-#720 := [rewrite]: #379
-#645 := [monotonicity #720 #664]: #653
-#655 := [trans #645 #644]: #654
-#378 := (iff #726 #383)
-#719 := [rewrite]: #378
-#728 := (iff #727 false)
-#718 := [rewrite]: #728
-#657 := [monotonicity #718 #719 #655]: #652
-#649 := [trans #657 #647]: #648
-#638 := [monotonicity #649]: #637
-#636 := [trans #638 #640]: #637
-#635 := [quant-inst]: #634
-#641 := [mp #635 #636]: #650
-[unit-resolution #641 #742 #170]: false
-unsat
 f77523eb025c744a899f70384fe23771c2e2df49 335 0
 #2 := false
 #179 := -2::int
@@ -39515,6 +37750,333 @@
 #666 := [mp #673 #665]: #671
 [unit-resolution #666 #759 #190]: 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
 51326f007e56e93f50afc5ba2d03581cdc697a2d 328 0
 #2 := false
 #11 := 0::int
@@ -39844,6 +38406,334 @@
 #688 := [mp #591 #684]: #582
 [unit-resolution #688 #766 #194]: 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
 0f74c41c938e9378ba24d8b61b7a4d0b50210370 356 0
 #2 := false
 #69 := -1::int
@@ -40201,6 +39091,318 @@
 #669 := [mp #663 #664]: #678
 [unit-resolution #669 #770 #199]: 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
 3c5a12d92626ff8ea2f3e7c2d04bebd45ef00916 350 0
 #2 := false
 #42 := 2::int
@@ -40552,6 +39754,345 @@
 #664 := [mp #674 #663]: #673
 [unit-resolution #664 #766 #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
 9ee8d0dc4ede7b68c19a265d2c4cda80156ee20b 331 0
 #2 := false
 #11 := 0::int
@@ -40884,6 +40425,599 @@
 #669 := [mp #594 #677]: #592
 [unit-resolution #669 #766 #194]: false
 unsat
+ab1c3c43720ded895d2a82ffbaf56e6367dbc9b8 592 0
+#2 := false
+#11 := 0::int
+decl f3 :: (-> int int int)
+#39 := 1::int
+decl f5 :: int
+#38 := f5
+#40 := (f3 f5 1::int)
+#704 := (>= #40 0::int)
+#752 := (= #40 0::int)
+#753 := (= f5 0::int)
+#591 := (not #753)
+#493 := [hypothesis]: #591
+#405 := (<= f5 0::int)
+#416 := (>= f5 0::int)
+#737 := (not #416)
+#451 := (not #405)
+#446 := (or #451 #737)
+#537 := (not #446)
+#69 := -1::int
+#427 := (* -1::int f5)
+#437 := (div #427 -1::int)
+#717 := (* -1::int #437)
+#715 := (+ #40 #717)
+#720 := (= #715 0::int)
+#428 := (div f5 1::int)
+#432 := (* -1::int #428)
+#411 := (+ #40 #432)
+#746 := (= #411 0::int)
+#711 := (ite #446 #746 #720)
+#483 := (or #753 #711)
+#712 := (ite #753 #752 #711)
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#10 := (f3 #8 #9)
+#754 := (pattern #10)
+#73 := (* -1::int #9)
+#70 := (* -1::int #8)
+#76 := (div #70 #73)
+#257 := (* -1::int #76)
+#258 := (+ #10 #257)
+#259 := (= #258 0::int)
+#21 := (div #8 #9)
+#254 := (* -1::int #21)
+#255 := (+ #10 #254)
+#256 := (= #255 0::int)
+#95 := (<= #9 0::int)
+#91 := (<= #8 0::int)
+#212 := (or #91 #95)
+#213 := (not #212)
+#102 := (>= #8 0::int)
+#204 := (or #95 #102)
+#205 := (not #204)
+#219 := (or #205 #213)
+#260 := (ite #219 #256 #259)
+#253 := (= #10 0::int)
+#13 := (= #9 0::int)
+#12 := (= #8 0::int)
+#14 := (or #12 #13)
+#261 := (ite #14 #253 #260)
+#755 := (forall (vars (?v0 int) (?v1 int)) (:pat #754) #261)
+#264 := (forall (vars (?v0 int) (?v1 int)) #261)
+#758 := (iff #264 #755)
+#756 := (iff #261 #261)
+#757 := [refl]: #756
+#759 := [quant-intro #757]: #758
+#224 := (ite #219 #21 #76)
+#227 := (ite #14 0::int #224)
+#230 := (= #10 #227)
+#233 := (forall (vars (?v0 int) (?v1 int)) #230)
+#265 := (iff #233 #264)
+#262 := (iff #230 #261)
+#263 := [rewrite]: #262
+#266 := [quant-intro #263]: #265
+#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)
+#234 := (iff #121 #233)
+#231 := (iff #118 #230)
+#228 := (= #115 #227)
+#225 := (= #112 #224)
+#222 := (iff #109 #219)
+#216 := (or #213 #205)
+#220 := (iff #216 #219)
+#221 := [rewrite]: #220
+#217 := (iff #109 #216)
+#214 := (iff #106 #205)
+#215 := [rewrite]: #214
+#202 := (iff #99 #213)
+#203 := [rewrite]: #202
+#218 := [monotonicity #203 #215]: #217
+#223 := [trans #218 #221]: #222
+#226 := [monotonicity #223]: #225
+#229 := [monotonicity #226]: #228
+#232 := [monotonicity #229]: #231
+#235 := [quant-intro #232]: #234
+#208 := (~ #121 #121)
+#206 := (~ #118 #118)
+#207 := [refl]: #206
+#209 := [nnf-pos #207]: #208
+#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
+#199 := [mp~ #126 #209]: #121
+#236 := [mp #199 #235]: #233
+#267 := [mp #236 #266]: #264
+#760 := [mp #267 #759]: #755
+#714 := (not #755)
+#555 := (or #714 #712)
+#426 := (* -1::int 1::int)
+#418 := (div #427 #426)
+#429 := (* -1::int #418)
+#430 := (+ #40 #429)
+#431 := (= #430 0::int)
+#748 := (<= 1::int 0::int)
+#535 := (or #405 #748)
+#742 := (not #535)
+#417 := (or #748 #416)
+#749 := (not #417)
+#750 := (or #749 #742)
+#751 := (ite #750 #746 #431)
+#747 := (= 1::int 0::int)
+#743 := (or #753 #747)
+#403 := (ite #743 #752 #751)
+#556 := (or #714 #403)
+#699 := (iff #556 #555)
+#701 := (iff #555 #555)
+#702 := [rewrite]: #701
+#713 := (iff #403 #712)
+#706 := (iff #751 #711)
+#709 := (iff #431 #720)
+#719 := (= #430 #715)
+#718 := (= #429 #717)
+#438 := (= #418 #437)
+#726 := (= #426 -1::int)
+#436 := [rewrite]: #726
+#439 := [monotonicity #436]: #438
+#433 := [monotonicity #439]: #718
+#716 := [monotonicity #433]: #719
+#710 := [monotonicity #716]: #709
+#725 := (iff #750 #446)
+#347 := (or #737 #451)
+#447 := (iff #347 #446)
+#721 := [rewrite]: #447
+#723 := (iff #750 #347)
+#452 := (iff #742 #451)
+#728 := (iff #535 #405)
+#380 := (or #405 false)
+#741 := (iff #380 #405)
+#727 := [rewrite]: #741
+#740 := (iff #535 #380)
+#395 := (iff #748 false)
+#731 := [rewrite]: #395
+#738 := [monotonicity #731]: #740
+#450 := [trans #738 #727]: #728
+#453 := [monotonicity #450]: #452
+#739 := (iff #749 #737)
+#736 := (iff #417 #416)
+#732 := (or false #416)
+#735 := (iff #732 #416)
+#730 := [rewrite]: #735
+#733 := (iff #417 #732)
+#734 := [monotonicity #731]: #733
+#374 := [trans #734 #730]: #736
+#379 := [monotonicity #374]: #739
+#724 := [monotonicity #379 #453]: #723
+#722 := [trans #724 #721]: #725
+#707 := [monotonicity #722 #710]: #706
+#393 := (iff #743 #753)
+#404 := (or #753 false)
+#729 := (iff #404 #753)
+#392 := [rewrite]: #729
+#745 := (iff #743 #404)
+#408 := (iff #747 false)
+#744 := [rewrite]: #408
+#388 := [monotonicity #744]: #745
+#394 := [trans #388 #392]: #393
+#708 := [monotonicity #394 #707]: #713
+#700 := [monotonicity #708]: #699
+#696 := [trans #700 #702]: #699
+#557 := [quant-inst]: #556
+#697 := [mp #557 #696]: #555
+#548 := [unit-resolution #697 #760]: #712
+#583 := (not #712)
+#594 := (or #583 #753 #711)
+#595 := [def-axiom]: #594
+#484 := [unit-resolution #595 #548]: #483
+#485 := [unit-resolution #484 #493]: #711
+#684 := (not #746)
+#692 := (>= #411 0::int)
+#422 := (not #692)
+decl f4 :: (-> int int int)
+#42 := (f4 f5 1::int)
+#600 := (>= #42 0::int)
+#682 := (= #42 0::int)
+#624 := (or #451 #737 #753)
+#481 := (or #624 #753)
+#599 := (not #624)
+#491 := [hypothesis]: #599
+#616 := (or #624 #405)
+#617 := [def-axiom]: #616
+#494 := [unit-resolution #617 #491]: #405
+#613 := (or #624 #416)
+#618 := [def-axiom]: #613
+#476 := [unit-resolution #618 #491]: #416
+#478 := (or #753 #451 #737)
+#479 := [th-lemma]: #478
+#480 := [unit-resolution #479 #476 #494 #493]: false
+#482 := [lemma #480]: #481
+#486 := [unit-resolution #482 #493]: #624
+#488 := (or #599 #682)
+#695 := (mod #427 -1::int)
+#641 := (+ #42 #695)
+#644 := (= #641 0::int)
+#627 := (ite #624 #682 #644)
+#29 := (f4 #8 #9)
+#761 := (pattern #29)
+#128 := (mod #70 #73)
+#273 := (+ #29 #128)
+#274 := (= #273 0::int)
+#30 := (mod #8 #9)
+#270 := (* -1::int #30)
+#271 := (+ #29 #270)
+#272 := (= #271 0::int)
+#275 := (ite #219 #272 #274)
+#269 := (= #29 0::int)
+#276 := (ite #12 #269 #275)
+#268 := (= #8 #29)
+#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
+#134 := (* -1::int #128)
+#237 := (ite #219 #30 #134)
+#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
+#154 := (ite #109 #30 #134)
+#157 := (ite #12 0::int #154)
+#160 := (ite #13 #8 #157)
+#163 := (= #29 #160)
+#166 := (forall (vars (?v0 int) (?v1 int)) #163)
+#250 := (iff #166 #249)
+#247 := (iff #163 #246)
+#244 := (= #160 #243)
+#241 := (= #157 #240)
+#238 := (= #154 #237)
+#239 := [monotonicity #223]: #238
+#242 := [monotonicity #239]: #241
+#245 := [monotonicity #242]: #244
+#248 := [monotonicity #245]: #247
+#251 := [quant-intro #248]: #250
+#200 := (~ #166 #166)
+#197 := (~ #163 #163)
+#210 := [refl]: #197
+#201 := [nnf-pos #210]: #200
+#31 := (mod #22 #23)
+#32 := (- #31)
+#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)
+#169 := (iff #37 #166)
+#139 := (ite #66 #30 #134)
+#142 := (ite #12 0::int #139)
+#145 := (ite #13 #8 #142)
+#148 := (= #29 #145)
+#151 := (forall (vars (?v0 int) (?v1 int)) #148)
+#167 := (iff #151 #166)
+#164 := (iff #148 #163)
+#161 := (= #145 #160)
+#158 := (= #142 #157)
+#155 := (= #139 #154)
+#156 := [monotonicity #111]: #155
+#159 := [monotonicity #156]: #158
+#162 := [monotonicity #159]: #161
+#165 := [monotonicity #162]: #164
+#168 := [quant-intro #165]: #167
+#152 := (iff #37 #151)
+#149 := (iff #36 #148)
+#146 := (= #35 #145)
+#143 := (= #34 #142)
+#140 := (= #33 #139)
+#137 := (= #32 #134)
+#131 := (- #128)
+#135 := (= #131 #134)
+#136 := [rewrite]: #135
+#132 := (= #32 #131)
+#129 := (= #31 #128)
+#130 := [monotonicity #72 #75]: #129
+#133 := [monotonicity #130]: #132
+#138 := [trans #133 #136]: #137
+#141 := [monotonicity #68 #138]: #140
+#144 := [monotonicity #141]: #143
+#147 := [monotonicity #144]: #146
+#150 := [monotonicity #147]: #149
+#153 := [quant-intro #150]: #152
+#170 := [trans #153 #168]: #169
+#127 := [asserted]: #37
+#171 := [mp #127 #170]: #166
+#211 := [mp~ #171 #201]: #166
+#252 := [mp #211 #251]: #249
+#283 := [mp #252 #282]: #280
+#767 := [mp #283 #766]: #762
+#603 := (not #762)
+#496 := (or #603 #627)
+#670 := (mod #427 #426)
+#667 := (+ #42 #670)
+#669 := (= #667 0::int)
+#693 := (mod f5 1::int)
+#679 := (* -1::int #693)
+#680 := (+ #42 #679)
+#681 := (= #680 0::int)
+#677 := (ite #750 #681 #669)
+#671 := (ite #753 #682 #677)
+#672 := (= f5 #42)
+#673 := (ite #747 #672 #671)
+#607 := (or #603 #673)
+#609 := (iff #607 #496)
+#610 := (iff #496 #496)
+#611 := [rewrite]: #610
+#601 := (iff #673 #627)
+#629 := (or #446 #753)
+#630 := (ite #629 #682 #644)
+#622 := (iff #630 #627)
+#625 := (iff #629 #624)
+#626 := [rewrite]: #625
+#628 := [monotonicity #626]: #622
+#620 := (iff #673 #630)
+#636 := (ite false #672 #630)
+#623 := (iff #636 #630)
+#619 := [rewrite]: #623
+#631 := (iff #673 #636)
+#634 := (iff #671 #630)
+#647 := (ite #446 #682 #644)
+#650 := (ite #753 #682 #647)
+#632 := (iff #650 #630)
+#633 := [rewrite]: #632
+#640 := (iff #671 #650)
+#648 := (iff #677 #647)
+#645 := (iff #669 #644)
+#642 := (= #667 #641)
+#638 := (= #670 #695)
+#639 := [monotonicity #436]: #638
+#643 := [monotonicity #639]: #642
+#646 := [monotonicity #643]: #645
+#654 := (iff #681 #682)
+#656 := (= #680 #42)
+#661 := (+ #42 0::int)
+#653 := (= #661 #42)
+#655 := [rewrite]: #653
+#666 := (= #680 #661)
+#664 := (= #679 0::int)
+#675 := (* -1::int 0::int)
+#662 := (= #675 0::int)
+#663 := [rewrite]: #662
+#659 := (= #679 #675)
+#674 := (= #693 0::int)
+#668 := [rewrite]: #674
+#660 := [monotonicity #668]: #659
+#665 := [trans #660 #663]: #664
+#652 := [monotonicity #665]: #666
+#657 := [trans #652 #655]: #656
+#658 := [monotonicity #657]: #654
+#649 := [monotonicity #722 #658 #646]: #648
+#651 := [monotonicity #649]: #640
+#635 := [trans #651 #633]: #634
+#637 := [monotonicity #744 #635]: #631
+#621 := [trans #637 #619]: #620
+#602 := [trans #621 #628]: #601
+#614 := [monotonicity #602]: #609
+#612 := [trans #614 #611]: #609
+#608 := [quant-inst]: #607
+#615 := [mp #608 #612]: #496
+#487 := [unit-resolution #615 #767]: #627
+#581 := (not #627)
+#540 := (or #581 #599 #682)
+#571 := [def-axiom]: #540
+#477 := [unit-resolution #571 #487]: #488
+#489 := [unit-resolution #477 #486]: #682
+#582 := (not #682)
+#396 := (or #582 #600)
+#397 := [th-lemma]: #396
+#398 := [unit-resolution #397 #489]: #600
+#421 := (not #600)
+#389 := (or #422 #421)
+#542 := (+ f5 #432)
+#541 := (= #542 0::int)
+#1 := true
+#60 := [true-axiom]: true
+#515 := (or false #541)
+#505 := [th-lemma]: #515
+#495 := [unit-resolution #505 #60]: #541
+#443 := (<= #542 0::int)
+#420 := (not #443)
+#185 := (* -1::int #42)
+#184 := (* -1::int #40)
+#186 := (+ #184 #185)
+#187 := (+ f5 #186)
+#341 := (<= #187 0::int)
+#444 := (not #341)
+#425 := (>= #187 0::int)
+#703 := (<= #40 0::int)
+#499 := (not #425)
+#507 := [hypothesis]: #499
+#461 := (or #753 #425)
+#694 := (<= #411 0::int)
+#504 := (not #694)
+#605 := (<= #42 0::int)
+#466 := (or #582 #605)
+#468 := [th-lemma]: #466
+#469 := [unit-resolution #468 #489]: #605
+#503 := (not #605)
+#490 := (or #503 #425 #504)
+#510 := [hypothesis]: #605
+#549 := (>= #542 0::int)
+#497 := (not #541)
+#498 := (or #497 #549)
+#501 := [th-lemma]: #498
+#502 := [unit-resolution #501 #495]: #549
+#506 := [hypothesis]: #694
+#500 := [th-lemma #507 #506 #502 #510]: false
+#492 := [lemma #500]: #490
+#470 := [unit-resolution #492 #469 #507]: #504
+#471 := (or #684 #694)
+#472 := [th-lemma]: #471
+#473 := [unit-resolution #472 #470]: #684
+#579 := (not #711)
+#538 := (or #579 #537 #746)
+#686 := [def-axiom]: #538
+#474 := [unit-resolution #686 #473 #485]: #537
+#698 := (or #446 #405)
+#705 := [def-axiom]: #698
+#467 := [unit-resolution #705 #474]: #405
+#534 := (or #446 #416)
+#536 := [def-axiom]: #534
+#475 := [unit-resolution #536 #474]: #416
+#460 := [unit-resolution #479 #475 #467 #493]: false
+#462 := [lemma #460]: #461
+#464 := [unit-resolution #462 #507]: #753
+#463 := (or #591 #752)
+#592 := (or #583 #591 #752)
+#593 := [def-axiom]: #592
+#465 := [unit-resolution #593 #548]: #463
+#454 := [unit-resolution #465 #464]: #752
+#688 := (not #752)
+#455 := (or #688 #703)
+#456 := [th-lemma]: #455
+#448 := [unit-resolution #456 #454]: #703
+#457 := (or #591 #416)
+#458 := [th-lemma]: #457
+#449 := [unit-resolution #458 #464]: #416
+#598 := (or #624 #591)
+#544 := [def-axiom]: #598
+#459 := [unit-resolution #544 #464]: #624
+#440 := [unit-resolution #477 #459]: #682
+#441 := [unit-resolution #468 #440]: #605
+#442 := [th-lemma #441 #507 #449 #448]: false
+#434 := [lemma #442]: #425
+#412 := (or #444 #499)
+#188 := (= #187 0::int)
+#191 := (not #188)
+#41 := (* #40 1::int)
+#43 := (+ #41 #42)
+#44 := (= f5 #43)
+#45 := (not #44)
+#194 := (iff #45 #191)
+#175 := (+ #40 #42)
+#178 := (= f5 #175)
+#181 := (not #178)
+#192 := (iff #181 #191)
+#189 := (iff #178 #188)
+#190 := [rewrite]: #189
+#193 := [monotonicity #190]: #192
+#182 := (iff #45 #181)
+#179 := (iff #44 #178)
+#176 := (= #43 #175)
+#173 := (= #41 #40)
+#174 := [rewrite]: #173
+#177 := [monotonicity #174]: #176
+#180 := [monotonicity #177]: #179
+#183 := [monotonicity #180]: #182
+#195 := [trans #183 #193]: #194
+#172 := [asserted]: #45
+#196 := [mp #172 #195]: #191
+#435 := (or #188 #444 #499)
+#445 := [th-lemma]: #435
+#414 := [unit-resolution #445 #196]: #412
+#415 := [unit-resolution #414 #434]: #444
+#406 := [hypothesis]: #600
+#419 := [hypothesis]: #692
+#423 := (or #420 #421 #341 #422)
+#413 := [th-lemma]: #423
+#424 := [unit-resolution #413 #419 #406 #415]: #420
+#407 := (or #497 #443)
+#409 := [th-lemma]: #407
+#410 := [unit-resolution #409 #424 #495]: false
+#391 := [lemma #410]: #389
+#399 := [unit-resolution #391 #398]: #422
+#400 := (or #684 #692)
+#401 := [th-lemma]: #400
+#390 := [unit-resolution #401 #399]: #684
+#402 := [unit-resolution #686 #390 #485]: #537
+#383 := [unit-resolution #705 #402]: #405
+#385 := [unit-resolution #536 #402]: #416
+#386 := [unit-resolution #479 #385 #383 #493]: false
+#384 := [lemma #386]: #753
+#387 := [unit-resolution #465 #384]: #752
+#375 := (or #688 #704)
+#377 := [th-lemma]: #375
+#378 := [unit-resolution #377 #387]: #704
+#381 := (or #591 #405)
+#376 := [th-lemma]: #381
+#382 := [unit-resolution #376 #384]: #405
+#357 := [unit-resolution #544 #384]: #624
+#361 := [unit-resolution #477 #357]: #682
+#362 := [unit-resolution #397 #361]: #600
+[th-lemma #362 #415 #382 #378]: false
+unsat
 0dff6836610ad24fd350792fdba18f67fffe78ae 353 0
 #2 := false
 #42 := 1::int
@@ -41238,6 +41372,556 @@
 #666 := [mp #660 #661]: #675
 [unit-resolution #666 #767 #196]: false
 unsat
+f7d21861165159375bc45e6364e84fc0183b7f2c 549 0
+#2 := false
+#11 := 0::int
+decl f3 :: (-> int int int)
+#39 := 3::int
+decl f5 :: int
+#38 := f5
+#40 := (f3 f5 3::int)
+#703 := (>= #40 0::int)
+#754 := (= #40 0::int)
+#750 := (= f5 0::int)
+#690 := (not #750)
+#562 := [hypothesis]: #690
+#745 := (>= f5 0::int)
+#751 := (<= f5 0::int)
+#453 := (not #751)
+#377 := (not #745)
+#456 := (or #377 #453)
+#706 := (not #456)
+#185 := -3::int
+#69 := -1::int
+#430 := (* -1::int f5)
+#450 := (div #430 -3::int)
+#725 := (* -1::int #450)
+#440 := (+ #40 #725)
+#720 := (= #440 0::int)
+#434 := (div f5 3::int)
+#431 := (* -1::int #434)
+#435 := (+ #40 #431)
+#414 := (= #435 0::int)
+#718 := (ite #456 #414 #720)
+#564 := (or #750 #718)
+#723 := (ite #750 #754 #718)
+#9 := (:var 0 int)
+#8 := (:var 1 int)
+#10 := (f3 #8 #9)
+#757 := (pattern #10)
+#73 := (* -1::int #9)
+#70 := (* -1::int #8)
+#76 := (div #70 #73)
+#259 := (* -1::int #76)
+#260 := (+ #10 #259)
+#261 := (= #260 0::int)
+#21 := (div #8 #9)
+#256 := (* -1::int #21)
+#257 := (+ #10 #256)
+#258 := (= #257 0::int)
+#95 := (<= #9 0::int)
+#91 := (<= #8 0::int)
+#214 := (or #91 #95)
+#215 := (not #214)
+#102 := (>= #8 0::int)
+#206 := (or #95 #102)
+#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)
+#758 := (forall (vars (?v0 int) (?v1 int)) (:pat #757) #263)
+#266 := (forall (vars (?v0 int) (?v1 int)) #263)
+#761 := (iff #266 #758)
+#759 := (iff #263 #263)
+#760 := [refl]: #759
+#762 := [quant-intro #760]: #761
+#226 := (ite #221 #21 #76)
+#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
+#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)
+#236 := (iff #121 #235)
+#233 := (iff #118 #232)
+#230 := (= #115 #229)
+#227 := (= #112 #226)
+#224 := (iff #109 #221)
+#218 := (or #215 #207)
+#222 := (iff #218 #221)
+#223 := [rewrite]: #222
+#219 := (iff #109 #218)
+#216 := (iff #106 #207)
+#217 := [rewrite]: #216
+#204 := (iff #99 #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 := (~ #121 #121)
+#208 := (~ #118 #118)
+#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)
+#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
+#201 := [mp~ #126 #211]: #121
+#238 := [mp #201 #237]: #235
+#269 := [mp #238 #268]: #266
+#763 := [mp #269 #762]: #758
+#714 := (not #758)
+#709 := (or #714 #723)
+#429 := (* -1::int 3::int)
+#421 := (div #430 #429)
+#432 := (* -1::int #421)
+#433 := (+ #40 #432)
+#361 := (= #433 0::int)
+#749 := (<= 3::int 0::int)
+#408 := (or #751 #749)
+#538 := (not #408)
+#419 := (or #749 #745)
+#420 := (not #419)
+#752 := (or #420 #538)
+#753 := (ite #752 #414 #361)
+#755 := (= 3::int 0::int)
+#756 := (or #750 #755)
+#746 := (ite #756 #754 #753)
+#710 := (or #714 #746)
+#716 := (iff #710 #709)
+#717 := (iff #709 #709)
+#558 := [rewrite]: #717
+#712 := (iff #746 #723)
+#722 := (iff #753 #718)
+#721 := (iff #361 #720)
+#441 := (= #433 #440)
+#729 := (= #432 #725)
+#724 := (= #421 #450)
+#727 := (= #429 -3::int)
+#449 := [rewrite]: #727
+#728 := [monotonicity #449]: #724
+#439 := [monotonicity #728]: #729
+#442 := [monotonicity #439]: #441
+#436 := [monotonicity #442]: #721
+#349 := (iff #752 #456)
+#454 := (iff #538 #453)
+#730 := (iff #408 #751)
+#382 := (or #751 false)
+#741 := (iff #382 #751)
+#744 := [rewrite]: #741
+#383 := (iff #408 #382)
+#397 := (iff #749 false)
+#398 := [rewrite]: #397
+#743 := [monotonicity #398]: #383
+#731 := [trans #743 #744]: #730
+#455 := [monotonicity #731]: #454
+#740 := (iff #420 #377)
+#733 := (iff #419 #745)
+#734 := (or false #745)
+#737 := (iff #734 #745)
+#738 := [rewrite]: #737
+#735 := (iff #419 #734)
+#736 := [monotonicity #398]: #735
+#739 := [trans #736 #738]: #733
+#742 := [monotonicity #739]: #740
+#726 := [monotonicity #742 #455]: #349
+#719 := [monotonicity #726 #436]: #722
+#395 := (iff #756 #750)
+#747 := (or #750 false)
+#391 := (iff #747 #750)
+#732 := [rewrite]: #391
+#407 := (iff #756 #747)
+#406 := (iff #755 false)
+#411 := [rewrite]: #406
+#748 := [monotonicity #411]: #407
+#396 := [trans #748 #732]: #395
+#713 := [monotonicity #396 #719]: #712
+#711 := [monotonicity #713]: #716
+#559 := [trans #711 #558]: #716
+#715 := [quant-inst]: #710
+#560 := [mp #715 #559]: #709
+#563 := [unit-resolution #560 #763]: #723
+#687 := (not #723)
+#592 := (or #687 #750 #718)
+#593 := [def-axiom]: #592
+#565 := [unit-resolution #593 #563]: #564
+#566 := [unit-resolution #565 #562]: #718
+#540 := (mod #430 -3::int)
+decl f4 :: (-> int int int)
+#42 := (f4 f5 3::int)
+#685 := (+ #42 #540)
+#676 := (= #685 0::int)
+#708 := (mod f5 3::int)
+#692 := (* -1::int #708)
+#679 := (+ #42 #692)
+#681 := (= #679 0::int)
+#678 := (ite #456 #681 #676)
+#525 := (or #750 #678)
+#670 := (= #42 0::int)
+#665 := (ite #750 #670 #678)
+#29 := (f4 #8 #9)
+#764 := (pattern #29)
+#128 := (mod #70 #73)
+#275 := (+ #29 #128)
+#276 := (= #275 0::int)
+#30 := (mod #8 #9)
+#272 := (* -1::int #30)
+#273 := (+ #29 #272)
+#274 := (= #273 0::int)
+#277 := (ite #221 #274 #276)
+#271 := (= #29 0::int)
+#278 := (ite #12 #271 #277)
+#270 := (= #8 #29)
+#279 := (ite #13 #270 #278)
+#765 := (forall (vars (?v0 int) (?v1 int)) (:pat #764) #279)
+#282 := (forall (vars (?v0 int) (?v1 int)) #279)
+#768 := (iff #282 #765)
+#766 := (iff #279 #279)
+#767 := [refl]: #766
+#769 := [quant-intro #767]: #768
+#134 := (* -1::int #128)
+#239 := (ite #221 #30 #134)
+#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
+#154 := (ite #109 #30 #134)
+#157 := (ite #12 0::int #154)
+#160 := (ite #13 #8 #157)
+#163 := (= #29 #160)
+#166 := (forall (vars (?v0 int) (?v1 int)) #163)
+#252 := (iff #166 #251)
+#249 := (iff #163 #248)
+#246 := (= #160 #245)
+#243 := (= #157 #242)
+#240 := (= #154 #239)
+#241 := [monotonicity #225]: #240
+#244 := [monotonicity #241]: #243
+#247 := [monotonicity #244]: #246
+#250 := [monotonicity #247]: #249
+#253 := [quant-intro #250]: #252
+#202 := (~ #166 #166)
+#199 := (~ #163 #163)
+#212 := [refl]: #199
+#203 := [nnf-pos #212]: #202
+#31 := (mod #22 #23)
+#32 := (- #31)
+#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)
+#169 := (iff #37 #166)
+#139 := (ite #66 #30 #134)
+#142 := (ite #12 0::int #139)
+#145 := (ite #13 #8 #142)
+#148 := (= #29 #145)
+#151 := (forall (vars (?v0 int) (?v1 int)) #148)
+#167 := (iff #151 #166)
+#164 := (iff #148 #163)
+#161 := (= #145 #160)
+#158 := (= #142 #157)
+#155 := (= #139 #154)
+#156 := [monotonicity #111]: #155
+#159 := [monotonicity #156]: #158
+#162 := [monotonicity #159]: #161
+#165 := [monotonicity #162]: #164
+#168 := [quant-intro #165]: #167
+#152 := (iff #37 #151)
+#149 := (iff #36 #148)
+#146 := (= #35 #145)
+#143 := (= #34 #142)
+#140 := (= #33 #139)
+#137 := (= #32 #134)
+#131 := (- #128)
+#135 := (= #131 #134)
+#136 := [rewrite]: #135
+#132 := (= #32 #131)
+#129 := (= #31 #128)
+#130 := [monotonicity #72 #75]: #129
+#133 := [monotonicity #130]: #132
+#138 := [trans #133 #136]: #137
+#141 := [monotonicity #68 #138]: #140
+#144 := [monotonicity #141]: #143
+#147 := [monotonicity #144]: #146
+#150 := [monotonicity #147]: #149
+#153 := [quant-intro #150]: #152
+#170 := [trans #153 #168]: #169
+#127 := [asserted]: #37
+#171 := [mp #127 #170]: #166
+#213 := [mp~ #171 #203]: #166
+#254 := [mp #213 #253]: #251
+#285 := [mp #254 #284]: #282
+#770 := [mp #285 #769]: #765
+#660 := (not #765)
+#657 := (or #660 #665)
+#598 := (mod #430 #429)
+#691 := (+ #42 #598)
+#688 := (= #691 0::int)
+#673 := (ite #752 #681 #688)
+#672 := (ite #750 #670 #673)
+#682 := (= f5 #42)
+#683 := (ite #755 #682 #672)
+#661 := (or #660 #683)
+#642 := (iff #661 #657)
+#645 := (iff #657 #657)
+#646 := [rewrite]: #645
+#658 := (iff #683 #665)
+#668 := (ite false #682 #665)
+#655 := (iff #668 #665)
+#656 := [rewrite]: #655
+#664 := (iff #683 #668)
+#666 := (iff #672 #665)
+#662 := (iff #673 #678)
+#677 := (iff #688 #676)
+#674 := (= #691 #685)
+#684 := (= #598 #540)
+#680 := [monotonicity #449]: #684
+#675 := [monotonicity #680]: #674
+#671 := [monotonicity #675]: #677
+#663 := [monotonicity #726 #671]: #662
+#667 := [monotonicity #663]: #666
+#669 := [monotonicity #411 #667]: #664
+#659 := [trans #669 #656]: #658
+#644 := [monotonicity #659]: #642
+#647 := [trans #644 #646]: #642
+#641 := [quant-inst]: #661
+#648 := [mp #641 #647]: #657
+#621 := [unit-resolution #648 #770]: #665
+#622 := (not #665)
+#627 := (or #622 #750 #678)
+#628 := [def-axiom]: #627
+#526 := [unit-resolution #628 #621]: #525
+#527 := [unit-resolution #526 #562]: #678
+#696 := (not #718)
+#654 := (not #678)
+#491 := (or #706 #654 #696)
+#652 := (>= #679 0::int)
+#567 := [hypothesis]: #678
+#568 := [hypothesis]: #456
+#632 := (or #654 #706 #681)
+#633 := [def-axiom]: #632
+#569 := [unit-resolution #633 #568 #567]: #681
+#637 := (not #681)
+#493 := (or #637 #652)
+#495 := [th-lemma]: #493
+#496 := [unit-resolution #495 #569]: #652
+#539 := (>= #435 0::int)
+#549 := [hypothesis]: #718
+#697 := (or #696 #706 #414)
+#695 := [def-axiom]: #697
+#550 := [unit-resolution #695 #568 #549]: #414
+#581 := (not #414)
+#494 := (or #581 #539)
+#497 := [th-lemma]: #494
+#479 := [unit-resolution #497 #550]: #539
+#187 := (* -1::int #42)
+#186 := (* -3::int #40)
+#188 := (+ #186 #187)
+#189 := (+ f5 #188)
+#343 := (<= #189 0::int)
+#481 := (not #343)
+#428 := (>= #189 0::int)
+#557 := (not #428)
+#573 := [hypothesis]: #557
+#533 := (or #750 #428)
+#554 := (or #706 #428 #654 #696)
+#651 := (<= #679 0::int)
+#570 := (or #637 #651)
+#571 := [th-lemma]: #570
+#572 := [unit-resolution #571 #569]: #651
+#615 := (* -3::int #434)
+#618 := (+ #615 #692)
+#619 := (+ f5 #618)
+#601 := (>= #619 0::int)
+#614 := (= #619 0::int)
+#1 := true
+#60 := [true-axiom]: true
+#534 := (or false #614)
+#535 := [th-lemma]: #534
+#542 := [unit-resolution #535 #60]: #614
+#544 := (not #614)
+#545 := (or #544 #601)
+#546 := [th-lemma]: #545
+#548 := [unit-resolution #546 #542]: #601
+#537 := (<= #435 0::int)
+#551 := (or #581 #537)
+#552 := [th-lemma]: #551
+#553 := [unit-resolution #552 #550]: #537
+#536 := [th-lemma #553 #548 #573 #572]: false
+#524 := [lemma #536]: #554
+#529 := [unit-resolution #524 #527 #573 #566]: #706
+#704 := (or #456 #745)
+#705 := [def-axiom]: #704
+#530 := [unit-resolution #705 #529]: #745
+#699 := (or #456 #751)
+#700 := [def-axiom]: #699
+#531 := [unit-resolution #700 #529]: #751
+#374 := (or #750 #453 #377)
+#532 := [th-lemma]: #374
+#528 := [unit-resolution #532 #531 #530 #562]: false
+#512 := [lemma #528]: #533
+#515 := [unit-resolution #512 #573]: #750
+#519 := (or #690 #745)
+#521 := [th-lemma]: #519
+#522 := [unit-resolution #521 #515]: #745
+#649 := (<= #42 0::int)
+#520 := (or #690 #670)
+#623 := (or #622 #690 #670)
+#624 := [def-axiom]: #623
+#523 := [unit-resolution #624 #621]: #520
+#511 := [unit-resolution #523 #515]: #670
+#629 := (not #670)
+#514 := (or #629 #649)
+#516 := [th-lemma]: #514
+#517 := [unit-resolution #516 #511]: #649
+#702 := (<= #40 0::int)
+#513 := (or #690 #754)
+#575 := (or #687 #690 #754)
+#590 := [def-axiom]: #575
+#518 := [unit-resolution #590 #563]: #513
+#508 := [unit-resolution #518 #515]: #754
+#586 := (not #754)
+#498 := (or #586 #702)
+#500 := [th-lemma]: #498
+#501 := [unit-resolution #500 #508]: #702
+#504 := [th-lemma #501 #517 #573 #522]: false
+#505 := [lemma #504]: #428
+#484 := (or #481 #557)
+#190 := (= #189 0::int)
+#193 := (not #190)
+#41 := (* #40 3::int)
+#43 := (+ #41 #42)
+#44 := (= f5 #43)
+#45 := (not #44)
+#196 := (iff #45 #193)
+#173 := (* 3::int #40)
+#176 := (+ #173 #42)
+#179 := (= f5 #176)
+#182 := (not #179)
+#194 := (iff #182 #193)
+#191 := (iff #179 #190)
+#192 := [rewrite]: #191
+#195 := [monotonicity #192]: #194
+#183 := (iff #45 #182)
+#180 := (iff #44 #179)
+#177 := (= #43 #176)
+#174 := (= #41 #173)
+#175 := [rewrite]: #174
+#178 := [monotonicity #175]: #177
+#181 := [monotonicity #178]: #180
+#184 := [monotonicity #181]: #183
+#197 := [trans #184 #195]: #196
+#172 := [asserted]: #45
+#198 := [mp #172 #197]: #193
+#482 := (or #190 #481 #557)
+#483 := [th-lemma]: #482
+#485 := [unit-resolution #483 #198]: #484
+#486 := [unit-resolution #485 #505]: #481
+#509 := (<= #619 0::int)
+#487 := (or #544 #509)
+#488 := [th-lemma]: #487
+#489 := [unit-resolution #488 #542]: #509
+#490 := [th-lemma #489 #486 #479 #496]: false
+#480 := [lemma #490]: #491
+#502 := [unit-resolution #480 #527 #566]: #706
+#510 := [unit-resolution #705 #502]: #745
+#503 := [unit-resolution #700 #502]: #751
+#506 := [unit-resolution #532 #503 #510 #562]: false
+#507 := [lemma #506]: #750
+#492 := [unit-resolution #518 #507]: #754
+#469 := (or #586 #703)
+#471 := [th-lemma]: #469
+#472 := [unit-resolution #471 #492]: #703
+#650 := (>= #42 0::int)
+#473 := [unit-resolution #523 #507]: #670
+#474 := (or #629 #650)
+#475 := [th-lemma]: #474
+#476 := [unit-resolution #475 #473]: #650
+#477 := (or #690 #751)
+#470 := [th-lemma]: #477
+#478 := [unit-resolution #470 #507]: #751
+[th-lemma #478 #476 #486 #472]: false
+unsat
 2131919aacf2b81143bdcdf048ab7c7b571f4dab 350 0
 #2 := false
 #68 := -1::int
@@ -41909,6 +42593,319 @@
 #589 := [mp #682 #588]: #685
 [unit-resolution #589 #766 #194]: false
 unsat
+b8a677f41acfea6ca1bd28f53fd89f2c10758abc 75 0
+#2 := false
+#8 := 0::int
+decl f3 :: int
+#9 := f3
+#32 := -1::int
+#33 := (* -1::int f3)
+#45 := (>= f3 0::int)
+#52 := (ite #45 f3 #33)
+#73 := (* -1::int #52)
+#81 := (+ f3 #73)
+#90 := (<= #81 0::int)
+#76 := (= f3 #52)
+#71 := (+ #33 #73)
+#80 := (<= #71 0::int)
+#77 := (= #33 #52)
+#46 := (not #45)
+#82 := [hypothesis]: #46
+#74 := (or #45 #77)
+#75 := [def-axiom]: #74
+#83 := [unit-resolution #75 #82]: #77
+#84 := (not #77)
+#85 := (or #84 #80)
+#86 := [th-lemma]: #85
+#87 := [unit-resolution #86 #83]: #80
+#61 := (>= #52 0::int)
+#65 := (not #61)
+#11 := (- f3)
+#10 := (< f3 0::int)
+#12 := (ite #10 #11 f3)
+#13 := (<= 0::int #12)
+#14 := (not #13)
+#68 := (iff #14 #65)
+#36 := (ite #10 #33 f3)
+#39 := (<= 0::int #36)
+#42 := (not #39)
+#66 := (iff #42 #65)
+#63 := (iff #39 #61)
+#57 := (<= 0::int #52)
+#60 := (iff #57 #61)
+#62 := [rewrite]: #60
+#58 := (iff #39 #57)
+#55 := (= #36 #52)
+#49 := (ite #46 #33 f3)
+#53 := (= #49 #52)
+#54 := [rewrite]: #53
+#50 := (= #36 #49)
+#47 := (iff #10 #46)
+#48 := [rewrite]: #47
+#51 := [monotonicity #48]: #50
+#56 := [trans #51 #54]: #55
+#59 := [monotonicity #56]: #58
+#64 := [trans #59 #62]: #63
+#67 := [monotonicity #64]: #66
+#43 := (iff #14 #42)
+#40 := (iff #13 #39)
+#37 := (= #12 #36)
+#34 := (= #11 #33)
+#35 := [rewrite]: #34
+#38 := [monotonicity #35]: #37
+#41 := [monotonicity #38]: #40
+#44 := [monotonicity #41]: #43
+#69 := [trans #44 #67]: #68
+#31 := [asserted]: #14
+#70 := [mp #31 #69]: #65
+#88 := [th-lemma #82 #70 #87]: false
+#89 := [lemma #88]: #45
+#78 := (or #46 #76)
+#79 := [def-axiom]: #78
+#92 := [unit-resolution #79 #89]: #76
+#93 := (not #76)
+#94 := (or #93 #90)
+#95 := [th-lemma]: #94
+#96 := [unit-resolution #95 #92]: #90
+[th-lemma #89 #70 #96]: false
+unsat
+957272aa2f77e61687a16c7aa490ba2664908958 132 0
+#2 := false
+#9 := 0::int
+decl f3 :: int
+#8 := f3
+#66 := (>= f3 0::int)
+#67 := (not #66)
+#34 := -1::int
+#35 := (* -1::int f3)
+#73 := (ite #66 f3 #35)
+#102 := (= f3 #73)
+#130 := (not #102)
+#14 := (= f3 0::int)
+#82 := (= #73 0::int)
+#124 := (iff #82 #14)
+#122 := (iff #14 #82)
+#121 := [hypothesis]: #102
+#123 := [monotonicity #121]: #122
+#125 := [symm #123]: #124
+#131 := (or #82 #130)
+#60 := (not #14)
+#99 := (not #82)
+#126 := (iff #99 #60)
+#127 := [monotonicity #125]: #126
+#119 := [hypothesis]: #99
+#128 := [mp #119 #127]: #60
+#112 := (or #14 #82)
+#89 := (iff #60 #82)
+#11 := (- f3)
+#10 := (< f3 0::int)
+#12 := (ite #10 #11 f3)
+#13 := (= #12 0::int)
+#15 := (iff #13 #14)
+#16 := (not #15)
+#94 := (iff #16 #89)
+#38 := (ite #10 #35 f3)
+#44 := (= 0::int #38)
+#61 := (iff #44 #60)
+#92 := (iff #61 #89)
+#86 := (iff #82 #60)
+#90 := (iff #86 #89)
+#91 := [rewrite]: #90
+#87 := (iff #61 #86)
+#84 := (iff #44 #82)
+#78 := (= 0::int #73)
+#81 := (iff #78 #82)
+#83 := [rewrite]: #81
+#79 := (iff #44 #78)
+#76 := (= #38 #73)
+#70 := (ite #67 #35 f3)
+#74 := (= #70 #73)
+#75 := [rewrite]: #74
+#71 := (= #38 #70)
+#68 := (iff #10 #67)
+#69 := [rewrite]: #68
+#72 := [monotonicity #69]: #71
+#77 := [trans #72 #75]: #76
+#80 := [monotonicity #77]: #79
+#85 := [trans #80 #83]: #84
+#88 := [monotonicity #85]: #87
+#93 := [trans #88 #91]: #92
+#64 := (iff #16 #61)
+#52 := (iff #14 #44)
+#57 := (not #52)
+#62 := (iff #57 #61)
+#63 := [rewrite]: #62
+#58 := (iff #16 #57)
+#55 := (iff #15 #52)
+#49 := (iff #44 #14)
+#53 := (iff #49 #52)
+#54 := [rewrite]: #53
+#50 := (iff #15 #49)
+#47 := (iff #13 #44)
+#41 := (= #38 0::int)
+#45 := (iff #41 #44)
+#46 := [rewrite]: #45
+#42 := (iff #13 #41)
+#39 := (= #12 #38)
+#36 := (= #11 #35)
+#37 := [rewrite]: #36
+#40 := [monotonicity #37]: #39
+#43 := [monotonicity #40]: #42
+#48 := [trans #43 #46]: #47
+#51 := [monotonicity #48]: #50
+#56 := [trans #51 #54]: #55
+#59 := [monotonicity #56]: #58
+#65 := [trans #59 #63]: #64
+#95 := [trans #65 #93]: #94
+#33 := [asserted]: #16
+#96 := [mp #33 #95]: #89
+#97 := (not #89)
+#110 := (or #14 #82 #97)
+#111 := [def-axiom]: #110
+#113 := [unit-resolution #111 #96]: #112
+#120 := [unit-resolution #113 #119]: #14
+#129 := [unit-resolution #120 #128]: false
+#132 := [lemma #129]: #131
+#133 := [unit-resolution #132 #121]: #82
+#135 := [mp #133 #125]: #14
+#108 := (or #60 #99)
+#106 := (or #60 #99 #97)
+#107 := [def-axiom]: #106
+#109 := [unit-resolution #107 #96]: #108
+#134 := [unit-resolution #109 #133]: #60
+#136 := [unit-resolution #134 #135]: false
+#137 := [lemma #136]: #130
+#104 := (or #67 #102)
+#105 := [def-axiom]: #104
+#143 := [unit-resolution #105 #137]: #67
+#138 := (= #35 0::int)
+#147 := (not #138)
+#157 := (iff #147 #99)
+#155 := (iff #138 #82)
+#103 := (= #35 #73)
+#100 := (or #66 #103)
+#101 := [def-axiom]: #100
+#154 := [unit-resolution #101 #143]: #103
+#156 := [monotonicity #154]: #155
+#158 := [monotonicity #156]: #157
+#139 := (<= #35 0::int)
+#145 := (not #139)
+#142 := [hypothesis]: #139
+#144 := [th-lemma #143 #142]: false
+#146 := [lemma #144]: #145
+#148 := (or #147 #139)
+#149 := [th-lemma]: #148
+#153 := [unit-resolution #149 #146]: #147
+#159 := [mp #153 #158]: #99
+#160 := [unit-resolution #113 #159]: #14
+#161 := (or #60 #66)
+#162 := [th-lemma]: #161
+[unit-resolution #162 #160 #143]: false
+unsat
+9485d306079aa8cdc08952543a8ac86391d45fc0 103 0
+#2 := false
+#8 := 0::int
+decl f3 :: int
+#9 := f3
+#34 := -1::int
+#35 := (* -1::int f3)
+#112 := (* -1::int #35)
+#113 := (+ f3 #112)
+#115 := (>= #113 0::int)
+#111 := (= f3 #35)
+#61 := (>= f3 0::int)
+#68 := (ite #61 f3 #35)
+#118 := (= #68 #35)
+#96 := (= #35 #68)
+#62 := (not #61)
+#107 := [hypothesis]: #61
+#73 := (= f3 #68)
+#97 := (or #62 #73)
+#98 := [def-axiom]: #97
+#108 := [unit-resolution #98 #107]: #73
+#95 := (not #73)
+#101 := (or #62 #95)
+#83 := (iff #62 #73)
+#12 := (- f3)
+#11 := (< f3 0::int)
+#13 := (ite #11 #12 f3)
+#14 := (= #13 f3)
+#10 := (<= 0::int f3)
+#15 := (iff #10 #14)
+#16 := (not #15)
+#88 := (iff #16 #83)
+#55 := (not #10)
+#38 := (ite #11 #35 f3)
+#44 := (= f3 #38)
+#56 := (iff #44 #55)
+#86 := (iff #56 #83)
+#80 := (iff #73 #62)
+#84 := (iff #80 #83)
+#85 := [rewrite]: #84
+#81 := (iff #56 #80)
+#78 := (iff #55 #62)
+#76 := (iff #10 #61)
+#77 := [rewrite]: #76
+#79 := [monotonicity #77]: #78
+#74 := (iff #44 #73)
+#71 := (= #38 #68)
+#65 := (ite #62 #35 f3)
+#69 := (= #65 #68)
+#70 := [rewrite]: #69
+#66 := (= #38 #65)
+#63 := (iff #11 #62)
+#64 := [rewrite]: #63
+#67 := [monotonicity #64]: #66
+#72 := [trans #67 #70]: #71
+#75 := [monotonicity #72]: #74
+#82 := [monotonicity #75 #79]: #81
+#87 := [trans #82 #85]: #86
+#59 := (iff #16 #56)
+#49 := (iff #10 #44)
+#52 := (not #49)
+#57 := (iff #52 #56)
+#58 := [rewrite]: #57
+#53 := (iff #16 #52)
+#50 := (iff #15 #49)
+#47 := (iff #14 #44)
+#41 := (= #38 f3)
+#45 := (iff #41 #44)
+#46 := [rewrite]: #45
+#42 := (iff #14 #41)
+#39 := (= #13 #38)
+#36 := (= #12 #35)
+#37 := [rewrite]: #36
+#40 := [monotonicity #37]: #39
+#43 := [monotonicity #40]: #42
+#48 := [trans #43 #46]: #47
+#51 := [monotonicity #48]: #50
+#54 := [monotonicity #51]: #53
+#60 := [trans #54 #58]: #59
+#89 := [trans #60 #87]: #88
+#33 := [asserted]: #16
+#90 := [mp #33 #89]: #83
+#93 := (not #83)
+#91 := (or #62 #95 #93)
+#100 := [def-axiom]: #91
+#102 := [unit-resolution #100 #90]: #101
+#109 := [unit-resolution #102 #108 #107]: false
+#110 := [lemma #109]: #62
+#99 := (or #61 #96)
+#94 := [def-axiom]: #99
+#116 := [unit-resolution #94 #110]: #96
+#119 := [symm #116]: #118
+#105 := (or #61 #73)
+#103 := (or #61 #73 #93)
+#104 := [def-axiom]: #103
+#106 := [unit-resolution #104 #90]: #105
+#117 := [unit-resolution #106 #110]: #73
+#120 := [trans #117 #119]: #111
+#121 := (not #111)
+#122 := (or #121 #115)
+#123 := [th-lemma]: #122
+#124 := [unit-resolution #123 #120]: #115
+[th-lemma #110 #124]: false
+unsat
 fc9f250b0540aecbc3a008bfe364829a1228d272 347 0
 #2 := false
 #183 := -2::int
@@ -42257,6 +43254,329 @@
 #661 := [mp #665 #660]: #670
 [unit-resolution #661 #764 #194]: false
 unsat
+443d726f158e992756000aa2462469d5494d62ee 149 0
+#2 := false
+#9 := 0::int
+decl f3 :: int
+#8 := f3
+#34 := -1::int
+#35 := (* -1::int f3)
+#61 := (>= f3 0::int)
+#68 := (ite #61 f3 #35)
+#111 := (* -1::int #68)
+#156 := (+ f3 #111)
+#157 := (<= #156 0::int)
+#98 := (= f3 #68)
+#10 := (<= f3 0::int)
+#55 := (not #10)
+#78 := (+ f3 #68)
+#77 := (= #78 0::int)
+#93 := (not #77)
+#115 := [hypothesis]: #93
+#95 := (>= #78 0::int)
+#112 := (+ #35 #111)
+#113 := (<= #112 0::int)
+#73 := (= #35 #68)
+#62 := (not #61)
+#131 := (or #62 #77)
+#118 := (= f3 0::int)
+#109 := (or #10 #77)
+#85 := (iff #55 #77)
+#12 := (- f3)
+#11 := (< f3 0::int)
+#13 := (ite #11 #12 f3)
+#14 := (= #13 #12)
+#15 := (iff #10 #14)
+#16 := (not #15)
+#90 := (iff #16 #85)
+#38 := (ite #11 #35 f3)
+#44 := (= #35 #38)
+#56 := (iff #44 #55)
+#88 := (iff #56 #85)
+#82 := (iff #77 #55)
+#86 := (iff #82 #85)
+#87 := [rewrite]: #86
+#83 := (iff #56 #82)
+#80 := (iff #44 #77)
+#76 := (iff #73 #77)
+#79 := [rewrite]: #76
+#74 := (iff #44 #73)
+#71 := (= #38 #68)
+#65 := (ite #62 #35 f3)
+#69 := (= #65 #68)
+#70 := [rewrite]: #69
+#66 := (= #38 #65)
+#63 := (iff #11 #62)
+#64 := [rewrite]: #63
+#67 := [monotonicity #64]: #66
+#72 := [trans #67 #70]: #71
+#75 := [monotonicity #72]: #74
+#81 := [trans #75 #79]: #80
+#84 := [monotonicity #81]: #83
+#89 := [trans #84 #87]: #88
+#59 := (iff #16 #56)
+#49 := (iff #10 #44)
+#52 := (not #49)
+#57 := (iff #52 #56)
+#58 := [rewrite]: #57
+#53 := (iff #16 #52)
+#50 := (iff #15 #49)
+#47 := (iff #14 #44)
+#41 := (= #38 #35)
+#45 := (iff #41 #44)
+#46 := [rewrite]: #45
+#42 := (iff #14 #41)
+#36 := (= #12 #35)
+#37 := [rewrite]: #36
+#39 := (= #13 #38)
+#40 := [monotonicity #37]: #39
+#43 := [monotonicity #40 #37]: #42
+#48 := [trans #43 #46]: #47
+#51 := [monotonicity #48]: #50
+#54 := [monotonicity #51]: #53
+#60 := [trans #54 #58]: #59
+#91 := [trans #60 #89]: #90
+#33 := [asserted]: #16
+#92 := [mp #33 #91]: #85
+#102 := (not #85)
+#107 := (or #10 #77 #102)
+#108 := [def-axiom]: #107
+#110 := [unit-resolution #108 #92]: #109
+#116 := [unit-resolution #110 #115]: #10
+#117 := [hypothesis]: #61
+#119 := [th-lemma #117 #116]: #118
+#127 := (= #78 f3)
+#125 := (= #68 f3)
+#99 := (or #62 #98)
+#100 := [def-axiom]: #99
+#120 := [unit-resolution #100 #117]: #98
+#126 := [symm #120]: #125
+#123 := (= #78 #68)
+#121 := (= #68 #78)
+#122 := [th-lemma #117 #116]: #121
+#124 := [symm #122]: #123
+#128 := [trans #124 #126]: #127
+#129 := [trans #128 #119]: #77
+#130 := [unit-resolution #115 #129]: false
+#132 := [lemma #130]: #131
+#133 := [unit-resolution #132 #115]: #62
+#101 := (or #61 #73)
+#96 := [def-axiom]: #101
+#134 := [unit-resolution #96 #133]: #73
+#135 := (not #73)
+#136 := (or #135 #113)
+#137 := [th-lemma]: #136
+#138 := [unit-resolution #137 #134]: #113
+#139 := (not #113)
+#140 := (or #95 #139)
+#141 := [th-lemma]: #140
+#142 := [unit-resolution #141 #138]: #95
+#97 := (<= #78 0::int)
+#114 := (>= #112 0::int)
+#143 := (or #135 #114)
+#144 := [th-lemma]: #143
+#145 := [unit-resolution #144 #134]: #114
+#146 := (not #114)
+#147 := (or #97 #146)
+#148 := [th-lemma]: #147
+#149 := [unit-resolution #148 #145]: #97
+#151 := (not #95)
+#150 := (not #97)
+#152 := (or #77 #150 #151)
+#153 := [th-lemma]: #152
+#154 := [unit-resolution #153 #149 #142 #115]: false
+#155 := [lemma #154]: #77
+#105 := (or #55 #93)
+#103 := (or #55 #93 #102)
+#104 := [def-axiom]: #103
+#106 := [unit-resolution #104 #92]: #105
+#159 := [unit-resolution #106 #155]: #55
+#160 := (or #61 #10)
+#161 := [th-lemma]: #160
+#162 := [unit-resolution #161 #159]: #61
+#163 := [unit-resolution #100 #162]: #98
+#164 := (not #98)
+#165 := (or #164 #157)
+#166 := [th-lemma]: #165
+#167 := [unit-resolution #166 #163]: #157
+#168 := (or #93 #97)
+#169 := [th-lemma]: #168
+#170 := [unit-resolution #169 #155]: #97
+[th-lemma #159 #170 #167]: false
+unsat
+5d75c12dca04a93af891e80a12133cb48b0e8ce3 114 0
+#2 := false
+#9 := 0::int
+decl f3 :: int
+#8 := f3
+#35 := -1::int
+#36 := (* -1::int f3)
+#67 := (>= f3 0::int)
+#74 := (ite #67 f3 #36)
+#88 := (* -1::int #74)
+#127 := (+ #36 #88)
+#137 := (<= #127 0::int)
+#114 := (= #36 #74)
+#68 := (not #67)
+#125 := (+ f3 #88)
+#126 := (<= #125 0::int)
+#113 := (= f3 #74)
+#128 := [hypothesis]: #67
+#115 := (or #68 #113)
+#116 := [def-axiom]: #115
+#129 := [unit-resolution #116 #128]: #113
+#130 := (not #113)
+#131 := (or #130 #126)
+#132 := [th-lemma]: #131
+#133 := [unit-resolution #132 #129]: #126
+#83 := (>= #74 0::int)
+#82 := (not #83)
+#94 := (ite #83 #74 #88)
+#99 := (= #74 #94)
+#102 := (not #99)
+#11 := (- f3)
+#10 := (< f3 0::int)
+#12 := (ite #10 #11 f3)
+#14 := (- #12)
+#13 := (< #12 0::int)
+#15 := (ite #13 #14 #12)
+#16 := (= #15 #12)
+#17 := (not #16)
+#105 := (iff #17 #102)
+#39 := (ite #10 #36 f3)
+#48 := (* -1::int #39)
+#42 := (< #39 0::int)
+#53 := (ite #42 #48 #39)
+#59 := (= #39 #53)
+#64 := (not #59)
+#103 := (iff #64 #102)
+#100 := (iff #59 #99)
+#97 := (= #53 #94)
+#91 := (ite #82 #88 #74)
+#95 := (= #91 #94)
+#96 := [rewrite]: #95
+#92 := (= #53 #91)
+#77 := (= #39 #74)
+#71 := (ite #68 #36 f3)
+#75 := (= #71 #74)
+#76 := [rewrite]: #75
+#72 := (= #39 #71)
+#69 := (iff #10 #68)
+#70 := [rewrite]: #69
+#73 := [monotonicity #70]: #72
+#78 := [trans #73 #76]: #77
+#89 := (= #48 #88)
+#90 := [monotonicity #78]: #89
+#86 := (iff #42 #82)
+#79 := (< #74 0::int)
+#84 := (iff #79 #82)
+#85 := [rewrite]: #84
+#80 := (iff #42 #79)
+#81 := [monotonicity #78]: #80
+#87 := [trans #81 #85]: #86
+#93 := [monotonicity #87 #90 #78]: #92
+#98 := [trans #93 #96]: #97
+#101 := [monotonicity #78 #98]: #100
+#104 := [monotonicity #101]: #103
+#65 := (iff #17 #64)
+#62 := (iff #16 #59)
+#56 := (= #53 #39)
+#60 := (iff #56 #59)
+#61 := [rewrite]: #60
+#57 := (iff #16 #56)
+#40 := (= #12 #39)
+#37 := (= #11 #36)
+#38 := [rewrite]: #37
+#41 := [monotonicity #38]: #40
+#54 := (= #15 #53)
+#51 := (= #14 #48)
+#45 := (- #39)
+#49 := (= #45 #48)
+#50 := [rewrite]: #49
+#46 := (= #14 #45)
+#47 := [monotonicity #41]: #46
+#52 := [trans #47 #50]: #51
+#43 := (iff #13 #42)
+#44 := [monotonicity #41]: #43
+#55 := [monotonicity #44 #52 #41]: #54
+#58 := [monotonicity #55 #41]: #57
+#63 := [trans #58 #61]: #62
+#66 := [monotonicity #63]: #65
+#106 := [trans #66 #104]: #105
+#34 := [asserted]: #17
+#107 := [mp #34 #106]: #102
+#108 := (or #82 #99)
+#117 := [def-axiom]: #108
+#134 := [unit-resolution #117 #107]: #82
+#135 := [th-lemma #128 #134 #133]: false
+#136 := [lemma #135]: #68
+#111 := (or #67 #114)
+#112 := [def-axiom]: #111
+#139 := [unit-resolution #112 #136]: #114
+#140 := (not #114)
+#141 := (or #140 #137)
+#142 := [th-lemma]: #141
+#143 := [unit-resolution #142 #139]: #137
+[th-lemma #136 #134 #143]: false
+unsat
+1c29e683939b29daf6c0c39d6dc5211f90494d2b 57 0
+#2 := false
+#36 := 0::int
+decl f4 :: int
+#9 := f4
+decl f3 :: int
+#8 := f3
+#33 := -1::int
+#34 := (* -1::int f4)
+#35 := (+ f3 #34)
+#37 := (<= #35 0::int)
+#40 := (ite #37 f3 f4)
+#48 := (* -1::int #40)
+#49 := (+ f3 #48)
+#47 := (>= #49 0::int)
+#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::int)
+#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
 66247d826364718443f0bc07b82495d982768eb5 327 0
 #2 := false
 #11 := 0::int
@@ -42585,2131 +43905,6 @@
 #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)
-#39 := 1::int
-decl f5 :: int
-#38 := f5
-#40 := (f3 f5 1::int)
-#704 := (>= #40 0::int)
-#752 := (= #40 0::int)
-#753 := (= f5 0::int)
-#591 := (not #753)
-#493 := [hypothesis]: #591
-#405 := (<= f5 0::int)
-#416 := (>= f5 0::int)
-#737 := (not #416)
-#451 := (not #405)
-#446 := (or #451 #737)
-#537 := (not #446)
-#69 := -1::int
-#427 := (* -1::int f5)
-#437 := (div #427 -1::int)
-#717 := (* -1::int #437)
-#715 := (+ #40 #717)
-#720 := (= #715 0::int)
-#428 := (div f5 1::int)
-#432 := (* -1::int #428)
-#411 := (+ #40 #432)
-#746 := (= #411 0::int)
-#711 := (ite #446 #746 #720)
-#483 := (or #753 #711)
-#712 := (ite #753 #752 #711)
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#10 := (f3 #8 #9)
-#754 := (pattern #10)
-#73 := (* -1::int #9)
-#70 := (* -1::int #8)
-#76 := (div #70 #73)
-#257 := (* -1::int #76)
-#258 := (+ #10 #257)
-#259 := (= #258 0::int)
-#21 := (div #8 #9)
-#254 := (* -1::int #21)
-#255 := (+ #10 #254)
-#256 := (= #255 0::int)
-#95 := (<= #9 0::int)
-#91 := (<= #8 0::int)
-#212 := (or #91 #95)
-#213 := (not #212)
-#102 := (>= #8 0::int)
-#204 := (or #95 #102)
-#205 := (not #204)
-#219 := (or #205 #213)
-#260 := (ite #219 #256 #259)
-#253 := (= #10 0::int)
-#13 := (= #9 0::int)
-#12 := (= #8 0::int)
-#14 := (or #12 #13)
-#261 := (ite #14 #253 #260)
-#755 := (forall (vars (?v0 int) (?v1 int)) (:pat #754) #261)
-#264 := (forall (vars (?v0 int) (?v1 int)) #261)
-#758 := (iff #264 #755)
-#756 := (iff #261 #261)
-#757 := [refl]: #756
-#759 := [quant-intro #757]: #758
-#224 := (ite #219 #21 #76)
-#227 := (ite #14 0::int #224)
-#230 := (= #10 #227)
-#233 := (forall (vars (?v0 int) (?v1 int)) #230)
-#265 := (iff #233 #264)
-#262 := (iff #230 #261)
-#263 := [rewrite]: #262
-#266 := [quant-intro #263]: #265
-#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)
-#234 := (iff #121 #233)
-#231 := (iff #118 #230)
-#228 := (= #115 #227)
-#225 := (= #112 #224)
-#222 := (iff #109 #219)
-#216 := (or #213 #205)
-#220 := (iff #216 #219)
-#221 := [rewrite]: #220
-#217 := (iff #109 #216)
-#214 := (iff #106 #205)
-#215 := [rewrite]: #214
-#202 := (iff #99 #213)
-#203 := [rewrite]: #202
-#218 := [monotonicity #203 #215]: #217
-#223 := [trans #218 #221]: #222
-#226 := [monotonicity #223]: #225
-#229 := [monotonicity #226]: #228
-#232 := [monotonicity #229]: #231
-#235 := [quant-intro #232]: #234
-#208 := (~ #121 #121)
-#206 := (~ #118 #118)
-#207 := [refl]: #206
-#209 := [nnf-pos #207]: #208
-#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
-#199 := [mp~ #126 #209]: #121
-#236 := [mp #199 #235]: #233
-#267 := [mp #236 #266]: #264
-#760 := [mp #267 #759]: #755
-#714 := (not #755)
-#555 := (or #714 #712)
-#426 := (* -1::int 1::int)
-#418 := (div #427 #426)
-#429 := (* -1::int #418)
-#430 := (+ #40 #429)
-#431 := (= #430 0::int)
-#748 := (<= 1::int 0::int)
-#535 := (or #405 #748)
-#742 := (not #535)
-#417 := (or #748 #416)
-#749 := (not #417)
-#750 := (or #749 #742)
-#751 := (ite #750 #746 #431)
-#747 := (= 1::int 0::int)
-#743 := (or #753 #747)
-#403 := (ite #743 #752 #751)
-#556 := (or #714 #403)
-#699 := (iff #556 #555)
-#701 := (iff #555 #555)
-#702 := [rewrite]: #701
-#713 := (iff #403 #712)
-#706 := (iff #751 #711)
-#709 := (iff #431 #720)
-#719 := (= #430 #715)
-#718 := (= #429 #717)
-#438 := (= #418 #437)
-#726 := (= #426 -1::int)
-#436 := [rewrite]: #726
-#439 := [monotonicity #436]: #438
-#433 := [monotonicity #439]: #718
-#716 := [monotonicity #433]: #719
-#710 := [monotonicity #716]: #709
-#725 := (iff #750 #446)
-#347 := (or #737 #451)
-#447 := (iff #347 #446)
-#721 := [rewrite]: #447
-#723 := (iff #750 #347)
-#452 := (iff #742 #451)
-#728 := (iff #535 #405)
-#380 := (or #405 false)
-#741 := (iff #380 #405)
-#727 := [rewrite]: #741
-#740 := (iff #535 #380)
-#395 := (iff #748 false)
-#731 := [rewrite]: #395
-#738 := [monotonicity #731]: #740
-#450 := [trans #738 #727]: #728
-#453 := [monotonicity #450]: #452
-#739 := (iff #749 #737)
-#736 := (iff #417 #416)
-#732 := (or false #416)
-#735 := (iff #732 #416)
-#730 := [rewrite]: #735
-#733 := (iff #417 #732)
-#734 := [monotonicity #731]: #733
-#374 := [trans #734 #730]: #736
-#379 := [monotonicity #374]: #739
-#724 := [monotonicity #379 #453]: #723
-#722 := [trans #724 #721]: #725
-#707 := [monotonicity #722 #710]: #706
-#393 := (iff #743 #753)
-#404 := (or #753 false)
-#729 := (iff #404 #753)
-#392 := [rewrite]: #729
-#745 := (iff #743 #404)
-#408 := (iff #747 false)
-#744 := [rewrite]: #408
-#388 := [monotonicity #744]: #745
-#394 := [trans #388 #392]: #393
-#708 := [monotonicity #394 #707]: #713
-#700 := [monotonicity #708]: #699
-#696 := [trans #700 #702]: #699
-#557 := [quant-inst]: #556
-#697 := [mp #557 #696]: #555
-#548 := [unit-resolution #697 #760]: #712
-#583 := (not #712)
-#594 := (or #583 #753 #711)
-#595 := [def-axiom]: #594
-#484 := [unit-resolution #595 #548]: #483
-#485 := [unit-resolution #484 #493]: #711
-#684 := (not #746)
-#692 := (>= #411 0::int)
-#422 := (not #692)
-decl f4 :: (-> int int int)
-#42 := (f4 f5 1::int)
-#600 := (>= #42 0::int)
-#682 := (= #42 0::int)
-#624 := (or #451 #737 #753)
-#481 := (or #624 #753)
-#599 := (not #624)
-#491 := [hypothesis]: #599
-#616 := (or #624 #405)
-#617 := [def-axiom]: #616
-#494 := [unit-resolution #617 #491]: #405
-#613 := (or #624 #416)
-#618 := [def-axiom]: #613
-#476 := [unit-resolution #618 #491]: #416
-#478 := (or #753 #451 #737)
-#479 := [th-lemma]: #478
-#480 := [unit-resolution #479 #476 #494 #493]: false
-#482 := [lemma #480]: #481
-#486 := [unit-resolution #482 #493]: #624
-#488 := (or #599 #682)
-#695 := (mod #427 -1::int)
-#641 := (+ #42 #695)
-#644 := (= #641 0::int)
-#627 := (ite #624 #682 #644)
-#29 := (f4 #8 #9)
-#761 := (pattern #29)
-#128 := (mod #70 #73)
-#273 := (+ #29 #128)
-#274 := (= #273 0::int)
-#30 := (mod #8 #9)
-#270 := (* -1::int #30)
-#271 := (+ #29 #270)
-#272 := (= #271 0::int)
-#275 := (ite #219 #272 #274)
-#269 := (= #29 0::int)
-#276 := (ite #12 #269 #275)
-#268 := (= #8 #29)
-#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
-#134 := (* -1::int #128)
-#237 := (ite #219 #30 #134)
-#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
-#154 := (ite #109 #30 #134)
-#157 := (ite #12 0::int #154)
-#160 := (ite #13 #8 #157)
-#163 := (= #29 #160)
-#166 := (forall (vars (?v0 int) (?v1 int)) #163)
-#250 := (iff #166 #249)
-#247 := (iff #163 #246)
-#244 := (= #160 #243)
-#241 := (= #157 #240)
-#238 := (= #154 #237)
-#239 := [monotonicity #223]: #238
-#242 := [monotonicity #239]: #241
-#245 := [monotonicity #242]: #244
-#248 := [monotonicity #245]: #247
-#251 := [quant-intro #248]: #250
-#200 := (~ #166 #166)
-#197 := (~ #163 #163)
-#210 := [refl]: #197
-#201 := [nnf-pos #210]: #200
-#31 := (mod #22 #23)
-#32 := (- #31)
-#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)
-#169 := (iff #37 #166)
-#139 := (ite #66 #30 #134)
-#142 := (ite #12 0::int #139)
-#145 := (ite #13 #8 #142)
-#148 := (= #29 #145)
-#151 := (forall (vars (?v0 int) (?v1 int)) #148)
-#167 := (iff #151 #166)
-#164 := (iff #148 #163)
-#161 := (= #145 #160)
-#158 := (= #142 #157)
-#155 := (= #139 #154)
-#156 := [monotonicity #111]: #155
-#159 := [monotonicity #156]: #158
-#162 := [monotonicity #159]: #161
-#165 := [monotonicity #162]: #164
-#168 := [quant-intro #165]: #167
-#152 := (iff #37 #151)
-#149 := (iff #36 #148)
-#146 := (= #35 #145)
-#143 := (= #34 #142)
-#140 := (= #33 #139)
-#137 := (= #32 #134)
-#131 := (- #128)
-#135 := (= #131 #134)
-#136 := [rewrite]: #135
-#132 := (= #32 #131)
-#129 := (= #31 #128)
-#130 := [monotonicity #72 #75]: #129
-#133 := [monotonicity #130]: #132
-#138 := [trans #133 #136]: #137
-#141 := [monotonicity #68 #138]: #140
-#144 := [monotonicity #141]: #143
-#147 := [monotonicity #144]: #146
-#150 := [monotonicity #147]: #149
-#153 := [quant-intro #150]: #152
-#170 := [trans #153 #168]: #169
-#127 := [asserted]: #37
-#171 := [mp #127 #170]: #166
-#211 := [mp~ #171 #201]: #166
-#252 := [mp #211 #251]: #249
-#283 := [mp #252 #282]: #280
-#767 := [mp #283 #766]: #762
-#603 := (not #762)
-#496 := (or #603 #627)
-#670 := (mod #427 #426)
-#667 := (+ #42 #670)
-#669 := (= #667 0::int)
-#693 := (mod f5 1::int)
-#679 := (* -1::int #693)
-#680 := (+ #42 #679)
-#681 := (= #680 0::int)
-#677 := (ite #750 #681 #669)
-#671 := (ite #753 #682 #677)
-#672 := (= f5 #42)
-#673 := (ite #747 #672 #671)
-#607 := (or #603 #673)
-#609 := (iff #607 #496)
-#610 := (iff #496 #496)
-#611 := [rewrite]: #610
-#601 := (iff #673 #627)
-#629 := (or #446 #753)
-#630 := (ite #629 #682 #644)
-#622 := (iff #630 #627)
-#625 := (iff #629 #624)
-#626 := [rewrite]: #625
-#628 := [monotonicity #626]: #622
-#620 := (iff #673 #630)
-#636 := (ite false #672 #630)
-#623 := (iff #636 #630)
-#619 := [rewrite]: #623
-#631 := (iff #673 #636)
-#634 := (iff #671 #630)
-#647 := (ite #446 #682 #644)
-#650 := (ite #753 #682 #647)
-#632 := (iff #650 #630)
-#633 := [rewrite]: #632
-#640 := (iff #671 #650)
-#648 := (iff #677 #647)
-#645 := (iff #669 #644)
-#642 := (= #667 #641)
-#638 := (= #670 #695)
-#639 := [monotonicity #436]: #638
-#643 := [monotonicity #639]: #642
-#646 := [monotonicity #643]: #645
-#654 := (iff #681 #682)
-#656 := (= #680 #42)
-#661 := (+ #42 0::int)
-#653 := (= #661 #42)
-#655 := [rewrite]: #653
-#666 := (= #680 #661)
-#664 := (= #679 0::int)
-#675 := (* -1::int 0::int)
-#662 := (= #675 0::int)
-#663 := [rewrite]: #662
-#659 := (= #679 #675)
-#674 := (= #693 0::int)
-#668 := [rewrite]: #674
-#660 := [monotonicity #668]: #659
-#665 := [trans #660 #663]: #664
-#652 := [monotonicity #665]: #666
-#657 := [trans #652 #655]: #656
-#658 := [monotonicity #657]: #654
-#649 := [monotonicity #722 #658 #646]: #648
-#651 := [monotonicity #649]: #640
-#635 := [trans #651 #633]: #634
-#637 := [monotonicity #744 #635]: #631
-#621 := [trans #637 #619]: #620
-#602 := [trans #621 #628]: #601
-#614 := [monotonicity #602]: #609
-#612 := [trans #614 #611]: #609
-#608 := [quant-inst]: #607
-#615 := [mp #608 #612]: #496
-#487 := [unit-resolution #615 #767]: #627
-#581 := (not #627)
-#540 := (or #581 #599 #682)
-#571 := [def-axiom]: #540
-#477 := [unit-resolution #571 #487]: #488
-#489 := [unit-resolution #477 #486]: #682
-#582 := (not #682)
-#396 := (or #582 #600)
-#397 := [th-lemma]: #396
-#398 := [unit-resolution #397 #489]: #600
-#421 := (not #600)
-#389 := (or #422 #421)
-#542 := (+ f5 #432)
-#541 := (= #542 0::int)
-#1 := true
-#60 := [true-axiom]: true
-#515 := (or false #541)
-#505 := [th-lemma]: #515
-#495 := [unit-resolution #505 #60]: #541
-#443 := (<= #542 0::int)
-#420 := (not #443)
-#185 := (* -1::int #42)
-#184 := (* -1::int #40)
-#186 := (+ #184 #185)
-#187 := (+ f5 #186)
-#341 := (<= #187 0::int)
-#444 := (not #341)
-#425 := (>= #187 0::int)
-#703 := (<= #40 0::int)
-#499 := (not #425)
-#507 := [hypothesis]: #499
-#461 := (or #753 #425)
-#694 := (<= #411 0::int)
-#504 := (not #694)
-#605 := (<= #42 0::int)
-#466 := (or #582 #605)
-#468 := [th-lemma]: #466
-#469 := [unit-resolution #468 #489]: #605
-#503 := (not #605)
-#490 := (or #503 #425 #504)
-#510 := [hypothesis]: #605
-#549 := (>= #542 0::int)
-#497 := (not #541)
-#498 := (or #497 #549)
-#501 := [th-lemma]: #498
-#502 := [unit-resolution #501 #495]: #549
-#506 := [hypothesis]: #694
-#500 := [th-lemma #507 #506 #502 #510]: false
-#492 := [lemma #500]: #490
-#470 := [unit-resolution #492 #469 #507]: #504
-#471 := (or #684 #694)
-#472 := [th-lemma]: #471
-#473 := [unit-resolution #472 #470]: #684
-#579 := (not #711)
-#538 := (or #579 #537 #746)
-#686 := [def-axiom]: #538
-#474 := [unit-resolution #686 #473 #485]: #537
-#698 := (or #446 #405)
-#705 := [def-axiom]: #698
-#467 := [unit-resolution #705 #474]: #405
-#534 := (or #446 #416)
-#536 := [def-axiom]: #534
-#475 := [unit-resolution #536 #474]: #416
-#460 := [unit-resolution #479 #475 #467 #493]: false
-#462 := [lemma #460]: #461
-#464 := [unit-resolution #462 #507]: #753
-#463 := (or #591 #752)
-#592 := (or #583 #591 #752)
-#593 := [def-axiom]: #592
-#465 := [unit-resolution #593 #548]: #463
-#454 := [unit-resolution #465 #464]: #752
-#688 := (not #752)
-#455 := (or #688 #703)
-#456 := [th-lemma]: #455
-#448 := [unit-resolution #456 #454]: #703
-#457 := (or #591 #416)
-#458 := [th-lemma]: #457
-#449 := [unit-resolution #458 #464]: #416
-#598 := (or #624 #591)
-#544 := [def-axiom]: #598
-#459 := [unit-resolution #544 #464]: #624
-#440 := [unit-resolution #477 #459]: #682
-#441 := [unit-resolution #468 #440]: #605
-#442 := [th-lemma #441 #507 #449 #448]: false
-#434 := [lemma #442]: #425
-#412 := (or #444 #499)
-#188 := (= #187 0::int)
-#191 := (not #188)
-#41 := (* #40 1::int)
-#43 := (+ #41 #42)
-#44 := (= f5 #43)
-#45 := (not #44)
-#194 := (iff #45 #191)
-#175 := (+ #40 #42)
-#178 := (= f5 #175)
-#181 := (not #178)
-#192 := (iff #181 #191)
-#189 := (iff #178 #188)
-#190 := [rewrite]: #189
-#193 := [monotonicity #190]: #192
-#182 := (iff #45 #181)
-#179 := (iff #44 #178)
-#176 := (= #43 #175)
-#173 := (= #41 #40)
-#174 := [rewrite]: #173
-#177 := [monotonicity #174]: #176
-#180 := [monotonicity #177]: #179
-#183 := [monotonicity #180]: #182
-#195 := [trans #183 #193]: #194
-#172 := [asserted]: #45
-#196 := [mp #172 #195]: #191
-#435 := (or #188 #444 #499)
-#445 := [th-lemma]: #435
-#414 := [unit-resolution #445 #196]: #412
-#415 := [unit-resolution #414 #434]: #444
-#406 := [hypothesis]: #600
-#419 := [hypothesis]: #692
-#423 := (or #420 #421 #341 #422)
-#413 := [th-lemma]: #423
-#424 := [unit-resolution #413 #419 #406 #415]: #420
-#407 := (or #497 #443)
-#409 := [th-lemma]: #407
-#410 := [unit-resolution #409 #424 #495]: false
-#391 := [lemma #410]: #389
-#399 := [unit-resolution #391 #398]: #422
-#400 := (or #684 #692)
-#401 := [th-lemma]: #400
-#390 := [unit-resolution #401 #399]: #684
-#402 := [unit-resolution #686 #390 #485]: #537
-#383 := [unit-resolution #705 #402]: #405
-#385 := [unit-resolution #536 #402]: #416
-#386 := [unit-resolution #479 #385 #383 #493]: false
-#384 := [lemma #386]: #753
-#387 := [unit-resolution #465 #384]: #752
-#375 := (or #688 #704)
-#377 := [th-lemma]: #375
-#378 := [unit-resolution #377 #387]: #704
-#381 := (or #591 #405)
-#376 := [th-lemma]: #381
-#382 := [unit-resolution #376 #384]: #405
-#357 := [unit-resolution #544 #384]: #624
-#361 := [unit-resolution #477 #357]: #682
-#362 := [unit-resolution #397 #361]: #600
-[th-lemma #362 #415 #382 #378]: false
-unsat
-f7d21861165159375bc45e6364e84fc0183b7f2c 549 0
-#2 := false
-#11 := 0::int
-decl f3 :: (-> int int int)
-#39 := 3::int
-decl f5 :: int
-#38 := f5
-#40 := (f3 f5 3::int)
-#703 := (>= #40 0::int)
-#754 := (= #40 0::int)
-#750 := (= f5 0::int)
-#690 := (not #750)
-#562 := [hypothesis]: #690
-#745 := (>= f5 0::int)
-#751 := (<= f5 0::int)
-#453 := (not #751)
-#377 := (not #745)
-#456 := (or #377 #453)
-#706 := (not #456)
-#185 := -3::int
-#69 := -1::int
-#430 := (* -1::int f5)
-#450 := (div #430 -3::int)
-#725 := (* -1::int #450)
-#440 := (+ #40 #725)
-#720 := (= #440 0::int)
-#434 := (div f5 3::int)
-#431 := (* -1::int #434)
-#435 := (+ #40 #431)
-#414 := (= #435 0::int)
-#718 := (ite #456 #414 #720)
-#564 := (or #750 #718)
-#723 := (ite #750 #754 #718)
-#9 := (:var 0 int)
-#8 := (:var 1 int)
-#10 := (f3 #8 #9)
-#757 := (pattern #10)
-#73 := (* -1::int #9)
-#70 := (* -1::int #8)
-#76 := (div #70 #73)
-#259 := (* -1::int #76)
-#260 := (+ #10 #259)
-#261 := (= #260 0::int)
-#21 := (div #8 #9)
-#256 := (* -1::int #21)
-#257 := (+ #10 #256)
-#258 := (= #257 0::int)
-#95 := (<= #9 0::int)
-#91 := (<= #8 0::int)
-#214 := (or #91 #95)
-#215 := (not #214)
-#102 := (>= #8 0::int)
-#206 := (or #95 #102)
-#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)
-#758 := (forall (vars (?v0 int) (?v1 int)) (:pat #757) #263)
-#266 := (forall (vars (?v0 int) (?v1 int)) #263)
-#761 := (iff #266 #758)
-#759 := (iff #263 #263)
-#760 := [refl]: #759
-#762 := [quant-intro #760]: #761
-#226 := (ite #221 #21 #76)
-#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
-#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)
-#236 := (iff #121 #235)
-#233 := (iff #118 #232)
-#230 := (= #115 #229)
-#227 := (= #112 #226)
-#224 := (iff #109 #221)
-#218 := (or #215 #207)
-#222 := (iff #218 #221)
-#223 := [rewrite]: #222
-#219 := (iff #109 #218)
-#216 := (iff #106 #207)
-#217 := [rewrite]: #216
-#204 := (iff #99 #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 := (~ #121 #121)
-#208 := (~ #118 #118)
-#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)
-#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
-#201 := [mp~ #126 #211]: #121
-#238 := [mp #201 #237]: #235
-#269 := [mp #238 #268]: #266
-#763 := [mp #269 #762]: #758
-#714 := (not #758)
-#709 := (or #714 #723)
-#429 := (* -1::int 3::int)
-#421 := (div #430 #429)
-#432 := (* -1::int #421)
-#433 := (+ #40 #432)
-#361 := (= #433 0::int)
-#749 := (<= 3::int 0::int)
-#408 := (or #751 #749)
-#538 := (not #408)
-#419 := (or #749 #745)
-#420 := (not #419)
-#752 := (or #420 #538)
-#753 := (ite #752 #414 #361)
-#755 := (= 3::int 0::int)
-#756 := (or #750 #755)
-#746 := (ite #756 #754 #753)
-#710 := (or #714 #746)
-#716 := (iff #710 #709)
-#717 := (iff #709 #709)
-#558 := [rewrite]: #717
-#712 := (iff #746 #723)
-#722 := (iff #753 #718)
-#721 := (iff #361 #720)
-#441 := (= #433 #440)
-#729 := (= #432 #725)
-#724 := (= #421 #450)
-#727 := (= #429 -3::int)
-#449 := [rewrite]: #727
-#728 := [monotonicity #449]: #724
-#439 := [monotonicity #728]: #729
-#442 := [monotonicity #439]: #441
-#436 := [monotonicity #442]: #721
-#349 := (iff #752 #456)
-#454 := (iff #538 #453)
-#730 := (iff #408 #751)
-#382 := (or #751 false)
-#741 := (iff #382 #751)
-#744 := [rewrite]: #741
-#383 := (iff #408 #382)
-#397 := (iff #749 false)
-#398 := [rewrite]: #397
-#743 := [monotonicity #398]: #383
-#731 := [trans #743 #744]: #730
-#455 := [monotonicity #731]: #454
-#740 := (iff #420 #377)
-#733 := (iff #419 #745)
-#734 := (or false #745)
-#737 := (iff #734 #745)
-#738 := [rewrite]: #737
-#735 := (iff #419 #734)
-#736 := [monotonicity #398]: #735
-#739 := [trans #736 #738]: #733
-#742 := [monotonicity #739]: #740
-#726 := [monotonicity #742 #455]: #349
-#719 := [monotonicity #726 #436]: #722
-#395 := (iff #756 #750)
-#747 := (or #750 false)
-#391 := (iff #747 #750)
-#732 := [rewrite]: #391
-#407 := (iff #756 #747)
-#406 := (iff #755 false)
-#411 := [rewrite]: #406
-#748 := [monotonicity #411]: #407
-#396 := [trans #748 #732]: #395
-#713 := [monotonicity #396 #719]: #712
-#711 := [monotonicity #713]: #716
-#559 := [trans #711 #558]: #716
-#715 := [quant-inst]: #710
-#560 := [mp #715 #559]: #709
-#563 := [unit-resolution #560 #763]: #723
-#687 := (not #723)
-#592 := (or #687 #750 #718)
-#593 := [def-axiom]: #592
-#565 := [unit-resolution #593 #563]: #564
-#566 := [unit-resolution #565 #562]: #718
-#540 := (mod #430 -3::int)
-decl f4 :: (-> int int int)
-#42 := (f4 f5 3::int)
-#685 := (+ #42 #540)
-#676 := (= #685 0::int)
-#708 := (mod f5 3::int)
-#692 := (* -1::int #708)
-#679 := (+ #42 #692)
-#681 := (= #679 0::int)
-#678 := (ite #456 #681 #676)
-#525 := (or #750 #678)
-#670 := (= #42 0::int)
-#665 := (ite #750 #670 #678)
-#29 := (f4 #8 #9)
-#764 := (pattern #29)
-#128 := (mod #70 #73)
-#275 := (+ #29 #128)
-#276 := (= #275 0::int)
-#30 := (mod #8 #9)
-#272 := (* -1::int #30)
-#273 := (+ #29 #272)
-#274 := (= #273 0::int)
-#277 := (ite #221 #274 #276)
-#271 := (= #29 0::int)
-#278 := (ite #12 #271 #277)
-#270 := (= #8 #29)
-#279 := (ite #13 #270 #278)
-#765 := (forall (vars (?v0 int) (?v1 int)) (:pat #764) #279)
-#282 := (forall (vars (?v0 int) (?v1 int)) #279)
-#768 := (iff #282 #765)
-#766 := (iff #279 #279)
-#767 := [refl]: #766
-#769 := [quant-intro #767]: #768
-#134 := (* -1::int #128)
-#239 := (ite #221 #30 #134)
-#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
-#154 := (ite #109 #30 #134)
-#157 := (ite #12 0::int #154)
-#160 := (ite #13 #8 #157)
-#163 := (= #29 #160)
-#166 := (forall (vars (?v0 int) (?v1 int)) #163)
-#252 := (iff #166 #251)
-#249 := (iff #163 #248)
-#246 := (= #160 #245)
-#243 := (= #157 #242)
-#240 := (= #154 #239)
-#241 := [monotonicity #225]: #240
-#244 := [monotonicity #241]: #243
-#247 := [monotonicity #244]: #246
-#250 := [monotonicity #247]: #249
-#253 := [quant-intro #250]: #252
-#202 := (~ #166 #166)
-#199 := (~ #163 #163)
-#212 := [refl]: #199
-#203 := [nnf-pos #212]: #202
-#31 := (mod #22 #23)
-#32 := (- #31)
-#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)
-#169 := (iff #37 #166)
-#139 := (ite #66 #30 #134)
-#142 := (ite #12 0::int #139)
-#145 := (ite #13 #8 #142)
-#148 := (= #29 #145)
-#151 := (forall (vars (?v0 int) (?v1 int)) #148)
-#167 := (iff #151 #166)
-#164 := (iff #148 #163)
-#161 := (= #145 #160)
-#158 := (= #142 #157)
-#155 := (= #139 #154)
-#156 := [monotonicity #111]: #155
-#159 := [monotonicity #156]: #158
-#162 := [monotonicity #159]: #161
-#165 := [monotonicity #162]: #164
-#168 := [quant-intro #165]: #167
-#152 := (iff #37 #151)
-#149 := (iff #36 #148)
-#146 := (= #35 #145)
-#143 := (= #34 #142)
-#140 := (= #33 #139)
-#137 := (= #32 #134)
-#131 := (- #128)
-#135 := (= #131 #134)
-#136 := [rewrite]: #135
-#132 := (= #32 #131)
-#129 := (= #31 #128)
-#130 := [monotonicity #72 #75]: #129
-#133 := [monotonicity #130]: #132
-#138 := [trans #133 #136]: #137
-#141 := [monotonicity #68 #138]: #140
-#144 := [monotonicity #141]: #143
-#147 := [monotonicity #144]: #146
-#150 := [monotonicity #147]: #149
-#153 := [quant-intro #150]: #152
-#170 := [trans #153 #168]: #169
-#127 := [asserted]: #37
-#171 := [mp #127 #170]: #166
-#213 := [mp~ #171 #203]: #166
-#254 := [mp #213 #253]: #251
-#285 := [mp #254 #284]: #282
-#770 := [mp #285 #769]: #765
-#660 := (not #765)
-#657 := (or #660 #665)
-#598 := (mod #430 #429)
-#691 := (+ #42 #598)
-#688 := (= #691 0::int)
-#673 := (ite #752 #681 #688)
-#672 := (ite #750 #670 #673)
-#682 := (= f5 #42)
-#683 := (ite #755 #682 #672)
-#661 := (or #660 #683)
-#642 := (iff #661 #657)
-#645 := (iff #657 #657)
-#646 := [rewrite]: #645
-#658 := (iff #683 #665)
-#668 := (ite false #682 #665)
-#655 := (iff #668 #665)
-#656 := [rewrite]: #655
-#664 := (iff #683 #668)
-#666 := (iff #672 #665)
-#662 := (iff #673 #678)
-#677 := (iff #688 #676)
-#674 := (= #691 #685)
-#684 := (= #598 #540)
-#680 := [monotonicity #449]: #684
-#675 := [monotonicity #680]: #674
-#671 := [monotonicity #675]: #677
-#663 := [monotonicity #726 #671]: #662
-#667 := [monotonicity #663]: #666
-#669 := [monotonicity #411 #667]: #664
-#659 := [trans #669 #656]: #658
-#644 := [monotonicity #659]: #642
-#647 := [trans #644 #646]: #642
-#641 := [quant-inst]: #661
-#648 := [mp #641 #647]: #657
-#621 := [unit-resolution #648 #770]: #665
-#622 := (not #665)
-#627 := (or #622 #750 #678)
-#628 := [def-axiom]: #627
-#526 := [unit-resolution #628 #621]: #525
-#527 := [unit-resolution #526 #562]: #678
-#696 := (not #718)
-#654 := (not #678)
-#491 := (or #706 #654 #696)
-#652 := (>= #679 0::int)
-#567 := [hypothesis]: #678
-#568 := [hypothesis]: #456
-#632 := (or #654 #706 #681)
-#633 := [def-axiom]: #632
-#569 := [unit-resolution #633 #568 #567]: #681
-#637 := (not #681)
-#493 := (or #637 #652)
-#495 := [th-lemma]: #493
-#496 := [unit-resolution #495 #569]: #652
-#539 := (>= #435 0::int)
-#549 := [hypothesis]: #718
-#697 := (or #696 #706 #414)
-#695 := [def-axiom]: #697
-#550 := [unit-resolution #695 #568 #549]: #414
-#581 := (not #414)
-#494 := (or #581 #539)
-#497 := [th-lemma]: #494
-#479 := [unit-resolution #497 #550]: #539
-#187 := (* -1::int #42)
-#186 := (* -3::int #40)
-#188 := (+ #186 #187)
-#189 := (+ f5 #188)
-#343 := (<= #189 0::int)
-#481 := (not #343)
-#428 := (>= #189 0::int)
-#557 := (not #428)
-#573 := [hypothesis]: #557
-#533 := (or #750 #428)
-#554 := (or #706 #428 #654 #696)
-#651 := (<= #679 0::int)
-#570 := (or #637 #651)
-#571 := [th-lemma]: #570
-#572 := [unit-resolution #571 #569]: #651
-#615 := (* -3::int #434)
-#618 := (+ #615 #692)
-#619 := (+ f5 #618)
-#601 := (>= #619 0::int)
-#614 := (= #619 0::int)
-#1 := true
-#60 := [true-axiom]: true
-#534 := (or false #614)
-#535 := [th-lemma]: #534
-#542 := [unit-resolution #535 #60]: #614
-#544 := (not #614)
-#545 := (or #544 #601)
-#546 := [th-lemma]: #545
-#548 := [unit-resolution #546 #542]: #601
-#537 := (<= #435 0::int)
-#551 := (or #581 #537)
-#552 := [th-lemma]: #551
-#553 := [unit-resolution #552 #550]: #537
-#536 := [th-lemma #553 #548 #573 #572]: false
-#524 := [lemma #536]: #554
-#529 := [unit-resolution #524 #527 #573 #566]: #706
-#704 := (or #456 #745)
-#705 := [def-axiom]: #704
-#530 := [unit-resolution #705 #529]: #745
-#699 := (or #456 #751)
-#700 := [def-axiom]: #699
-#531 := [unit-resolution #700 #529]: #751
-#374 := (or #750 #453 #377)
-#532 := [th-lemma]: #374
-#528 := [unit-resolution #532 #531 #530 #562]: false
-#512 := [lemma #528]: #533
-#515 := [unit-resolution #512 #573]: #750
-#519 := (or #690 #745)
-#521 := [th-lemma]: #519
-#522 := [unit-resolution #521 #515]: #745
-#649 := (<= #42 0::int)
-#520 := (or #690 #670)
-#623 := (or #622 #690 #670)
-#624 := [def-axiom]: #623
-#523 := [unit-resolution #624 #621]: #520
-#511 := [unit-resolution #523 #515]: #670
-#629 := (not #670)
-#514 := (or #629 #649)
-#516 := [th-lemma]: #514
-#517 := [unit-resolution #516 #511]: #649
-#702 := (<= #40 0::int)
-#513 := (or #690 #754)
-#575 := (or #687 #690 #754)
-#590 := [def-axiom]: #575
-#518 := [unit-resolution #590 #563]: #513
-#508 := [unit-resolution #518 #515]: #754
-#586 := (not #754)
-#498 := (or #586 #702)
-#500 := [th-lemma]: #498
-#501 := [unit-resolution #500 #508]: #702
-#504 := [th-lemma #501 #517 #573 #522]: false
-#505 := [lemma #504]: #428
-#484 := (or #481 #557)
-#190 := (= #189 0::int)
-#193 := (not #190)
-#41 := (* #40 3::int)
-#43 := (+ #41 #42)
-#44 := (= f5 #43)
-#45 := (not #44)
-#196 := (iff #45 #193)
-#173 := (* 3::int #40)
-#176 := (+ #173 #42)
-#179 := (= f5 #176)
-#182 := (not #179)
-#194 := (iff #182 #193)
-#191 := (iff #179 #190)
-#192 := [rewrite]: #191
-#195 := [monotonicity #192]: #194
-#183 := (iff #45 #182)
-#180 := (iff #44 #179)
-#177 := (= #43 #176)
-#174 := (= #41 #173)
-#175 := [rewrite]: #174
-#178 := [monotonicity #175]: #177
-#181 := [monotonicity #178]: #180
-#184 := [monotonicity #181]: #183
-#197 := [trans #184 #195]: #196
-#172 := [asserted]: #45
-#198 := [mp #172 #197]: #193
-#482 := (or #190 #481 #557)
-#483 := [th-lemma]: #482
-#485 := [unit-resolution #483 #198]: #484
-#486 := [unit-resolution #485 #505]: #481
-#509 := (<= #619 0::int)
-#487 := (or #544 #509)
-#488 := [th-lemma]: #487
-#489 := [unit-resolution #488 #542]: #509
-#490 := [th-lemma #489 #486 #479 #496]: false
-#480 := [lemma #490]: #491
-#502 := [unit-resolution #480 #527 #566]: #706
-#510 := [unit-resolution #705 #502]: #745
-#503 := [unit-resolution #700 #502]: #751
-#506 := [unit-resolution #532 #503 #510 #562]: false
-#507 := [lemma #506]: #750
-#492 := [unit-resolution #518 #507]: #754
-#469 := (or #586 #703)
-#471 := [th-lemma]: #469
-#472 := [unit-resolution #471 #492]: #703
-#650 := (>= #42 0::int)
-#473 := [unit-resolution #523 #507]: #670
-#474 := (or #629 #650)
-#475 := [th-lemma]: #474
-#476 := [unit-resolution #475 #473]: #650
-#477 := (or #690 #751)
-#470 := [th-lemma]: #477
-#478 := [unit-resolution #470 #507]: #751
-[th-lemma #478 #476 #486 #472]: false
-unsat
-b8a677f41acfea6ca1bd28f53fd89f2c10758abc 75 0
-#2 := false
-#8 := 0::int
-decl f3 :: int
-#9 := f3
-#32 := -1::int
-#33 := (* -1::int f3)
-#45 := (>= f3 0::int)
-#52 := (ite #45 f3 #33)
-#73 := (* -1::int #52)
-#81 := (+ f3 #73)
-#90 := (<= #81 0::int)
-#76 := (= f3 #52)
-#71 := (+ #33 #73)
-#80 := (<= #71 0::int)
-#77 := (= #33 #52)
-#46 := (not #45)
-#82 := [hypothesis]: #46
-#74 := (or #45 #77)
-#75 := [def-axiom]: #74
-#83 := [unit-resolution #75 #82]: #77
-#84 := (not #77)
-#85 := (or #84 #80)
-#86 := [th-lemma]: #85
-#87 := [unit-resolution #86 #83]: #80
-#61 := (>= #52 0::int)
-#65 := (not #61)
-#11 := (- f3)
-#10 := (< f3 0::int)
-#12 := (ite #10 #11 f3)
-#13 := (<= 0::int #12)
-#14 := (not #13)
-#68 := (iff #14 #65)
-#36 := (ite #10 #33 f3)
-#39 := (<= 0::int #36)
-#42 := (not #39)
-#66 := (iff #42 #65)
-#63 := (iff #39 #61)
-#57 := (<= 0::int #52)
-#60 := (iff #57 #61)
-#62 := [rewrite]: #60
-#58 := (iff #39 #57)
-#55 := (= #36 #52)
-#49 := (ite #46 #33 f3)
-#53 := (= #49 #52)
-#54 := [rewrite]: #53
-#50 := (= #36 #49)
-#47 := (iff #10 #46)
-#48 := [rewrite]: #47
-#51 := [monotonicity #48]: #50
-#56 := [trans #51 #54]: #55
-#59 := [monotonicity #56]: #58
-#64 := [trans #59 #62]: #63
-#67 := [monotonicity #64]: #66
-#43 := (iff #14 #42)
-#40 := (iff #13 #39)
-#37 := (= #12 #36)
-#34 := (= #11 #33)
-#35 := [rewrite]: #34
-#38 := [monotonicity #35]: #37
-#41 := [monotonicity #38]: #40
-#44 := [monotonicity #41]: #43
-#69 := [trans #44 #67]: #68
-#31 := [asserted]: #14
-#70 := [mp #31 #69]: #65
-#88 := [th-lemma #82 #70 #87]: false
-#89 := [lemma #88]: #45
-#78 := (or #46 #76)
-#79 := [def-axiom]: #78
-#92 := [unit-resolution #79 #89]: #76
-#93 := (not #76)
-#94 := (or #93 #90)
-#95 := [th-lemma]: #94
-#96 := [unit-resolution #95 #92]: #90
-[th-lemma #89 #70 #96]: false
-unsat
-957272aa2f77e61687a16c7aa490ba2664908958 132 0
-#2 := false
-#9 := 0::int
-decl f3 :: int
-#8 := f3
-#66 := (>= f3 0::int)
-#67 := (not #66)
-#34 := -1::int
-#35 := (* -1::int f3)
-#73 := (ite #66 f3 #35)
-#102 := (= f3 #73)
-#130 := (not #102)
-#14 := (= f3 0::int)
-#82 := (= #73 0::int)
-#124 := (iff #82 #14)
-#122 := (iff #14 #82)
-#121 := [hypothesis]: #102
-#123 := [monotonicity #121]: #122
-#125 := [symm #123]: #124
-#131 := (or #82 #130)
-#60 := (not #14)
-#99 := (not #82)
-#126 := (iff #99 #60)
-#127 := [monotonicity #125]: #126
-#119 := [hypothesis]: #99
-#128 := [mp #119 #127]: #60
-#112 := (or #14 #82)
-#89 := (iff #60 #82)
-#11 := (- f3)
-#10 := (< f3 0::int)
-#12 := (ite #10 #11 f3)
-#13 := (= #12 0::int)
-#15 := (iff #13 #14)
-#16 := (not #15)
-#94 := (iff #16 #89)
-#38 := (ite #10 #35 f3)
-#44 := (= 0::int #38)
-#61 := (iff #44 #60)
-#92 := (iff #61 #89)
-#86 := (iff #82 #60)
-#90 := (iff #86 #89)
-#91 := [rewrite]: #90
-#87 := (iff #61 #86)
-#84 := (iff #44 #82)
-#78 := (= 0::int #73)
-#81 := (iff #78 #82)
-#83 := [rewrite]: #81
-#79 := (iff #44 #78)
-#76 := (= #38 #73)
-#70 := (ite #67 #35 f3)
-#74 := (= #70 #73)
-#75 := [rewrite]: #74
-#71 := (= #38 #70)
-#68 := (iff #10 #67)
-#69 := [rewrite]: #68
-#72 := [monotonicity #69]: #71
-#77 := [trans #72 #75]: #76
-#80 := [monotonicity #77]: #79
-#85 := [trans #80 #83]: #84
-#88 := [monotonicity #85]: #87
-#93 := [trans #88 #91]: #92
-#64 := (iff #16 #61)
-#52 := (iff #14 #44)
-#57 := (not #52)
-#62 := (iff #57 #61)
-#63 := [rewrite]: #62
-#58 := (iff #16 #57)
-#55 := (iff #15 #52)
-#49 := (iff #44 #14)
-#53 := (iff #49 #52)
-#54 := [rewrite]: #53
-#50 := (iff #15 #49)
-#47 := (iff #13 #44)
-#41 := (= #38 0::int)
-#45 := (iff #41 #44)
-#46 := [rewrite]: #45
-#42 := (iff #13 #41)
-#39 := (= #12 #38)
-#36 := (= #11 #35)
-#37 := [rewrite]: #36
-#40 := [monotonicity #37]: #39
-#43 := [monotonicity #40]: #42
-#48 := [trans #43 #46]: #47
-#51 := [monotonicity #48]: #50
-#56 := [trans #51 #54]: #55
-#59 := [monotonicity #56]: #58
-#65 := [trans #59 #63]: #64
-#95 := [trans #65 #93]: #94
-#33 := [asserted]: #16
-#96 := [mp #33 #95]: #89
-#97 := (not #89)
-#110 := (or #14 #82 #97)
-#111 := [def-axiom]: #110
-#113 := [unit-resolution #111 #96]: #112
-#120 := [unit-resolution #113 #119]: #14
-#129 := [unit-resolution #120 #128]: false
-#132 := [lemma #129]: #131
-#133 := [unit-resolution #132 #121]: #82
-#135 := [mp #133 #125]: #14
-#108 := (or #60 #99)
-#106 := (or #60 #99 #97)
-#107 := [def-axiom]: #106
-#109 := [unit-resolution #107 #96]: #108
-#134 := [unit-resolution #109 #133]: #60
-#136 := [unit-resolution #134 #135]: false
-#137 := [lemma #136]: #130
-#104 := (or #67 #102)
-#105 := [def-axiom]: #104
-#143 := [unit-resolution #105 #137]: #67
-#138 := (= #35 0::int)
-#147 := (not #138)
-#157 := (iff #147 #99)
-#155 := (iff #138 #82)
-#103 := (= #35 #73)
-#100 := (or #66 #103)
-#101 := [def-axiom]: #100
-#154 := [unit-resolution #101 #143]: #103
-#156 := [monotonicity #154]: #155
-#158 := [monotonicity #156]: #157
-#139 := (<= #35 0::int)
-#145 := (not #139)
-#142 := [hypothesis]: #139
-#144 := [th-lemma #143 #142]: false
-#146 := [lemma #144]: #145
-#148 := (or #147 #139)
-#149 := [th-lemma]: #148
-#153 := [unit-resolution #149 #146]: #147
-#159 := [mp #153 #158]: #99
-#160 := [unit-resolution #113 #159]: #14
-#161 := (or #60 #66)
-#162 := [th-lemma]: #161
-[unit-resolution #162 #160 #143]: false
-unsat
-9485d306079aa8cdc08952543a8ac86391d45fc0 103 0
-#2 := false
-#8 := 0::int
-decl f3 :: int
-#9 := f3
-#34 := -1::int
-#35 := (* -1::int f3)
-#112 := (* -1::int #35)
-#113 := (+ f3 #112)
-#115 := (>= #113 0::int)
-#111 := (= f3 #35)
-#61 := (>= f3 0::int)
-#68 := (ite #61 f3 #35)
-#118 := (= #68 #35)
-#96 := (= #35 #68)
-#62 := (not #61)
-#107 := [hypothesis]: #61
-#73 := (= f3 #68)
-#97 := (or #62 #73)
-#98 := [def-axiom]: #97
-#108 := [unit-resolution #98 #107]: #73
-#95 := (not #73)
-#101 := (or #62 #95)
-#83 := (iff #62 #73)
-#12 := (- f3)
-#11 := (< f3 0::int)
-#13 := (ite #11 #12 f3)
-#14 := (= #13 f3)
-#10 := (<= 0::int f3)
-#15 := (iff #10 #14)
-#16 := (not #15)
-#88 := (iff #16 #83)
-#55 := (not #10)
-#38 := (ite #11 #35 f3)
-#44 := (= f3 #38)
-#56 := (iff #44 #55)
-#86 := (iff #56 #83)
-#80 := (iff #73 #62)
-#84 := (iff #80 #83)
-#85 := [rewrite]: #84
-#81 := (iff #56 #80)
-#78 := (iff #55 #62)
-#76 := (iff #10 #61)
-#77 := [rewrite]: #76
-#79 := [monotonicity #77]: #78
-#74 := (iff #44 #73)
-#71 := (= #38 #68)
-#65 := (ite #62 #35 f3)
-#69 := (= #65 #68)
-#70 := [rewrite]: #69
-#66 := (= #38 #65)
-#63 := (iff #11 #62)
-#64 := [rewrite]: #63
-#67 := [monotonicity #64]: #66
-#72 := [trans #67 #70]: #71
-#75 := [monotonicity #72]: #74
-#82 := [monotonicity #75 #79]: #81
-#87 := [trans #82 #85]: #86
-#59 := (iff #16 #56)
-#49 := (iff #10 #44)
-#52 := (not #49)
-#57 := (iff #52 #56)
-#58 := [rewrite]: #57
-#53 := (iff #16 #52)
-#50 := (iff #15 #49)
-#47 := (iff #14 #44)
-#41 := (= #38 f3)
-#45 := (iff #41 #44)
-#46 := [rewrite]: #45
-#42 := (iff #14 #41)
-#39 := (= #13 #38)
-#36 := (= #12 #35)
-#37 := [rewrite]: #36
-#40 := [monotonicity #37]: #39
-#43 := [monotonicity #40]: #42
-#48 := [trans #43 #46]: #47
-#51 := [monotonicity #48]: #50
-#54 := [monotonicity #51]: #53
-#60 := [trans #54 #58]: #59
-#89 := [trans #60 #87]: #88
-#33 := [asserted]: #16
-#90 := [mp #33 #89]: #83
-#93 := (not #83)
-#91 := (or #62 #95 #93)
-#100 := [def-axiom]: #91
-#102 := [unit-resolution #100 #90]: #101
-#109 := [unit-resolution #102 #108 #107]: false
-#110 := [lemma #109]: #62
-#99 := (or #61 #96)
-#94 := [def-axiom]: #99
-#116 := [unit-resolution #94 #110]: #96
-#119 := [symm #116]: #118
-#105 := (or #61 #73)
-#103 := (or #61 #73 #93)
-#104 := [def-axiom]: #103
-#106 := [unit-resolution #104 #90]: #105
-#117 := [unit-resolution #106 #110]: #73
-#120 := [trans #117 #119]: #111
-#121 := (not #111)
-#122 := (or #121 #115)
-#123 := [th-lemma]: #122
-#124 := [unit-resolution #123 #120]: #115
-[th-lemma #110 #124]: false
-unsat
-443d726f158e992756000aa2462469d5494d62ee 149 0
-#2 := false
-#9 := 0::int
-decl f3 :: int
-#8 := f3
-#34 := -1::int
-#35 := (* -1::int f3)
-#61 := (>= f3 0::int)
-#68 := (ite #61 f3 #35)
-#111 := (* -1::int #68)
-#156 := (+ f3 #111)
-#157 := (<= #156 0::int)
-#98 := (= f3 #68)
-#10 := (<= f3 0::int)
-#55 := (not #10)
-#78 := (+ f3 #68)
-#77 := (= #78 0::int)
-#93 := (not #77)
-#115 := [hypothesis]: #93
-#95 := (>= #78 0::int)
-#112 := (+ #35 #111)
-#113 := (<= #112 0::int)
-#73 := (= #35 #68)
-#62 := (not #61)
-#131 := (or #62 #77)
-#118 := (= f3 0::int)
-#109 := (or #10 #77)
-#85 := (iff #55 #77)
-#12 := (- f3)
-#11 := (< f3 0::int)
-#13 := (ite #11 #12 f3)
-#14 := (= #13 #12)
-#15 := (iff #10 #14)
-#16 := (not #15)
-#90 := (iff #16 #85)
-#38 := (ite #11 #35 f3)
-#44 := (= #35 #38)
-#56 := (iff #44 #55)
-#88 := (iff #56 #85)
-#82 := (iff #77 #55)
-#86 := (iff #82 #85)
-#87 := [rewrite]: #86
-#83 := (iff #56 #82)
-#80 := (iff #44 #77)
-#76 := (iff #73 #77)
-#79 := [rewrite]: #76
-#74 := (iff #44 #73)
-#71 := (= #38 #68)
-#65 := (ite #62 #35 f3)
-#69 := (= #65 #68)
-#70 := [rewrite]: #69
-#66 := (= #38 #65)
-#63 := (iff #11 #62)
-#64 := [rewrite]: #63
-#67 := [monotonicity #64]: #66
-#72 := [trans #67 #70]: #71
-#75 := [monotonicity #72]: #74
-#81 := [trans #75 #79]: #80
-#84 := [monotonicity #81]: #83
-#89 := [trans #84 #87]: #88
-#59 := (iff #16 #56)
-#49 := (iff #10 #44)
-#52 := (not #49)
-#57 := (iff #52 #56)
-#58 := [rewrite]: #57
-#53 := (iff #16 #52)
-#50 := (iff #15 #49)
-#47 := (iff #14 #44)
-#41 := (= #38 #35)
-#45 := (iff #41 #44)
-#46 := [rewrite]: #45
-#42 := (iff #14 #41)
-#36 := (= #12 #35)
-#37 := [rewrite]: #36
-#39 := (= #13 #38)
-#40 := [monotonicity #37]: #39
-#43 := [monotonicity #40 #37]: #42
-#48 := [trans #43 #46]: #47
-#51 := [monotonicity #48]: #50
-#54 := [monotonicity #51]: #53
-#60 := [trans #54 #58]: #59
-#91 := [trans #60 #89]: #90
-#33 := [asserted]: #16
-#92 := [mp #33 #91]: #85
-#102 := (not #85)
-#107 := (or #10 #77 #102)
-#108 := [def-axiom]: #107
-#110 := [unit-resolution #108 #92]: #109
-#116 := [unit-resolution #110 #115]: #10
-#117 := [hypothesis]: #61
-#119 := [th-lemma #117 #116]: #118
-#127 := (= #78 f3)
-#125 := (= #68 f3)
-#99 := (or #62 #98)
-#100 := [def-axiom]: #99
-#120 := [unit-resolution #100 #117]: #98
-#126 := [symm #120]: #125
-#123 := (= #78 #68)
-#121 := (= #68 #78)
-#122 := [th-lemma #117 #116]: #121
-#124 := [symm #122]: #123
-#128 := [trans #124 #126]: #127
-#129 := [trans #128 #119]: #77
-#130 := [unit-resolution #115 #129]: false
-#132 := [lemma #130]: #131
-#133 := [unit-resolution #132 #115]: #62
-#101 := (or #61 #73)
-#96 := [def-axiom]: #101
-#134 := [unit-resolution #96 #133]: #73
-#135 := (not #73)
-#136 := (or #135 #113)
-#137 := [th-lemma]: #136
-#138 := [unit-resolution #137 #134]: #113
-#139 := (not #113)
-#140 := (or #95 #139)
-#141 := [th-lemma]: #140
-#142 := [unit-resolution #141 #138]: #95
-#97 := (<= #78 0::int)
-#114 := (>= #112 0::int)
-#143 := (or #135 #114)
-#144 := [th-lemma]: #143
-#145 := [unit-resolution #144 #134]: #114
-#146 := (not #114)
-#147 := (or #97 #146)
-#148 := [th-lemma]: #147
-#149 := [unit-resolution #148 #145]: #97
-#151 := (not #95)
-#150 := (not #97)
-#152 := (or #77 #150 #151)
-#153 := [th-lemma]: #152
-#154 := [unit-resolution #153 #149 #142 #115]: false
-#155 := [lemma #154]: #77
-#105 := (or #55 #93)
-#103 := (or #55 #93 #102)
-#104 := [def-axiom]: #103
-#106 := [unit-resolution #104 #92]: #105
-#159 := [unit-resolution #106 #155]: #55
-#160 := (or #61 #10)
-#161 := [th-lemma]: #160
-#162 := [unit-resolution #161 #159]: #61
-#163 := [unit-resolution #100 #162]: #98
-#164 := (not #98)
-#165 := (or #164 #157)
-#166 := [th-lemma]: #165
-#167 := [unit-resolution #166 #163]: #157
-#168 := (or #93 #97)
-#169 := [th-lemma]: #168
-#170 := [unit-resolution #169 #155]: #97
-[th-lemma #159 #170 #167]: false
-unsat
-5d75c12dca04a93af891e80a12133cb48b0e8ce3 114 0
-#2 := false
-#9 := 0::int
-decl f3 :: int
-#8 := f3
-#35 := -1::int
-#36 := (* -1::int f3)
-#67 := (>= f3 0::int)
-#74 := (ite #67 f3 #36)
-#88 := (* -1::int #74)
-#127 := (+ #36 #88)
-#137 := (<= #127 0::int)
-#114 := (= #36 #74)
-#68 := (not #67)
-#125 := (+ f3 #88)
-#126 := (<= #125 0::int)
-#113 := (= f3 #74)
-#128 := [hypothesis]: #67
-#115 := (or #68 #113)
-#116 := [def-axiom]: #115
-#129 := [unit-resolution #116 #128]: #113
-#130 := (not #113)
-#131 := (or #130 #126)
-#132 := [th-lemma]: #131
-#133 := [unit-resolution #132 #129]: #126
-#83 := (>= #74 0::int)
-#82 := (not #83)
-#94 := (ite #83 #74 #88)
-#99 := (= #74 #94)
-#102 := (not #99)
-#11 := (- f3)
-#10 := (< f3 0::int)
-#12 := (ite #10 #11 f3)
-#14 := (- #12)
-#13 := (< #12 0::int)
-#15 := (ite #13 #14 #12)
-#16 := (= #15 #12)
-#17 := (not #16)
-#105 := (iff #17 #102)
-#39 := (ite #10 #36 f3)
-#48 := (* -1::int #39)
-#42 := (< #39 0::int)
-#53 := (ite #42 #48 #39)
-#59 := (= #39 #53)
-#64 := (not #59)
-#103 := (iff #64 #102)
-#100 := (iff #59 #99)
-#97 := (= #53 #94)
-#91 := (ite #82 #88 #74)
-#95 := (= #91 #94)
-#96 := [rewrite]: #95
-#92 := (= #53 #91)
-#77 := (= #39 #74)
-#71 := (ite #68 #36 f3)
-#75 := (= #71 #74)
-#76 := [rewrite]: #75
-#72 := (= #39 #71)
-#69 := (iff #10 #68)
-#70 := [rewrite]: #69
-#73 := [monotonicity #70]: #72
-#78 := [trans #73 #76]: #77
-#89 := (= #48 #88)
-#90 := [monotonicity #78]: #89
-#86 := (iff #42 #82)
-#79 := (< #74 0::int)
-#84 := (iff #79 #82)
-#85 := [rewrite]: #84
-#80 := (iff #42 #79)
-#81 := [monotonicity #78]: #80
-#87 := [trans #81 #85]: #86
-#93 := [monotonicity #87 #90 #78]: #92
-#98 := [trans #93 #96]: #97
-#101 := [monotonicity #78 #98]: #100
-#104 := [monotonicity #101]: #103
-#65 := (iff #17 #64)
-#62 := (iff #16 #59)
-#56 := (= #53 #39)
-#60 := (iff #56 #59)
-#61 := [rewrite]: #60
-#57 := (iff #16 #56)
-#40 := (= #12 #39)
-#37 := (= #11 #36)
-#38 := [rewrite]: #37
-#41 := [monotonicity #38]: #40
-#54 := (= #15 #53)
-#51 := (= #14 #48)
-#45 := (- #39)
-#49 := (= #45 #48)
-#50 := [rewrite]: #49
-#46 := (= #14 #45)
-#47 := [monotonicity #41]: #46
-#52 := [trans #47 #50]: #51
-#43 := (iff #13 #42)
-#44 := [monotonicity #41]: #43
-#55 := [monotonicity #44 #52 #41]: #54
-#58 := [monotonicity #55 #41]: #57
-#63 := [trans #58 #61]: #62
-#66 := [monotonicity #63]: #65
-#106 := [trans #66 #104]: #105
-#34 := [asserted]: #17
-#107 := [mp #34 #106]: #102
-#108 := (or #82 #99)
-#117 := [def-axiom]: #108
-#134 := [unit-resolution #117 #107]: #82
-#135 := [th-lemma #128 #134 #133]: false
-#136 := [lemma #135]: #68
-#111 := (or #67 #114)
-#112 := [def-axiom]: #111
-#139 := [unit-resolution #112 #136]: #114
-#140 := (not #114)
-#141 := (or #140 #137)
-#142 := [th-lemma]: #141
-#143 := [unit-resolution #142 #139]: #137
-[th-lemma #136 #134 #143]: false
-unsat
-1c29e683939b29daf6c0c39d6dc5211f90494d2b 57 0
-#2 := false
-#36 := 0::int
-decl f4 :: int
-#9 := f4
-decl f3 :: int
-#8 := f3
-#33 := -1::int
-#34 := (* -1::int f4)
-#35 := (+ f3 #34)
-#37 := (<= #35 0::int)
-#40 := (ite #37 f3 f4)
-#48 := (* -1::int #40)
-#49 := (+ f3 #48)
-#47 := (>= #49 0::int)
-#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::int)
-#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
 f5bdd04ade0ff62ddedb44ca71a4ab32cea43ed1 57 0
 #2 := false
 #36 := 0::int
@@ -45153,6 +44348,352 @@
 #145 := [unit-resolution #117 #144]: #106
 [th-lemma #131 #143 #141 #85 #145]: 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
 de67b37b9c3eba477099ff5d9b09e840de376762 57 0
 #2 := false
 #36 := 0::int
@@ -45462,6 +45003,63 @@
 #109 := [trans #108 #83]: #51
 [unit-resolution #57 #109]: false
 unsat
+efdfb7635dd41e99f117d8463d9d75f3897207dd 56 0
+#2 := false
+#11 := 1::int
+decl f3 :: int
+#9 := f3
+#14 := (= f3 1::int)
+#12 := (<= f3 1::int)
+#8 := 0::int
+#41 := (<= f3 0::int)
+#42 := (not #41)
+#48 := (and #12 #42)
+#53 := (not #48)
+#59 := (or #14 #53)
+#64 := (not #59)
+#10 := (< 0::int f3)
+#13 := (and #10 #12)
+#15 := (implies #13 #14)
+#16 := (not #15)
+#67 := (iff #16 #64)
+#34 := (not #13)
+#35 := (or #34 #14)
+#38 := (not #35)
+#65 := (iff #38 #64)
+#62 := (iff #35 #59)
+#56 := (or #53 #14)
+#60 := (iff #56 #59)
+#61 := [rewrite]: #60
+#57 := (iff #35 #56)
+#54 := (iff #34 #53)
+#51 := (iff #13 #48)
+#45 := (and #42 #12)
+#49 := (iff #45 #48)
+#50 := [rewrite]: #49
+#46 := (iff #13 #45)
+#43 := (iff #10 #42)
+#44 := [rewrite]: #43
+#47 := [monotonicity #44]: #46
+#52 := [trans #47 #50]: #51
+#55 := [monotonicity #52]: #54
+#58 := [monotonicity #55]: #57
+#63 := [trans #58 #61]: #62
+#66 := [monotonicity #63]: #65
+#39 := (iff #16 #38)
+#36 := (iff #15 #35)
+#37 := [rewrite]: #36
+#40 := [monotonicity #37]: #39
+#68 := [trans #40 #66]: #67
+#33 := [asserted]: #16
+#69 := [mp #33 #68]: #64
+#72 := [not-or-elim #69]: #48
+#73 := [and-elim #72]: #12
+#74 := [and-elim #72]: #42
+#85 := [th-lemma #74 #73]: #14
+#70 := (not #14)
+#71 := [not-or-elim #69]: #70
+[unit-resolution #71 #85]: false
+unsat
 66ca5933e35ee3faa0336268f65c331335330f57 66 0
 #2 := false
 decl f3 :: int
@@ -45529,6 +45127,25 @@
 #86 := [trans #83 #85]: #33
 [unit-resolution #69 #86]: false
 unsat
+6a196ac012aba8c7b41b1d8437b14fc3fb23e5a4 18 0
+#2 := false
+decl f3 :: int
+#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
 a176c257f1c89ffea971e08639242e005fb1222f 228 0
 #2 := false
 #9 := 0::int
@@ -45758,82 +45375,6 @@
 #250 := [unit-resolution #196 #249]: #190
 [th-lemma #236 #250 #131 #234 #248 #238]: false
 unsat
-efdfb7635dd41e99f117d8463d9d75f3897207dd 56 0
-#2 := false
-#11 := 1::int
-decl f3 :: int
-#9 := f3
-#14 := (= f3 1::int)
-#12 := (<= f3 1::int)
-#8 := 0::int
-#41 := (<= f3 0::int)
-#42 := (not #41)
-#48 := (and #12 #42)
-#53 := (not #48)
-#59 := (or #14 #53)
-#64 := (not #59)
-#10 := (< 0::int f3)
-#13 := (and #10 #12)
-#15 := (implies #13 #14)
-#16 := (not #15)
-#67 := (iff #16 #64)
-#34 := (not #13)
-#35 := (or #34 #14)
-#38 := (not #35)
-#65 := (iff #38 #64)
-#62 := (iff #35 #59)
-#56 := (or #53 #14)
-#60 := (iff #56 #59)
-#61 := [rewrite]: #60
-#57 := (iff #35 #56)
-#54 := (iff #34 #53)
-#51 := (iff #13 #48)
-#45 := (and #42 #12)
-#49 := (iff #45 #48)
-#50 := [rewrite]: #49
-#46 := (iff #13 #45)
-#43 := (iff #10 #42)
-#44 := [rewrite]: #43
-#47 := [monotonicity #44]: #46
-#52 := [trans #47 #50]: #51
-#55 := [monotonicity #52]: #54
-#58 := [monotonicity #55]: #57
-#63 := [trans #58 #61]: #62
-#66 := [monotonicity #63]: #65
-#39 := (iff #16 #38)
-#36 := (iff #15 #35)
-#37 := [rewrite]: #36
-#40 := [monotonicity #37]: #39
-#68 := [trans #40 #66]: #67
-#33 := [asserted]: #16
-#69 := [mp #33 #68]: #64
-#72 := [not-or-elim #69]: #48
-#73 := [and-elim #72]: #12
-#74 := [and-elim #72]: #42
-#85 := [th-lemma #74 #73]: #14
-#70 := (not #14)
-#71 := [not-or-elim #69]: #70
-[unit-resolution #71 #85]: false
-unsat
-6a196ac012aba8c7b41b1d8437b14fc3fb23e5a4 18 0
-#2 := false
-decl f3 :: int
-#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
 4afe860c901fd901bf34dc83b799fd986f9e4155 51 0
 #2 := false
 decl f4 :: int
@@ -46070,6 +45611,24 @@
 #78 := [unit-resolution #82 #74]: #83
 [unit-resolution #78 #92 #89]: false
 unsat
+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
 65c77124096528f9cb7c1cb6b89625b62f5eeb2a 15 0
 #2 := false
 decl f3 :: int
@@ -46086,6 +45645,32 @@
 #28 := [asserted]: #11
 [mp #28 #34]: 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
 8aa9d7db6e905a29f4a4232a6a82f51c63e2641b 75 0
 #2 := false
 #53 := 0::int
@@ -46162,6 +45747,24 @@
 #89 := [not-or-elim #88]: #54
 [th-lemma #89 #91 #92]: 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
 625ca444862ba21a1fa0a52470751952d4f5acee 80 0
 #2 := false
 #53 := 0::int
@@ -46243,6 +45846,60 @@
 #94 := [not-or-elim #93]: #54
 [th-lemma #94 #96 #97]: 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
 59582f385e59c3a49a302bb1b9701c92bb368fac 80 0
 #2 := false
 #52 := 0::int
@@ -46324,6 +45981,25 @@
 #96 := [not-or-elim #93]: #95
 [th-lemma #96 #94 #97]: 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
 6dfe89c45e68ba796109229b1c54d23603c64ded 85 0
 #2 := false
 #52 := 0::int
@@ -46410,6 +46086,47 @@
 #101 := [not-or-elim #99]: #66
 [th-lemma #101 #100 #102]: 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
 650b9ceb8c2cf801e6ceee5375af886f013abe2c 69 0
 #2 := false
 #44 := 0::int
@@ -46480,182 +46197,6 @@
 #85 := [and-elim #83]: #53
 [th-lemma #85 #84 #86]: false
 unsat
-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
@@ -46760,116 +46301,6 @@
 #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
-decl f4 :: real
-#9 := f4
-decl f3 :: real
-#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
-afbfb539cd215cee99e6f9076de9621f41db722f 45 0
-#2 := false
-decl f4 :: real
-#9 := f4
-#13 := (- f4)
-decl f3 :: real
-#8 := f3
-#14 := (= f3 #13)
-#11 := 0::real
-#10 := (+ f3 f4)
-#12 := (= #10 0::real)
-#15 := (iff #12 #14)
-#16 := (not #15)
-#62 := (iff #16 false)
-#47 := (not #12)
-#34 := -1::real
-#35 := (* -1::real 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
 4a497c9573e0d4cf02c29e65ea385ac694a5d2a9 58 0
 #2 := false
 decl f3 :: real
@@ -46929,6 +46360,36 @@
 #31 := [asserted]: #14
 [mp #31 #77]: 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
 5310d952762e8bd0d9db1c0127d9100532f42cf4 61 0
 #2 := false
 #8 := 0::real
@@ -46991,6 +46452,40 @@
 #31 := [asserted]: #14
 [mp #31 #79]: false
 unsat
+08018537eda07e4214b99470b7aa75706e7cd28c 33 0
+#2 := false
+decl f5 :: real
+#10 := f5
+decl f4 :: real
+#9 := f4
+decl f3 :: real
+#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
 aed50da01a066fa4f6c8e7af4a461d101bc09de0 61 0
 #2 := false
 decl f3 :: real
@@ -47053,6 +46548,52 @@
 #31 := [asserted]: #14
 [mp #31 #79]: false
 unsat
+afbfb539cd215cee99e6f9076de9621f41db722f 45 0
+#2 := false
+decl f4 :: real
+#9 := f4
+#13 := (- f4)
+decl f3 :: real
+#8 := f3
+#14 := (= f3 #13)
+#11 := 0::real
+#10 := (+ f3 f4)
+#12 := (= #10 0::real)
+#15 := (iff #12 #14)
+#16 := (not #15)
+#62 := (iff #16 false)
+#47 := (not #12)
+#34 := -1::real
+#35 := (* -1::real 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
 4c33a6158d9a09229cc4bfd8f383004f3cebc48a 26 0
 #2 := false
 decl f3 :: real
@@ -47726,132 +47267,6 @@
 #30 := [asserted]: #13
 [mp #30 #53]: false
 unsat
-8140105cc276b32acd77e4f2cd380fea1f24efd8 62 0
-#2 := false
-#8 := 0::real
-decl f3 :: real
-#9 := f3
-#67 := 1/3::real
-#68 := (* 1/3::real f3)
-#69 := (<= #68 0::real)
-#65 := (not #69)
-#54 := (<= f3 0::real)
-#71 := (or #54 #65)
-#74 := (not #71)
-#13 := 3::real
-#11 := 2::real
-#12 := (* 2::real f3)
-#14 := (/ #12 3::real)
-#15 := (< #14 f3)
-#10 := (< 0::real f3)
-#16 := (implies #10 #15)
-#17 := (not #16)
-#77 := (iff #17 #74)
-#36 := 2/3::real
-#37 := (* 2/3::real f3)
-#39 := (< #37 f3)
-#45 := (not #10)
-#46 := (or #45 #39)
-#51 := (not #46)
-#75 := (iff #51 #74)
-#72 := (iff #46 #71)
-#66 := (iff #39 #65)
-#70 := [rewrite]: #66
-#63 := (iff #45 #54)
-#55 := (not #54)
-#58 := (not #55)
-#61 := (iff #58 #54)
-#62 := [rewrite]: #61
-#59 := (iff #45 #58)
-#56 := (iff #10 #55)
-#57 := [rewrite]: #56
-#60 := [monotonicity #57]: #59
-#64 := [trans #60 #62]: #63
-#73 := [monotonicity #64 #70]: #72
-#76 := [monotonicity #73]: #75
-#52 := (iff #17 #51)
-#49 := (iff #16 #46)
-#42 := (implies #10 #39)
-#47 := (iff #42 #46)
-#48 := [rewrite]: #47
-#43 := (iff #16 #42)
-#40 := (iff #15 #39)
-#35 := (= #14 #37)
-#38 := [rewrite]: #35
-#41 := [monotonicity #38]: #40
-#44 := [monotonicity #41]: #43
-#50 := [trans #44 #48]: #49
-#53 := [monotonicity #50]: #52
-#78 := [trans #53 #76]: #77
-#34 := [asserted]: #17
-#79 := [mp #34 #78]: #74
-#81 := [not-or-elim #79]: #69
-#80 := [not-or-elim #79]: #55
-[th-lemma #80 #81]: false
-unsat
-9a571f31a85693be816acc5c7b7aa7362cb820bb 62 0
-#2 := false
-#9 := 0::real
-decl f3 :: real
-#8 := f3
-#69 := 1/3::real
-#70 := (* 1/3::real f3)
-#67 := (>= #70 0::real)
-#65 := (not #67)
-#56 := (>= f3 0::real)
-#71 := (or #56 #65)
-#74 := (not #71)
-#13 := 3::real
-#11 := 2::real
-#12 := (* 2::real f3)
-#14 := (/ #12 3::real)
-#15 := (< f3 #14)
-#10 := (< f3 0::real)
-#16 := (implies #10 #15)
-#17 := (not #16)
-#77 := (iff #17 #74)
-#36 := 2/3::real
-#37 := (* 2/3::real f3)
-#39 := (< f3 #37)
-#45 := (not #10)
-#46 := (or #45 #39)
-#51 := (not #46)
-#75 := (iff #51 #74)
-#72 := (iff #46 #71)
-#66 := (iff #39 #65)
-#68 := [rewrite]: #66
-#63 := (iff #45 #56)
-#54 := (not #56)
-#58 := (not #54)
-#61 := (iff #58 #56)
-#62 := [rewrite]: #61
-#59 := (iff #45 #58)
-#55 := (iff #10 #54)
-#57 := [rewrite]: #55
-#60 := [monotonicity #57]: #59
-#64 := [trans #60 #62]: #63
-#73 := [monotonicity #64 #68]: #72
-#76 := [monotonicity #73]: #75
-#52 := (iff #17 #51)
-#49 := (iff #16 #46)
-#42 := (implies #10 #39)
-#47 := (iff #42 #46)
-#48 := [rewrite]: #47
-#43 := (iff #16 #42)
-#40 := (iff #15 #39)
-#35 := (= #14 #37)
-#38 := [rewrite]: #35
-#41 := [monotonicity #38]: #40
-#44 := [monotonicity #41]: #43
-#50 := [trans #44 #48]: #49
-#53 := [monotonicity #50]: #52
-#78 := [trans #53 #76]: #77
-#34 := [asserted]: #17
-#79 := [mp #34 #78]: #74
-#81 := [not-or-elim #79]: #67
-#80 := [not-or-elim #79]: #54
-[th-lemma #80 #81]: false
-unsat
 2f685a966903afb267e7499dd8ad6b2b43caca93 75 0
 #2 := false
 #8 := 0::real
@@ -47928,6 +47343,69 @@
 #96 := [unit-resolution #95 #92]: #90
 [th-lemma #89 #70 #96]: false
 unsat
+8140105cc276b32acd77e4f2cd380fea1f24efd8 62 0
+#2 := false
+#8 := 0::real
+decl f3 :: real
+#9 := f3
+#67 := 1/3::real
+#68 := (* 1/3::real f3)
+#69 := (<= #68 0::real)
+#65 := (not #69)
+#54 := (<= f3 0::real)
+#71 := (or #54 #65)
+#74 := (not #71)
+#13 := 3::real
+#11 := 2::real
+#12 := (* 2::real f3)
+#14 := (/ #12 3::real)
+#15 := (< #14 f3)
+#10 := (< 0::real f3)
+#16 := (implies #10 #15)
+#17 := (not #16)
+#77 := (iff #17 #74)
+#36 := 2/3::real
+#37 := (* 2/3::real f3)
+#39 := (< #37 f3)
+#45 := (not #10)
+#46 := (or #45 #39)
+#51 := (not #46)
+#75 := (iff #51 #74)
+#72 := (iff #46 #71)
+#66 := (iff #39 #65)
+#70 := [rewrite]: #66
+#63 := (iff #45 #54)
+#55 := (not #54)
+#58 := (not #55)
+#61 := (iff #58 #54)
+#62 := [rewrite]: #61
+#59 := (iff #45 #58)
+#56 := (iff #10 #55)
+#57 := [rewrite]: #56
+#60 := [monotonicity #57]: #59
+#64 := [trans #60 #62]: #63
+#73 := [monotonicity #64 #70]: #72
+#76 := [monotonicity #73]: #75
+#52 := (iff #17 #51)
+#49 := (iff #16 #46)
+#42 := (implies #10 #39)
+#47 := (iff #42 #46)
+#48 := [rewrite]: #47
+#43 := (iff #16 #42)
+#40 := (iff #15 #39)
+#35 := (= #14 #37)
+#38 := [rewrite]: #35
+#41 := [monotonicity #38]: #40
+#44 := [monotonicity #41]: #43
+#50 := [trans #44 #48]: #49
+#53 := [monotonicity #50]: #52
+#78 := [trans #53 #76]: #77
+#34 := [asserted]: #17
+#79 := [mp #34 #78]: #74
+#81 := [not-or-elim #79]: #69
+#80 := [not-or-elim #79]: #55
+[th-lemma #80 #81]: false
+unsat
 4a38ab99c62d0654d0c5bc149db3e769cf921224 132 0
 #2 := false
 #9 := 0::real
@@ -48061,6 +47539,69 @@
 #162 := [th-lemma]: #161
 [unit-resolution #162 #160 #143]: false
 unsat
+9a571f31a85693be816acc5c7b7aa7362cb820bb 62 0
+#2 := false
+#9 := 0::real
+decl f3 :: real
+#8 := f3
+#69 := 1/3::real
+#70 := (* 1/3::real f3)
+#67 := (>= #70 0::real)
+#65 := (not #67)
+#56 := (>= f3 0::real)
+#71 := (or #56 #65)
+#74 := (not #71)
+#13 := 3::real
+#11 := 2::real
+#12 := (* 2::real f3)
+#14 := (/ #12 3::real)
+#15 := (< f3 #14)
+#10 := (< f3 0::real)
+#16 := (implies #10 #15)
+#17 := (not #16)
+#77 := (iff #17 #74)
+#36 := 2/3::real
+#37 := (* 2/3::real f3)
+#39 := (< f3 #37)
+#45 := (not #10)
+#46 := (or #45 #39)
+#51 := (not #46)
+#75 := (iff #51 #74)
+#72 := (iff #46 #71)
+#66 := (iff #39 #65)
+#68 := [rewrite]: #66
+#63 := (iff #45 #56)
+#54 := (not #56)
+#58 := (not #54)
+#61 := (iff #58 #56)
+#62 := [rewrite]: #61
+#59 := (iff #45 #58)
+#55 := (iff #10 #54)
+#57 := [rewrite]: #55
+#60 := [monotonicity #57]: #59
+#64 := [trans #60 #62]: #63
+#73 := [monotonicity #64 #68]: #72
+#76 := [monotonicity #73]: #75
+#52 := (iff #17 #51)
+#49 := (iff #16 #46)
+#42 := (implies #10 #39)
+#47 := (iff #42 #46)
+#48 := [rewrite]: #47
+#43 := (iff #16 #42)
+#40 := (iff #15 #39)
+#35 := (= #14 #37)
+#38 := [rewrite]: #35
+#41 := [monotonicity #38]: #40
+#44 := [monotonicity #41]: #43
+#50 := [trans #44 #48]: #49
+#53 := [monotonicity #50]: #52
+#78 := [trans #53 #76]: #77
+#34 := [asserted]: #17
+#79 := [mp #34 #78]: #74
+#81 := [not-or-elim #79]: #67
+#80 := [not-or-elim #79]: #54
+[th-lemma #80 #81]: false
+unsat
 704113f68324742364e301291c46543b32d30adc 103 0
 #2 := false
 #8 := 0::real
@@ -48165,6 +47706,64 @@
 #124 := [unit-resolution #123 #120]: #115
 [th-lemma #110 #124]: false
 unsat
+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
 099809f4b951e6c4b8dd87d068c12977d185ca9a 149 0
 #2 := false
 #9 := 0::real
@@ -48315,6 +47914,64 @@
 #170 := [unit-resolution #169 #155]: #97
 [th-lemma #159 #170 #167]: 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
 29b8934bcf2c51c8a551e0b17bf814bf8deedd61 114 0
 #2 := false
 #9 := 0::real
@@ -48430,122 +48087,6 @@
 #143 := [unit-resolution #142 #139]: #137
 [th-lemma #136 #134 #143]: false
 unsat
-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
@@ -48650,6 +48191,64 @@
 #122 := [unit-resolution #121 #118]: #115
 [th-lemma #122 #92 #93]: false
 unsat
+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
 23d9458c2d94638d741588b7f5f53da3b743fc4b 86 0
 #2 := false
 decl f3 :: real
@@ -48737,6 +48336,64 @@
 #105 := [trans #104 #83]: #51
 [unit-resolution #57 #105]: 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
 ca3670773dec1aa185941c6109c12ef74c091595 68 0
 #2 := false
 #8 := 0::real
@@ -48806,6 +48463,110 @@
 #88 := [trans #87 #84]: #13
 [unit-resolution #71 #88]: 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
 a5005d5e2177d7e219c0762e500aaee9c6b18ed1 124 0
 #2 := false
 #13 := 0::real
@@ -48931,226 +48692,6 @@
 #145 := [unit-resolution #117 #144]: #106
 [th-lemma #131 #143 #141 #85 #145]: false
 unsat
-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
@@ -49307,6 +48848,25 @@
 #86 := [trans #83 #85]: #33
 [unit-resolution #69 #86]: false
 unsat
+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
 3ea0cce22deeb1cf0a91c70b01f38d938cbfca2b 228 0
 #2 := false
 #9 := 0::real
@@ -49536,25 +49096,6 @@
 #250 := [unit-resolution #196 #249]: #190
 [th-lemma #236 #250 #131 #234 #248 #238]: false
 unsat
-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
@@ -49715,6 +49256,50 @@
 #77 := [unit-resolution #76 #65]: #53
 [unit-resolution #77 #67]: false
 unsat
+7db3ee0924a4cb1b14a4cebbf1c598c0b8dbec57 43 0
+#2 := false
+decl f5 :: (-> S2 S2 S3)
+#15 := (:var 0 S2)
+#14 := (:var 1 S2)
+#16 := (f5 #14 #15)
+#561 := (pattern #16)
+decl f4 :: (-> S3 S2)
+#17 := (f4 #16)
+#47 := (= #14 #17)
+#562 := (forall (vars (?v0 S2) (?v1 S2)) (:pat #561) #47)
+#50 := (forall (vars (?v0 S2) (?v1 S2)) #47)
+#565 := (iff #50 #562)
+#563 := (iff #47 #47)
+#564 := [refl]: #563
+#566 := [quant-intro #564]: #565
+#83 := (~ #50 #50)
+#81 := (~ #47 #47)
+#82 := [refl]: #81
+#84 := [nnf-pos #82]: #83
+#18 := (= #17 #14)
+#19 := (forall (vars (?v0 S2) (?v1 S2)) #18)
+#51 := (iff #19 #50)
+#48 := (iff #18 #47)
+#49 := [rewrite]: #48
+#52 := [quant-intro #49]: #51
+#46 := [asserted]: #19
+#55 := [mp #46 #52]: #50
+#73 := [mp~ #55 #84]: #50
+#567 := [mp #73 #566]: #562
+decl f6 :: S2
+#9 := f6
+decl f3 :: S2
+#8 := f3
+#10 := (f5 f3 f6)
+#11 := (f4 #10)
+#12 := (= f3 #11)
+#13 := (not #12)
+#45 := [asserted]: #13
+#144 := (not #562)
+#231 := (or #144 #12)
+#145 := [quant-inst]: #231
+[unit-resolution #145 #45 #567]: false
+unsat
 5631a67f2c650efd9761a68c96a2b28d9ed4cdd0 15 0
 #2 := false
 decl f3 :: real
@@ -49731,6 +49316,50 @@
 #28 := [asserted]: #11
 [mp #28 #34]: false
 unsat
+8502f9cde674da83d9cbc8a977644ab6e85265ed 43 0
+#2 := false
+decl f5 :: (-> S2 S2 S3)
+#15 := (:var 0 S2)
+#14 := (:var 1 S2)
+#16 := (f5 #14 #15)
+#561 := (pattern #16)
+decl f4 :: (-> S3 S2)
+#20 := (f4 #16)
+#54 := (= #15 #20)
+#568 := (forall (vars (?v0 S2) (?v1 S2)) (:pat #561) #54)
+#58 := (forall (vars (?v0 S2) (?v1 S2)) #54)
+#571 := (iff #58 #568)
+#569 := (iff #54 #54)
+#570 := [refl]: #569
+#572 := [quant-intro #570]: #571
+#75 := (~ #58 #58)
+#74 := (~ #54 #54)
+#72 := [refl]: #74
+#76 := [nnf-pos #72]: #75
+#21 := (= #20 #15)
+#22 := (forall (vars (?v0 S2) (?v1 S2)) #21)
+#59 := (iff #22 #58)
+#56 := (iff #21 #54)
+#57 := [rewrite]: #56
+#60 := [quant-intro #57]: #59
+#53 := [asserted]: #22
+#63 := [mp #53 #60]: #58
+#69 := [mp~ #63 #76]: #58
+#573 := [mp #69 #572]: #568
+decl f3 :: S2
+#8 := f3
+decl f6 :: S2
+#9 := f6
+#10 := (f5 f6 f3)
+#11 := (f4 #10)
+#12 := (= f3 #11)
+#13 := (not #12)
+#45 := [asserted]: #13
+#234 := (not #568)
+#235 := (or #234 #12)
+#163 := [quant-inst]: #235
+[unit-resolution #163 #45 #573]: false
+unsat
 c41bf12f912d9026a2697f2c615b98bf21c74fe7 75 0
 #2 := false
 #53 := 0::real
@@ -49807,412 +49436,6 @@
 #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)
-#14 := (:var 1 S2)
-#16 := (f5 #14 #15)
-#561 := (pattern #16)
-decl f4 :: (-> S3 S2)
-#17 := (f4 #16)
-#47 := (= #14 #17)
-#562 := (forall (vars (?v0 S2) (?v1 S2)) (:pat #561) #47)
-#50 := (forall (vars (?v0 S2) (?v1 S2)) #47)
-#565 := (iff #50 #562)
-#563 := (iff #47 #47)
-#564 := [refl]: #563
-#566 := [quant-intro #564]: #565
-#83 := (~ #50 #50)
-#81 := (~ #47 #47)
-#82 := [refl]: #81
-#84 := [nnf-pos #82]: #83
-#18 := (= #17 #14)
-#19 := (forall (vars (?v0 S2) (?v1 S2)) #18)
-#51 := (iff #19 #50)
-#48 := (iff #18 #47)
-#49 := [rewrite]: #48
-#52 := [quant-intro #49]: #51
-#46 := [asserted]: #19
-#55 := [mp #46 #52]: #50
-#73 := [mp~ #55 #84]: #50
-#567 := [mp #73 #566]: #562
-decl f6 :: S2
-#9 := f6
-decl f3 :: S2
-#8 := f3
-#10 := (f5 f3 f6)
-#11 := (f4 #10)
-#12 := (= f3 #11)
-#13 := (not #12)
-#45 := [asserted]: #13
-#144 := (not #562)
-#231 := (or #144 #12)
-#145 := [quant-inst]: #231
-[unit-resolution #145 #45 #567]: false
-unsat
-8502f9cde674da83d9cbc8a977644ab6e85265ed 43 0
-#2 := false
-decl f5 :: (-> S2 S2 S3)
-#15 := (:var 0 S2)
-#14 := (:var 1 S2)
-#16 := (f5 #14 #15)
-#561 := (pattern #16)
-decl f4 :: (-> S3 S2)
-#20 := (f4 #16)
-#54 := (= #15 #20)
-#568 := (forall (vars (?v0 S2) (?v1 S2)) (:pat #561) #54)
-#58 := (forall (vars (?v0 S2) (?v1 S2)) #54)
-#571 := (iff #58 #568)
-#569 := (iff #54 #54)
-#570 := [refl]: #569
-#572 := [quant-intro #570]: #571
-#75 := (~ #58 #58)
-#74 := (~ #54 #54)
-#72 := [refl]: #74
-#76 := [nnf-pos #72]: #75
-#21 := (= #20 #15)
-#22 := (forall (vars (?v0 S2) (?v1 S2)) #21)
-#59 := (iff #22 #58)
-#56 := (iff #21 #54)
-#57 := [rewrite]: #56
-#60 := [quant-intro #57]: #59
-#53 := [asserted]: #22
-#63 := [mp #53 #60]: #58
-#69 := [mp~ #63 #76]: #58
-#573 := [mp #69 #572]: #568
-decl f3 :: S2
-#8 := f3
-decl f6 :: S2
-#9 := f6
-#10 := (f5 f6 f3)
-#11 := (f4 #10)
-#12 := (= f3 #11)
-#13 := (not #12)
-#45 := [asserted]: #13
-#234 := (not #568)
-#235 := (or #234 #12)
-#163 := [quant-inst]: #235
-[unit-resolution #163 #45 #573]: false
-unsat
 7028fd42bf273f5da876a633b9b2fac12fd44679 85 0
 #2 := false
 decl f5 :: S2
@@ -50299,6 +49522,87 @@
 #206 := [unit-resolution #170 #205]: #239
 [unit-resolution #206 #186]: 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
 ab91d7002afe61891f92aa68c03a93dbeac150bb 155 0
 #2 := false
 decl f7 :: S2
@@ -50455,6 +49759,87 @@
 #263 := [def-axiom]: #176
 [unit-resolution #263 #289 #215 #569]: 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
 c53da5ad16e14d4123258f05451b97645b05e47b 91 0
 #2 := false
 decl f9 :: S3
@@ -50547,6 +49932,92 @@
 #320 := [unit-resolution #230 #598]: #299
 [unit-resolution #320 #321]: 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
 4d6ae030a899d6a058af8429edba2402745a216b 210 0
 #2 := false
 decl f11 :: S5
@@ -50758,6 +50229,76 @@
 #321 := [def-axiom]: #234
 [unit-resolution #321 #597 #608 #332]: 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
 003bb45597da3e4b1800d61bab0b4e1d95dbdf99 144 0
 #2 := false
 decl f11 :: S3
@@ -50903,6 +50444,57 @@
 #432 := [unit-resolution #232 #431]: #301
 [unit-resolution #432 #433]: false
 unsat
+6cc8ed53fe294f253a561b590f79ef0e0271d1c2 50 0
+#2 := false
+decl f3 :: (-> S2 S3 S4 S3 S4)
+#15 := (:var 1 S3)
+#16 := (:var 0 S4)
+#14 := (:var 2 S2)
+#17 := (f3 #14 #15 #16 #15)
+#571 := (pattern #17)
+#54 := (= #16 #17)
+#572 := (forall (vars (?v0 S2) (?v1 S3) (?v2 S4)) (:pat #571) #54)
+#58 := (forall (vars (?v0 S2) (?v1 S3) (?v2 S4)) #54)
+#575 := (iff #58 #572)
+#573 := (iff #54 #54)
+#574 := [refl]: #573
+#576 := [quant-intro #574]: #575
+#87 := (~ #58 #58)
+#85 := (~ #54 #54)
+#86 := [refl]: #85
+#88 := [nnf-pos #86]: #87
+#18 := (= #17 #16)
+#19 := (forall (vars (?v0 S2) (?v1 S3) (?v2 S4)) #18)
+#59 := (iff #19 #58)
+#56 := (iff #18 #54)
+#57 := [rewrite]: #56
+#60 := [quant-intro #57]: #59
+#53 := [asserted]: #19
+#63 := [mp #53 #60]: #58
+#79 := [mp~ #63 #88]: #58
+#577 := [mp #79 #576]: #572
+decl f5 :: S3
+#9 := f5
+decl f6 :: S4
+#10 := f6
+decl f4 :: S2
+#8 := f4
+#11 := (f3 f4 f5 f6 f5)
+#47 := (= f6 #11)
+#50 := (not #47)
+#12 := (= #11 f6)
+#13 := (not #12)
+#51 := (iff #13 #50)
+#48 := (iff #12 #47)
+#49 := [rewrite]: #48
+#52 := [monotonicity #49]: #51
+#46 := [asserted]: #13
+#55 := [mp #46 #52]: #50
+#154 := (not #572)
+#241 := (or #154 #47)
+#155 := [quant-inst]: #241
+[unit-resolution #155 #55 #577]: false
+unsat
 31b5dadcaa7e7a3e69848f4809b04b844a6e8b56 144 0
 #2 := false
 decl f12 :: S3
@@ -51048,6 +50640,92 @@
 #432 := [unit-resolution #232 #431]: #301
 [unit-resolution #432 #433]: false
 unsat
+d2a264ef3e43d9b7d009fa11ec172212f0bd6279 85 0
+#2 := false
+decl f8 :: (-> S3 S2 S4)
+decl f4 :: S2
+#9 := f4
+decl f6 :: S3
+#12 := f6
+#15 := (f8 f6 f4)
+decl f5 :: (-> S3 S2 S4 S2 S4)
+decl f7 :: S4
+#13 := f7
+decl f3 :: S2
+#8 := f3
+#14 := (f5 f6 f3 f7 f4)
+#16 := (= #14 #15)
+#161 := (= f7 #14)
+#10 := (= f3 f4)
+#248 := (ite #10 #161 #16)
+#252 := (not #248)
+#59 := (not #16)
+#52 := (or #10 #16)
+#55 := (not #52)
+#11 := (not #10)
+#17 := (implies #11 #16)
+#18 := (not #17)
+#56 := (iff #18 #55)
+#53 := (iff #17 #52)
+#54 := [rewrite]: #53
+#57 := [monotonicity #54]: #56
+#51 := [asserted]: #18
+#60 := [mp #51 #57]: #55
+#61 := [not-or-elim #60]: #59
+#58 := [not-or-elim #60]: #11
+#254 := (or #252 #10 #16)
+#251 := [def-axiom]: #254
+#162 := [unit-resolution #251 #58 #61]: #252
+#28 := (:var 0 S2)
+#27 := (:var 1 S4)
+#26 := (:var 2 S2)
+#25 := (:var 3 S3)
+#29 := (f5 #25 #26 #27 #28)
+#586 := (pattern #29)
+#31 := (f8 #25 #28)
+#83 := (= #29 #31)
+#86 := (= #27 #29)
+#70 := (= #26 #28)
+#93 := (ite #70 #86 #83)
+#587 := (forall (vars (?v0 S3) (?v1 S2) (?v2 S4) (?v3 S2)) (:pat #586) #93)
+#100 := (forall (vars (?v0 S3) (?v1 S2) (?v2 S4) (?v3 S2)) #93)
+#590 := (iff #100 #587)
+#588 := (iff #93 #93)
+#589 := [refl]: #588
+#591 := [quant-intro #589]: #590
+#74 := (ite #70 #27 #31)
+#77 := (= #29 #74)
+#80 := (forall (vars (?v0 S3) (?v1 S2) (?v2 S4) (?v3 S2)) #77)
+#101 := (iff #80 #100)
+#94 := (iff #77 #93)
+#99 := [rewrite]: #94
+#102 := [quant-intro #99]: #101
+#91 := (~ #80 #80)
+#90 := (~ #77 #77)
+#87 := [refl]: #90
+#92 := [nnf-pos #87]: #91
+#30 := (= #28 #26)
+#32 := (ite #30 #27 #31)
+#33 := (= #29 #32)
+#34 := (forall (vars (?v0 S3) (?v1 S2) (?v2 S4) (?v3 S2)) #33)
+#81 := (iff #34 #80)
+#78 := (iff #33 #77)
+#75 := (= #32 #74)
+#72 := (iff #30 #70)
+#73 := [rewrite]: #72
+#76 := [monotonicity #73]: #75
+#79 := [monotonicity #76]: #78
+#82 := [quant-intro #79]: #81
+#69 := [asserted]: #34
+#85 := [mp #69 #82]: #80
+#88 := [mp~ #85 #92]: #80
+#103 := [mp #88 #102]: #100
+#592 := [mp #103 #591]: #587
+#163 := (not #587)
+#250 := (or #163 #248)
+#241 := [quant-inst]: #250
+[unit-resolution #241 #592 #162]: false
+unsat
 87752431ac232f724f4719a9dd746411dc8e0334 103 0
 #2 := false
 decl f6 :: S3
@@ -51152,373 +50830,6 @@
 #205 := [unit-resolution #171 #544]: #240
 [unit-resolution #205 #206]: false
 unsat
-de0be68d62a4eef58cc08ca9b43494f9219ad474 111 0
-#2 := false
-decl f9 :: (-> S2 S3)
-decl f7 :: S2
-#13 := f7
-#18 := (f9 f7)
-decl f8 :: (-> S2 S3)
-decl f3 :: S2
-#8 := f3
-#17 := (f8 f3)
-#19 := (= #17 #18)
-decl f4 :: (-> S3 S3 S2)
-decl f5 :: S3
-#9 := f5
-decl f6 :: S3
-#10 := f6
-#14 := (f4 f6 f5)
-#258 := (f9 #14)
-#220 := (= #258 #18)
-#230 := (= #18 #258)
-#15 := (= f7 #14)
-#11 := (f4 f5 f6)
-#12 := (= f3 #11)
-#16 := (and #12 #15)
-#54 := (not #16)
-#55 := (or #54 #19)
-#58 := (not #55)
-#20 := (implies #16 #19)
-#21 := (not #20)
-#59 := (iff #21 #58)
-#56 := (iff #20 #55)
-#57 := [rewrite]: #56
-#60 := [monotonicity #57]: #59
-#53 := [asserted]: #21
-#63 := [mp #53 #60]: #58
-#61 := [not-or-elim #63]: #16
-#64 := [and-elim #61]: #15
-#573 := [monotonicity #64]: #230
-#221 := [symm #573]: #220
-#561 := (= #17 #258)
-#237 := (= f5 #258)
-#23 := (:var 0 S3)
-#22 := (:var 1 S3)
-#24 := (f4 #22 #23)
-#582 := (pattern #24)
-#28 := (f9 #24)
-#75 := (= #23 #28)
-#589 := (forall (vars (?v0 S3) (?v1 S3)) (:pat #582) #75)
-#79 := (forall (vars (?v0 S3) (?v1 S3)) #75)
-#592 := (iff #79 #589)
-#590 := (iff #75 #75)
-#591 := [refl]: #590
-#593 := [quant-intro #591]: #592
-#100 := (~ #79 #79)
-#99 := (~ #75 #75)
-#96 := [refl]: #99
-#101 := [nnf-pos #96]: #100
-#29 := (= #28 #23)
-#30 := (forall (vars (?v0 S3) (?v1 S3)) #29)
-#80 := (iff #30 #79)
-#77 := (iff #29 #75)
-#78 := [rewrite]: #77
-#81 := [quant-intro #78]: #80
-#74 := [asserted]: #30
-#84 := [mp #74 #81]: #79
-#97 := [mp~ #84 #101]: #79
-#594 := [mp #97 #593]: #589
-#184 := (not #589)
-#570 := (or #184 #237)
-#242 := [quant-inst]: #570
-#581 := [unit-resolution #242 #594]: #237
-#559 := (= #17 f5)
-#164 := (f8 #11)
-#218 := (= #164 f5)
-#251 := (= f5 #164)
-#25 := (f8 #24)
-#68 := (= #22 #25)
-#583 := (forall (vars (?v0 S3) (?v1 S3)) (:pat #582) #68)
-#71 := (forall (vars (?v0 S3) (?v1 S3)) #68)
-#586 := (iff #71 #583)
-#584 := (iff #68 #68)
-#585 := [refl]: #584
-#587 := [quant-intro #585]: #586
-#108 := (~ #71 #71)
-#106 := (~ #68 #68)
-#107 := [refl]: #106
-#109 := [nnf-pos #107]: #108
-#26 := (= #25 #22)
-#27 := (forall (vars (?v0 S3) (?v1 S3)) #26)
-#72 := (iff #27 #71)
-#69 := (iff #26 #68)
-#70 := [rewrite]: #69
-#73 := [quant-intro #70]: #72
-#67 := [asserted]: #27
-#76 := [mp #67 #73]: #71
-#98 := [mp~ #76 #109]: #71
-#588 := [mp #98 #587]: #583
-#166 := (not #583)
-#253 := (or #166 #251)
-#244 := [quant-inst]: #253
-#575 := [unit-resolution #244 #588]: #251
-#219 := [symm #575]: #218
-#234 := (= #17 #164)
-#62 := [and-elim #61]: #12
-#572 := [monotonicity #62]: #234
-#560 := [trans #572 #219]: #559
-#562 := [trans #560 #581]: #561
-#563 := [trans #562 #221]: #19
-#65 := (not #19)
-#66 := [not-or-elim #63]: #65
-[unit-resolution #66 #563]: false
-unsat
-c992b1f07c8690f317843ec597ee745825b8d2b2 117 0
-#2 := false
-decl f6 :: (-> S3 S3 S2)
-decl f3 :: (-> S2 S3)
-decl f4 :: S2
-#8 := f4
-#9 := (f3 f4)
-decl f5 :: (-> S2 S3)
-#10 := (f5 f4)
-#12 := (f6 #10 #9)
-#13 := (= f4 #12)
-#349 := (f6 #9 #10)
-#550 := (= #349 #12)
-#549 := (= #12 #349)
-#11 := (= #9 #10)
-#243 := (f3 #12)
-#543 := (= #243 #10)
-#240 := (= #10 #243)
-#17 := (:var 0 S3)
-#16 := (:var 1 S3)
-#18 := (f6 #16 #17)
-#568 := (pattern #18)
-#19 := (f3 #18)
-#53 := (= #16 #19)
-#569 := (forall (vars (?v0 S3) (?v1 S3)) (:pat #568) #53)
-#57 := (forall (vars (?v0 S3) (?v1 S3)) #53)
-#572 := (iff #57 #569)
-#570 := (iff #53 #53)
-#571 := [refl]: #570
-#573 := [quant-intro #571]: #572
-#90 := (~ #57 #57)
-#88 := (~ #53 #53)
-#89 := [refl]: #88
-#91 := [nnf-pos #89]: #90
-#20 := (= #19 #16)
-#21 := (forall (vars (?v0 S3) (?v1 S3)) #20)
-#58 := (iff #21 #57)
-#55 := (iff #20 #53)
-#56 := [rewrite]: #55
-#59 := [quant-intro #56]: #58
-#52 := [asserted]: #21
-#62 := [mp #52 #59]: #57
-#80 := [mp~ #62 #91]: #57
-#574 := [mp #80 #573]: #569
-#560 := (not #569)
-#562 := (or #560 #240)
-#217 := [quant-inst]: #562
-#220 := [unit-resolution #217 #574]: #240
-#204 := [symm #220]: #543
-#559 := (= #9 #243)
-#558 := (= #243 #9)
-#557 := (= #12 f4)
-#48 := (not #11)
-#564 := [hypothesis]: #48
-#238 := (or #13 #11)
-#49 := (iff #13 #48)
-#14 := (iff #11 #13)
-#15 := (not #14)
-#50 := (iff #15 #49)
-#51 := [rewrite]: #50
-#47 := [asserted]: #15
-#54 := [mp #47 #51]: #49
-#150 := (not #49)
-#237 := (or #13 #11 #150)
-#151 := [def-axiom]: #237
-#152 := [unit-resolution #151 #54]: #238
-#565 := [unit-resolution #152 #564]: #13
-#215 := [symm #565]: #557
-#216 := [monotonicity #215]: #558
-#200 := [symm #216]: #559
-#205 := [trans #200 #204]: #11
-#206 := [unit-resolution #564 #205]: false
-#207 := [lemma #206]: #11
-#546 := (= #10 #9)
-#547 := [symm #207]: #546
-#544 := [monotonicity #547 #207]: #549
-#186 := [symm #544]: #550
-#556 := (= f4 #349)
-#25 := (:var 0 S2)
-#27 := (f5 #25)
-#582 := (pattern #27)
-#26 := (f3 #25)
-#581 := (pattern #26)
-#28 := (f6 #26 #27)
-#69 := (= #25 #28)
-#583 := (forall (vars (?v0 S2)) (:pat #581 #582) #69)
-#73 := (forall (vars (?v0 S2)) #69)
-#586 := (iff #73 #583)
-#584 := (iff #69 #69)
-#585 := [refl]: #584
-#587 := [quant-intro #585]: #586
-#84 := (~ #73 #73)
-#92 := (~ #69 #69)
-#93 := [refl]: #92
-#85 := [nnf-pos #93]: #84
-#29 := (= #28 #25)
-#30 := (forall (vars (?v0 S2)) #29)
-#74 := (iff #30 #73)
-#71 := (iff #29 #69)
-#72 := [rewrite]: #71
-#75 := [quant-intro #72]: #74
-#68 := [asserted]: #30
-#78 := [mp #68 #75]: #73
-#94 := [mp~ #78 #85]: #73
-#588 := [mp #94 #587]: #583
-#566 := (not #583)
-#561 := (or #566 #556)
-#567 := [quant-inst]: #561
-#548 := [unit-resolution #567 #588]: #556
-#551 := [trans #548 #186]: #13
-#239 := (not #13)
-#242 := (or #239 #48)
-#230 := (or #239 #48 #150)
-#241 := [def-axiom]: #230
-#170 := [unit-resolution #241 #54]: #242
-#545 := [unit-resolution #170 #207]: #239
-[unit-resolution #545 #551]: false
-unsat
-6cc8ed53fe294f253a561b590f79ef0e0271d1c2 50 0
-#2 := false
-decl f3 :: (-> S2 S3 S4 S3 S4)
-#15 := (:var 1 S3)
-#16 := (:var 0 S4)
-#14 := (:var 2 S2)
-#17 := (f3 #14 #15 #16 #15)
-#571 := (pattern #17)
-#54 := (= #16 #17)
-#572 := (forall (vars (?v0 S2) (?v1 S3) (?v2 S4)) (:pat #571) #54)
-#58 := (forall (vars (?v0 S2) (?v1 S3) (?v2 S4)) #54)
-#575 := (iff #58 #572)
-#573 := (iff #54 #54)
-#574 := [refl]: #573
-#576 := [quant-intro #574]: #575
-#87 := (~ #58 #58)
-#85 := (~ #54 #54)
-#86 := [refl]: #85
-#88 := [nnf-pos #86]: #87
-#18 := (= #17 #16)
-#19 := (forall (vars (?v0 S2) (?v1 S3) (?v2 S4)) #18)
-#59 := (iff #19 #58)
-#56 := (iff #18 #54)
-#57 := [rewrite]: #56
-#60 := [quant-intro #57]: #59
-#53 := [asserted]: #19
-#63 := [mp #53 #60]: #58
-#79 := [mp~ #63 #88]: #58
-#577 := [mp #79 #576]: #572
-decl f5 :: S3
-#9 := f5
-decl f6 :: S4
-#10 := f6
-decl f4 :: S2
-#8 := f4
-#11 := (f3 f4 f5 f6 f5)
-#47 := (= f6 #11)
-#50 := (not #47)
-#12 := (= #11 f6)
-#13 := (not #12)
-#51 := (iff #13 #50)
-#48 := (iff #12 #47)
-#49 := [rewrite]: #48
-#52 := [monotonicity #49]: #51
-#46 := [asserted]: #13
-#55 := [mp #46 #52]: #50
-#154 := (not #572)
-#241 := (or #154 #47)
-#155 := [quant-inst]: #241
-[unit-resolution #155 #55 #577]: false
-unsat
-d2a264ef3e43d9b7d009fa11ec172212f0bd6279 85 0
-#2 := false
-decl f8 :: (-> S3 S2 S4)
-decl f4 :: S2
-#9 := f4
-decl f6 :: S3
-#12 := f6
-#15 := (f8 f6 f4)
-decl f5 :: (-> S3 S2 S4 S2 S4)
-decl f7 :: S4
-#13 := f7
-decl f3 :: S2
-#8 := f3
-#14 := (f5 f6 f3 f7 f4)
-#16 := (= #14 #15)
-#161 := (= f7 #14)
-#10 := (= f3 f4)
-#248 := (ite #10 #161 #16)
-#252 := (not #248)
-#59 := (not #16)
-#52 := (or #10 #16)
-#55 := (not #52)
-#11 := (not #10)
-#17 := (implies #11 #16)
-#18 := (not #17)
-#56 := (iff #18 #55)
-#53 := (iff #17 #52)
-#54 := [rewrite]: #53
-#57 := [monotonicity #54]: #56
-#51 := [asserted]: #18
-#60 := [mp #51 #57]: #55
-#61 := [not-or-elim #60]: #59
-#58 := [not-or-elim #60]: #11
-#254 := (or #252 #10 #16)
-#251 := [def-axiom]: #254
-#162 := [unit-resolution #251 #58 #61]: #252
-#28 := (:var 0 S2)
-#27 := (:var 1 S4)
-#26 := (:var 2 S2)
-#25 := (:var 3 S3)
-#29 := (f5 #25 #26 #27 #28)
-#586 := (pattern #29)
-#31 := (f8 #25 #28)
-#83 := (= #29 #31)
-#86 := (= #27 #29)
-#70 := (= #26 #28)
-#93 := (ite #70 #86 #83)
-#587 := (forall (vars (?v0 S3) (?v1 S2) (?v2 S4) (?v3 S2)) (:pat #586) #93)
-#100 := (forall (vars (?v0 S3) (?v1 S2) (?v2 S4) (?v3 S2)) #93)
-#590 := (iff #100 #587)
-#588 := (iff #93 #93)
-#589 := [refl]: #588
-#591 := [quant-intro #589]: #590
-#74 := (ite #70 #27 #31)
-#77 := (= #29 #74)
-#80 := (forall (vars (?v0 S3) (?v1 S2) (?v2 S4) (?v3 S2)) #77)
-#101 := (iff #80 #100)
-#94 := (iff #77 #93)
-#99 := [rewrite]: #94
-#102 := [quant-intro #99]: #101
-#91 := (~ #80 #80)
-#90 := (~ #77 #77)
-#87 := [refl]: #90
-#92 := [nnf-pos #87]: #91
-#30 := (= #28 #26)
-#32 := (ite #30 #27 #31)
-#33 := (= #29 #32)
-#34 := (forall (vars (?v0 S3) (?v1 S2) (?v2 S4) (?v3 S2)) #33)
-#81 := (iff #34 #80)
-#78 := (iff #33 #77)
-#75 := (= #32 #74)
-#72 := (iff #30 #70)
-#73 := [rewrite]: #72
-#76 := [monotonicity #73]: #75
-#79 := [monotonicity #76]: #78
-#82 := [quant-intro #79]: #81
-#69 := [asserted]: #34
-#85 := [mp #69 #82]: #80
-#88 := [mp~ #85 #92]: #80
-#103 := [mp #88 #102]: #100
-#592 := [mp #103 #591]: #587
-#163 := (not #587)
-#250 := (or #163 #248)
-#241 := [quant-inst]: #250
-[unit-resolution #241 #592 #162]: false
-unsat
 13cd2a6082f8fe0947641619628aa7b793fa5eb0 158 0
 #2 := false
 decl f5 :: (-> S3 S2 S4)
@@ -51678,6 +50989,118 @@
 #73 := [not-or-elim #72]: #71
 [unit-resolution #73 #578]: false
 unsat
+de0be68d62a4eef58cc08ca9b43494f9219ad474 111 0
+#2 := false
+decl f9 :: (-> S2 S3)
+decl f7 :: S2
+#13 := f7
+#18 := (f9 f7)
+decl f8 :: (-> S2 S3)
+decl f3 :: S2
+#8 := f3
+#17 := (f8 f3)
+#19 := (= #17 #18)
+decl f4 :: (-> S3 S3 S2)
+decl f5 :: S3
+#9 := f5
+decl f6 :: S3
+#10 := f6
+#14 := (f4 f6 f5)
+#258 := (f9 #14)
+#220 := (= #258 #18)
+#230 := (= #18 #258)
+#15 := (= f7 #14)
+#11 := (f4 f5 f6)
+#12 := (= f3 #11)
+#16 := (and #12 #15)
+#54 := (not #16)
+#55 := (or #54 #19)
+#58 := (not #55)
+#20 := (implies #16 #19)
+#21 := (not #20)
+#59 := (iff #21 #58)
+#56 := (iff #20 #55)
+#57 := [rewrite]: #56
+#60 := [monotonicity #57]: #59
+#53 := [asserted]: #21
+#63 := [mp #53 #60]: #58
+#61 := [not-or-elim #63]: #16
+#64 := [and-elim #61]: #15
+#573 := [monotonicity #64]: #230
+#221 := [symm #573]: #220
+#561 := (= #17 #258)
+#237 := (= f5 #258)
+#23 := (:var 0 S3)
+#22 := (:var 1 S3)
+#24 := (f4 #22 #23)
+#582 := (pattern #24)
+#28 := (f9 #24)
+#75 := (= #23 #28)
+#589 := (forall (vars (?v0 S3) (?v1 S3)) (:pat #582) #75)
+#79 := (forall (vars (?v0 S3) (?v1 S3)) #75)
+#592 := (iff #79 #589)
+#590 := (iff #75 #75)
+#591 := [refl]: #590
+#593 := [quant-intro #591]: #592
+#100 := (~ #79 #79)
+#99 := (~ #75 #75)
+#96 := [refl]: #99
+#101 := [nnf-pos #96]: #100
+#29 := (= #28 #23)
+#30 := (forall (vars (?v0 S3) (?v1 S3)) #29)
+#80 := (iff #30 #79)
+#77 := (iff #29 #75)
+#78 := [rewrite]: #77
+#81 := [quant-intro #78]: #80
+#74 := [asserted]: #30
+#84 := [mp #74 #81]: #79
+#97 := [mp~ #84 #101]: #79
+#594 := [mp #97 #593]: #589
+#184 := (not #589)
+#570 := (or #184 #237)
+#242 := [quant-inst]: #570
+#581 := [unit-resolution #242 #594]: #237
+#559 := (= #17 f5)
+#164 := (f8 #11)
+#218 := (= #164 f5)
+#251 := (= f5 #164)
+#25 := (f8 #24)
+#68 := (= #22 #25)
+#583 := (forall (vars (?v0 S3) (?v1 S3)) (:pat #582) #68)
+#71 := (forall (vars (?v0 S3) (?v1 S3)) #68)
+#586 := (iff #71 #583)
+#584 := (iff #68 #68)
+#585 := [refl]: #584
+#587 := [quant-intro #585]: #586
+#108 := (~ #71 #71)
+#106 := (~ #68 #68)
+#107 := [refl]: #106
+#109 := [nnf-pos #107]: #108
+#26 := (= #25 #22)
+#27 := (forall (vars (?v0 S3) (?v1 S3)) #26)
+#72 := (iff #27 #71)
+#69 := (iff #26 #68)
+#70 := [rewrite]: #69
+#73 := [quant-intro #70]: #72
+#67 := [asserted]: #27
+#76 := [mp #67 #73]: #71
+#98 := [mp~ #76 #109]: #71
+#588 := [mp #98 #587]: #583
+#166 := (not #583)
+#253 := (or #166 #251)
+#244 := [quant-inst]: #253
+#575 := [unit-resolution #244 #588]: #251
+#219 := [symm #575]: #218
+#234 := (= #17 #164)
+#62 := [and-elim #61]: #12
+#572 := [monotonicity #62]: #234
+#560 := [trans #572 #219]: #559
+#562 := [trans #560 #581]: #561
+#563 := [trans #562 #221]: #19
+#65 := (not #19)
+#66 := [not-or-elim #63]: #65
+[unit-resolution #66 #563]: false
+unsat
 c3052f64b6a04dd1c8c5afaf04fa56dc411bc19c 72 0
 #2 := false
 decl f6 :: (-> S3 S2 S4 S3)
@@ -51833,6 +51256,124 @@
 #73 := [not-or-elim #72]: #71
 [unit-resolution #73 #230]: false
 unsat
+c992b1f07c8690f317843ec597ee745825b8d2b2 117 0
+#2 := false
+decl f6 :: (-> S3 S3 S2)
+decl f3 :: (-> S2 S3)
+decl f4 :: S2
+#8 := f4
+#9 := (f3 f4)
+decl f5 :: (-> S2 S3)
+#10 := (f5 f4)
+#12 := (f6 #10 #9)
+#13 := (= f4 #12)
+#349 := (f6 #9 #10)
+#550 := (= #349 #12)
+#549 := (= #12 #349)
+#11 := (= #9 #10)
+#243 := (f3 #12)
+#543 := (= #243 #10)
+#240 := (= #10 #243)
+#17 := (:var 0 S3)
+#16 := (:var 1 S3)
+#18 := (f6 #16 #17)
+#568 := (pattern #18)
+#19 := (f3 #18)
+#53 := (= #16 #19)
+#569 := (forall (vars (?v0 S3) (?v1 S3)) (:pat #568) #53)
+#57 := (forall (vars (?v0 S3) (?v1 S3)) #53)
+#572 := (iff #57 #569)
+#570 := (iff #53 #53)
+#571 := [refl]: #570
+#573 := [quant-intro #571]: #572
+#90 := (~ #57 #57)
+#88 := (~ #53 #53)
+#89 := [refl]: #88
+#91 := [nnf-pos #89]: #90
+#20 := (= #19 #16)
+#21 := (forall (vars (?v0 S3) (?v1 S3)) #20)
+#58 := (iff #21 #57)
+#55 := (iff #20 #53)
+#56 := [rewrite]: #55
+#59 := [quant-intro #56]: #58
+#52 := [asserted]: #21
+#62 := [mp #52 #59]: #57
+#80 := [mp~ #62 #91]: #57
+#574 := [mp #80 #573]: #569
+#560 := (not #569)
+#562 := (or #560 #240)
+#217 := [quant-inst]: #562
+#220 := [unit-resolution #217 #574]: #240
+#204 := [symm #220]: #543
+#559 := (= #9 #243)
+#558 := (= #243 #9)
+#557 := (= #12 f4)
+#48 := (not #11)
+#564 := [hypothesis]: #48
+#238 := (or #13 #11)
+#49 := (iff #13 #48)
+#14 := (iff #11 #13)
+#15 := (not #14)
+#50 := (iff #15 #49)
+#51 := [rewrite]: #50
+#47 := [asserted]: #15
+#54 := [mp #47 #51]: #49
+#150 := (not #49)
+#237 := (or #13 #11 #150)
+#151 := [def-axiom]: #237
+#152 := [unit-resolution #151 #54]: #238
+#565 := [unit-resolution #152 #564]: #13
+#215 := [symm #565]: #557
+#216 := [monotonicity #215]: #558
+#200 := [symm #216]: #559
+#205 := [trans #200 #204]: #11
+#206 := [unit-resolution #564 #205]: false
+#207 := [lemma #206]: #11
+#546 := (= #10 #9)
+#547 := [symm #207]: #546
+#544 := [monotonicity #547 #207]: #549
+#186 := [symm #544]: #550
+#556 := (= f4 #349)
+#25 := (:var 0 S2)
+#27 := (f5 #25)
+#582 := (pattern #27)
+#26 := (f3 #25)
+#581 := (pattern #26)
+#28 := (f6 #26 #27)
+#69 := (= #25 #28)
+#583 := (forall (vars (?v0 S2)) (:pat #581 #582) #69)
+#73 := (forall (vars (?v0 S2)) #69)
+#586 := (iff #73 #583)
+#584 := (iff #69 #69)
+#585 := [refl]: #584
+#587 := [quant-intro #585]: #586
+#84 := (~ #73 #73)
+#92 := (~ #69 #69)
+#93 := [refl]: #92
+#85 := [nnf-pos #93]: #84
+#29 := (= #28 #25)
+#30 := (forall (vars (?v0 S2)) #29)
+#74 := (iff #30 #73)
+#71 := (iff #29 #69)
+#72 := [rewrite]: #71
+#75 := [quant-intro #72]: #74
+#68 := [asserted]: #30
+#78 := [mp #68 #75]: #73
+#94 := [mp~ #78 #85]: #73
+#588 := [mp #94 #587]: #583
+#566 := (not #583)
+#561 := (or #566 #556)
+#567 := [quant-inst]: #561
+#548 := [unit-resolution #567 #588]: #556
+#551 := [trans #548 #186]: #13
+#239 := (not #13)
+#242 := (or #239 #48)
+#230 := (or #239 #48 #150)
+#241 := [def-axiom]: #230
+#170 := [unit-resolution #241 #54]: #242
+#545 := [unit-resolution #170 #207]: #239
+[unit-resolution #545 #551]: false
+unsat
 d43a3750252b86094159a0508ea79def7392e6da 187 0
 #2 := false
 decl f6 :: (-> S3 S2 S4)
--- a/src/HOL/SMT_Examples/SMT_Tests.thy	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/SMT_Examples/SMT_Tests.thy	Sat Oct 30 12:25:18 2010 -0700
@@ -104,15 +104,15 @@
   by smt+
 
 lemma
-  "distinct []"
-  "distinct [a]"
-  "distinct [a, b, c] \<longrightarrow> a \<noteq> c"
-  "distinct [a, b, c] \<longrightarrow> d = b \<longrightarrow> a \<noteq> d"
-  "\<not> distinct [a, b, a, b]"
-  "a = b \<longrightarrow> \<not>distinct [a, b]"
-  "a = b \<and> a = c \<longrightarrow> \<not>distinct [a, b, c]"
-  "distinct [a, b, c, d] \<longrightarrow> distinct [d, b, c, a]"
-  "distinct [a, b, c, d] \<longrightarrow> distinct [a, b, c] \<and> distinct [b, c, d]"
+  "SMT.distinct []"
+  "SMT.distinct [a]"
+  "SMT.distinct [a, b, c] \<longrightarrow> a \<noteq> c"
+  "SMT.distinct [a, b, c] \<longrightarrow> d = b \<longrightarrow> a \<noteq> d"
+  "\<not> SMT.distinct [a, b, a, b]"
+  "a = b \<longrightarrow> \<not> SMT.distinct [a, b]"
+  "a = b \<and> a = c \<longrightarrow> \<not> SMT.distinct [a, b, c]"
+  "SMT.distinct [a, b, c, d] \<longrightarrow> SMT.distinct [d, b, c, a]"
+  "SMT.distinct [a, b, c, d] \<longrightarrow> SMT.distinct [a, b, c] \<and> SMT.distinct [b, c, d]"
   by smt+
 
 lemma
@@ -193,7 +193,7 @@
   by smt+
 
 lemma
-  "distinct [a, b, c] \<and> (\<forall>x y. f x = f y \<longrightarrow> y = x) \<longrightarrow> f a \<noteq> f b"
+  "SMT.distinct [a, b, c] \<and> (\<forall>x y. f x = f y \<longrightarrow> y = x) \<longrightarrow> f a \<noteq> f b"
   sorry  (* FIXME: injective function *)
 
 lemma
@@ -636,7 +636,7 @@
   "i1 \<noteq> i2 \<longrightarrow> (f (i1 := v1, i2 := v2)) i2 = v2"
   "i1 = i2 \<longrightarrow> (f (i1 := v1, i2 := v2)) i1 = v2"
   "i1 = i2 \<longrightarrow> (f (i1 := v1, i2 := v2)) i1 = v2"
-  "distinct [i1, i2, i3] \<longrightarrow> (f (i1 := v1, i2 := v2)) i3 = f i3"
+  "SMT.distinct [i1, i2, i3] \<longrightarrow> (f (i1 := v1, i2 := v2)) i3 = f i3"
   by smt+
 
 
--- a/src/HOL/Sum_Type.thy	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Sum_Type.thy	Sat Oct 30 12:25:18 2010 -0700
@@ -162,6 +162,8 @@
 definition Plus :: "'a set \<Rightarrow> 'b set \<Rightarrow> ('a + 'b) set" (infixr "<+>" 65) where
   "A <+> B = Inl ` A \<union> Inr ` B"
 
+hide_const (open) Plus --"Valuable identifier"
+
 lemma InlI [intro!]: "a \<in> A \<Longrightarrow> Inl a \<in> A <+> B"
 by (simp add: Plus_def)
 
--- a/src/HOL/Tools/Datatype/datatype.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Tools/Datatype/datatype.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -703,7 +703,7 @@
     val (descr, _) = unfold_datatypes tmp_thy dts' sorts dt_info dts' i;
     val _ = check_nonempty descr handle (exn as Datatype_Empty s) =>
       if #strict config then error ("Nonemptiness check failed for datatype " ^ s)
-      else raise exn;
+      else reraise exn;
 
     val _ = message config ("Constructing datatype(s) " ^ commas_quote new_type_names);
 
--- a/src/HOL/Tools/Meson/meson.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Tools/Meson/meson.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -111,39 +111,57 @@
     SOME th => th
   | NONE => raise THM ("first_order_resolve", 0, [thA, thB]))
 
-(* Applying "choice" swaps the bound variable names. We tweak
-   "Thm.rename_boundvars"'s input to get the desired names. *)
-fun fix_bounds (_ $ (Const (@{const_name Ex}, _)
-                     $ Abs (_, _, Const (@{const_name All}, _) $ _)))
-               (t0 $ (Const (@{const_name All}, T1)
-                      $ Abs (a1, T1', Const (@{const_name Ex}, T2)
-                                      $ Abs (a2, T2', t')))) =
-    t0 $ (Const (@{const_name All}, T1)
-          $ Abs (a2, T1', Const (@{const_name Ex}, T2) $ Abs (a1, T2', t')))
-  | fix_bounds _ t = t
+(* Hack to make it less likely that we lose our precious bound variable names in
+   "rename_bound_vars_RS" below, because of a clash. *)
+val protect_prefix = "Meson_xyzzy"
+
+fun protect_bound_var_names (t $ u) =
+    protect_bound_var_names t $ protect_bound_var_names u
+  | protect_bound_var_names (Abs (s, T, t')) =
+    Abs (protect_prefix ^ s, T, protect_bound_var_names t')
+  | protect_bound_var_names t = t
 
-(* Hack to make it less likely that we lose our precious bound variable names in
-   "rename_bvs_RS" below, because of a clash. *)
-val protect_prefix = "_"
+fun fix_bound_var_names old_t new_t =
+  let
+    fun quant_of @{const_name All} = SOME true
+      | quant_of @{const_name Ball} = SOME true
+      | quant_of @{const_name Ex} = SOME false
+      | quant_of @{const_name Bex} = SOME false
+      | quant_of _ = NONE
+    val flip_quant = Option.map not
+    fun some_eq (SOME x) (SOME y) = x = y
+      | some_eq _ _ = false
+    fun add_names quant (Const (quant_s, _) $ Abs (s, _, t')) =
+        add_names quant t' #> some_eq quant (quant_of quant_s) ? cons s
+      | add_names quant (@{const Not} $ t) = add_names (flip_quant quant) t
+      | add_names quant (@{const implies} $ t1 $ t2) =
+        add_names (flip_quant quant) t1 #> add_names quant t2
+      | add_names quant (t1 $ t2) = fold (add_names quant) [t1, t2]
+      | add_names _ _ = I
+    fun lost_names quant =
+      subtract (op =) (add_names quant new_t []) (add_names quant old_t [])
+    fun aux ((t1 as Const (quant_s, _)) $ (Abs (s, T, t'))) =
+      t1 $ Abs (s |> String.isPrefix protect_prefix s
+                   ? perhaps (try (fn _ => hd (lost_names (quant_of quant_s)))),
+                T, aux t')
+      | aux (t1 $ t2) = aux t1 $ aux t2
+      | aux t = t
+  in aux new_t end
 
-fun protect_bounds (t $ u) = protect_bounds t $ protect_bounds u
-  | protect_bounds (Abs (s, T, t')) =
-    Abs (protect_prefix ^ s, T, protect_bounds t')
-  | protect_bounds t = t
-
-(* Forward proof while preserving bound variables names*)
-fun rename_bvs_RS th rl =
+(* Forward proof while preserving bound variables names *)
+fun rename_bound_vars_RS th rl =
   let
     val t = concl_of th
     val r = concl_of rl
-    val th' = th RS Thm.rename_boundvars r (protect_bounds r) rl
+    val th' = th RS Thm.rename_boundvars r (protect_bound_var_names r) rl
     val t' = concl_of th'
-  in Thm.rename_boundvars t' (fix_bounds t' t) th' end
+  in Thm.rename_boundvars t' (fix_bound_var_names t t') th' end
 
 (*raises exception if no rules apply*)
 fun tryres (th, rls) =
   let fun tryall [] = raise THM("tryres", 0, th::rls)
-        | tryall (rl::rls) = (rename_bvs_RS th rl handle THM _ => tryall rls)
+        | tryall (rl::rls) =
+          (rename_bound_vars_RS th rl handle THM _ => tryall rls)
   in  tryall rls  end;
 
 (*Permits forward proof from rules that discharge assumptions. The supplied proof state st,
@@ -577,7 +595,7 @@
                |> forward_res ctxt aux
                |> aux
                handle THM ("tryres", _, _) =>
-                      rename_bvs_RS th ex_forward
+                      rename_bound_vars_RS th ex_forward
                       |> forward_res ctxt aux
   in aux o make_nnf ctxt end
 
--- a/src/HOL/Tools/Meson/meson_clausify.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Tools/Meson/meson_clausify.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -233,7 +233,7 @@
 fun zapped_var_name ((ax_no, cluster_no), skolem) index_no s =
   (if skolem then new_skolem_var_prefix else new_nonskolem_var_prefix) ^
   "_" ^ string_of_int ax_no ^ "_" ^ string_of_int cluster_no ^ "_" ^
-  string_of_int index_no ^ "_" ^ s
+  string_of_int index_no ^ "_" ^ Name.desymbolize false s
 
 fun cluster_of_zapped_var_name s =
   let val get_int = the o Int.fromString o nth (space_explode "_" s) in
@@ -241,10 +241,11 @@
      String.isPrefix new_skolem_var_prefix s)
   end
 
-fun zap (cluster as (cluster_no, cluster_skolem)) index_no pos ct =
-  ct
-  |> (case term_of ct of
-        Const (s, _) $ Abs (s', _, _) =>
+fun rename_bound_vars_to_be_zapped ax_no =
+  let
+    fun aux (cluster as (cluster_no, cluster_skolem)) index_no pos t =
+      case t of
+        (t1 as Const (s, _)) $ Abs (s', T, t') =>
         if s = @{const_name all} orelse s = @{const_name All} orelse
            s = @{const_name Ex} then
           let
@@ -252,34 +253,53 @@
             val (cluster, index_no) =
               if skolem = cluster_skolem then (cluster, index_no)
               else ((cluster_no ||> cluster_skolem ? Integer.add 1, skolem), 0)
-          in
-            Thm.dest_comb #> snd
-            #> Thm.dest_abs (SOME (zapped_var_name cluster index_no s'))
-            #> snd #> zap cluster (index_no + 1) pos
-          end
+            val s' = zapped_var_name cluster index_no s'
+          in t1 $ Abs (s', T, aux cluster (index_no + 1) pos t') end
+        else
+          t
+      | (t1 as Const (s, _)) $ t2 $ t3 =>
+        if s = @{const_name "==>"} orelse s = @{const_name HOL.implies} then
+          t1 $ aux cluster index_no (not pos) t2 $ aux cluster index_no pos t3
+        else if s = @{const_name HOL.conj} orelse
+                s = @{const_name HOL.disj} then
+          t1 $ aux cluster index_no pos t2 $ aux cluster index_no pos t3
+        else
+          t
+      | (t1 as Const (s, _)) $ t2 =>
+        if s = @{const_name Trueprop} then
+          t1 $ aux cluster index_no pos t2
+        else if s = @{const_name Not} then
+          t1 $ aux cluster index_no (not pos) t2
+        else
+          t
+      | _ => t
+  in aux ((ax_no, 0), true) 0 true end
+
+fun zap pos ct =
+  ct
+  |> (case term_of ct of
+        Const (s, _) $ Abs (s', _, _) =>
+        if s = @{const_name all} orelse s = @{const_name All} orelse
+           s = @{const_name Ex} then
+          Thm.dest_comb #> snd #> Thm.dest_abs (SOME s') #> snd #> zap pos
         else
           Conv.all_conv
       | Const (s, _) $ _ $ _ =>
         if s = @{const_name "==>"} orelse s = @{const_name implies} then
-          Conv.combination_conv (Conv.arg_conv (zap cluster index_no (not pos)))
-                                (zap cluster index_no pos)
+          Conv.combination_conv (Conv.arg_conv (zap (not pos))) (zap pos)
         else if s = @{const_name conj} orelse s = @{const_name disj} then
-          Conv.combination_conv (Conv.arg_conv (zap cluster index_no pos))
-                                (zap cluster index_no pos)
+          Conv.combination_conv (Conv.arg_conv (zap pos)) (zap pos)
         else
           Conv.all_conv
       | Const (s, _) $ _ =>
-        if s = @{const_name Trueprop} then
-          Conv.arg_conv (zap cluster index_no pos)
-        else if s = @{const_name Not} then
-          Conv.arg_conv (zap cluster index_no (not pos))
-        else
-          Conv.all_conv
+        if s = @{const_name Trueprop} then Conv.arg_conv (zap pos)
+        else if s = @{const_name Not} then Conv.arg_conv (zap (not pos))
+        else Conv.all_conv
       | _ => Conv.all_conv)
 
 fun ss_only ths = MetaSimplifier.clear_ss HOL_basic_ss addsimps ths
 
-val no_choice =
+val cheat_choice =
   @{prop "ALL x. EX y. Q x y ==> EX f. ALL x. Q x (f x)"}
   |> Logic.varify_global
   |> Skip_Proof.make_thm @{theory}
@@ -304,29 +324,34 @@
           #> simplify (ss_only @{thms all_simps[symmetric]})
         val pull_out =
           simplify (ss_only @{thms all_simps[symmetric] ex_simps[symmetric]})
-        val (discharger_th, fully_skolemized_th) =
-          if null choice_ths then
-            th |> `I |>> pull_out ||> skolemize [no_choice]
-          else
-            th |> skolemize choice_ths |> `I
-        val t =
-          fully_skolemized_th |> cprop_of
-          |> zap ((ax_no, 0), true) 0 true |> Drule.export_without_context
+        val no_choice = null choice_ths
+        val discharger_th =
+          th |> (if no_choice then pull_out else skolemize choice_ths)
+        val fully_skolemized_t =
+          discharger_th |> prop_of |> rename_bound_vars_to_be_zapped ax_no
+          |> (if no_choice then
+                Skip_Proof.make_thm thy #> skolemize [cheat_choice] #> cprop_of
+              else
+                cterm_of thy)
+          |> zap true |> Drule.export_without_context
           |> cprop_of |> Thm.dest_equals |> snd |> term_of
       in
         if exists_subterm (fn Var ((s, _), _) =>
                               String.isPrefix new_skolem_var_prefix s
-                            | _ => false) t then
+                            | _ => false) fully_skolemized_t then
           let
-            val (ct, ctxt) =
-              Variable.import_terms true [t] ctxt
+            val (fully_skolemized_ct, ctxt) =
+              Variable.import_terms true [fully_skolemized_t] ctxt
               |>> the_single |>> cterm_of thy
-          in (SOME (discharger_th, ct), Thm.assume ct, ctxt) end
+          in
+            (SOME (discharger_th, fully_skolemized_ct),
+             (Thm.assume fully_skolemized_ct, ctxt))
+          end
        else
-         (NONE, th, ctxt)
+         (NONE, (th, ctxt))
       end
     else
-      (NONE, th, ctxt)
+      (NONE, (th, ctxt))
   end
 
 (* Convert a theorem to CNF, with additional premises due to skolemization. *)
@@ -334,7 +359,8 @@
   let
     val thy = ProofContext.theory_of ctxt0
     val choice_ths = choice_theorems thy
-    val (opt, nnf_th, ctxt) = nnf_axiom choice_ths new_skolemizer ax_no th ctxt0
+    val (opt, (nnf_th, ctxt)) =
+      nnf_axiom choice_ths new_skolemizer ax_no th ctxt0
     fun clausify th =
       make_cnf (if new_skolemizer orelse null choice_ths then
                   []
@@ -344,6 +370,7 @@
     val (cnf_ths, ctxt) =
       clausify nnf_th
       |> (fn ([], _) =>
+             (* FIXME: This probably doesn't work with the new Skolemizer *)
              clausify (to_definitional_cnf_with_quantifiers thy nnf_th)
            | p => p)
     fun intr_imp ct th =
--- a/src/HOL/Tools/Metis/metis_reconstruct.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Tools/Metis/metis_reconstruct.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -16,7 +16,7 @@
   val lookth : (Metis_Thm.thm * 'a) list -> Metis_Thm.thm -> 'a
   val untyped_aconv : term -> term -> bool
   val replay_one_inference :
-    Proof.context -> mode -> (string * term) list
+    Proof.context -> mode -> (string * term) list * int Unsynchronized.ref
     -> Metis_Thm.thm * Metis_Proof.inference -> (Metis_Thm.thm * thm) list
     -> (Metis_Thm.thm * thm) list
   val discharge_skolem_premises :
@@ -93,7 +93,7 @@
         | NONE    => raise Fail ("hol_type_from_metis_term: " ^ x));
 
 (*Maps metis terms to isabelle terms*)
-fun hol_term_from_metis_PT ctxt fol_tm =
+fun hol_term_from_metis_PT ctxt new_skolem_var_count fol_tm =
   let val thy = ProofContext.theory_of ctxt
       val _ = trace_msg ctxt (fn () => "hol_term_from_metis_PT: " ^
                                        Metis_Term.toString fol_tm)
@@ -126,9 +126,11 @@
                     val nterms = length ts - ntypes
                     val tts = map tm_to_tt ts
                     val tys = types_of (List.take(tts,ntypes))
+                    val j = !new_skolem_var_count + 1
+                    val _ = new_skolem_var_count := j
                     val t =
                       if String.isPrefix new_skolem_const_prefix c then
-                        Var (new_skolem_var_from_const c,
+                        Var ((new_skolem_var_name_from_const c, j),
                              Type_Infer.paramify_vars (tl tys ---> hd tys))
                       else
                         Const (c, dummyT)
@@ -162,7 +164,7 @@
   end
 
 (*Maps fully-typed metis terms to isabelle terms*)
-fun hol_term_from_metis_FT ctxt fol_tm =
+fun hol_term_from_metis_FT ctxt new_skolem_var_count fol_tm =
   let val _ = trace_msg ctxt (fn () => "hol_term_from_metis_FT: " ^
                                        Metis_Term.toString fol_tm)
       fun cvt (Metis_Term.Fn ("ti", [Metis_Term.Var v, _])) =
@@ -193,17 +195,17 @@
                 SOME v => Free (v, dummyT)
               | NONE => (trace_msg ctxt (fn () =>
                                       "hol_term_from_metis_FT bad const: " ^ x);
-                         hol_term_from_metis_PT ctxt t))
+                         hol_term_from_metis_PT ctxt new_skolem_var_count t))
         | cvt t = (trace_msg ctxt (fn () =>
                    "hol_term_from_metis_FT bad term: " ^ Metis_Term.toString t);
-                   hol_term_from_metis_PT ctxt t)
+                   hol_term_from_metis_PT ctxt new_skolem_var_count t)
   in fol_tm |> cvt end
 
 fun hol_term_from_metis FT = hol_term_from_metis_FT
   | hol_term_from_metis _ = hol_term_from_metis_PT
 
-fun hol_terms_from_fol ctxt mode old_skolems fol_tms =
-  let val ts = map (hol_term_from_metis mode ctxt) fol_tms
+fun hol_terms_from_metis ctxt mode (old_skolems, new_skolem_var_count) fol_tms =
+  let val ts = map (hol_term_from_metis mode ctxt new_skolem_var_count) fol_tms
       val _ = trace_msg ctxt (fn () => "  calling type inference:")
       val _ = app (fn t => trace_msg ctxt
                                      (fn () => Syntax.string_of_term ctxt t)) ts
@@ -249,17 +251,18 @@
       val substs = [(cterm_of thy (Var vx), cterm_of thy i_atm)]
   in  cterm_instantiate substs th  end;
 
-fun assume_inf ctxt mode old_skolems atm =
+fun assume_inf ctxt mode skolem_params atm =
   inst_excluded_middle
       (ProofContext.theory_of ctxt)
-      (singleton (hol_terms_from_fol ctxt mode old_skolems) (Metis_Term.Fn atm))
+      (singleton (hol_terms_from_metis ctxt mode skolem_params)
+                 (Metis_Term.Fn atm))
 
 (* INFERENCE RULE: INSTANTIATE (SUBST). Type instantiations are ignored. Trying
    to reconstruct them admits new possibilities of errors, e.g. concerning
    sorts. Instead we try to arrange that new TVars are distinct and that types
    can be inferred from terms. *)
 
-fun inst_inf ctxt mode old_skolems thpairs fsubst th =
+fun inst_inf ctxt mode (old_skolems, new_skolem_var_count) thpairs fsubst th =
   let val thy = ProofContext.theory_of ctxt
       val i_th = lookth thpairs th
       val i_th_vars = Term.add_vars (prop_of i_th) []
@@ -267,7 +270,7 @@
       fun subst_translation (x,y) =
         let val v = find_var x
             (* We call "reveal_old_skolem_terms" and "infer_types" below. *)
-            val t = hol_term_from_metis mode ctxt y
+            val t = hol_term_from_metis mode ctxt new_skolem_var_count y
         in  SOME (cterm_of thy (Var v), t)  end
         handle Option.Option =>
                (trace_msg ctxt (fn () => "\"find_var\" failed for " ^ x ^
@@ -336,8 +339,10 @@
            | ps => aux (instantiate (ps, []) tha) (instantiate (ps, []) thb)
   end
 
-fun mk_not (Const (@{const_name Not}, _) $ b) = b
-  | mk_not b = HOLogic.mk_not b
+fun s_not (@{const Not} $ t) = t
+  | s_not t = HOLogic.mk_not t
+fun simp_not_not (@{const Not} $ t) = s_not (simp_not_not t)
+  | simp_not_not t = t
 
 (* Match untyped terms. *)
 fun untyped_aconv (Const (a, _)) (Const(b, _)) = (a = b)
@@ -351,16 +356,12 @@
   | untyped_aconv _ _ = false
 
 (* Finding the relative location of an untyped term within a list of terms *)
-fun literal_index lit =
+fun index_of_literal lit haystack =
   let
-    val lit = Envir.eta_contract lit
-    fun get _ [] = raise Empty
-      | get n (x :: xs) =
-        if untyped_aconv lit (Envir.eta_contract (HOLogic.dest_Trueprop x)) then
-          n
-        else
-          get (n+1) xs
-  in get 1 end
+    val normalize = simp_not_not o Envir.eta_contract
+    val match_lit =
+      HOLogic.dest_Trueprop #> normalize #> untyped_aconv (lit |> normalize)
+  in case find_index match_lit haystack of ~1 => raise Empty | n => n + 1 end
 
 (* Permute a rule's premises to move the i-th premise to the last position. *)
 fun make_last i th =
@@ -377,7 +378,7 @@
 (* Maps the clause  [P1,...Pn]==>False to [P1,...,P(i-1),P(i+1),...Pn] ==> ~P *)
 val select_literal = negate_head oo make_last
 
-fun resolve_inf ctxt mode old_skolems thpairs atm th1 th2 =
+fun resolve_inf ctxt mode skolem_params thpairs atm th1 th2 =
   let
     val thy = ProofContext.theory_of ctxt
     val i_th1 = lookth thpairs th1 and i_th2 = lookth thpairs th2
@@ -391,16 +392,19 @@
       i_th1
     else
       let
-        val i_atm = singleton (hol_terms_from_fol ctxt mode old_skolems)
-                              (Metis_Term.Fn atm)
+        val i_atm =
+          singleton (hol_terms_from_metis ctxt mode skolem_params)
+                    (Metis_Term.Fn atm)
         val _ = trace_msg ctxt (fn () => "  atom: " ^ Syntax.string_of_term ctxt i_atm)
         val prems_th1 = prems_of i_th1
         val prems_th2 = prems_of i_th2
-        val index_th1 = literal_index (mk_not i_atm) prems_th1
-              handle Empty => raise Fail "Failed to find literal in th1"
+        val index_th1 =
+          index_of_literal (s_not i_atm) prems_th1
+          handle Empty => raise Fail "Failed to find literal in th1"
         val _ = trace_msg ctxt (fn () => "  index_th1: " ^ Int.toString index_th1)
-        val index_th2 = literal_index i_atm prems_th2
-              handle Empty => raise Fail "Failed to find literal in th2"
+        val index_th2 =
+          index_of_literal i_atm prems_th2
+          handle Empty => raise Fail "Failed to find literal in th2"
         val _ = trace_msg ctxt (fn () => "  index_th2: " ^ Int.toString index_th2)
     in
       resolve_inc_tyvars thy (select_literal index_th1 i_th1) index_th2 i_th2
@@ -414,9 +418,9 @@
 val refl_x = cterm_of @{theory} (Var (hd (Term.add_vars (prop_of REFL_THM) [])));
 val refl_idx = 1 + Thm.maxidx_of REFL_THM;
 
-fun refl_inf ctxt mode old_skolems t =
+fun refl_inf ctxt mode skolem_params t =
   let val thy = ProofContext.theory_of ctxt
-      val i_t = singleton (hol_terms_from_fol ctxt mode old_skolems) t
+      val i_t = singleton (hol_terms_from_metis ctxt mode skolem_params) t
       val _ = trace_msg ctxt (fn () => "  term: " ^ Syntax.string_of_term ctxt i_t)
       val c_t = cterm_incr_types thy refl_idx i_t
   in  cterm_instantiate [(refl_x, c_t)] REFL_THM  end;
@@ -432,10 +436,10 @@
   | get_ty_arg_size thy (Const (c, _)) = (num_type_args thy c handle TYPE _ => 0)
   | get_ty_arg_size _ _ = 0;
 
-fun equality_inf ctxt mode old_skolems (pos, atm) fp fr =
+fun equality_inf ctxt mode skolem_params (pos, atm) fp fr =
   let val thy = ProofContext.theory_of ctxt
       val m_tm = Metis_Term.Fn atm
-      val [i_atm,i_tm] = hol_terms_from_fol ctxt mode old_skolems [m_tm, fr]
+      val [i_atm,i_tm] = hol_terms_from_metis ctxt mode skolem_params [m_tm, fr]
       val _ = trace_msg ctxt (fn () => "sign of the literal: " ^ Bool.toString pos)
       fun replace_item_list lx 0 (_::ls) = lx::ls
         | replace_item_list lx i (l::ls) = l :: replace_item_list lx (i-1) ls
@@ -514,17 +518,17 @@
 
 val factor = Seq.hd o distinct_subgoals_tac;
 
-fun step ctxt mode old_skolems thpairs p =
+fun step ctxt mode skolem_params thpairs p =
   case p of
     (fol_th, Metis_Proof.Axiom _) => factor (axiom_inf thpairs fol_th)
-  | (_, Metis_Proof.Assume f_atm) => assume_inf ctxt mode old_skolems f_atm
+  | (_, Metis_Proof.Assume f_atm) => assume_inf ctxt mode skolem_params f_atm
   | (_, Metis_Proof.Metis_Subst (f_subst, f_th1)) =>
-    factor (inst_inf ctxt mode old_skolems thpairs f_subst f_th1)
+    factor (inst_inf ctxt mode skolem_params thpairs f_subst f_th1)
   | (_, Metis_Proof.Resolve(f_atm, f_th1, f_th2)) =>
-    factor (resolve_inf ctxt mode old_skolems thpairs f_atm f_th1 f_th2)
-  | (_, Metis_Proof.Refl f_tm) => refl_inf ctxt mode old_skolems f_tm
+    factor (resolve_inf ctxt mode skolem_params thpairs f_atm f_th1 f_th2)
+  | (_, Metis_Proof.Refl f_tm) => refl_inf ctxt mode skolem_params f_tm
   | (_, Metis_Proof.Equality (f_lit, f_p, f_r)) =>
-    equality_inf ctxt mode old_skolems f_lit f_p f_r
+    equality_inf ctxt mode skolem_params f_lit f_p f_r
 
 fun flexflex_first_order th =
   case Thm.tpairs_of th of
@@ -544,12 +548,12 @@
 
 fun count p xs = fold (fn x => if p x then Integer.add 1 else I) xs 0
 
-fun replay_one_inference ctxt mode old_skolems (fol_th, inf) thpairs =
+fun replay_one_inference ctxt mode skolem_params (fol_th, inf) thpairs =
   let
     val _ = trace_msg ctxt (fn () => "=============================================")
     val _ = trace_msg ctxt (fn () => "METIS THM: " ^ Metis_Thm.toString fol_th)
     val _ = trace_msg ctxt (fn () => "INFERENCE: " ^ Metis_Proof.inferenceToString inf)
-    val th = step ctxt mode old_skolems thpairs (fol_th, inf)
+    val th = step ctxt mode skolem_params thpairs (fol_th, inf)
              |> flexflex_first_order
     val _ = trace_msg ctxt (fn () => "ISABELLE THM: " ^ Display.string_of_thm ctxt th)
     val _ = trace_msg ctxt (fn () => "=============================================")
@@ -608,10 +612,6 @@
       | _ => if untyped_aconv t u then I else raise TERM ("unify_terms", [t, u])
   in th |> term_instantiate thy (unify_terms (prem, concl) []) end
 
-fun shuffle_key (((axiom_no, (_, index_no)), _), _) = (index_no, axiom_no)
-fun shuffle_ord p =
-  rev_order (prod_ord int_ord int_ord (pairself shuffle_key p))
-
 val copy_prem = @{lemma "P ==> (P ==> P ==> Q) ==> Q" by fast}
 
 fun copy_prems_tac [] ns i =
@@ -621,13 +621,21 @@
   | copy_prems_tac (m :: ms) ns i =
     etac copy_prem i THEN copy_prems_tac ms (m div 2 :: (m + 1) div 2 :: ns) i
 
-fun instantiate_forall_tac thy params t i =
+fun instantiate_forall_tac thy t i st =
   let
+    val params = Logic.strip_params (Logic.get_goal (prop_of st) i) |> rev
     fun repair (t as (Var ((s, _), _))) =
-        (case find_index (fn ((s', _), _) => s' = s) params of
+        (case find_index (fn (s', _) => s' = s) params of
            ~1 => t
          | j => Bound j)
-      | repair (t $ u) = repair t $ repair u
+      | repair (t $ u) =
+        (case (repair t, repair u) of
+           (t as Bound j, u as Bound k) =>
+           (* This is a rather subtle trick to repair the discrepancy between
+              the fully skolemized term that MESON gives us (where existentials
+              were pulled out) and the reality. *)
+           if k > j then t else t $ u
+         | (t, u) => t $ u)
       | repair t = t
     val t' = t |> repair |> fold (curry absdummy) (map snd params)
     fun do_instantiate th =
@@ -635,41 +643,50 @@
         th |> term_instantiate thy [(Var var, t')]
       end
   in
-    etac @{thm allE} i
-    THEN rotate_tac ~1 i
-    THEN PRIMITIVE do_instantiate
+    (etac @{thm allE} i
+     THEN rotate_tac ~1 i
+     THEN PRIMITIVE do_instantiate) st
   end
 
-fun release_clusters_tac _ _ _ _ [] = K all_tac
-  | release_clusters_tac thy ax_counts substs params
+fun fix_exists_tac thy t =
+  etac @{thm exE}
+  THEN' rename_tac [t |> dest_Var |> fst |> fst]
+
+fun release_quantifier_tac thy (skolem, t) =
+  (if skolem then fix_exists_tac else instantiate_forall_tac) thy t
+
+fun release_clusters_tac _ _ _ [] = K all_tac
+  | release_clusters_tac thy ax_counts substs
                          ((ax_no, cluster_no) :: clusters) =
     let
-      fun in_right_cluster s =
-        (s |> Meson_Clausify.cluster_of_zapped_var_name |> fst |> snd |> fst)
-        = cluster_no
+      val cluster_of_var =
+        Meson_Clausify.cluster_of_zapped_var_name o fst o fst o dest_Var
+      fun in_right_cluster ((_, (cluster_no', _)), _) = cluster_no' = cluster_no
       val cluster_substs =
         substs
         |> map_filter (fn (ax_no', (_, (_, tsubst))) =>
                           if ax_no' = ax_no then
-                            tsubst |> filter (in_right_cluster
-                                              o fst o fst o dest_Var o fst)
-                                   |> map snd |> SOME
-                           else
-                             NONE)
-      val n = length cluster_substs
+                            tsubst |> map (apfst cluster_of_var)
+                                   |> filter (in_right_cluster o fst)
+                                   |> map (apfst snd)
+                                   |> SOME
+                          else
+                            NONE)
       fun do_cluster_subst cluster_subst =
-        map (instantiate_forall_tac thy params) cluster_subst @ [rotate_tac 1]
-      val params' = params (* FIXME ### existentials! *)
+        map (release_quantifier_tac thy) cluster_subst @ [rotate_tac 1]
       val first_prem = find_index (fn (ax_no', _) => ax_no' = ax_no) substs
     in
       rotate_tac first_prem
       THEN' (EVERY' (maps do_cluster_subst cluster_substs))
       THEN' rotate_tac (~ first_prem - length cluster_substs)
-      THEN' release_clusters_tac thy ax_counts substs params' clusters
+      THEN' release_clusters_tac thy ax_counts substs clusters
     end
 
-val cluster_ord =
-  prod_ord (prod_ord int_ord (prod_ord int_ord int_ord)) bool_ord
+fun cluster_key ((ax_no, (cluster_no, index_no)), skolem) =
+  (ax_no, (cluster_no, (skolem, index_no)))
+fun cluster_ord p =
+  prod_ord int_ord (prod_ord int_ord (prod_ord bool_ord int_ord))
+           (pairself cluster_key p)
 
 val tysubst_ord =
   list_ord (prod_ord Term_Ord.fast_indexname_ord
@@ -682,6 +699,9 @@
 structure Int_Pair_Graph =
   Graph(type key = int * int val ord = prod_ord int_ord int_ord)
 
+fun shuffle_key (((axiom_no, (_, index_no)), _), _) = (index_no, axiom_no)
+fun shuffle_ord p = prod_ord int_ord int_ord (pairself shuffle_key p)
+
 (* Attempts to derive the theorem "False" from a theorem of the form
    "P1 ==> ... ==> Pn ==> False", where the "Pi"s are to be discharged using the
    specified axioms. The axioms have leading "All" and "Ex" quantifiers, which
@@ -693,6 +713,7 @@
     let
       val thy = ProofContext.theory_of ctxt
       (* distinguish variables with same name but different types *)
+      (* ### FIXME: needed? *)
       val prems_imp_false' =
         prems_imp_false |> try (forall_intr_vars #> gen_all)
                         |> the_default prems_imp_false
@@ -756,12 +777,12 @@
         handle Int_Pair_Graph.CYCLES _ =>
                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))
+      val outer_param_names =
+        [] |> fold (Term.add_var_names o snd) (map_filter I axioms)
+           |> map (`(Meson_Clausify.cluster_of_zapped_var_name o fst))
            |> filter (fn (((_, (cluster_no, _)), skolem), _) =>
                          cluster_no = 0 andalso skolem)
-           |> sort shuffle_ord |> map snd
+           |> sort shuffle_ord |> map (fst o snd)
       val ax_counts =
         Int_Tysubst_Table.empty
         |> fold (fn (ax_no, (_, (tysubst, _))) =>
@@ -774,11 +795,10 @@
         "; tysubst: " ^ PolyML.makestring tysubst ^ "; tsubst: {" ^
         commas (map ((fn (s, t) => s ^ " |-> " ^ t)
                      o pairself (Syntax.string_of_term ctxt)) tsubst) ^ "}"
+      val _ = tracing ("ORDERED CLUSTERS: " ^ PolyML.makestring ordered_clusters)
+      val _ = tracing ("AXIOM COUNTS: " ^ PolyML.makestring ax_counts)
       val _ = tracing ("SUBSTS (" ^ string_of_int (length substs) ^ "):\n" ^
                        cat_lines (map string_for_subst_info substs))
-      val _ = tracing ("OUTERMOST SKOLEMS: " ^ PolyML.makestring params0)
-      val _ = tracing ("ORDERED CLUSTERS: " ^ PolyML.makestring ordered_clusters)
-      val _ = tracing ("AXIOM COUNTS: " ^ PolyML.makestring ax_counts)
 *)
       fun rotation_for_subgoal i =
         find_index (fn (_, (subgoal_no, _)) => subgoal_no = i) substs
@@ -787,14 +807,14 @@
           (K (cut_rules_tac
                   (map (fst o the o nth axioms o fst o fst) ax_counts) 1
               THEN TRY (REPEAT_ALL_NEW (etac @{thm exE}) 1)
+              THEN rename_tac outer_param_names 1
               THEN copy_prems_tac (map snd ax_counts) [] 1
-              THEN release_clusters_tac thy ax_counts substs params0
-                                        ordered_clusters 1
+              THEN release_clusters_tac thy ax_counts substs ordered_clusters 1
               THEN match_tac [prems_imp_false] 1
               THEN ALLGOALS (fn i =>
                        rtac @{thm Meson.skolem_COMBK_I} i
                        THEN rotate_tac (rotation_for_subgoal i) i
-(*                       THEN PRIMITIVE (unify_first_prem_with_concl thy i) ###*)
+(*                       THEN PRIMITIVE (unify_first_prem_with_concl thy i) ### FIXME: needed? *)
                        THEN assume_tac i)))
       handle ERROR _ =>
              error ("Cannot replay Metis proof in Isabelle:\n\
--- a/src/HOL/Tools/Metis/metis_tactics.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Tools/Metis/metis_tactics.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -91,9 +91,11 @@
                           Metis_Thm.toString mth)
                 val ctxt' = fold Variable.declare_constraints (map prop_of cls) ctxt
                              (*add constraints arising from converting goal to clause form*)
+                val new_skolem_var_count = Unsynchronized.ref 1
                 val proof = Metis_Proof.proof mth
                 val result =
-                  fold (replay_one_inference ctxt' mode old_skolems) proof axioms
+                  fold (replay_one_inference ctxt' mode
+                              (old_skolems, new_skolem_var_count)) proof axioms
                 and used = map_filter (used_axioms axioms) proof
                 val _ = trace_msg ctxt (fn () => "METIS COMPLETED...clauses actually used:")
                 val _ = app (fn th => trace_msg ctxt (fn () => Display.string_of_thm ctxt th)) used
--- a/src/HOL/Tools/Metis/metis_translate.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Tools/Metis/metis_translate.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -59,7 +59,7 @@
   val make_fixed_type_const : string -> string
   val make_type_class : string -> string
   val num_type_args: theory -> string -> int
-  val new_skolem_var_from_const: string -> indexname
+  val new_skolem_var_name_from_const : string -> string
   val type_literals_for_types : typ list -> type_literal list
   val make_class_rel_clauses :
     theory -> class list -> class list -> class_rel_clause list
@@ -214,9 +214,9 @@
   else
     (s, Sign.the_const_type thy s) |> Sign.const_typargs thy |> length
 
-fun new_skolem_var_from_const s =
+fun new_skolem_var_name_from_const s =
   let val ss = s |> space_explode Long_Name.separator in
-    (nth ss (length ss - 2), 0)
+    nth ss (length ss - 2)
   end
 
 
--- a/src/HOL/Tools/Nitpick/kodkod.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Tools/Nitpick/kodkod.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -1036,12 +1036,7 @@
           val outcome =
             let
               val code =
-                bash ("cd " ^ File.shell_quote temp_dir ^ ";\n" ^
-                      "env CLASSPATH=\"$KODKODI_CLASSPATH:$CLASSPATH\" \
-                      \JAVA_LIBRARY_PATH=\"$KODKODI_JAVA_LIBRARY_PATH:\
-                      \$JAVA_LIBRARY_PATH\" \
-                      \LD_LIBRARY_PATH=\"$KODKODI_JAVA_LIBRARY_PATH:\
-                      \$LD_LIBRARY_PATH\" \
+                bash ("cd " ^ File.shell_quote temp_dir ^ ";\n\
                       \\"$KODKODI\"/bin/kodkodi" ^
                       (if ms >= 0 then " -max-msecs " ^ string_of_int ms
                        else "") ^
--- a/src/HOL/Tools/Nitpick/nitpick.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Tools/Nitpick/nitpick.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -629,7 +629,7 @@
       in
         (pprint (Pretty.chunks
              [Pretty.blk (0,
-                  (pstrs ("Nitpick found a" ^
+                  (pstrs ((if auto then "Auto " else "") ^ "Nitpick found a" ^
                           (if not genuine then " potential "
                            else if genuine_means_genuine then " "
                            else " quasi genuine ") ^ das_wort_model) @
--- a/src/HOL/Tools/Predicate_Compile/etc/settings	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Tools/Predicate_Compile/etc/settings	Sat Oct 30 12:25:18 2010 -0700
@@ -1,13 +1,16 @@
 
 EXEC_SWIPL=$(choosefrom \
-  "$ISABELLE_HOME/contrib/swipl/bin/swipl" \
-  "$ISABELLE_HOME/contrib_devel/swipl/bin/swipl" \
+  "$ISABELLE_HOME/contrib/swipl/$ISABELLE_PLATFORM/bin/swipl" \
+  "$ISABELLE_HOME/contrib_devel/swipl/$ISABELLE_PLATFORM/bin/swipl" \
   "$ISABELLE_HOME/../swipl" \
   $(type -p swipl) \
   "")
 
 EXEC_YAP=$(choosefrom \
-  "$ISABELLE_HOME/contrib/yap" \
+  "$ISABELLE_HOME/contrib/yap/$ISABELLE_PLATFORM/bin/yap" \
+  "$ISABELLE_HOME/contrib_devel/yap/$ISABELLE_PLATFORM/bin/yap" \
   "$ISABELLE_HOME/../yap" \
   $(type -p yap) \
   "")
+
+SWIPL_VERSION=$("$COMPONENT/lib/scripts/swipl_version")
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Predicate_Compile/lib/scripts/swipl_version	Sat Oct 30 12:25:18 2010 -0700
@@ -0,0 +1,12 @@
+#!/usr/bin/env bash
+#
+# Author: Lukas Bulwahn, TU Muenchen, 2010
+#
+# Determine SWI-Prolog version
+
+if [ "$EXEC_SWIPL" = "" ]; then
+  echo ""
+else
+  VERSION=`$EXEC_SWIPL --version`
+  echo `expr match "$VERSION" 'SWI-Prolog version \([0-9\.]*\)'`
+fi
--- a/src/HOL/Tools/Quotient/quotient_info.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Tools/Quotient/quotient_info.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -6,10 +6,11 @@
 
 signature QUOTIENT_INFO =
 sig
-  exception NotFound
+  exception NotFound  (* FIXME complicates signatures unnecessarily *)
 
   type maps_info = {mapfun: string, relmap: string}
   val maps_defined: theory -> string -> bool
+  (* FIXME functions called "lookup" must return option, not raise exception *)
   val maps_lookup: theory -> string -> maps_info     (* raises NotFound *)
   val maps_update_thy: string -> maps_info -> theory -> theory
   val maps_update: string -> maps_info -> Proof.context -> Proof.context
--- a/src/HOL/Tools/Quotient/quotient_term.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Tools/Quotient/quotient_term.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -67,8 +67,8 @@
 fun get_mapfun ctxt s =
 let
   val thy = ProofContext.theory_of ctxt
-  val exn = LIFT_MATCH ("No map function for type " ^ quote s ^ " found.")
-  val mapfun = #mapfun (maps_lookup thy s) handle NotFound => raise exn
+  val mapfun = #mapfun (maps_lookup thy s) handle Quotient_Info.NotFound =>
+    raise LIFT_MATCH ("No map function for type " ^ quote s ^ " found.")
 in
   Const (mapfun, dummyT)
 end
@@ -104,8 +104,8 @@
 fun get_rty_qty ctxt s =
 let
   val thy = ProofContext.theory_of ctxt
-  val exn = LIFT_MATCH ("No quotient type " ^ quote s ^ " found.")
-  val qdata = quotdata_lookup thy s handle NotFound => raise exn
+  val qdata = quotdata_lookup thy s handle Quotient_Info.NotFound =>
+    raise LIFT_MATCH ("No quotient type " ^ quote s ^ " found.")
 in
   (#rtyp qdata, #qtyp qdata)
 end
@@ -127,7 +127,7 @@
   val thy = ProofContext.theory_of ctxt
 in
   Sign.typ_match thy (ty_pat, ty) Vartab.empty
-  handle MATCH_TYPE => err ctxt ty_pat ty
+  handle Type.TYPE_MATCH => err ctxt ty_pat ty
 end
 
 (* produces the rep or abs constant for a qty *)
@@ -276,8 +276,8 @@
 fun get_relmap ctxt s =
 let
   val thy = ProofContext.theory_of ctxt
-  val exn = LIFT_MATCH ("get_relmap (no relation map function found for type " ^ s ^ ")")
-  val relmap = #relmap (maps_lookup thy s) handle NotFound => raise exn
+  val relmap = #relmap (maps_lookup thy s) handle Quotient_Info.NotFound =>
+    raise LIFT_MATCH ("get_relmap (no relation map function found for type " ^ s ^ ")")
 in
   Const (relmap, dummyT)
 end
@@ -299,9 +299,9 @@
 fun get_equiv_rel ctxt s =
 let
   val thy = ProofContext.theory_of ctxt
-  val exn = LIFT_MATCH ("get_quotdata (no quotient found for type " ^ s ^ ")")
 in
-  #equiv_rel (quotdata_lookup thy s) handle NotFound => raise exn
+  #equiv_rel (quotdata_lookup thy s) handle Quotient_Info.NotFound =>
+    raise LIFT_MATCH ("get_quotdata (no quotient found for type " ^ s ^ ")")
 end
 
 fun equiv_match_err ctxt ty_pat ty =
@@ -563,7 +563,8 @@
          else
            let
              val rtrm' = #rconst (qconsts_lookup thy qtrm)
-               handle NotFound => term_mismatch "regularize (constant not found)" ctxt rtrm qtrm
+               handle Quotient_Info.NotFound =>
+                term_mismatch "regularize (constant not found)" ctxt rtrm qtrm
            in
              if Pattern.matches thy (rtrm', rtrm)
              then rtrm else term_mismatch "regularize (constant mismatch)" ctxt rtrm qtrm
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/SMT/smt_builtin.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -0,0 +1,110 @@
+(*  Title:      HOL/Tools/SMT/smt_builtin.ML
+    Author:     Sascha Boehme, TU Muenchen
+
+Crafted collection of SMT built-in symbols.
+
+FIXME: This list is currently not automatically kept synchronized with the
+remaining implementation.
+*)
+
+signature SMT_BUILTIN =
+sig
+  val is_partially_builtin: Proof.context -> string * typ -> bool
+  val is_builtin: Proof.context -> string * typ -> bool
+end
+
+structure SMT_Builtin: SMT_BUILTIN =
+struct
+
+fun is_arithT (Type (@{type_name fun}, [T, _])) =
+      (case T of
+        @{typ int} => true
+      | @{typ nat} => true
+      | Type ("RealDef.real", []) => true
+      | _ => false)
+  | is_arithT _ = false
+
+val builtins = Symtab.make [
+
+  (* Pure constants *)
+  (@{const_name all}, K true),
+  (@{const_name "=="}, K true),
+  (@{const_name "==>"}, K true),
+  (@{const_name Trueprop}, K true),
+
+  (* logical constants *)
+  (@{const_name All}, K true),
+  (@{const_name Ex}, K true),
+  (@{const_name Ball}, K true),
+  (@{const_name Bex}, K true),
+  (@{const_name Ex1}, K true),
+  (@{const_name True}, K true),
+  (@{const_name False}, K true),
+  (@{const_name Not}, K true),
+  (@{const_name HOL.conj}, K true),
+  (@{const_name HOL.disj}, K true),
+  (@{const_name HOL.implies}, K true),
+
+  (* term abbreviations *)
+  (@{const_name If}, K true),
+  (@{const_name bool.bool_case}, K true),
+  (@{const_name Let}, K true),
+  (@{const_name SMT.distinct}, K true),
+
+  (* technicalities *)
+  (@{const_name SMT.pat}, K true),
+  (@{const_name SMT.nopat}, K true),
+  (@{const_name SMT.trigger}, K true),
+  (@{const_name SMT.fun_app}, K true),
+  (@{const_name SMT.z3div}, K true),
+  (@{const_name SMT.z3mod}, K true),
+
+  (* equality *)
+  (@{const_name HOL.eq}, K true),
+
+  (* numerals *)
+  (@{const_name zero_class.zero}, K true),
+  (@{const_name one_class.one}, K true),
+  (@{const_name Int.Pls}, K true),
+  (@{const_name Int.Min}, K true),
+  (@{const_name Int.Bit0}, K true),
+  (@{const_name Int.Bit1}, K true),
+  (@{const_name number_of}, K true),
+
+  (* arithmetic *)
+  (@{const_name nat}, K true),
+  (@{const_name of_nat}, K true),
+  (@{const_name Suc}, K true),
+  (@{const_name plus}, is_arithT),
+  (@{const_name uminus}, is_arithT),
+  (@{const_name minus}, is_arithT),
+  (@{const_name abs}, is_arithT),
+  (@{const_name min}, is_arithT),
+  (@{const_name max}, is_arithT),
+  (@{const_name less}, is_arithT),
+  (@{const_name less_eq}, is_arithT),
+  
+  (* pairs *)
+  (@{const_name fst}, K true),
+  (@{const_name snd}, K true),
+  (@{const_name Pair}, K true)
+
+  ]
+
+val all_builtins =
+  builtins
+  |> Symtab.update (@{const_name times}, is_arithT)
+  |> Symtab.update (@{const_name div_class.div}, is_arithT)
+  |> Symtab.update (@{const_name div_class.mod}, is_arithT)
+  |> Symtab.update ("Rings.inverse_class.divide", is_arithT)
+
+fun lookup_builtin bs (name, T) =
+  (case Symtab.lookup bs name of
+    SOME proper_type => proper_type T
+  | NONE => false)
+
+fun is_partially_builtin _ = lookup_builtin builtins
+
+fun is_builtin _ = lookup_builtin all_builtins
+
+end
--- a/src/HOL/Tools/SMT/smt_normalize.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Tools/SMT/smt_normalize.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -19,7 +19,8 @@
 sig
   type extra_norm = bool -> (int * thm) list -> Proof.context ->
     (int * thm) list * Proof.context
-  val normalize: extra_norm -> bool -> (int * thm) list -> Proof.context ->
+  val normalize: (Proof.context -> (thm -> string) -> thm -> unit) -> bool ->
+    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
@@ -40,15 +41,18 @@
    three elements in the argument list) *)
 
 local
-  fun is_trivial_distinct (Const (@{const_name distinct}, _) $ t) =
-        length (HOLogic.dest_list t) <= 2
+  fun is_trivial_distinct (Const (@{const_name SMT.distinct}, _) $ t) =
+       (length (HOLogic.dest_list t) <= 2
+        handle TERM _ => error ("SMT: constant " ^
+          quote @{const_name SMT.distinct} ^ " must be applied to " ^
+          "an explicit list."))
     | is_trivial_distinct _ = false
 
   val thms = map mk_meta_eq @{lemma
-    "distinct [] = True"
-    "distinct [x] = True"
-    "distinct [x, y] = (x ~= y)"
-    by simp_all}
+    "SMT.distinct [] = True"
+    "SMT.distinct [x] = True"
+    "SMT.distinct [x, y] = (x ~= y)"
+    by (simp_all add: distinct_def)}
   fun distinct_conv _ =
     if_true_conv is_trivial_distinct (Conv.rewrs_conv thms)
 in
@@ -66,11 +70,9 @@
       Const (@{const_name "bool.bool_case"}, _) $ _ $ _ $ _ => true
     | _ => false)
 
-  val thms = map mk_meta_eq @{lemma
-    "(case P of True => x | False => y) = (if P then x else y)"
-    "(case P of False => y | True => x) = (if P then x else y)"
-    by simp_all}
-  val unfold_conv = if_true_conv is_bool_case (Conv.rewrs_conv thms)
+  val thm = mk_meta_eq @{lemma
+    "(case P of True => x | False => y) = (if P then x else y)" by simp}
+  val unfold_conv = if_true_conv is_bool_case (Conv.rewr_conv thm)
 in
 fun rewrite_bool_cases ctxt =
   map (apsnd ((Term.exists_subterm is_bool_case o Thm.prop_of) ??
@@ -127,27 +129,29 @@
   val nat_rewriting = @{lemma
     "0 = nat 0"
     "1 = nat 1"
-    "number_of i = nat (number_of i)"
+    "(number_of :: int => nat) = (%i. nat (number_of i))"
     "int (nat 0) = 0"
     "int (nat 1) = 1"
-    "a < b = (int a < int b)"
-    "a <= b = (int a <= int b)"
-    "Suc a = nat (int a + 1)"
-    "a + b = nat (int a + int b)"
-    "a - b = nat (int a - int b)"
-    "a * b = nat (int a * int b)"
-    "a div b = nat (int a div int b)"
-    "a mod b = nat (int a mod int b)"
-    "min a b = nat (min (int a) (int b))"
-    "max a b = nat (max (int a) (int b))"
+    "op < = (%a b. int a < int b)"
+    "op <= = (%a b. int a <= int b)"
+    "Suc = (%a. nat (int a + 1))"
+    "op + = (%a b. nat (int a + int b))"
+    "op - = (%a b. nat (int a - int b))"
+    "op * = (%a b. nat (int a * int b))"
+    "op div = (%a b. nat (int a div int b))"
+    "op mod = (%a b. nat (int a mod int b))"
+    "min = (%a b. nat (min (int a) (int b)))"
+    "max = (%a b. nat (max (int a) (int b)))"
     "int (nat (int a + int b)) = int a + int b"
+    "int (nat (int a + 1)) = int a + 1"  (* special rule due to Suc above *)
     "int (nat (int a * int b)) = int a * int b"
     "int (nat (int a div int b)) = int a div int b"
     "int (nat (int a mod int b)) = int a mod int b"
     "int (nat (min (int a) (int b))) = min (int a) (int b)"
     "int (nat (max (int a) (int b))) = max (int a) (int b)"
-    by (simp_all add: nat_mult_distrib nat_div_distrib nat_mod_distrib
-      int_mult[symmetric] zdiv_int[symmetric] zmod_int[symmetric])}
+    by (auto intro!: ext simp add: nat_mult_distrib nat_div_distrib
+      nat_mod_distrib int_mult[symmetric] zdiv_int[symmetric]
+      zmod_int[symmetric])}
 
   fun on_positive num f x = 
     (case try HOLogic.dest_number (Thm.term_of num) of
@@ -169,9 +173,10 @@
 
   val nat_ss = HOL_ss
     addsimps nat_rewriting
-    addsimprocs [Simplifier.make_simproc {
-      name = "cancel_int_nat_num", lhss = [@{cpat "int (nat _)"}],
-      proc = cancel_int_nat_simproc, identifier = [] }]
+    addsimprocs [
+      Simplifier.make_simproc {
+        name = "cancel_int_nat_num", lhss = [@{cpat "int (nat _)"}],
+        proc = cancel_int_nat_simproc, identifier = [] }]
 
   fun conv ctxt = Simplifier.rewrite (Simplifier.context ctxt nat_ss)
 
@@ -196,6 +201,14 @@
 local
   val eta_conv = eta_expand_conv
 
+  fun args_conv cv ct =
+    (case Thm.term_of ct of
+      _ $ _ => Conv.combination_conv (args_conv cv) cv
+    | _ => Conv.all_conv) ct
+
+  fun eta_args_conv cv 0 = args_conv o cv
+    | eta_args_conv cv i = eta_conv (eta_args_conv cv (i-1))
+
   fun keep_conv ctxt = Conv.binder_conv (norm_conv o snd) ctxt
   and eta_binder_conv ctxt = Conv.arg_conv (eta_conv norm_conv ctxt)
   and keep_let_conv ctxt = Conv.combination_conv
@@ -227,30 +240,48 @@
     | Const (@{const_name Bex}, _) $ _ => eta_conv unfold_bex_conv
     | Const (@{const_name Bex}, _) => eta_conv (eta_conv unfold_bex_conv)
     | Abs _ => Conv.abs_conv (norm_conv o snd)
-    | _ $ _ => Conv.comb_conv o norm_conv
-    | _ => K Conv.all_conv) ctxt ct
+    | _ =>
+        (case Term.strip_comb (Thm.term_of ct) of
+          (Const (c as (_, T)), ts) =>
+            if SMT_Builtin.is_builtin ctxt c
+            then eta_args_conv norm_conv
+              (length (Term.binder_types T) - length ts)
+            else args_conv o norm_conv
+        | (_, ts) => args_conv o norm_conv)) ctxt ct
 
-  fun is_normed t =
+  fun is_normed ctxt t =
     (case t of
-      Const (@{const_name All}, _) $ Abs (_, _, u) => is_normed u
+      Const (@{const_name All}, _) $ Abs (_, _, u) => is_normed ctxt u
     | Const (@{const_name All}, _) $ _ => false
     | Const (@{const_name All}, _) => false
-    | Const (@{const_name Ex}, _) $ Abs (_, _, u) => is_normed u
+    | Const (@{const_name Ex}, _) $ Abs (_, _, u) => is_normed ctxt u
     | Const (@{const_name Ex}, _) $ _ => false
     | Const (@{const_name Ex}, _) => false
     | Const (@{const_name Let}, _) $ u1 $ Abs (_, _, u2) =>
-        is_normed u1 andalso is_normed u2
+        is_normed ctxt u1 andalso is_normed ctxt u2
     | Const (@{const_name Let}, _) $ _ $ _ => false
     | Const (@{const_name Let}, _) $ _ => false
     | Const (@{const_name Let}, _) => false
+    | Const (@{const_name Ex1}, _) $ _ => false
     | Const (@{const_name Ex1}, _) => false
+    | Const (@{const_name Ball}, _) $ _ $ _ => false
+    | Const (@{const_name Ball}, _) $ _ => false
     | Const (@{const_name Ball}, _) => false
+    | Const (@{const_name Bex}, _) $ _ $ _ => false
+    | Const (@{const_name Bex}, _) $ _ => false
     | Const (@{const_name Bex}, _) => false
-    | Abs (_, _, u) => is_normed u
-    | u1 $ u2 => is_normed u1 andalso is_normed u2
-    | _ => true)
+    | Abs (_, _, u) => is_normed ctxt u
+    | _ =>
+        (case Term.strip_comb t of
+          (Const (c as (_, T)), ts) =>
+            if SMT_Builtin.is_builtin ctxt c
+            then length (Term.binder_types T) = length ts andalso
+              forall (is_normed ctxt) ts
+            else forall (is_normed ctxt) ts
+        | (_, ts) => forall (is_normed ctxt) ts))
 in
-fun norm_binder_conv ctxt = if_conv is_normed Conv.all_conv (norm_conv ctxt)
+fun norm_binder_conv ctxt =
+  if_conv (is_normed ctxt) Conv.all_conv (norm_conv ctxt)
 end
 
 fun norm_def ctxt thm =
@@ -485,18 +516,28 @@
 
 fun with_context f irules ctxt = (f ctxt irules, ctxt)
 
-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_datatypes
-  |-> with_context (fn cx => map (apsnd (normalize_rule cx)))
-  |-> SMT_Monomorph.monomorph
-  |-> lift_lambdas
-  |-> with_context explicit_application
-  |-> (if with_datatypes then datatype_selectors else pair)
+fun normalize trace keep_assms extra_norm with_datatypes irules ctxt =
+  let
+    fun norm f ctxt' (i, thm) =
+      if keep_assms then SOME (i, f ctxt' thm)
+      else
+        (case try (f ctxt') thm of
+          SOME thm' => SOME (i, thm')
+        | NONE => (trace ctxt' (prefix ("SMT warning: " ^
+            "dropping assumption: ") o Display.string_of_thm ctxt') thm; NONE))
+  in
+    irules
+    |> trivial_distinct ctxt
+    |> rewrite_bool_cases ctxt
+    |> normalize_numerals ctxt
+    |> nat_as_int ctxt
+    |> rpair ctxt
+    |-> extra_norm with_datatypes
+    |-> with_context (map_filter o norm normalize_rule)
+    |-> SMT_Monomorph.monomorph
+    |-> lift_lambdas
+    |-> with_context explicit_application
+    |-> (if with_datatypes then datatype_selectors else pair)
+  end
 
 end
--- a/src/HOL/Tools/SMT/smt_setup_solvers.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Tools/SMT/smt_setup_solvers.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -17,7 +17,9 @@
   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."))
+    quote solver_name ^ " failed. Enable SMT tracing by setting the " ^
+    "configuration option " ^ quote SMT_Solver.traceN ^ " and " ^
+    " see the trace for details."))
 
 fun on_first_line test_outcome solver_name lines =
   let
--- a/src/HOL/Tools/SMT/smt_solver.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Tools/SMT/smt_solver.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -34,8 +34,10 @@
   val oracle: bool Config.T
   val filter_only: bool Config.T
   val datatypes: bool Config.T
+  val keep_assms: bool Config.T
   val timeout: int Config.T
   val with_timeout: Proof.context -> ('a -> 'b) -> 'a -> 'b
+  val traceN: string
   val trace: bool Config.T
   val trace_msg: Proof.context -> ('a -> string) -> 'a -> unit
   val trace_used_facts: bool Config.T
@@ -122,13 +124,17 @@
 
 val (datatypes, setup_datatypes) = Attrib.config_bool "smt_datatypes" (K false)
 
+val (keep_assms, setup_keep_assms) =
+  Attrib.config_bool "smt_keep_assms" (K true)
+
 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 Time_Out
 
-val (trace, setup_trace) = Attrib.config_bool "smt_trace" (K false)
+val traceN = "smt_trace"
+val (trace, setup_trace) = Attrib.config_bool traceN (K false)
 
 fun trace_msg ctxt f x =
   if Config.get ctxt trace then tracing (f x) else ()
@@ -292,9 +298,10 @@
     val (with_datatypes, translate') =
       set_has_datatypes (Config.get ctxt datatypes) translate
     val cmd = (rm, env_var, is_remote, name)
+    val keep = Config.get ctxt keep_assms
   in
     (irules, ctxt)
-    |-> SMT_Normalize.normalize extra_norm with_datatypes
+    |-> SMT_Normalize.normalize trace_msg keep extra_norm with_datatypes
     |-> invoke translate' name cmd more_options options
     |-> reconstruct
     |-> (fn (idxs, thm) => fn ctxt' => thm
@@ -429,6 +436,7 @@
       |> Config.put timeout (Time.toSeconds time_limit)
       |> Config.put oracle false
       |> Config.put filter_only true
+      |> Config.put keep_assms false
     val cprop =
       Thm.cprem_of goal i
       |> Thm.rhs_of o SMT_Normalize.atomize_conv ctxt
@@ -481,6 +489,7 @@
   setup_oracle #>
   setup_filter_only #>
   setup_datatypes #>
+  setup_keep_assms #>
   setup_timeout #>
   setup_trace #>
   setup_trace_used_facts #>
--- a/src/HOL/Tools/SMT/smt_translate.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Tools/SMT/smt_translate.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -233,7 +233,7 @@
         (q as Const (qn, _), [Abs (n, T, t')]) =>
           if is_some (quantifier qn) then q $ Abs (n, T, in_trig t')
           else as_term (in_term t)
-      | (Const (c as (@{const_name distinct}, T)), [t']) =>
+      | (Const (c as (@{const_name SMT.distinct}, T)), [t']) =>
           if is_builtin_distinct then Const (pred c) $ in_list T in_term t'
           else as_term (in_term t)
       | (Const c, ts) =>
@@ -377,7 +377,7 @@
       | (Const (@{const_name Let}, _), [t1, Abs (_, T, t2)]) =>
           transT T ##>> trans t1 ##>> trans t2 #>>
           (fn ((U, u1), u2) => SLet (U, u1, u2))
-      | (h as Const (c as (@{const_name distinct}, T)), [t1]) =>
+      | (h as Const (c as (@{const_name SMT.distinct}, T)), [t1]) =>
           (case builtin_fun ctxt c (HOLogic.dest_list t1) of
             SOME (n, ts) => fold_map trans ts #>> app n
           | NONE => transs h T [t1])
--- a/src/HOL/Tools/SMT/smtlib_interface.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Tools/SMT/smtlib_interface.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -167,7 +167,7 @@
   | conn @{const_name If} = SOME "if_then_else"
   | conn _ = NONE
 
-fun pred @{const_name distinct} _ = SOME "distinct"
+fun pred @{const_name SMT.distinct} _ = SOME "distinct"
   | pred @{const_name HOL.eq} _ = SOME "="
   | pred @{const_name term_eq} _ = SOME "="
   | pred @{const_name less} T = if_int_type T "<"
--- a/src/HOL/Tools/SMT/z3_interface.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Tools/SMT/z3_interface.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -65,7 +65,7 @@
 (** interface **)
 
 local
-  val {extra_norm, translate} = SMTLIB_Interface.interface
+  val {translate, ...} = SMTLIB_Interface.interface
   val {prefixes, strict, header, builtins, serialize} = translate
   val {is_builtin_pred, ...}= the strict
   val {builtin_typ, builtin_num, builtin_fun, ...} = builtins
@@ -189,7 +189,7 @@
 fun mk_list cT cts =
   fold_rev (Thm.mk_binop (instT cT cons_term)) cts (instT cT nil_term)
 
-val distinct = mk_inst_pair (destT1 o destT1) @{cpat distinct}
+val distinct = mk_inst_pair (destT1 o destT1) @{cpat SMT.distinct}
 fun mk_distinct [] = mk_true
   | mk_distinct (cts as (ct :: _)) =
       Thm.capply (instT' ct distinct) (mk_list (Thm.ctyp_of_term ct) cts)
--- a/src/HOL/Tools/SMT/z3_proof_reconstruction.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Tools/SMT/z3_proof_reconstruction.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -303,14 +303,14 @@
           prove disjI3 (Thm.capply @{cterm Not} ct2, false) (ct1, false)
       | @{term HOL.disj} $ _ $ _ =>
           prove disjI2 (Thm.capply @{cterm Not} ct1, false) (ct2, true)
-      | Const (@{const_name distinct}, _) $ _ =>
+      | Const (@{const_name SMT.distinct}, _) $ _ =>
           let
             fun dis_conv cv = Conv.arg_conv (Conv.arg1_conv cv)
             fun prv cu =
               let val (cu1, cu2) = Thm.dest_binop (Thm.dest_arg cu)
               in prove disjI4 (Thm.dest_arg cu2, true) (cu1, true) end
           in T.with_conv (dis_conv T.unfold_distinct_conv) prv ct end
-      | @{term Not} $ (Const (@{const_name distinct}, _) $ _) =>
+      | @{term Not} $ (Const (@{const_name SMT.distinct}, _) $ _) =>
           let
             fun dis_conv cv = Conv.arg_conv (Conv.arg1_conv (Conv.arg_conv cv))
             fun prv cu =
@@ -476,7 +476,7 @@
     (case Term.head_of (Thm.term_of cl) of
       @{term HOL.conj} => prove_nary L.is_conj (prove_eq_exn f)
     | @{term HOL.disj} => prove_nary L.is_disj (prove_eq_exn f)
-    | Const (@{const_name distinct}, _) => prove_distinct (prove_eq_safe f)
+    | Const (@{const_name SMT.distinct}, _) => prove_distinct (prove_eq_safe f)
     | _ => prove (prove_eq_safe f)) cp
 in
 fun monotonicity eqs ct =
--- a/src/HOL/Tools/SMT/z3_proof_tools.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Tools/SMT/z3_proof_tools.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -204,7 +204,7 @@
       | @{term HOL.disj} $ _ $ _ => abstr2 cvs ct
       | @{term HOL.implies} $ _ $ _ => abstr2 cvs ct
       | Const (@{const_name HOL.eq}, _) $ _ $ _ => abstr2 cvs ct
-      | Const (@{const_name distinct}, _) $ _ =>
+      | Const (@{const_name SMT.distinct}, _) $ _ =>
           if ext_logic then abs_arg (abs_list abstr fresh_abstraction) cvs ct
           else fresh_abstraction cvs ct
       | Const (@{const_name If}, _) $ _ $ _ $ _ =>
@@ -263,10 +263,10 @@
     (Conv.rewrs_conv [set1, set2, set3, set4] else_conv
     (Conv.rewr_conv set5 then_conv Conv.arg_conv set_conv)) ct
 
-  val dist1 = @{lemma "distinct [] == ~False" by simp}
-  val dist2 = @{lemma "distinct [x] == ~False" by simp}
-  val dist3 = @{lemma "distinct (x # xs) == x ~: set xs & distinct xs"
-    by simp}
+  val dist1 = @{lemma "SMT.distinct [] == ~False" by (simp add: distinct_def)}
+  val dist2 = @{lemma "SMT.distinct [x] == ~False" by (simp add: distinct_def)}
+  val dist3 = @{lemma "SMT.distinct (x # xs) == x ~: set xs & distinct xs"
+    by (simp add: distinct_def)}
 
   fun binop_conv cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2
 in
--- a/src/HOL/Tools/Sledgehammer/sledgehammer.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -266,7 +266,7 @@
      "\n" ^ Syntax.string_of_term ctxt (Thm.term_of (Thm.cprem_of goal i)))
 
 fun proof_banner auto =
-  if auto then "Sledgehammer found a proof" else "Try this command"
+  if auto then "Auto Sledgehammer found a proof" else "Try this command"
 
 (* generic TPTP-based ATPs *)
 
@@ -399,7 +399,7 @@
     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
+      if not auto andalso random () <= important_message_keep_factor then
         extract_important_message output
       else
         ""
@@ -432,7 +432,7 @@
   | failure_from_smt_failure SMT_Solver.Out_Of_Memory = OutOfResources
   | failure_from_smt_failure _ = UnknownError
 
-fun run_smt_solver remote ({timeout, ...} : params) minimize_command
+fun run_smt_solver auto remote ({timeout, ...} : params) minimize_command
                    ({state, subgoal, subgoal_count, facts, ...}
                     : prover_problem) =
   let
@@ -443,7 +443,7 @@
     val message =
       case outcome of
         NONE =>
-        try_command_line (proof_banner false)
+        try_command_line (proof_banner auto)
                          (apply_on_subgoal subgoal subgoal_count ^
                           command_call smtN (map fst used_facts)) ^
         minimize_line minimize_command (map fst used_facts)
@@ -458,8 +458,10 @@
   end
 
 fun get_prover thy auto name =
-  if is_smt_prover name then run_smt_solver (String.isPrefix remote_prefix name)
-  else run_atp auto name (get_atp thy name)
+  if is_smt_prover name then
+    run_smt_solver auto (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
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_filter.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_filter.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -86,8 +86,12 @@
 val skolem_prefix = sledgehammer_prefix ^ "sko"
 val theory_const_suffix = Long_Name.separator ^ " 1"
 
+fun needs_quoting reserved s =
+  Symtab.defined reserved s (* FIXME: orelse
+  exists (not o Syntax.is_identifier) (Long_Name.explode s) *)
+
 fun repair_name reserved multi j name =
-  (name |> Symtab.defined reserved name ? quote) ^
+  (name |> needs_quoting reserved name ? quote) ^
   (if multi then "(" ^ Int.toString j ^ ")" else "")
 
 fun fact_from_ref ctxt reserved chained_ths (xthm as (xref, args)) =
--- a/src/HOL/Tools/sat_funcs.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Tools/sat_funcs.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -274,22 +274,6 @@
   | string_of_prop_formula (PropLogic.And (fm1, fm2)) = "(" ^ string_of_prop_formula fm1 ^ " & " ^ string_of_prop_formula fm2 ^ ")";
 
 (* ------------------------------------------------------------------------- *)
-(* take_prefix:                                                              *)
-(*      take_prefix n [x_1, ..., x_k] = ([x_1, ..., x_n], [x_n+1, ..., x_k]) *)
-(* ------------------------------------------------------------------------- *)
-
-(* int -> 'a list -> 'a list * 'a list *)
-
-fun take_prefix n xs =
-let
-	fun take 0 (rxs, xs)      = (rev rxs, xs)
-	  | take _ (rxs, [])      = (rev rxs, [])
-	  | take n (rxs, x :: xs) = take (n-1) (x :: rxs, xs)
-in
-	take n ([], xs)
-end;
-
-(* ------------------------------------------------------------------------- *)
 (* rawsat_thm: run external SAT solver with the given clauses.  Reconstructs *)
 (*      a proof from the resulting proof trace of the SAT solver.  The       *)
 (*      theorem returned is just "False" (with some of the given clauses as  *)
--- a/src/HOL/Tools/try.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/Tools/try.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -2,33 +2,23 @@
     Author:     Jasmin Blanchette, TU Muenchen
 
 Try a combination of proof methods.
-
-FIXME: reintroduce or remove commented code (see also HOL.thy)
 *)
 
 signature TRY =
 sig
-(*
   val auto : bool Unsynchronized.ref
-*)
   val invoke_try : Time.time option -> Proof.state -> bool
-(*
   val setup : theory -> theory
-*)
 end;
 
 structure Try : TRY =
 struct
 
-(*
 val auto = Unsynchronized.ref false
 
 val _ =
   ProofGeneralPgip.add_preference Preferences.category_tracing
-  (Unsynchronized.setmp auto true (fn () =>
-    Preferences.bool_pref auto
-      "auto-try" "Try standard proof methods.") ());
-*)
+      (Preferences.bool_pref auto "auto-try" "Try standard proof methods.")
 
 val default_timeout = Time.fromSeconds 5
 
@@ -57,26 +47,37 @@
   let val thy = ProofContext.theory_of ctxt in
     Proof.refine (Method.SelectGoals (1, Method.Basic (named_method thy name)))
   end
+  handle ERROR _ => K Seq.empty (* e.g., the method isn't available yet *)
 
-fun do_named_method (name, all_goals) timeout_opt st =
-  do_generic timeout_opt
-             (name ^ (if all_goals andalso
-                         nprems_of (#goal (Proof.goal st)) > 1 then
-                        "[1]"
-                      else
-                        "")) I (#goal o Proof.goal)
-             (apply_named_method_on_first_goal name (Proof.context_of st)) st
+fun do_named_method (name, (all_goals, run_if_auto)) auto timeout_opt st =
+  if not auto orelse run_if_auto then
+    do_generic timeout_opt
+               (name ^ (if all_goals andalso
+                           nprems_of (#goal (Proof.goal st)) > 1 then
+                          "[1]"
+                        else
+                          "")) I (#goal o Proof.goal)
+               (apply_named_method_on_first_goal name (Proof.context_of st)) st
+  else
+    NONE
 
+(* name * (all_goals, run_if_auto) *)
 val named_methods =
-  [("simp", false), ("auto", true), ("fast", false), ("fastsimp", false),
-   ("force", false), ("blast", false), ("metis", false), ("linarith", false),
-   ("presburger", false)]
+  [("simp", (false, true)),
+   ("auto", (true, true)),
+   ("fast", (false, false)),
+   ("fastsimp", (false, false)),
+   ("force", (false, false)),
+   ("blast", (false, true)),
+   ("metis", (false, true)),
+   ("linarith", (false, true)),
+   ("presburger", (false, true))]
 val do_methods = map do_named_method named_methods
 
 fun time_string (s, ms) = s ^ ": " ^ string_of_int ms ^ " ms"
 
 fun do_try auto timeout_opt st =
-  case do_methods |> Par_List.map (fn f => f timeout_opt st)
+  case do_methods |> Par_List.map (fn f => f auto timeout_opt st)
                   |> map_filter I |> sort (int_ord o pairself snd) of
     [] => (if auto then () else writeln "No proof found."; (false, st))
   | xs as (s, _) :: _ =>
@@ -88,7 +89,7 @@
         Markup.markup Markup.sendback
             ((if nprems_of (#goal (Proof.goal st)) = 1 then "by" else "apply") ^
              " " ^ s) ^
-        ".\n(" ^ space_implode "; " (map time_string xs) ^ ")\n"
+        "\n(" ^ space_implode "; " (map time_string xs) ^ ").\n"
     in
       (true, st |> (if auto then
                       Proof.goal_message
@@ -109,10 +110,8 @@
       (Scan.succeed (Toplevel.keep (K () o do_try false (SOME default_timeout)
                                     o Toplevel.proof_of)))
 
-(*
 fun auto_try st = if not (!auto) then (false, st) else do_try true NONE st
 
 val setup = Auto_Tools.register_tool (tryN, auto_try)
-*)
 
 end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/Coercion_Examples.thy	Sat Oct 30 12:25:18 2010 -0700
@@ -0,0 +1,180 @@
+(*  Title:      HOL/ex/Coercion_Examples.thy
+    Author:     Dmitriy Traytel, TU Muenchen
+
+Examples for coercive subtyping via subtype constraints.
+*)
+
+theory Coercion_Examples
+imports Main
+uses "~~/src/Tools/subtyping.ML"
+begin
+
+setup Subtyping.setup
+
+(* Coercion/type maps definitions*)
+
+consts func :: "(nat \<Rightarrow> int) \<Rightarrow> nat"
+consts arg :: "int \<Rightarrow> nat"
+(* Invariant arguments
+term "func arg"
+*)
+(* No subtype relation - constraint
+term "(1::nat)::int"
+*)
+consts func' :: "int \<Rightarrow> int"
+consts arg' :: "nat"
+(* No subtype relation - function application
+term "func' arg'"
+*)
+(* Uncomparable types in bound
+term "arg' = True"
+*)
+(* Unfullfilled type class requirement
+term "1 = True"
+*)
+(* Different constructors
+term "[1::int] = func"
+*)
+
+primrec nat_of_bool :: "bool \<Rightarrow> nat"
+where
+  "nat_of_bool False = 0"
+| "nat_of_bool True = 1"
+
+declare [[coercion nat_of_bool]]
+
+declare [[coercion int]]
+
+declare [[map_function map]]
+
+definition map_fun :: "('a \<Rightarrow> 'b) \<Rightarrow> ('c \<Rightarrow> 'd) \<Rightarrow> ('b \<Rightarrow> 'c) \<Rightarrow> ('a \<Rightarrow> 'd)" where
+ "map_fun f g h = g o h o f"
+
+declare [[map_function "\<lambda> f g h . g o h o f"]]
+
+primrec map_pair :: "('a \<Rightarrow> 'c) \<Rightarrow> ('b \<Rightarrow> 'd) \<Rightarrow> ('a * 'b) \<Rightarrow> ('c * 'd)" where
+ "map_pair f g (x,y) = (f x, g y)"
+
+declare [[map_function map_pair]]
+
+(* Examples taken from the haskell draft implementation *)
+
+term "(1::nat) = True"
+
+term "True = (1::nat)"
+
+term "(1::nat) = (True = (1::nat))"
+
+term "op = (True = (1::nat))"
+
+term "[1::nat,True]"
+
+term "[True,1::nat]"
+
+term "[1::nat] = [True]"
+
+term "[True] = [1::nat]"
+
+term "[[True]] = [[1::nat]]"
+
+term "[[[[[[[[[[True]]]]]]]]]] = [[[[[[[[[[1::nat]]]]]]]]]]"
+
+term "[[True],[42::nat]] = rev [[True]]"
+
+term "rev [10000::nat] = [False, 420000::nat, True]"
+
+term "\<lambda> x . x = (3::nat)"
+
+term "(\<lambda> x . x = (3::nat)) True"
+
+term "map (\<lambda> x . x = (3::nat))"
+
+term "map (\<lambda> x . x = (3::nat)) [True,1::nat]"
+
+consts bnn :: "(bool \<Rightarrow> nat) \<Rightarrow> nat"
+consts nb :: "nat \<Rightarrow> bool"
+consts ab :: "'a \<Rightarrow> bool"
+
+term "bnn nb"
+
+term "bnn ab"
+
+term "\<lambda> x . x = (3::int)"
+
+term "map (\<lambda> x . x = (3::int)) [True]"
+
+term "map (\<lambda> x . x = (3::int)) [True,1::nat]"
+
+term "map (\<lambda> x . x = (3::int)) [True,1::nat,1::int]"
+
+term "[1::nat,True,1::int,False]"
+
+term "map (map (\<lambda> x . x = (3::int))) [[True],[1::nat],[True,1::int]]"
+
+consts cbool :: "'a \<Rightarrow> bool"
+consts cnat :: "'a \<Rightarrow> nat"
+consts cint :: "'a \<Rightarrow> int"
+
+term "[id, cbool, cnat, cint]"
+
+consts funfun :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
+consts flip :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'c"
+
+term "flip funfun"
+
+term "map funfun [id,cnat,cint,cbool]"
+
+term "map (flip funfun True)"
+
+term "map (flip funfun True) [id,cnat,cint,cbool]"
+
+consts ii :: "int \<Rightarrow> int"
+consts aaa :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
+consts nlist :: "nat list"
+consts ilil :: "int list \<Rightarrow> int list"
+
+term "ii (aaa (1::nat) True)"
+
+term "map ii nlist"
+
+term "ilil nlist"
+
+(***************************************************)
+
+(* Other examples *)
+
+definition xs :: "bool list" where "xs = [True]"
+
+term "(xs::nat list)"
+
+term "(1::nat) = True"
+
+term "True = (1::nat)"
+
+term "int (1::nat)"
+
+term "((True::nat)::int)"
+
+term "1::nat"
+
+term "nat 1"
+
+definition C :: nat
+where "C = 123"
+
+consts g :: "int \<Rightarrow> int"
+consts h :: "nat \<Rightarrow> nat"
+
+term "(g (1::nat)) + (h 2)"
+
+term "g 1"
+
+term "1+(1::nat)"
+
+term "((1::int) + (1::nat),(1::int))"
+
+definition ys :: "bool list list list list list" where "ys=[[[[[True]]]]]"
+
+term "ys=[[[[[1::nat]]]]]"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/Iff_Oracle.thy	Sat Oct 30 12:25:18 2010 -0700
@@ -0,0 +1,76 @@
+(*  Title:      HOL/ex/Iff_Oracle.thy
+    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
+    Author:     Makarius
+*)
+
+header {* Example of Declaring an Oracle *}
+
+theory Iff_Oracle
+imports Main
+begin
+
+subsection {* Oracle declaration *}
+
+text {*
+  This oracle makes tautologies of the form @{prop "P <-> P <-> P <-> P"}.
+  The length is specified by an integer, which is checked to be even
+  and positive.
+*}
+
+oracle iff_oracle = {*
+  let
+    fun mk_iff 1 = Var (("P", 0), @{typ bool})
+      | mk_iff n = HOLogic.mk_eq (Var (("P", 0), @{typ bool}), mk_iff (n - 1));
+  in
+    fn (thy, n) =>
+      if n > 0 andalso n mod 2 = 0
+      then Thm.cterm_of thy (HOLogic.mk_Trueprop (mk_iff n))
+      else raise Fail ("iff_oracle: " ^ string_of_int n)
+  end
+*}
+
+
+subsection {* Oracle as low-level rule *}
+
+ML {* iff_oracle (@{theory}, 2) *}
+ML {* iff_oracle (@{theory}, 10) *}
+ML {* Thm.status_of (iff_oracle (@{theory}, 10)) *}
+
+text {* These oracle calls had better fail. *}
+
+ML {*
+  (iff_oracle (@{theory}, 5); error "Bad oracle")
+    handle Fail _ => warning "Oracle failed, as expected"
+*}
+
+ML {*
+  (iff_oracle (@{theory}, 1); error "Bad oracle")
+    handle Fail _ => warning "Oracle failed, as expected"
+*}
+
+
+subsection {* Oracle as proof method *}
+
+method_setup iff = {*
+  Scan.lift Parse.nat >> (fn n => fn ctxt =>
+    SIMPLE_METHOD
+      (HEADGOAL (rtac (iff_oracle (ProofContext.theory_of ctxt, n)))
+        handle Fail _ => no_tac))
+*} "iff oracle"
+
+
+lemma "A <-> A"
+  by (iff 2)
+
+lemma "A <-> A <-> A <-> A <-> A <-> A <-> A <-> A <-> A <-> A"
+  by (iff 10)
+
+lemma "A <-> A <-> A <-> A <-> A"
+  apply (iff 5)?
+  oops
+
+lemma A
+  apply (iff 1)?
+  oops
+
+end
--- a/src/HOL/ex/ROOT.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/HOL/ex/ROOT.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -12,6 +12,8 @@
 ];
 
 use_thys [
+  "Iff_Oracle",
+  "Coercion_Examples",
   "Numeral",
   "Higher_Order_Logic",
   "Abstract_NAT",
--- a/src/Pure/Concurrent/simple_thread.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/Pure/Concurrent/simple_thread.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -32,21 +32,21 @@
 fun synchronized name lock e =
   if Multithreading.available then
     Exn.release (uninterruptible (fn restore_attributes => fn () =>
-    let
-      val immediate =
-        if Mutex.trylock lock then true
-        else
-          let
-            val _ = Multithreading.tracing 5 (fn () => name ^ ": locking ...");
-            val time = Multithreading.real_time Mutex.lock lock;
-            val _ = Multithreading.tracing_time true time
-              (fn () => name ^ ": locked after " ^ Time.toString time);
-          in false end;
-      val result = Exn.capture (restore_attributes e) ();
-      val _ =
-        if immediate then () else Multithreading.tracing 5 (fn () => name ^ ": unlocking ...");
-      val _ = Mutex.unlock lock;
-    in result end) ())
+      let
+        val immediate =
+          if Mutex.trylock lock then true
+          else
+            let
+              val _ = Multithreading.tracing 5 (fn () => name ^ ": locking ...");
+              val time = Multithreading.real_time Mutex.lock lock;
+              val _ = Multithreading.tracing_time true time
+                (fn () => name ^ ": locked after " ^ Time.toString time);
+            in false end;
+        val result = Exn.capture (restore_attributes e) ();
+        val _ =
+          if immediate then () else Multithreading.tracing 5 (fn () => name ^ ": unlocking ...");
+        val _ = Mutex.unlock lock;
+      in result end) ())
   else e ();
 
 end;
--- a/src/Pure/General/exn.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/Pure/General/exn.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -11,12 +11,13 @@
   val get_exn: 'a result -> exn option
   val capture: ('a -> 'b) -> 'a -> 'b result
   val release: 'a result -> 'a
-  val map_result: ('a -> 'a) -> 'a result -> 'a result
+  val map_result: ('a -> 'b) -> 'a result -> 'b result
   exception Interrupt
   val interrupt: unit -> 'a
   val is_interrupt: exn -> bool
   val interrupt_exn: 'a result
   val is_interrupt_exn: 'a result -> bool
+  val interruptible_capture: ('a -> 'b) -> 'a -> 'b result
   exception EXCEPTIONS of exn list
   val flatten: exn -> exn list
   val flatten_list: exn list -> exn list
@@ -45,7 +46,7 @@
   | release (Exn e) = reraise e;
 
 fun map_result f (Result x) = Result (f x)
-  | map_result f e = e;
+  | map_result f (Exn e) = (Exn e);
 
 
 (* interrupts *)
@@ -55,7 +56,7 @@
 fun interrupt () = raise Interrupt;
 
 fun is_interrupt Interrupt = true
-  | is_interrupt (IO.Io {cause = Interrupt, ...}) = true
+  | is_interrupt (IO.Io {cause, ...}) = is_interrupt cause
   | is_interrupt _ = false;
 
 val interrupt_exn = Exn Interrupt;
@@ -63,6 +64,9 @@
 fun is_interrupt_exn (Exn exn) = is_interrupt exn
   | is_interrupt_exn _ = false;
 
+fun interruptible_capture f x =
+  Result (f x) handle e => if is_interrupt e then reraise e else Exn e;
+
 
 (* nested exceptions *)
 
--- a/src/Pure/General/scan.scala	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/Pure/General/scan.scala	Sat Oct 30 12:25:18 2010 -0700
@@ -262,6 +262,7 @@
     {
       val id = one(symbols.is_letter) ~ many(symbols.is_letdig) ^^ { case x ~ y => x + y }
       val nat = many1(symbols.is_digit)
+      val natdot = nat ~ "." ~ nat ^^ { case x ~ y ~ z => x + y + z }
       val id_nat = id ~ opt("." ~ nat) ^^ { case x ~ Some(y ~ z) => x + y + z case x ~ None => x }
 
       val ident = id ~ rep("." ~> id) ^^
@@ -272,6 +273,8 @@
       val type_ident = "'" ~ id ^^ { case x ~ y => Token(Token.Kind.TYPE_IDENT, x + y) }
       val type_var = "?'" ~ id_nat ^^ { case x ~ y => Token(Token.Kind.TYPE_VAR, x + y) }
       val nat_ = nat ^^ (x => Token(Token.Kind.NAT, x))
+      val float =
+        ("-" ~ natdot ^^ { case x ~ y => x + y } | natdot) ^^ (x => Token(Token.Kind.FLOAT, x))
 
       val sym_ident =
         (many1(symbols.is_symbolic_char) |
@@ -294,7 +297,7 @@
       (space | (string | (alt_string | (verbatim ^^ (x => Token(Token.Kind.VERBATIM, x)) |
         comment ^^ (x => Token(Token.Kind.COMMENT, x)))))) |
       bad_delimiter |
-      ((ident | (var_ | (type_ident | (type_var | (nat_ | sym_ident))))) |||
+      ((ident | (var_ | (type_ident | (type_var | (float | (nat_ | sym_ident)))))) |||
         keyword ^^ (x => Token(if (is_command(x)) Token.Kind.COMMAND else Token.Kind.KEYWORD, x))) |
       bad
     }
--- a/src/Pure/IsaMakefile	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/Pure/IsaMakefile	Sat Oct 30 12:25:18 2010 -0700
@@ -20,6 +20,7 @@
 ## Pure
 
 BOOTSTRAP_FILES = 					\
+  General/exn.ML					\
   ML-Systems/bash.ML 					\
   ML-Systems/compiler_polyml-5.2.ML			\
   ML-Systems/compiler_polyml-5.3.ML			\
@@ -73,7 +74,6 @@
   General/basics.ML					\
   General/binding.ML					\
   General/buffer.ML					\
-  General/exn.ML					\
   General/file.ML					\
   General/graph.ML					\
   General/heap.ML					\
--- a/src/Pure/Isar/args.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/Pure/Isar/args.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -224,7 +224,7 @@
   let
     val keyword_symid = token (Parse.keyword_with is_symid);
     fun atom blk = Parse.group "argument"
-      (ident || keyword_symid || string || alt_string ||
+      (ident || keyword_symid || string || alt_string || token Parse.float_number ||
         (if blk then token (Parse.$$$ ",") else Scan.fail));
 
     fun args blk x = Scan.optional (args1 blk) [] x
--- a/src/Pure/Isar/attrib.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/Pure/Isar/attrib.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -38,10 +38,13 @@
   val internal: (morphism -> attribute) -> src
   val config_bool: bstring -> (Context.generic -> bool) -> bool Config.T * (theory -> theory)
   val config_int: bstring -> (Context.generic -> int) -> int Config.T * (theory -> theory)
+  val config_real: bstring -> (Context.generic -> real) -> real Config.T * (theory -> theory)
   val config_string: bstring -> (Context.generic -> string) -> string Config.T * (theory -> theory)
   val config_bool_global: bstring -> (Context.generic -> bool) -> bool Config.T * (theory -> theory)
   val config_int_global: bstring -> (Context.generic -> int) -> int Config.T * (theory -> theory)
-  val config_string_global: bstring -> (Context.generic -> string) -> string Config.T * (theory -> theory)
+  val config_real_global: bstring -> (Context.generic -> real) -> real Config.T * (theory -> theory)
+  val config_string_global:
+    bstring -> (Context.generic -> string) -> string Config.T * (theory -> theory)
 end;
 
 structure Attrib: ATTRIB =
@@ -353,6 +356,7 @@
       equals -- Args.$$$ "true" >> K (Config.Bool true) ||
       Scan.succeed (Config.Bool true)
   | scan_value (Config.Int _) = equals |-- Parse.int >> Config.Int
+  | scan_value (Config.Real _) = equals |-- Parse.real >> Config.Real
   | scan_value (Config.String _) = equals |-- Args.name >> Config.String;
 
 fun scan_config thy config =
@@ -380,10 +384,12 @@
 
 val config_bool   = declare_config Config.Bool Config.bool false;
 val config_int    = declare_config Config.Int Config.int false;
+val config_real   = declare_config Config.Real Config.real false;
 val config_string = declare_config Config.String Config.string false;
 
 val config_bool_global   = declare_config Config.Bool Config.bool true;
 val config_int_global    = declare_config Config.Int Config.int true;
+val config_real_global   = declare_config Config.Real Config.real true;
 val config_string_global = declare_config Config.String Config.string true;
 
 end;
--- a/src/Pure/Isar/local_defs.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/Pure/Isar/local_defs.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -44,8 +44,9 @@
 
 fun cert_def ctxt eq =
   let
-    fun err msg = cat_error msg ("The error(s) above occurred in definition: " ^
-      quote (Syntax.string_of_term ctxt eq));
+    fun err msg =
+      cat_error msg ("The error(s) above occurred in definition:\n" ^
+        quote (Syntax.string_of_term ctxt eq));
     val ((lhs, _), eq') = eq
       |> Sign.no_vars ctxt
       |> Primitive_Defs.dest_def ctxt Term.is_Free (Variable.is_fixed ctxt) (K true)
@@ -245,7 +246,7 @@
           (CONVERSION (meta_rewrite_conv ctxt') THEN'
             MetaSimplifier.rewrite_goal_tac [def] THEN'
             resolve_tac [Drule.reflexive_thm])))
-        handle ERROR msg => cat_error msg "Failed to prove definitional specification."
+        handle ERROR msg => cat_error msg "Failed to prove definitional specification"
       end;
   in (((c, T), rhs), prove) end;
 
--- a/src/Pure/Isar/parse.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/Pure/Isar/parse.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -26,6 +26,7 @@
   val type_ident: string parser
   val type_var: string parser
   val number: string parser
+  val float_number: string parser
   val string: string parser
   val alt_string: string parser
   val verbatim: string parser
@@ -46,6 +47,7 @@
   val opt_begin: bool parser
   val nat: int parser
   val int: int parser
+  val real: real parser
   val enum: string -> 'a parser -> 'a list parser
   val enum1: string -> 'a parser -> 'a list parser
   val and_list: 'a parser -> 'a list parser
@@ -168,6 +170,7 @@
 val type_ident = kind Token.TypeIdent;
 val type_var = kind Token.TypeVar;
 val number = kind Token.Nat;
+val float_number = kind Token.Float;
 val string = kind Token.String;
 val alt_string = kind Token.AltString;
 val verbatim = kind Token.Verbatim;
@@ -192,6 +195,7 @@
 
 val nat = number >> (#1 o Library.read_int o Symbol.explode);
 val int = Scan.optional (minus >> K ~1) 1 -- nat >> op *;
+val real = float_number >> (the o Real.fromString);
 
 val tag_name = group "tag name" (short_ident || string);
 val tags = Scan.repeat ($$$ "%" |-- !!! tag_name);
--- a/src/Pure/Isar/token.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/Pure/Isar/token.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -8,7 +8,7 @@
 sig
   datatype kind =
     Command | Keyword | Ident | LongIdent | SymIdent | Var | TypeIdent | TypeVar |
-    Nat | String | AltString | Verbatim | Space | Comment | InternalValue |
+    Nat | Float | String | AltString | Verbatim | Space | Comment | InternalValue |
     Malformed | Error of string | Sync | EOF
   datatype value =
     Text of string | Typ of typ | Term of term | Fact of thm list |
@@ -89,7 +89,7 @@
 
 datatype kind =
   Command | Keyword | Ident | LongIdent | SymIdent | Var | TypeIdent | TypeVar |
-  Nat | String | AltString | Verbatim | Space | Comment | InternalValue |
+  Nat | Float | String | AltString | Verbatim | Space | Comment | InternalValue |
   Malformed | Error of string | Sync | EOF;
 
 datatype T = Token of (Symbol_Pos.text * Position.range) * (kind * string) * slot;
@@ -103,7 +103,8 @@
   | Var => "schematic variable"
   | TypeIdent => "type variable"
   | TypeVar => "schematic type variable"
-  | Nat => "number"
+  | Nat => "natural number"
+  | Float => "floating-point number"
   | String => "string"
   | AltString => "back-quoted string"
   | Verbatim => "verbatim text"
@@ -351,6 +352,7 @@
         Syntax.scan_var >> pair Var ||
         Syntax.scan_tid >> pair TypeIdent ||
         Syntax.scan_tvar >> pair TypeVar ||
+        Syntax.scan_float >> pair Float ||
         Syntax.scan_nat >> pair Nat ||
         scan_symid >> pair SymIdent) >> uncurry token));
 
--- a/src/Pure/Isar/token.scala	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/Pure/Isar/token.scala	Sat Oct 30 12:25:18 2010 -0700
@@ -21,7 +21,8 @@
     val VAR = Value("schematic variable")
     val TYPE_IDENT = Value("type variable")
     val TYPE_VAR = Value("schematic type variable")
-    val NAT = Value("number")
+    val NAT = Value("natural number")
+    val FLOAT = Value("floating-point number")
     val STRING = Value("string")
     val ALT_STRING = Value("back-quoted string")
     val VERBATIM = Value("verbatim text")
--- a/src/Pure/ML/ml_antiquote.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/Pure/ML/ml_antiquote.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -71,12 +71,6 @@
     "Context.get_theory (ML_Context.the_global_context ()) " ^ ML_Syntax.print_string name)
   || Scan.succeed "ML_Context.the_global_context ()");
 
-val _ = value "theory_ref"
-  (Scan.lift Args.name >> (fn name =>
-    "Theory.check_thy (Context.get_theory (ML_Context.the_global_context ()) " ^
-      ML_Syntax.print_string name ^ ")")
-  || Scan.succeed "Theory.check_thy (ML_Context.the_global_context ())");
-
 val _ = value "context" (Scan.succeed "ML_Context.the_local_context ()");
 
 val _ = inline "typ" (Args.typ >> (ML_Syntax.atomic o ML_Syntax.print_typ));
--- a/src/Pure/Syntax/lexicon.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/Pure/Syntax/lexicon.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -14,6 +14,7 @@
   val scan_tid: Symbol_Pos.T list -> Symbol_Pos.T list * Symbol_Pos.T list
   val scan_nat: Symbol_Pos.T list -> Symbol_Pos.T list * Symbol_Pos.T list
   val scan_int: Symbol_Pos.T list -> Symbol_Pos.T list * Symbol_Pos.T list
+  val scan_float: Symbol_Pos.T list -> Symbol_Pos.T list * Symbol_Pos.T list
   val scan_hex: Symbol_Pos.T list -> Symbol_Pos.T list * Symbol_Pos.T list
   val scan_bin: Symbol_Pos.T list -> Symbol_Pos.T list * Symbol_Pos.T list
   val scan_var: Symbol_Pos.T list -> Symbol_Pos.T list * Symbol_Pos.T list
--- a/src/Pure/Syntax/syn_trans.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/Pure/Syntax/syn_trans.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -547,7 +547,7 @@
     fun ast_of (Parser.Node (a, pts)) = trans a (map ast_of pts)
       | ast_of (Parser.Tip tok) = Ast.Variable (Lexicon.str_of_token tok);
 
-    val exn_results = map (Exn.capture ast_of) pts;
+    val exn_results = map (Exn.interruptible_capture ast_of) pts;
     val exns = map_filter Exn.get_exn exn_results;
     val results = map_filter Exn.get_result exn_results
   in (case (results, exns) of ([], exn :: _) => reraise exn | _ => results) end;
@@ -571,7 +571,7 @@
           Term.list_comb (term_of ast, map term_of asts)
       | term_of (ast as Ast.Appl _) = raise Ast.AST ("ast_to_term: malformed ast", [ast]);
 
-    val exn_results = map (Exn.capture term_of) asts;
+    val exn_results = map (Exn.interruptible_capture term_of) asts;
     val exns = map_filter Exn.get_exn exn_results;
     val results = map_filter Exn.get_result exn_results
   in (case (results, exns) of ([], exn :: _) => reraise exn | _ => results) end;
--- a/src/Pure/Thy/thy_syntax.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/Pure/Thy/thy_syntax.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -56,6 +56,7 @@
   | Token.TypeIdent     => Markup.tfree
   | Token.TypeVar       => Markup.tvar
   | Token.Nat           => Markup.ident
+  | Token.Float         => Markup.ident
   | Token.String        => Markup.string
   | Token.AltString     => Markup.altstring
   | Token.Verbatim      => Markup.verbatim
--- a/src/Pure/config.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/Pure/config.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -6,13 +6,14 @@
 
 signature CONFIG =
 sig
-  datatype value = Bool of bool | Int of int | String of string
+  datatype value = Bool of bool | Int of int | Real of real | String of string
   val print_value: value -> string
   val print_type: value -> string
   type 'a T
   type raw = value T
   val bool: raw -> bool T
   val int: raw -> int T
+  val real: raw -> real T
   val string: raw -> string T
   val get: Proof.context -> 'a T -> 'a
   val map: 'a T -> ('a -> 'a) -> Proof.context -> Proof.context
@@ -37,19 +38,23 @@
 datatype value =
   Bool of bool |
   Int of int |
+  Real of real |
   String of string;
 
 fun print_value (Bool true) = "true"
   | print_value (Bool false) = "false"
   | print_value (Int i) = signed_string_of_int i
+  | print_value (Real x) = signed_string_of_real x
   | print_value (String s) = quote s;
 
 fun print_type (Bool _) = "bool"
   | print_type (Int _) = "int"
+  | print_type (Real _) = "real"
   | print_type (String _) = "string";
 
 fun same_type (Bool _) (Bool _) = true
   | same_type (Int _) (Int _) = true
+  | same_type (Real _) (Real _) = true
   | same_type (String _) (String _) = true
   | same_type _ _ = false;
 
@@ -78,6 +83,7 @@
 
 val bool = coerce Bool (fn Bool b => b);
 val int = coerce Int (fn Int i => i);
+val real = coerce Real (fn Real x => x);
 val string = coerce String (fn String s => s);
 
 fun get_generic context (Config {get_value, ...}) = get_value context;
@@ -118,7 +124,8 @@
     fun map_value f (context as Context.Proof _) =
           let val context' = update_value f context in
             if global andalso
-              get_value (Context.Theory (Context.theory_of context')) <> get_value context'
+              print_value (get_value (Context.Theory (Context.theory_of context'))) <>
+                print_value (get_value context')
             then (warning ("Ignoring local change of global option " ^ quote name); context)
             else context'
           end
--- a/src/Pure/library.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/Pure/library.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -130,6 +130,10 @@
   val read_int: string list -> int * string list
   val oct_char: string -> string
 
+  (*reals*)
+  val string_of_real: real -> string
+  val signed_string_of_real: real -> string
+
   (*strings*)
   val nth_string: string -> int -> string
   val fold_string: (string -> 'a -> 'a) -> string -> 'a -> 'a
@@ -668,6 +672,15 @@
 
 
 
+(** reals **)
+
+val string_of_real = Real.fmt (StringCvt.GEN NONE);
+
+fun signed_string_of_real x =
+  if x < 0.0 then "-" ^ string_of_real (~ x) else string_of_real x;
+
+
+
 (** strings **)
 
 (* functions tuned for strings, avoiding explode *)
--- a/src/Pure/more_thm.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/Pure/more_thm.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -12,6 +12,7 @@
   structure Ctermtab: TABLE
   structure Thmtab: TABLE
   val aconvc: cterm * cterm -> bool
+  type attribute = Context.generic * thm -> Context.generic * thm
 end;
 
 signature THM =
@@ -64,6 +65,7 @@
   val close_derivation: thm -> thm
   val add_axiom: binding * term -> theory -> (string * thm) * theory
   val add_def: bool -> bool -> binding * term -> theory -> (string * thm) * theory
+  type attribute = Context.generic * thm -> Context.generic * thm
   type binding = binding * attribute list
   val empty_binding: binding
   val rule_attribute: (Context.generic -> thm -> thm) -> attribute
@@ -382,6 +384,9 @@
 
 (** attributes **)
 
+(*attributes subsume any kind of rules or context modifiers*)
+type attribute = Context.generic * thm -> Context.generic * thm;
+
 type binding = binding * attribute list;
 val empty_binding: binding = (Binding.empty, []);
 
--- a/src/Pure/thm.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/Pure/thm.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -39,7 +39,6 @@
   (*theorems*)
   type thm
   type conv = cterm -> thm
-  type attribute = Context.generic * thm -> Context.generic * thm
   val rep_thm: thm ->
    {thy_ref: theory_ref,
     tags: Properties.T,
@@ -350,9 +349,6 @@
 
 type conv = cterm -> thm;
 
-(*attributes subsume any kind of rules or context modifiers*)
-type attribute = Context.generic * thm -> Context.generic * thm;
-
 (*errors involving theorems*)
 exception THM of string * int * thm list;
 
--- a/src/Pure/type_infer.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/Pure/type_infer.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -9,6 +9,7 @@
   val is_param: indexname -> bool
   val is_paramT: typ -> bool
   val param: int -> string * sort -> typ
+  val mk_param: int -> sort -> typ
   val anyT: sort -> typ
   val paramify_vars: typ -> typ
   val paramify_dummies: typ -> int -> typ * int
--- a/src/Tools/Code/code_runtime.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/Tools/Code/code_runtime.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -64,7 +64,7 @@
 
 fun exec verbose code =
   (if ! trace then tracing code else ();
-  ML_Context.exec (fn () => Secure.use_text ML_Env.local_context (0, "generated code") false code));
+  ML_Context.exec (fn () => Secure.use_text ML_Env.local_context (0, "generated code") verbose code));
 
 fun value ctxt (get, put, put_ml) (prelude, value) =
   let
@@ -105,7 +105,7 @@
     val (program_code, [SOME value_name']) = serializer naming program' [value_name];
     val value_code = space_implode " "
       (value_name' :: "()" :: map (enclose "(" ")") args);
-  in Exn.capture (value ctxt cookie) (program_code, value_code) end;
+  in Exn.interruptible_capture (value ctxt cookie) (program_code, value_code) end;
 
 fun partiality_as_none e = SOME (Exn.release e)
   handle General.Match => NONE
--- a/src/Tools/WWW_Find/scgi_req.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/Tools/WWW_Find/scgi_req.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -106,8 +106,9 @@
        | SOME wv => Byte.unpackStringVec wv);
 
     val content_length =
-      (the o Int.fromString o field) "CONTENT_LENGTH"
-      handle _ => raise InvalidReq "Bad CONTENT_LENGTH";
+      (case Int.fromString (field "CONTENT_LENGTH") of
+        SOME n => n
+      | NONE => raise InvalidReq "Bad CONTENT_LENGTH");
 
     val req = Req {
         path_info = field "PATH_INFO",
--- a/src/Tools/jEdit/README	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/Tools/jEdit/README	Sat Oct 30 12:25:18 2010 -0700
@@ -55,6 +55,10 @@
   Linux, but probably also the deeper reason for the other oddities of
   Apple font management.
 
+- The native font renderer of -Dapple.awt.graphics.UseQuartz=true
+  fails to handle the infinitesimal font size of "hidden" text (e.g.
+  used in Isabelle sub/superscripts).
+
 
 Windows/Linux
 =============
--- a/src/Tools/jEdit/dist-template/etc/settings	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/Tools/jEdit/dist-template/etc/settings	Sat Oct 30 12:25:18 2010 -0700
@@ -6,7 +6,7 @@
 JEDIT_OPTIONS="-reuseview -noserver -nobackground"
 JEDIT_JAVA_OPTIONS="-Xms128m -Xmx512m -Xss2m -Dactors.corePoolSize=4 -Dactors.enableForkJoin=false"
 #JEDIT_JAVA_OPTIONS="-server -Xms128m -Xmx1024m -Xss4m -Dactors.corePoolSize=8 -Dactors.enableForkJoin=false"
-JEDIT_SYSTEM_OPTIONS="-Dapple.awt.graphics.UseQuartz=true -Dapple.laf.useScreenMenuBar=true -Dcom.apple.mrj.application.apple.menu.about.name=Isabelle/jEdit"
+JEDIT_SYSTEM_OPTIONS="-Dapple.laf.useScreenMenuBar=true -Dcom.apple.mrj.application.apple.menu.about.name=Isabelle/jEdit"
 
 JEDIT_STYLE_SHEETS="$ISABELLE_HOME/lib/html/isabelle.css:$JEDIT_HOME/etc/isabelle-jedit.css:$ISABELLE_HOME_USER/etc/isabelle.css:$ISABELLE_HOME_USER/etc/isabelle-jedit.css"
 
--- a/src/Tools/jEdit/dist-template/lib/Tools/jedit	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/Tools/jEdit/dist-template/lib/Tools/jedit	Sat Oct 30 12:25:18 2010 -0700
@@ -106,7 +106,7 @@
 
 if [ ! -e "$JEDIT_SETTINGS/perspective.xml" ]; then
   cat > "$JEDIT_SETTINGS/DockableWindowManager/perspective-view0.xml" <<EOF
-<DOCKING LEFT="" TOP="" RIGHT="isabelle-session" BOTTOM="" LEFT_POS="0" TOP_POS="0" RIGHT_POS="250" BOTTOM_POS="250" />
+<DOCKING LEFT="" TOP="" RIGHT="" BOTTOM="isabelle-session" LEFT_POS="0" TOP_POS="0" RIGHT_POS="250" BOTTOM_POS="250" />
 EOF
   cat > "$JEDIT_SETTINGS/perspective.xml" <<EOF
 <?xml version="1.0" encoding="UTF-8" ?>
--- a/src/Tools/jEdit/dist-template/properties/jedit.props	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/Tools/jEdit/dist-template/properties/jedit.props	Sat Oct 30 12:25:18 2010 -0700
@@ -178,9 +178,9 @@
 insert-newline-indent.shortcut=
 insert-newline.shortcut=ENTER
 isabelle-output.dock-position=bottom
+isabelle-session.dock-position=bottom
 isabelle-output.height=174
 isabelle-output.width=412
-isabelle-session.dock-position=right
 line-end.shortcut=END
 line-home.shortcut=HOME
 lookAndFeel=com.sun.java.swing.plaf.nimbus.NimbusLookAndFeel
--- a/src/Tools/jEdit/src/jedit/document_view.scala	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/Tools/jEdit/src/jedit/document_view.scala	Sat Oct 30 12:25:18 2010 -0700
@@ -130,6 +130,7 @@
   private def init_popup(snapshot: Document.Snapshot, x: Int, y: Int)
   {
     exit_popup()
+/* FIXME broken
     val offset = text_area.xyToOffset(x, y)
     val p = new Point(x, y); SwingUtilities.convertPointToScreen(p, text_area.getPainter)
 
@@ -143,6 +144,7 @@
         html_popup = Some(popup)
       case _ =>
     }
+*/
   }
 
 
--- a/src/Tools/quickcheck.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/Tools/quickcheck.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -1,5 +1,5 @@
 (*  Title:      Tools/quickcheck.ML
-    Author:     Stefan Berghofer, Florian Haftmann, TU Muenchen
+    Author:     Stefan Berghofer, Florian Haftmann, Lukas Bulwahn, TU Muenchen
 
 Generic counterexample search engine.
 *)
@@ -16,18 +16,19 @@
   datatype expectation = No_Expectation | No_Counterexample | Counterexample;
   datatype test_params = Test_Params of
   { size: int, iterations: int, default_type: typ list, no_assms: bool,
-    expect : expectation, report: bool, quiet : bool};
-  val test_params_of: Proof.context -> test_params 
+    expect : expectation, report: bool, quiet : bool, timeout : int};
+  val test_params_of : Proof.context -> test_params
   val report : Proof.context -> bool
-  val set_reporting : bool -> Context.generic -> Context.generic
+  val map_test_params :
+    ((int * int) * ((typ list * bool) * ((expectation * bool) * (bool * int)))
+      -> (int * int) * ((typ list * bool) * ((expectation * bool) * (bool * int))))
+    -> Context.generic -> Context.generic
   val add_generator:
     string * (Proof.context -> term -> int -> term list option * (bool list * bool))
       -> Context.generic -> Context.generic
   (* testing terms and proof states *)
-  val gen_test_term: Proof.context -> bool -> string option -> int -> int -> term ->
+  val test_term: Proof.context -> string option -> term ->
     (string * term) list option * ((string * int) list * ((int * report list) list) option)
-  val test_term: Proof.context -> bool -> string option -> int -> int -> term ->
-    (string * term) list option
   val quickcheck: (string * string list) list -> int -> Proof.state -> (string * term) list option
 end;
 
@@ -80,26 +81,31 @@
 
 datatype test_params = Test_Params of
   { size: int, iterations: int, default_type: typ list, no_assms: bool,
-    expect : expectation, report: bool, quiet : bool};
+    expect : expectation, report: bool, quiet : bool, timeout : int};
 
-fun dest_test_params (Test_Params { size, iterations, default_type, no_assms, expect, report, quiet }) =
-  ((size, iterations), ((default_type, no_assms), ((expect, report), quiet)));
+fun dest_test_params
+    (Test_Params { size, iterations, default_type, no_assms, expect, report, quiet, timeout }) =
+  ((size, iterations), ((default_type, no_assms), ((expect, report), (quiet, timeout))));
 
-fun make_test_params ((size, iterations), ((default_type, no_assms), ((expect, report), quiet))) =
+fun make_test_params
+    ((size, iterations), ((default_type, no_assms), ((expect, report), (quiet, timeout)))) =
   Test_Params { size = size, iterations = iterations, default_type = default_type,
-    no_assms = no_assms, expect = expect, report = report, quiet = quiet };
+    no_assms = no_assms, expect = expect, report = report, quiet = quiet, timeout = timeout };
 
-fun map_test_params f (Test_Params { size, iterations, default_type, no_assms, expect, report, quiet }) =
-  make_test_params (f ((size, iterations), ((default_type, no_assms), ((expect, report), quiet))));
+fun map_test_params' f
+    (Test_Params { size, iterations, default_type, no_assms, expect, report, quiet, timeout }) =
+  make_test_params
+    (f ((size, iterations), ((default_type, no_assms), ((expect, report), (quiet, timeout)))));
 
 fun merge_test_params
  (Test_Params { size = size1, iterations = iterations1, default_type = default_type1,
-    no_assms = no_assms1, expect = expect1, report = report1, quiet = quiet1 },
+    no_assms = no_assms1, expect = expect1, report = report1, quiet = quiet1, timeout = timeout1 },
   Test_Params { size = size2, iterations = iterations2, default_type = default_type2,
-    no_assms = no_assms2, expect = expect2, report = report2, quiet = quiet2 }) =
+    no_assms = no_assms2, expect = expect2, report = report2, quiet = quiet2, timeout = timeout2}) =
   make_test_params ((Int.max (size1, size2), Int.max (iterations1, iterations2)),
     ((merge (op =) (default_type1, default_type2), no_assms1 orelse no_assms2),
-    ((merge_expectation (expect1, expect2), report1 orelse report2), quiet1 orelse quiet2)));
+    ((merge_expectation (expect1, expect2), report1 orelse report2),
+    (quiet1 orelse quiet2, Int.min (timeout1, timeout2)))));
 
 structure Data = Generic_Data
 (
@@ -108,7 +114,7 @@
       * test_params;
   val empty = ([], Test_Params
     { size = 10, iterations = 100, default_type = [], no_assms = false,
-      expect = No_Expectation, report = false, quiet = false});
+      expect = No_Expectation, report = false, quiet = false, timeout = 30});
   val extend = I;
   fun merge ((generators1, params1), (generators2, params2)) : T =
     (AList.merge (op =) (K true) (generators1, generators2),
@@ -117,12 +123,37 @@
 
 val test_params_of = snd o Data.get o Context.Proof;
 
+val size = fst o fst o dest_test_params o test_params_of
+
+val iterations = snd o fst o dest_test_params o test_params_of
+
+val default_type = fst o fst o snd o dest_test_params o test_params_of
+
+val no_assms = snd o fst o snd o dest_test_params o test_params_of
+
+val expect = fst o fst o snd o snd o dest_test_params o test_params_of
+
 val report = snd o fst o snd o snd o dest_test_params o test_params_of
 
-fun map_report f (Test_Params { size, iterations, default_type, no_assms, expect, report, quiet }) =
-  make_test_params ((size, iterations), ((default_type, no_assms), ((expect, f report), quiet)));
+val quiet = fst o snd o snd o snd o dest_test_params o test_params_of
+
+val timeout = snd o snd o snd o snd o dest_test_params o test_params_of
+
+fun map_report f
+  (Test_Params { size, iterations, default_type, no_assms, expect, report, quiet, timeout }) =
+    make_test_params
+      ((size, iterations), ((default_type, no_assms), ((expect, f report), (quiet, timeout))));
 
-fun set_reporting report = Data.map (apsnd (map_report (K report)))
+fun map_quiet f
+  (Test_Params { size, iterations, default_type, no_assms, expect, report, quiet, timeout }) =
+    make_test_params
+      ((size, iterations), ((default_type, no_assms), ((expect, report), (f quiet, timeout))));
+    
+fun set_report report = Data.map (apsnd (map_report (K report)))
+
+fun set_quiet quiet = Data.map (apsnd (map_quiet (K quiet)))
+
+val map_test_params = Data.map o apsnd o map_test_params'
 
 val add_generator = Data.map o apfst o AList.update (op =);
 
@@ -139,9 +170,10 @@
 
 fun mk_testers_strict ctxt t =
   let
-    val generators = ((map snd o fst o Data.get  o Context.Proof) ctxt)
-    val testers = map (fn generator => Exn.capture (generator ctxt) t) generators;
-  in if forall (is_none o Exn.get_result) testers
+    val generators = ((map snd o fst o Data.get o Context.Proof) ctxt)
+    val testers = map (fn generator => Exn.interruptible_capture (generator ctxt) t) generators;
+  in
+    if forall (is_none o Exn.get_result) testers
     then [(Exn.release o snd o split_last) testers]
     else map_filter Exn.get_result testers
   end;
@@ -165,17 +197,17 @@
     val time = Time.toMilliseconds (#cpu (end_timing start))
   in (Exn.release result, (description, time)) end
 
-fun gen_test_term ctxt quiet generator_name size i t =
+fun test_term ctxt generator_name t =
   let
     val (names, t') = prep_test_term t;
     val (testers, comp_time) = cpu_time "quickcheck compilation"
       (fn () => (case generator_name
-       of NONE => if quiet then mk_testers ctxt t' else mk_testers_strict ctxt t'
+       of NONE => if quiet ctxt then mk_testers ctxt t' else mk_testers_strict ctxt t'
         | SOME name => [mk_tester_select name ctxt t']));
     fun iterate f 0 report = (NONE, report)
       | iterate f j report =
         let
-          val (test_result, single_report) = apsnd Run (f ()) handle Match => (if quiet then ()
+          val (test_result, single_report) = apsnd Run (f ()) handle Match => (if quiet ctxt then ()
              else warning "Exception Match raised during quickcheck"; (NONE, MatchExc))
           val report = collect_single_report single_report report
         in
@@ -185,19 +217,19 @@
       satisfied_assms = [], positive_concl_tests = 0 }
     fun with_testers k [] = (NONE, [])
       | with_testers k (tester :: testers) =
-          case iterate (fn () => tester (k - 1)) i empty_report
+          case iterate (fn () => tester (k - 1)) (iterations ctxt) 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)
+      if k > size ctxt then (NONE, reports)
       else
-       (if quiet then () else Output.urgent_message ("Test data size: " ^ string_of_int k);
+       (if quiet ctxt 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) =
-      TimeLimit.timeLimit (Time.fromSeconds 20) (fn () => cpu_time "quickcheck execution"
+      TimeLimit.timeLimit (Time.fromSeconds (timeout ctxt)) (fn () => cpu_time "quickcheck execution"
       (fn () => apfst
          (fn result => case result of NONE => NONE
         | SOME ts => SOME (names ~~ ts)) (with_size 1 []))) ()
@@ -206,11 +238,6 @@
     (result, ([exec_time, comp_time], if report ctxt then SOME reports else NONE))
   end;
 
-fun test_term ctxt quiet generator_name size i t =
-  ctxt
-  |> Context.proof_map (set_reporting false)
-  |> (fn ctxt' => fst (gen_test_term ctxt' quiet generator_name size i t))
-
 exception WELLSORTED of string
 
 fun monomorphic_term thy insts default_T = 
@@ -231,7 +258,7 @@
 
 datatype wellsorted_error = Wellsorted_Error of string | Term of term
 
-fun test_goal quiet generator_name size iterations default_Ts no_assms insts i state =
+fun test_goal generator_name insts i state =
   let
     val lthy = Proof.context_of state;
     val thy = Proof.theory_of state;
@@ -243,34 +270,37 @@
      of NONE => NONE
       | SOME "" => NONE
       | SOME locale => SOME locale;
-    val assms = if no_assms then [] else case some_locale
+    val assms = if no_assms lthy then [] else case some_locale
      of NONE => Assumption.all_assms_of lthy
       | SOME locale => Assumption.local_assms_of lthy (Locale.init locale thy);
     val proto_goal = Logic.list_implies (map Thm.term_of assms, subst_bounds (frees, strip gi));
     val check_goals = case some_locale
      of NONE => [proto_goal]
-      | SOME locale => map (fn (_, phi) => Morphism.term phi proto_goal) (Locale.registrations_of (Context.Theory thy) (*FIXME*) locale);
+      | SOME locale => map (fn (_, phi) => Morphism.term phi proto_goal)
+        (Locale.registrations_of (Context.Theory thy) (*FIXME*) locale);
     val inst_goals = maps (fn check_goal => map (fn T =>
       Term ((Object_Logic.atomize_term thy o monomorphic_term thy insts T) check_goal)
-        handle WELLSORTED s => Wellsorted_Error s) default_Ts) check_goals
+        handle WELLSORTED s => Wellsorted_Error s) (default_type lthy)) check_goals
     val error_msg = cat_lines (map_filter (fn Term t => NONE | Wellsorted_Error s => SOME s) inst_goals)
     val correct_inst_goals =
       case map_filter (fn Term t => SOME t | Wellsorted_Error s => NONE) inst_goals of
         [] => error error_msg
       | xs => xs
-    val _ = if quiet then () else warning error_msg
+    val _ = if quiet lthy then () else warning error_msg
     fun collect_results f reports [] = (NONE, rev reports)
       | collect_results f reports (t :: ts) =
         case f t of
           (SOME res, report) => (SOME res, rev (report :: reports))
         | (NONE, report) => collect_results f (report :: reports) ts
-  in collect_results (gen_test_term lthy quiet generator_name size iterations) [] correct_inst_goals end;
+  in collect_results (test_term lthy generator_name) [] correct_inst_goals end;
 
 (* pretty printing *)
 
-fun pretty_counterex ctxt NONE = Pretty.str "Quickcheck found no counterexample."
-  | pretty_counterex ctxt (SOME cex) =
-      Pretty.chunks (Pretty.str "Quickcheck found a counterexample:\n" ::
+fun tool_name auto = (if auto then "Auto " else "") ^ "Quickcheck"
+
+fun pretty_counterex ctxt auto NONE = Pretty.str (tool_name auto ^ " found no counterexample.")
+  | pretty_counterex ctxt auto (SOME cex) =
+      Pretty.chunks (Pretty.str (tool_name auto ^ " found a counterexample:\n") ::
         map (fn (s, t) =>
           Pretty.block [Pretty.str (s ^ " ="), Pretty.brk 1, Syntax.pretty_term ctxt t]) cex);
 
@@ -299,8 +329,9 @@
       (rev reports))
   | pretty_reports ctxt NONE = Pretty.str ""
 
-fun pretty_counterex_and_reports ctxt (cex, timing_and_reports) =
-  Pretty.chunks (pretty_counterex ctxt cex :: map (pretty_reports ctxt) (map snd timing_and_reports))
+fun pretty_counterex_and_reports ctxt auto (cex, timing_and_reports) =
+  Pretty.chunks (pretty_counterex ctxt auto cex ::
+    map (pretty_reports ctxt) (map snd timing_and_reports))
 
 (* automatic testing *)
 
@@ -310,18 +341,16 @@
   else
     let
       val ctxt = Proof.context_of state;
-      val Test_Params {size, iterations, default_type, no_assms, ...} =
-        (snd o Data.get o Context.Proof) ctxt;
       val res =
         state
-        |> Proof.map_context (Context.proof_map (set_reporting false))
-        |> try (test_goal true NONE size iterations default_type no_assms [] 1);
+        |> Proof.map_context (Context.proof_map (set_report false #> set_quiet true))
+        |> try (test_goal NONE [] 1);
     in
       case res of
         NONE => (false, state)
       | SOME (NONE, report) => (false, state)
       | SOME (cex, report) => (true, Proof.goal_message (K (Pretty.chunks [Pretty.str "",
-          Pretty.mark Markup.hilite (pretty_counterex ctxt cex)])) state)
+          Pretty.mark Markup.hilite (pretty_counterex ctxt true cex)])) state)
     end
 
 val setup = Auto_Tools.register_tool ("quickcheck", auto_quickcheck)
@@ -356,7 +385,9 @@
   | parse_test_param ctxt ("report", [arg]) =
       (apsnd o apsnd o apfst o apsnd o K) (read_bool arg)
   | parse_test_param ctxt ("quiet", [arg]) =
-      (apsnd o apsnd o apsnd o K) (read_bool arg)       
+      (apsnd o apsnd o apsnd o apfst o K) (read_bool arg)
+  | parse_test_param ctxt ("timeout", [arg]) =
+      (apsnd o apsnd o apsnd o apsnd o K) (read_nat arg)
   | parse_test_param ctxt (name, _) =
       error ("Unknown test parameter: " ^ name);
 
@@ -374,7 +405,7 @@
     val f = fold (parse_test_param ctxt) args;
   in
     thy
-    |> (Context.theory_map o Data.map o apsnd o map_test_params) f
+    |> (Context.theory_map o map_test_params) f
   end;
 
 fun gen_quickcheck args i state =
@@ -382,23 +413,22 @@
     val ctxt = Proof.context_of state;
     val default_params = (dest_test_params o snd o Data.get o Context.Proof) ctxt;
     val f = fold (parse_test_param_inst ctxt) args;
-    val (((size, iterations), ((default_type, no_assms), ((expect, report), quiet))), (generator_name, insts)) =
-      f (default_params, (NONE, []));
+    val (test_params, (generator_name, insts)) = f (default_params, (NONE, []));
   in
     state
-    |> Proof.map_context (Context.proof_map (set_reporting report))
-    |> test_goal quiet generator_name size iterations default_type no_assms insts i
-    |> tap (fn (SOME x, _) => if expect = No_Counterexample then
+    |> Proof.map_context (Context.proof_map (map_test_params (K test_params)))
+    |> (fn state' => test_goal generator_name insts i state'
+      |> tap (fn (SOME x, _) => if expect (Proof.context_of state') = No_Counterexample then
                  error ("quickcheck expected to find no counterexample but found one") else ()
-             | (NONE, _) => if expect = Counterexample then
-                 error ("quickcheck expected to find a counterexample but did not find one") else ())
+             | (NONE, _) => if expect (Proof.context_of state') = Counterexample then
+                 error ("quickcheck expected to find a counterexample but did not find one") else ()))
   end;
 
 fun quickcheck args i state = fst (gen_quickcheck args i state);
 
 fun quickcheck_cmd args i state =
   gen_quickcheck args i (Toplevel.proof_of state)
-  |> Pretty.writeln o pretty_counterex_and_reports (Toplevel.context_of state);
+  |> Pretty.writeln o pretty_counterex_and_reports (Toplevel.context_of state) false;
 
 val parse_arg = Parse.name -- (Scan.optional (Parse.$$$ "=" |-- 
   ((Parse.name >> single) || (Parse.$$$ "[" |-- Parse.list1 Parse.name --| Parse.$$$ "]"))) ["true"]);
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/subtyping.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -0,0 +1,769 @@
+(*  Title:      Tools/subtyping.ML
+    Author:     Dmitriy Traytel, TU Muenchen
+
+Coercive subtyping via subtype constraints.
+*)
+
+signature SUBTYPING =
+sig
+  datatype variance = COVARIANT | CONTRAVARIANT | INVARIANT
+  val infer_types: Proof.context -> (string -> typ option) -> (indexname -> typ option) ->
+    term list -> term list
+  val add_type_map: term -> Context.generic -> Context.generic
+  val add_coercion: term -> Context.generic -> Context.generic
+  val setup: theory -> theory
+end;
+
+structure Subtyping: SUBTYPING =
+struct
+
+(** coercions data **)
+
+datatype variance = COVARIANT | CONTRAVARIANT | INVARIANT
+
+datatype data = Data of
+  {coes: term Symreltab.table,  (*coercions table*)
+   coes_graph: unit Graph.T,  (*coercions graph*)
+   tmaps: (term * variance list) Symtab.table};  (*map functions*)
+
+fun make_data (coes, coes_graph, tmaps) =
+  Data {coes = coes, coes_graph = coes_graph, tmaps = tmaps};
+
+structure Data = Generic_Data
+(
+  type T = data;
+  val empty = make_data (Symreltab.empty, Graph.empty, Symtab.empty);
+  val extend = I;
+  fun merge
+    (Data {coes = coes1, coes_graph = coes_graph1, tmaps = tmaps1},
+      Data {coes = coes2, coes_graph = coes_graph2, tmaps = tmaps2}) =
+    make_data (Symreltab.merge (op aconv) (coes1, coes2),
+      Graph.merge (op =) (coes_graph1, coes_graph2),
+      Symtab.merge (eq_pair (op aconv) (op =)) (tmaps1, tmaps2));
+);
+
+fun map_data f =
+  Data.map (fn Data {coes, coes_graph, tmaps} =>
+    make_data (f (coes, coes_graph, tmaps)));
+
+fun map_coes f =
+  map_data (fn (coes, coes_graph, tmaps) =>
+    (f coes, coes_graph, tmaps));
+
+fun map_coes_graph f =
+  map_data (fn (coes, coes_graph, tmaps) =>
+    (coes, f coes_graph, tmaps));
+
+fun map_coes_and_graph f =
+  map_data (fn (coes, coes_graph, tmaps) =>
+    let val (coes', coes_graph') = f (coes, coes_graph);
+    in (coes', coes_graph', tmaps) end);
+
+fun map_tmaps f =
+  map_data (fn (coes, coes_graph, tmaps) =>
+    (coes, coes_graph, f tmaps));
+
+val rep_data = (fn Data args => args) o Data.get o Context.Proof;
+
+val coes_of = #coes o rep_data;
+val coes_graph_of = #coes_graph o rep_data;
+val tmaps_of = #tmaps o rep_data;
+
+
+
+(** utils **)
+
+fun nameT (Type (s, [])) = s;
+fun t_of s = Type (s, []);
+
+fun sort_of (TFree (_, S)) = SOME S
+  | sort_of (TVar (_, S)) = SOME S
+  | sort_of _ = NONE;
+
+val is_typeT = fn (Type _) => true | _ => false;
+val is_compT = fn (Type (_, _ :: _)) => true | _ => false;
+val is_freeT = fn (TFree _) => true | _ => false;
+val is_fixedvarT = fn (TVar (xi, _)) => not (Type_Infer.is_param xi) | _ => false;
+
+
+(* unification *)  (* TODO dup? needed for weak unification *)
+
+exception NO_UNIFIER of string * typ Vartab.table;
+
+fun unify weak ctxt =
+  let
+    val thy = ProofContext.theory_of ctxt;
+    val pp = Syntax.pp ctxt;
+    val arity_sorts = Type.arity_sorts pp (Sign.tsig_of thy);
+
+
+    (* adjust sorts of parameters *)
+
+    fun not_of_sort x S' S =
+      "Variable " ^ x ^ "::" ^ Syntax.string_of_sort ctxt S' ^ " not of sort " ^
+        Syntax.string_of_sort ctxt S;
+
+    fun meet (_, []) tye_idx = tye_idx
+      | meet (Type (a, Ts), S) (tye_idx as (tye, _)) =
+          meets (Ts, arity_sorts a S handle ERROR msg => raise NO_UNIFIER (msg, tye)) tye_idx
+      | meet (TFree (x, S'), S) (tye_idx as (tye, _)) =
+          if Sign.subsort thy (S', S) then tye_idx
+          else raise NO_UNIFIER (not_of_sort x S' S, tye)
+      | meet (TVar (xi, S'), S) (tye_idx as (tye, idx)) =
+          if Sign.subsort thy (S', S) then tye_idx
+          else if Type_Infer.is_param xi then
+            (Vartab.update_new
+              (xi, Type_Infer.mk_param idx (Sign.inter_sort thy (S', S))) tye, idx + 1)
+          else raise NO_UNIFIER (not_of_sort (Term.string_of_vname xi) S' S, tye)
+    and meets (T :: Ts, S :: Ss) (tye_idx as (tye, _)) =
+          meets (Ts, Ss) (meet (Type_Infer.deref tye T, S) tye_idx)
+      | meets _ tye_idx = tye_idx;
+
+    val weak_meet = if weak then fn _ => I else meet
+
+
+    (* occurs check and assignment *)
+
+    fun occurs_check tye xi (TVar (xi', _)) =
+          if xi = xi' then raise NO_UNIFIER ("Occurs check!", tye)
+          else
+            (case Vartab.lookup tye xi' of
+              NONE => ()
+            | SOME T => occurs_check tye xi T)
+      | occurs_check tye xi (Type (_, Ts)) = List.app (occurs_check tye xi) Ts
+      | occurs_check _ _ _ = ();
+
+    fun assign xi (T as TVar (xi', _)) S env =
+          if xi = xi' then env
+          else env |> weak_meet (T, S) |>> Vartab.update_new (xi, T)
+      | assign xi T S (env as (tye, _)) =
+          (occurs_check tye xi T; env |> weak_meet (T, S) |>> Vartab.update_new (xi, T));
+
+
+    (* unification *)
+
+    fun show_tycon (a, Ts) =
+      quote (Syntax.string_of_typ ctxt (Type (a, replicate (length Ts) dummyT)));
+
+    fun unif (T1, T2) (env as (tye, _)) =
+      (case pairself (`Type_Infer.is_paramT o Type_Infer.deref tye) (T1, T2) of
+        ((true, TVar (xi, S)), (_, T)) => assign xi T S env
+      | ((_, T), (true, TVar (xi, S))) => assign xi T S env
+      | ((_, Type (a, Ts)), (_, Type (b, Us))) =>
+          if weak andalso null Ts andalso null Us then env
+          else if a <> b then
+            raise NO_UNIFIER
+              ("Clash of types " ^ show_tycon (a, Ts) ^ " and " ^ show_tycon (b, Us), tye)
+          else fold unif (Ts ~~ Us) env
+      | ((_, T), (_, U)) => if T = U then env else raise NO_UNIFIER ("", tye));
+
+  in unif end;
+
+val weak_unify = unify true;
+val strong_unify = unify false;
+
+
+(* Typ_Graph shortcuts *)
+
+val add_edge = Typ_Graph.add_edge_acyclic;
+fun get_preds G T = Typ_Graph.all_preds G [T];
+fun get_succs G T = Typ_Graph.all_succs G [T];
+fun maybe_new_typnode T G = perhaps (try (Typ_Graph.new_node (T, ()))) G;
+fun maybe_new_typnodes Ts G = fold maybe_new_typnode Ts G;
+fun new_imm_preds G Ts =
+  subtract (op =) Ts (distinct (op =) (maps (Typ_Graph.imm_preds G) Ts));
+fun new_imm_succs G Ts =
+  subtract op= Ts (distinct (op =) (maps (Typ_Graph.imm_succs G) Ts));
+
+
+(* Graph shortcuts *)
+
+fun maybe_new_node s G = perhaps (try (Graph.new_node (s, ()))) G
+fun maybe_new_nodes ss G = fold maybe_new_node ss G
+
+
+
+(** error messages **)
+
+fun prep_output ctxt tye bs ts Ts =
+  let
+    val (Ts_bTs', ts') = Type_Infer.finish ctxt tye (Ts @ map snd bs, ts);
+    val (Ts', Ts'') = chop (length Ts) Ts_bTs';
+    fun prep t =
+      let val xs = rev (Term.variant_frees t (rev (map fst bs ~~ Ts'')))
+      in Term.subst_bounds (map Syntax.mark_boundT xs, t) end;
+  in (map prep ts', Ts') end;
+
+fun err_loose i = error ("Loose bound variable: B." ^ string_of_int i);
+
+fun inf_failed msg =
+  "Subtype inference failed" ^ (if msg = "" then "" else ": " ^ msg) ^ "\n\n";
+
+fun err_appl ctxt msg tye bs t T u U =
+  let val ([t', u'], [T', U']) = prep_output ctxt tye bs [t, u] [T, U]
+  in error (inf_failed msg ^ Type.appl_error (Syntax.pp ctxt) t' T' u' U' ^ "\n") end;
+
+fun err_subtype ctxt msg tye (bs, t $ u, U, V, U') =
+  err_appl ctxt msg tye bs t (U --> V) u U';
+
+fun err_list ctxt msg tye Ts =
+  let
+    val (_, Ts') = prep_output ctxt tye [] [] Ts;
+    val text = cat_lines ([inf_failed msg,
+      "Cannot unify a list of types that should be the same,",
+      "according to suptype dependencies:",
+      (Pretty.string_of (Pretty.list "[" "]" (map (Pretty.typ (Syntax.pp ctxt)) Ts')))]);
+  in
+    error text
+  end;
+
+fun err_bound ctxt msg tye packs =
+  let
+    val pp = Syntax.pp ctxt;
+    val (ts, Ts) = fold
+      (fn (bs, t $ u, U, _, U') => fn (ts, Ts) =>
+        let val (t', T') = prep_output ctxt tye bs [t, u] [U, U']
+        in (t' :: ts, T' :: Ts) end)
+      packs ([], []);
+    val text = cat_lines ([inf_failed msg, "Cannot fullfill subtype constraints:"] @
+        (map2 (fn [t, u] => fn [T, U] => Pretty.string_of (
+          Pretty.block [
+            Pretty.typ pp T, Pretty.brk 2, Pretty.str "<:", Pretty.brk 2, Pretty.typ pp U,
+            Pretty.brk 3, Pretty.str "from function application", Pretty.brk 2,
+            Pretty.block [Pretty.term pp t, Pretty.brk 1, Pretty.term pp u]]))
+        ts Ts))
+  in
+    error text
+  end;
+
+
+
+(** constraint generation **)
+
+fun generate_constraints ctxt =
+  let
+    fun gen cs _ (Const (_, T)) tye_idx = (T, tye_idx, cs)
+      | gen cs _ (Free (_, T)) tye_idx = (T, tye_idx, cs)
+      | gen cs _ (Var (_, T)) tye_idx = (T, tye_idx, cs)
+      | gen cs bs (Bound i) tye_idx =
+          (snd (nth bs i handle Subscript => err_loose i), tye_idx, cs)
+      | gen cs bs (Abs (x, T, t)) tye_idx =
+          let val (U, tye_idx', cs') = gen cs ((x, T) :: bs) t tye_idx
+          in (T --> U, tye_idx', cs') end
+      | gen cs bs (t $ u) tye_idx =
+          let
+            val (T, tye_idx', cs') = gen cs bs t tye_idx;
+            val (U', (tye, idx), cs'') = gen cs' bs u tye_idx';
+            val U = Type_Infer.mk_param idx [];
+            val V = Type_Infer.mk_param (idx + 1) [];
+            val tye_idx''= strong_unify ctxt (U --> V, T) (tye, idx + 2)
+              handle NO_UNIFIER (msg, tye') => err_appl ctxt msg tye' bs t T u U;
+            val error_pack = (bs, t $ u, U, V, U');
+          in (V, tye_idx'', ((U', U), error_pack) :: cs'') end;
+  in
+    gen [] []
+  end;
+
+
+
+(** constraint resolution **)
+
+exception BOUND_ERROR of string;
+
+fun process_constraints ctxt cs tye_idx =
+  let
+    val coes_graph = coes_graph_of ctxt;
+    val tmaps = tmaps_of ctxt;
+    val tsig = Sign.tsig_of (ProofContext.theory_of ctxt);
+    val pp = Syntax.pp ctxt;
+    val arity_sorts = Type.arity_sorts pp tsig;
+    val subsort = Type.subsort tsig;
+
+    fun split_cs _ [] = ([], [])
+      | split_cs f (c :: cs) =
+          (case pairself f (fst c) of
+            (false, false) => apsnd (cons c) (split_cs f cs)
+          | _ => apfst (cons c) (split_cs f cs));
+
+
+    (* check whether constraint simplification will terminate using weak unification *)
+
+    val _ = fold (fn (TU, error_pack) => fn tye_idx =>
+      (weak_unify ctxt TU tye_idx handle NO_UNIFIER (msg, tye) =>
+        err_subtype ctxt ("Weak unification of subtype constraints fails:\n" ^ msg)
+          tye error_pack)) cs tye_idx;
+
+
+    (* simplify constraints *)
+
+    fun simplify_constraints cs tye_idx =
+      let
+        fun contract a Ts Us error_pack done todo tye idx =
+          let
+            val arg_var =
+              (case Symtab.lookup tmaps a of
+                (*everything is invariant for unknown constructors*)
+                NONE => replicate (length Ts) INVARIANT
+              | SOME av => snd av);
+            fun new_constraints (variance, constraint) (cs, tye_idx) =
+              (case variance of
+                COVARIANT => (constraint :: cs, tye_idx)
+              | CONTRAVARIANT => (swap constraint :: cs, tye_idx)
+              | INVARIANT => (cs, strong_unify ctxt constraint tye_idx
+                  handle NO_UNIFIER (msg, tye) => err_subtype ctxt msg tye error_pack));
+            val (new, (tye', idx')) = apfst (fn cs => (cs ~~ replicate (length cs) error_pack))
+              (fold new_constraints (arg_var ~~ (Ts ~~ Us)) ([], (tye, idx)));
+            val test_update = is_compT orf is_freeT orf is_fixedvarT;
+            val (ch, done') =
+              if not (null new) then ([], done)
+              else split_cs (test_update o Type_Infer.deref tye') done;
+            val todo' = ch @ todo;
+          in
+            simplify done' (new @ todo') (tye', idx')
+          end
+        (*xi is definitely a parameter*)
+        and expand varleq xi S a Ts error_pack done todo tye idx =
+          let
+            val n = length Ts;
+            val args = map2 Type_Infer.mk_param (idx upto idx + n - 1) (arity_sorts a S);
+            val tye' = Vartab.update_new (xi, Type(a, args)) tye;
+            val (ch, done') = split_cs (is_compT o Type_Infer.deref tye') done;
+            val todo' = ch @ todo;
+            val new =
+              if varleq then (Type(a, args), Type (a, Ts))
+              else (Type (a, Ts), Type (a, args));
+          in
+            simplify done' ((new, error_pack) :: todo') (tye', idx + n)
+          end
+        (*TU is a pair of a parameter and a free/fixed variable*)
+        and eliminate TU error_pack done todo tye idx =
+          let
+            val [TVar (xi, S)] = filter Type_Infer.is_paramT TU;
+            val [T] = filter_out Type_Infer.is_paramT TU;
+            val SOME S' = sort_of T;
+            val test_update = if is_freeT T then is_freeT else is_fixedvarT;
+            val tye' = Vartab.update_new (xi, T) tye;
+            val (ch, done') = split_cs (test_update o Type_Infer.deref tye') done;
+            val todo' = ch @ todo;
+          in
+            if subsort (S', S) (*TODO check this*)
+            then simplify done' todo' (tye', idx)
+            else err_subtype ctxt "Sort mismatch" tye error_pack
+          end
+        and simplify done [] tye_idx = (done, tye_idx)
+          | simplify done (((T, U), error_pack) :: todo) (tye_idx as (tye, idx)) =
+              (case (Type_Infer.deref tye T, Type_Infer.deref tye U) of
+                (Type (a, []), Type (b, [])) =>
+                  if a = b then simplify done todo tye_idx
+                  else if Graph.is_edge coes_graph (a, b) then simplify done todo tye_idx
+                  else err_subtype ctxt (a ^ " is not a subtype of " ^ b) (fst tye_idx) error_pack
+              | (Type (a, Ts), Type (b, Us)) =>
+                  if a <> b then err_subtype ctxt "Different constructors" (fst tye_idx) error_pack
+                  else contract a Ts Us error_pack done todo tye idx
+              | (TVar (xi, S), Type (a, Ts as (_ :: _))) =>
+                  expand true xi S a Ts error_pack done todo tye idx
+              | (Type (a, Ts as (_ :: _)), TVar (xi, S)) =>
+                  expand false xi S a Ts error_pack done todo tye idx
+              | (T, U) =>
+                  if T = U then simplify done todo tye_idx
+                  else if exists (is_freeT orf is_fixedvarT) [T, U] andalso
+                    exists Type_Infer.is_paramT [T, U]
+                  then eliminate [T, U] error_pack done todo tye idx
+                  else if exists (is_freeT orf is_fixedvarT) [T, U]
+                  then err_subtype ctxt "Not eliminated free/fixed variables"
+                        (fst tye_idx) error_pack
+                  else simplify (((T, U), error_pack) :: done) todo tye_idx);
+      in
+        simplify [] cs tye_idx
+      end;
+
+
+    (* do simplification *)
+
+    val (cs', tye_idx') = simplify_constraints cs tye_idx;
+
+    fun find_error_pack lower T' =
+      map snd (filter (fn ((T, U), _) => if lower then T' = U else T' = T) cs');
+
+    fun unify_list (T :: Ts) tye_idx =
+      fold (fn U => fn tye_idx => strong_unify ctxt (T, U) tye_idx
+        handle NO_UNIFIER (msg, tye) => err_list ctxt msg tye (T :: Ts))
+      Ts tye_idx;
+
+    (*styps stands either for supertypes or for subtypes of a type T
+      in terms of the subtype-relation (excluding T itself)*)
+    fun styps super T =
+      (if super then Graph.imm_succs else Graph.imm_preds) coes_graph T
+        handle Graph.UNDEF _ => [];
+
+    fun minmax sup (T :: Ts) =
+      let
+        fun adjust T U = if sup then (T, U) else (U, T);
+        fun extract T [] = T
+          | extract T (U :: Us) =
+              if Graph.is_edge coes_graph (adjust T U) then extract T Us
+              else if Graph.is_edge coes_graph (adjust U T) then extract U Us
+              else raise BOUND_ERROR "Uncomparable types in type list";
+      in
+        t_of (extract T Ts)
+      end;
+
+    fun ex_styp_of_sort super T styps_and_sorts =
+      let
+        fun adjust T U = if super then (T, U) else (U, T);
+        fun styp_test U Ts = forall
+          (fn T => T = U orelse Graph.is_edge coes_graph (adjust U T)) Ts;
+        fun fitting Ts S U = Type.of_sort tsig (t_of U, S) andalso styp_test U Ts
+      in
+        forall (fn (Ts, S) => exists (fitting Ts S) (T :: styps super T)) styps_and_sorts
+      end;
+
+    (* computes the tightest possible, correct assignment for 'a::S
+       e.g. in the supremum case (sup = true):
+               ------- 'a::S---
+              /        /    \  \
+             /        /      \  \
+        'b::C1   'c::C2 ...  T1 T2 ...
+
+       sorts - list of sorts [C1, C2, ...]
+       T::Ts - non-empty list of base types [T1, T2, ...]
+    *)
+    fun tightest sup S styps_and_sorts (T :: Ts) =
+      let
+        fun restriction T = Type.of_sort tsig (t_of T, S)
+          andalso ex_styp_of_sort (not sup) T styps_and_sorts;
+        fun candidates T = inter (op =) (filter restriction (T :: styps sup T));
+      in
+        (case fold candidates Ts (filter restriction (T :: styps sup T)) of
+          [] => raise BOUND_ERROR ("No " ^ (if sup then "supremum" else "infimum"))
+        | [T] => t_of T
+        | Ts => minmax sup Ts)
+      end;
+
+    fun build_graph G [] tye_idx = (G, tye_idx)
+      | build_graph G ((T, U) :: cs) tye_idx =
+        if T = U then build_graph G cs tye_idx
+        else
+          let
+            val G' = maybe_new_typnodes [T, U] G;
+            val (G'', tye_idx') = (add_edge (T, U) G', tye_idx)
+              handle Typ_Graph.CYCLES cycles =>
+                let
+                  val (tye, idx) = fold unify_list cycles tye_idx
+                in
+                  (*all cycles collapse to one node,
+                    because all of them share at least the nodes x and y*)
+                  collapse (tye, idx) (distinct (op =) (flat cycles)) G
+                end;
+          in
+            build_graph G'' cs tye_idx'
+          end
+    and collapse (tye, idx) nodes G = (*nodes non-empty list*)
+      let
+        val T = hd nodes;
+        val P = new_imm_preds G nodes;
+        val S = new_imm_succs G nodes;
+        val G' = Typ_Graph.del_nodes (tl nodes) G;
+      in
+        build_graph G' (map (fn x => (x, T)) P @ map (fn x => (T, x)) S) (tye, idx)
+      end;
+
+    fun assign_bound lower G key (tye_idx as (tye, _)) =
+      if Type_Infer.is_paramT (Type_Infer.deref tye key) then
+        let
+          val TVar (xi, S) = Type_Infer.deref tye key;
+          val get_bound = if lower then get_preds else get_succs;
+          val raw_bound = get_bound G key;
+          val bound = map (Type_Infer.deref tye) raw_bound;
+          val not_params = filter_out Type_Infer.is_paramT bound;
+          fun to_fulfil T =
+            (case sort_of T of
+              NONE => NONE
+            | SOME S =>
+                SOME
+                  (map nameT
+                    (filter_out Type_Infer.is_paramT (map (Type_Infer.deref tye) (get_bound G T))),
+                      S));
+          val styps_and_sorts = distinct (op =) (map_filter to_fulfil raw_bound);
+          val assignment =
+            if null bound orelse null not_params then NONE
+            else SOME (tightest lower S styps_and_sorts (map nameT not_params)
+                handle BOUND_ERROR msg => err_bound ctxt msg tye (find_error_pack lower key))
+        in
+          (case assignment of
+            NONE => tye_idx
+          | SOME T =>
+              if Type_Infer.is_paramT T then tye_idx
+              else if lower then (*upper bound check*)
+                let
+                  val other_bound = map (Type_Infer.deref tye) (get_succs G key);
+                  val s = nameT T;
+                in
+                  if subset (op = o apfst nameT) (filter is_typeT other_bound, s :: styps true s)
+                  then apfst (Vartab.update (xi, T)) tye_idx
+                  else err_bound ctxt ("Assigned simple type " ^ s ^
+                    " clashes with the upper bound of variable " ^
+                    Syntax.string_of_typ ctxt (TVar(xi, S))) tye (find_error_pack (not lower) key)
+                end
+              else apfst (Vartab.update (xi, T)) tye_idx)
+        end
+      else tye_idx;
+
+    val assign_lb = assign_bound true;
+    val assign_ub = assign_bound false;
+
+    fun assign_alternating ts' ts G tye_idx =
+      if ts' = ts then tye_idx
+      else
+        let
+          val (tye_idx' as (tye, _)) = fold (assign_lb G) ts tye_idx
+            |> fold (assign_ub G) ts;
+        in
+          assign_alternating ts (filter (Type_Infer.is_paramT o Type_Infer.deref tye) ts) G tye_idx'
+        end;
+
+    (*Unify all weakly connected components of the constraint forest,
+      that contain only params. These are the only WCCs that contain
+      params anyway.*)
+    fun unify_params G (tye_idx as (tye, _)) =
+      let
+        val max_params =
+          filter (Type_Infer.is_paramT o Type_Infer.deref tye) (Typ_Graph.maximals G);
+        val to_unify = map (fn T => T :: get_preds G T) max_params;
+      in
+        fold unify_list to_unify tye_idx
+      end;
+
+    fun solve_constraints G tye_idx = tye_idx
+      |> assign_alternating [] (Typ_Graph.keys G) G
+      |> unify_params G;
+  in
+    build_graph Typ_Graph.empty (map fst cs') tye_idx'
+      |-> solve_constraints
+  end;
+
+
+
+(** coercion insertion **)
+
+fun insert_coercions ctxt tye ts =
+  let
+    fun deep_deref T =
+      (case Type_Infer.deref tye T of
+        Type (a, Ts) => Type (a, map deep_deref Ts)
+      | U => U);
+
+    fun gen_coercion ((Type (a, [])), (Type (b, []))) =
+          if a = b
+          then Abs (Name.uu, Type (a, []), Bound 0)
+          else
+            (case Symreltab.lookup (coes_of ctxt) (a, b) of
+              NONE => raise Fail (a ^ " is not a subtype of " ^ b)
+            | SOME co => co)
+      | gen_coercion ((Type (a, Ts)), (Type (b, Us))) =
+          if a <> b
+          then raise raise Fail ("Different constructors: " ^ a ^ " and " ^ b)
+          else
+            let
+              fun inst t Ts =
+                Term.subst_vars
+                  (((Term.add_tvar_namesT (fastype_of t) []) ~~ rev Ts), []) t;
+              fun sub_co (COVARIANT, TU) = gen_coercion TU
+                | sub_co (CONTRAVARIANT, TU) = gen_coercion (swap TU);
+              fun ts_of [] = []
+                | ts_of (Type ("fun", [x1, x2]) :: xs) = x1 :: x2 :: (ts_of xs);
+            in
+              (case Symtab.lookup (tmaps_of ctxt) a of
+                NONE => raise Fail ("No map function for " ^ a ^ " known")
+              | SOME tmap =>
+                  let
+                    val used_coes = map sub_co ((snd tmap) ~~ (Ts ~~ Us));
+                  in
+                    Term.list_comb
+                      (inst (fst tmap) (ts_of (map fastype_of used_coes)), used_coes)
+                  end)
+            end
+      | gen_coercion (T, U) =
+          if Type.could_unify (T, U)
+          then Abs (Name.uu, T, Bound 0)
+          else raise Fail ("Cannot generate coercion from "
+            ^ Syntax.string_of_typ ctxt T ^ " to " ^ Syntax.string_of_typ ctxt U);
+
+    fun insert _ (Const (c, T)) =
+          let val T' = deep_deref T;
+          in (Const (c, T'), T') end
+      | insert _ (Free (x, T)) =
+          let val T' = deep_deref T;
+          in (Free (x, T'), T') end
+      | insert _ (Var (xi, T)) =
+          let val T' = deep_deref T;
+          in (Var (xi, T'), T') end
+      | insert bs (Bound i) =
+          let val T = nth bs i handle Subscript =>
+            raise TYPE ("Loose bound variable: B." ^ string_of_int i, [], []);
+          in (Bound i, T) end
+      | insert bs (Abs (x, T, t)) =
+          let
+            val T' = deep_deref T;
+            val (t', T'') = insert (T' :: bs) t;
+          in
+            (Abs (x, T', t'), T' --> T'')
+          end
+      | insert bs (t $ u) =
+          let
+            val (t', Type ("fun", [U, T])) = insert bs t;
+            val (u', U') = insert bs u;
+          in
+            if U <> U'
+            then (t' $ (gen_coercion (U', U) $ u'), T)
+            else (t' $ u', T)
+          end
+  in
+    map (fst o insert []) ts
+  end;
+
+
+
+(** assembling the pipeline **)
+
+fun infer_types ctxt const_type var_type raw_ts =
+  let
+    val (idx, ts) = Type_Infer.prepare ctxt const_type var_type raw_ts;
+
+    fun gen_all t (tye_idx, constraints) =
+      let
+        val (_, tye_idx', constraints') = generate_constraints ctxt t tye_idx
+      in (tye_idx', constraints' @ constraints) end;
+
+    val (tye_idx, constraints) = fold gen_all ts ((Vartab.empty, idx), []);
+    val (tye, _) = process_constraints ctxt constraints tye_idx;
+    val ts' = insert_coercions ctxt tye ts;
+
+    val (_, ts'') = Type_Infer.finish ctxt tye ([], ts');
+  in ts'' end;
+
+
+
+(** installation **)
+
+(* term check *)
+
+fun coercion_infer_types ctxt =
+  infer_types ctxt
+    (try (Consts.the_constraint (ProofContext.consts_of ctxt)))
+    (ProofContext.def_type ctxt);
+
+val add_term_check =
+  Syntax.add_term_check ~100 "coercions"
+    (fn xs => fn ctxt =>
+      let val xs' = coercion_infer_types ctxt xs
+      in if eq_list (op aconv) (xs, xs') then NONE else SOME (xs', ctxt) end);
+
+
+(* declarations *)
+
+fun add_type_map raw_t context =
+  let
+    val ctxt = Context.proof_of context;
+    val t = singleton (Variable.polymorphic ctxt) raw_t;
+
+    fun err_str () = "\n\nthe general type signature for a map function is" ^
+      "\nf1 => f2 => ... => fn => C [x1, ..., xn] => C [x1, ..., xn]" ^
+      "\nwhere C is a constructor and fi is of type (xi => yi) or (yi => xi)";
+
+    fun gen_arg_var ([], []) = []
+      | gen_arg_var ((T, T') :: Ts, (U, U') :: Us) =
+          if T = U andalso T' = U' then COVARIANT :: gen_arg_var (Ts, Us)
+          else if T = U' andalso T' = U then CONTRAVARIANT :: gen_arg_var (Ts, Us)
+          else error ("Functions do not apply to arguments correctly:" ^ err_str ())
+      | gen_arg_var (_, _) =
+          error ("Different numbers of functions and arguments\n" ^ err_str ());
+
+    (* TODO: This function is only needed to introde the fun type map
+      function: "% f g h . g o h o f". There must be a better solution. *)
+    fun balanced (Type (_, [])) (Type (_, [])) = true
+      | balanced (Type (a, Ts)) (Type (b, Us)) =
+          a = b andalso forall I (map2 balanced Ts Us)
+      | balanced (TFree _) (TFree _) = true
+      | balanced (TVar _) (TVar _) = true
+      | balanced _ _ = false;
+
+    fun check_map_fun (pairs, []) (Type ("fun", [T as Type (C, Ts), U as Type (_, Us)])) =
+          if balanced T U
+          then ((pairs, Ts ~~ Us), C)
+          else if C = "fun"
+            then check_map_fun (pairs @ [(hd Ts, hd (tl Ts))], []) U
+            else error ("Not a proper map function:" ^ err_str ())
+      | check_map_fun _ _ = error ("Not a proper map function:" ^ err_str ());
+
+    val res = check_map_fun ([], []) (fastype_of t);
+    val res_av = gen_arg_var (fst res);
+  in
+    map_tmaps (Symtab.update (snd res, (t, res_av))) context
+  end;
+
+fun add_coercion raw_t context =
+  let
+    val ctxt = Context.proof_of context;
+    val t = singleton (Variable.polymorphic ctxt) raw_t;
+
+    fun err_coercion () = error ("Bad type for coercion " ^
+        Syntax.string_of_term ctxt t ^ ":\n" ^
+        Syntax.string_of_typ ctxt (fastype_of t));
+
+    val (Type ("fun", [T1, T2])) = fastype_of t
+      handle Bind => err_coercion ();
+
+    val a =
+      (case T1 of
+        Type (x, []) => x
+      | _ => err_coercion ());
+
+    val b =
+      (case T2 of
+        Type (x, []) => x
+      | _ => err_coercion ());
+
+    fun coercion_data_update (tab, G) =
+      let
+        val G' = maybe_new_nodes [a, b] G
+        val G'' = Graph.add_edge_trans_acyclic (a, b) G'
+          handle Graph.CYCLES _ => error (a ^ " is already a subtype of " ^ b ^
+            "!\n\nCannot add coercion of type: " ^ a ^ " => " ^ b);
+        val new_edges =
+          flat (Graph.dest G'' |> map (fn (x, ys) => ys |> map_filter (fn y =>
+            if Graph.is_edge G' (x, y) then NONE else SOME (x, y))));
+        val G_and_new = Graph.add_edge (a, b) G';
+
+        fun complex_coercion tab G (a, b) =
+          let
+            val path = hd (Graph.irreducible_paths G (a, b))
+            val path' = (fst (split_last path)) ~~ tl path
+          in Abs (Name.uu, Type (a, []),
+              fold (fn t => fn u => t $ u) (map (the o Symreltab.lookup tab) path') (Bound 0))
+          end;
+
+        val tab' = fold
+          (fn pair => fn tab => Symreltab.update (pair, complex_coercion tab G_and_new pair) tab)
+          (filter (fn pair => pair <> (a, b)) new_edges)
+          (Symreltab.update ((a, b), t) tab);
+      in
+        (tab', G'')
+      end;
+  in
+    map_coes_and_graph coercion_data_update context
+  end;
+
+
+(* theory setup *)
+
+val setup =
+  Context.theory_map add_term_check #>
+  Attrib.setup @{binding coercion}
+    (Args.term >> (fn t => Thm.declaration_attribute (K (add_coercion t))))
+    "declaration of new coercions" #>
+  Attrib.setup @{binding map_function}
+    (Args.term >> (fn t => Thm.declaration_attribute (K (add_type_map t))))
+    "declaration of new map functions";
+
+end;
--- a/src/ZF/simpdata.ML	Fri Oct 29 17:15:28 2010 -0700
+++ b/src/ZF/simpdata.ML	Sat Oct 30 12:25:18 2010 -0700
@@ -59,10 +59,10 @@
 
 in
 
-val defBEX_regroup = Simplifier.simproc_global (Theory.deref @{theory_ref})
+val defBEX_regroup = Simplifier.simproc_global @{theory}
   "defined BEX" ["EX x:A. P(x) & Q(x)"] rearrange_bex;
 
-val defBALL_regroup = Simplifier.simproc_global (Theory.deref @{theory_ref})
+val defBALL_regroup = Simplifier.simproc_global @{theory}
   "defined BALL" ["ALL x:A. P(x) --> Q(x)"] rearrange_ball;
 
 end;