--- a/doc-src/Nitpick/nitpick.tex Sun Jan 31 15:22:40 2010 +0100
+++ b/doc-src/Nitpick/nitpick.tex Tue Feb 02 18:16:48 2010 +0100
@@ -25,8 +25,8 @@
\def\lparr{\mathopen{(\mkern-4mu\mid}}
\def\rparr{\mathclose{\mid\mkern-4mu)}}
-\def\undef{\textit{undefined}}
\def\unk{{?}}
+\def\undef{(\lambda x.\; \unk)}
%\def\unr{\textit{others}}
\def\unr{\ldots}
\def\Abs#1{\hbox{\rm{\flqq}}{\,#1\,}\hbox{\rm{\frqq}}}
@@ -188,8 +188,8 @@
\prew
\textbf{apply}~\textit{auto} \\[2\smallskipamount]
{\slshape goal (2 subgoals): \\
-\ 1. $P\,\Longrightarrow\, Q$ \\
-\ 2. $Q\,\Longrightarrow\, P$} \\[2\smallskipamount]
+\phantom{0}1. $P\,\Longrightarrow\, Q$ \\
+\phantom{0}2. $Q\,\Longrightarrow\, P$} \\[2\smallskipamount]
\textbf{nitpick}~1 \\[2\smallskipamount]
{\slshape Nitpick found a counterexample: \\[2\smallskipamount]
\hbox{}\qquad Free variables: \nopagebreak \\
@@ -317,11 +317,9 @@
\postw
Notice that $\textit{The}~(\lambda y.\;P~y) = \textit{The}~\{a_2, a_3\} = a_1$,
-just like before.\footnote{The \undef{} symbol's presence is explained as
-follows: In higher-order logic, any function can be built from the undefined
-function using repeated applications of the function update operator $f(x :=
-y)$, just like any list can be built from the empty list using $x \mathbin{\#}
-xs$.}
+just like before.\footnote{The Isabelle/HOL notation $f(x :=
+y)$ denotes the function that maps $x$ to $y$ and that otherwise behaves like
+$f$.}
Our misadventures with THE suggest adding `$\exists!x{.}$' (``there exists a
unique $x$ such that'') at the front of our putative lemma's assumption:
@@ -562,7 +560,7 @@
{\slshape Datatype:} \\
\hbox{}\qquad $'a$~\textit{list}~= $\{[],\, [a_3, a_3],\, [a_3],\, \unr\}$ \\
{\slshape Constants:} \\
-\hbox{}\qquad $\lambda x_1.\; x_1 \mathbin{@} [y, y] = \undef([] := [a_3, a_3],\> [a_3, a_3] := \unk,\> [a_3] := \unk)$ \\
+\hbox{}\qquad $\lambda x_1.\; x_1 \mathbin{@} [y, y] = \undef([] := [a_3, a_3])$ \\
\hbox{}\qquad $\textit{hd} = \undef([] := a_2,\> [a_3, a_3] := a_3,\> [a_3] := a_3)$
\postw
@@ -1243,6 +1241,208 @@
\textit{int~list}, \textit{rat}, and \textit{rat~list} will lead Nitpick to
consider only 8~scopes instead of $32\,768$.
+\subsection{Inductive Properties}
+\label{inductive-properties}
+
+Inductive properties are a particular pain to prove, because the failure to
+establish an induction step can mean several things:
+%
+\begin{enumerate}
+\item The property is invalid.
+\item The property is valid but is too weak to support the induction step.
+\item The property is valid and strong enough; it's just that we haven't found
+the proof yet.
+\end{enumerate}
+%
+Depending on which scenario applies, we would take the appropriate course of
+action:
+%
+\begin{enumerate}
+\item Repair the statement of the property so that it becomes valid.
+\item Generalize the property and/or prove auxiliary properties.
+\item Work harder on a proof.
+\end{enumerate}
+%
+How can we distinguish between the three scenarios? Nitpick's normal mode of
+operation can often detect scenario 1, and Isabelle's automatic tactics help with
+scenario 3. Using appropriate techniques, it is also often possible to use
+Nitpick to identify scenario 2. Consider the following transition system,
+in which natural numbers represent states:
+
+\prew
+\textbf{inductive\_set}~\textit{reach}~\textbf{where} \\
+``$(4\Colon\textit{nat}) \in \textit{reach\/}$'' $\mid$ \\
+``$\lbrakk n < 4;\> n \in \textit{reach\/}\rbrakk \,\Longrightarrow\, 3 * n + 1 \in \textit{reach\/}$'' $\mid$ \\
+``$n \in \textit{reach} \,\Longrightarrow n + 2 \in \textit{reach\/}$''
+\postw
+
+We will try to prove that only even numbers are reachable:
+
+\prew
+\textbf{lemma}~``$n \in \textit{reach} \,\Longrightarrow\, 2~\textrm{dvd}~n$''
+\postw
+
+Does this property hold? Nitpick cannot find a counterexample within 30 seconds,
+so let's attempt a proof by induction:
+
+\prew
+\textbf{apply}~(\textit{induct~set}{:}~\textit{reach\/}) \\
+\textbf{apply}~\textit{auto}
+\postw
+
+This leaves us in the following proof state:
+
+\prew
+{\slshape goal (2 subgoals): \\
+\phantom{0}1. ${\bigwedge}n.\;\, \lbrakk n \in \textit{reach\/};\, n < 4;\, 2~\textsl{dvd}~n\rbrakk \,\Longrightarrow\, 2~\textsl{dvd}~\textit{Suc}~(3 * n)$ \\
+\phantom{0}2. ${\bigwedge}n.\;\, \lbrakk n \in \textit{reach\/};\, 2~\textsl{dvd}~n\rbrakk \,\Longrightarrow\, 2~\textsl{dvd}~\textit{Suc}~(\textit{Suc}~n)$
+}
+\postw
+
+If we run Nitpick on the first subgoal, it still won't find any
+counterexample; and yet, \textit{auto} fails to go further, and \textit{arith}
+is helpless. However, notice the $n \in \textit{reach}$ assumption, which
+strengthens the induction hypothesis but is not immediately usable in the proof.
+If we remove it and invoke Nitpick, this time we get a counterexample:
+
+\prew
+\textbf{apply}~(\textit{thin\_tac}~``$n \in \textit{reach\/}$'') \\
+\textbf{nitpick} \\[2\smallskipamount]
+\slshape Nitpick found a counterexample: \\[2\smallskipamount]
+\hbox{}\qquad Skolem constant: \nopagebreak \\
+\hbox{}\qquad\qquad $n = 0$
+\postw
+
+Indeed, 0 < 4, 2 divides 0, but 2 does not divide 1. We can use this information
+to strength the lemma:
+
+\prew
+\textbf{lemma}~``$n \in \textit{reach} \,\Longrightarrow\, 2~\textrm{dvd}~n \mathrel{\lor} n \not= 0$''
+\postw
+
+Unfortunately, the proof by induction still gets stuck, except that Nitpick now
+finds the counterexample $n = 2$. We generalize the lemma further to
+
+\prew
+\textbf{lemma}~``$n \in \textit{reach} \,\Longrightarrow\, 2~\textrm{dvd}~n \mathrel{\lor} n \ge 4$''
+\postw
+
+and this time \textit{arith} can finish off the subgoals.
+
+A similar technique can be employed for structural induction. The
+following mini-formalization of full binary trees will serve as illustration:
+
+\prew
+\textbf{datatype} $\kern1pt'a$~\textit{bin\_tree} = $\textit{Leaf}~{\kern1pt'a}$ $\mid$ $\textit{Branch}$ ``\kern1pt$'a$ \textit{bin\_tree}'' ``\kern1pt$'a$ \textit{bin\_tree}'' \\[2\smallskipamount]
+\textbf{primrec}~\textit{labels}~\textbf{where} \\
+``$\textit{labels}~(\textit{Leaf}~a) = \{a\}$'' $\mid$ \\
+``$\textit{labels}~(\textit{Branch}~t~u) = \textit{labels}~t \mathrel{\cup} \textit{labels}~u$'' \\[2\smallskipamount]
+\textbf{primrec}~\textit{swap}~\textbf{where} \\
+``$\textit{swap}~(\textit{Leaf}~c)~a~b =$ \\
+\phantom{``}$(\textrm{if}~c = a~\textrm{then}~\textit{Leaf}~b~\textrm{else~if}~c = b~\textrm{then}~\textit{Leaf}~a~\textrm{else}~\textit{Leaf}~c)$'' $\mid$ \\
+``$\textit{swap}~(\textit{Branch}~t~u)~a~b = \textit{Branch}~(\textit{swap}~t~a~b)~(\textit{swap}~u~a~b)$''
+\postw
+
+The \textit{labels} function returns the set of labels occurring on leaves of a
+tree, and \textit{swap} exchanges two labels. Intuitively, if two distinct
+labels $a$ and $b$ occur in a tree $t$, they should also occur in the tree
+obtained by swapping $a$ and $b$:
+
+\prew
+\textbf{lemma} $``\lbrakk a \in \textit{labels}~t;\, b \in \textit{labels}~t;\, a \not= b\rbrakk {}$ \\
+\phantom{\textbf{lemma} ``}$\,{\Longrightarrow}{\;\,} \textit{labels}~(\textit{swap}~t~a~b) = \textit{labels}~t$''
+\postw
+
+Nitpick can't find any counterexample, so we proceed with induction
+(this time favoring a more structured style):
+
+\prew
+\textbf{proof}~(\textit{induct}~$t$) \\
+\hbox{}\quad \textbf{case}~\textit{Leaf}~\textbf{thus}~\textit{?case}~\textbf{by}~\textit{simp} \\
+\textbf{next} \\
+\hbox{}\quad \textbf{case}~$(\textit{Branch}~t~u)$~\textbf{thus} \textit{?case}
+\postw
+
+Nitpick can't find any counterexample at this point either, but it makes the
+following suggestion:
+
+\prew
+\slshape
+Hint: To check that the induction hypothesis is general enough, try the following command:
+\textbf{nitpick}~[\textit{non\_std} ``${\kern1pt'a}~\textit{bin\_tree}$'', \textit{show\_consts}].
+\postw
+
+If we follow the hint, we get a ``nonstandard'' counterexample for the step:
+
+\prew
+\slshape Nitpick found a nonstandard counterexample for \textit{card} $'a$ = 4: \\[2\smallskipamount]
+\hbox{}\qquad Free variables: \nopagebreak \\
+\hbox{}\qquad\qquad $a = a_4$ \\
+\hbox{}\qquad\qquad $b = a_3$ \\
+\hbox{}\qquad\qquad $t = \xi_3$ \\
+\hbox{}\qquad\qquad $u = \xi_4$ \\
+\hbox{}\qquad {\slshape Constants:} \nopagebreak \\
+\hbox{}\qquad\qquad $\textit{labels} = \undef
+ (\!\begin{aligned}[t]%
+ & \xi_3 := \{a_4\},\> \xi_4 := \{a_1, a_3\}, \\[-2pt] %% TYPESETTING
+ & \textit{Branch}~\xi_3~\xi_3 := \{a_4\}, \\[-2pt]
+ & \textit{Branch}~\xi_3~\xi_4 := \{a_1, a_3, a_4\})\end{aligned}$ \\
+\hbox{}\qquad\qquad $\lambda x_1.\> \textit{swap}~x_1~a~b = \undef
+ (\!\begin{aligned}[t]%
+ & \xi_3 := \xi_3,\> \xi_4 := \xi_3, \\[-2pt]
+ & \textit{Branch}~\xi_3~\xi_3 := \textit{Branch}~\xi_3~\xi_3, \\[-2pt]
+ & \textit{Branch}~\xi_4~\xi_3 := \textit{Branch}~\xi_3~\xi_3)\end{aligned}$ \\[2\smallskipamount]
+The existence of a nonstandard model suggests that the induction hypothesis is not general enough or perhaps
+even wrong. See the ``Inductive Properties'' section of the Nitpick manual for details (``\textit{isabelle doc nitpick}'').
+\postw
+
+Reading the Nitpick manual is a most excellent idea.
+But what's going on? The \textit{non\_std} ``${\kern1pt'a}~\textit{bin\_tree}$''
+option told the tool to look for nonstandard models of binary trees, which
+means that new ``nonstandard'' trees $\xi_1, \xi_2, \ldots$, are now allowed in
+addition to the standard trees generated by the \textit{Leaf} and
+\textit{Branch} constructors.%
+\footnote{Notice the similarity between allowing nonstandard trees here and
+allowing unreachable states in the preceding example (by removing the ``$n \in
+\textit{reach\/}$'' assumption). In both cases, we effectively enlarge the
+set of objects over which the induction is performed while doing the step
+so as to test the induction hypothesis's strength.}
+The new trees are so nonstandard that we know nothing about them, except what
+the induction hypothesis states and what can be proved about all trees without
+relying on induction or case distinction. The key observation is,
+%
+\begin{quote}
+\textsl{If the induction
+hypothesis is strong enough, the induction step will hold even for nonstandard
+objects, and Nitpick won't find any nonstandard counterexample.}
+\end{quote}
+%
+But here, Nitpick did find some nonstandard trees $t = \xi_3$
+and $u = \xi_4$ such that $a \in \textit{labels}~t$, $b \notin
+\textit{labels}~t$, $a \notin \textit{labels}~u$, and $b \in \textit{labels}~u$.
+Because neither tree contains both $a$ and $b$, the induction hypothesis tells
+us nothing about the labels of $\textit{swap}~t~a~b$ and $\textit{swap}~u~a~b$,
+and as a result we know nothing about the labels of the tree
+$\textit{swap}~(\textit{Branch}~t~u)~a~b$, which by definition equals
+$\textit{Branch}$ $(\textit{swap}~t~a~b)$ $(\textit{swap}~u~a~b)$, whose
+labels are $\textit{labels}$ $(\textit{swap}~t~a~b) \mathrel{\cup}
+\textit{labels}$ $(\textit{swap}~u~a~b)$.
+
+The solution is to ensure that we always know what the labels of the subtrees
+are in the inductive step, by covering the cases where $a$ and/or~$b$ is not in
+$t$ in the statement of the lemma:
+
+\prew
+\textbf{lemma} ``$\textit{labels}~(\textit{swap}~t~a~b) = {}$ \\
+\phantom{\textbf{lemma} ``}$(\textrm{if}~a \in \textit{labels}~t~\textrm{then}$ \nopagebreak \\
+\phantom{\textbf{lemma} ``(\quad}$\textrm{if}~b \in \textit{labels}~t~\textrm{then}~\textit{labels}~t~\textrm{else}~(\textit{labels}~t - \{a\}) \mathrel{\cup} \{b\}$ \\
+\phantom{\textbf{lemma} ``(}$\textrm{else}$ \\
+\phantom{\textbf{lemma} ``(\quad}$\textrm{if}~b \in \textit{labels}~t~\textrm{then}~(\textit{labels}~t - \{b\}) \mathrel{\cup} \{a\}~\textrm{else}~\textit{labels}~t)$''
+\postw
+
+This time, Nitpick won't find any nonstandard counterexample, and we can perform
+the induction step using \textbf{auto}.
+
\section{Case Studies}
\label{case-studies}
@@ -1398,7 +1598,8 @@
functions:
\prew
-\textbf{datatype} $'a$~\textit{tree} = $\Lambda$ $\mid$ $N$ ``\kern1pt$'a\Colon \textit{linorder}$'' \textit{nat} ``\kern1pt$'a$ \textit{tree}'' ``\kern1pt$'a$ \textit{tree}'' \\[2\smallskipamount]
+\textbf{datatype} $'a$~\textit{aa\_tree} = \\
+\hbox{}\quad $\Lambda$ $\mid$ $N$ ``\kern1pt$'a\Colon \textit{linorder}$'' \textit{nat} ``\kern1pt$'a$ \textit{aa\_tree}'' ``\kern1pt$'a$ \textit{aa\_tree}'' \\[2\smallskipamount]
\textbf{primrec} \textit{data} \textbf{where} \\
``$\textit{data}~\Lambda = \undef$'' $\,\mid$ \\
``$\textit{data}~(N~x~\_~\_~\_) = x$'' \\[2\smallskipamount]
@@ -1526,7 +1727,7 @@
\slshape Nitpick found a counterexample for \textit{card} $'a$ = 4: \\[2\smallskipamount]
\hbox{}\qquad Free variables: \nopagebreak \\
\hbox{}\qquad\qquad $t = N~a_3~1~\Lambda~\Lambda$ \\
-\hbox{}\qquad\qquad $x = a_4$ \\[2\smallskipamount]
+\hbox{}\qquad\qquad $x = a_4$
\postw
It's hard to see why this is a counterexample. To improve readability, we will
@@ -1583,10 +1784,11 @@
\def\optrue#1#2{\flushitem{\textit{#1} $\bigl[$= \qtybf{bool}$\bigr]$\quad [\textit{true}]\hfill (neg.: \textit{#2})}\nopagebreak\\[\parskip]}
\def\opfalse#1#2{\flushitem{\textit{#1} $\bigl[$= \qtybf{bool}$\bigr]$\quad [\textit{false}]\hfill (neg.: \textit{#2})}\nopagebreak\\[\parskip]}
\def\opsmart#1#2{\flushitem{\textit{#1} $\bigl[$= \qtybf{bool\_or\_smart}$\bigr]$\quad [\textit{smart}]\hfill (neg.: \textit{#2})}\nopagebreak\\[\parskip]}
-\def\ops#1#2{\flushitem{\textit{#1} = \qtybf{#2}} \nopagebreak\\[\parskip]}
-\def\opt#1#2#3{\flushitem{\textit{#1} = \qtybf{#2}\quad [\textit{#3}]} \nopagebreak\\[\parskip]}
-\def\opu#1#2#3{\flushitem{\textit{#1} \qtybf{#2} = \qtybf{#3}} \nopagebreak\\[\parskip]}
-\def\opusmart#1#2#3{\flushitem{\textit{#1} \qtybf{#2} $\bigl[$= \qtybf{bool\_or\_smart}$\bigr]$\hfill (neg.: \textit{#3})}\nopagebreak\\[\parskip]}
+\def\opnodefault#1#2{\flushitem{\textit{#1} = \qtybf{#2}} \nopagebreak\\[\parskip]}
+\def\opdefault#1#2#3{\flushitem{\textit{#1} = \qtybf{#2}\quad [\textit{#3}]} \nopagebreak\\[\parskip]}
+\def\oparg#1#2#3{\flushitem{\textit{#1} \qtybf{#2} = \qtybf{#3}} \nopagebreak\\[\parskip]}
+\def\opargbool#1#2#3{\flushitem{\textit{#1} \qtybf{#2} $\bigl[$= \qtybf{bool}$\bigr]$\hfill (neg.: \textit{#3})}\nopagebreak\\[\parskip]}
+\def\opargboolorsmart#1#2#3{\flushitem{\textit{#1} \qtybf{#2} $\bigl[$= \qtybf{bool\_or\_smart}$\bigr]$\hfill (neg.: \textit{#3})}\nopagebreak\\[\parskip]}
Nitpick's behavior can be influenced by various options, which can be specified
in brackets after the \textbf{nitpick} command. Default values can be set
@@ -1696,7 +1898,7 @@
\label{scope-of-search}
\begin{enum}
-\opu{card}{type}{int\_seq}
+\oparg{card}{type}{int\_seq}
Specifies the sequence of cardinalities to use for a given type.
For free types, and often also for \textbf{typedecl}'d types, it usually makes
sense to specify cardinalities as a range of the form \textit{$1$--$n$}.
@@ -1709,18 +1911,18 @@
\nopagebreak
{\small See also \textit{mono} (\S\ref{scope-of-search}).}
-\opt{card}{int\_seq}{$\mathbf{1}$--$\mathbf{8}$}
+\opdefault{card}{int\_seq}{$\mathbf{1}$--$\mathbf{8}$}
Specifies the default sequence of cardinalities to use. This can be overridden
on a per-type basis using the \textit{card}~\qty{type} option described above.
-\opu{max}{const}{int\_seq}
+\oparg{max}{const}{int\_seq}
Specifies the sequence of maximum multiplicities to use for a given
(co)in\-duc\-tive datatype constructor. A constructor's multiplicity is the
number of distinct values that it can construct. Nonsensical values (e.g.,
\textit{max}~[]~$=$~2) are silently repaired. This option is only available for
datatypes equipped with several constructors.
-\ops{max}{int\_seq}
+\opnodefault{max}{int\_seq}
Specifies the default sequence of maximum multiplicities to use for
(co)in\-duc\-tive datatype constructors. This can be overridden on a per-constructor
basis using the \textit{max}~\qty{const} option described above.
@@ -1757,13 +1959,13 @@
{\small See also \textit{bits} (\S\ref{scope-of-search}) and
\textit{show\_datatypes} (\S\ref{output-format}).}
-\opt{bits}{int\_seq}{$\mathbf{1},\mathbf{2},\mathbf{3},\mathbf{4},\mathbf{6},\mathbf{8},\mathbf{10},\mathbf{12}$}
+\opdefault{bits}{int\_seq}{$\mathbf{1},\mathbf{2},\mathbf{3},\mathbf{4},\mathbf{6},\mathbf{8},\mathbf{10},\mathbf{12}$}
Specifies the number of bits to use to represent natural numbers and integers in
binary, excluding the sign bit. The minimum is 1 and the maximum is 31.
{\small See also \textit{binary\_ints} (\S\ref{scope-of-search}).}
-\opusmart{wf}{const}{non\_wf}
+\opargboolorsmart{wf}{const}{non\_wf}
Specifies whether the specified (co)in\-duc\-tively defined predicate is
well-founded. The option can take the following values:
@@ -1780,7 +1982,7 @@
\item[$\bullet$] \textbf{\textit{smart}}: Try to prove that the inductive
predicate is well-founded using Isabelle's \textit{lexicographic\_order} and
-\textit{sizechange} tactics. If this succeeds (or the predicate occurs with an
+\textit{size\_change} tactics. If this succeeds (or the predicate occurs with an
appropriate polarity in the formula to falsify), use an efficient fixed point
equation as specification of the predicate; otherwise, unroll the predicates
according to the \textit{iter}~\qty{const} and \textit{iter} options.
@@ -1795,7 +1997,7 @@
Specifies the default wellfoundedness setting to use. This can be overridden on
a per-predicate basis using the \textit{wf}~\qty{const} option above.
-\opu{iter}{const}{int\_seq}
+\oparg{iter}{const}{int\_seq}
Specifies the sequence of iteration counts to use when unrolling a given
(co)in\-duc\-tive predicate. By default, unrolling is applied for inductive
predicates that occur negatively and coinductive predicates that occur
@@ -1807,12 +2009,12 @@
{\small See also \textit{wf} (\S\ref{scope-of-search}) and
\textit{star\_linear\_preds} (\S\ref{optimizations}).}
-\opt{iter}{int\_seq}{$\mathbf{1{,}2{,}4{,}8{,}12{,}16{,}24{,}32}$}
+\opdefault{iter}{int\_seq}{$\mathbf{1{,}2{,}4{,}8{,}12{,}16{,}24{,}32}$}
Specifies the sequence of iteration counts to use when unrolling (co)in\-duc\-tive
predicates. This can be overridden on a per-predicate basis using the
\textit{iter} \qty{const} option above.
-\opt{bisim\_depth}{int\_seq}{$\mathbf{7}$}
+\opdefault{bisim\_depth}{int\_seq}{$\mathbf{7}$}
Specifies the sequence of iteration counts to use when unrolling the
bisimilarity predicate generated by Nitpick for coinductive datatypes. A value
of $-1$ means that no predicate is generated, in which case Nitpick performs an
@@ -1822,7 +2024,7 @@
the sum of the cardinalities of the coinductive datatypes occurring in the
formula to falsify.
-\opusmart{box}{type}{dont\_box}
+\opargboolorsmart{box}{type}{dont\_box}
Specifies whether Nitpick should attempt to wrap (``box'') a given function or
product type in an isomorphic datatype internally. Boxing is an effective mean
to reduce the search space and speed up Nitpick, because the isomorphic datatype
@@ -1850,8 +2052,8 @@
Specifies the default boxing setting to use. This can be overridden on a
per-type basis using the \textit{box}~\qty{type} option described above.
-\opusmart{mono}{type}{non\_mono}
-Specifies whether the specified type should be considered monotonic when
+\opargboolorsmart{mono}{type}{non\_mono}
+Specifies whether the given type should be considered monotonic when
enumerating scopes. If the option is set to \textit{smart}, Nitpick performs a
monotonicity check on the type. Setting this option to \textit{true} can reduce
the number of scopes tried, but it also diminishes the theoretical chance of
@@ -1873,6 +2075,15 @@
theoretical chance of finding a counterexample.
{\small See also \textit{mono} (\S\ref{scope-of-search}).}
+
+\opargbool{std}{type}{non\_std}
+Specifies whether the given type should be given standard models.
+Nonstandard models are unsound but can help debug inductive arguments,
+as explained in \S\ref{inductive-properties}.
+
+\optrue{std}{non\_std}
+Specifies the default standardness to use. This can be overridden on a per-type
+basis using the \textit{std}~\qty{type} option described above.
\end{enum}
\subsection{Output Format}
@@ -1919,7 +2130,7 @@
Enabling this option effectively enables \textit{show\_skolems},
\textit{show\_datatypes}, and \textit{show\_consts}.
-\opt{max\_potential}{int}{$\mathbf{1}$}
+\opdefault{max\_potential}{int}{$\mathbf{1}$}
Specifies the maximum number of potential counterexamples to display. Setting
this option to 0 speeds up the search for a genuine counterexample. This option
is implicitly set to 0 for automatic runs. If you set this option to a value
@@ -1933,7 +2144,7 @@
{\small See also \textit{check\_potential} (\S\ref{authentication}) and
\textit{sat\_solver} (\S\ref{optimizations}).}
-\opt{max\_genuine}{int}{$\mathbf{1}$}
+\opdefault{max\_genuine}{int}{$\mathbf{1}$}
Specifies the maximum number of genuine counterexamples to display. If you set
this option to a value greater than 1, you will need an incremental SAT solver:
For efficiency, it is recommended to install the JNI version of MiniSat and set
@@ -1945,12 +2156,12 @@
{\small See also \textit{check\_genuine} (\S\ref{authentication}) and
\textit{sat\_solver} (\S\ref{optimizations}).}
-\ops{eval}{term\_list}
+\opnodefault{eval}{term\_list}
Specifies the list of terms whose values should be displayed along with
counterexamples. This option suffers from an ``observer effect'': Nitpick might
find different counterexamples for different values of this option.
-\opu{format}{term}{int\_seq}
+\oparg{format}{term}{int\_seq}
Specifies how to uncurry the value displayed for a variable or constant.
Uncurrying sometimes increases the readability of the output for high-arity
functions. For example, given the variable $y \mathbin{\Colon} {'a}\Rightarrow
@@ -1966,7 +2177,7 @@
\nopagebreak
{\small See also \textit{uncurry} (\S\ref{optimizations}).}
-\opt{format}{int\_seq}{$\mathbf{1}$}
+\opdefault{format}{int\_seq}{$\mathbf{1}$}
Specifies the default format to use. Irrespective of the default format, the
extra arguments to a Skolem constant corresponding to the outer bound variables
are kept separated from the remaining arguments, the \textbf{for} arguments of
@@ -1999,7 +2210,7 @@
\nopagebreak
{\small See also \textit{max\_genuine} (\S\ref{output-format}).}
-\ops{expect}{string}
+\opnodefault{expect}{string}
Specifies the expected outcome, which must be one of the following:
\begin{enum}
@@ -2025,7 +2236,7 @@
\sloppy
\begin{enum}
-\opt{sat\_solver}{string}{smart}
+\opdefault{sat\_solver}{string}{smart}
Specifies which SAT solver to use. SAT solvers implemented in C or \cpp{} tend
to be faster than their Java counterparts, but they can be more difficult to
install. Also, if you set the \textit{max\_potential} (\S\ref{output-format}) or
@@ -2118,7 +2329,7 @@
\end{enum}
\fussy
-\opt{batch\_size}{int\_or\_smart}{smart}
+\opdefault{batch\_size}{int\_or\_smart}{smart}
Specifies the maximum number of Kodkod problems that should be lumped together
when invoking Kodkodi. Each problem corresponds to one scope. Lumping problems
together ensures that Kodkodi is launched less often, but it makes the verbose
@@ -2188,7 +2399,7 @@
Unless you are tracking down a bug in Nitpick or distrust the peephole
optimizer, you should leave this option enabled.
-\opt{sym\_break}{int}{20}
+\opdefault{sym\_break}{int}{20}
Specifies an upper bound on the number of relations for which Kodkod generates
symmetry breaking predicates. According to the Kodkod documentation
\cite{kodkod-2009-options}, ``in general, the higher this value, the more
@@ -2196,7 +2407,7 @@
setting the value too high may have the opposite effect and slow down the
solving.''
-\opt{sharing\_depth}{int}{3}
+\opdefault{sharing\_depth}{int}{3}
Specifies the depth to which Kodkod should check circuits for equivalence during
the translation to SAT. The default of 3 is the same as in Alloy. The minimum
allowed depth is 1. Increasing the sharing may result in a smaller SAT problem,
@@ -2207,7 +2418,7 @@
Although this might sound like a good idea, in practice it can drastically slow
down Kodkod.
-\opt{max\_threads}{int}{0}
+\opdefault{max\_threads}{int}{0}
Specifies the maximum number of threads to use in Kodkod. If this option is set
to 0, Kodkod will compute an appropriate value based on the number of processor
cores available.
@@ -2221,7 +2432,7 @@
\label{timeouts}
\begin{enum}
-\opt{timeout}{time}{$\mathbf{30}$ s}
+\opdefault{timeout}{time}{$\mathbf{30}$ s}
Specifies the maximum amount of time that the \textbf{nitpick} command should
spend looking for a counterexample. Nitpick tries to honor this constraint as
well as it can but offers no guarantees. For automatic runs,
@@ -2232,10 +2443,10 @@
\nopagebreak
{\small See also \textit{max\_threads} (\S\ref{optimizations}).}
-\opt{tac\_timeout}{time}{$\mathbf{500}$\,ms}
+\opdefault{tac\_timeout}{time}{$\mathbf{500}$\,ms}
Specifies the maximum amount of time that the \textit{auto} tactic should use
when checking a counterexample, and similarly that \textit{lexicographic\_order}
-and \textit{sizechange} should use when checking whether a (co)in\-duc\-tive
+and \textit{size\_change} should use when checking whether a (co)in\-duc\-tive
predicate is well-founded. Nitpick tries to honor this constraint as well as it
can but offers no guarantees.
@@ -2477,7 +2688,7 @@
\nopagebreak
\\[2\smallskipamount]
\hbox{}\qquad Empty assignment} \nopagebreak\\[2\smallskipamount]
-\textbf{by}~(\textit{auto simp}: \textit{prec\_def})
+\textbf{by}~(\textit{auto simp}:~\textit{prec\_def})
\postw
Such theorems are considered bad style because they rely on the internal
--- a/src/HOL/Divides.thy Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/Divides.thy Tue Feb 02 18:16:48 2010 +0100
@@ -2451,7 +2451,8 @@
in subst [OF mod_div_equality [of _ n]])
arith
-lemmas [nitpick_def] = mod_div_equality' [THEN eq_reflection]
+lemmas [nitpick_def] = dvd_eq_mod_eq_0 [THEN eq_reflection]
+ mod_div_equality' [THEN eq_reflection]
zmod_zdiv_equality' [THEN eq_reflection]
subsubsection {* Code generation *}
--- a/src/HOL/Library/AssocList.thy Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/Library/AssocList.thy Tue Feb 02 18:16:48 2010 +0100
@@ -1,11 +1,11 @@
(* Title: HOL/Library/AssocList.thy
- Author: Norbert Schirmer, Tobias Nipkow, Martin Wildmoser
+ Author: Norbert Schirmer, Tobias Nipkow, Martin Wildmoser, TU Muenchen
*)
header {* Map operations implemented on association lists*}
theory AssocList
-imports Map Main
+imports Main
begin
text {*
@@ -15,318 +15,42 @@
to establish the invariant, e.g. for inductive proofs.
*}
-primrec
- delete :: "'key \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
-where
- "delete k [] = []"
- | "delete k (p#ps) = (if fst p = k then delete k ps else p # delete k ps)"
+subsection {* @{text update} and @{text updates} *}
-primrec
- update :: "'key \<Rightarrow> 'val \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
-where
+primrec update :: "'key \<Rightarrow> 'val \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list" where
"update k v [] = [(k, v)]"
| "update k v (p#ps) = (if fst p = k then (k, v) # ps else p # update k v ps)"
-primrec
- updates :: "'key list \<Rightarrow> 'val list \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
-where
- "updates [] vs ps = ps"
- | "updates (k#ks) vs ps = (case vs
- of [] \<Rightarrow> ps
- | (v#vs') \<Rightarrow> updates ks vs' (update k v ps))"
-
-primrec
- merge :: "('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
-where
- "merge qs [] = qs"
- | "merge qs (p#ps) = update (fst p) (snd p) (merge qs ps)"
-
-lemma length_delete_le: "length (delete k al) \<le> length al"
-proof (induct al)
- case Nil thus ?case by simp
-next
- case (Cons a al)
- note length_filter_le [of "\<lambda>p. fst p \<noteq> fst a" al]
- also have "\<And>n. n \<le> Suc n"
- by simp
- finally have "length [p\<leftarrow>al . fst p \<noteq> fst a] \<le> Suc (length al)" .
- with Cons show ?case
- by auto
-qed
-
-lemma compose_hint [simp]:
- "length (delete k al) < Suc (length al)"
-proof -
- note length_delete_le
- also have "\<And>n. n < Suc n"
- by simp
- finally show ?thesis .
-qed
-
-fun
- compose :: "('key \<times> 'a) list \<Rightarrow> ('a \<times> 'b) list \<Rightarrow> ('key \<times> 'b) list"
-where
- "compose [] ys = []"
- | "compose (x#xs) ys = (case map_of ys (snd x)
- of None \<Rightarrow> compose (delete (fst x) xs) ys
- | Some v \<Rightarrow> (fst x, v) # compose xs ys)"
-
-primrec
- restrict :: "'key set \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
-where
- "restrict A [] = []"
- | "restrict A (p#ps) = (if fst p \<in> A then p#restrict A ps else restrict A ps)"
-
-primrec
- map_ran :: "('key \<Rightarrow> 'val \<Rightarrow> 'val) \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
-where
- "map_ran f [] = []"
- | "map_ran f (p#ps) = (fst p, f (fst p) (snd p)) # map_ran f ps"
-
-fun
- clearjunk :: "('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list"
-where
- "clearjunk [] = []"
- | "clearjunk (p#ps) = p # clearjunk (delete (fst p) ps)"
-
-lemmas [simp del] = compose_hint
-
-
-subsection {* @{const delete} *}
-
-lemma delete_eq:
- "delete k xs = filter (\<lambda>p. fst p \<noteq> k) xs"
- by (induct xs) auto
-
-lemma delete_id [simp]: "k \<notin> fst ` set al \<Longrightarrow> delete k al = al"
- by (induct al) auto
-
-lemma delete_conv: "map_of (delete k al) k' = ((map_of al)(k := None)) k'"
- by (induct al) auto
-
-lemma delete_conv': "map_of (delete k al) = ((map_of al)(k := None))"
- by (rule ext) (rule delete_conv)
-
-lemma delete_idem: "delete k (delete k al) = delete k al"
- by (induct al) auto
-
-lemma map_of_delete [simp]:
- "k' \<noteq> k \<Longrightarrow> map_of (delete k al) k' = map_of al k'"
- by (induct al) auto
-
-lemma delete_notin_dom: "k \<notin> fst ` set (delete k al)"
- by (induct al) auto
-
-lemma dom_delete_subset: "fst ` set (delete k al) \<subseteq> fst ` set al"
- by (induct al) auto
-
-lemma distinct_delete:
- assumes "distinct (map fst al)"
- shows "distinct (map fst (delete k al))"
-using assms
-proof (induct al)
- case Nil thus ?case by simp
-next
- case (Cons a al)
- from Cons.prems obtain
- a_notin_al: "fst a \<notin> fst ` set al" and
- dist_al: "distinct (map fst al)"
- by auto
- show ?case
- proof (cases "fst a = k")
- case True
- with Cons dist_al show ?thesis by simp
- next
- case False
- from dist_al
- have "distinct (map fst (delete k al))"
- by (rule Cons.hyps)
- moreover from a_notin_al dom_delete_subset [of k al]
- have "fst a \<notin> fst ` set (delete k al)"
- by blast
- ultimately show ?thesis using False by simp
- qed
-qed
-
-lemma delete_twist: "delete x (delete y al) = delete y (delete x al)"
- by (induct al) auto
-
-lemma clearjunk_delete: "clearjunk (delete x al) = delete x (clearjunk al)"
- by (induct al rule: clearjunk.induct) (auto simp add: delete_idem delete_twist)
-
+lemma update_conv': "map_of (update k v al) = (map_of al)(k\<mapsto>v)"
+ by (induct al) (auto simp add: expand_fun_eq)
-subsection {* @{const clearjunk} *}
-
-lemma insert_fst_filter:
- "insert a(fst ` {x \<in> set ps. fst x \<noteq> a}) = insert a (fst ` set ps)"
- by (induct ps) auto
-
-lemma dom_clearjunk: "fst ` set (clearjunk al) = fst ` set al"
- by (induct al rule: clearjunk.induct) (simp_all add: insert_fst_filter delete_eq)
-
-lemma notin_filter_fst: "a \<notin> fst ` {x \<in> set ps. fst x \<noteq> a}"
- by (induct ps) auto
-
-lemma distinct_clearjunk [simp]: "distinct (map fst (clearjunk al))"
- by (induct al rule: clearjunk.induct)
- (simp_all add: dom_clearjunk notin_filter_fst delete_eq)
-
-lemma map_of_filter: "k \<noteq> a \<Longrightarrow> map_of [q\<leftarrow>ps . fst q \<noteq> a] k = map_of ps k"
- by (induct ps) auto
-
-lemma map_of_clearjunk: "map_of (clearjunk al) = map_of al"
- apply (rule ext)
- apply (induct al rule: clearjunk.induct)
- apply simp
- apply (simp add: map_of_filter)
- done
-
-lemma length_clearjunk: "length (clearjunk al) \<le> length al"
-proof (induct al rule: clearjunk.induct [case_names Nil Cons])
- case Nil thus ?case by simp
-next
- case (Cons p ps)
- from Cons have "length (clearjunk [q\<leftarrow>ps . fst q \<noteq> fst p]) \<le> length [q\<leftarrow>ps . fst q \<noteq> fst p]"
- by (simp add: delete_eq)
- also have "\<dots> \<le> length ps"
- by simp
- finally show ?case
- by (simp add: delete_eq)
-qed
-
-lemma notin_fst_filter: "a \<notin> fst ` set ps \<Longrightarrow> [q\<leftarrow>ps . fst q \<noteq> a] = ps"
- by (induct ps) auto
-
-lemma distinct_clearjunk_id [simp]: "distinct (map fst al) \<Longrightarrow> clearjunk al = al"
- by (induct al rule: clearjunk.induct) (auto simp add: notin_fst_filter)
-
-lemma clearjunk_idem: "clearjunk (clearjunk al) = clearjunk al"
- by simp
-
-
-subsection {* @{const dom} and @{term "ran"} *}
-
-lemma dom_map_of': "fst ` set al = dom (map_of al)"
- by (induct al) auto
-
-lemmas dom_map_of = dom_map_of' [symmetric]
-
-lemma ran_clearjunk: "ran (map_of (clearjunk al)) = ran (map_of al)"
- by (simp add: map_of_clearjunk)
-
-lemma ran_distinct:
- assumes dist: "distinct (map fst al)"
- shows "ran (map_of al) = snd ` set al"
-using dist
-proof (induct al)
- case Nil
- thus ?case by simp
-next
- case (Cons a al)
- hence hyp: "snd ` set al = ran (map_of al)"
- by simp
-
- have "ran (map_of (a # al)) = {snd a} \<union> ran (map_of al)"
- proof
- show "ran (map_of (a # al)) \<subseteq> {snd a} \<union> ran (map_of al)"
- proof
- fix v
- assume "v \<in> ran (map_of (a#al))"
- then obtain x where "map_of (a#al) x = Some v"
- by (auto simp add: ran_def)
- then show "v \<in> {snd a} \<union> ran (map_of al)"
- by (auto split: split_if_asm simp add: ran_def)
- qed
- next
- show "{snd a} \<union> ran (map_of al) \<subseteq> ran (map_of (a # al))"
- proof
- fix v
- assume v_in: "v \<in> {snd a} \<union> ran (map_of al)"
- show "v \<in> ran (map_of (a#al))"
- proof (cases "v=snd a")
- case True
- with v_in show ?thesis
- by (auto simp add: ran_def)
- next
- case False
- with v_in have "v \<in> ran (map_of al)" by auto
- then obtain x where al_x: "map_of al x = Some v"
- by (auto simp add: ran_def)
- from map_of_SomeD [OF this]
- have "x \<in> fst ` set al"
- by (force simp add: image_def)
- with Cons.prems have "x\<noteq>fst a"
- by - (rule ccontr,simp)
- with al_x
- show ?thesis
- by (auto simp add: ran_def)
- qed
- qed
- qed
- with hyp show ?case
- by (simp only:) auto
-qed
-
-lemma ran_map_of: "ran (map_of al) = snd ` set (clearjunk al)"
-proof -
- have "ran (map_of al) = ran (map_of (clearjunk al))"
- by (simp add: ran_clearjunk)
- also have "\<dots> = snd ` set (clearjunk al)"
- by (simp add: ran_distinct)
- finally show ?thesis .
-qed
-
-
-subsection {* @{const update} *}
-
-lemma update_conv: "map_of (update k v al) k' = ((map_of al)(k\<mapsto>v)) k'"
- by (induct al) auto
-
-lemma update_conv': "map_of (update k v al) = ((map_of al)(k\<mapsto>v))"
- by (rule ext) (rule update_conv)
+corollary update_conv: "map_of (update k v al) k' = ((map_of al)(k\<mapsto>v)) k'"
+ by (simp add: update_conv')
lemma dom_update: "fst ` set (update k v al) = {k} \<union> fst ` set al"
by (induct al) auto
+lemma update_keys:
+ "map fst (update k v al) =
+ (if k \<in> set (map fst al) then map fst al else map fst al @ [k])"
+ by (induct al) simp_all
+
lemma distinct_update:
assumes "distinct (map fst al)"
shows "distinct (map fst (update k v al))"
-using assms
-proof (induct al)
- case Nil thus ?case by simp
-next
- case (Cons a al)
- from Cons.prems obtain
- a_notin_al: "fst a \<notin> fst ` set al" and
- dist_al: "distinct (map fst al)"
- by auto
- show ?case
- proof (cases "fst a = k")
- case True
- from True dist_al a_notin_al show ?thesis by simp
- next
- case False
- from dist_al
- have "distinct (map fst (update k v al))"
- by (rule Cons.hyps)
- with False a_notin_al show ?thesis by (simp add: dom_update)
- qed
-qed
+ using assms by (simp add: update_keys)
lemma update_filter:
"a\<noteq>k \<Longrightarrow> update k v [q\<leftarrow>ps . fst q \<noteq> a] = [q\<leftarrow>update k v ps . fst q \<noteq> a]"
by (induct ps) auto
-lemma clearjunk_update: "clearjunk (update k v al) = update k v (clearjunk al)"
- by (induct al rule: clearjunk.induct) (auto simp add: update_filter delete_eq)
-
lemma update_triv: "map_of al k = Some v \<Longrightarrow> update k v al = al"
by (induct al) auto
lemma update_nonempty [simp]: "update k v al \<noteq> []"
by (induct al) auto
-lemma update_eqD: "update k v al = update k v' al' \<Longrightarrow> v=v'"
+lemma update_eqD: "update k v al = update k v' al' \<Longrightarrow> v = v'"
proof (induct al arbitrary: al')
case Nil thus ?case
by (cases al') (auto split: split_if_asm)
@@ -339,59 +63,61 @@
by (induct al) auto
text {* Note that the lists are not necessarily the same:
- @{term "update k v (update k' v' []) = [(k',v'),(k,v)]"} and
- @{term "update k' v' (update k v []) = [(k,v),(k',v')]"}.*}
+ @{term "update k v (update k' v' []) = [(k', v'), (k, v)]"} and
+ @{term "update k' v' (update k v []) = [(k, v), (k', v')]"}.*}
lemma update_swap: "k\<noteq>k'
\<Longrightarrow> map_of (update k v (update k' v' al)) = map_of (update k' v' (update k v al))"
- by (auto simp add: update_conv' intro: ext)
+ by (simp add: update_conv' expand_fun_eq)
lemma update_Some_unfold:
- "(map_of (update k v al) x = Some y) =
- (x = k \<and> v = y \<or> x \<noteq> k \<and> map_of al x = Some y)"
+ "map_of (update k v al) x = Some y \<longleftrightarrow>
+ x = k \<and> v = y \<or> x \<noteq> k \<and> map_of al x = Some y"
by (simp add: update_conv' map_upd_Some_unfold)
-lemma image_update[simp]: "x \<notin> A \<Longrightarrow> map_of (update x y al) ` A = map_of al ` A"
+lemma image_update [simp]:
+ "x \<notin> A \<Longrightarrow> map_of (update x y al) ` A = map_of al ` A"
by (simp add: update_conv' image_map_upd)
+definition updates :: "'key list \<Rightarrow> 'val list \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list" where
+ "updates ks vs al = foldl (\<lambda>al (k, v). update k v al) al (zip ks vs)"
-subsection {* @{const updates} *}
+lemma updates_simps [simp]:
+ "updates [] vs ps = ps"
+ "updates ks [] ps = ps"
+ "updates (k#ks) (v#vs) ps = updates ks vs (update k v ps)"
+ by (simp_all add: updates_def)
+
+lemma updates_key_simp [simp]:
+ "updates (k # ks) vs ps =
+ (case vs of [] \<Rightarrow> ps | v # vs \<Rightarrow> updates ks vs (update k v ps))"
+ by (cases vs) simp_all
+
+lemma updates_conv': "map_of (updates ks vs al) = (map_of al)(ks[\<mapsto>]vs)"
+proof -
+ have "foldl (\<lambda>f (k, v). f(k \<mapsto> v)) (map_of al) (zip ks vs) =
+ map_of (foldl (\<lambda>al (k, v). update k v al) al (zip ks vs))"
+ by (rule foldl_apply) (auto simp add: expand_fun_eq update_conv')
+ then show ?thesis
+ by (simp add: updates_def map_upds_fold_map_upd)
+qed
lemma updates_conv: "map_of (updates ks vs al) k = ((map_of al)(ks[\<mapsto>]vs)) k"
-proof (induct ks arbitrary: vs al)
- case Nil
- thus ?case by simp
-next
- case (Cons k ks)
- show ?case
- proof (cases vs)
- case Nil
- with Cons show ?thesis by simp
- next
- case (Cons k ks')
- with Cons.hyps show ?thesis
- by (simp add: update_conv fun_upd_def)
- qed
-qed
-
-lemma updates_conv': "map_of (updates ks vs al) = ((map_of al)(ks[\<mapsto>]vs))"
- by (rule ext) (rule updates_conv)
+ by (simp add: updates_conv')
lemma distinct_updates:
assumes "distinct (map fst al)"
shows "distinct (map fst (updates ks vs al))"
- using assms
- by (induct ks arbitrary: vs al)
- (auto simp add: distinct_update split: list.splits)
-
-lemma clearjunk_updates:
- "clearjunk (updates ks vs al) = updates ks vs (clearjunk al)"
- by (induct ks arbitrary: vs al) (auto simp add: clearjunk_update split: list.splits)
-
-lemma updates_empty[simp]: "updates vs [] al = al"
- by (induct vs) auto
-
-lemma updates_Cons: "updates (k#ks) (v#vs) al = updates ks vs (update k v al)"
- by simp
+proof -
+ from assms have "distinct (foldl
+ (\<lambda>al (k, v). if k \<in> set al then al else al @ [k])
+ (map fst al) (zip ks vs))"
+ by (rule foldl_invariant) auto
+ moreover have "foldl (\<lambda>al (k, v). if k \<in> set al then al else al @ [k])
+ (map fst al) (zip ks vs)
+ = map fst (foldl (\<lambda>al (k, v). update k v al) al (zip ks vs))"
+ by (rule foldl_apply) (simp add: update_keys split_def comp_def)
+ ultimately show ?thesis by (simp add: updates_def)
+qed
lemma updates_append1[simp]: "size ks < size vs \<Longrightarrow>
updates (ks@[k]) vs al = update k (vs!size ks) (updates ks vs al)"
@@ -426,25 +152,285 @@
by (induct xs arbitrary: ys al) (auto split: list.splits)
-subsection {* @{const map_ran} *}
+subsection {* @{text delete} *}
+
+definition delete :: "'key \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list" where
+ delete_eq: "delete k = filter (\<lambda>(k', _). k \<noteq> k')"
+
+lemma delete_simps [simp]:
+ "delete k [] = []"
+ "delete k (p#ps) = (if fst p = k then delete k ps else p # delete k ps)"
+ by (auto simp add: delete_eq)
+
+lemma delete_conv': "map_of (delete k al) = (map_of al)(k := None)"
+ by (induct al) (auto simp add: expand_fun_eq)
+
+corollary delete_conv: "map_of (delete k al) k' = ((map_of al)(k := None)) k'"
+ by (simp add: delete_conv')
+
+lemma delete_keys:
+ "map fst (delete k al) = removeAll k (map fst al)"
+ by (simp add: delete_eq removeAll_filter_not_eq filter_map split_def comp_def)
+
+lemma distinct_delete:
+ assumes "distinct (map fst al)"
+ shows "distinct (map fst (delete k al))"
+ using assms by (simp add: delete_keys distinct_removeAll)
+
+lemma delete_id [simp]: "k \<notin> fst ` set al \<Longrightarrow> delete k al = al"
+ by (auto simp add: image_iff delete_eq filter_id_conv)
+
+lemma delete_idem: "delete k (delete k al) = delete k al"
+ by (simp add: delete_eq)
+
+lemma map_of_delete [simp]:
+ "k' \<noteq> k \<Longrightarrow> map_of (delete k al) k' = map_of al k'"
+ by (simp add: delete_conv')
+
+lemma delete_notin_dom: "k \<notin> fst ` set (delete k al)"
+ by (auto simp add: delete_eq)
+
+lemma dom_delete_subset: "fst ` set (delete k al) \<subseteq> fst ` set al"
+ by (auto simp add: delete_eq)
+
+lemma delete_update_same:
+ "delete k (update k v al) = delete k al"
+ by (induct al) simp_all
+
+lemma delete_update:
+ "k \<noteq> l \<Longrightarrow> delete l (update k v al) = update k v (delete l al)"
+ by (induct al) simp_all
+
+lemma delete_twist: "delete x (delete y al) = delete y (delete x al)"
+ by (simp add: delete_eq conj_commute)
+
+lemma length_delete_le: "length (delete k al) \<le> length al"
+ by (simp add: delete_eq)
+
+
+subsection {* @{text restrict} *}
+
+definition restrict :: "'key set \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list" where
+ restrict_eq: "restrict A = filter (\<lambda>(k, v). k \<in> A)"
+
+lemma restr_simps [simp]:
+ "restrict A [] = []"
+ "restrict A (p#ps) = (if fst p \<in> A then p # restrict A ps else restrict A ps)"
+ by (auto simp add: restrict_eq)
+
+lemma restr_conv': "map_of (restrict A al) = ((map_of al)|` A)"
+proof
+ fix k
+ show "map_of (restrict A al) k = ((map_of al)|` A) k"
+ by (induct al) (simp, cases "k \<in> A", auto)
+qed
+
+corollary restr_conv: "map_of (restrict A al) k = ((map_of al)|` A) k"
+ by (simp add: restr_conv')
+
+lemma distinct_restr:
+ "distinct (map fst al) \<Longrightarrow> distinct (map fst (restrict A al))"
+ by (induct al) (auto simp add: restrict_eq)
+
+lemma restr_empty [simp]:
+ "restrict {} al = []"
+ "restrict A [] = []"
+ by (induct al) (auto simp add: restrict_eq)
+
+lemma restr_in [simp]: "x \<in> A \<Longrightarrow> map_of (restrict A al) x = map_of al x"
+ by (simp add: restr_conv')
+
+lemma restr_out [simp]: "x \<notin> A \<Longrightarrow> map_of (restrict A al) x = None"
+ by (simp add: restr_conv')
+
+lemma dom_restr [simp]: "fst ` set (restrict A al) = fst ` set al \<inter> A"
+ by (induct al) (auto simp add: restrict_eq)
+
+lemma restr_upd_same [simp]: "restrict (-{x}) (update x y al) = restrict (-{x}) al"
+ by (induct al) (auto simp add: restrict_eq)
+
+lemma restr_restr [simp]: "restrict A (restrict B al) = restrict (A\<inter>B) al"
+ by (induct al) (auto simp add: restrict_eq)
+
+lemma restr_update[simp]:
+ "map_of (restrict D (update x y al)) =
+ map_of ((if x \<in> D then (update x y (restrict (D-{x}) al)) else restrict D al))"
+ by (simp add: restr_conv' update_conv')
+
+lemma restr_delete [simp]:
+ "(delete x (restrict D al)) =
+ (if x \<in> D then restrict (D - {x}) al else restrict D al)"
+apply (simp add: delete_eq restrict_eq)
+apply (auto simp add: split_def)
+proof -
+ have "\<And>y. y \<noteq> x \<longleftrightarrow> x \<noteq> y" by auto
+ then show "[p \<leftarrow> al. fst p \<in> D \<and> x \<noteq> fst p] = [p \<leftarrow> al. fst p \<in> D \<and> fst p \<noteq> x]"
+ by simp
+ assume "x \<notin> D"
+ then have "\<And>y. y \<in> D \<longleftrightarrow> y \<in> D \<and> x \<noteq> y" by auto
+ then show "[p \<leftarrow> al . fst p \<in> D \<and> x \<noteq> fst p] = [p \<leftarrow> al . fst p \<in> D]"
+ by simp
+qed
+
+lemma update_restr:
+ "map_of (update x y (restrict D al)) = map_of (update x y (restrict (D-{x}) al))"
+ by (simp add: update_conv' restr_conv') (rule fun_upd_restrict)
-lemma map_ran_conv: "map_of (map_ran f al) k = Option.map (f k) (map_of al k)"
- by (induct al) auto
+lemma upate_restr_conv [simp]:
+ "x \<in> D \<Longrightarrow>
+ map_of (update x y (restrict D al)) = map_of (update x y (restrict (D-{x}) al))"
+ by (simp add: update_conv' restr_conv')
+
+lemma restr_updates [simp]: "
+ \<lbrakk> length xs = length ys; set xs \<subseteq> D \<rbrakk>
+ \<Longrightarrow> map_of (restrict D (updates xs ys al)) =
+ map_of (updates xs ys (restrict (D - set xs) al))"
+ by (simp add: updates_conv' restr_conv')
+
+lemma restr_delete_twist: "(restrict A (delete a ps)) = delete a (restrict A ps)"
+ by (induct ps) auto
+
+
+subsection {* @{text clearjunk} *}
+
+function clearjunk :: "('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list" where
+ "clearjunk [] = []"
+ | "clearjunk (p#ps) = p # clearjunk (delete (fst p) ps)"
+ by pat_completeness auto
+termination by (relation "measure length")
+ (simp_all add: less_Suc_eq_le length_delete_le)
+
+lemma map_of_clearjunk:
+ "map_of (clearjunk al) = map_of al"
+ by (induct al rule: clearjunk.induct)
+ (simp_all add: expand_fun_eq)
+
+lemma clearjunk_keys_set:
+ "set (map fst (clearjunk al)) = set (map fst al)"
+ by (induct al rule: clearjunk.induct)
+ (simp_all add: delete_keys)
+
+lemma dom_clearjunk:
+ "fst ` set (clearjunk al) = fst ` set al"
+ using clearjunk_keys_set by simp
+
+lemma distinct_clearjunk [simp]:
+ "distinct (map fst (clearjunk al))"
+ by (induct al rule: clearjunk.induct)
+ (simp_all del: set_map add: clearjunk_keys_set delete_keys)
+
+lemma ran_clearjunk:
+ "ran (map_of (clearjunk al)) = ran (map_of al)"
+ by (simp add: map_of_clearjunk)
+
+lemma ran_map_of:
+ "ran (map_of al) = snd ` set (clearjunk al)"
+proof -
+ have "ran (map_of al) = ran (map_of (clearjunk al))"
+ by (simp add: ran_clearjunk)
+ also have "\<dots> = snd ` set (clearjunk al)"
+ by (simp add: ran_distinct)
+ finally show ?thesis .
+qed
+
+lemma clearjunk_update:
+ "clearjunk (update k v al) = update k v (clearjunk al)"
+ by (induct al rule: clearjunk.induct)
+ (simp_all add: delete_update)
-lemma dom_map_ran: "fst ` set (map_ran f al) = fst ` set al"
+lemma clearjunk_updates:
+ "clearjunk (updates ks vs al) = updates ks vs (clearjunk al)"
+proof -
+ have "foldl (\<lambda>al (k, v). update k v al) (clearjunk al) (zip ks vs) =
+ clearjunk (foldl (\<lambda>al (k, v). update k v al) al (zip ks vs))"
+ by (rule foldl_apply) (simp add: clearjunk_update expand_fun_eq split_def)
+ then show ?thesis by (simp add: updates_def)
+qed
+
+lemma clearjunk_delete:
+ "clearjunk (delete x al) = delete x (clearjunk al)"
+ by (induct al rule: clearjunk.induct) (auto simp add: delete_idem delete_twist)
+
+lemma clearjunk_restrict:
+ "clearjunk (restrict A al) = restrict A (clearjunk al)"
+ by (induct al rule: clearjunk.induct) (auto simp add: restr_delete_twist)
+
+lemma distinct_clearjunk_id [simp]:
+ "distinct (map fst al) \<Longrightarrow> clearjunk al = al"
+ by (induct al rule: clearjunk.induct) auto
+
+lemma clearjunk_idem:
+ "clearjunk (clearjunk al) = clearjunk al"
+ by simp
+
+lemma length_clearjunk:
+ "length (clearjunk al) \<le> length al"
+proof (induct al rule: clearjunk.induct [case_names Nil Cons])
+ case Nil then show ?case by simp
+next
+ case (Cons kv al)
+ moreover have "length (delete (fst kv) al) \<le> length al" by (fact length_delete_le)
+ ultimately have "length (clearjunk (delete (fst kv) al)) \<le> length al" by (rule order_trans)
+ then show ?case by simp
+qed
+
+lemma delete_map:
+ assumes "\<And>kv. fst (f kv) = fst kv"
+ shows "delete k (map f ps) = map f (delete k ps)"
+ by (simp add: delete_eq filter_map comp_def split_def assms)
+
+lemma clearjunk_map:
+ assumes "\<And>kv. fst (f kv) = fst kv"
+ shows "clearjunk (map f ps) = map f (clearjunk ps)"
+ by (induct ps rule: clearjunk.induct [case_names Nil Cons])
+ (simp_all add: clearjunk_delete delete_map assms)
+
+
+subsection {* @{text map_ran} *}
+
+definition map_ran :: "('key \<Rightarrow> 'val \<Rightarrow> 'val) \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list" where
+ "map_ran f = map (\<lambda>(k, v). (k, f k v))"
+
+lemma map_ran_simps [simp]:
+ "map_ran f [] = []"
+ "map_ran f ((k, v) # ps) = (k, f k v) # map_ran f ps"
+ by (simp_all add: map_ran_def)
+
+lemma dom_map_ran:
+ "fst ` set (map_ran f al) = fst ` set al"
+ by (simp add: map_ran_def image_image split_def)
+
+lemma map_ran_conv:
+ "map_of (map_ran f al) k = Option.map (f k) (map_of al k)"
by (induct al) auto
-lemma distinct_map_ran: "distinct (map fst al) \<Longrightarrow> distinct (map fst (map_ran f al))"
- by (induct al) (auto simp add: dom_map_ran)
+lemma distinct_map_ran:
+ "distinct (map fst al) \<Longrightarrow> distinct (map fst (map_ran f al))"
+ by (simp add: map_ran_def split_def comp_def)
-lemma map_ran_filter: "map_ran f [p\<leftarrow>ps. fst p \<noteq> a] = [p\<leftarrow>map_ran f ps. fst p \<noteq> a]"
- by (induct ps) auto
+lemma map_ran_filter:
+ "map_ran f [p\<leftarrow>ps. fst p \<noteq> a] = [p\<leftarrow>map_ran f ps. fst p \<noteq> a]"
+ by (simp add: map_ran_def filter_map split_def comp_def)
-lemma clearjunk_map_ran: "clearjunk (map_ran f al) = map_ran f (clearjunk al)"
- by (induct al rule: clearjunk.induct) (auto simp add: delete_eq map_ran_filter)
+lemma clearjunk_map_ran:
+ "clearjunk (map_ran f al) = map_ran f (clearjunk al)"
+ by (simp add: map_ran_def split_def clearjunk_map)
-subsection {* @{const merge} *}
+subsection {* @{text merge} *}
+
+definition merge :: "('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list \<Rightarrow> ('key \<times> 'val) list" where
+ "merge qs ps = foldr (\<lambda>(k, v). update k v) ps qs"
+
+lemma merge_simps [simp]:
+ "merge qs [] = qs"
+ "merge qs (p#ps) = update (fst p) (snd p) (merge qs ps)"
+ by (simp_all add: merge_def split_def)
+
+lemma merge_updates:
+ "merge qs ps = updates (rev (map fst ps)) (rev (map snd ps)) qs"
+ by (simp add: merge_def updates_def split_def
+ foldr_foldl zip_rev zip_map_fst_snd)
lemma dom_merge: "fst ` set (merge xs ys) = fst ` set xs \<union> fst ` set ys"
by (induct ys arbitrary: xs) (auto simp add: dom_update)
@@ -452,34 +438,27 @@
lemma distinct_merge:
assumes "distinct (map fst xs)"
shows "distinct (map fst (merge xs ys))"
- using assms
-by (induct ys arbitrary: xs) (auto simp add: dom_merge distinct_update)
+using assms by (simp add: merge_updates distinct_updates)
lemma clearjunk_merge:
- "clearjunk (merge xs ys) = merge (clearjunk xs) ys"
- by (induct ys) (auto simp add: clearjunk_update)
+ "clearjunk (merge xs ys) = merge (clearjunk xs) ys"
+ by (simp add: merge_updates clearjunk_updates)
-lemma merge_conv: "map_of (merge xs ys) k = (map_of xs ++ map_of ys) k"
-proof (induct ys)
- case Nil thus ?case by simp
-next
- case (Cons y ys)
- show ?case
- proof (cases "k = fst y")
- case True
- from True show ?thesis
- by (simp add: update_conv)
- next
- case False
- from False show ?thesis
- by (auto simp add: update_conv Cons.hyps map_add_def)
- qed
+lemma merge_conv':
+ "map_of (merge xs ys) = map_of xs ++ map_of ys"
+proof -
+ have "foldl (\<lambda>m (k, v). m(k \<mapsto> v)) (map_of xs) (rev ys) =
+ map_of (foldl (\<lambda>xs (k, v). update k v xs) xs (rev ys)) "
+ by (rule foldl_apply) (simp add: expand_fun_eq split_def update_conv')
+ then show ?thesis
+ by (simp add: merge_def map_add_map_of_foldr foldr_foldl split_def)
qed
-lemma merge_conv': "map_of (merge xs ys) = (map_of xs ++ map_of ys)"
- by (rule ext) (rule merge_conv)
+corollary merge_conv:
+ "map_of (merge xs ys) k = (map_of xs ++ map_of ys) k"
+ by (simp add: merge_conv')
-lemma merge_emty: "map_of (merge [] ys) = map_of ys"
+lemma merge_empty: "map_of (merge [] ys) = map_of ys"
by (simp add: merge_conv')
lemma merge_assoc[simp]: "map_of (merge m1 (merge m2 m3)) =
@@ -491,8 +470,7 @@
(map_of n k = Some x \<or> map_of n k = None \<and> map_of m k = Some x)"
by (simp add: merge_conv' map_add_Some_iff)
-lemmas merge_SomeD = merge_Some_iff [THEN iffD1, standard]
-declare merge_SomeD [dest!]
+lemmas merge_SomeD [dest!] = merge_Some_iff [THEN iffD1, standard]
lemma merge_find_right[simp]: "map_of n k = Some v \<Longrightarrow> map_of (merge m n) k = Some v"
by (simp add: merge_conv')
@@ -513,7 +491,16 @@
by (simp add: merge_conv')
-subsection {* @{const compose} *}
+subsection {* @{text compose} *}
+
+function compose :: "('key \<times> 'a) list \<Rightarrow> ('a \<times> 'b) list \<Rightarrow> ('key \<times> 'b) list" where
+ "compose [] ys = []"
+ | "compose (x#xs) ys = (case map_of ys (snd x)
+ of None \<Rightarrow> compose (delete (fst x) xs) ys
+ | Some v \<Rightarrow> (fst x, v) # compose xs ys)"
+ by pat_completeness auto
+termination by (relation "measure (length \<circ> fst)")
+ (simp_all add: less_Suc_eq_le length_delete_le)
lemma compose_first_None [simp]:
assumes "map_of xs k = None"
@@ -669,126 +656,4 @@
(map_of xs k = None \<or> (\<exists>k'. map_of xs k = Some k' \<and> map_of ys k' = None)) "
by (simp add: compose_conv map_comp_None_iff)
-
-subsection {* @{const restrict} *}
-
-lemma restrict_eq:
- "restrict A = filter (\<lambda>p. fst p \<in> A)"
-proof
- fix xs
- show "restrict A xs = filter (\<lambda>p. fst p \<in> A) xs"
- by (induct xs) auto
-qed
-
-lemma distinct_restr: "distinct (map fst al) \<Longrightarrow> distinct (map fst (restrict A al))"
- by (induct al) (auto simp add: restrict_eq)
-
-lemma restr_conv: "map_of (restrict A al) k = ((map_of al)|` A) k"
- apply (induct al)
- apply (simp add: restrict_eq)
- apply (cases "k\<in>A")
- apply (auto simp add: restrict_eq)
- done
-
-lemma restr_conv': "map_of (restrict A al) = ((map_of al)|` A)"
- by (rule ext) (rule restr_conv)
-
-lemma restr_empty [simp]:
- "restrict {} al = []"
- "restrict A [] = []"
- by (induct al) (auto simp add: restrict_eq)
-
-lemma restr_in [simp]: "x \<in> A \<Longrightarrow> map_of (restrict A al) x = map_of al x"
- by (simp add: restr_conv')
-
-lemma restr_out [simp]: "x \<notin> A \<Longrightarrow> map_of (restrict A al) x = None"
- by (simp add: restr_conv')
-
-lemma dom_restr [simp]: "fst ` set (restrict A al) = fst ` set al \<inter> A"
- by (induct al) (auto simp add: restrict_eq)
-
-lemma restr_upd_same [simp]: "restrict (-{x}) (update x y al) = restrict (-{x}) al"
- by (induct al) (auto simp add: restrict_eq)
-
-lemma restr_restr [simp]: "restrict A (restrict B al) = restrict (A\<inter>B) al"
- by (induct al) (auto simp add: restrict_eq)
-
-lemma restr_update[simp]:
- "map_of (restrict D (update x y al)) =
- map_of ((if x \<in> D then (update x y (restrict (D-{x}) al)) else restrict D al))"
- by (simp add: restr_conv' update_conv')
-
-lemma restr_delete [simp]:
- "(delete x (restrict D al)) =
- (if x\<in> D then restrict (D - {x}) al else restrict D al)"
-proof (induct al)
- case Nil thus ?case by simp
-next
- case (Cons a al)
- show ?case
- proof (cases "x \<in> D")
- case True
- note x_D = this
- with Cons have hyp: "delete x (restrict D al) = restrict (D - {x}) al"
- by simp
- show ?thesis
- proof (cases "fst a = x")
- case True
- from Cons.hyps
- show ?thesis
- using x_D True
- by simp
- next
- case False
- note not_fst_a_x = this
- show ?thesis
- proof (cases "fst a \<in> D")
- case True
- with not_fst_a_x
- have "delete x (restrict D (a#al)) = a#(delete x (restrict D al))"
- by (cases a) (simp add: restrict_eq)
- also from not_fst_a_x True hyp have "\<dots> = restrict (D - {x}) (a # al)"
- by (cases a) (simp add: restrict_eq)
- finally show ?thesis
- using x_D by simp
- next
- case False
- hence "delete x (restrict D (a#al)) = delete x (restrict D al)"
- by (cases a) (simp add: restrict_eq)
- moreover from False not_fst_a_x
- have "restrict (D - {x}) (a # al) = restrict (D - {x}) al"
- by (cases a) (simp add: restrict_eq)
- ultimately
- show ?thesis using x_D hyp by simp
- qed
- qed
- next
- case False
- from False Cons show ?thesis
- by simp
- qed
-qed
-
-lemma update_restr:
- "map_of (update x y (restrict D al)) = map_of (update x y (restrict (D-{x}) al))"
- by (simp add: update_conv' restr_conv') (rule fun_upd_restrict)
-
-lemma upate_restr_conv [simp]:
- "x \<in> D \<Longrightarrow>
- map_of (update x y (restrict D al)) = map_of (update x y (restrict (D-{x}) al))"
- by (simp add: update_conv' restr_conv')
-
-lemma restr_updates [simp]: "
- \<lbrakk> length xs = length ys; set xs \<subseteq> D \<rbrakk>
- \<Longrightarrow> map_of (restrict D (updates xs ys al)) =
- map_of (updates xs ys (restrict (D - set xs) al))"
- by (simp add: updates_conv' restr_conv')
-
-lemma restr_delete_twist: "(restrict A (delete a ps)) = delete a (restrict A ps)"
- by (induct ps) auto
-
-lemma clearjunk_restrict:
- "clearjunk (restrict A al) = restrict A (clearjunk al)"
- by (induct al rule: clearjunk.induct) (auto simp add: restr_delete_twist)
-
end
--- a/src/HOL/Library/Executable_Set.thy Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/Library/Executable_Set.thy Tue Feb 02 18:16:48 2010 +0100
@@ -93,14 +93,14 @@
by simp
lemma insert_Set [code]:
- "insert x (Set xs) = Set (List_Set.insert x xs)"
- "insert x (Coset xs) = Coset (remove_all x xs)"
- by (simp_all add: insert_set insert_set_compl)
+ "insert x (Set xs) = Set (List.insert x xs)"
+ "insert x (Coset xs) = Coset (removeAll x xs)"
+ by (simp_all add: set_insert)
lemma remove_Set [code]:
- "remove x (Set xs) = Set (remove_all x xs)"
- "remove x (Coset xs) = Coset (List_Set.insert x xs)"
- by (simp_all add:remove_set remove_set_compl)
+ "remove x (Set xs) = Set (removeAll x xs)"
+ "remove x (Coset xs) = Coset (List.insert x xs)"
+ by (auto simp add: set_insert remove_def)
lemma image_Set [code]:
"image f (Set xs) = Set (remdups (map f xs))"
--- a/src/HOL/Library/Fset.thy Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/Library/Fset.thy Tue Feb 02 18:16:48 2010 +0100
@@ -126,17 +126,18 @@
[simp]: "insert x A = Fset (Set.insert x (member A))"
lemma insert_Set [code]:
- "insert x (Set xs) = Set (List_Set.insert x xs)"
- "insert x (Coset xs) = Coset (remove_all x xs)"
- by (simp_all add: Set_def Coset_def insert_set insert_set_compl)
+ "insert x (Set xs) = Set (List.insert x xs)"
+ "insert x (Coset xs) = Coset (removeAll x xs)"
+ by (simp_all add: Set_def Coset_def set_insert)
definition remove :: "'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset" where
[simp]: "remove x A = Fset (List_Set.remove x (member A))"
lemma remove_Set [code]:
- "remove x (Set xs) = Set (remove_all x xs)"
- "remove x (Coset xs) = Coset (List_Set.insert x xs)"
- by (simp_all add: Set_def Coset_def remove_set remove_set_compl)
+ "remove x (Set xs) = Set (removeAll x xs)"
+ "remove x (Coset xs) = Coset (List.insert x xs)"
+ by (simp_all add: Set_def Coset_def remove_set_compl)
+ (simp add: List_Set.remove_def)
definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset" where
[simp]: "map f A = Fset (image f (member A))"
@@ -203,7 +204,7 @@
by (simp add: inter project_def Set_def)
have "foldl (\<lambda>A x. List_Set.remove x A) (member A) xs =
member (foldl (\<lambda>A x. Fset (List_Set.remove x (member A))) A xs)"
- by (rule foldl_apply_inv) simp
+ by (rule foldl_apply) (simp add: expand_fun_eq)
then show "inf A (Coset xs) = foldl (\<lambda>A x. remove x A) A xs"
by (simp add: Diff_eq [symmetric] minus_set)
qed
@@ -214,7 +215,7 @@
proof -
have "foldl (\<lambda>A x. List_Set.remove x A) (member A) xs =
member (foldl (\<lambda>A x. Fset (List_Set.remove x (member A))) A xs)"
- by (rule foldl_apply_inv) simp
+ by (rule foldl_apply) (simp add: expand_fun_eq)
then show "A - Set xs = foldl (\<lambda>A x. remove x A) A xs"
by (simp add: minus_set)
show "A - Coset xs = Set (List.filter (member A) xs)"
@@ -227,7 +228,7 @@
proof -
have "foldl (\<lambda>A x. Set.insert x A) (member A) xs =
member (foldl (\<lambda>A x. Fset (Set.insert x (member A))) A xs)"
- by (rule foldl_apply_inv) simp
+ by (rule foldl_apply) (simp add: expand_fun_eq)
then show "sup (Set xs) A = foldl (\<lambda>A x. insert x A) A xs"
by (simp add: union_set)
show "sup (Coset xs) A = Coset (List.filter (Not \<circ> member A) xs)"
--- a/src/HOL/Library/List_Set.thy Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/Library/List_Set.thy Tue Feb 02 18:16:48 2010 +0100
@@ -7,15 +7,6 @@
imports Main
begin
-subsection {* Various additional list functions *}
-
-definition insert :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
- "insert x xs = (if x \<in> set xs then xs else x # xs)"
-
-definition remove_all :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
- "remove_all x xs = filter (Not o op = x) xs"
-
-
subsection {* Various additional set functions *}
definition is_empty :: "'a set \<Rightarrow> bool" where
@@ -61,21 +52,13 @@
"{} = set []"
by simp
-lemma insert_set:
- "Set.insert x (set xs) = set (insert x xs)"
- by (auto simp add: insert_def)
-
lemma insert_set_compl:
- "Set.insert x (- set xs) = - set (List_Set.remove_all x xs)"
- by (auto simp del: mem_def simp add: remove_all_def)
-
-lemma remove_set:
- "remove x (set xs) = set (remove_all x xs)"
- by (auto simp add: remove_def remove_all_def)
+ "insert x (- set xs) = - set (removeAll x xs)"
+ by auto
lemma remove_set_compl:
- "List_Set.remove x (- set xs) = - set (List_Set.insert x xs)"
- by (auto simp del: mem_def simp add: remove_def List_Set.insert_def)
+ "remove x (- set xs) = - set (List.insert x xs)"
+ by (auto simp del: mem_def simp add: remove_def List.insert_def)
lemma image_set:
"image f (set xs) = set (map f xs)"
@@ -128,7 +111,4 @@
"A \<inter> B = project (\<lambda>x. x \<in> A) B"
by (auto simp add: project_def)
-
-hide (open) const insert
-
end
\ No newline at end of file
--- a/src/HOL/List.thy Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/List.thy Tue Feb 02 18:16:48 2010 +0100
@@ -167,6 +167,12 @@
"remdups [] = []"
| "remdups (x # xs) = (if x \<in> set xs then remdups xs else x # remdups xs)"
+definition
+ insert :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
+ "insert x xs = (if x \<in> set xs then xs else x # xs)"
+
+hide (open) const insert hide (open) fact insert_def
+
primrec
remove1 :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
"remove1 x [] = []"
@@ -242,6 +248,8 @@
@{lemma "dropWhile (%n::nat. n<3) [1,2,3,0] = [3,0]" by simp}\\
@{lemma "distinct [2,0,1::nat]" by simp}\\
@{lemma "remdups [2,0,2,1::nat,2] = [0,1,2]" by simp}\\
+@{lemma "List.insert 2 [0::nat,1,2] = [0,1,2]" by (simp add: List.insert_def)}\\
+@{lemma "List.insert 3 [0::nat,1,2] = [3, 0,1,2]" by (simp add: List.insert_def)}\\
@{lemma "remove1 2 [2,0,2,1::nat,2] = [0,2,1,2]" by simp}\\
@{lemma "removeAll 2 [2,0,2,1::nat,2] = [0,1]" by simp}\\
@{lemma "nth [a,b,c,d] 2 = c" by simp}\\
@@ -1899,6 +1907,23 @@
"length (zip xs ys) = min (length xs) (length ys)"
by (induct xs ys rule:list_induct2') auto
+lemma zip_obtain_same_length:
+ assumes "\<And>zs ws n. length zs = length ws \<Longrightarrow> n = min (length xs) (length ys)
+ \<Longrightarrow> zs = take n xs \<Longrightarrow> ws = take n ys \<Longrightarrow> P (zip zs ws)"
+ shows "P (zip xs ys)"
+proof -
+ let ?n = "min (length xs) (length ys)"
+ have "P (zip (take ?n xs) (take ?n ys))"
+ by (rule assms) simp_all
+ moreover have "zip xs ys = zip (take ?n xs) (take ?n ys)"
+ proof (induct xs arbitrary: ys)
+ case Nil then show ?case by simp
+ next
+ case (Cons x xs) then show ?case by (cases ys) simp_all
+ qed
+ ultimately show ?thesis by simp
+qed
+
lemma zip_append1:
"zip (xs @ ys) zs =
zip xs (take (length xs) zs) @ zip ys (drop (length xs) zs)"
@@ -2213,10 +2238,10 @@
"foldl g a (map f xs) = foldl (%a x. g a (f x)) a xs"
by(induct xs arbitrary:a) simp_all
-lemma foldl_apply_inv:
- assumes "\<And>x. g (h x) = x"
- shows "foldl f (g s) xs = g (foldl (\<lambda>s x. h (f (g s) x)) s xs)"
- by (rule sym, induct xs arbitrary: s) (simp_all add: assms)
+lemma foldl_apply:
+ assumes "\<And>x. x \<in> set xs \<Longrightarrow> f x \<circ> h = h \<circ> g x"
+ shows "foldl (\<lambda>s x. f x s) (h s) xs = h (foldl (\<lambda>s x. g x s) s xs)"
+ by (rule sym, insert assms, induct xs arbitrary: s) (simp_all add: expand_fun_eq)
lemma foldl_cong [fundef_cong, recdef_cong]:
"[| a = b; l = k; !!a x. x : set l ==> f a x = g a x |]
@@ -2282,6 +2307,12 @@
"\<lbrakk>Q x ; \<forall> x\<in> set xs. P x; \<forall> x y. P x \<and> Q y \<longrightarrow> Q (f y x) \<rbrakk> \<Longrightarrow> Q (foldl f x xs)"
by (induct xs arbitrary: x, simp_all)
+lemma foldl_weak_invariant:
+ assumes "P s"
+ and "\<And>s x. x \<in> set xs \<Longrightarrow> P s \<Longrightarrow> P (f s x)"
+ shows "P (foldl f s xs)"
+ using assms by (induct xs arbitrary: s) simp_all
+
text {* @{const foldl} and @{const concat} *}
lemma foldl_conv_concat:
@@ -2804,6 +2835,25 @@
from length_remdups_concat[of "[xs]"] show ?thesis unfolding xs by simp
qed
+subsubsection {* @{const insert} *}
+
+lemma in_set_insert [simp]:
+ "x \<in> set xs \<Longrightarrow> List.insert x xs = xs"
+ by (simp add: List.insert_def)
+
+lemma not_in_set_insert [simp]:
+ "x \<notin> set xs \<Longrightarrow> List.insert x xs = x # xs"
+ by (simp add: List.insert_def)
+
+lemma insert_Nil [simp]:
+ "List.insert x [] = [x]"
+ by simp
+
+lemma set_insert:
+ "set (List.insert x xs) = insert x (set xs)"
+ by (auto simp add: List.insert_def)
+
+
subsubsection {* @{text remove1} *}
lemma remove1_append:
@@ -2852,6 +2902,14 @@
subsubsection {* @{text removeAll} *}
+lemma removeAll_filter_not_eq:
+ "removeAll x = filter (\<lambda>y. x \<noteq> y)"
+proof
+ fix xs
+ show "removeAll x xs = filter (\<lambda>y. x \<noteq> y) xs"
+ by (induct xs) auto
+qed
+
lemma removeAll_append[simp]:
"removeAll x (xs @ ys) = removeAll x xs @ removeAll x ys"
by (induct xs) auto
@@ -2871,6 +2929,9 @@
"\<not> P x \<Longrightarrow> removeAll x (filter P xs) = filter P xs"
by(induct xs) auto
+lemma distinct_removeAll:
+ "distinct xs \<Longrightarrow> distinct (removeAll x xs)"
+ by (simp add: removeAll_filter_not_eq)
lemma distinct_remove1_removeAll:
"distinct xs ==> remove1 x xs = removeAll x xs"
--- a/src/HOL/Map.thy Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/Map.thy Tue Feb 02 18:16:48 2010 +0100
@@ -331,6 +331,19 @@
"inj_on (m ++ m') (dom m') = inj_on m' (dom m')"
by (fastsimp simp: map_add_def dom_def inj_on_def split: option.splits)
+lemma map_upds_fold_map_upd:
+ "f(ks[\<mapsto>]vs) = foldl (\<lambda>f (k, v). f(k\<mapsto>v)) f (zip ks vs)"
+unfolding map_upds_def proof (rule sym, rule zip_obtain_same_length)
+ fix ks :: "'a list" and vs :: "'b list"
+ assume "length ks = length vs"
+ then show "foldl (\<lambda>f (k, v). f(k\<mapsto>v)) f (zip ks vs) = f ++ map_of (rev (zip ks vs))"
+ by (induct arbitrary: f rule: list_induct2) simp_all
+qed
+
+lemma map_add_map_of_foldr:
+ "m ++ map_of ps = foldr (\<lambda>(k, v) m. m(k \<mapsto> v)) ps m"
+ by (induct ps) (auto simp add: expand_fun_eq map_add_def)
+
subsection {* @{term [source] restrict_map} *}
@@ -480,12 +493,13 @@
"dom(f(x := y)) = (if y=None then dom f - {x} else insert x (dom f))"
by(auto simp add:dom_def)
-lemma dom_map_of: "dom(map_of xys) = {x. \<exists>y. (x,y) : set xys}"
-by (induct xys) (auto simp del: fun_upd_apply)
+lemma dom_if:
+ "dom (\<lambda>x. if P x then f x else g x) = dom f \<inter> {x. P x} \<union> dom g \<inter> {x. \<not> P x}"
+ by (auto split: if_splits)
lemma dom_map_of_conv_image_fst:
- "dom(map_of xys) = fst ` (set xys)"
-by(force simp: dom_map_of)
+ "dom (map_of xys) = fst ` set xys"
+ by (induct xys) (auto simp add: dom_if)
lemma dom_map_of_zip [simp]: "[| length xs = length ys; distinct xs |] ==>
dom(map_of(zip xs ys)) = set xs"
@@ -523,11 +537,6 @@
"dom (\<lambda>x. Some y) = UNIV"
by auto
-lemma dom_if:
- "dom (\<lambda>x. if P x then f x else g x) = dom f \<inter> {x. P x} \<union> dom g \<inter> {x. \<not> P x}"
- by (auto split: if_splits)
-
-
(* Due to John Matthews - could be rephrased with dom *)
lemma finite_map_freshness:
"finite (dom (f :: 'a \<rightharpoonup> 'b)) \<Longrightarrow> \<not> finite (UNIV :: 'a set) \<Longrightarrow>
@@ -559,6 +568,19 @@
apply auto
done
+lemma ran_distinct:
+ assumes dist: "distinct (map fst al)"
+ shows "ran (map_of al) = snd ` set al"
+using assms proof (induct al)
+ case Nil then show ?case by simp
+next
+ case (Cons kv al)
+ then have "ran (map_of al) = snd ` set al" by simp
+ moreover from Cons.prems have "map_of al (fst kv) = None"
+ by (simp add: map_of_eq_None_iff)
+ ultimately show ?case by (simp only: map_of.simps ran_map_upd) simp
+qed
+
subsection {* @{text "map_le"} *}
@@ -610,7 +632,6 @@
lemma map_add_le_mapI: "\<lbrakk> f \<subseteq>\<^sub>m h; g \<subseteq>\<^sub>m h; f \<subseteq>\<^sub>m f++g \<rbrakk> \<Longrightarrow> f++g \<subseteq>\<^sub>m h"
by (clarsimp simp add: map_le_def map_add_def dom_def split: option.splits)
-
lemma dom_eq_singleton_conv: "dom f = {x} \<longleftrightarrow> (\<exists>v. f = [x \<mapsto> v])"
proof(rule iffI)
assume "\<exists>v. f = [x \<mapsto> v]"
@@ -626,3 +647,4 @@
qed
end
+
--- a/src/HOL/Multivariate_Analysis/Derivative.thy Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/Multivariate_Analysis/Derivative.thy Tue Feb 02 18:16:48 2010 +0100
@@ -1252,22 +1252,21 @@
using vector_derivative_works[unfolded differentiable_def]
using assms by(auto simp add:has_vector_derivative_def)
-lemma has_vector_derivative_within_subset: fixes f::"real \<Rightarrow> real^'a" shows
+lemma has_vector_derivative_within_subset:
"(f has_vector_derivative f') (at x within s) \<Longrightarrow> t \<subseteq> s \<Longrightarrow> (f has_vector_derivative f') (at x within t)"
unfolding has_vector_derivative_def apply(rule has_derivative_within_subset) by auto
-lemma has_vector_derivative_const: fixes c::"real^'n" shows
+lemma has_vector_derivative_const:
"((\<lambda>x. c) has_vector_derivative 0) net"
unfolding has_vector_derivative_def using has_derivative_const by auto
lemma has_vector_derivative_id: "((\<lambda>x::real. x) has_vector_derivative 1) net"
unfolding has_vector_derivative_def using has_derivative_id by auto
-lemma has_vector_derivative_cmul: fixes f::"real \<Rightarrow> real^'a"
- shows "(f has_vector_derivative f') net \<Longrightarrow> ((\<lambda>x. c *\<^sub>R f x) has_vector_derivative (c *\<^sub>R f')) net"
+lemma has_vector_derivative_cmul: "(f has_vector_derivative f') net \<Longrightarrow> ((\<lambda>x. c *\<^sub>R f x) has_vector_derivative (c *\<^sub>R f')) net"
unfolding has_vector_derivative_def apply(drule has_derivative_cmul) by(auto simp add:group_simps)
-lemma has_vector_derivative_cmul_eq: fixes f::"real \<Rightarrow> real^'a" assumes "c \<noteq> 0"
+lemma has_vector_derivative_cmul_eq: assumes "c \<noteq> 0"
shows "(((\<lambda>x. c *\<^sub>R f x) has_vector_derivative (c *\<^sub>R f')) net \<longleftrightarrow> (f has_vector_derivative f') net)"
apply rule apply(drule has_vector_derivative_cmul[where c="1/c"]) defer
apply(rule has_vector_derivative_cmul) using assms by auto
--- a/src/HOL/Nitpick.thy Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/Nitpick.thy Tue Feb 02 18:16:48 2010 +0100
@@ -36,6 +36,7 @@
and bisim_iterator_max :: bisim_iterator
and Quot :: "'a \<Rightarrow> 'b"
and quot_normal :: "'a \<Rightarrow> 'a"
+ and NonStd :: "'a \<Rightarrow> 'b"
and Tha :: "('a \<Rightarrow> bool) \<Rightarrow> 'a"
datatype ('a, 'b) pair_box = PairBox 'a 'b
@@ -43,6 +44,7 @@
typedecl unsigned_bit
typedecl signed_bit
+typedecl \<xi>
datatype 'a word = Word "('a set)"
@@ -250,12 +252,12 @@
setup {* Nitpick_Isar.setup *}
hide (open) const unknown is_unknown undefined_fast_The undefined_fast_Eps bisim
- bisim_iterator_max Quot quot_normal Tha PairBox FunBox Word refl' wf'
+ bisim_iterator_max Quot quot_normal NonStd Tha PairBox FunBox Word refl' wf'
wf_wfrec wf_wfrec' wfrec' card' setsum' fold_graph' nat_gcd nat_lcm int_gcd
int_lcm Frac Abs_Frac Rep_Frac zero_frac one_frac num denom norm_frac frac
plus_frac times_frac uminus_frac number_of_frac inverse_frac less_eq_frac
of_frac
-hide (open) type bisim_iterator pair_box fun_box unsigned_bit signed_bit word
+hide (open) type bisim_iterator pair_box fun_box unsigned_bit signed_bit \<xi> word
hide (open) fact If_def Ex1_def rtrancl_def rtranclp_def tranclp_def refl'_def
wf'_def wf_wfrec'_def wfrec'_def card'_def setsum'_def fold_graph'_def
The_psimp Eps_psimp unit_case_def nat_case_def list_size_simp nat_gcd_def
--- a/src/HOL/Nitpick_Examples/Manual_Nits.thy Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/Nitpick_Examples/Manual_Nits.thy Tue Feb 02 18:16:48 2010 +0100
@@ -220,6 +220,69 @@
nitpick
oops
+subsection {* 3.12. Inductive Properties *}
+
+inductive_set reach where
+"(4\<Colon>nat) \<in> reach" |
+"n \<in> reach \<Longrightarrow> n < 4 \<Longrightarrow> 3 * n + 1 \<in> reach" |
+"n \<in> reach \<Longrightarrow> n + 2 \<in> reach"
+
+lemma "n \<in> reach \<Longrightarrow> 2 dvd n"
+nitpick
+apply (induct set: reach)
+ apply auto
+ nitpick
+ apply (thin_tac "n \<in> reach")
+ nitpick
+oops
+
+lemma "n \<in> reach \<Longrightarrow> 2 dvd n \<and> n \<noteq> 0"
+nitpick
+apply (induct set: reach)
+ apply auto
+ nitpick
+ apply (thin_tac "n \<in> reach")
+ nitpick
+oops
+
+lemma "n \<in> reach \<Longrightarrow> 2 dvd n \<and> n \<ge> 4"
+by (induct set: reach) arith+
+
+datatype 'a bin_tree = Leaf 'a | Branch "'a bin_tree" "'a bin_tree"
+
+primrec labels where
+"labels (Leaf a) = {a}" |
+"labels (Branch t u) = labels t \<union> labels u"
+
+primrec swap where
+"swap (Leaf c) a b =
+ (if c = a then Leaf b else if c = b then Leaf a else Leaf c)" |
+"swap (Branch t u) a b = Branch (swap t a b) (swap u a b)"
+
+lemma "\<lbrakk>a \<in> labels t; b \<in> labels t; a \<noteq> b\<rbrakk> \<Longrightarrow> labels (swap t a b) = labels t"
+nitpick
+proof (induct t)
+ case Leaf thus ?case by simp
+next
+ case (Branch t u) thus ?case
+ nitpick
+ nitpick [non_std "'a bin_tree", show_consts]
+oops
+
+lemma "labels (swap t a b) =
+ (if a \<in> labels t then
+ if b \<in> labels t then labels t else (labels t - {a}) \<union> {b}
+ else
+ if b \<in> labels t then (labels t - {b}) \<union> {a} else labels t)"
+nitpick
+proof (induct t)
+ case Leaf thus ?case by simp
+next
+ case (Branch t u) thus ?case
+ nitpick [non_std "'a bin_tree", show_consts]
+ by auto
+qed
+
section {* 4. Case Studies *}
nitpick_params [max_potential = 0, max_threads = 2]
@@ -300,7 +363,7 @@
subsection {* 4.2. AA Trees *}
-datatype 'a tree = \<Lambda> | N "'a\<Colon>linorder" nat "'a tree" "'a tree"
+datatype 'a aa_tree = \<Lambda> | N "'a\<Colon>linorder" nat "'a aa_tree" "'a aa_tree"
primrec data where
"data \<Lambda> = undefined" |
--- a/src/HOL/Nitpick_Examples/Mono_Nits.thy Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/Nitpick_Examples/Mono_Nits.thy Tue Feb 02 18:16:48 2010 +0100
@@ -18,18 +18,19 @@
val def_table = Nitpick_HOL.const_def_table @{context} defs
val ext_ctxt : Nitpick_HOL.extended_context =
{thy = @{theory}, ctxt = @{context}, max_bisim_depth = ~1, boxes = [],
- user_axioms = NONE, debug = false, wfs = [], binary_ints = SOME false,
- destroy_constrs = false, specialize = false, skolemize = false,
- star_linear_preds = false, uncurry = false, fast_descrs = false,
- tac_timeout = NONE, evals = [], case_names = [], def_table = def_table,
- nondef_table = Symtab.empty, user_nondefs = [],
+ stds = [(NONE, true)], wfs = [], user_axioms = NONE, debug = false,
+ binary_ints = SOME false, destroy_constrs = false, specialize = false,
+ skolemize = false, star_linear_preds = false, uncurry = false,
+ fast_descrs = false, tac_timeout = NONE, evals = [], case_names = [],
+ def_table = def_table, nondef_table = Symtab.empty, user_nondefs = [],
simp_table = Unsynchronized.ref Symtab.empty, psimp_table = Symtab.empty,
intro_table = Symtab.empty, ground_thm_table = Inttab.empty,
ersatz_table = [], skolems = Unsynchronized.ref [],
special_funs = Unsynchronized.ref [], unrolled_preds = Unsynchronized.ref [],
wf_cache = Unsynchronized.ref [], constr_cache = Unsynchronized.ref []}
(* term -> bool *)
-val is_mono = Nitpick_Mono.formulas_monotonic ext_ctxt @{typ 'a} [] []
+val is_mono = Nitpick_Mono.formulas_monotonic ext_ctxt @{typ 'a}
+ Nitpick_Mono.Plus [] []
fun is_const t =
let val T = fastype_of t in
is_mono (Logic.mk_implies (Logic.mk_equals (Free ("dummyP", T), t),
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SMT/Examples/SMT_Examples.certs Tue Feb 02 18:16:48 2010 +0100
@@ -0,0 +1,13573 @@
+yknPpoG47N1CXnUaEL9RvQ 133
+#2 := false
+#1 := true
+#4 := (not true)
+#21 := (iff #4 false)
+#22 := [rewrite]: #21
+#20 := [asserted]: #4
+[mp #20 #22]: false
+unsat
+
+ymB2ZiCSl9aUjMXP3tIpZA 359
+#2 := false
+decl up_1 :: bool
+#4 := up_1
+#5 := (not up_1)
+#6 := (or up_1 #5)
+#7 := (not #6)
+#31 := (iff #7 false)
+#1 := true
+#26 := (not true)
+#29 := (iff #26 false)
+#30 := [rewrite]: #29
+#27 := (iff #7 #26)
+#24 := (iff #6 true)
+#25 := [rewrite]: #24
+#28 := [monotonicity #25]: #27
+#32 := [trans #28 #30]: #31
+#23 := [asserted]: #7
+[mp #23 #32]: false
+unsat
+
+XC3j0tGm4Y5klDm8sMkzVg 510
+#2 := false
+decl up_1 :: bool
+#4 := up_1
+#1 := true
+#5 := (and up_1 true)
+#6 := (iff #5 up_1)
+#7 := (not #6)
+#37 := (iff #7 false)
+#32 := (not true)
+#35 := (iff #32 false)
+#36 := [rewrite]: #35
+#33 := (iff #7 #32)
+#30 := (iff #6 true)
+#25 := (iff up_1 up_1)
+#28 := (iff #25 true)
+#29 := [rewrite]: #28
+#26 := (iff #6 #25)
+#24 := [rewrite]: #6
+#27 := [monotonicity #24]: #26
+#31 := [trans #27 #29]: #30
+#34 := [monotonicity #31]: #33
+#38 := [trans #34 #36]: #37
+#23 := [asserted]: #7
+[mp #23 #38]: false
+unsat
+
+y5d/Jtt47lXm/wUEvH5fHw 794
+#2 := false
+decl up_2 :: bool
+#5 := up_2
+decl up_1 :: bool
+#4 := up_1
+#6 := (or up_1 up_2)
+#51 := (iff #6 false)
+#46 := (or false false)
+#49 := (iff #46 false)
+#50 := [rewrite]: #49
+#47 := (iff #6 #46)
+#40 := (iff up_2 false)
+#9 := (not up_2)
+#43 := (iff #9 #40)
+#41 := (iff #40 #9)
+#42 := [rewrite]: #41
+#44 := [symm #42]: #43
+#32 := [asserted]: #9
+#45 := [mp #32 #44]: #40
+#35 := (iff up_1 false)
+#7 := (not up_1)
+#37 := (iff #7 #35)
+#33 := (iff #35 #7)
+#36 := [rewrite]: #33
+#38 := [symm #36]: #37
+#26 := (and #7 #6)
+#8 := (and #6 #7)
+#27 := (iff #8 #26)
+#28 := [rewrite]: #27
+#25 := [asserted]: #8
+#31 := [mp #25 #28]: #26
+#29 := [and-elim #31]: #7
+#39 := [mp #29 #38]: #35
+#48 := [monotonicity #39 #45]: #47
+#52 := [trans #48 #50]: #51
+#30 := [and-elim #31]: #6
+[mp #30 #52]: false
+unsat
+
+mDaukNkyA4glYbkfEOtcAw 7
+unsat
+
+TmB9YjKjdtDMIh0j9rMVPw 1530
+#2 := false
+decl up_1 :: bool
+#4 := up_1
+decl up_3 :: bool
+#7 := up_3
+#10 := (and up_1 up_3)
+decl up_2 :: bool
+#5 := up_2
+#9 := (and up_3 up_2)
+#11 := (or #9 #10)
+#12 := (implies up_1 #11)
+#13 := (or #12 up_1)
+#6 := (and up_1 up_2)
+#8 := (or #6 up_3)
+#14 := (implies #8 #13)
+#15 := (not #14)
+#81 := (iff #15 false)
+#32 := (and up_2 up_3)
+#38 := (or #10 #32)
+#46 := (not up_1)
+#47 := (or #46 #38)
+#55 := (or up_1 #47)
+#63 := (not #8)
+#64 := (or #63 #55)
+#69 := (not #64)
+#79 := (iff #69 false)
+#1 := true
+#74 := (not true)
+#77 := (iff #74 false)
+#78 := [rewrite]: #77
+#75 := (iff #69 #74)
+#72 := (iff #64 true)
+#73 := [rewrite]: #72
+#76 := [monotonicity #73]: #75
+#80 := [trans #76 #78]: #79
+#70 := (iff #15 #69)
+#67 := (iff #14 #64)
+#60 := (implies #8 #55)
+#65 := (iff #60 #64)
+#66 := [rewrite]: #65
+#61 := (iff #14 #60)
+#58 := (iff #13 #55)
+#52 := (or #47 up_1)
+#56 := (iff #52 #55)
+#57 := [rewrite]: #56
+#53 := (iff #13 #52)
+#50 := (iff #12 #47)
+#43 := (implies up_1 #38)
+#48 := (iff #43 #47)
+#49 := [rewrite]: #48
+#44 := (iff #12 #43)
+#41 := (iff #11 #38)
+#35 := (or #32 #10)
+#39 := (iff #35 #38)
+#40 := [rewrite]: #39
+#36 := (iff #11 #35)
+#33 := (iff #9 #32)
+#34 := [rewrite]: #33
+#37 := [monotonicity #34]: #36
+#42 := [trans #37 #40]: #41
+#45 := [monotonicity #42]: #44
+#51 := [trans #45 #49]: #50
+#54 := [monotonicity #51]: #53
+#59 := [trans #54 #57]: #58
+#62 := [monotonicity #59]: #61
+#68 := [trans #62 #66]: #67
+#71 := [monotonicity #68]: #70
+#82 := [trans #71 #80]: #81
+#31 := [asserted]: #15
+[mp #31 #82]: false
+unsat
+
+olufSxMlwMAAqyArYWPVOA 1300
+#2 := false
+decl up_1 :: bool
+#4 := up_1
+#5 := (iff up_1 up_1)
+#6 := (iff #5 up_1)
+#7 := (iff #6 up_1)
+#8 := (iff #7 up_1)
+#9 := (iff #8 up_1)
+#10 := (iff #9 up_1)
+#11 := (iff #10 up_1)
+#12 := (iff #11 up_1)
+#13 := (iff #12 up_1)
+#14 := (not #13)
+#69 := (iff #14 false)
+#1 := true
+#64 := (not true)
+#67 := (iff #64 false)
+#68 := [rewrite]: #67
+#65 := (iff #14 #64)
+#62 := (iff #13 true)
+#31 := (iff #5 true)
+#32 := [rewrite]: #31
+#60 := (iff #13 #5)
+#33 := (iff true up_1)
+#36 := (iff #33 up_1)
+#37 := [rewrite]: #36
+#57 := (iff #12 #33)
+#55 := (iff #11 true)
+#53 := (iff #11 #5)
+#50 := (iff #10 #33)
+#48 := (iff #9 true)
+#46 := (iff #9 #5)
+#43 := (iff #8 #33)
+#41 := (iff #7 true)
+#39 := (iff #7 #5)
+#34 := (iff #6 #33)
+#35 := [monotonicity #32]: #34
+#38 := [trans #35 #37]: #7
+#40 := [monotonicity #38]: #39
+#42 := [trans #40 #32]: #41
+#44 := [monotonicity #42]: #43
+#45 := [trans #44 #37]: #9
+#47 := [monotonicity #45]: #46
+#49 := [trans #47 #32]: #48
+#51 := [monotonicity #49]: #50
+#52 := [trans #51 #37]: #11
+#54 := [monotonicity #52]: #53
+#56 := [trans #54 #32]: #55
+#58 := [monotonicity #56]: #57
+#59 := [trans #58 #37]: #13
+#61 := [monotonicity #59]: #60
+#63 := [trans #61 #32]: #62
+#66 := [monotonicity #63]: #65
+#70 := [trans #66 #68]: #69
+#30 := [asserted]: #14
+[mp #30 #70]: false
+unsat
+
+agKJ550QwZ1mvlK8gw+tjQ 4627
+#2 := false
+decl up_1 :: bool
+#4 := up_1
+#75 := (not up_1)
+#246 := (iff #75 false)
+#1 := true
+#214 := (not true)
+#217 := (iff #214 false)
+#218 := [rewrite]: #217
+#244 := (iff #75 #214)
+#238 := (iff up_1 true)
+#241 := (iff up_1 #238)
+#239 := (iff #238 up_1)
+#240 := [rewrite]: #239
+#242 := [symm #240]: #241
+decl up_4 :: bool
+#7 := up_4
+decl up_2 :: bool
+#5 := up_2
+#161 := (or up_1 up_2 up_4)
+#200 := (iff #161 up_1)
+#195 := (or up_1 false false)
+#198 := (iff #195 up_1)
+#199 := [rewrite]: #198
+#196 := (iff #161 #195)
+#189 := (iff up_4 false)
+#102 := (not up_4)
+#192 := (iff #102 #189)
+#190 := (iff #189 #102)
+#191 := [rewrite]: #190
+#193 := [symm #191]: #192
+decl up_3 :: bool
+#6 := up_3
+#108 := (or up_3 #102)
+#180 := (iff #108 #102)
+#175 := (or false #102)
+#178 := (iff #175 #102)
+#179 := [rewrite]: #178
+#176 := (iff #108 #175)
+#152 := (iff up_3 false)
+#16 := (not up_3)
+#155 := (iff #16 #152)
+#153 := (iff #152 #16)
+#154 := [rewrite]: #153
+#156 := [symm #154]: #155
+decl up_9 :: bool
+#32 := up_9
+#33 := (not up_9)
+#34 := (and up_9 #33)
+decl up_8 :: bool
+#30 := up_8
+#35 := (or up_8 #34)
+#31 := (not up_8)
+#36 := (and #31 #35)
+#37 := (or up_3 #36)
+#38 := (not #37)
+#138 := (iff #38 #16)
+#136 := (iff #37 up_3)
+#131 := (or up_3 false)
+#134 := (iff #131 up_3)
+#135 := [rewrite]: #134
+#132 := (iff #37 #131)
+#129 := (iff #36 false)
+#124 := (and #31 up_8)
+#127 := (iff #124 false)
+#128 := [rewrite]: #127
+#125 := (iff #36 #124)
+#122 := (iff #35 up_8)
+#117 := (or up_8 false)
+#120 := (iff #117 up_8)
+#121 := [rewrite]: #120
+#118 := (iff #35 #117)
+#114 := (iff #34 false)
+#116 := [rewrite]: #114
+#119 := [monotonicity #116]: #118
+#123 := [trans #119 #121]: #122
+#126 := [monotonicity #123]: #125
+#130 := [trans #126 #128]: #129
+#133 := [monotonicity #130]: #132
+#137 := [trans #133 #135]: #136
+#139 := [monotonicity #137]: #138
+#113 := [asserted]: #38
+#142 := [mp #113 #139]: #16
+#157 := [mp #142 #156]: #152
+#177 := [monotonicity #157]: #176
+#181 := [trans #177 #179]: #180
+#27 := (or up_4 false)
+#28 := (not #27)
+#29 := (or #28 up_3)
+#111 := (iff #29 #108)
+#105 := (or #102 up_3)
+#109 := (iff #105 #108)
+#110 := [rewrite]: #109
+#106 := (iff #29 #105)
+#103 := (iff #28 #102)
+#99 := (iff #27 up_4)
+#101 := [rewrite]: #99
+#104 := [monotonicity #101]: #103
+#107 := [monotonicity #104]: #106
+#112 := [trans #107 #110]: #111
+#98 := [asserted]: #29
+#115 := [mp #98 #112]: #108
+#182 := [mp #115 #181]: #102
+#194 := [mp #182 #193]: #189
+#183 := (iff up_2 false)
+#92 := (not up_2)
+#186 := (iff #92 #183)
+#184 := (iff #183 #92)
+#185 := [rewrite]: #184
+#187 := [symm #185]: #186
+#95 := (or #92 up_3)
+#172 := (iff #95 #92)
+#167 := (or #92 false)
+#170 := (iff #167 #92)
+#171 := [rewrite]: #170
+#168 := (iff #95 #167)
+#169 := [monotonicity #157]: #168
+#173 := [trans #169 #171]: #172
+decl up_7 :: bool
+#21 := up_7
+#22 := (not up_7)
+#23 := (or up_7 #22)
+#24 := (and up_2 #23)
+#25 := (not #24)
+#26 := (or #25 up_3)
+#96 := (iff #26 #95)
+#93 := (iff #25 #92)
+#90 := (iff #24 up_2)
+#85 := (and up_2 true)
+#88 := (iff #85 up_2)
+#89 := [rewrite]: #88
+#86 := (iff #24 #85)
+#82 := (iff #23 true)
+#84 := [rewrite]: #82
+#87 := [monotonicity #84]: #86
+#91 := [trans #87 #89]: #90
+#94 := [monotonicity #91]: #93
+#97 := [monotonicity #94]: #96
+#81 := [asserted]: #26
+#100 := [mp #81 #97]: #95
+#174 := [mp #100 #173]: #92
+#188 := [mp #174 #187]: #183
+#197 := [monotonicity #188 #194]: #196
+#201 := [trans #197 #199]: #200
+#58 := (or up_1 up_2 up_3 up_4)
+#164 := (iff #58 #161)
+#158 := (or up_1 up_2 false up_4)
+#162 := (iff #158 #161)
+#163 := [rewrite]: #162
+#159 := (iff #58 #158)
+#160 := [monotonicity #157]: #159
+#165 := [trans #160 #163]: #164
+#8 := (or up_3 up_4)
+#9 := (or up_2 #8)
+#10 := (or up_1 #9)
+#59 := (iff #10 #58)
+#60 := [rewrite]: #59
+#55 := [asserted]: #10
+#61 := [mp #55 #60]: #58
+#166 := [mp #61 #165]: #161
+#202 := [mp #166 #201]: up_1
+#243 := [mp #202 #242]: #238
+#245 := [monotonicity #243]: #244
+#247 := [trans #245 #218]: #246
+#78 := (or #75 up_2)
+#235 := (iff #78 #75)
+#230 := (or #75 false)
+#233 := (iff #230 #75)
+#234 := [rewrite]: #233
+#231 := (iff #78 #230)
+#232 := [monotonicity #188]: #231
+#236 := [trans #232 #234]: #235
+#17 := (and up_3 #16)
+#18 := (or up_1 #17)
+#19 := (not #18)
+#20 := (or #19 up_2)
+#79 := (iff #20 #78)
+#76 := (iff #19 #75)
+#73 := (iff #18 up_1)
+#68 := (or up_1 false)
+#71 := (iff #68 up_1)
+#72 := [rewrite]: #71
+#69 := (iff #18 #68)
+#62 := (iff #17 false)
+#67 := [rewrite]: #62
+#70 := [monotonicity #67]: #69
+#74 := [trans #70 #72]: #73
+#77 := [monotonicity #74]: #76
+#80 := [monotonicity #77]: #79
+#57 := [asserted]: #20
+#83 := [mp #57 #80]: #78
+#237 := [mp #83 #236]: #75
+[mp #237 #247]: false
+unsat
+
++R/oj2I+uFZAw/Eu+3ULdw 1160
+#2 := false
+decl uf_1 :: (-> T1 T1 T1)
+decl uf_2 :: T1
+#10 := uf_2
+decl uf_3 :: T1
+#12 := uf_3
+#14 := (uf_1 uf_3 uf_2)
+#13 := (uf_1 uf_2 uf_3)
+#15 := (= #13 #14)
+#44 := (not #15)
+#11 := (= uf_2 uf_2)
+#16 := (and #11 #15)
+#17 := (not #16)
+#45 := (iff #17 #44)
+#42 := (iff #16 #15)
+#1 := true
+#37 := (and true #15)
+#40 := (iff #37 #15)
+#41 := [rewrite]: #40
+#38 := (iff #16 #37)
+#35 := (iff #11 true)
+#36 := [rewrite]: #35
+#39 := [monotonicity #36]: #38
+#43 := [trans #39 #41]: #42
+#46 := [monotonicity #43]: #45
+#34 := [asserted]: #17
+#49 := [mp #34 #46]: #44
+#4 := (:var 1 T1)
+#5 := (:var 0 T1)
+#7 := (uf_1 #5 #4)
+#530 := (pattern #7)
+#6 := (uf_1 #4 #5)
+#529 := (pattern #6)
+#8 := (= #6 #7)
+#531 := (forall (vars (?x1 T1) (?x2 T1)) (:pat #529 #530) #8)
+#9 := (forall (vars (?x1 T1) (?x2 T1)) #8)
+#534 := (iff #9 #531)
+#532 := (iff #8 #8)
+#533 := [refl]: #532
+#535 := [quant-intro #533]: #534
+#55 := (~ #9 #9)
+#53 := (~ #8 #8)
+#54 := [refl]: #53
+#56 := [nnf-pos #54]: #55
+#33 := [asserted]: #9
+#57 := [mp~ #33 #56]: #9
+#536 := [mp #57 #535]: #531
+#112 := (not #531)
+#199 := (or #112 #15)
+#113 := [quant-inst]: #199
+[unit-resolution #113 #536 #49]: false
+unsat
+
+c67Ar5f1aFkzAZ2wYy62Wg 56943
+#2 := false
+decl up_54 :: bool
+#126 := up_54
+#317 := (not up_54)
+decl up_60 :: bool
+#145 := up_60
+decl up_56 :: bool
+#131 := up_56
+#325 := (not up_56)
+decl up_55 :: bool
+#130 := up_55
+decl up_46 :: bool
+#108 := up_46
+#291 := (not up_46)
+decl up_35 :: bool
+#81 := up_35
+decl up_29 :: bool
+#66 := up_29
+decl up_32 :: bool
+#72 := up_32
+#235 := (not up_32)
+decl up_34 :: bool
+#77 := up_34
+#243 := (not up_34)
+decl up_33 :: bool
+#76 := up_33
+#250 := (not up_35)
+#1611 := [hypothesis]: #250
+decl up_24 :: bool
+#54 := up_24
+#209 := (not up_24)
+decl up_13 :: bool
+#28 := up_13
+decl up_11 :: bool
+#24 := up_11
+#165 := (not up_11)
+decl up_12 :: bool
+#25 := up_12
+#2327 := (or up_12 up_35)
+#345 := (not up_60)
+decl up_59 :: bool
+#142 := up_59
+decl up_19 :: bool
+#40 := up_19
+decl up_8 :: bool
+#17 := up_8
+#156 := (not up_8)
+decl up_7 :: bool
+#16 := up_7
+#166 := (not up_12)
+#1457 := [hypothesis]: #166
+#2183 := (or up_7 up_12 up_35)
+#155 := (not up_7)
+#1612 := [hypothesis]: #155
+decl up_10 :: bool
+#21 := up_10
+#161 := (not up_10)
+decl up_20 :: bool
+#44 := up_20
+decl up_23 :: bool
+#50 := up_23
+#202 := (not up_23)
+#2170 := (or up_34 up_7 up_35 up_12)
+#1605 := [hypothesis]: #243
+#2164 := (or up_29 up_34 up_7 up_35 up_12)
+decl up_42 :: bool
+#98 := up_42
+#275 := (not up_42)
+#226 := (not up_29)
+#907 := [hypothesis]: #226
+#2136 := (or up_29 up_12 up_7 up_35 up_32)
+decl up_22 :: bool
+#49 := up_22
+#895 := [hypothesis]: #235
+#1624 := (or up_29 up_22 up_12 up_32 up_35 up_7)
+decl up_21 :: bool
+#45 := up_21
+decl up_31 :: bool
+#71 := up_31
+#234 := (not up_31)
+decl up_9 :: bool
+#20 := up_9
+#201 := (not up_22)
+#1456 := [hypothesis]: #201
+#847 := (or #161 up_32 up_29 up_22 up_12)
+#193 := (not up_20)
+#1400 := [hypothesis]: up_10
+#964 := (or #161 #193)
+#197 := (or #193 #161)
+#966 := (iff #197 #964)
+#967 := [rewrite]: #966
+#963 := [asserted]: #197
+#970 := [mp #963 #967]: #964
+#1399 := [unit-resolution #970 #1400]: #193
+#500 := (or up_12 up_20 up_22 up_23)
+#51 := (or up_20 up_12)
+#52 := (or up_23 #51)
+#53 := (or up_22 #52)
+#503 := (iff #53 #500)
+#491 := (or up_12 up_20)
+#494 := (or up_23 #491)
+#497 := (or up_22 #494)
+#501 := (iff #497 #500)
+#502 := [rewrite]: #501
+#498 := (iff #53 #497)
+#495 := (iff #52 #494)
+#492 := (iff #51 #491)
+#493 := [rewrite]: #492
+#496 := [monotonicity #493]: #495
+#499 := [monotonicity #496]: #498
+#504 := [trans #499 #502]: #503
+#490 := [asserted]: #53
+#505 := [mp #490 #504]: #500
+#900 := [unit-resolution #505 #1399 #1456 #1457]: up_23
+#194 := (not up_21)
+#974 := (or #161 #194)
+#199 := (or #194 #161)
+#976 := (iff #199 #974)
+#977 := [rewrite]: #976
+#973 := [asserted]: #199
+#980 := [mp #973 #977]: #974
+#902 := [unit-resolution #980 #1400]: #194
+#574 := (or up_21 up_29 up_31 up_32)
+#73 := (or up_29 up_21)
+#74 := (or up_32 #73)
+#75 := (or up_31 #74)
+#577 := (iff #75 #574)
+#565 := (or up_21 up_29)
+#568 := (or up_32 #565)
+#571 := (or up_31 #568)
+#575 := (iff #571 #574)
+#576 := [rewrite]: #575
+#572 := (iff #75 #571)
+#569 := (iff #74 #568)
+#566 := (iff #73 #565)
+#567 := [rewrite]: #566
+#570 := [monotonicity #567]: #569
+#573 := [monotonicity #570]: #572
+#578 := [trans #573 #576]: #577
+#564 := [asserted]: #75
+#579 := [mp #564 #578]: #574
+#851 := [unit-resolution #579 #902 #895 #907]: up_31
+#1135 := (or #202 #234)
+#249 := (or #234 #202)
+#1137 := (iff #249 #1135)
+#1138 := [rewrite]: #1137
+#1134 := [asserted]: #249
+#1141 := [mp #1134 #1138]: #1135
+#858 := [unit-resolution #1141 #851 #900]: false
+#853 := [lemma #858]: #847
+#1613 := [unit-resolution #853 #907 #1456 #895 #1457]: #161
+#405 := (or up_7 up_9 up_10)
+#22 := (or up_10 up_7)
+#23 := (or up_9 #22)
+#408 := (iff #23 #405)
+#399 := (or up_7 up_10)
+#402 := (or up_9 #399)
+#406 := (iff #402 #405)
+#407 := [rewrite]: #406
+#403 := (iff #23 #402)
+#400 := (iff #22 #399)
+#401 := [rewrite]: #400
+#404 := [monotonicity #401]: #403
+#409 := [trans #404 #407]: #408
+#398 := [asserted]: #23
+#410 := [mp #398 #409]: #405
+#1614 := [unit-resolution #410 #1613 #1612]: up_9
+#160 := (not up_9)
+#881 := (or #160 #165)
+#168 := (or #165 #160)
+#882 := (iff #168 #881)
+#883 := [rewrite]: #882
+#879 := [asserted]: #168
+#886 := [mp #879 #883]: #881
+#1615 := [unit-resolution #886 #1614]: #165
+#425 := (or up_11 up_13)
+#29 := (or up_13 up_11)
+#426 := (iff #29 #425)
+#427 := [rewrite]: #426
+#424 := [asserted]: #29
+#430 := [mp #424 #427]: #425
+#1616 := [unit-resolution #430 #1615]: up_13
+#170 := (not up_13)
+#1015 := (or #170 #209)
+#211 := (or #209 #170)
+#1017 := (iff #211 #1015)
+#1018 := [rewrite]: #1017
+#1014 := [asserted]: #211
+#1021 := [mp #1014 #1018]: #1015
+#1617 := [unit-resolution #1021 #1616]: #209
+#603 := (or up_24 up_33 up_35)
+#82 := (or up_33 up_24)
+#83 := (or up_35 #82)
+#606 := (iff #83 #603)
+#597 := (or up_24 up_33)
+#600 := (or up_35 #597)
+#604 := (iff #600 #603)
+#605 := [rewrite]: #604
+#601 := (iff #83 #600)
+#598 := (iff #82 #597)
+#599 := [rewrite]: #598
+#602 := [monotonicity #599]: #601
+#607 := [trans #602 #605]: #606
+#596 := [asserted]: #83
+#608 := [mp #596 #607]: #603
+#1618 := [unit-resolution #608 #1617 #1611]: up_33
+#242 := (not up_33)
+#1116 := (or #234 #242)
+#245 := (or #242 #234)
+#1117 := (iff #245 #1116)
+#1118 := [rewrite]: #1117
+#1114 := [asserted]: #245
+#1121 := [mp #1114 #1118]: #1116
+#1619 := [unit-resolution #1121 #1618]: #234
+#1620 := [unit-resolution #579 #1619 #895 #907]: up_21
+#1120 := (or #202 #242)
+#246 := (or #242 #202)
+#1122 := (iff #246 #1120)
+#1123 := [rewrite]: #1122
+#1119 := [asserted]: #246
+#1126 := [mp #1119 #1123]: #1120
+#1621 := [unit-resolution #1126 #1618]: #202
+#1622 := [unit-resolution #505 #1621 #1456 #1457]: up_20
+#195 := (or #193 #194)
+#957 := [asserted]: #195
+#1623 := [unit-resolution #957 #1622 #1620]: false
+#1625 := [lemma #1623]: #1624
+#2132 := [unit-resolution #1625 #907 #1611 #1457 #895 #1612]: up_22
+#1978 := (or up_32 up_35 up_29 up_21 up_12 up_7)
+#1972 := [unit-resolution #1625 #895 #907 #1457 #1611 #1612]: up_22
+#1010 := (or #201 #209)
+#210 := (or #209 #201)
+#1012 := (iff #210 #1010)
+#1013 := [rewrite]: #1012
+#1009 := [asserted]: #210
+#1016 := [mp #1009 #1013]: #1010
+#1973 := [unit-resolution #1016 #1972]: #209
+#1974 := [hypothesis]: #194
+#1975 := [unit-resolution #579 #895 #907 #1974]: up_31
+#1976 := [unit-resolution #1121 #1975]: #242
+#1977 := [unit-resolution #608 #1976 #1973 #1611]: false
+#1979 := [lemma #1977]: #1978
+#2133 := [unit-resolution #1979 #907 #1611 #1457 #895 #1612]: up_21
+#1682 := (or #194 up_7 up_12 up_23)
+#1673 := [hypothesis]: #202
+#1674 := [hypothesis]: up_21
+#1675 := [unit-resolution #957 #1674]: #193
+#1676 := [unit-resolution #505 #1675 #1457 #1673]: up_22
+#1020 := (or #170 #201)
+#212 := (or #201 #170)
+#1022 := (iff #212 #1020)
+#1023 := [rewrite]: #1022
+#1019 := [asserted]: #212
+#1026 := [mp #1019 #1023]: #1020
+#1677 := [unit-resolution #1026 #1676]: #170
+#1678 := [unit-resolution #980 #1674]: #161
+#1679 := [unit-resolution #410 #1678 #1612]: up_9
+#1680 := [unit-resolution #886 #1679]: #165
+#1681 := [unit-resolution #430 #1680 #1677]: false
+#1683 := [lemma #1681]: #1682
+#2134 := [unit-resolution #1683 #2133 #1457 #1612]: up_23
+#203 := (or #201 #202)
+#983 := [asserted]: #203
+#2135 := [unit-resolution #983 #2134 #2132]: false
+#2137 := [lemma #2135]: #2136
+#2156 := [unit-resolution #2137 #907 #1612 #1611 #1457]: up_32
+#1224 := (or #235 #275)
+#279 := (or #275 #235)
+#1226 := (iff #279 #1224)
+#1227 := [rewrite]: #1226
+#1223 := [asserted]: #279
+#1230 := [mp #1223 #1227]: #1224
+#2157 := [unit-resolution #1230 #2156]: #275
+#2158 := (or up_12 up_29 up_7 up_54)
+decl up_26 :: bool
+#58 := up_26
+#214 := (not up_26)
+decl up_15 :: bool
+#31 := up_15
+decl up_14 :: bool
+#30 := up_14
+#172 := (not up_14)
+decl up_6 :: bool
+#13 := up_6
+decl up_5 :: bool
+#12 := up_5
+#150 := (not up_5)
+decl up_25 :: bool
+#57 := up_25
+#2099 := [hypothesis]: up_5
+#859 := (or #150 #155)
+#158 := (or #155 #150)
+#860 := (iff #158 #859)
+#861 := [rewrite]: #860
+#857 := [asserted]: #158
+#864 := [mp #857 #861]: #859
+#2100 := [unit-resolution #864 #2099]: #155
+#863 := (or #150 #156)
+#159 := (or #156 #150)
+#865 := (iff #159 #863)
+#866 := [rewrite]: #865
+#862 := [asserted]: #159
+#869 := [mp #862 #866]: #863
+#2101 := [unit-resolution #869 #2099]: #156
+#2097 := (or up_12 up_7 up_8)
+#1626 := [hypothesis]: #156
+#2054 := (or up_54 up_7 up_8)
+decl up_16 :: bool
+#34 := up_16
+#1597 := [hypothesis]: #317
+#1888 := (or up_16 up_8 up_7 up_54)
+decl up_45 :: bool
+#104 := up_45
+#284 := (not up_45)
+decl up_52 :: bool
+#121 := up_52
+#309 := (not up_52)
+decl up_51 :: bool
+#120 := up_51
+#177 := (not up_16)
+#1627 := [hypothesis]: #177
+#1733 := (or up_51 up_7 up_54 up_8 up_16)
+decl up_53 :: bool
+#125 := up_53
+#308 := (not up_51)
+#1598 := [hypothesis]: #308
+decl up_43 :: bool
+#99 := up_43
+#276 := (not up_43)
+#1710 := (or up_32 up_16 up_8 up_7 up_51 up_54)
+#1671 := (or up_35 up_16 up_8 up_32 up_7)
+#1655 := (or #166 up_32 up_16 up_8 up_35 up_7)
+#1642 := [hypothesis]: up_12
+#885 := (or #160 #166)
+#169 := (or #166 #160)
+#887 := (iff #169 #885)
+#888 := [rewrite]: #887
+#884 := [asserted]: #169
+#891 := [mp #884 #888]: #885
+#1643 := [unit-resolution #891 #1642]: #160
+#1644 := [unit-resolution #410 #1643 #1612]: up_10
+#1645 := [unit-resolution #980 #1644]: #194
+#167 := (or #165 #166)
+#878 := [asserted]: #167
+#1646 := [unit-resolution #878 #1642]: #165
+#1647 := [unit-resolution #430 #1646]: up_13
+#1648 := [unit-resolution #1021 #1647]: #209
+#1649 := [unit-resolution #608 #1648 #1611]: up_33
+#1650 := [unit-resolution #1121 #1649]: #234
+decl up_18 :: bool
+#39 := up_18
+#185 := (not up_18)
+#979 := (or #161 #185)
+#200 := (or #185 #161)
+#981 := (iff #200 #979)
+#982 := [rewrite]: #981
+#978 := [asserted]: #200
+#985 := [mp #978 #982]: #979
+#1651 := [unit-resolution #985 #1644]: #185
+#468 := (or up_8 up_16 up_18 up_19)
+#41 := (or up_16 up_8)
+#42 := (or up_19 #41)
+#43 := (or up_18 #42)
+#471 := (iff #43 #468)
+#459 := (or up_8 up_16)
+#462 := (or up_19 #459)
+#465 := (or up_18 #462)
+#469 := (iff #465 #468)
+#470 := [rewrite]: #469
+#466 := (iff #43 #465)
+#463 := (iff #42 #462)
+#460 := (iff #41 #459)
+#461 := [rewrite]: #460
+#464 := [monotonicity #461]: #463
+#467 := [monotonicity #464]: #466
+#472 := [trans #467 #470]: #471
+#458 := [asserted]: #43
+#473 := [mp #458 #472]: #468
+#1652 := [unit-resolution #473 #1651 #1627 #1626]: up_19
+#186 := (not up_19)
+#1068 := (or #186 #226)
+#230 := (or #226 #186)
+#1070 := (iff #230 #1068)
+#1071 := [rewrite]: #1070
+#1067 := [asserted]: #230
+#1074 := [mp #1067 #1071]: #1068
+#1653 := [unit-resolution #1074 #1652]: #226
+#1654 := [unit-resolution #579 #1653 #1650 #895 #1645]: false
+#1656 := [lemma #1654]: #1655
+#1657 := [unit-resolution #1656 #1611 #1627 #1626 #895 #1612]: #166
+#1640 := (or up_12 up_35 up_7 up_22 up_16 up_8 up_32)
+#1628 := [unit-resolution #1625 #1457 #1456 #895 #1611 #1612]: up_29
+#1629 := [unit-resolution #1074 #1628]: #186
+#1630 := [unit-resolution #473 #1629 #1627 #1626]: up_18
+#960 := (or #185 #193)
+#196 := (or #193 #185)
+#961 := (iff #196 #960)
+#962 := [rewrite]: #961
+#958 := [asserted]: #196
+#965 := [mp #958 #962]: #960
+#1631 := [unit-resolution #965 #1630]: #193
+#1632 := [unit-resolution #505 #1631 #1456 #1457]: up_23
+#1633 := [unit-resolution #1126 #1632]: #242
+#1634 := [unit-resolution #608 #1633 #1611]: up_24
+#1635 := [unit-resolution #985 #1630]: #161
+#1636 := [unit-resolution #410 #1635 #1612]: up_9
+#1637 := [unit-resolution #886 #1636]: #165
+#1638 := [unit-resolution #430 #1637]: up_13
+#1639 := [unit-resolution #1021 #1638 #1634]: false
+#1641 := [lemma #1639]: #1640
+#1658 := [unit-resolution #1641 #1657 #1612 #1611 #1627 #1626 #895]: up_22
+#1659 := [unit-resolution #1016 #1658]: #209
+#1660 := [unit-resolution #608 #1659 #1611]: up_33
+#1661 := [unit-resolution #1121 #1660]: #234
+#1662 := [unit-resolution #1026 #1658]: #170
+#1663 := [unit-resolution #430 #1662]: up_11
+#1664 := [unit-resolution #886 #1663]: #160
+#1665 := [unit-resolution #410 #1664 #1612]: up_10
+#1666 := [unit-resolution #980 #1665]: #194
+#1667 := [unit-resolution #579 #1666 #895 #1661]: up_29
+#1668 := [unit-resolution #985 #1665]: #185
+#1669 := [unit-resolution #473 #1668 #1627 #1626]: up_19
+#1670 := [unit-resolution #1074 #1669 #1667]: false
+#1672 := [lemma #1670]: #1671
+#1698 := [unit-resolution #1672 #895 #1626 #1627 #1612]: up_35
+#1609 := (or #250 up_34 up_51 up_54)
+#316 := (not up_53)
+#1599 := [hypothesis]: up_35
+#1275 := (or #250 #291)
+#293 := (or #291 #250)
+#1277 := (iff #293 #1275)
+#1278 := [rewrite]: #1277
+#1274 := [asserted]: #293
+#1281 := [mp #1274 #1278]: #1275
+#1600 := [unit-resolution #1281 #1599]: #291
+#777 := (or up_46 up_55)
+decl up_4 :: bool
+#10 := up_4
+#783 := (or up_4 up_46 up_55)
+#1514 := (iff #783 #777)
+#1509 := (or false up_46 up_55)
+#1512 := (iff #1509 #777)
+#1513 := [rewrite]: #1512
+#1510 := (iff #783 #1509)
+#1485 := (iff up_4 false)
+#11 := (not up_4)
+#1488 := (iff #11 #1485)
+#1486 := (iff #1485 #11)
+#1487 := [rewrite]: #1486
+#1489 := [symm #1487]: #1488
+#371 := [asserted]: #11
+#1490 := [mp #371 #1489]: #1485
+#1511 := [monotonicity #1490]: #1510
+#1515 := [trans #1511 #1513]: #1514
+#135 := (or up_55 up_46)
+#136 := (or up_4 #135)
+#786 := (iff #136 #783)
+#780 := (or up_4 #777)
+#784 := (iff #780 #783)
+#785 := [rewrite]: #784
+#781 := (iff #136 #780)
+#778 := (iff #135 #777)
+#779 := [rewrite]: #778
+#782 := [monotonicity #779]: #781
+#787 := [trans #782 #785]: #786
+#776 := [asserted]: #136
+#788 := [mp #776 #787]: #783
+#1516 := [mp #788 #1515]: #777
+#1601 := [unit-resolution #1516 #1600]: up_55
+#324 := (not up_55)
+#1376 := (or #316 #324)
+#327 := (or #324 #316)
+#1377 := (iff #327 #1376)
+#1378 := [rewrite]: #1377
+#1374 := [asserted]: #327
+#1381 := [mp #1374 #1378]: #1376
+#1602 := [unit-resolution #1381 #1601]: #316
+#754 := (or up_43 up_51 up_53 up_54)
+#127 := (or up_51 up_43)
+#128 := (or up_54 #127)
+#129 := (or up_53 #128)
+#757 := (iff #129 #754)
+#745 := (or up_43 up_51)
+#748 := (or up_54 #745)
+#751 := (or up_53 #748)
+#755 := (iff #751 #754)
+#756 := [rewrite]: #755
+#752 := (iff #129 #751)
+#749 := (iff #128 #748)
+#746 := (iff #127 #745)
+#747 := [rewrite]: #746
+#750 := [monotonicity #747]: #749
+#753 := [monotonicity #750]: #752
+#758 := [trans #753 #756]: #757
+#744 := [asserted]: #129
+#759 := [mp #744 #758]: #754
+#1603 := [unit-resolution #759 #1602 #1598 #1597]: up_43
+decl up_44 :: bool
+#103 := up_44
+#283 := (not up_44)
+#1280 := (or #250 #283)
+#294 := (or #283 #250)
+#1282 := (iff #294 #1280)
+#1283 := [rewrite]: #1282
+#1279 := [asserted]: #294
+#1286 := [mp #1279 #1283]: #1280
+#1604 := [unit-resolution #1286 #1599]: #283
+#1380 := (or #284 #324)
+#328 := (or #324 #284)
+#1382 := (iff #328 #1380)
+#1383 := [rewrite]: #1382
+#1379 := [asserted]: #328
+#1386 := [mp #1379 #1383]: #1380
+#1606 := [unit-resolution #1386 #1601]: #284
+#680 := (or up_34 up_42 up_44 up_45)
+#105 := (or up_42 up_34)
+#106 := (or up_45 #105)
+#107 := (or up_44 #106)
+#683 := (iff #107 #680)
+#671 := (or up_34 up_42)
+#674 := (or up_45 #671)
+#677 := (or up_44 #674)
+#681 := (iff #677 #680)
+#682 := [rewrite]: #681
+#678 := (iff #107 #677)
+#675 := (iff #106 #674)
+#672 := (iff #105 #671)
+#673 := [rewrite]: #672
+#676 := [monotonicity #673]: #675
+#679 := [monotonicity #676]: #678
+#684 := [trans #679 #682]: #683
+#670 := [asserted]: #107
+#685 := [mp #670 #684]: #680
+#1607 := [unit-resolution #685 #1606 #1605 #1604]: up_42
+#277 := (or #275 #276)
+#1217 := [asserted]: #277
+#1608 := [unit-resolution #1217 #1607 #1603]: false
+#1610 := [lemma #1608]: #1609
+#1699 := [unit-resolution #1610 #1698 #1598 #1597]: up_34
+#1125 := (or #234 #243)
+#247 := (or #243 #234)
+#1127 := (iff #247 #1125)
+#1128 := [rewrite]: #1127
+#1124 := [asserted]: #247
+#1131 := [mp #1124 #1128]: #1125
+#1700 := [unit-resolution #1131 #1699]: #234
+#1130 := (or #202 #243)
+#248 := (or #243 #202)
+#1132 := (iff #248 #1130)
+#1133 := [rewrite]: #1132
+#1129 := [asserted]: #248
+#1136 := [mp #1129 #1133]: #1130
+#1701 := [unit-resolution #1136 #1699]: #202
+#1696 := (or up_12 up_7 up_23 up_16 up_8 up_32 up_31)
+#1684 := [hypothesis]: #234
+#1685 := [unit-resolution #1683 #1457 #1612 #1673]: #194
+#1686 := [unit-resolution #579 #1685 #895 #1684]: up_29
+#1687 := [unit-resolution #1074 #1686]: #186
+#1688 := [unit-resolution #473 #1687 #1627 #1626]: up_18
+#1689 := [unit-resolution #965 #1688]: #193
+#1690 := [unit-resolution #505 #1689 #1457 #1673]: up_22
+#1691 := [unit-resolution #1026 #1690]: #170
+#1692 := [unit-resolution #985 #1688]: #161
+#1693 := [unit-resolution #410 #1692 #1612]: up_9
+#1694 := [unit-resolution #886 #1693]: #165
+#1695 := [unit-resolution #430 #1694 #1691]: false
+#1697 := [lemma #1695]: #1696
+#1702 := [unit-resolution #1697 #1701 #1612 #1627 #1626 #895 #1700]: up_12
+#1703 := [unit-resolution #891 #1702]: #160
+#1704 := [unit-resolution #410 #1703 #1612]: up_10
+#1705 := [unit-resolution #980 #1704]: #194
+#1706 := [unit-resolution #579 #1705 #895 #1700]: up_29
+#1707 := [unit-resolution #985 #1704]: #185
+#1708 := [unit-resolution #473 #1707 #1627 #1626]: up_19
+#1709 := [unit-resolution #1074 #1708 #1706]: false
+#1711 := [lemma #1709]: #1710
+#1712 := [unit-resolution #1711 #1598 #1626 #1612 #1627 #1597]: up_32
+#1234 := (or #235 #276)
+#281 := (or #276 #235)
+#1236 := (iff #281 #1234)
+#1237 := [rewrite]: #1236
+#1233 := [asserted]: #281
+#1240 := [mp #1233 #1237]: #1234
+#1713 := [unit-resolution #1240 #1712]: #276
+#1714 := [unit-resolution #759 #1713 #1598 #1597]: up_53
+#1395 := (or #284 #316)
+#331 := (or #316 #284)
+#1397 := (iff #331 #1395)
+#1398 := [rewrite]: #1397
+#1394 := [asserted]: #331
+#1401 := [mp #1394 #1398]: #1395
+#1715 := [unit-resolution #1401 #1714]: #284
+#1716 := [unit-resolution #1230 #1712]: #275
+#1717 := [unit-resolution #1381 #1714]: #324
+#1718 := [unit-resolution #1516 #1717]: up_46
+#1270 := (or #283 #291)
+#292 := (or #291 #283)
+#1272 := (iff #292 #1270)
+#1273 := [rewrite]: #1272
+#1269 := [asserted]: #292
+#1276 := [mp #1269 #1273]: #1270
+#1719 := [unit-resolution #1276 #1718]: #283
+#1720 := [unit-resolution #685 #1719 #1716 #1715]: up_34
+#1721 := [unit-resolution #1136 #1720]: #202
+#1722 := [unit-resolution #1281 #1718]: #250
+#244 := (or #242 #243)
+#1113 := [asserted]: #244
+#1723 := [unit-resolution #1113 #1720]: #242
+#1724 := [unit-resolution #608 #1723 #1722]: up_24
+#1725 := [unit-resolution #1016 #1724]: #201
+#1726 := [unit-resolution #1021 #1724]: #170
+#1727 := [unit-resolution #430 #1726]: up_11
+#1728 := [unit-resolution #878 #1727]: #166
+#1729 := [unit-resolution #505 #1728 #1725 #1721]: up_20
+#1730 := [unit-resolution #886 #1727]: #160
+#1731 := [unit-resolution #410 #1730 #1612]: up_10
+#1732 := [unit-resolution #970 #1731 #1729]: false
+#1734 := [lemma #1732]: #1733
+#1858 := [unit-resolution #1734 #1627 #1597 #1626 #1612]: up_51
+#310 := (or #308 #309)
+#1321 := [asserted]: #310
+#1859 := [unit-resolution #1321 #1858]: #309
+decl up_58 :: bool
+#139 := up_58
+#337 := (not up_58)
+decl up_49 :: bool
+#115 := up_49
+#300 := (not up_49)
+#1324 := (or #300 #308)
+#311 := (or #308 #300)
+#1325 := (iff #311 #1324)
+#1326 := [rewrite]: #1325
+#1322 := [asserted]: #311
+#1329 := [mp #1322 #1326]: #1324
+#1860 := [unit-resolution #1329 #1858]: #300
+decl up_39 :: bool
+#89 := up_39
+#260 := (not up_39)
+decl up_38 :: bool
+#88 := up_38
+decl up_40 :: bool
+#93 := up_40
+#267 := (not up_40)
+decl up_41 :: bool
+#94 := up_41
+#268 := (not up_41)
+#1328 := (or #268 #308)
+#312 := (or #308 #268)
+#1330 := (iff #312 #1328)
+#1331 := [rewrite]: #1330
+#1327 := [asserted]: #312
+#1334 := [mp #1327 #1331]: #1328
+#1861 := [unit-resolution #1334 #1858]: #268
+#1771 := (or up_32 up_16 up_8 up_41 up_49 up_52 up_7)
+#1735 := [unit-resolution #1281 #1698]: #291
+#1736 := [unit-resolution #1516 #1735]: up_55
+#1737 := [unit-resolution #1386 #1736]: #284
+#1738 := [unit-resolution #1286 #1698]: #283
+#259 := (not up_38)
+decl up_50 :: bool
+#116 := up_50
+#301 := (not up_50)
+#1739 := [hypothesis]: #309
+#341 := (not up_59)
+#326 := (or #324 #325)
+#1373 := [asserted]: #326
+#1740 := [unit-resolution #1373 #1736]: #325
+#834 := (or up_56 up_60)
+decl up_3 :: bool
+#8 := up_3
+#840 := (or up_3 up_56 up_60)
+#1522 := (iff #840 #834)
+#1517 := (or false up_56 up_60)
+#1520 := (iff #1517 #834)
+#1521 := [rewrite]: #1520
+#1518 := (iff #840 #1517)
+#1479 := (iff up_3 false)
+#9 := (not up_3)
+#1482 := (iff #9 #1479)
+#1480 := (iff #1479 #9)
+#1481 := [rewrite]: #1480
+#1483 := [symm #1481]: #1482
+#370 := [asserted]: #9
+#1484 := [mp #370 #1483]: #1479
+#1519 := [monotonicity #1484]: #1518
+#1523 := [trans #1519 #1521]: #1522
+#148 := (or up_60 up_56)
+#149 := (or up_3 #148)
+#843 := (iff #149 #840)
+#837 := (or up_3 #834)
+#841 := (iff #837 #840)
+#842 := [rewrite]: #841
+#838 := (iff #149 #837)
+#835 := (iff #148 #834)
+#836 := [rewrite]: #835
+#839 := [monotonicity #836]: #838
+#844 := [trans #839 #842]: #843
+#833 := [asserted]: #149
+#845 := [mp #833 #844]: #840
+#1524 := [mp #845 #1523]: #834
+#1741 := [unit-resolution #1524 #1740]: up_60
+#1442 := (or #341 #345)
+#346 := (or #345 #341)
+#1444 := (iff #346 #1442)
+#1445 := [rewrite]: #1444
+#1441 := [asserted]: #346
+#1448 := [mp #1441 #1445]: #1442
+#1742 := [unit-resolution #1448 #1741]: #341
+#814 := (or up_52 up_58 up_59)
+#143 := (or up_58 up_52)
+#144 := (or up_59 #143)
+#817 := (iff #144 #814)
+#808 := (or up_52 up_58)
+#811 := (or up_59 #808)
+#815 := (iff #811 #814)
+#816 := [rewrite]: #815
+#812 := (iff #144 #811)
+#809 := (iff #143 #808)
+#810 := [rewrite]: #809
+#813 := [monotonicity #810]: #812
+#818 := [trans #813 #816]: #817
+#807 := [asserted]: #144
+#819 := [mp #807 #818]: #814
+#1743 := [unit-resolution #819 #1742 #1739]: up_58
+#1417 := (or #301 #337)
+#339 := (or #337 #301)
+#1419 := (iff #339 #1417)
+#1420 := [rewrite]: #1419
+#1416 := [asserted]: #339
+#1423 := [mp #1416 #1420]: #1417
+#1744 := [unit-resolution #1423 #1743]: #301
+#1745 := [hypothesis]: #300
+decl up_47 :: bool
+#111 := up_47
+#295 := (not up_47)
+decl up_48 :: bool
+#112 := up_48
+decl up_57 :: bool
+#137 := up_57
+#335 := (not up_57)
+#1412 := (or #335 #337)
+#338 := (or #337 #335)
+#1414 := (iff #338 #1412)
+#1415 := [rewrite]: #1414
+#1411 := [asserted]: #338
+#1418 := [mp #1411 #1415]: #1412
+#1746 := [unit-resolution #1418 #1743]: #335
+#790 := (or up_48 up_57)
+#138 := (or up_57 up_48)
+#791 := (iff #138 #790)
+#792 := [rewrite]: #791
+#789 := [asserted]: #138
+#795 := [mp #789 #792]: #790
+#1747 := [unit-resolution #795 #1746]: up_48
+#296 := (not up_48)
+#297 := (or #295 #296)
+#1284 := [asserted]: #297
+#1748 := [unit-resolution #1284 #1747]: #295
+#722 := (or up_39 up_47 up_49 up_50)
+#117 := (or up_47 up_39)
+#118 := (or up_50 #117)
+#119 := (or up_49 #118)
+#725 := (iff #119 #722)
+#713 := (or up_39 up_47)
+#716 := (or up_50 #713)
+#719 := (or up_49 #716)
+#723 := (iff #719 #722)
+#724 := [rewrite]: #723
+#720 := (iff #119 #719)
+#717 := (iff #118 #716)
+#714 := (iff #117 #713)
+#715 := [rewrite]: #714
+#718 := [monotonicity #715]: #717
+#721 := [monotonicity #718]: #720
+#726 := [trans #721 #724]: #725
+#712 := [asserted]: #119
+#727 := [mp #712 #726]: #722
+#1749 := [unit-resolution #727 #1748 #1745 #1744]: up_39
+#261 := (or #259 #260)
+#1165 := [asserted]: #261
+#1750 := [unit-resolution #1165 #1749]: #259
+#1751 := [hypothesis]: #268
+decl up_30 :: bool
+#67 := up_30
+#227 := (not up_30)
+decl up_27 :: bool
+#61 := up_27
+#213 := (not up_25)
+decl up_37 :: bool
+#85 := up_37
+#255 := (not up_37)
+#1291 := (or #255 #296)
+#299 := (or #296 #255)
+#1293 := (iff #299 #1291)
+#1294 := [rewrite]: #1293
+#1290 := [asserted]: #299
+#1297 := [mp #1290 #1294]: #1291
+#1752 := [unit-resolution #1297 #1747]: #255
+decl up_36 :: bool
+#84 := up_36
+#254 := (not up_36)
+#1177 := (or #254 #260)
+#264 := (or #260 #254)
+#1179 := (iff #264 #1177)
+#1180 := [rewrite]: #1179
+#1176 := [asserted]: #264
+#1183 := [mp #1176 #1180]: #1177
+#1753 := [unit-resolution #1183 #1749]: #254
+#616 := (or up_26 up_36 up_37)
+#86 := (or up_37 up_26)
+#87 := (or up_36 #86)
+#619 := (iff #87 #616)
+#610 := (or up_26 up_37)
+#613 := (or up_36 #610)
+#617 := (iff #613 #616)
+#618 := [rewrite]: #617
+#614 := (iff #87 #613)
+#611 := (iff #86 #610)
+#612 := [rewrite]: #611
+#615 := [monotonicity #612]: #614
+#620 := [trans #615 #618]: #619
+#609 := [asserted]: #87
+#621 := [mp #609 #620]: #616
+#1754 := [unit-resolution #621 #1753 #1752]: up_26
+#215 := (or #213 #214)
+#1024 := [asserted]: #215
+#1755 := [unit-resolution #1024 #1754]: #213
+decl up_28 :: bool
+#62 := up_28
+#219 := (not up_28)
+#1182 := (or #219 #260)
+#265 := (or #260 #219)
+#1184 := (iff #265 #1182)
+#1185 := [rewrite]: #1184
+#1181 := [asserted]: #265
+#1188 := [mp #1181 #1185]: #1182
+#1756 := [unit-resolution #1188 #1749]: #219
+decl up_17 :: bool
+#35 := up_17
+#178 := (not up_17)
+#173 := (not up_15)
+#1031 := (or #173 #214)
+#217 := (or #214 #173)
+#1033 := (iff #217 #1031)
+#1034 := [rewrite]: #1033
+#1030 := [asserted]: #217
+#1037 := [mp #1030 #1034]: #1031
+#1757 := [unit-resolution #1037 #1754]: #173
+#1503 := (or up_14 up_15)
+decl up_2 :: bool
+#6 := up_2
+#436 := (or up_2 up_14 up_15)
+#1506 := (iff #436 #1503)
+#1500 := (or false up_14 up_15)
+#1504 := (iff #1500 #1503)
+#1505 := [rewrite]: #1504
+#1501 := (iff #436 #1500)
+#1473 := (iff up_2 false)
+#7 := (not up_2)
+#1476 := (iff #7 #1473)
+#1474 := (iff #1473 #7)
+#1475 := [rewrite]: #1474
+#1477 := [symm #1475]: #1476
+#369 := [asserted]: #7
+#1478 := [mp #369 #1477]: #1473
+#1502 := [monotonicity #1478]: #1501
+#1507 := [trans #1502 #1505]: #1506
+#32 := (or up_15 up_2)
+#33 := (or up_14 #32)
+#439 := (iff #33 #436)
+#429 := (or up_2 up_15)
+#433 := (or up_14 #429)
+#437 := (iff #433 #436)
+#438 := [rewrite]: #437
+#434 := (iff #33 #433)
+#431 := (iff #32 #429)
+#432 := [rewrite]: #431
+#435 := [monotonicity #432]: #434
+#440 := [trans #435 #438]: #439
+#428 := [asserted]: #33
+#441 := [mp #428 #440]: #436
+#1508 := [mp #441 #1507]: #1503
+#1758 := [unit-resolution #1508 #1757]: up_14
+#917 := (or #172 #178)
+#182 := (or #178 #172)
+#919 := (iff #182 #917)
+#920 := [rewrite]: #919
+#916 := [asserted]: #182
+#923 := [mp #916 #920]: #917
+#1759 := [unit-resolution #923 #1758]: #178
+#542 := (or up_17 up_25 up_27 up_28)
+#63 := (or up_25 up_17)
+#64 := (or up_28 #63)
+#65 := (or up_27 #64)
+#545 := (iff #65 #542)
+#533 := (or up_17 up_25)
+#536 := (or up_28 #533)
+#539 := (or up_27 #536)
+#543 := (iff #539 #542)
+#544 := [rewrite]: #543
+#540 := (iff #65 #539)
+#537 := (iff #64 #536)
+#534 := (iff #63 #533)
+#535 := [rewrite]: #534
+#538 := [monotonicity #535]: #537
+#541 := [monotonicity #538]: #540
+#546 := [trans #541 #544]: #545
+#532 := [asserted]: #65
+#547 := [mp #532 #546]: #542
+#1760 := [unit-resolution #547 #1759 #1756 #1755]: up_27
+#218 := (not up_27)
+#1073 := (or #218 #227)
+#231 := (or #227 #218)
+#1075 := (iff #231 #1073)
+#1076 := [rewrite]: #1075
+#1072 := [asserted]: #231
+#1079 := [mp #1072 #1076]: #1073
+#1761 := [unit-resolution #1079 #1760]: #227
+#648 := (or up_30 up_38 up_40 up_41)
+#95 := (or up_38 up_30)
+#96 := (or up_41 #95)
+#97 := (or up_40 #96)
+#651 := (iff #97 #648)
+#639 := (or up_30 up_38)
+#642 := (or up_41 #639)
+#645 := (or up_40 #642)
+#649 := (iff #645 #648)
+#650 := [rewrite]: #649
+#646 := (iff #97 #645)
+#643 := (iff #96 #642)
+#640 := (iff #95 #639)
+#641 := [rewrite]: #640
+#644 := [monotonicity #641]: #643
+#647 := [monotonicity #644]: #646
+#652 := [trans #647 #650]: #651
+#638 := [asserted]: #97
+#653 := [mp #638 #652]: #648
+#1762 := [unit-resolution #653 #1761 #1751 #1750]: up_40
+#1220 := (or #267 #275)
+#278 := (or #275 #267)
+#1221 := (iff #278 #1220)
+#1222 := [rewrite]: #1221
+#1218 := [asserted]: #278
+#1225 := [mp #1218 #1222]: #1220
+#1763 := [unit-resolution #1225 #1762]: #275
+#1764 := [unit-resolution #685 #1763 #1738 #1737]: up_34
+#1064 := (or #218 #226)
+#229 := (or #226 #218)
+#1065 := (iff #229 #1064)
+#1066 := [rewrite]: #1065
+#1062 := [asserted]: #229
+#1069 := [mp #1062 #1066]: #1064
+#1765 := [unit-resolution #1069 #1760]: #226
+#1083 := (or #186 #218)
+#233 := (or #218 #186)
+#1085 := (iff #233 #1083)
+#1086 := [rewrite]: #1085
+#1082 := [asserted]: #233
+#1089 := [mp #1082 #1086]: #1083
+#1766 := [unit-resolution #1089 #1760]: #186
+#1767 := [unit-resolution #473 #1766 #1627 #1626]: up_18
+#969 := (or #185 #194)
+#198 := (or #194 #185)
+#971 := (iff #198 #969)
+#972 := [rewrite]: #971
+#968 := [asserted]: #198
+#975 := [mp #968 #972]: #969
+#1768 := [unit-resolution #975 #1767]: #194
+#1769 := [unit-resolution #579 #1768 #895 #1765]: up_31
+#1770 := [unit-resolution #1131 #1769 #1764]: false
+#1772 := [lemma #1770]: #1771
+#1862 := [unit-resolution #1772 #1627 #1626 #1861 #1860 #1859 #1612]: up_32
+#1239 := (or #235 #267)
+#282 := (or #267 #235)
+#1241 := (iff #282 #1239)
+#1242 := [rewrite]: #1241
+#1238 := [asserted]: #282
+#1245 := [mp #1238 #1242]: #1239
+#1863 := [unit-resolution #1245 #1862]: #267
+#1856 := (or up_12 up_52 up_49 up_41 up_16 up_8 up_7)
+#1828 := [unit-resolution #1772 #1627 #1626 #1751 #1745 #1739 #1612]: up_32
+#1829 := [unit-resolution #1245 #1828]: #267
+#1830 := [unit-resolution #1230 #1828]: #275
+#1826 := (or #170 up_41 up_40 up_16 up_8 up_49 up_12 up_52 up_42)
+#1804 := [hypothesis]: up_13
+#1805 := [unit-resolution #1026 #1804]: #201
+#1806 := [unit-resolution #1021 #1804]: #209
+#1798 := [hypothesis]: #275
+#1782 := [hypothesis]: #267
+#1802 := (or #242 up_42 up_52 up_49 up_41 up_40 up_16 up_8 up_12 up_22)
+#1783 := [hypothesis]: up_33
+#1784 := [unit-resolution #1126 #1783]: #202
+#1785 := [unit-resolution #505 #1784 #1457 #1456]: up_20
+#1786 := [unit-resolution #965 #1785]: #185
+#1787 := [unit-resolution #473 #1786 #1627 #1626]: up_19
+#1078 := (or #186 #227)
+#232 := (or #227 #186)
+#1080 := (iff #232 #1078)
+#1081 := [rewrite]: #1080
+#1077 := [asserted]: #232
+#1084 := [mp #1077 #1081]: #1078
+#1788 := [unit-resolution #1084 #1787]: #227
+#1789 := [unit-resolution #653 #1788 #1751 #1782]: up_38
+#1790 := [unit-resolution #1165 #1789]: #260
+#1780 := (or #337 up_49 up_39)
+#1773 := [hypothesis]: up_58
+#1774 := [unit-resolution #1418 #1773]: #335
+#1775 := [unit-resolution #795 #1774]: up_48
+#1776 := [hypothesis]: #260
+#1777 := [unit-resolution #1423 #1773]: #301
+#1778 := [unit-resolution #727 #1777 #1745 #1776]: up_47
+#1779 := [unit-resolution #1284 #1778 #1775]: false
+#1781 := [lemma #1779]: #1780
+#1791 := [unit-resolution #1781 #1790 #1745]: #337
+#1792 := [unit-resolution #819 #1791 #1739]: up_59
+#1793 := [unit-resolution #1448 #1792]: #345
+#1794 := [unit-resolution #1524 #1793]: up_56
+#1795 := [unit-resolution #1373 #1794]: #324
+#1796 := [unit-resolution #1516 #1795]: up_46
+#1797 := [unit-resolution #1113 #1783]: #243
+#1390 := (or #284 #325)
+#330 := (or #325 #284)
+#1392 := (iff #330 #1390)
+#1393 := [rewrite]: #1392
+#1389 := [asserted]: #330
+#1396 := [mp #1389 #1393]: #1390
+#1799 := [unit-resolution #1396 #1794]: #284
+#1800 := [unit-resolution #685 #1799 #1798 #1797]: up_44
+#1801 := [unit-resolution #1276 #1800 #1796]: false
+#1803 := [lemma #1801]: #1802
+#1807 := [unit-resolution #1803 #1805 #1739 #1745 #1751 #1782 #1627 #1626 #1457 #1798]: #242
+#1808 := [unit-resolution #608 #1807 #1806]: up_35
+#1809 := [unit-resolution #1286 #1808]: #283
+#1810 := [unit-resolution #1281 #1808]: #291
+#1811 := [unit-resolution #1516 #1810]: up_55
+#1812 := [unit-resolution #1386 #1811]: #284
+#1813 := [unit-resolution #685 #1812 #1798 #1809]: up_34
+#1814 := [unit-resolution #1136 #1813]: #202
+#1815 := [unit-resolution #505 #1814 #1457 #1805]: up_20
+#1816 := [unit-resolution #965 #1815]: #185
+#1817 := [unit-resolution #473 #1816 #1627 #1626]: up_19
+#1818 := [unit-resolution #1373 #1811]: #325
+#1819 := [unit-resolution #1524 #1818]: up_60
+#1820 := [unit-resolution #1448 #1819]: #341
+#1821 := [unit-resolution #819 #1820 #1739]: up_58
+#1822 := [unit-resolution #1781 #1821 #1745]: up_39
+#1823 := [unit-resolution #1165 #1822]: #259
+#1824 := [unit-resolution #653 #1823 #1751 #1782]: up_30
+#1825 := [unit-resolution #1084 #1824 #1817]: false
+#1827 := [lemma #1825]: #1826
+#1831 := [unit-resolution #1827 #1457 #1829 #1627 #1626 #1745 #1751 #1739 #1830]: #170
+#1832 := [unit-resolution #430 #1831]: up_11
+#1833 := [unit-resolution #886 #1832]: #160
+#1834 := [unit-resolution #410 #1833 #1612]: up_10
+#1835 := [unit-resolution #985 #1834]: #185
+#1836 := [unit-resolution #473 #1835 #1627 #1626]: up_19
+#1837 := [unit-resolution #1084 #1836]: #227
+#1838 := [unit-resolution #653 #1837 #1751 #1829]: up_38
+#1839 := [unit-resolution #1165 #1838]: #260
+#1840 := [unit-resolution #1781 #1839 #1745]: #337
+#1841 := [unit-resolution #819 #1840 #1739]: up_59
+#1842 := [unit-resolution #1448 #1841]: #345
+#1843 := [unit-resolution #1524 #1842]: up_56
+#1844 := [unit-resolution #1373 #1843]: #324
+#1845 := [unit-resolution #1516 #1844]: up_46
+#1846 := [unit-resolution #1281 #1845]: #250
+#1847 := [unit-resolution #1396 #1843]: #284
+#1848 := [unit-resolution #1276 #1845]: #283
+#1849 := [unit-resolution #685 #1848 #1830 #1847]: up_34
+#1850 := [unit-resolution #1113 #1849]: #242
+#1851 := [unit-resolution #608 #1850 #1846]: up_24
+#1852 := [unit-resolution #970 #1834]: #193
+#1853 := [unit-resolution #1136 #1849]: #202
+#1854 := [unit-resolution #505 #1853 #1457 #1852]: up_22
+#1855 := [unit-resolution #1016 #1854 #1851]: false
+#1857 := [lemma #1855]: #1856
+#1864 := [unit-resolution #1857 #1859 #1860 #1861 #1627 #1626 #1612]: up_12
+#1865 := [unit-resolution #891 #1864]: #160
+#1866 := [unit-resolution #410 #1865 #1612]: up_10
+#1867 := [unit-resolution #985 #1866]: #185
+#1868 := [unit-resolution #473 #1867 #1627 #1626]: up_19
+#1869 := [unit-resolution #1084 #1868]: #227
+#1870 := [unit-resolution #653 #1869 #1861 #1863]: up_38
+#1871 := [unit-resolution #1165 #1870]: #260
+#1872 := [unit-resolution #1781 #1871 #1860]: #337
+#1873 := [unit-resolution #819 #1872 #1859]: up_59
+#1874 := [unit-resolution #1448 #1873]: #345
+#1875 := [unit-resolution #1524 #1874]: up_56
+#1876 := [unit-resolution #1396 #1875]: #284
+#1877 := [unit-resolution #1230 #1862]: #275
+#1878 := [unit-resolution #1373 #1875]: #324
+#1879 := [unit-resolution #1516 #1878]: up_46
+#1880 := [unit-resolution #1276 #1879]: #283
+#1881 := [unit-resolution #685 #1880 #1877 #1876]: up_34
+#1882 := [unit-resolution #878 #1864]: #165
+#1883 := [unit-resolution #430 #1882]: up_13
+#1884 := [unit-resolution #1021 #1883]: #209
+#1885 := [unit-resolution #1281 #1879]: #250
+#1886 := [unit-resolution #608 #1885 #1884]: up_33
+#1887 := [unit-resolution #1113 #1886 #1881]: false
+#1889 := [lemma #1887]: #1888
+#2026 := [unit-resolution #1889 #1597 #1612 #1626]: up_16
+#908 := (or #172 #177)
+#180 := (or #177 #172)
+#909 := (iff #180 #908)
+#910 := [rewrite]: #909
+#906 := [asserted]: #180
+#913 := [mp #906 #910]: #908
+#2027 := [unit-resolution #913 #2026]: #172
+#2028 := [unit-resolution #1508 #2027]: up_15
+#2029 := [unit-resolution #1037 #2028]: #214
+#1027 := (or #173 #213)
+#216 := (or #213 #173)
+#1028 := (iff #216 #1027)
+#1029 := [rewrite]: #1028
+#1025 := [asserted]: #216
+#1032 := [mp #1025 #1029]: #1027
+#2030 := [unit-resolution #1032 #2028]: #213
+#179 := (or #177 #178)
+#905 := [asserted]: #179
+#2031 := [unit-resolution #905 #2026]: #178
+#1917 := (or #226 up_54 up_26 up_17 up_25)
+#1890 := [hypothesis]: #214
+#1891 := [hypothesis]: #213
+#1892 := [hypothesis]: #178
+#1893 := [hypothesis]: up_29
+#1894 := [unit-resolution #1069 #1893]: #218
+#1895 := [unit-resolution #547 #1894 #1892 #1891]: up_28
+#1187 := (or #219 #254)
+#266 := (or #254 #219)
+#1189 := (iff #266 #1187)
+#1190 := [rewrite]: #1189
+#1186 := [asserted]: #266
+#1193 := [mp #1186 #1190]: #1187
+#1896 := [unit-resolution #1193 #1895]: #254
+#1897 := [unit-resolution #621 #1896 #1890]: up_37
+#1898 := [unit-resolution #1297 #1897]: #296
+#1899 := [unit-resolution #795 #1898]: up_57
+#1900 := [unit-resolution #1418 #1899]: #337
+#1901 := [unit-resolution #1188 #1895]: #260
+#1287 := (or #255 #295)
+#298 := (or #295 #255)
+#1288 := (iff #298 #1287)
+#1289 := [rewrite]: #1288
+#1285 := [asserted]: #298
+#1292 := [mp #1285 #1289]: #1287
+#1902 := [unit-resolution #1292 #1897]: #295
+#1422 := (or #301 #335)
+#340 := (or #335 #301)
+#1424 := (iff #340 #1422)
+#1425 := [rewrite]: #1424
+#1421 := [asserted]: #340
+#1428 := [mp #1421 #1425]: #1422
+#1903 := [unit-resolution #1428 #1899]: #301
+#1904 := [unit-resolution #727 #1903 #1902 #1901]: up_49
+#1333 := (or #300 #309)
+#313 := (or #309 #300)
+#1335 := (iff #313 #1333)
+#1336 := [rewrite]: #1335
+#1332 := [asserted]: #313
+#1339 := [mp #1332 #1336]: #1333
+#1905 := [unit-resolution #1339 #1904]: #309
+#1906 := [unit-resolution #819 #1905 #1900]: up_59
+#1907 := [unit-resolution #1448 #1906]: #345
+#1908 := [unit-resolution #1524 #1907]: up_56
+#1909 := [unit-resolution #1329 #1904]: #308
+#1172 := (or #219 #259)
+#263 := (or #259 #219)
+#1174 := (iff #263 #1172)
+#1175 := [rewrite]: #1174
+#1171 := [asserted]: #263
+#1178 := [mp #1171 #1175]: #1172
+#1910 := [unit-resolution #1178 #1895]: #259
+#228 := (or #226 #227)
+#1061 := [asserted]: #228
+#1911 := [unit-resolution #1061 #1893]: #227
+#1343 := (or #268 #300)
+#315 := (or #300 #268)
+#1345 := (iff #315 #1343)
+#1346 := [rewrite]: #1345
+#1342 := [asserted]: #315
+#1349 := [mp #1342 #1346]: #1343
+#1912 := [unit-resolution #1349 #1904]: #268
+#1913 := [unit-resolution #653 #1912 #1911 #1910]: up_40
+#1229 := (or #267 #276)
+#280 := (or #276 #267)
+#1231 := (iff #280 #1229)
+#1232 := [rewrite]: #1231
+#1228 := [asserted]: #280
+#1235 := [mp #1228 #1232]: #1229
+#1914 := [unit-resolution #1235 #1913]: #276
+#1915 := [unit-resolution #759 #1914 #1909 #1597]: up_53
+#1385 := (or #316 #325)
+#329 := (or #325 #316)
+#1387 := (iff #329 #1385)
+#1388 := [rewrite]: #1387
+#1384 := [asserted]: #329
+#1391 := [mp #1384 #1388]: #1385
+#1916 := [unit-resolution #1391 #1915 #1908]: false
+#1918 := [lemma #1916]: #1917
+#2032 := [unit-resolution #1918 #1597 #2029 #2031 #2030]: #226
+#2010 := (or up_12 up_29 up_7 up_54 up_26)
+#1993 := (or up_35 up_12 up_54 up_26 up_29 up_7)
+#1955 := (or #170 up_54 up_26 up_29 up_12 up_35 up_7)
+#1940 := [unit-resolution #1625 #1805 #907 #1457 #1611 #1612]: up_32
+#1941 := [unit-resolution #1240 #1940]: #276
+#1942 := [unit-resolution #1230 #1940]: #275
+#1943 := [unit-resolution #608 #1806 #1611]: up_33
+#1944 := [unit-resolution #1113 #1943]: #243
+#1925 := (or #325 up_34 up_42)
+#1919 := [hypothesis]: up_56
+#1920 := [unit-resolution #1373 #1919]: #324
+#1921 := [unit-resolution #1516 #1920]: up_46
+#1922 := [unit-resolution #1396 #1919]: #284
+#1923 := [unit-resolution #685 #1922 #1605 #1798]: up_44
+#1924 := [unit-resolution #1276 #1923 #1921]: false
+#1926 := [lemma #1924]: #1925
+#1945 := [unit-resolution #1926 #1944 #1942]: #325
+#1946 := [unit-resolution #1524 #1945]: up_60
+#1947 := [unit-resolution #1448 #1946]: #341
+#1938 := (or #308 up_26 up_59)
+#1927 := [hypothesis]: up_51
+#1928 := [unit-resolution #1329 #1927]: #300
+#1929 := [hypothesis]: #341
+#1930 := [unit-resolution #1321 #1927]: #309
+#1931 := [unit-resolution #819 #1930 #1929]: up_58
+#1932 := [unit-resolution #1781 #1931 #1928]: up_39
+#1933 := [unit-resolution #1183 #1932]: #254
+#1934 := [unit-resolution #1418 #1931]: #335
+#1935 := [unit-resolution #795 #1934]: up_48
+#1936 := [unit-resolution #1297 #1935]: #255
+#1937 := [unit-resolution #621 #1936 #1933 #1890]: false
+#1939 := [lemma #1937]: #1938
+#1948 := [unit-resolution #1939 #1947 #1890]: #308
+#1949 := [unit-resolution #759 #1948 #1941 #1597]: up_53
+#1950 := [unit-resolution #1381 #1949]: #324
+#1951 := [unit-resolution #1516 #1950]: up_46
+#1952 := [unit-resolution #1401 #1949]: #284
+#1953 := [unit-resolution #685 #1952 #1944 #1942]: up_44
+#1954 := [unit-resolution #1276 #1953 #1951]: false
+#1956 := [lemma #1954]: #1955
+#1980 := [unit-resolution #1956 #1611 #1890 #907 #1457 #1597 #1612]: #170
+#1981 := [unit-resolution #430 #1980]: up_11
+#1982 := [unit-resolution #886 #1981]: #160
+#1983 := [unit-resolution #410 #1982 #1612]: up_10
+#1984 := [unit-resolution #980 #1983]: #194
+#1985 := [unit-resolution #1979 #1611 #907 #1984 #1457 #1612]: up_32
+#1970 := (or #235 up_34 up_54 up_26)
+#1957 := [hypothesis]: up_32
+#1958 := [unit-resolution #1240 #1957]: #276
+#1959 := [unit-resolution #1230 #1957]: #275
+#1960 := [unit-resolution #1926 #1959 #1605]: #325
+#1961 := [unit-resolution #1524 #1960]: up_60
+#1962 := [unit-resolution #1448 #1961]: #341
+#1963 := [unit-resolution #1939 #1962 #1890]: #308
+#1964 := [unit-resolution #759 #1963 #1958 #1597]: up_53
+#1965 := [unit-resolution #1381 #1964]: #324
+#1966 := [unit-resolution #1516 #1965]: up_46
+#1967 := [unit-resolution #1401 #1964]: #284
+#1968 := [unit-resolution #685 #1967 #1605 #1959]: up_44
+#1969 := [unit-resolution #1276 #1968 #1966]: false
+#1971 := [lemma #1969]: #1970
+#1986 := [unit-resolution #1971 #1985 #1597 #1890]: up_34
+#1987 := [unit-resolution #1113 #1986]: #242
+#1988 := [unit-resolution #608 #1987 #1611]: up_24
+#1989 := [unit-resolution #970 #1983]: #193
+#1990 := [unit-resolution #1136 #1986]: #202
+#1991 := [unit-resolution #505 #1990 #1457 #1989]: up_22
+#1992 := [unit-resolution #1016 #1991 #1988]: false
+#1994 := [lemma #1992]: #1993
+#1995 := [unit-resolution #1994 #1457 #1597 #1890 #907 #1612]: up_35
+#1996 := [unit-resolution #1281 #1995]: #291
+#1997 := [unit-resolution #1516 #1996]: up_55
+#1998 := [unit-resolution #1373 #1997]: #325
+#1999 := [unit-resolution #1524 #1998]: up_60
+#2000 := [unit-resolution #1448 #1999]: #341
+#2001 := [unit-resolution #1939 #2000 #1890]: #308
+#2002 := [unit-resolution #1610 #2001 #1995 #1597]: up_34
+#2003 := [unit-resolution #1131 #2002]: #234
+#2004 := [unit-resolution #1381 #1997]: #316
+#2005 := [unit-resolution #759 #2001 #2004 #1597]: up_43
+#2006 := [unit-resolution #1240 #2005]: #235
+#2007 := [unit-resolution #1136 #2002]: #202
+#2008 := [unit-resolution #1683 #2007 #1612 #1457]: #194
+#2009 := [unit-resolution #579 #2008 #2006 #907 #2003]: false
+#2011 := [lemma #2009]: #2010
+#2033 := [unit-resolution #2011 #2032 #1612 #1597 #2029]: up_12
+#2034 := [unit-resolution #891 #2033]: #160
+#2035 := [unit-resolution #410 #2034 #1612]: up_10
+#2036 := [unit-resolution #980 #2035]: #194
+#2037 := [unit-resolution #878 #2033]: #165
+#2038 := [unit-resolution #430 #2037]: up_13
+#2039 := [unit-resolution #1021 #2038]: #209
+#2024 := (or #234 up_26 up_54 up_24)
+#2012 := [hypothesis]: #209
+#2013 := [hypothesis]: up_31
+#2014 := [unit-resolution #1121 #2013]: #242
+#2015 := [unit-resolution #608 #2014 #2012]: up_35
+#2016 := [unit-resolution #1131 #2013]: #243
+#2017 := [unit-resolution #1610 #2016 #2015 #1597]: up_51
+#2018 := [unit-resolution #1939 #2017 #1890]: up_59
+#2019 := [unit-resolution #1448 #2018]: #345
+#2020 := [unit-resolution #1281 #2015]: #291
+#2021 := [unit-resolution #1516 #2020]: up_55
+#2022 := [unit-resolution #1373 #2021]: #325
+#2023 := [unit-resolution #1524 #2022 #2019]: false
+#2025 := [lemma #2023]: #2024
+#2040 := [unit-resolution #2025 #2029 #1597 #2039]: #234
+#2041 := [unit-resolution #579 #2040 #2032 #2036]: up_32
+#2042 := [unit-resolution #1240 #2041]: #276
+#2043 := [unit-resolution #1971 #2041 #1597 #2029]: up_34
+#2044 := [unit-resolution #1113 #2043]: #242
+#2045 := [unit-resolution #608 #2044 #2039]: up_35
+#2046 := [unit-resolution #1281 #2045]: #291
+#2047 := [unit-resolution #1516 #2046]: up_55
+#2048 := [unit-resolution #1381 #2047]: #316
+#2049 := [unit-resolution #759 #2048 #2042 #1597]: up_51
+#2050 := [unit-resolution #1373 #2047]: #325
+#2051 := [unit-resolution #1524 #2050]: up_60
+#2052 := [unit-resolution #1448 #2051]: #341
+#2053 := [unit-resolution #1939 #2052 #2049 #2029]: false
+#2055 := [lemma #2053]: #2054
+#2065 := [unit-resolution #2055 #1612 #1626]: up_54
+#1447 := (or #317 #345)
+#347 := (or #345 #317)
+#1449 := (iff #347 #1447)
+#1450 := [rewrite]: #1449
+#1446 := [asserted]: #347
+#1453 := [mp #1446 #1450]: #1447
+#2066 := [unit-resolution #1453 #2065]: #345
+#2067 := [unit-resolution #1524 #2066]: up_56
+#2083 := (or #275 up_7 up_12 up_8)
+#2063 := [hypothesis]: up_42
+#2064 := [unit-resolution #1230 #2063]: #235
+#2068 := [unit-resolution #1373 #2067]: #324
+#2069 := [unit-resolution #1516 #2068]: up_46
+#2070 := [unit-resolution #1281 #2069]: #250
+#2071 := [unit-resolution #1672 #2064 #1626 #2070 #1612]: up_16
+#2072 := [unit-resolution #913 #2071]: #172
+#2073 := [unit-resolution #1508 #2072]: up_15
+#2074 := [unit-resolution #1032 #2073]: #213
+#2075 := [unit-resolution #905 #2071]: #178
+#1452 := (or #317 #341)
+#348 := (or #341 #317)
+#1454 := (iff #348 #1452)
+#1455 := [rewrite]: #1454
+#1451 := [asserted]: #348
+#1458 := [mp #1451 #1455]: #1452
+#2076 := [unit-resolution #1458 #2065]: #341
+#2077 := [unit-resolution #1225 #2063]: #267
+#2061 := (or #226 up_59 up_40 up_17 up_25)
+#2056 := [unit-resolution #653 #1910 #1782 #1911]: up_41
+#2057 := [unit-resolution #1349 #2056]: #300
+#1338 := (or #268 #309)
+#314 := (or #309 #268)
+#1340 := (iff #314 #1338)
+#1341 := [rewrite]: #1340
+#1337 := [asserted]: #314
+#1344 := [mp #1337 #1341]: #1338
+#2058 := [unit-resolution #1344 #2056]: #309
+#2059 := [unit-resolution #819 #2058 #1929]: up_58
+#2060 := [unit-resolution #1781 #2059 #2057 #1901]: false
+#2062 := [lemma #2060]: #2061
+#2078 := [unit-resolution #2062 #2077 #2076 #2075 #2074]: #226
+#2079 := [unit-resolution #1625 #2078 #2070 #1457 #2064 #1612]: up_22
+#2080 := [unit-resolution #1979 #2078 #2070 #1457 #2064 #1612]: up_21
+#2081 := [unit-resolution #1683 #2080 #1612 #1457]: up_23
+#2082 := [unit-resolution #983 #2081 #2079]: false
+#2084 := [lemma #2082]: #2083
+#2085 := [unit-resolution #2084 #1457 #1612 #1626]: #275
+#2086 := [unit-resolution #1926 #2085 #2067]: up_34
+#2087 := [unit-resolution #1136 #2086]: #202
+#2088 := [unit-resolution #1113 #2086]: #242
+#2089 := [unit-resolution #608 #2088 #2070]: up_24
+#2090 := [unit-resolution #1016 #2089]: #201
+#2091 := [unit-resolution #505 #2090 #1457 #2087]: up_20
+#2092 := [unit-resolution #970 #2091]: #161
+#2093 := [unit-resolution #1021 #2089]: #170
+#2094 := [unit-resolution #430 #2093]: up_11
+#2095 := [unit-resolution #886 #2094]: #160
+#2096 := [unit-resolution #410 #2095 #2092 #1612]: false
+#2098 := [lemma #2096]: #2097
+#2102 := [unit-resolution #2098 #2100 #2101]: up_12
+#2103 := [unit-resolution #891 #2102]: #160
+#2104 := [unit-resolution #410 #2103 #2100]: up_10
+#2105 := [unit-resolution #980 #2104]: #194
+#2106 := [unit-resolution #2055 #2100 #2101]: up_54
+#2107 := [unit-resolution #1453 #2106]: #345
+#2108 := [unit-resolution #1524 #2107]: up_56
+#2109 := [unit-resolution #1373 #2108]: #324
+#2110 := [unit-resolution #1516 #2109]: up_46
+#2111 := [unit-resolution #1281 #2110]: #250
+#2112 := [unit-resolution #878 #2102]: #165
+#2113 := [unit-resolution #430 #2112]: up_13
+#2114 := [unit-resolution #1021 #2113]: #209
+#2115 := [unit-resolution #608 #2114 #2111]: up_33
+#2116 := [unit-resolution #1121 #2115]: #234
+#2117 := [unit-resolution #1276 #2110]: #283
+#2118 := [unit-resolution #1396 #2108]: #284
+#2119 := [unit-resolution #1113 #2115]: #243
+#2120 := [unit-resolution #685 #2119 #2118 #2117]: up_42
+#2121 := [unit-resolution #1230 #2120]: #235
+#2122 := [unit-resolution #579 #2121 #2116 #2105]: up_29
+#2123 := [unit-resolution #1225 #2120]: #267
+#2124 := [unit-resolution #1458 #2106]: #341
+#2125 := [unit-resolution #1672 #2121 #2101 #2111 #2100]: up_16
+#2126 := [unit-resolution #905 #2125]: #178
+#2127 := [unit-resolution #2062 #2126 #2124 #2123 #2122]: up_25
+#2128 := [unit-resolution #913 #2125]: #172
+#2129 := [unit-resolution #1508 #2128]: up_15
+#2130 := [unit-resolution #1032 #2129 #2127]: false
+#2131 := [lemma #2130]: #150
+#1494 := (or up_5 up_6)
+decl up_1 :: bool
+#4 := up_1
+#379 := (or up_1 up_5 up_6)
+#1497 := (iff #379 #1494)
+#1491 := (or false up_5 up_6)
+#1495 := (iff #1491 #1494)
+#1496 := [rewrite]: #1495
+#1492 := (iff #379 #1491)
+#1467 := (iff up_1 false)
+#5 := (not up_1)
+#1470 := (iff #5 #1467)
+#1463 := (iff #1467 #5)
+#1468 := [rewrite]: #1463
+#1471 := [symm #1468]: #1470
+#368 := [asserted]: #5
+#1472 := [mp #368 #1471]: #1467
+#1493 := [monotonicity #1472]: #1492
+#1498 := [trans #1493 #1496]: #1497
+#14 := (or up_6 up_1)
+#15 := (or up_5 #14)
+#382 := (iff #15 #379)
+#373 := (or up_1 up_6)
+#376 := (or up_5 #373)
+#380 := (iff #376 #379)
+#381 := [rewrite]: #380
+#377 := (iff #15 #376)
+#374 := (iff #14 #373)
+#375 := [rewrite]: #374
+#378 := [monotonicity #375]: #377
+#383 := [trans #378 #381]: #382
+#372 := [asserted]: #15
+#384 := [mp #372 #383]: #379
+#1499 := [mp #384 #1498]: #1494
+#2138 := [unit-resolution #1499 #2131]: up_6
+#151 := (not up_6)
+#927 := (or #151 #172)
+#184 := (or #172 #151)
+#929 := (iff #184 #927)
+#930 := [rewrite]: #929
+#926 := [asserted]: #184
+#933 := [mp #926 #930]: #927
+#2139 := [unit-resolution #933 #2138]: #172
+#2140 := [unit-resolution #1508 #2139]: up_15
+#2147 := [unit-resolution #1037 #2140]: #214
+#2159 := [unit-resolution #2011 #2147]: #2158
+#2160 := [unit-resolution #2159 #907 #1612 #1457]: up_54
+#2161 := [unit-resolution #1453 #2160]: #345
+#2162 := [unit-resolution #1524 #2161]: up_56
+#2163 := [unit-resolution #1926 #2162 #2157 #1605]: false
+#2165 := [lemma #2163]: #2164
+#2166 := [unit-resolution #2165 #1605 #1612 #1611 #1457]: up_29
+#2148 := (or #226 up_54)
+#2141 := [unit-resolution #1032 #2140]: #213
+#922 := (or #151 #178)
+#183 := (or #178 #151)
+#924 := (iff #183 #922)
+#925 := [rewrite]: #924
+#921 := [asserted]: #183
+#928 := [mp #921 #925]: #922
+#2142 := [unit-resolution #928 #2138]: #178
+#2149 := [unit-resolution #1918 #2147 #2142 #2141]: #2148
+#2167 := [unit-resolution #2149 #2166]: up_54
+#2154 := (or #226 up_34 up_59)
+#2143 := (or #226 up_59 up_40)
+#2144 := [unit-resolution #2062 #2142 #2141]: #2143
+#2145 := [unit-resolution #2144 #1893 #1929]: up_40
+#2146 := [unit-resolution #1225 #2145]: #275
+#2150 := [unit-resolution #2149 #1893]: up_54
+#2151 := [unit-resolution #1453 #2150]: #345
+#2152 := [unit-resolution #1524 #2151]: up_56
+#2153 := [unit-resolution #1926 #2152 #2146 #1605]: false
+#2155 := [lemma #2153]: #2154
+#2168 := [unit-resolution #2155 #2166 #1605]: up_59
+#2169 := [unit-resolution #1458 #2168 #2167]: false
+#2171 := [lemma #2169]: #2170
+#2172 := [unit-resolution #2171 #1612 #1611 #1457]: up_34
+#2173 := [unit-resolution #1136 #2172]: #202
+#2174 := [unit-resolution #1113 #2172]: #242
+#2175 := [unit-resolution #608 #2174 #1611]: up_24
+#2176 := [unit-resolution #1016 #2175]: #201
+#2177 := [unit-resolution #505 #2176 #1457 #2173]: up_20
+#2178 := [unit-resolution #970 #2177]: #161
+#2179 := [unit-resolution #1021 #2175]: #170
+#2180 := [unit-resolution #430 #2179]: up_11
+#2181 := [unit-resolution #886 #2180]: #160
+#2182 := [unit-resolution #410 #2181 #2178 #1612]: false
+#2184 := [lemma #2182]: #2183
+#2235 := [unit-resolution #2184 #1457 #1611]: up_7
+#157 := (or #155 #156)
+#856 := [asserted]: #157
+#2236 := [unit-resolution #856 #2235]: #156
+#2299 := (or up_34 up_35 up_12)
+#2283 := (or #186 up_34)
+#2185 := [hypothesis]: up_19
+#2191 := [unit-resolution #1084 #2185]: #227
+#2186 := [unit-resolution #1089 #2185]: #218
+#2187 := (or up_27 up_28)
+#2188 := [unit-resolution #547 #2142 #2141]: #2187
+#2189 := [unit-resolution #2188 #2186]: up_28
+#2192 := [unit-resolution #1178 #2189]: #259
+#2265 := [unit-resolution #1193 #2189]: #254
+#2266 := (or up_36 up_37)
+#2267 := [unit-resolution #621 #2147]: #2266
+#2268 := [unit-resolution #2267 #2265]: up_37
+#2269 := [unit-resolution #1292 #2268]: #295
+#2190 := [unit-resolution #1188 #2189]: #260
+#2270 := [unit-resolution #1297 #2268]: #296
+#2271 := [unit-resolution #795 #2270]: up_57
+#2272 := [unit-resolution #1428 #2271]: #301
+#2273 := [unit-resolution #727 #2272 #2190 #2269]: up_49
+#2274 := [unit-resolution #1349 #2273]: #268
+#2275 := [unit-resolution #653 #2274 #2192 #2191]: up_40
+#2276 := [unit-resolution #1225 #2275]: #275
+#2277 := [unit-resolution #1418 #2271]: #337
+#2278 := [unit-resolution #1339 #2273]: #309
+#2279 := [unit-resolution #819 #2278 #2277]: up_59
+#2280 := [unit-resolution #1448 #2279]: #345
+#2281 := [unit-resolution #1524 #2280]: up_56
+#2282 := [unit-resolution #1926 #2281 #2276 #1605]: false
+#2284 := [lemma #2282]: #2283
+#2292 := [unit-resolution #2284 #1605]: #186
+#2223 := (or up_8 up_18 up_19)
+#912 := (or #151 #177)
+#181 := (or #177 #151)
+#914 := (iff #181 #912)
+#915 := [rewrite]: #914
+#911 := [asserted]: #181
+#918 := [mp #911 #915]: #912
+#2222 := [unit-resolution #918 #2138]: #177
+#2224 := [unit-resolution #473 #2222]: #2223
+#2293 := [unit-resolution #2224 #2292 #2236]: up_18
+#2257 := (or #235 up_34)
+#2252 := (or #235 up_34 up_54)
+#2253 := [unit-resolution #1971 #2147]: #2252
+#2254 := [unit-resolution #2253 #1957 #1605]: up_54
+#2255 := [unit-resolution #1453 #2254]: #345
+#2256 := [unit-resolution #1524 #2255 #1960]: false
+#2258 := [lemma #2256]: #2257
+#2294 := [unit-resolution #2258 #1605]: #235
+#2290 := (or up_29 up_35 up_32 up_12 #185)
+#2200 := [hypothesis]: up_18
+#2206 := (or #185 up_29 up_32 up_12 up_22)
+#2201 := [unit-resolution #965 #2200]: #193
+#2202 := [unit-resolution #505 #2201 #1457 #1456]: up_23
+#2203 := [unit-resolution #975 #2200]: #194
+#2204 := [unit-resolution #579 #2203 #907 #895]: up_31
+#2205 := [unit-resolution #1141 #2204 #2202]: false
+#2207 := [lemma #2205]: #2206
+#2285 := [unit-resolution #2207 #907 #895 #1457 #2200]: up_22
+#2286 := [unit-resolution #1016 #2285]: #209
+#2287 := [unit-resolution #579 #907 #895 #2203]: up_31
+#2288 := [unit-resolution #1121 #2287]: #242
+#2289 := [unit-resolution #608 #2288 #2286 #1611]: false
+#2291 := [lemma #2289]: #2290
+#2295 := [unit-resolution #2291 #2294 #1611 #1457 #2293]: up_29
+#2296 := [unit-resolution #2149 #2295]: up_54
+#2297 := [unit-resolution #2155 #2295 #1605]: up_59
+#2298 := [unit-resolution #1458 #2297 #2296]: false
+#2300 := [lemma #2298]: #2299
+#2301 := [unit-resolution #2300 #1457 #1611]: up_34
+#2302 := [unit-resolution #1136 #2301]: #202
+#2303 := [unit-resolution #1113 #2301]: #242
+#2304 := [unit-resolution #608 #2303 #1611]: up_24
+#2305 := [unit-resolution #1016 #2304]: #201
+#2306 := [unit-resolution #505 #2305 #1457 #2302]: up_20
+#2307 := [unit-resolution #965 #2306]: #185
+#2308 := [unit-resolution #2224 #2307 #2236]: up_19
+#2309 := [unit-resolution #957 #2306]: #194
+#2310 := [unit-resolution #1131 #2301]: #234
+#2311 := [unit-resolution #1074 #2308]: #226
+#2312 := [unit-resolution #579 #2311 #2310 #2309]: up_32
+#2313 := [unit-resolution #1245 #2312]: #267
+#2198 := (or #186 up_59 up_40)
+#2193 := [unit-resolution #653 #2192 #1782 #2191]: up_41
+#2194 := [unit-resolution #1349 #2193]: #300
+#2195 := [unit-resolution #1344 #2193]: #309
+#2196 := [unit-resolution #819 #2195 #1929]: up_58
+#2197 := [unit-resolution #1781 #2196 #2194 #2190]: false
+#2199 := [lemma #2197]: #2198
+#2314 := [unit-resolution #2199 #2313 #2308]: up_59
+#2315 := [unit-resolution #1448 #2314]: #345
+#2316 := [unit-resolution #1524 #2315]: up_56
+#2317 := [unit-resolution #1084 #2308]: #227
+#2318 := [unit-resolution #1089 #2308]: #218
+#2319 := [unit-resolution #2188 #2318]: up_28
+#2320 := [unit-resolution #1178 #2319]: #259
+#2321 := [unit-resolution #653 #2313 #2320 #2317]: up_41
+#2322 := [unit-resolution #1334 #2321]: #308
+#2323 := [unit-resolution #1240 #2312]: #276
+#2324 := [unit-resolution #1458 #2314]: #317
+#2325 := [unit-resolution #759 #2324 #2323 #2322]: up_53
+#2326 := [unit-resolution #1391 #2325 #2316]: false
+#2328 := [lemma #2326]: #2327
+#2337 := [unit-resolution #2328 #1611]: up_12
+#2338 := [unit-resolution #878 #2337]: #165
+#2339 := [unit-resolution #430 #2338]: up_13
+#2340 := [unit-resolution #1021 #2339]: #209
+#2341 := [unit-resolution #608 #2340 #1611]: up_33
+#2342 := [unit-resolution #1113 #2341]: #243
+#2343 := [unit-resolution #2258 #2342]: #235
+#2344 := [unit-resolution #1121 #2341]: #234
+#2345 := [unit-resolution #2284 #2342]: #186
+#2346 := [unit-resolution #891 #2337]: #160
+#2335 := (or #194 up_9 up_19)
+#2329 := [hypothesis]: #186
+#2330 := [unit-resolution #975 #1674]: #185
+#2331 := [unit-resolution #2224 #2330 #2329]: up_8
+#2332 := [hypothesis]: #160
+#2333 := [unit-resolution #410 #1678 #2332]: up_7
+#2334 := [unit-resolution #856 #2333 #2331]: false
+#2336 := [lemma #2334]: #2335
+#2347 := [unit-resolution #2336 #2346 #2345]: #194
+#2348 := [unit-resolution #579 #2347 #2344 #2343]: up_29
+#2349 := [unit-resolution #2149 #2348]: up_54
+#2350 := [unit-resolution #2155 #2348 #2342]: up_59
+#2351 := [unit-resolution #1458 #2350 #2349]: false
+#2352 := [lemma #2351]: up_35
+#2353 := [unit-resolution #1281 #2352]: #291
+#2354 := [unit-resolution #1516 #2353]: up_55
+#2355 := [unit-resolution #1373 #2354]: #325
+#2356 := [unit-resolution #1524 #2355]: up_60
+#2357 := [unit-resolution #1453 #2356]: #317
+#2358 := [unit-resolution #2149 #2357]: #226
+#2359 := [unit-resolution #1448 #2356]: #341
+#2217 := (or #308 up_59)
+#2218 := [unit-resolution #1939 #2147]: #2217
+#2360 := [unit-resolution #2218 #2359]: #308
+#2361 := [unit-resolution #1381 #2354]: #316
+#2362 := [unit-resolution #759 #2357 #2361 #2360]: up_43
+#2363 := [unit-resolution #1235 #2362]: #267
+#2364 := [unit-resolution #2199 #2363 #2359]: #186
+#1145 := (or #209 #250)
+#252 := (or #250 #209)
+#1147 := (iff #252 #1145)
+#1148 := [rewrite]: #1147
+#1144 := [asserted]: #252
+#1151 := [mp #1144 #1148]: #1145
+#2365 := [unit-resolution #1151 #2352]: #209
+#2230 := (or #234 up_54 up_24)
+#2231 := [unit-resolution #2025 #2147]: #2230
+#2366 := [unit-resolution #2231 #2357 #2365]: #234
+#2367 := [unit-resolution #1240 #2362]: #235
+#2368 := [unit-resolution #579 #2367 #2366 #2358]: up_21
+#2369 := [unit-resolution #2336 #2368 #2364]: up_9
+#870 := (or #155 #160)
+#163 := (or #160 #155)
+#871 := (iff #163 #870)
+#872 := [rewrite]: #871
+#868 := [asserted]: #163
+#875 := [mp #868 #872]: #870
+#2370 := [unit-resolution #875 #2369]: #155
+#2371 := [unit-resolution #891 #2369]: #166
+[unit-resolution #2159 #2371 #2370 #2358 #2357]: false
+unsat
+
+NdJwa8pysYWhn57eCXiqFA 1731
+#2 := false
+decl up_1 :: (-> int bool)
+decl ?x1!0 :: int
+#54 := ?x1!0
+#55 := (up_1 ?x1!0)
+#58 := (not #55)
+decl ?x2!1 :: int
+#66 := ?x2!1
+#67 := (up_1 ?x2!1)
+#85 := (or #55 #67)
+#88 := (not #85)
+#91 := (and #55 #88)
+#68 := (or #67 #55)
+#69 := (not #68)
+#63 := (not #58)
+#75 := (and #63 #69)
+#92 := (iff #75 #91)
+#89 := (iff #69 #88)
+#86 := (iff #68 #85)
+#87 := [rewrite]: #86
+#90 := [monotonicity #87]: #89
+#83 := (iff #63 #55)
+#84 := [rewrite]: #83
+#93 := [monotonicity #84 #90]: #92
+#6 := (:var 1 int)
+#7 := (up_1 #6)
+#4 := (:var 0 int)
+#5 := (up_1 #4)
+#29 := (or #5 #7)
+#32 := (forall (vars (?x2 int)) #29)
+#38 := (not #5)
+#39 := (or #38 #32)
+#44 := (forall (vars (?x1 int)) #39)
+#47 := (not #44)
+#78 := (~ #47 #75)
+#56 := (or #5 #55)
+#57 := (forall (vars (?x2 int)) #56)
+#59 := (or #58 #57)
+#60 := (not #59)
+#76 := (~ #60 #75)
+#70 := (not #57)
+#71 := (~ #70 #69)
+#72 := [sk]: #71
+#64 := (~ #63 #63)
+#65 := [refl]: #64
+#77 := [nnf-neg #65 #72]: #76
+#61 := (~ #47 #60)
+#62 := [sk]: #61
+#79 := [trans #62 #77]: #78
+#8 := (or #7 #5)
+#9 := (forall (vars (?x2 int)) #8)
+#10 := (implies #5 #9)
+#11 := (forall (vars (?x1 int)) #10)
+#12 := (not #11)
+#48 := (iff #12 #47)
+#45 := (iff #11 #44)
+#42 := (iff #10 #39)
+#35 := (implies #5 #32)
+#40 := (iff #35 #39)
+#41 := [rewrite]: #40
+#36 := (iff #10 #35)
+#33 := (iff #9 #32)
+#30 := (iff #8 #29)
+#31 := [rewrite]: #30
+#34 := [quant-intro #31]: #33
+#37 := [monotonicity #34]: #36
+#43 := [trans #37 #41]: #42
+#46 := [quant-intro #43]: #45
+#49 := [monotonicity #46]: #48
+#28 := [asserted]: #12
+#52 := [mp #28 #49]: #47
+#80 := [mp~ #52 #79]: #75
+#81 := [mp #80 #93]: #91
+#94 := [and-elim #81]: #88
+#95 := [not-or-elim #94]: #58
+#82 := [and-elim #81]: #55
+[unit-resolution #82 #95]: false
+unsat
+
+dWWXDEA5bUZjEafLPXbSkA 3321
+#2 := false
+decl up_1 :: (-> T1 T2 bool)
+#5 := (:var 0 T2)
+decl uf_3 :: T1
+#11 := uf_3
+#12 := (up_1 uf_3 #5)
+#560 := (pattern #12)
+#57 := (not #12)
+#561 := (forall (vars (?x3 T2)) (:pat #560) #57)
+decl uf_4 :: T2
+#14 := uf_4
+#15 := (up_1 uf_3 uf_4)
+decl uf_2 :: T1
+#7 := uf_2
+#136 := (= uf_2 uf_3)
+#543 := (iff #15 #136)
+#4 := (:var 1 T1)
+#6 := (up_1 #4 #5)
+#553 := (pattern #6)
+#8 := (= #4 uf_2)
+#9 := (iff #6 #8)
+#554 := (forall (vars (?x1 T1) (?x2 T2)) (:pat #553) #9)
+#10 := (forall (vars (?x1 T1) (?x2 T2)) #9)
+#557 := (iff #10 #554)
+#555 := (iff #9 #9)
+#556 := [refl]: #555
+#558 := [quant-intro #556]: #557
+#47 := (~ #10 #10)
+#45 := (~ #9 #9)
+#46 := [refl]: #45
+#48 := [nnf-pos #46]: #47
+#33 := [asserted]: #10
+#49 := [mp~ #33 #48]: #10
+#559 := [mp #49 #558]: #554
+#227 := (not #554)
+#185 := (or #227 #543)
+#135 := (= uf_3 uf_2)
+#205 := (iff #15 #135)
+#528 := (or #227 #205)
+#190 := (iff #528 #185)
+#192 := (iff #185 #185)
+#530 := [rewrite]: #192
+#201 := (iff #205 #543)
+#223 := (iff #135 #136)
+#137 := [rewrite]: #223
+#544 := [monotonicity #137]: #201
+#191 := [monotonicity #544]: #190
+#531 := [trans #191 #530]: #190
+#189 := [quant-inst]: #528
+#532 := [mp #189 #531]: #185
+#539 := [unit-resolution #532 #559]: #543
+decl ?x3!0 :: T2
+#50 := ?x3!0
+#51 := (up_1 uf_3 ?x3!0)
+#224 := (iff #51 #136)
+#155 := (or #227 #224)
+#222 := (iff #51 #135)
+#228 := (or #227 #222)
+#229 := (iff #228 #155)
+#545 := (iff #155 #155)
+#547 := [rewrite]: #545
+#215 := (iff #222 #224)
+#226 := [monotonicity #137]: #215
+#208 := [monotonicity #226]: #229
+#202 := [trans #208 #547]: #229
+#225 := [quant-inst]: #228
+#334 := [mp #225 #202]: #155
+#537 := [unit-resolution #334 #559]: #224
+#541 := (not #224)
+#527 := (or #541 #136)
+#63 := (not #15)
+#540 := [hypothesis]: #63
+#68 := (or #15 #51)
+#60 := (forall (vars (?x3 T2)) #57)
+#69 := (or #63 #60)
+#76 := (and #68 #69)
+#70 := (and #69 #68)
+#77 := (iff #70 #76)
+#78 := [rewrite]: #77
+#13 := (exists (vars (?x3 T2)) #12)
+#35 := (not #13)
+#36 := (iff #15 #35)
+#71 := (~ #36 #70)
+#61 := (~ #35 #60)
+#58 := (~ #57 #57)
+#59 := [refl]: #58
+#62 := [nnf-neg #59]: #61
+#54 := (not #35)
+#55 := (~ #54 #51)
+#42 := (~ #13 #51)
+#39 := [sk]: #42
+#56 := [nnf-neg #39]: #55
+#66 := (~ #15 #15)
+#67 := [refl]: #66
+#64 := (~ #63 #63)
+#65 := [refl]: #64
+#72 := [nnf-pos #65 #67 #56 #62]: #71
+#16 := (iff #13 #15)
+#17 := (not #16)
+#37 := (iff #17 #36)
+#38 := [rewrite]: #37
+#34 := [asserted]: #17
+#41 := [mp #34 #38]: #36
+#73 := [mp~ #41 #72]: #70
+#74 := [mp #73 #78]: #76
+#75 := [and-elim #74]: #68
+#526 := [unit-resolution #75 #540]: #51
+#549 := (not #51)
+#550 := (or #541 #549 #136)
+#551 := [def-axiom]: #550
+#233 := [unit-resolution #551 #526]: #527
+#249 := [unit-resolution #233 #537]: #136
+#213 := (not #136)
+#533 := (not #543)
+#250 := (or #533 #213)
+#534 := (or #533 #15 #213)
+#529 := [def-axiom]: #534
+#251 := [unit-resolution #529 #540]: #250
+#237 := [unit-resolution #251 #249 #539]: false
+#252 := [lemma #237]: #15
+#566 := (or #63 #561)
+#567 := (iff #69 #566)
+#564 := (iff #60 #561)
+#562 := (iff #57 #57)
+#563 := [refl]: #562
+#565 := [quant-intro #563]: #564
+#568 := [monotonicity #565]: #567
+#79 := [and-elim #74]: #69
+#569 := [mp #79 #568]: #566
+#535 := [unit-resolution #569 #252]: #561
+#536 := (not #561)
+#538 := (or #536 #63)
+#176 := [quant-inst]: #538
+[unit-resolution #176 #252 #535]: false
+unsat
+
+iGZ7b2jaCnn82lPL6oIDZA 3465
+WARNING: failed to find a pattern for quantifier (quantifier id: k!11)
+#2 := false
+decl up_1 :: (-> T1 T2 bool)
+#5 := (:var 0 T2)
+decl uf_4 :: T1
+#18 := uf_4
+#19 := (up_1 uf_4 #5)
+#635 := (pattern #19)
+#116 := (not #19)
+#636 := (forall (vars (?x6 T2)) (:pat #635) #116)
+decl uf_3 :: T2
+#14 := uf_3
+#21 := (up_1 uf_4 uf_3)
+decl uf_2 :: T1
+#7 := uf_2
+#195 := (= uf_2 uf_4)
+#602 := (iff #21 #195)
+#4 := (:var 1 T1)
+#6 := (up_1 #4 #5)
+#612 := (pattern #6)
+#8 := (= #4 uf_2)
+#9 := (iff #6 #8)
+#613 := (forall (vars (?x1 T1) (?x2 T2)) (:pat #612) #9)
+#10 := (forall (vars (?x1 T1) (?x2 T2)) #9)
+#616 := (iff #10 #613)
+#614 := (iff #9 #9)
+#615 := [refl]: #614
+#617 := [quant-intro #615]: #616
+#56 := (~ #10 #10)
+#54 := (~ #9 #9)
+#55 := [refl]: #54
+#57 := [nnf-pos #55]: #56
+#39 := [asserted]: #10
+#58 := [mp~ #39 #57]: #10
+#618 := [mp #58 #617]: #613
+#286 := (not #613)
+#244 := (or #286 #602)
+#194 := (= uf_4 uf_2)
+#264 := (iff #21 #194)
+#587 := (or #286 #264)
+#249 := (iff #587 #244)
+#251 := (iff #244 #244)
+#589 := [rewrite]: #251
+#260 := (iff #264 #602)
+#282 := (iff #194 #195)
+#196 := [rewrite]: #282
+#603 := [monotonicity #196]: #260
+#250 := [monotonicity #603]: #249
+#590 := [trans #250 #589]: #249
+#248 := [quant-inst]: #587
+#591 := [mp #248 #590]: #244
+#598 := [unit-resolution #591 #618]: #602
+decl ?x6!3 :: T2
+#63 := ?x6!3
+#64 := (up_1 uf_4 ?x6!3)
+#283 := (iff #64 #195)
+#214 := (or #286 #283)
+#281 := (iff #64 #194)
+#287 := (or #286 #281)
+#288 := (iff #287 #214)
+#604 := (iff #214 #214)
+#606 := [rewrite]: #604
+#274 := (iff #281 #283)
+#285 := [monotonicity #196]: #274
+#267 := [monotonicity #285]: #288
+#261 := [trans #267 #606]: #288
+#284 := [quant-inst]: #287
+#393 := [mp #284 #261]: #214
+#596 := [unit-resolution #393 #618]: #283
+#600 := (not #283)
+#586 := (or #600 #195)
+#122 := (not #21)
+#599 := [hypothesis]: #122
+#127 := (or #21 #64)
+#119 := (forall (vars (?x6 T2)) #116)
+#128 := (or #122 #119)
+#135 := (and #127 #128)
+#129 := (and #128 #127)
+#136 := (iff #129 #135)
+#137 := [rewrite]: #136
+#20 := (exists (vars (?x6 T2)) #19)
+#42 := (not #20)
+#43 := (iff #21 #42)
+#130 := (~ #43 #129)
+#120 := (~ #42 #119)
+#117 := (~ #116 #116)
+#118 := [refl]: #117
+#121 := [nnf-neg #118]: #120
+#113 := (not #42)
+#114 := (~ #113 #64)
+#88 := (~ #20 #64)
+#89 := [sk]: #88
+#115 := [nnf-neg #89]: #114
+#125 := (~ #21 #21)
+#126 := [refl]: #125
+#123 := (~ #122 #122)
+#124 := [refl]: #123
+#131 := [nnf-pos #124 #126 #115 #121]: #130
+#22 := (iff #20 #21)
+#23 := (not #22)
+#44 := (iff #23 #43)
+#45 := [rewrite]: #44
+#41 := [asserted]: #23
+#48 := [mp #41 #45]: #43
+#132 := [mp~ #48 #131]: #129
+#133 := [mp #132 #137]: #135
+#134 := [and-elim #133]: #127
+#585 := [unit-resolution #134 #599]: #64
+#608 := (not #64)
+#609 := (or #600 #608 #195)
+#610 := [def-axiom]: #609
+#292 := [unit-resolution #610 #585]: #586
+#308 := [unit-resolution #292 #596]: #195
+#272 := (not #195)
+#592 := (not #602)
+#309 := (or #592 #272)
+#593 := (or #592 #21 #272)
+#588 := [def-axiom]: #593
+#310 := [unit-resolution #588 #599]: #309
+#296 := [unit-resolution #310 #308 #598]: false
+#311 := [lemma #296]: #21
+#641 := (or #122 #636)
+#642 := (iff #128 #641)
+#639 := (iff #119 #636)
+#637 := (iff #116 #116)
+#638 := [refl]: #637
+#640 := [quant-intro #638]: #639
+#643 := [monotonicity #640]: #642
+#138 := [and-elim #133]: #128
+#644 := [mp #138 #643]: #641
+#594 := [unit-resolution #644 #311]: #636
+#595 := (not #636)
+#597 := (or #595 #122)
+#235 := [quant-inst]: #597
+[unit-resolution #235 #311 #594]: false
+unsat
+
+eTjcfu6S5eyz+xNJ7SVluw 1246
+#2 := false
+decl up_1 :: (-> T1 bool)
+decl uf_2 :: T1
+#4 := uf_2
+#5 := (up_1 uf_2)
+decl uf_3 :: T1
+#13 := uf_3
+#14 := (up_1 uf_3)
+#34 := (not #5)
+#35 := (or #34 #14)
+#38 := (not #35)
+#15 := (implies #5 #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]: #5
+#6 := (:var 0 T1)
+#7 := (up_1 #6)
+#536 := (pattern #7)
+#10 := (not #7)
+#537 := (forall (vars (?x2 T1)) (:pat #536) #10)
+#11 := (forall (vars (?x2 T1)) #10)
+#540 := (iff #11 #537)
+#538 := (iff #10 #10)
+#539 := [refl]: #538
+#541 := [quant-intro #539]: #540
+#8 := (exists (vars (?x1 T1)) #7)
+#9 := (not #8)
+#45 := (~ #9 #11)
+#50 := (~ #10 #10)
+#51 := [refl]: #50
+#59 := [nnf-neg #51]: #45
+#12 := (ite #5 #9 #11)
+#57 := (iff #12 #9)
+#1 := true
+#52 := (ite true #9 #11)
+#55 := (iff #52 #9)
+#56 := [rewrite]: #55
+#53 := (iff #12 #52)
+#48 := (iff #5 true)
+#49 := [iff-true #41]: #48
+#54 := [monotonicity #49]: #53
+#58 := [trans #54 #56]: #57
+#32 := [asserted]: #12
+#47 := [mp #32 #58]: #9
+#60 := [mp~ #47 #59]: #11
+#542 := [mp #60 #541]: #537
+#119 := (not #537)
+#206 := (or #119 #34)
+#120 := [quant-inst]: #206
+[unit-resolution #120 #542 #41]: false
+unsat
+
+anG1bKU/YVTHEmc1Eh/ZXw 331
+#2 := false
+#4 := 3::int
+#5 := (= 3::int 3::int)
+#6 := (not #5)
+#30 := (iff #6 false)
+#1 := true
+#25 := (not true)
+#28 := (iff #25 false)
+#29 := [rewrite]: #28
+#26 := (iff #6 #25)
+#23 := (iff #5 true)
+#24 := [rewrite]: #23
+#27 := [monotonicity #24]: #26
+#31 := [trans #27 #29]: #30
+#22 := [asserted]: #6
+[mp #22 #31]: false
+unsat
+
+lHpRCTa744ODgmii2zARAw 334
+#2 := false
+#4 := 3::real
+#5 := (= 3::real 3::real)
+#6 := (not #5)
+#30 := (iff #6 false)
+#1 := true
+#25 := (not true)
+#28 := (iff #25 false)
+#29 := [rewrite]: #28
+#26 := (iff #6 #25)
+#23 := (iff #5 true)
+#24 := [rewrite]: #23
+#27 := [monotonicity #24]: #26
+#31 := [trans #27 #29]: #30
+#22 := [asserted]: #6
+[mp #22 #31]: false
+unsat
+
+AinXomcY4W1L/t0ZtkDhBg 524
+#2 := false
+#7 := 4::int
+#5 := 1::int
+#4 := 3::int
+#6 := (+ 3::int 1::int)
+#8 := (= #6 4::int)
+#9 := (not #8)
+#39 := (iff #9 false)
+#1 := true
+#34 := (not true)
+#37 := (iff #34 false)
+#38 := [rewrite]: #37
+#35 := (iff #9 #34)
+#32 := (iff #8 true)
+#27 := (= 4::int 4::int)
+#30 := (iff #27 true)
+#31 := [rewrite]: #30
+#28 := (iff #8 #27)
+#26 := [rewrite]: #8
+#29 := [monotonicity #26]: #28
+#33 := [trans #29 #31]: #32
+#36 := [monotonicity #33]: #35
+#40 := [trans #36 #38]: #39
+#25 := [asserted]: #9
+[mp #25 #40]: false
+unsat
+
+WxMdOusjxqQwBPorpXBkFQ 815
+#2 := false
+decl uf_1 :: int
+#4 := uf_1
+decl uf_3 :: int
+#6 := uf_3
+#9 := (+ uf_3 uf_1)
+decl uf_2 :: int
+#5 := uf_2
+#10 := (+ uf_2 #9)
+#7 := (+ uf_2 uf_3)
+#8 := (+ uf_1 #7)
+#11 := (= #8 #10)
+#12 := (not #11)
+#51 := (iff #12 false)
+#1 := true
+#46 := (not true)
+#49 := (iff #46 false)
+#50 := [rewrite]: #49
+#47 := (iff #12 #46)
+#44 := (iff #11 true)
+#39 := (= #8 #8)
+#42 := (iff #39 true)
+#43 := [rewrite]: #42
+#40 := (iff #11 #39)
+#37 := (= #10 #8)
+#29 := (+ uf_1 uf_3)
+#32 := (+ uf_2 #29)
+#35 := (= #32 #8)
+#36 := [rewrite]: #35
+#33 := (= #10 #32)
+#30 := (= #9 #29)
+#31 := [rewrite]: #30
+#34 := [monotonicity #31]: #33
+#38 := [trans #34 #36]: #37
+#41 := [monotonicity #38]: #40
+#45 := [trans #41 #43]: #44
+#48 := [monotonicity #45]: #47
+#52 := [trans #48 #50]: #51
+#28 := [asserted]: #12
+[mp #28 #52]: false
+unsat
+
+K7g37p4yZoVyQcabYS4I2w 754
+#2 := false
+#5 := 3::int
+#6 := 8::int
+#7 := (<= 3::int 8::int)
+#8 := (ite #7 8::int 3::int)
+#4 := 5::int
+#9 := (< 5::int #8)
+#10 := (not #9)
+#50 := (iff #10 false)
+#1 := true
+#45 := (not true)
+#48 := (iff #45 false)
+#49 := [rewrite]: #48
+#46 := (iff #10 #45)
+#43 := (iff #9 true)
+#38 := (< 5::int 8::int)
+#41 := (iff #38 true)
+#42 := [rewrite]: #41
+#39 := (iff #9 #38)
+#36 := (= #8 8::int)
+#31 := (ite true 8::int 3::int)
+#34 := (= #31 8::int)
+#35 := [rewrite]: #34
+#32 := (= #8 #31)
+#29 := (iff #7 true)
+#30 := [rewrite]: #29
+#33 := [monotonicity #30]: #32
+#37 := [trans #33 #35]: #36
+#40 := [monotonicity #37]: #39
+#44 := [trans #40 #42]: #43
+#47 := [monotonicity #44]: #46
+#51 := [trans #47 #49]: #50
+#26 := [asserted]: #10
+[mp #26 #51]: false
+unsat
+
+eCmVy21SUmWImXZDJNOfzA 6496
+#2 := false
+#7 := 0::real
+decl uf_2 :: real
+#5 := uf_2
+#143 := 2::real
+#144 := (* 2::real uf_2)
+#165 := (<= #144 0::real)
+#188 := (not #165)
+#88 := (>= uf_2 0::real)
+#166 := (or #88 #165)
+#191 := (not #166)
+decl uf_1 :: real
+#4 := uf_1
+#76 := (>= uf_1 0::real)
+#89 := (not #88)
+#146 := (* 2::real uf_1)
+#167 := (<= #146 0::real)
+#199 := (not #167)
+#263 := [hypothesis]: #88
+#147 := (+ #146 #144)
+#168 := (<= #147 0::real)
+#169 := (ite #88 #167 #168)
+#194 := (not #169)
+#186 := (or #166 #89)
+#187 := [def-axiom]: #186
+#271 := [unit-resolution #187 #263]: #166
+#170 := (ite #76 #166 #169)
+#205 := (not #170)
+#6 := (+ uf_1 uf_2)
+#64 := (>= #6 0::real)
+#269 := (or #64 #89)
+#65 := (not #64)
+#262 := [hypothesis]: #65
+#174 := (>= #144 0::real)
+#175 := (or #89 #174)
+#230 := (not #175)
+#257 := [hypothesis]: #230
+#225 := (or #175 #88)
+#226 := [def-axiom]: #225
+#258 := [unit-resolution #226 #257]: #88
+#227 := (not #174)
+#228 := (or #175 #227)
+#229 := [def-axiom]: #228
+#259 := [unit-resolution #229 #257]: #227
+#260 := [th-lemma #259 #258]: false
+#261 := [lemma #260]: #175
+#172 := (>= #146 0::real)
+#171 := (>= #147 0::real)
+#173 := (ite #88 #171 #172)
+#176 := (ite #76 #173 #175)
+#233 := (not #176)
+#264 := (or #64 #233)
+#177 := (ite #64 #170 #176)
+#182 := (not #177)
+#36 := -1::real
+#38 := (* -1::real uf_2)
+#95 := (ite #88 uf_2 #38)
+#107 := (* -1::real #95)
+#37 := (* -1::real uf_1)
+#83 := (ite #76 uf_1 #37)
+#106 := (* -1::real #83)
+#108 := (+ #106 #107)
+#39 := (+ #37 #38)
+#71 := (ite #64 #6 #39)
+#109 := (+ #71 #108)
+#110 := (<= #109 0::real)
+#115 := (not #110)
+#183 := (iff #115 #182)
+#180 := (iff #110 #177)
+#150 := -2::real
+#152 := (* -2::real uf_2)
+#155 := (ite #88 #152 0::real)
+#151 := (* -2::real uf_1)
+#153 := (+ #151 #152)
+#154 := (ite #88 #153 #151)
+#156 := (ite #76 #154 #155)
+#148 := (ite #88 #146 #147)
+#145 := (ite #88 0::real #144)
+#149 := (ite #76 #145 #148)
+#157 := (ite #64 #149 #156)
+#162 := (<= #157 0::real)
+#178 := (iff #162 #177)
+#179 := [rewrite]: #178
+#163 := (iff #110 #162)
+#160 := (= #109 #157)
+#133 := (+ uf_1 #38)
+#134 := (ite #88 #133 #6)
+#131 := (+ #37 uf_2)
+#132 := (ite #88 #39 #131)
+#135 := (ite #76 #132 #134)
+#140 := (+ #71 #135)
+#158 := (= #140 #157)
+#159 := [rewrite]: #158
+#141 := (= #109 #140)
+#138 := (= #108 #135)
+#125 := (ite #88 #38 uf_2)
+#123 := (ite #76 #37 uf_1)
+#128 := (+ #123 #125)
+#136 := (= #128 #135)
+#137 := [rewrite]: #136
+#129 := (= #108 #128)
+#126 := (= #107 #125)
+#127 := [rewrite]: #126
+#121 := (= #106 #123)
+#124 := [rewrite]: #121
+#130 := [monotonicity #124 #127]: #129
+#139 := [trans #130 #137]: #138
+#142 := [monotonicity #139]: #141
+#161 := [trans #142 #159]: #160
+#164 := [monotonicity #161]: #163
+#181 := [trans #164 #179]: #180
+#184 := [monotonicity #181]: #183
+#15 := (- uf_2)
+#14 := (< uf_2 0::real)
+#16 := (ite #14 #15 uf_2)
+#12 := (- uf_1)
+#11 := (< uf_1 0::real)
+#13 := (ite #11 #12 uf_1)
+#17 := (+ #13 #16)
+#9 := (- #6)
+#8 := (< #6 0::real)
+#10 := (ite #8 #9 #6)
+#18 := (<= #10 #17)
+#19 := (not #18)
+#118 := (iff #19 #115)
+#52 := (ite #14 #38 uf_2)
+#47 := (ite #11 #37 uf_1)
+#55 := (+ #47 #52)
+#42 := (ite #8 #39 #6)
+#58 := (<= #42 #55)
+#61 := (not #58)
+#116 := (iff #61 #115)
+#113 := (iff #58 #110)
+#100 := (+ #83 #95)
+#103 := (<= #71 #100)
+#111 := (iff #103 #110)
+#112 := [rewrite]: #111
+#104 := (iff #58 #103)
+#101 := (= #55 #100)
+#98 := (= #52 #95)
+#92 := (ite #89 #38 uf_2)
+#96 := (= #92 #95)
+#97 := [rewrite]: #96
+#93 := (= #52 #92)
+#90 := (iff #14 #89)
+#91 := [rewrite]: #90
+#94 := [monotonicity #91]: #93
+#99 := [trans #94 #97]: #98
+#86 := (= #47 #83)
+#77 := (not #76)
+#80 := (ite #77 #37 uf_1)
+#84 := (= #80 #83)
+#85 := [rewrite]: #84
+#81 := (= #47 #80)
+#78 := (iff #11 #77)
+#79 := [rewrite]: #78
+#82 := [monotonicity #79]: #81
+#87 := [trans #82 #85]: #86
+#102 := [monotonicity #87 #99]: #101
+#74 := (= #42 #71)
+#68 := (ite #65 #39 #6)
+#72 := (= #68 #71)
+#73 := [rewrite]: #72
+#69 := (= #42 #68)
+#66 := (iff #8 #65)
+#67 := [rewrite]: #66
+#70 := [monotonicity #67]: #69
+#75 := [trans #70 #73]: #74
+#105 := [monotonicity #75 #102]: #104
+#114 := [trans #105 #112]: #113
+#117 := [monotonicity #114]: #116
+#62 := (iff #19 #61)
+#59 := (iff #18 #58)
+#56 := (= #17 #55)
+#53 := (= #16 #52)
+#50 := (= #15 #38)
+#51 := [rewrite]: #50
+#54 := [monotonicity #51]: #53
+#48 := (= #13 #47)
+#45 := (= #12 #37)
+#46 := [rewrite]: #45
+#49 := [monotonicity #46]: #48
+#57 := [monotonicity #49 #54]: #56
+#43 := (= #10 #42)
+#40 := (= #9 #39)
+#41 := [rewrite]: #40
+#44 := [monotonicity #41]: #43
+#60 := [monotonicity #44 #57]: #59
+#63 := [monotonicity #60]: #62
+#119 := [trans #63 #117]: #118
+#35 := [asserted]: #19
+#120 := [mp #35 #119]: #115
+#185 := [mp #120 #184]: #182
+#248 := (or #177 #64 #233)
+#249 := [def-axiom]: #248
+#265 := [unit-resolution #249 #185]: #264
+#266 := [unit-resolution #265 #262]: #233
+#240 := (or #176 #76 #230)
+#241 := [def-axiom]: #240
+#267 := [unit-resolution #241 #266 #261]: #76
+#268 := [th-lemma #267 #263 #262]: false
+#270 := [lemma #268]: #269
+#272 := [unit-resolution #270 #263]: #64
+#273 := (or #65 #205)
+#246 := (or #177 #65 #205)
+#247 := [def-axiom]: #246
+#274 := [unit-resolution #247 #185]: #273
+#275 := [unit-resolution #274 #272]: #205
+#255 := (or #170 #194 #191)
+#250 := [hypothesis]: #169
+#251 := [hypothesis]: #205
+#252 := [hypothesis]: #166
+#210 := (or #170 #77 #191)
+#211 := [def-axiom]: #210
+#253 := [unit-resolution #211 #251 #252]: #77
+#212 := (or #170 #76 #194)
+#213 := [def-axiom]: #212
+#254 := [unit-resolution #213 #253 #251 #250]: false
+#256 := [lemma #254]: #255
+#276 := [unit-resolution #256 #275 #271]: #194
+#200 := (or #169 #89 #199)
+#201 := [def-axiom]: #200
+#277 := [unit-resolution #201 #276 #263]: #199
+#278 := [unit-resolution #211 #275 #271]: #77
+#279 := [th-lemma #278 #277]: false
+#280 := [lemma #279]: #89
+#281 := [hypothesis]: #77
+#282 := [unit-resolution #241 #281 #261]: #176
+#283 := [unit-resolution #265 #282]: #64
+#284 := [th-lemma #281 #283 #280]: false
+#285 := [lemma #284]: #76
+#222 := (not #172)
+#286 := [hypothesis]: #222
+#287 := [th-lemma #285 #286]: false
+#288 := [lemma #287]: #172
+#223 := (or #173 #88 #222)
+#224 := [def-axiom]: #223
+#289 := [unit-resolution #224 #288 #280]: #173
+#214 := (not #173)
+#238 := (or #176 #77 #214)
+#239 := [def-axiom]: #238
+#290 := [unit-resolution #239 #289 #285]: #176
+#291 := [unit-resolution #265 #290]: #64
+#292 := [unit-resolution #274 #291]: #205
+#293 := [unit-resolution #211 #292 #285]: #191
+#189 := (or #166 #188)
+#190 := [def-axiom]: #189
+#294 := [unit-resolution #190 #293]: #188
+[th-lemma #280 #294]: false
+unsat
+
+eBRZKSXriNPK3BNu3AWMmQ 3017
+#2 := false
+decl uf_1 :: (-> T1 T2)
+decl uf_3 :: T1
+#8 := uf_3
+#9 := (uf_1 uf_3)
+decl uf_2 :: (-> int int T1)
+#5 := 3::int
+#4 := 2::int
+#6 := (uf_2 2::int 3::int)
+#7 := (uf_1 #6)
+#10 := (= #7 #9)
+#225 := (= #6 uf_3)
+#13 := (:var 0 int)
+#12 := (:var 1 int)
+#14 := (uf_2 #12 #13)
+#549 := (pattern #14)
+#52 := 0::int
+#50 := -1::int
+#54 := (* -1::int #13)
+#55 := (+ #12 #54)
+#53 := (>= #55 0::int)
+#51 := (not #53)
+#36 := (= uf_3 #14)
+#61 := (iff #36 #51)
+#550 := (forall (vars (?x1 int) (?x2 int)) (:pat #549) #61)
+#66 := (forall (vars (?x1 int) (?x2 int)) #61)
+#553 := (iff #66 #550)
+#551 := (iff #61 #61)
+#552 := [refl]: #551
+#554 := [quant-intro #552]: #553
+#79 := (~ #66 #66)
+#77 := (~ #61 #61)
+#78 := [refl]: #77
+#80 := [nnf-pos #78]: #79
+#16 := (< #12 #13)
+#15 := (= #14 uf_3)
+#17 := (iff #15 #16)
+#18 := (forall (vars (?x1 int) (?x2 int)) #17)
+#69 := (iff #18 #66)
+#42 := (iff #16 #36)
+#47 := (forall (vars (?x1 int) (?x2 int)) #42)
+#67 := (iff #47 #66)
+#64 := (iff #42 #61)
+#58 := (iff #51 #36)
+#62 := (iff #58 #61)
+#63 := [rewrite]: #62
+#59 := (iff #42 #58)
+#56 := (iff #16 #51)
+#57 := [rewrite]: #56
+#60 := [monotonicity #57]: #59
+#65 := [trans #60 #63]: #64
+#68 := [quant-intro #65]: #67
+#48 := (iff #18 #47)
+#45 := (iff #17 #42)
+#39 := (iff #36 #16)
+#43 := (iff #39 #42)
+#44 := [rewrite]: #43
+#40 := (iff #17 #39)
+#37 := (iff #15 #36)
+#38 := [rewrite]: #37
+#41 := [monotonicity #38]: #40
+#46 := [trans #41 #44]: #45
+#49 := [quant-intro #46]: #48
+#70 := [trans #49 #68]: #69
+#35 := [asserted]: #18
+#71 := [mp #35 #70]: #66
+#74 := [mp~ #71 #80]: #66
+#555 := [mp #74 #554]: #550
+#529 := (not #550)
+#530 := (or #529 #225)
+#220 := (* -1::int 3::int)
+#221 := (+ 2::int #220)
+#222 := (>= #221 0::int)
+#213 := (not #222)
+#135 := (= uf_3 #6)
+#224 := (iff #135 #213)
+#525 := (or #529 #224)
+#169 := (iff #525 #530)
+#534 := (iff #530 #530)
+#174 := [rewrite]: #534
+#527 := (iff #224 #225)
+#1 := true
+#187 := (iff #225 true)
+#190 := (iff #187 #225)
+#526 := [rewrite]: #190
+#188 := (iff #224 #187)
+#183 := (iff #213 true)
+#198 := (not false)
+#199 := (iff #198 true)
+#540 := [rewrite]: #199
+#203 := (iff #213 #198)
+#548 := (iff #222 false)
+#544 := (>= -1::int 0::int)
+#547 := (iff #544 false)
+#542 := [rewrite]: #547
+#545 := (iff #222 #544)
+#211 := (= #221 -1::int)
+#223 := -3::int
+#541 := (+ 2::int -3::int)
+#330 := (= #541 -1::int)
+#537 := [rewrite]: #330
+#543 := (= #221 #541)
+#227 := (= #220 -3::int)
+#206 := [rewrite]: #227
+#200 := [monotonicity #206]: #543
+#212 := [trans #200 #537]: #211
+#546 := [monotonicity #212]: #545
+#538 := [trans #546 #542]: #548
+#539 := [monotonicity #538]: #203
+#524 := [trans #539 #540]: #183
+#153 := (iff #135 #225)
+#226 := [rewrite]: #153
+#189 := [monotonicity #226 #524]: #188
+#528 := [trans #189 #526]: #527
+#532 := [monotonicity #528]: #169
+#175 := [trans #532 #174]: #169
+#531 := [quant-inst]: #525
+#535 := [mp #531 #175]: #530
+#533 := [unit-resolution #535 #555]: #225
+#536 := [monotonicity #533]: #10
+#11 := (not #10)
+#34 := [asserted]: #11
+[unit-resolution #34 #536]: false
+unsat
+
+CjDkjvq1e9i+SJ3L9ESARg 1146
+#2 := false
+#9 := 1::int
+decl uf_1 :: int
+#5 := uf_1
+#10 := (< uf_1 1::int)
+#6 := 3::int
+#7 := (+ uf_1 3::int)
+#4 := 4::int
+#8 := (<= 4::int #7)
+#11 := (or #8 #10)
+#12 := (not #11)
+#66 := (iff #12 false)
+#29 := (+ 3::int uf_1)
+#32 := (<= 4::int #29)
+#38 := (or #10 #32)
+#43 := (not #38)
+#64 := (iff #43 false)
+#1 := true
+#59 := (not true)
+#62 := (iff #59 false)
+#63 := [rewrite]: #62
+#60 := (iff #43 #59)
+#57 := (iff #38 true)
+#48 := (>= uf_1 1::int)
+#46 := (not #48)
+#52 := (or #46 #48)
+#55 := (iff #52 true)
+#56 := [rewrite]: #55
+#53 := (iff #38 #52)
+#50 := (iff #32 #48)
+#51 := [rewrite]: #50
+#47 := (iff #10 #46)
+#49 := [rewrite]: #47
+#54 := [monotonicity #49 #51]: #53
+#58 := [trans #54 #56]: #57
+#61 := [monotonicity #58]: #60
+#65 := [trans #61 #63]: #64
+#44 := (iff #12 #43)
+#41 := (iff #11 #38)
+#35 := (or #32 #10)
+#39 := (iff #35 #38)
+#40 := [rewrite]: #39
+#36 := (iff #11 #35)
+#33 := (iff #8 #32)
+#30 := (= #7 #29)
+#31 := [rewrite]: #30
+#34 := [monotonicity #31]: #33
+#37 := [monotonicity #34]: #36
+#42 := [trans #37 #40]: #41
+#45 := [monotonicity #42]: #44
+#67 := [trans #45 #65]: #66
+#28 := [asserted]: #12
+[mp #28 #67]: false
+unsat
+
+nonk4MmmwlBqL8YtiJY/Qw 1339
+#2 := false
+#11 := 0::int
+decl uf_2 :: int
+#7 := uf_2
+#42 := -1::int
+#45 := (* -1::int uf_2)
+decl uf_1 :: int
+#5 := uf_1
+#46 := (+ uf_1 #45)
+#63 := (>= #46 0::int)
+#83 := (iff #63 false)
+#44 := -4::int
+#79 := (>= -4::int 0::int)
+#81 := (iff #79 false)
+#82 := [rewrite]: #81
+#77 := (iff #63 #79)
+#47 := (= #46 -4::int)
+#8 := 4::int
+#9 := (+ uf_1 4::int)
+#10 := (= uf_2 #9)
+#49 := (iff #10 #47)
+#32 := (+ 4::int uf_1)
+#39 := (= uf_2 #32)
+#43 := (iff #39 #47)
+#48 := [rewrite]: #43
+#40 := (iff #10 #39)
+#37 := (= #9 #32)
+#38 := [rewrite]: #37
+#41 := [monotonicity #38]: #40
+#50 := [trans #41 #48]: #49
+#31 := [asserted]: #10
+#51 := [mp #31 #50]: #47
+#80 := [monotonicity #51]: #77
+#84 := [trans #80 #82]: #83
+#12 := (- uf_2 uf_1)
+#13 := (< 0::int #12)
+#14 := (not #13)
+#74 := (iff #14 #63)
+#53 := (* -1::int uf_1)
+#54 := (+ #53 uf_2)
+#57 := (< 0::int #54)
+#60 := (not #57)
+#72 := (iff #60 #63)
+#64 := (not #63)
+#67 := (not #64)
+#70 := (iff #67 #63)
+#71 := [rewrite]: #70
+#68 := (iff #60 #67)
+#65 := (iff #57 #64)
+#66 := [rewrite]: #65
+#69 := [monotonicity #66]: #68
+#73 := [trans #69 #71]: #72
+#61 := (iff #14 #60)
+#58 := (iff #13 #57)
+#55 := (= #12 #54)
+#56 := [rewrite]: #55
+#59 := [monotonicity #56]: #58
+#62 := [monotonicity #59]: #61
+#75 := [trans #62 #73]: #74
+#52 := [asserted]: #14
+#76 := [mp #52 #75]: #63
+[mp #76 #84]: false
+unsat
+
+dCX9jxibjKl6gmr8okzk0w 727
+#2 := false
+#6 := 5::int
+#4 := 2::int
+#5 := (+ 2::int 2::int)
+#7 := (= #5 5::int)
+#8 := (not #7)
+#9 := (not #8)
+#48 := (iff #9 false)
+#1 := true
+#43 := (not true)
+#46 := (iff #43 false)
+#47 := [rewrite]: #46
+#44 := (iff #9 #43)
+#41 := (iff #8 true)
+#36 := (not false)
+#39 := (iff #36 true)
+#40 := [rewrite]: #39
+#37 := (iff #8 #36)
+#34 := (iff #7 false)
+#26 := 4::int
+#29 := (= 4::int 5::int)
+#32 := (iff #29 false)
+#33 := [rewrite]: #32
+#30 := (iff #7 #29)
+#27 := (= #5 4::int)
+#28 := [rewrite]: #27
+#31 := [monotonicity #28]: #30
+#35 := [trans #31 #33]: #34
+#38 := [monotonicity #35]: #37
+#42 := [trans #38 #40]: #41
+#45 := [monotonicity #42]: #44
+#49 := [trans #45 #47]: #48
+#25 := [asserted]: #9
+[mp #25 #49]: false
+unsat
+
+/kLzs8f/jQjEM38PdppYPA 912
+#2 := false
+#11 := 4::real
+decl uf_2 :: real
+#8 := uf_2
+#7 := 7::real
+#9 := (* 7::real uf_2)
+decl uf_1 :: real
+#5 := uf_1
+#4 := 3::real
+#6 := (* 3::real uf_1)
+#10 := (+ #6 #9)
+#41 := (>= #10 4::real)
+#39 := (not #41)
+#12 := (< #10 4::real)
+#40 := (iff #12 #39)
+#37 := [rewrite]: #40
+#34 := [asserted]: #12
+#38 := [mp #34 #37]: #39
+#13 := 2::real
+#14 := (* 2::real uf_1)
+#43 := (<= #14 3::real)
+#44 := (not #43)
+#15 := (< 3::real #14)
+#45 := (iff #15 #44)
+#46 := [rewrite]: #45
+#35 := [asserted]: #15
+#47 := [mp #35 #46]: #44
+#16 := 0::real
+#51 := (>= uf_2 0::real)
+#17 := (< uf_2 0::real)
+#18 := (not #17)
+#58 := (iff #18 #51)
+#49 := (not #51)
+#53 := (not #49)
+#56 := (iff #53 #51)
+#57 := [rewrite]: #56
+#54 := (iff #18 #53)
+#50 := (iff #17 #49)
+#52 := [rewrite]: #50
+#55 := [monotonicity #52]: #54
+#59 := [trans #55 #57]: #58
+#36 := [asserted]: #18
+#60 := [mp #36 #59]: #51
+[th-lemma #60 #47 #38]: false
+unsat
+
+iT8vKYi503k30rQLczD7yw 1245
+#2 := false
+#16 := (not false)
+decl uf_2 :: int
+#8 := uf_2
+#4 := 0::int
+#12 := (<= 0::int uf_2)
+#13 := (not #12)
+#14 := (or #13 #12)
+#6 := 1::int
+#7 := (- 1::int)
+#9 := (* #7 uf_2)
+decl uf_1 :: int
+#5 := uf_1
+#10 := (+ uf_1 #9)
+#11 := (<= 0::int #10)
+#15 := (or #11 #14)
+#17 := (iff #15 #16)
+#18 := (not #17)
+#70 := (iff #18 false)
+#1 := true
+#65 := (not true)
+#68 := (iff #65 false)
+#69 := [rewrite]: #68
+#66 := (iff #18 #65)
+#63 := (iff #17 true)
+#58 := (iff true true)
+#61 := (iff #58 true)
+#62 := [rewrite]: #61
+#59 := (iff #17 #58)
+#56 := (iff #16 true)
+#57 := [rewrite]: #56
+#54 := (iff #15 true)
+#35 := -1::int
+#38 := (* -1::int uf_2)
+#41 := (+ uf_1 #38)
+#44 := (<= 0::int #41)
+#49 := (or #44 true)
+#52 := (iff #49 true)
+#53 := [rewrite]: #52
+#50 := (iff #15 #49)
+#47 := (iff #14 true)
+#48 := [rewrite]: #47
+#45 := (iff #11 #44)
+#42 := (= #10 #41)
+#39 := (= #9 #38)
+#36 := (= #7 -1::int)
+#37 := [rewrite]: #36
+#40 := [monotonicity #37]: #39
+#43 := [monotonicity #40]: #42
+#46 := [monotonicity #43]: #45
+#51 := [monotonicity #46 #48]: #50
+#55 := [trans #51 #53]: #54
+#60 := [monotonicity #55 #57]: #59
+#64 := [trans #60 #62]: #63
+#67 := [monotonicity #64]: #66
+#71 := [trans #67 #69]: #70
+#34 := [asserted]: #18
+[mp #34 #71]: false
+unsat
+
+6R4JcV7tL9QRH7WWPAKM5g 5413
+#2 := false
+decl uf_4 :: T1
+#13 := uf_4
+decl uf_1 :: (-> int int T1)
+#5 := 3::int
+decl uf_2 :: int
+#4 := uf_2
+#6 := (uf_1 uf_2 3::int)
+#559 := (= #6 uf_4)
+decl uf_3 :: (-> int int T1)
+#7 := (uf_3 3::int uf_2)
+#254 := (= #7 uf_4)
+#524 := (iff #254 #559)
+#529 := (iff #559 #254)
+#39 := (= #6 #7)
+#8 := (distinct #6 #7)
+#9 := (not #8)
+#48 := (iff #9 #39)
+#40 := (not #39)
+#43 := (not #40)
+#46 := (iff #43 #39)
+#47 := [rewrite]: #46
+#44 := (iff #9 #43)
+#41 := (iff #8 #40)
+#42 := [rewrite]: #41
+#45 := [monotonicity #42]: #44
+#49 := [trans #45 #47]: #48
+#38 := [asserted]: #9
+#52 := [mp #38 #49]: #39
+#523 := [monotonicity #52]: #529
+#530 := [symm #523]: #524
+#547 := (not #559)
+#570 := (not #254)
+#531 := (iff #570 #547)
+#525 := [monotonicity #530]: #531
+#540 := [hypothesis]: #570
+#532 := [mp #540 #525]: #547
+#256 := (>= uf_2 3::int)
+#579 := (not #256)
+#541 := (or #254 #579)
+#258 := (iff #254 #256)
+#11 := (:var 0 int)
+#10 := (:var 1 int)
+#12 := (uf_3 #10 #11)
+#581 := (pattern #12)
+#57 := 0::int
+#54 := -1::int
+#55 := (* -1::int #11)
+#56 := (+ #10 #55)
+#58 := (<= #56 0::int)
+#14 := (= #12 uf_4)
+#61 := (iff #14 #58)
+#582 := (forall (vars (?x1 int) (?x2 int)) (:pat #581) #61)
+#64 := (forall (vars (?x1 int) (?x2 int)) #61)
+#585 := (iff #64 #582)
+#583 := (iff #61 #61)
+#584 := [refl]: #583
+#586 := [quant-intro #584]: #585
+#108 := (~ #64 #64)
+#106 := (~ #61 #61)
+#107 := [refl]: #106
+#109 := [nnf-pos #107]: #108
+#15 := (<= #10 #11)
+#16 := (iff #14 #15)
+#17 := (forall (vars (?x1 int) (?x2 int)) #16)
+#65 := (iff #17 #64)
+#62 := (iff #16 #61)
+#59 := (iff #15 #58)
+#60 := [rewrite]: #59
+#63 := [monotonicity #60]: #62
+#66 := [quant-intro #63]: #65
+#50 := [asserted]: #17
+#67 := [mp #50 #66]: #64
+#101 := [mp~ #67 #109]: #64
+#587 := [mp #101 #586]: #582
+#238 := (not #582)
+#573 := (or #238 #258)
+#167 := (* -1::int uf_2)
+#252 := (+ 3::int #167)
+#253 := (<= #252 0::int)
+#245 := (iff #254 #253)
+#575 := (or #238 #245)
+#362 := (iff #575 #573)
+#243 := (iff #573 #573)
+#244 := [rewrite]: #243
+#255 := (iff #245 #258)
+#257 := (iff #253 #256)
+#185 := [rewrite]: #257
+#259 := [monotonicity #185]: #255
+#569 := [monotonicity #259]: #362
+#576 := [trans #569 #244]: #362
+#232 := [quant-inst]: #575
+#577 := [mp #232 #576]: #573
+#535 := [unit-resolution #577 #587]: #258
+#578 := (not #258)
+#574 := (or #578 #254 #579)
+#580 := [def-axiom]: #574
+#382 := [unit-resolution #580 #535]: #541
+#383 := [unit-resolution #382 #540]: #579
+#526 := (or #559 #256)
+#273 := (iff #559 #579)
+#18 := (uf_1 #10 #11)
+#588 := (pattern #18)
+#82 := (>= #56 0::int)
+#81 := (not #82)
+#53 := (= uf_4 #18)
+#88 := (iff #53 #81)
+#589 := (forall (vars (?x3 int) (?x4 int)) (:pat #588) #88)
+#93 := (forall (vars (?x3 int) (?x4 int)) #88)
+#592 := (iff #93 #589)
+#590 := (iff #88 #88)
+#591 := [refl]: #590
+#593 := [quant-intro #591]: #592
+#102 := (~ #93 #93)
+#99 := (~ #88 #88)
+#110 := [refl]: #99
+#103 := [nnf-pos #110]: #102
+#20 := (< #10 #11)
+#19 := (= #18 uf_4)
+#21 := (iff #19 #20)
+#22 := (forall (vars (?x3 int) (?x4 int)) #21)
+#96 := (iff #22 #93)
+#73 := (iff #20 #53)
+#78 := (forall (vars (?x3 int) (?x4 int)) #73)
+#94 := (iff #78 #93)
+#91 := (iff #73 #88)
+#85 := (iff #81 #53)
+#89 := (iff #85 #88)
+#90 := [rewrite]: #89
+#86 := (iff #73 #85)
+#83 := (iff #20 #81)
+#84 := [rewrite]: #83
+#87 := [monotonicity #84]: #86
+#92 := [trans #87 #90]: #91
+#95 := [quant-intro #92]: #94
+#79 := (iff #22 #78)
+#76 := (iff #21 #73)
+#70 := (iff #53 #20)
+#74 := (iff #70 #73)
+#75 := [rewrite]: #74
+#71 := (iff #21 #70)
+#68 := (iff #19 #53)
+#69 := [rewrite]: #68
+#72 := [monotonicity #69]: #71
+#77 := [trans #72 #75]: #76
+#80 := [quant-intro #77]: #79
+#97 := [trans #80 #95]: #96
+#51 := [asserted]: #22
+#98 := [mp #51 #97]: #93
+#111 := [mp~ #98 #103]: #93
+#594 := [mp #111 #593]: #589
+#552 := (not #589)
+#549 := (or #552 #273)
+#219 := (* -1::int 3::int)
+#220 := (+ uf_2 #219)
+#221 := (>= #220 0::int)
+#222 := (not #221)
+#556 := (= uf_4 #6)
+#558 := (iff #556 #222)
+#553 := (or #552 #558)
+#264 := (iff #553 #549)
+#266 := (iff #549 #549)
+#544 := [rewrite]: #266
+#274 := (iff #558 #273)
+#550 := (iff #222 #579)
+#280 := (iff #221 #256)
+#562 := -3::int
+#206 := (+ -3::int uf_2)
+#554 := (>= #206 0::int)
+#278 := (iff #554 #256)
+#279 := [rewrite]: #278
+#555 := (iff #221 #554)
+#565 := (= #220 #206)
+#201 := (+ uf_2 -3::int)
+#207 := (= #201 #206)
+#567 := [rewrite]: #207
+#564 := (= #220 #201)
+#557 := (= #219 -3::int)
+#563 := [rewrite]: #557
+#566 := [monotonicity #563]: #564
+#568 := [trans #566 #567]: #565
+#277 := [monotonicity #568]: #555
+#173 := [trans #277 #279]: #280
+#551 := [monotonicity #173]: #550
+#560 := (iff #556 #559)
+#561 := [rewrite]: #560
+#548 := [monotonicity #561 #551]: #274
+#265 := [monotonicity #548]: #264
+#545 := [trans #265 #544]: #264
+#263 := [quant-inst]: #553
+#260 := [mp #263 #545]: #549
+#384 := [unit-resolution #260 #594]: #273
+#542 := (not #273)
+#546 := (or #542 #559 #256)
+#543 := [def-axiom]: #546
+#527 := [unit-resolution #543 #384]: #526
+#528 := [unit-resolution #527 #383]: #559
+#361 := [unit-resolution #528 #532]: false
+#363 := [lemma #361]: #254
+#522 := [mp #363 #530]: #559
+#364 := (or #570 #256)
+#230 := (or #578 #570 #256)
+#235 := [def-axiom]: #230
+#517 := [unit-resolution #235 #535]: #364
+#518 := [unit-resolution #517 #363]: #256
+#520 := (or #547 #579)
+#536 := (or #542 #547 #579)
+#537 := [def-axiom]: #536
+#521 := [unit-resolution #537 #384]: #520
+#519 := [unit-resolution #521 #518]: #547
+[unit-resolution #519 #522]: false
+unsat
+
+eOXl5Nf4A2Sq4Q+Wh5XNfA 2026
+#2 := false
+decl uf_1 :: int
+#5 := uf_1
+#7 := 2::int
+#29 := (* 2::int uf_1)
+#4 := 0::int
+#54 := (= 0::int #29)
+#55 := (not #54)
+#61 := (= #29 0::int)
+#104 := (not #61)
+#110 := (iff #104 #55)
+#108 := (iff #61 #54)
+#109 := [commutativity]: #108
+#111 := [monotonicity #109]: #110
+#62 := (<= #29 0::int)
+#100 := (not #62)
+#30 := (<= uf_1 0::int)
+#31 := (not #30)
+#6 := (< 0::int uf_1)
+#32 := (iff #6 #31)
+#33 := [rewrite]: #32
+#27 := [asserted]: #6
+#34 := [mp #27 #33]: #31
+#101 := (or #100 #30)
+#102 := [th-lemma]: #101
+#103 := [unit-resolution #102 #34]: #100
+#105 := (or #104 #62)
+#106 := [th-lemma]: #105
+#107 := [unit-resolution #106 #103]: #104
+#112 := [mp #107 #111]: #55
+#56 := (= uf_1 #29)
+#57 := (not #56)
+#53 := (= 0::int uf_1)
+#50 := (not #53)
+#58 := (and #50 #55 #57)
+#69 := (not #58)
+#42 := (distinct 0::int uf_1 #29)
+#47 := (not #42)
+#9 := (- uf_1 uf_1)
+#8 := (* uf_1 2::int)
+#10 := (distinct uf_1 #8 #9)
+#11 := (not #10)
+#48 := (iff #11 #47)
+#45 := (iff #10 #42)
+#39 := (distinct uf_1 #29 0::int)
+#43 := (iff #39 #42)
+#44 := [rewrite]: #43
+#40 := (iff #10 #39)
+#37 := (= #9 0::int)
+#38 := [rewrite]: #37
+#35 := (= #8 #29)
+#36 := [rewrite]: #35
+#41 := [monotonicity #36 #38]: #40
+#46 := [trans #41 #44]: #45
+#49 := [monotonicity #46]: #48
+#28 := [asserted]: #11
+#52 := [mp #28 #49]: #47
+#80 := (or #42 #69)
+#81 := [def-axiom]: #80
+#82 := [unit-resolution #81 #52]: #69
+#59 := (= uf_1 0::int)
+#83 := (not #59)
+#89 := (iff #83 #50)
+#87 := (iff #59 #53)
+#88 := [commutativity]: #87
+#90 := [monotonicity #88]: #89
+#84 := (or #83 #30)
+#85 := [th-lemma]: #84
+#86 := [unit-resolution #85 #34]: #83
+#91 := [mp #86 #90]: #50
+#64 := -1::int
+#65 := (* -1::int #29)
+#66 := (+ uf_1 #65)
+#68 := (>= #66 0::int)
+#92 := (not #68)
+#93 := (or #92 #30)
+#94 := [th-lemma]: #93
+#95 := [unit-resolution #94 #34]: #92
+#96 := (or #57 #68)
+#97 := [th-lemma]: #96
+#98 := [unit-resolution #97 #95]: #57
+#76 := (or #58 #53 #54 #56)
+#77 := [def-axiom]: #76
+#99 := [unit-resolution #77 #98 #91 #82]: #54
+[unit-resolution #99 #112]: false
+unsat
+
+TwJgkTydtls9Q94iw4a3jw 17377
+#2 := false
+#169 := 0::int
+decl uf_2 :: int
+#5 := uf_2
+#166 := -1::int
+#202 := (* -1::int uf_2)
+decl uf_1 :: int
+#4 := uf_1
+#203 := (+ uf_1 #202)
+#218 := (>= #203 0::int)
+decl uf_3 :: int
+#7 := uf_3
+#167 := (* -1::int uf_3)
+#168 := (+ uf_1 #167)
+#178 := (>= #168 0::int)
+#217 := (not #218)
+#204 := (<= #203 0::int)
+#205 := (not #204)
+#692 := [hypothesis]: #205
+#177 := (not #178)
+#693 := (or #177 #204)
+#170 := (<= #168 0::int)
+#191 := (+ uf_2 #167)
+#237 := (<= #191 0::int)
+#238 := (not #237)
+#171 := (not #170)
+#685 := [hypothesis]: #171
+#190 := (>= #191 0::int)
+#455 := (or #170 #190)
+#189 := (not #190)
+#197 := (and #171 #189)
+#354 := (not #197)
+#464 := (iff #354 #455)
+#456 := (not #455)
+#459 := (not #456)
+#462 := (iff #459 #455)
+#463 := [rewrite]: #462
+#460 := (iff #354 #459)
+#457 := (iff #197 #456)
+#458 := [rewrite]: #457
+#461 := [monotonicity #458]: #460
+#465 := [trans #461 #463]: #464
+#287 := (and #189 #217)
+#10 := (= uf_2 uf_3)
+#279 := (and #10 #217)
+#273 := (and #177 #238)
+#15 := (= uf_1 uf_3)
+#268 := (and #15 #238)
+#17 := (= uf_1 uf_2)
+#260 := (and #17 #189)
+#252 := (and #205 #238)
+#244 := (and #17 #238)
+#232 := (and #171 #217)
+#224 := (and #15 #217)
+#214 := (and #10 #205)
+#211 := (and #177 #205)
+#208 := (and #15 #205)
+#184 := (and #17 #177)
+#174 := (and #10 #171)
+#115 := (and #10 #17)
+#337 := (or #115 #174 #184 #197 #208 #211 #214 #224 #232 #244 #252 #260 #268 #273 #279 #287)
+#342 := (not #337)
+#21 := (= uf_2 uf_1)
+#27 := (= uf_3 uf_2)
+#34 := (and #27 #21)
+#23 := (< uf_3 uf_1)
+#33 := (and #10 #23)
+#35 := (or #33 #34)
+#12 := (< uf_1 uf_3)
+#32 := (and #21 #12)
+#36 := (or #32 #35)
+#8 := (< uf_2 uf_3)
+#31 := (and #8 #23)
+#37 := (or #31 #36)
+#25 := (= uf_3 uf_1)
+#19 := (< uf_2 uf_1)
+#30 := (and #19 #25)
+#38 := (or #30 #37)
+#29 := (and #19 #12)
+#39 := (or #29 #38)
+#28 := (and #27 #19)
+#40 := (or #28 #39)
+#6 := (< uf_1 uf_2)
+#26 := (and #25 #6)
+#41 := (or #26 #40)
+#24 := (and #23 #6)
+#42 := (or #24 #41)
+#13 := (< uf_3 uf_2)
+#22 := (and #13 #21)
+#43 := (or #22 #42)
+#20 := (and #13 #19)
+#44 := (or #20 #43)
+#18 := (and #17 #8)
+#45 := (or #18 #44)
+#16 := (and #15 #13)
+#46 := (or #16 #45)
+#14 := (and #12 #13)
+#47 := (or #14 #46)
+#11 := (and #6 #10)
+#48 := (or #11 #47)
+#9 := (and #6 #8)
+#49 := (or #9 #48)
+#50 := (not #49)
+#345 := (iff #50 #342)
+#118 := (or #33 #115)
+#110 := (and #12 #17)
+#121 := (or #110 #118)
+#124 := (or #31 #121)
+#102 := (and #15 #19)
+#127 := (or #102 #124)
+#96 := (and #12 #19)
+#130 := (or #96 #127)
+#93 := (and #10 #19)
+#133 := (or #93 #130)
+#86 := (and #6 #15)
+#136 := (or #86 #133)
+#78 := (and #6 #23)
+#139 := (or #78 #136)
+#75 := (and #13 #17)
+#142 := (or #75 #139)
+#145 := (or #20 #142)
+#70 := (and #8 #17)
+#148 := (or #70 #145)
+#67 := (and #13 #15)
+#151 := (or #67 #148)
+#154 := (or #14 #151)
+#157 := (or #11 #154)
+#160 := (or #9 #157)
+#163 := (not #160)
+#343 := (iff #163 #342)
+#340 := (iff #160 #337)
+#292 := (or #174 #115)
+#295 := (or #184 #292)
+#298 := (or #197 #295)
+#301 := (or #208 #298)
+#304 := (or #211 #301)
+#307 := (or #214 #304)
+#310 := (or #224 #307)
+#313 := (or #232 #310)
+#316 := (or #244 #313)
+#319 := (or #252 #316)
+#322 := (or #260 #319)
+#325 := (or #268 #322)
+#328 := (or #273 #325)
+#331 := (or #279 #328)
+#334 := (or #287 #331)
+#338 := (iff #334 #337)
+#339 := [rewrite]: #338
+#335 := (iff #160 #334)
+#332 := (iff #157 #331)
+#329 := (iff #154 #328)
+#326 := (iff #151 #325)
+#323 := (iff #148 #322)
+#320 := (iff #145 #319)
+#317 := (iff #142 #316)
+#314 := (iff #139 #313)
+#311 := (iff #136 #310)
+#308 := (iff #133 #307)
+#305 := (iff #130 #304)
+#302 := (iff #127 #301)
+#299 := (iff #124 #298)
+#296 := (iff #121 #295)
+#293 := (iff #118 #292)
+#175 := (iff #33 #174)
+#172 := (iff #23 #171)
+#173 := [rewrite]: #172
+#176 := [monotonicity #173]: #175
+#294 := [monotonicity #176]: #293
+#187 := (iff #110 #184)
+#181 := (and #177 #17)
+#185 := (iff #181 #184)
+#186 := [rewrite]: #185
+#182 := (iff #110 #181)
+#179 := (iff #12 #177)
+#180 := [rewrite]: #179
+#183 := [monotonicity #180]: #182
+#188 := [trans #183 #186]: #187
+#297 := [monotonicity #188 #294]: #296
+#200 := (iff #31 #197)
+#194 := (and #189 #171)
+#198 := (iff #194 #197)
+#199 := [rewrite]: #198
+#195 := (iff #31 #194)
+#192 := (iff #8 #189)
+#193 := [rewrite]: #192
+#196 := [monotonicity #193 #173]: #195
+#201 := [trans #196 #199]: #200
+#300 := [monotonicity #201 #297]: #299
+#209 := (iff #102 #208)
+#206 := (iff #19 #205)
+#207 := [rewrite]: #206
+#210 := [monotonicity #207]: #209
+#303 := [monotonicity #210 #300]: #302
+#212 := (iff #96 #211)
+#213 := [monotonicity #180 #207]: #212
+#306 := [monotonicity #213 #303]: #305
+#215 := (iff #93 #214)
+#216 := [monotonicity #207]: #215
+#309 := [monotonicity #216 #306]: #308
+#227 := (iff #86 #224)
+#221 := (and #217 #15)
+#225 := (iff #221 #224)
+#226 := [rewrite]: #225
+#222 := (iff #86 #221)
+#219 := (iff #6 #217)
+#220 := [rewrite]: #219
+#223 := [monotonicity #220]: #222
+#228 := [trans #223 #226]: #227
+#312 := [monotonicity #228 #309]: #311
+#235 := (iff #78 #232)
+#229 := (and #217 #171)
+#233 := (iff #229 #232)
+#234 := [rewrite]: #233
+#230 := (iff #78 #229)
+#231 := [monotonicity #220 #173]: #230
+#236 := [trans #231 #234]: #235
+#315 := [monotonicity #236 #312]: #314
+#247 := (iff #75 #244)
+#241 := (and #238 #17)
+#245 := (iff #241 #244)
+#246 := [rewrite]: #245
+#242 := (iff #75 #241)
+#239 := (iff #13 #238)
+#240 := [rewrite]: #239
+#243 := [monotonicity #240]: #242
+#248 := [trans #243 #246]: #247
+#318 := [monotonicity #248 #315]: #317
+#255 := (iff #20 #252)
+#249 := (and #238 #205)
+#253 := (iff #249 #252)
+#254 := [rewrite]: #253
+#250 := (iff #20 #249)
+#251 := [monotonicity #240 #207]: #250
+#256 := [trans #251 #254]: #255
+#321 := [monotonicity #256 #318]: #320
+#263 := (iff #70 #260)
+#257 := (and #189 #17)
+#261 := (iff #257 #260)
+#262 := [rewrite]: #261
+#258 := (iff #70 #257)
+#259 := [monotonicity #193]: #258
+#264 := [trans #259 #262]: #263
+#324 := [monotonicity #264 #321]: #323
+#271 := (iff #67 #268)
+#265 := (and #238 #15)
+#269 := (iff #265 #268)
+#270 := [rewrite]: #269
+#266 := (iff #67 #265)
+#267 := [monotonicity #240]: #266
+#272 := [trans #267 #270]: #271
+#327 := [monotonicity #272 #324]: #326
+#274 := (iff #14 #273)
+#275 := [monotonicity #180 #240]: #274
+#330 := [monotonicity #275 #327]: #329
+#282 := (iff #11 #279)
+#276 := (and #217 #10)
+#280 := (iff #276 #279)
+#281 := [rewrite]: #280
+#277 := (iff #11 #276)
+#278 := [monotonicity #220]: #277
+#283 := [trans #278 #281]: #282
+#333 := [monotonicity #283 #330]: #332
+#290 := (iff #9 #287)
+#284 := (and #217 #189)
+#288 := (iff #284 #287)
+#289 := [rewrite]: #288
+#285 := (iff #9 #284)
+#286 := [monotonicity #220 #193]: #285
+#291 := [trans #286 #289]: #290
+#336 := [monotonicity #291 #333]: #335
+#341 := [trans #336 #339]: #340
+#344 := [monotonicity #341]: #343
+#164 := (iff #50 #163)
+#161 := (iff #49 #160)
+#158 := (iff #48 #157)
+#155 := (iff #47 #154)
+#152 := (iff #46 #151)
+#149 := (iff #45 #148)
+#146 := (iff #44 #145)
+#143 := (iff #43 #142)
+#140 := (iff #42 #139)
+#137 := (iff #41 #136)
+#134 := (iff #40 #133)
+#131 := (iff #39 #130)
+#128 := (iff #38 #127)
+#125 := (iff #37 #124)
+#122 := (iff #36 #121)
+#119 := (iff #35 #118)
+#116 := (iff #34 #115)
+#73 := (iff #21 #17)
+#74 := [rewrite]: #73
+#91 := (iff #27 #10)
+#92 := [rewrite]: #91
+#117 := [monotonicity #92 #74]: #116
+#120 := [monotonicity #117]: #119
+#113 := (iff #32 #110)
+#107 := (and #17 #12)
+#111 := (iff #107 #110)
+#112 := [rewrite]: #111
+#108 := (iff #32 #107)
+#109 := [monotonicity #74]: #108
+#114 := [trans #109 #112]: #113
+#123 := [monotonicity #114 #120]: #122
+#126 := [monotonicity #123]: #125
+#105 := (iff #30 #102)
+#99 := (and #19 #15)
+#103 := (iff #99 #102)
+#104 := [rewrite]: #103
+#100 := (iff #30 #99)
+#81 := (iff #25 #15)
+#82 := [rewrite]: #81
+#101 := [monotonicity #82]: #100
+#106 := [trans #101 #104]: #105
+#129 := [monotonicity #106 #126]: #128
+#97 := (iff #29 #96)
+#98 := [rewrite]: #97
+#132 := [monotonicity #98 #129]: #131
+#94 := (iff #28 #93)
+#95 := [monotonicity #92]: #94
+#135 := [monotonicity #95 #132]: #134
+#89 := (iff #26 #86)
+#83 := (and #15 #6)
+#87 := (iff #83 #86)
+#88 := [rewrite]: #87
+#84 := (iff #26 #83)
+#85 := [monotonicity #82]: #84
+#90 := [trans #85 #88]: #89
+#138 := [monotonicity #90 #135]: #137
+#79 := (iff #24 #78)
+#80 := [rewrite]: #79
+#141 := [monotonicity #80 #138]: #140
+#76 := (iff #22 #75)
+#77 := [monotonicity #74]: #76
+#144 := [monotonicity #77 #141]: #143
+#147 := [monotonicity #144]: #146
+#71 := (iff #18 #70)
+#72 := [rewrite]: #71
+#150 := [monotonicity #72 #147]: #149
+#68 := (iff #16 #67)
+#69 := [rewrite]: #68
+#153 := [monotonicity #69 #150]: #152
+#156 := [monotonicity #153]: #155
+#159 := [monotonicity #156]: #158
+#162 := [monotonicity #159]: #161
+#165 := [monotonicity #162]: #164
+#346 := [trans #165 #344]: #345
+#66 := [asserted]: #50
+#347 := [mp #66 #346]: #342
+#355 := [not-or-elim #347]: #354
+#466 := [mp #355 #465]: #455
+#686 := [unit-resolution #466 #685]: #190
+#427 := (or #170 #189 #238)
+#350 := (not #174)
+#430 := (iff #350 #427)
+#382 := (or #189 #238)
+#414 := (or #170 #382)
+#428 := (iff #414 #427)
+#429 := [rewrite]: #428
+#425 := (iff #350 #414)
+#415 := (not #414)
+#420 := (not #415)
+#423 := (iff #420 #414)
+#424 := [rewrite]: #423
+#421 := (iff #350 #420)
+#418 := (iff #174 #415)
+#380 := (not #382)
+#411 := (and #380 #171)
+#416 := (iff #411 #415)
+#417 := [rewrite]: #416
+#412 := (iff #174 #411)
+#383 := (iff #10 #380)
+#384 := [rewrite]: #383
+#413 := [monotonicity #384]: #412
+#419 := [trans #413 #417]: #418
+#422 := [monotonicity #419]: #421
+#426 := [trans #422 #424]: #425
+#431 := [trans #426 #429]: #430
+#351 := [not-or-elim #347]: #350
+#432 := [mp #351 #431]: #427
+#687 := [unit-resolution #432 #686 #685]: #238
+#549 := (or #170 #218)
+#364 := (not #232)
+#558 := (iff #364 #549)
+#550 := (not #549)
+#553 := (not #550)
+#556 := (iff #553 #549)
+#557 := [rewrite]: #556
+#554 := (iff #364 #553)
+#551 := (iff #232 #550)
+#552 := [rewrite]: #551
+#555 := [monotonicity #552]: #554
+#559 := [trans #555 #557]: #558
+#365 := [not-or-elim #347]: #364
+#560 := [mp #365 #559]: #549
+#688 := [unit-resolution #560 #685]: #218
+#577 := (or #205 #217 #237)
+#366 := (not #244)
+#580 := (iff #366 #577)
+#385 := (or #205 #217)
+#564 := (or #237 #385)
+#578 := (iff #564 #577)
+#579 := [rewrite]: #578
+#575 := (iff #366 #564)
+#565 := (not #564)
+#570 := (not #565)
+#573 := (iff #570 #564)
+#574 := [rewrite]: #573
+#571 := (iff #366 #570)
+#568 := (iff #244 #565)
+#386 := (not #385)
+#561 := (and #386 #238)
+#566 := (iff #561 #565)
+#567 := [rewrite]: #566
+#562 := (iff #244 #561)
+#387 := (iff #17 #386)
+#388 := [rewrite]: #387
+#563 := [monotonicity #388]: #562
+#569 := [trans #563 #567]: #568
+#572 := [monotonicity #569]: #571
+#576 := [trans #572 #574]: #575
+#581 := [trans #576 #579]: #580
+#367 := [not-or-elim #347]: #366
+#582 := [mp #367 #581]: #577
+#689 := [unit-resolution #582 #688 #687]: #205
+#583 := (or #204 #237)
+#368 := (not #252)
+#592 := (iff #368 #583)
+#584 := (not #583)
+#587 := (not #584)
+#590 := (iff #587 #583)
+#591 := [rewrite]: #590
+#588 := (iff #368 #587)
+#585 := (iff #252 #584)
+#586 := [rewrite]: #585
+#589 := [monotonicity #586]: #588
+#593 := [trans #589 #591]: #592
+#369 := [not-or-elim #347]: #368
+#594 := [mp #369 #593]: #583
+#690 := [unit-resolution #594 #689 #687]: false
+#691 := [lemma #690]: #170
+#487 := (or #171 #177 #204)
+#356 := (not #208)
+#490 := (iff #356 #487)
+#467 := (or #171 #177)
+#474 := (or #204 #467)
+#488 := (iff #474 #487)
+#489 := [rewrite]: #488
+#485 := (iff #356 #474)
+#475 := (not #474)
+#480 := (not #475)
+#483 := (iff #480 #474)
+#484 := [rewrite]: #483
+#481 := (iff #356 #480)
+#478 := (iff #208 #475)
+#468 := (not #467)
+#471 := (and #468 #205)
+#476 := (iff #471 #475)
+#477 := [rewrite]: #476
+#472 := (iff #208 #471)
+#469 := (iff #15 #468)
+#470 := [rewrite]: #469
+#473 := [monotonicity #470]: #472
+#479 := [trans #473 #477]: #478
+#482 := [monotonicity #479]: #481
+#486 := [trans #482 #484]: #485
+#491 := [trans #486 #489]: #490
+#357 := [not-or-elim #347]: #356
+#492 := [mp #357 #491]: #487
+#694 := [unit-resolution #492 #691]: #693
+#695 := [unit-resolution #694 #692]: #177
+#493 := (or #178 #204)
+#358 := (not #211)
+#502 := (iff #358 #493)
+#494 := (not #493)
+#497 := (not #494)
+#500 := (iff #497 #493)
+#501 := [rewrite]: #500
+#498 := (iff #358 #497)
+#495 := (iff #211 #494)
+#496 := [rewrite]: #495
+#499 := [monotonicity #496]: #498
+#503 := [trans #499 #501]: #502
+#359 := [not-or-elim #347]: #358
+#504 := [mp #359 #503]: #493
+#696 := [unit-resolution #504 #695 #692]: false
+#697 := [lemma #696]: #204
+#698 := [hypothesis]: #177
+#449 := (or #178 #205 #217)
+#352 := (not #184)
+#452 := (iff #352 #449)
+#436 := (or #178 #385)
+#450 := (iff #436 #449)
+#451 := [rewrite]: #450
+#447 := (iff #352 #436)
+#437 := (not #436)
+#442 := (not #437)
+#445 := (iff #442 #436)
+#446 := [rewrite]: #445
+#443 := (iff #352 #442)
+#440 := (iff #184 #437)
+#433 := (and #386 #177)
+#438 := (iff #433 #437)
+#439 := [rewrite]: #438
+#434 := (iff #184 #433)
+#435 := [monotonicity #388]: #434
+#441 := [trans #435 #439]: #440
+#444 := [monotonicity #441]: #443
+#448 := [trans #444 #446]: #447
+#453 := [trans #448 #451]: #452
+#353 := [not-or-elim #347]: #352
+#454 := [mp #353 #453]: #449
+#699 := [unit-resolution #454 #698 #697]: #217
+#639 := (or #178 #237)
+#374 := (not #273)
+#648 := (iff #374 #639)
+#640 := (not #639)
+#643 := (not #640)
+#646 := (iff #643 #639)
+#647 := [rewrite]: #646
+#644 := (iff #374 #643)
+#641 := (iff #273 #640)
+#642 := [rewrite]: #641
+#645 := [monotonicity #642]: #644
+#649 := [trans #645 #647]: #648
+#375 := [not-or-elim #347]: #374
+#650 := [mp #375 #649]: #639
+#700 := [unit-resolution #650 #698]: #237
+#667 := (or #189 #218 #238)
+#376 := (not #279)
+#670 := (iff #376 #667)
+#654 := (or #218 #382)
+#668 := (iff #654 #667)
+#669 := [rewrite]: #668
+#665 := (iff #376 #654)
+#655 := (not #654)
+#660 := (not #655)
+#663 := (iff #660 #654)
+#664 := [rewrite]: #663
+#661 := (iff #376 #660)
+#658 := (iff #279 #655)
+#651 := (and #380 #217)
+#656 := (iff #651 #655)
+#657 := [rewrite]: #656
+#652 := (iff #279 #651)
+#653 := [monotonicity #384]: #652
+#659 := [trans #653 #657]: #658
+#662 := [monotonicity #659]: #661
+#666 := [trans #662 #664]: #665
+#671 := [trans #666 #669]: #670
+#377 := [not-or-elim #347]: #376
+#672 := [mp #377 #671]: #667
+#701 := [unit-resolution #672 #699 #700]: #189
+#673 := (or #190 #218)
+#378 := (not #287)
+#682 := (iff #378 #673)
+#674 := (not #673)
+#677 := (not #674)
+#680 := (iff #677 #673)
+#681 := [rewrite]: #680
+#678 := (iff #378 #677)
+#675 := (iff #287 #674)
+#676 := [rewrite]: #675
+#679 := [monotonicity #676]: #678
+#683 := [trans #679 #681]: #682
+#379 := [not-or-elim #347]: #378
+#684 := [mp #379 #683]: #673
+#702 := [unit-resolution #684 #701 #699]: false
+#703 := [lemma #702]: #178
+#704 := (or #177 #218)
+#543 := (or #171 #177 #218)
+#362 := (not #224)
+#546 := (iff #362 #543)
+#530 := (or #218 #467)
+#544 := (iff #530 #543)
+#545 := [rewrite]: #544
+#541 := (iff #362 #530)
+#531 := (not #530)
+#536 := (not #531)
+#539 := (iff #536 #530)
+#540 := [rewrite]: #539
+#537 := (iff #362 #536)
+#534 := (iff #224 #531)
+#527 := (and #468 #217)
+#532 := (iff #527 #531)
+#533 := [rewrite]: #532
+#528 := (iff #224 #527)
+#529 := [monotonicity #470]: #528
+#535 := [trans #529 #533]: #534
+#538 := [monotonicity #535]: #537
+#542 := [trans #538 #540]: #541
+#547 := [trans #542 #545]: #546
+#363 := [not-or-elim #347]: #362
+#548 := [mp #363 #547]: #543
+#705 := [unit-resolution #548 #691]: #704
+#706 := [unit-resolution #705 #703]: #218
+#707 := (or #177 #237)
+#633 := (or #171 #177 #237)
+#372 := (not #268)
+#636 := (iff #372 #633)
+#620 := (or #237 #467)
+#634 := (iff #620 #633)
+#635 := [rewrite]: #634
+#631 := (iff #372 #620)
+#621 := (not #620)
+#626 := (not #621)
+#629 := (iff #626 #620)
+#630 := [rewrite]: #629
+#627 := (iff #372 #626)
+#624 := (iff #268 #621)
+#617 := (and #468 #238)
+#622 := (iff #617 #621)
+#623 := [rewrite]: #622
+#618 := (iff #268 #617)
+#619 := [monotonicity #470]: #618
+#625 := [trans #619 #623]: #624
+#628 := [monotonicity #625]: #627
+#632 := [trans #628 #630]: #631
+#637 := [trans #632 #635]: #636
+#373 := [not-or-elim #347]: #372
+#638 := [mp #373 #637]: #633
+#708 := [unit-resolution #638 #691]: #707
+#709 := [unit-resolution #708 #703]: #237
+#611 := (or #190 #205 #217)
+#370 := (not #260)
+#614 := (iff #370 #611)
+#598 := (or #190 #385)
+#612 := (iff #598 #611)
+#613 := [rewrite]: #612
+#609 := (iff #370 #598)
+#599 := (not #598)
+#604 := (not #599)
+#607 := (iff #604 #598)
+#608 := [rewrite]: #607
+#605 := (iff #370 #604)
+#602 := (iff #260 #599)
+#595 := (and #386 #189)
+#600 := (iff #595 #599)
+#601 := [rewrite]: #600
+#596 := (iff #260 #595)
+#597 := [monotonicity #388]: #596
+#603 := [trans #597 #601]: #602
+#606 := [monotonicity #603]: #605
+#610 := [trans #606 #608]: #609
+#615 := [trans #610 #613]: #614
+#371 := [not-or-elim #347]: #370
+#616 := [mp #371 #615]: #611
+#710 := [unit-resolution #616 #706 #697]: #190
+#405 := (or #189 #205 #217 #238)
+#348 := (not #115)
+#408 := (iff #348 #405)
+#392 := (or #382 #385)
+#406 := (iff #392 #405)
+#407 := [rewrite]: #406
+#403 := (iff #348 #392)
+#393 := (not #392)
+#398 := (not #393)
+#401 := (iff #398 #392)
+#402 := [rewrite]: #401
+#399 := (iff #348 #398)
+#396 := (iff #115 #393)
+#389 := (and #380 #386)
+#394 := (iff #389 #393)
+#395 := [rewrite]: #394
+#390 := (iff #115 #389)
+#391 := [monotonicity #384 #388]: #390
+#397 := [trans #391 #395]: #396
+#400 := [monotonicity #397]: #399
+#404 := [trans #400 #402]: #403
+#409 := [trans #404 #407]: #408
+#349 := [not-or-elim #347]: #348
+#410 := [mp #349 #409]: #405
+[unit-resolution #410 #710 #709 #697 #706]: false
+unsat
+
+ib5n9nvBAC5jXuKIpV/54g 82870
+#2 := false
+#6 := 0::int
+decl z3name!0 :: int
+#647 := z3name!0
+#81 := -1::int
+#656 := (* -1::int z3name!0)
+decl uf_2 :: int
+#5 := uf_2
+#882 := (+ uf_2 #656)
+#883 := (<= #882 0::int)
+#885 := (not #883)
+#881 := (>= #882 0::int)
+#884 := (not #881)
+#886 := (or #884 #885)
+decl uf_11 :: int
+#55 := uf_11
+#513 := (* -1::int uf_11)
+#514 := (+ uf_2 #513)
+#515 := (<= #514 0::int)
+decl z3name!5 :: int
+#777 := z3name!5
+decl uf_7 :: int
+#31 := uf_7
+#1083 := (+ uf_7 z3name!5)
+#1084 := (<= #1083 0::int)
+#335 := (>= uf_7 0::int)
+#1085 := (>= #1083 0::int)
+#1087 := (not #1085)
+#1086 := (not #1084)
+#1088 := (or #1086 #1087)
+#2302 := [hypothesis]: #1086
+#1289 := (or #1088 #1084)
+#1290 := [def-axiom]: #1289
+#2303 := [unit-resolution #1290 #2302]: #1088
+#1089 := (not #1088)
+#1092 := (or #335 #1089)
+#1099 := (not #1092)
+#786 := (* -1::int z3name!5)
+#1072 := (+ uf_7 #786)
+#1073 := (<= #1072 0::int)
+#1075 := (not #1073)
+#1071 := (>= #1072 0::int)
+#1074 := (not #1071)
+#1076 := (or #1074 #1075)
+#1077 := (not #1076)
+#336 := (not #335)
+#1080 := (or #336 #1077)
+#1098 := (not #1080)
+#1100 := (or #1098 #1099)
+#1101 := (not #1100)
+#318 := (* -1::int uf_7)
+#780 := (= z3name!5 #318)
+#781 := (or #335 #780)
+#778 := (= z3name!5 uf_7)
+#779 := (or #336 #778)
+#782 := (and #779 #781)
+#1104 := (iff #782 #1101)
+#1095 := (and #1080 #1092)
+#1102 := (iff #1095 #1101)
+#1103 := [rewrite]: #1102
+#1096 := (iff #782 #1095)
+#1093 := (iff #781 #1092)
+#1090 := (iff #780 #1089)
+#1091 := [rewrite]: #1090
+#1094 := [monotonicity #1091]: #1093
+#1081 := (iff #779 #1080)
+#1078 := (iff #778 #1077)
+#1079 := [rewrite]: #1078
+#1082 := [monotonicity #1079]: #1081
+#1097 := [monotonicity #1082 #1094]: #1096
+#1105 := [trans #1097 #1103]: #1104
+#783 := [intro-def]: #782
+#1106 := [mp #783 #1105]: #1101
+#1108 := [not-or-elim #1106]: #1092
+#2304 := [unit-resolution #1108 #2303]: #335
+decl uf_4 :: int
+#13 := uf_4
+#194 := (>= uf_4 0::int)
+decl uf_10 :: int
+#49 := uf_10
+#459 := (* -1::int uf_10)
+decl uf_3 :: int
+#10 := uf_3
+#508 := (+ uf_3 #459)
+#509 := (>= #508 0::int)
+decl z3name!1 :: int
+#673 := z3name!1
+#682 := (* -1::int z3name!1)
+decl uf_1 :: int
+#4 := uf_1
+#920 := (+ uf_1 #682)
+#921 := (<= #920 0::int)
+#931 := (+ uf_1 z3name!1)
+#933 := (>= #931 0::int)
+#935 := (not #933)
+#932 := (<= #931 0::int)
+#934 := (not #932)
+#936 := (or #934 #935)
+#937 := (not #936)
+#147 := (>= uf_1 0::int)
+#148 := (not #147)
+#923 := (not #921)
+#919 := (>= #920 0::int)
+#922 := (not #919)
+#924 := (or #922 #923)
+#2022 := [hypothesis]: #923
+#1237 := (or #924 #921)
+#1238 := [def-axiom]: #1237
+#2023 := [unit-resolution #1238 #2022]: #924
+#925 := (not #924)
+#928 := (or #148 #925)
+#940 := (or #147 #937)
+#947 := (not #940)
+#946 := (not #928)
+#948 := (or #946 #947)
+#949 := (not #948)
+#130 := (* -1::int uf_1)
+#676 := (= z3name!1 #130)
+#677 := (or #147 #676)
+#674 := (= z3name!1 uf_1)
+#675 := (or #148 #674)
+#678 := (and #675 #677)
+#952 := (iff #678 #949)
+#943 := (and #928 #940)
+#950 := (iff #943 #949)
+#951 := [rewrite]: #950
+#944 := (iff #678 #943)
+#941 := (iff #677 #940)
+#938 := (iff #676 #937)
+#939 := [rewrite]: #938
+#942 := [monotonicity #939]: #941
+#929 := (iff #675 #928)
+#926 := (iff #674 #925)
+#927 := [rewrite]: #926
+#930 := [monotonicity #927]: #929
+#945 := [monotonicity #930 #942]: #944
+#953 := [trans #945 #951]: #952
+#679 := [intro-def]: #678
+#954 := [mp #679 #953]: #949
+#955 := [not-or-elim #954]: #928
+#2024 := [unit-resolution #955 #2023]: #148
+#956 := [not-or-elim #954]: #940
+#2025 := [unit-resolution #956 #2024]: #937
+#2026 := (or #921 #919)
+#2027 := [th-lemma]: #2026
+#2028 := [unit-resolution #2027 #2022]: #919
+#2029 := (or #922 #147 #935)
+#2030 := [th-lemma]: #2029
+#2031 := [unit-resolution #2030 #2024 #2028]: #935
+#1243 := (or #936 #933)
+#1244 := [def-axiom]: #1243
+#2032 := [unit-resolution #1244 #2031 #2025]: false
+#2033 := [lemma #2032]: #921
+decl z3name!7 :: int
+#829 := z3name!7
+decl uf_9 :: int
+#43 := uf_9
+#1159 := (+ uf_9 z3name!7)
+#1160 := (<= #1159 0::int)
+#838 := (* -1::int z3name!7)
+#1148 := (+ uf_9 #838)
+#1147 := (>= #1148 0::int)
+decl z3name!4 :: int
+#751 := z3name!4
+#760 := (* -1::int z3name!4)
+decl uf_6 :: int
+#25 := uf_6
+#1034 := (+ uf_6 #760)
+#1033 := (>= #1034 0::int)
+#1035 := (<= #1034 0::int)
+#1037 := (not #1035)
+#1036 := (not #1033)
+#1038 := (or #1036 #1037)
+#1039 := (not #1038)
+#288 := (>= uf_6 0::int)
+#893 := (+ uf_2 z3name!0)
+#895 := (>= #893 0::int)
+#897 := (not #895)
+#894 := (<= #893 0::int)
+#896 := (not #894)
+#898 := (or #896 #897)
+#899 := (not #898)
+#100 := (>= uf_2 0::int)
+#101 := (not #100)
+#1736 := [hypothesis]: #885
+#1225 := (or #886 #883)
+#1226 := [def-axiom]: #1225
+#1737 := [unit-resolution #1226 #1736]: #886
+#887 := (not #886)
+#890 := (or #101 #887)
+#902 := (or #100 #899)
+#909 := (not #902)
+#908 := (not #890)
+#910 := (or #908 #909)
+#911 := (not #910)
+#82 := (* -1::int uf_2)
+#650 := (= z3name!0 #82)
+#651 := (or #100 #650)
+#648 := (= z3name!0 uf_2)
+#649 := (or #101 #648)
+#652 := (and #649 #651)
+#914 := (iff #652 #911)
+#905 := (and #890 #902)
+#912 := (iff #905 #911)
+#913 := [rewrite]: #912
+#906 := (iff #652 #905)
+#903 := (iff #651 #902)
+#900 := (iff #650 #899)
+#901 := [rewrite]: #900
+#904 := [monotonicity #901]: #903
+#891 := (iff #649 #890)
+#888 := (iff #648 #887)
+#889 := [rewrite]: #888
+#892 := [monotonicity #889]: #891
+#907 := [monotonicity #892 #904]: #906
+#915 := [trans #907 #913]: #914
+#653 := [intro-def]: #652
+#916 := [mp #653 #915]: #911
+#917 := [not-or-elim #916]: #890
+#1738 := [unit-resolution #917 #1737]: #101
+#918 := [not-or-elim #916]: #902
+#1739 := [unit-resolution #918 #1738]: #899
+#1231 := (or #898 #895)
+#1232 := [def-axiom]: #1231
+#1740 := [unit-resolution #1232 #1739]: #895
+#1741 := [th-lemma #1736 #1738 #1740]: false
+#1742 := [lemma #1741]: #883
+#1149 := (<= #1148 0::int)
+#1151 := (not #1149)
+#1150 := (not #1147)
+#1152 := (or #1150 #1151)
+#1153 := (not #1152)
+#429 := (>= uf_9 0::int)
+decl z3name!6 :: int
+#803 := z3name!6
+#812 := (* -1::int z3name!6)
+decl uf_8 :: int
+#37 := uf_8
+#1110 := (+ uf_8 #812)
+#1111 := (<= #1110 0::int)
+#1113 := (not #1111)
+#1109 := (>= #1110 0::int)
+#1112 := (not #1109)
+#1114 := (or #1112 #1113)
+#1865 := [hypothesis]: #1113
+#1297 := (or #1114 #1111)
+#1298 := [def-axiom]: #1297
+#1866 := [unit-resolution #1298 #1865]: #1114
+#382 := (>= uf_8 0::int)
+#1685 := (or #1111 #1109)
+#1686 := [th-lemma]: #1685
+#1867 := [unit-resolution #1686 #1865]: #1109
+#1734 := (or #382 #1112)
+#1121 := (+ uf_8 z3name!6)
+#1123 := (>= #1121 0::int)
+#1125 := (not #1123)
+#1122 := (<= #1121 0::int)
+#1124 := (not #1122)
+#1126 := (or #1124 #1125)
+#1127 := (not #1126)
+#383 := (not #382)
+#1428 := [hypothesis]: #383
+#1130 := (or #382 #1127)
+#1137 := (not #1130)
+#1115 := (not #1114)
+#1118 := (or #383 #1115)
+#1136 := (not #1118)
+#1138 := (or #1136 #1137)
+#1139 := (not #1138)
+#365 := (* -1::int uf_8)
+#806 := (= z3name!6 #365)
+#807 := (or #382 #806)
+#804 := (= z3name!6 uf_8)
+#805 := (or #383 #804)
+#808 := (and #805 #807)
+#1142 := (iff #808 #1139)
+#1133 := (and #1118 #1130)
+#1140 := (iff #1133 #1139)
+#1141 := [rewrite]: #1140
+#1134 := (iff #808 #1133)
+#1131 := (iff #807 #1130)
+#1128 := (iff #806 #1127)
+#1129 := [rewrite]: #1128
+#1132 := [monotonicity #1129]: #1131
+#1119 := (iff #805 #1118)
+#1116 := (iff #804 #1115)
+#1117 := [rewrite]: #1116
+#1120 := [monotonicity #1117]: #1119
+#1135 := [monotonicity #1120 #1132]: #1134
+#1143 := [trans #1135 #1141]: #1142
+#809 := [intro-def]: #808
+#1144 := [mp #809 #1143]: #1139
+#1146 := [not-or-elim #1144]: #1130
+#1729 := [unit-resolution #1146 #1428]: #1127
+#1637 := [hypothesis]: #1109
+#1730 := (or #1112 #1125 #382)
+#1731 := [th-lemma]: #1730
+#1732 := [unit-resolution #1731 #1428 #1637]: #1125
+#1303 := (or #1126 #1123)
+#1304 := [def-axiom]: #1303
+#1733 := [unit-resolution #1304 #1732 #1729]: false
+#1735 := [lemma #1733]: #1734
+#1868 := [unit-resolution #1735 #1867]: #382
+#1145 := [not-or-elim #1144]: #1118
+#1869 := [unit-resolution #1145 #1868 #1866]: false
+#1870 := [lemma #1869]: #1111
+#289 := (not #288)
+#1405 := [hypothesis]: #289
+#1688 := (or #288 #429 #1113)
+#815 := (+ uf_9 #812)
+#818 := (+ uf_7 #815)
+#825 := (>= #818 0::int)
+#389 := (ite #382 uf_8 #365)
+#400 := (* -1::int #389)
+#401 := (+ uf_9 #400)
+#402 := (+ uf_7 #401)
+#599 := (>= #402 0::int)
+#826 := (= #599 #825)
+#819 := (~ #402 #818)
+#816 := (~ #401 #815)
+#813 := (~ #400 #812)
+#810 := (~ #389 z3name!6)
+#811 := [apply-def #809]: #810
+#814 := [monotonicity #811]: #813
+#817 := [monotonicity #814]: #816
+#820 := [monotonicity #817]: #819
+#827 := [monotonicity #820]: #826
+#601 := (not #599)
+#598 := (<= #402 0::int)
+#600 := (not #598)
+#602 := (or #600 #601)
+#603 := (not #602)
+#403 := (= #402 0::int)
+#604 := (iff #403 #603)
+#605 := [rewrite]: #604
+#45 := (- uf_8)
+#44 := (< uf_8 0::int)
+#46 := (ite #44 #45 uf_8)
+#47 := (- #46 uf_7)
+#48 := (= uf_9 #47)
+#408 := (iff #48 #403)
+#368 := (ite #44 #365 uf_8)
+#374 := (+ #318 #368)
+#379 := (= uf_9 #374)
+#406 := (iff #379 #403)
+#394 := (+ #318 #389)
+#397 := (= uf_9 #394)
+#404 := (iff #397 #403)
+#405 := [rewrite]: #404
+#398 := (iff #379 #397)
+#395 := (= #374 #394)
+#392 := (= #368 #389)
+#386 := (ite #383 #365 uf_8)
+#390 := (= #386 #389)
+#391 := [rewrite]: #390
+#387 := (= #368 #386)
+#384 := (iff #44 #383)
+#385 := [rewrite]: #384
+#388 := [monotonicity #385]: #387
+#393 := [trans #388 #391]: #392
+#396 := [monotonicity #393]: #395
+#399 := [monotonicity #396]: #398
+#407 := [trans #399 #405]: #406
+#380 := (iff #48 #379)
+#377 := (= #47 #374)
+#371 := (- #368 uf_7)
+#375 := (= #371 #374)
+#376 := [rewrite]: #375
+#372 := (= #47 #371)
+#369 := (= #46 #368)
+#366 := (= #45 #365)
+#367 := [rewrite]: #366
+#370 := [monotonicity #367]: #369
+#373 := [monotonicity #370]: #372
+#378 := [trans #373 #376]: #377
+#381 := [monotonicity #378]: #380
+#409 := [trans #381 #407]: #408
+#364 := [asserted]: #48
+#410 := [mp #364 #409]: #403
+#606 := [mp #410 #605]: #603
+#608 := [not-or-elim #606]: #599
+#828 := [mp~ #608 #827]: #825
+#1441 := [hypothesis]: #1075
+#1285 := (or #1076 #1073)
+#1286 := [def-axiom]: #1285
+#1442 := [unit-resolution #1286 #1441]: #1076
+#1107 := [not-or-elim #1106]: #1080
+#1443 := [unit-resolution #1107 #1442]: #336
+#1444 := [unit-resolution #1108 #1443]: #1089
+#1291 := (or #1088 #1085)
+#1292 := [def-axiom]: #1291
+#1445 := [unit-resolution #1292 #1444]: #1085
+#1446 := [th-lemma #1441 #1445 #1443]: false
+#1447 := [lemma #1446]: #1073
+#789 := (+ uf_8 #786)
+#792 := (+ uf_6 #789)
+#799 := (>= #792 0::int)
+#342 := (ite #335 uf_7 #318)
+#353 := (* -1::int #342)
+#354 := (+ uf_8 #353)
+#355 := (+ uf_6 #354)
+#588 := (>= #355 0::int)
+#800 := (= #588 #799)
+#793 := (~ #355 #792)
+#790 := (~ #354 #789)
+#787 := (~ #353 #786)
+#784 := (~ #342 z3name!5)
+#785 := [apply-def #783]: #784
+#788 := [monotonicity #785]: #787
+#791 := [monotonicity #788]: #790
+#794 := [monotonicity #791]: #793
+#801 := [monotonicity #794]: #800
+#590 := (not #588)
+#587 := (<= #355 0::int)
+#589 := (not #587)
+#591 := (or #589 #590)
+#592 := (not #591)
+#356 := (= #355 0::int)
+#593 := (iff #356 #592)
+#594 := [rewrite]: #593
+#39 := (- uf_7)
+#38 := (< uf_7 0::int)
+#40 := (ite #38 #39 uf_7)
+#41 := (- #40 uf_6)
+#42 := (= uf_8 #41)
+#361 := (iff #42 #356)
+#321 := (ite #38 #318 uf_7)
+#271 := (* -1::int uf_6)
+#327 := (+ #271 #321)
+#332 := (= uf_8 #327)
+#359 := (iff #332 #356)
+#347 := (+ #271 #342)
+#350 := (= uf_8 #347)
+#357 := (iff #350 #356)
+#358 := [rewrite]: #357
+#351 := (iff #332 #350)
+#348 := (= #327 #347)
+#345 := (= #321 #342)
+#339 := (ite #336 #318 uf_7)
+#343 := (= #339 #342)
+#344 := [rewrite]: #343
+#340 := (= #321 #339)
+#337 := (iff #38 #336)
+#338 := [rewrite]: #337
+#341 := [monotonicity #338]: #340
+#346 := [trans #341 #344]: #345
+#349 := [monotonicity #346]: #348
+#352 := [monotonicity #349]: #351
+#360 := [trans #352 #358]: #359
+#333 := (iff #42 #332)
+#330 := (= #41 #327)
+#324 := (- #321 uf_6)
+#328 := (= #324 #327)
+#329 := [rewrite]: #328
+#325 := (= #41 #324)
+#322 := (= #40 #321)
+#319 := (= #39 #318)
+#320 := [rewrite]: #319
+#323 := [monotonicity #320]: #322
+#326 := [monotonicity #323]: #325
+#331 := [trans #326 #329]: #330
+#334 := [monotonicity #331]: #333
+#362 := [trans #334 #360]: #361
+#317 := [asserted]: #42
+#363 := [mp #317 #362]: #356
+#595 := [mp #363 #594]: #592
+#597 := [not-or-elim #595]: #588
+#802 := [mp~ #597 #801]: #799
+#1343 := (not #825)
+#1350 := (not #799)
+#1351 := (or #288 #1075 #1350 #429 #1113 #1343)
+#1352 := [th-lemma]: #1351
+#1689 := [unit-resolution #1352 #802 #1447 #828]: #1688
+#2046 := [unit-resolution #1689 #1405 #1870]: #429
+#430 := (not #429)
+#1156 := (or #430 #1153)
+#1161 := (>= #1159 0::int)
+#1163 := (not #1161)
+#1162 := (not #1160)
+#1164 := (or #1162 #1163)
+#1165 := (not #1164)
+#1168 := (or #429 #1165)
+#1175 := (not #1168)
+#1174 := (not #1156)
+#1176 := (or #1174 #1175)
+#1177 := (not #1176)
+#412 := (* -1::int uf_9)
+#832 := (= z3name!7 #412)
+#833 := (or #429 #832)
+#830 := (= z3name!7 uf_9)
+#831 := (or #430 #830)
+#834 := (and #831 #833)
+#1180 := (iff #834 #1177)
+#1171 := (and #1156 #1168)
+#1178 := (iff #1171 #1177)
+#1179 := [rewrite]: #1178
+#1172 := (iff #834 #1171)
+#1169 := (iff #833 #1168)
+#1166 := (iff #832 #1165)
+#1167 := [rewrite]: #1166
+#1170 := [monotonicity #1167]: #1169
+#1157 := (iff #831 #1156)
+#1154 := (iff #830 #1153)
+#1155 := [rewrite]: #1154
+#1158 := [monotonicity #1155]: #1157
+#1173 := [monotonicity #1158 #1170]: #1172
+#1181 := [trans #1173 #1179]: #1180
+#835 := [intro-def]: #834
+#1182 := [mp #835 #1181]: #1177
+#1183 := [not-or-elim #1182]: #1156
+#2047 := [unit-resolution #1183 #2046]: #1153
+#1307 := (or #1152 #1147)
+#1308 := [def-axiom]: #1307
+#2112 := [unit-resolution #1308 #2047]: #1147
+#2009 := (or #288 #382)
+#1998 := (or #1036 #288)
+#1045 := (+ uf_6 z3name!4)
+#1047 := (>= #1045 0::int)
+#1049 := (not #1047)
+#1046 := (<= #1045 0::int)
+#1048 := (not #1046)
+#1050 := (or #1048 #1049)
+#1460 := [hypothesis]: #1049
+#1279 := (or #1050 #1047)
+#1280 := [def-axiom]: #1279
+#1461 := [unit-resolution #1280 #1460]: #1050
+#1464 := (or #1047 #289)
+#1051 := (not #1050)
+#1448 := [hypothesis]: #1037
+#1273 := (or #1038 #1035)
+#1274 := [def-axiom]: #1273
+#1449 := [unit-resolution #1274 #1448]: #1038
+#1042 := (or #289 #1039)
+#1054 := (or #288 #1051)
+#1061 := (not #1054)
+#1060 := (not #1042)
+#1062 := (or #1060 #1061)
+#1063 := (not #1062)
+#754 := (= z3name!4 #271)
+#755 := (or #288 #754)
+#752 := (= z3name!4 uf_6)
+#753 := (or #289 #752)
+#756 := (and #753 #755)
+#1066 := (iff #756 #1063)
+#1057 := (and #1042 #1054)
+#1064 := (iff #1057 #1063)
+#1065 := [rewrite]: #1064
+#1058 := (iff #756 #1057)
+#1055 := (iff #755 #1054)
+#1052 := (iff #754 #1051)
+#1053 := [rewrite]: #1052
+#1056 := [monotonicity #1053]: #1055
+#1043 := (iff #753 #1042)
+#1040 := (iff #752 #1039)
+#1041 := [rewrite]: #1040
+#1044 := [monotonicity #1041]: #1043
+#1059 := [monotonicity #1044 #1056]: #1058
+#1067 := [trans #1059 #1065]: #1066
+#757 := [intro-def]: #756
+#1068 := [mp #757 #1067]: #1063
+#1069 := [not-or-elim #1068]: #1042
+#1450 := [unit-resolution #1069 #1449]: #289
+#1070 := [not-or-elim #1068]: #1054
+#1451 := [unit-resolution #1070 #1450]: #1051
+#1452 := (or #1035 #1033)
+#1453 := [th-lemma]: #1452
+#1454 := [unit-resolution #1453 #1448]: #1033
+#1455 := (or #1036 #288 #1049)
+#1456 := [th-lemma]: #1455
+#1457 := [unit-resolution #1456 #1450 #1454]: #1049
+#1458 := [unit-resolution #1280 #1457 #1451]: false
+#1459 := [lemma #1458]: #1035
+#1462 := (or #1047 #1037 #289)
+#1463 := [th-lemma]: #1462
+#1465 := [unit-resolution #1463 #1459]: #1464
+#1466 := [unit-resolution #1465 #1460]: #289
+#1467 := [unit-resolution #1070 #1466 #1461]: false
+#1468 := [lemma #1467]: #1047
+#1999 := [unit-resolution #1456 #1468]: #1998
+#2000 := [unit-resolution #1999 #1405]: #1036
+#1407 := [unit-resolution #1070 #1405]: #1051
+#1277 := (or #1050 #1046)
+#1278 := [def-axiom]: #1277
+#1497 := [unit-resolution #1278 #1407]: #1046
+#2001 := (or #336 #1048 #1033 #382 #1350 #1075)
+#2002 := [th-lemma]: #2001
+#2003 := [unit-resolution #2002 #1497 #2000 #1447 #802 #1428]: #336
+#2004 := (or #1087 #1075 #1048 #1033 #382 #1350)
+#2005 := [th-lemma]: #2004
+#2006 := [unit-resolution #2005 #1497 #1447 #2000 #802 #1428]: #1087
+#2007 := [unit-resolution #1292 #2006]: #1088
+#2008 := [unit-resolution #1108 #2007 #2003]: false
+#2010 := [lemma #2008]: #2009
+#2113 := [unit-resolution #2010 #1405]: #382
+#2114 := [unit-resolution #1145 #2113]: #1115
+#1295 := (or #1114 #1109)
+#1296 := [def-axiom]: #1295
+#2115 := [unit-resolution #1296 #2114]: #1109
+decl z3name!2 :: int
+#699 := z3name!2
+#708 := (* -1::int z3name!2)
+#958 := (+ uf_4 #708)
+#957 := (>= #958 0::int)
+#959 := (<= #958 0::int)
+#961 := (not #959)
+#960 := (not #957)
+#962 := (or #960 #961)
+#963 := (not #962)
+decl uf_5 :: int
+#19 := uf_5
+#241 := (>= uf_5 0::int)
+#242 := (not #241)
+#1406 := [hypothesis]: #242
+#1579 := (or #1048 #241)
+#516 := (>= #514 0::int)
+#476 := (>= uf_10 0::int)
+#477 := (not #476)
+#1484 := (or #382 #241)
+#1430 := (or #382 #241 #1075 #1037)
+#1421 := [hypothesis]: #1035
+#1427 := [hypothesis]: #1073
+#763 := (+ uf_7 #760)
+#766 := (+ uf_5 #763)
+#773 := (>= #766 0::int)
+#295 := (ite #288 uf_6 #271)
+#306 := (* -1::int #295)
+#307 := (+ uf_7 #306)
+#308 := (+ uf_5 #307)
+#577 := (>= #308 0::int)
+#774 := (= #577 #773)
+#767 := (~ #308 #766)
+#764 := (~ #307 #763)
+#761 := (~ #306 #760)
+#758 := (~ #295 z3name!4)
+#759 := [apply-def #757]: #758
+#762 := [monotonicity #759]: #761
+#765 := [monotonicity #762]: #764
+#768 := [monotonicity #765]: #767
+#775 := [monotonicity #768]: #774
+#579 := (not #577)
+#576 := (<= #308 0::int)
+#578 := (not #576)
+#580 := (or #578 #579)
+#581 := (not #580)
+#309 := (= #308 0::int)
+#582 := (iff #309 #581)
+#583 := [rewrite]: #582
+#33 := (- uf_6)
+#32 := (< uf_6 0::int)
+#34 := (ite #32 #33 uf_6)
+#35 := (- #34 uf_5)
+#36 := (= uf_7 #35)
+#314 := (iff #36 #309)
+#274 := (ite #32 #271 uf_6)
+#224 := (* -1::int uf_5)
+#280 := (+ #224 #274)
+#285 := (= uf_7 #280)
+#312 := (iff #285 #309)
+#300 := (+ #224 #295)
+#303 := (= uf_7 #300)
+#310 := (iff #303 #309)
+#311 := [rewrite]: #310
+#304 := (iff #285 #303)
+#301 := (= #280 #300)
+#298 := (= #274 #295)
+#292 := (ite #289 #271 uf_6)
+#296 := (= #292 #295)
+#297 := [rewrite]: #296
+#293 := (= #274 #292)
+#290 := (iff #32 #289)
+#291 := [rewrite]: #290
+#294 := [monotonicity #291]: #293
+#299 := [trans #294 #297]: #298
+#302 := [monotonicity #299]: #301
+#305 := [monotonicity #302]: #304
+#313 := [trans #305 #311]: #312
+#286 := (iff #36 #285)
+#283 := (= #35 #280)
+#277 := (- #274 uf_5)
+#281 := (= #277 #280)
+#282 := [rewrite]: #281
+#278 := (= #35 #277)
+#275 := (= #34 #274)
+#272 := (= #33 #271)
+#273 := [rewrite]: #272
+#276 := [monotonicity #273]: #275
+#279 := [monotonicity #276]: #278
+#284 := [trans #279 #282]: #283
+#287 := [monotonicity #284]: #286
+#315 := [trans #287 #313]: #314
+#270 := [asserted]: #36
+#316 := [mp #270 #315]: #309
+#584 := [mp #316 #583]: #581
+#586 := [not-or-elim #584]: #577
+#776 := [mp~ #586 #775]: #773
+#1429 := [th-lemma #776 #1406 #1428 #1427 #802 #1421]: false
+#1431 := [lemma #1429]: #1430
+#1485 := [unit-resolution #1431 #1447 #1459]: #1484
+#1486 := [unit-resolution #1485 #1406]: #382
+#1487 := [unit-resolution #1145 #1486]: #1115
+#1496 := [unit-resolution #1298 #1487]: #1111
+#1545 := [hypothesis]: #1046
+#1548 := (or #1048 #1113 #429)
+#1546 := (or #1048 #1113 #429 #1343 #1075 #1350 #1037)
+#1547 := [th-lemma]: #1546
+#1549 := [unit-resolution #1547 #1447 #802 #1459 #828]: #1548
+#1550 := [unit-resolution #1549 #1545 #1496]: #429
+#1551 := [unit-resolution #1183 #1550]: #1153
+#1552 := [unit-resolution #1308 #1551]: #1147
+#1543 := (or #477 #241 #1150)
+#1488 := [unit-resolution #1296 #1487]: #1109
+#821 := (<= #818 0::int)
+#822 := (= #598 #821)
+#823 := [monotonicity #820]: #822
+#607 := [not-or-elim #606]: #598
+#824 := [mp~ #607 #823]: #821
+#841 := (+ uf_10 #838)
+#844 := (+ uf_8 #841)
+#847 := (<= #844 0::int)
+#436 := (ite #429 uf_9 #412)
+#447 := (* -1::int #436)
+#448 := (+ uf_10 #447)
+#449 := (+ uf_8 #448)
+#609 := (<= #449 0::int)
+#848 := (= #609 #847)
+#845 := (~ #449 #844)
+#842 := (~ #448 #841)
+#839 := (~ #447 #838)
+#836 := (~ #436 z3name!7)
+#837 := [apply-def #835]: #836
+#840 := [monotonicity #837]: #839
+#843 := [monotonicity #840]: #842
+#846 := [monotonicity #843]: #845
+#849 := [monotonicity #846]: #848
+#610 := (>= #449 0::int)
+#612 := (not #610)
+#611 := (not #609)
+#613 := (or #611 #612)
+#614 := (not #613)
+#450 := (= #449 0::int)
+#615 := (iff #450 #614)
+#616 := [rewrite]: #615
+#51 := (- uf_9)
+#50 := (< uf_9 0::int)
+#52 := (ite #50 #51 uf_9)
+#53 := (- #52 uf_8)
+#54 := (= uf_10 #53)
+#455 := (iff #54 #450)
+#415 := (ite #50 #412 uf_9)
+#421 := (+ #365 #415)
+#426 := (= uf_10 #421)
+#453 := (iff #426 #450)
+#441 := (+ #365 #436)
+#444 := (= uf_10 #441)
+#451 := (iff #444 #450)
+#452 := [rewrite]: #451
+#445 := (iff #426 #444)
+#442 := (= #421 #441)
+#439 := (= #415 #436)
+#433 := (ite #430 #412 uf_9)
+#437 := (= #433 #436)
+#438 := [rewrite]: #437
+#434 := (= #415 #433)
+#431 := (iff #50 #430)
+#432 := [rewrite]: #431
+#435 := [monotonicity #432]: #434
+#440 := [trans #435 #438]: #439
+#443 := [monotonicity #440]: #442
+#446 := [monotonicity #443]: #445
+#454 := [trans #446 #452]: #453
+#427 := (iff #54 #426)
+#424 := (= #53 #421)
+#418 := (- #415 uf_8)
+#422 := (= #418 #421)
+#423 := [rewrite]: #422
+#419 := (= #53 #418)
+#416 := (= #52 #415)
+#413 := (= #51 #412)
+#414 := [rewrite]: #413
+#417 := [monotonicity #414]: #416
+#420 := [monotonicity #417]: #419
+#425 := [trans #420 #423]: #424
+#428 := [monotonicity #425]: #427
+#456 := [trans #428 #454]: #455
+#411 := [asserted]: #54
+#457 := [mp #411 #456]: #450
+#617 := [mp #457 #616]: #614
+#618 := [not-or-elim #617]: #609
+#850 := [mp~ #618 #849]: #847
+#1540 := [hypothesis]: #1147
+#1541 := [hypothesis]: #476
+#1542 := [th-lemma #1468 #1406 #1541 #1540 #850 #824 #1488 #776 #1459]: false
+#1544 := [lemma #1542]: #1543
+#1553 := [unit-resolution #1544 #1552 #1406]: #477
+#851 := (>= #844 0::int)
+#852 := (= #610 #851)
+#853 := [monotonicity #846]: #852
+#619 := [not-or-elim #617]: #610
+#854 := [mp~ #619 #853]: #851
+#1309 := (or #1152 #1149)
+#1310 := [def-axiom]: #1309
+#1554 := [unit-resolution #1310 #1551]: #1149
+#769 := (<= #766 0::int)
+#770 := (= #576 #769)
+#771 := [monotonicity #768]: #770
+#585 := [not-or-elim #584]: #576
+#772 := [mp~ #585 #771]: #769
+decl z3name!3 :: int
+#725 := z3name!3
+#1007 := (+ uf_5 z3name!3)
+#1009 := (>= #1007 0::int)
+#1011 := (not #1009)
+#1398 := [hypothesis]: #1011
+#734 := (* -1::int z3name!3)
+#996 := (+ uf_5 #734)
+#997 := (<= #996 0::int)
+#999 := (not #997)
+#995 := (>= #996 0::int)
+#998 := (not #995)
+#1000 := (or #998 #999)
+#1001 := (not #1000)
+#1008 := (<= #1007 0::int)
+#1010 := (not #1008)
+#1012 := (or #1010 #1011)
+#1267 := (or #1012 #1009)
+#1268 := [def-axiom]: #1267
+#1399 := [unit-resolution #1268 #1398]: #1012
+#1013 := (not #1012)
+#1016 := (or #241 #1013)
+#1023 := (not #1016)
+#1004 := (or #242 #1001)
+#1022 := (not #1004)
+#1024 := (or #1022 #1023)
+#1025 := (not #1024)
+#728 := (= z3name!3 #224)
+#729 := (or #241 #728)
+#726 := (= z3name!3 uf_5)
+#727 := (or #242 #726)
+#730 := (and #727 #729)
+#1028 := (iff #730 #1025)
+#1019 := (and #1004 #1016)
+#1026 := (iff #1019 #1025)
+#1027 := [rewrite]: #1026
+#1020 := (iff #730 #1019)
+#1017 := (iff #729 #1016)
+#1014 := (iff #728 #1013)
+#1015 := [rewrite]: #1014
+#1018 := [monotonicity #1015]: #1017
+#1005 := (iff #727 #1004)
+#1002 := (iff #726 #1001)
+#1003 := [rewrite]: #1002
+#1006 := [monotonicity #1003]: #1005
+#1021 := [monotonicity #1006 #1018]: #1020
+#1029 := [trans #1021 #1027]: #1028
+#731 := [intro-def]: #730
+#1030 := [mp #731 #1029]: #1025
+#1032 := [not-or-elim #1030]: #1016
+#1400 := [unit-resolution #1032 #1399]: #241
+#1031 := [not-or-elim #1030]: #1004
+#1401 := [unit-resolution #1031 #1400]: #1001
+#1261 := (or #1000 #997)
+#1262 := [def-axiom]: #1261
+#1402 := [unit-resolution #1262 #1401]: #997
+#1403 := [th-lemma #1400 #1402 #1398]: false
+#1404 := [lemma #1403]: #1009
+#737 := (+ uf_6 #734)
+#740 := (+ uf_4 #737)
+#747 := (>= #740 0::int)
+#248 := (ite #241 uf_5 #224)
+#259 := (* -1::int #248)
+#260 := (+ uf_6 #259)
+#261 := (+ uf_4 #260)
+#566 := (>= #261 0::int)
+#748 := (= #566 #747)
+#741 := (~ #261 #740)
+#738 := (~ #260 #737)
+#735 := (~ #259 #734)
+#732 := (~ #248 z3name!3)
+#733 := [apply-def #731]: #732
+#736 := [monotonicity #733]: #735
+#739 := [monotonicity #736]: #738
+#742 := [monotonicity #739]: #741
+#749 := [monotonicity #742]: #748
+#568 := (not #566)
+#565 := (<= #261 0::int)
+#567 := (not #565)
+#569 := (or #567 #568)
+#570 := (not #569)
+#262 := (= #261 0::int)
+#571 := (iff #262 #570)
+#572 := [rewrite]: #571
+#27 := (- uf_5)
+#26 := (< uf_5 0::int)
+#28 := (ite #26 #27 uf_5)
+#29 := (- #28 uf_4)
+#30 := (= uf_6 #29)
+#267 := (iff #30 #262)
+#227 := (ite #26 #224 uf_5)
+#177 := (* -1::int uf_4)
+#233 := (+ #177 #227)
+#238 := (= uf_6 #233)
+#265 := (iff #238 #262)
+#253 := (+ #177 #248)
+#256 := (= uf_6 #253)
+#263 := (iff #256 #262)
+#264 := [rewrite]: #263
+#257 := (iff #238 #256)
+#254 := (= #233 #253)
+#251 := (= #227 #248)
+#245 := (ite #242 #224 uf_5)
+#249 := (= #245 #248)
+#250 := [rewrite]: #249
+#246 := (= #227 #245)
+#243 := (iff #26 #242)
+#244 := [rewrite]: #243
+#247 := [monotonicity #244]: #246
+#252 := [trans #247 #250]: #251
+#255 := [monotonicity #252]: #254
+#258 := [monotonicity #255]: #257
+#266 := [trans #258 #264]: #265
+#239 := (iff #30 #238)
+#236 := (= #29 #233)
+#230 := (- #227 uf_4)
+#234 := (= #230 #233)
+#235 := [rewrite]: #234
+#231 := (= #29 #230)
+#228 := (= #28 #227)
+#225 := (= #27 #224)
+#226 := [rewrite]: #225
+#229 := [monotonicity #226]: #228
+#232 := [monotonicity #229]: #231
+#237 := [trans #232 #235]: #236
+#240 := [monotonicity #237]: #239
+#268 := [trans #240 #266]: #267
+#223 := [asserted]: #30
+#269 := [mp #223 #268]: #262
+#573 := [mp #269 #572]: #570
+#575 := [not-or-elim #573]: #566
+#750 := [mp~ #575 #749]: #747
+#1364 := (not #747)
+#1357 := (not #769)
+#1337 := (not #851)
+#1555 := (or #194 #476 #1151 #1337 #1343 #1113 #1048 #1357 #1364 #1011)
+#1556 := [th-lemma]: #1555
+#1557 := [unit-resolution #1556 #1545 #750 #1404 #772 #1496 #828 #1554 #854 #1553]: #194
+#195 := (not #194)
+#966 := (or #195 #963)
+#969 := (+ uf_4 z3name!2)
+#971 := (>= #969 0::int)
+#973 := (not #971)
+#970 := (<= #969 0::int)
+#972 := (not #970)
+#974 := (or #972 #973)
+#975 := (not #974)
+#978 := (or #194 #975)
+#985 := (not #978)
+#984 := (not #966)
+#986 := (or #984 #985)
+#987 := (not #986)
+#702 := (= z3name!2 #177)
+#703 := (or #194 #702)
+#700 := (= z3name!2 uf_4)
+#701 := (or #195 #700)
+#704 := (and #701 #703)
+#990 := (iff #704 #987)
+#981 := (and #966 #978)
+#988 := (iff #981 #987)
+#989 := [rewrite]: #988
+#982 := (iff #704 #981)
+#979 := (iff #703 #978)
+#976 := (iff #702 #975)
+#977 := [rewrite]: #976
+#980 := [monotonicity #977]: #979
+#967 := (iff #701 #966)
+#964 := (iff #700 #963)
+#965 := [rewrite]: #964
+#968 := [monotonicity #965]: #967
+#983 := [monotonicity #968 #980]: #982
+#991 := [trans #983 #989]: #990
+#705 := [intro-def]: #704
+#992 := [mp #705 #991]: #987
+#993 := [not-or-elim #992]: #966
+#1558 := [unit-resolution #993 #1557]: #963
+#1249 := (or #962 #959)
+#1250 := [def-axiom]: #1249
+#1559 := [unit-resolution #1250 #1558]: #959
+decl z3name!8 :: int
+#855 := z3name!8
+#864 := (* -1::int z3name!8)
+#867 := (+ uf_11 #864)
+#870 := (+ uf_9 #867)
+#873 := (<= #870 0::int)
+#483 := (ite #476 uf_10 #459)
+#494 := (* -1::int #483)
+#495 := (+ uf_11 #494)
+#496 := (+ uf_9 #495)
+#620 := (<= #496 0::int)
+#874 := (= #620 #873)
+#871 := (~ #496 #870)
+#868 := (~ #495 #867)
+#865 := (~ #494 #864)
+#862 := (~ #483 z3name!8)
+#858 := (= z3name!8 #459)
+#859 := (or #476 #858)
+#856 := (= z3name!8 uf_10)
+#857 := (or #477 #856)
+#860 := (and #857 #859)
+#861 := [intro-def]: #860
+#863 := [apply-def #861]: #862
+#866 := [monotonicity #863]: #865
+#869 := [monotonicity #866]: #868
+#872 := [monotonicity #869]: #871
+#875 := [monotonicity #872]: #874
+#621 := (>= #496 0::int)
+#623 := (not #621)
+#622 := (not #620)
+#624 := (or #622 #623)
+#625 := (not #624)
+#497 := (= #496 0::int)
+#626 := (iff #497 #625)
+#627 := [rewrite]: #626
+#57 := (- uf_10)
+#56 := (< uf_10 0::int)
+#58 := (ite #56 #57 uf_10)
+#59 := (- #58 uf_9)
+#60 := (= uf_11 #59)
+#502 := (iff #60 #497)
+#462 := (ite #56 #459 uf_10)
+#468 := (+ #412 #462)
+#473 := (= uf_11 #468)
+#500 := (iff #473 #497)
+#488 := (+ #412 #483)
+#491 := (= uf_11 #488)
+#498 := (iff #491 #497)
+#499 := [rewrite]: #498
+#492 := (iff #473 #491)
+#489 := (= #468 #488)
+#486 := (= #462 #483)
+#480 := (ite #477 #459 uf_10)
+#484 := (= #480 #483)
+#485 := [rewrite]: #484
+#481 := (= #462 #480)
+#478 := (iff #56 #477)
+#479 := [rewrite]: #478
+#482 := [monotonicity #479]: #481
+#487 := [trans #482 #485]: #486
+#490 := [monotonicity #487]: #489
+#493 := [monotonicity #490]: #492
+#501 := [trans #493 #499]: #500
+#474 := (iff #60 #473)
+#471 := (= #59 #468)
+#465 := (- #462 uf_9)
+#469 := (= #465 #468)
+#470 := [rewrite]: #469
+#466 := (= #59 #465)
+#463 := (= #58 #462)
+#460 := (= #57 #459)
+#461 := [rewrite]: #460
+#464 := [monotonicity #461]: #463
+#467 := [monotonicity #464]: #466
+#472 := [trans #467 #470]: #471
+#475 := [monotonicity #472]: #474
+#503 := [trans #475 #501]: #502
+#458 := [asserted]: #60
+#504 := [mp #458 #503]: #497
+#628 := [mp #504 #627]: #625
+#629 := [not-or-elim #628]: #620
+#876 := [mp~ #629 #875]: #873
+#1197 := (+ uf_10 z3name!8)
+#1198 := (<= #1197 0::int)
+#1199 := (>= #1197 0::int)
+#1201 := (not #1199)
+#1200 := (not #1198)
+#1202 := (or #1200 #1201)
+#1203 := (not #1202)
+#1206 := (or #476 #1203)
+#1213 := (not #1206)
+#1186 := (+ uf_10 #864)
+#1187 := (<= #1186 0::int)
+#1189 := (not #1187)
+#1185 := (>= #1186 0::int)
+#1188 := (not #1185)
+#1190 := (or #1188 #1189)
+#1191 := (not #1190)
+#1194 := (or #477 #1191)
+#1212 := (not #1194)
+#1214 := (or #1212 #1213)
+#1215 := (not #1214)
+#1218 := (iff #860 #1215)
+#1209 := (and #1194 #1206)
+#1216 := (iff #1209 #1215)
+#1217 := [rewrite]: #1216
+#1210 := (iff #860 #1209)
+#1207 := (iff #859 #1206)
+#1204 := (iff #858 #1203)
+#1205 := [rewrite]: #1204
+#1208 := [monotonicity #1205]: #1207
+#1195 := (iff #857 #1194)
+#1192 := (iff #856 #1191)
+#1193 := [rewrite]: #1192
+#1196 := [monotonicity #1193]: #1195
+#1211 := [monotonicity #1196 #1208]: #1210
+#1219 := [trans #1211 #1217]: #1218
+#1220 := [mp #861 #1219]: #1215
+#1222 := [not-or-elim #1220]: #1206
+#1560 := [unit-resolution #1222 #1553]: #1203
+#1325 := (or #1202 #1198)
+#1326 := [def-axiom]: #1325
+#1561 := [unit-resolution #1326 #1560]: #1198
+#711 := (+ uf_5 #708)
+#714 := (+ uf_1 #711)
+#721 := (>= #714 0::int)
+#201 := (ite #194 uf_4 #177)
+#212 := (* -1::int #201)
+#213 := (+ uf_5 #212)
+#214 := (+ uf_1 #213)
+#555 := (>= #214 0::int)
+#722 := (= #555 #721)
+#715 := (~ #214 #714)
+#712 := (~ #213 #711)
+#709 := (~ #212 #708)
+#706 := (~ #201 z3name!2)
+#707 := [apply-def #705]: #706
+#710 := [monotonicity #707]: #709
+#713 := [monotonicity #710]: #712
+#716 := [monotonicity #713]: #715
+#723 := [monotonicity #716]: #722
+#557 := (not #555)
+#554 := (<= #214 0::int)
+#556 := (not #554)
+#558 := (or #556 #557)
+#559 := (not #558)
+#215 := (= #214 0::int)
+#560 := (iff #215 #559)
+#561 := [rewrite]: #560
+#21 := (- uf_4)
+#20 := (< uf_4 0::int)
+#22 := (ite #20 #21 uf_4)
+#23 := (- #22 uf_1)
+#24 := (= uf_5 #23)
+#220 := (iff #24 #215)
+#180 := (ite #20 #177 uf_4)
+#186 := (+ #130 #180)
+#191 := (= uf_5 #186)
+#218 := (iff #191 #215)
+#206 := (+ #130 #201)
+#209 := (= uf_5 #206)
+#216 := (iff #209 #215)
+#217 := [rewrite]: #216
+#210 := (iff #191 #209)
+#207 := (= #186 #206)
+#204 := (= #180 #201)
+#198 := (ite #195 #177 uf_4)
+#202 := (= #198 #201)
+#203 := [rewrite]: #202
+#199 := (= #180 #198)
+#196 := (iff #20 #195)
+#197 := [rewrite]: #196
+#200 := [monotonicity #197]: #199
+#205 := [trans #200 #203]: #204
+#208 := [monotonicity #205]: #207
+#211 := [monotonicity #208]: #210
+#219 := [trans #211 #217]: #218
+#192 := (iff #24 #191)
+#189 := (= #23 #186)
+#183 := (- #180 uf_1)
+#187 := (= #183 #186)
+#188 := [rewrite]: #187
+#184 := (= #23 #183)
+#181 := (= #22 #180)
+#178 := (= #21 #177)
+#179 := [rewrite]: #178
+#182 := [monotonicity #179]: #181
+#185 := [monotonicity #182]: #184
+#190 := [trans #185 #188]: #189
+#193 := [monotonicity #190]: #192
+#221 := [trans #193 #219]: #220
+#176 := [asserted]: #24
+#222 := [mp #176 #221]: #215
+#562 := [mp #222 #561]: #559
+#564 := [not-or-elim #562]: #555
+#724 := [mp~ #564 #723]: #721
+#685 := (+ uf_4 #682)
+#688 := (+ uf_2 #685)
+#695 := (>= #688 0::int)
+#154 := (ite #147 uf_1 #130)
+#165 := (* -1::int #154)
+#166 := (+ uf_4 #165)
+#167 := (+ uf_2 #166)
+#544 := (>= #167 0::int)
+#696 := (= #544 #695)
+#689 := (~ #167 #688)
+#686 := (~ #166 #685)
+#683 := (~ #165 #682)
+#680 := (~ #154 z3name!1)
+#681 := [apply-def #679]: #680
+#684 := [monotonicity #681]: #683
+#687 := [monotonicity #684]: #686
+#690 := [monotonicity #687]: #689
+#697 := [monotonicity #690]: #696
+#546 := (not #544)
+#543 := (<= #167 0::int)
+#545 := (not #543)
+#547 := (or #545 #546)
+#548 := (not #547)
+#168 := (= #167 0::int)
+#549 := (iff #168 #548)
+#550 := [rewrite]: #549
+#15 := (- uf_1)
+#14 := (< uf_1 0::int)
+#16 := (ite #14 #15 uf_1)
+#17 := (- #16 uf_2)
+#18 := (= uf_4 #17)
+#173 := (iff #18 #168)
+#133 := (ite #14 #130 uf_1)
+#139 := (+ #82 #133)
+#144 := (= uf_4 #139)
+#171 := (iff #144 #168)
+#159 := (+ #82 #154)
+#162 := (= uf_4 #159)
+#169 := (iff #162 #168)
+#170 := [rewrite]: #169
+#163 := (iff #144 #162)
+#160 := (= #139 #159)
+#157 := (= #133 #154)
+#151 := (ite #148 #130 uf_1)
+#155 := (= #151 #154)
+#156 := [rewrite]: #155
+#152 := (= #133 #151)
+#149 := (iff #14 #148)
+#150 := [rewrite]: #149
+#153 := [monotonicity #150]: #152
+#158 := [trans #153 #156]: #157
+#161 := [monotonicity #158]: #160
+#164 := [monotonicity #161]: #163
+#172 := [trans #164 #170]: #171
+#145 := (iff #18 #144)
+#142 := (= #17 #139)
+#136 := (- #133 uf_2)
+#140 := (= #136 #139)
+#141 := [rewrite]: #140
+#137 := (= #17 #136)
+#134 := (= #16 #133)
+#131 := (= #15 #130)
+#132 := [rewrite]: #131
+#135 := [monotonicity #132]: #134
+#138 := [monotonicity #135]: #137
+#143 := [trans #138 #141]: #142
+#146 := [monotonicity #143]: #145
+#174 := [trans #146 #172]: #173
+#129 := [asserted]: #18
+#175 := [mp #129 #174]: #168
+#551 := [mp #175 #550]: #548
+#553 := [not-or-elim #551]: #544
+#698 := [mp~ #553 #697]: #695
+#1373 := (not #721)
+#1562 := (or #147 #1373 #961 #241 #195)
+#1563 := [th-lemma]: #1562
+#1564 := [unit-resolution #1563 #1559 #1557 #724 #1406]: #147
+#1565 := [unit-resolution #955 #1564]: #925
+#1566 := [unit-resolution #1238 #1565]: #921
+#1372 := (not #873)
+#1371 := (not #695)
+#1498 := (or #516 #923 #1373 #1371 #1372 #1343 #1200 #1075 #1350 #1113 #961 #1151 #1337 #1048 #1357)
+#1499 := [th-lemma]: #1498
+#1567 := [unit-resolution #1499 #1566 #698 #724 #1545 #772 #1447 #802 #1496 #828 #1554 #854 #1561 #876 #1559]: #516
+#1247 := (or #962 #957)
+#1248 := [def-axiom]: #1247
+#1568 := [unit-resolution #1248 #1558]: #957
+#877 := (>= #870 0::int)
+#878 := (= #621 #877)
+#879 := [monotonicity #872]: #878
+#630 := [not-or-elim #628]: #621
+#880 := [mp~ #630 #879]: #877
+#1327 := (or #1202 #1199)
+#1328 := [def-axiom]: #1327
+#1569 := [unit-resolution #1328 #1560]: #1199
+#795 := (<= #792 0::int)
+#796 := (= #587 #795)
+#797 := [monotonicity #794]: #796
+#596 := [not-or-elim #595]: #587
+#798 := [mp~ #596 #797]: #795
+#1503 := (or #335 #1049 #241)
+#1425 := (or #335 #1049 #241 #1037)
+#1422 := [hypothesis]: #336
+#1423 := [hypothesis]: #1047
+#1424 := [th-lemma #1423 #1422 #776 #1406 #1421]: false
+#1426 := [lemma #1424]: #1425
+#1504 := [unit-resolution #1426 #1459]: #1503
+#1505 := [unit-resolution #1504 #1406 #1468]: #335
+#1506 := [unit-resolution #1107 #1505]: #1077
+#1283 := (or #1076 #1071)
+#1284 := [def-axiom]: #1283
+#1507 := [unit-resolution #1284 #1506]: #1071
+#717 := (<= #714 0::int)
+#718 := (= #554 #717)
+#719 := [monotonicity #716]: #718
+#563 := [not-or-elim #562]: #554
+#720 := [mp~ #563 #719]: #717
+#691 := (<= #688 0::int)
+#692 := (= #543 #691)
+#693 := [monotonicity #690]: #692
+#552 := [not-or-elim #551]: #543
+#694 := [mp~ #552 #693]: #691
+#1235 := (or #924 #919)
+#1236 := [def-axiom]: #1235
+#1570 := [unit-resolution #1236 #1565]: #919
+#1409 := (not #773)
+#1489 := (not #847)
+#1358 := (not #795)
+#1365 := (not #821)
+#1511 := (not #877)
+#1510 := (not #691)
+#1509 := (not #717)
+#1512 := (or #515 #922 #1509 #1510 #1511 #1365 #1201 #1074 #1358 #1112 #960 #1150 #1489 #1049 #1409)
+#1513 := [th-lemma]: #1512
+#1571 := [unit-resolution #1513 #1570 #694 #720 #1468 #776 #1507 #798 #1488 #824 #1552 #850 #1569 #880 #1568]: #515
+#506 := (<= #508 0::int)
+#659 := (+ uf_3 #656)
+#662 := (+ uf_1 #659)
+#665 := (<= #662 0::int)
+#107 := (ite #100 uf_2 #82)
+#118 := (* -1::int #107)
+#119 := (+ uf_3 #118)
+#120 := (+ uf_1 #119)
+#532 := (<= #120 0::int)
+#666 := (= #532 #665)
+#663 := (~ #120 #662)
+#660 := (~ #119 #659)
+#657 := (~ #118 #656)
+#654 := (~ #107 z3name!0)
+#655 := [apply-def #653]: #654
+#658 := [monotonicity #655]: #657
+#661 := [monotonicity #658]: #660
+#664 := [monotonicity #661]: #663
+#667 := [monotonicity #664]: #666
+#533 := (>= #120 0::int)
+#535 := (not #533)
+#534 := (not #532)
+#536 := (or #534 #535)
+#537 := (not #536)
+#121 := (= #120 0::int)
+#538 := (iff #121 #537)
+#539 := [rewrite]: #538
+#8 := (- uf_2)
+#7 := (< uf_2 0::int)
+#9 := (ite #7 #8 uf_2)
+#11 := (- #9 uf_3)
+#12 := (= uf_1 #11)
+#126 := (iff #12 #121)
+#85 := (ite #7 #82 uf_2)
+#91 := (* -1::int uf_3)
+#92 := (+ #91 #85)
+#97 := (= uf_1 #92)
+#124 := (iff #97 #121)
+#112 := (+ #91 #107)
+#115 := (= uf_1 #112)
+#122 := (iff #115 #121)
+#123 := [rewrite]: #122
+#116 := (iff #97 #115)
+#113 := (= #92 #112)
+#110 := (= #85 #107)
+#104 := (ite #101 #82 uf_2)
+#108 := (= #104 #107)
+#109 := [rewrite]: #108
+#105 := (= #85 #104)
+#102 := (iff #7 #101)
+#103 := [rewrite]: #102
+#106 := [monotonicity #103]: #105
+#111 := [trans #106 #109]: #110
+#114 := [monotonicity #111]: #113
+#117 := [monotonicity #114]: #116
+#125 := [trans #117 #123]: #124
+#98 := (iff #12 #97)
+#95 := (= #11 #92)
+#88 := (- #85 uf_3)
+#93 := (= #88 #92)
+#94 := [rewrite]: #93
+#89 := (= #11 #88)
+#86 := (= #9 #85)
+#83 := (= #8 #82)
+#84 := [rewrite]: #83
+#87 := [monotonicity #84]: #86
+#90 := [monotonicity #87]: #89
+#96 := [trans #90 #94]: #95
+#99 := [monotonicity #96]: #98
+#127 := [trans #99 #125]: #126
+#80 := [asserted]: #12
+#128 := [mp #80 #127]: #121
+#540 := [mp #128 #539]: #537
+#541 := [not-or-elim #540]: #532
+#668 := [mp~ #541 #667]: #665
+#1515 := (or #100 #241 #923 #1373 #1371 #961)
+#1516 := [th-lemma]: #1515
+#1572 := [unit-resolution #1516 #1566 #698 #1559 #724 #1406]: #100
+#1573 := [unit-resolution #917 #1572]: #887
+#1223 := (or #886 #881)
+#1224 := [def-axiom]: #1223
+#1574 := [unit-resolution #1224 #1573]: #881
+#1528 := (not #665)
+#1529 := (or #506 #884 #1528 #1364 #1011 #1343 #1113 #1151 #1337 #1048 #1357 #922 #1510)
+#1530 := [th-lemma]: #1529
+#1575 := [unit-resolution #1530 #1574 #668 #694 #1404 #750 #1545 #772 #1496 #828 #1554 #854 #1570]: #506
+#743 := (<= #740 0::int)
+#744 := (= #565 #743)
+#745 := [monotonicity #742]: #744
+#574 := [not-or-elim #573]: #565
+#746 := [mp~ #574 #745]: #743
+#1520 := [unit-resolution #1032 #1406]: #1013
+#1265 := (or #1012 #1008)
+#1266 := [def-axiom]: #1265
+#1521 := [unit-resolution #1266 #1520]: #1008
+#669 := (>= #662 0::int)
+#670 := (= #533 #669)
+#671 := [monotonicity #664]: #670
+#542 := [not-or-elim #540]: #533
+#672 := [mp~ #542 #671]: #669
+#1576 := [unit-resolution #1226 #1573]: #883
+#1523 := (not #743)
+#1522 := (not #669)
+#1524 := (or #509 #885 #1522 #1523 #1010 #1365 #1112 #1150 #1489 #1049 #1409 #923 #1371)
+#1525 := [th-lemma]: #1524
+#1577 := [unit-resolution #1525 #1576 #672 #698 #1521 #746 #1468 #776 #1488 #824 #1552 #850 #1566]: #509
+#634 := (not #516)
+#633 := (not #515)
+#632 := (not #509)
+#631 := (not #506)
+#635 := (or #631 #632 #633 #634)
+#523 := (and #506 #509 #515 #516)
+#528 := (not #523)
+#644 := (iff #528 #635)
+#636 := (not #635)
+#639 := (not #636)
+#642 := (iff #639 #635)
+#643 := [rewrite]: #642
+#640 := (iff #528 #639)
+#637 := (iff #523 #636)
+#638 := [rewrite]: #637
+#641 := [monotonicity #638]: #640
+#645 := [trans #641 #643]: #644
+#62 := (= uf_2 uf_11)
+#61 := (= uf_3 uf_10)
+#63 := (and #61 #62)
+#64 := (not #63)
+#529 := (iff #64 #528)
+#526 := (iff #63 #523)
+#517 := (and #515 #516)
+#510 := (and #506 #509)
+#520 := (and #510 #517)
+#524 := (iff #520 #523)
+#525 := [rewrite]: #524
+#521 := (iff #63 #520)
+#518 := (iff #62 #517)
+#519 := [rewrite]: #518
+#511 := (iff #61 #510)
+#512 := [rewrite]: #511
+#522 := [monotonicity #512 #519]: #521
+#527 := [trans #522 #525]: #526
+#530 := [monotonicity #527]: #529
+#505 := [asserted]: #64
+#531 := [mp #505 #530]: #528
+#646 := [mp #531 #645]: #635
+#1578 := [unit-resolution #646 #1577 #1575 #1571 #1567]: false
+#1580 := [lemma #1578]: #1579
+#1657 := [unit-resolution #1580 #1406]: #1048
+#1625 := (or #194 #241)
+#1535 := [hypothesis]: #195
+#1538 := (or #194 #960)
+#1432 := [hypothesis]: #973
+#1255 := (or #974 #971)
+#1256 := [def-axiom]: #1255
+#1433 := [unit-resolution #1256 #1432]: #974
+#994 := [not-or-elim #992]: #978
+#1434 := [unit-resolution #994 #1433]: #194
+#1435 := [unit-resolution #993 #1434]: #963
+#1436 := (or #971 #195 #961)
+#1437 := [th-lemma]: #1436
+#1438 := [unit-resolution #1437 #1434 #1432]: #961
+#1439 := [unit-resolution #1250 #1438 #1435]: false
+#1440 := [lemma #1439]: #971
+#1536 := [hypothesis]: #957
+#1537 := [th-lemma #1536 #1535 #1440]: false
+#1539 := [lemma #1537]: #1538
+#1581 := [unit-resolution #1539 #1535]: #960
+#1582 := (or #959 #957)
+#1583 := [th-lemma]: #1582
+#1584 := [unit-resolution #1583 #1581]: #959
+#1585 := (or #147 #1373 #241 #194 #973)
+#1586 := [th-lemma]: #1585
+#1587 := [unit-resolution #1586 #1535 #1440 #724 #1406]: #147
+#1588 := [unit-resolution #955 #1587]: #925
+#1589 := [unit-resolution #1238 #1588]: #921
+#1590 := [unit-resolution #1516 #1589 #698 #1584 #724 #1406]: #100
+#1591 := [unit-resolution #917 #1590]: #887
+#1592 := [unit-resolution #1224 #1591]: #881
+#1593 := (or #430 #1365 #1074 #1358 #1112 #194 #1364 #1011 #241)
+#1594 := [th-lemma]: #1593
+#1595 := [unit-resolution #1594 #1535 #1404 #750 #1507 #798 #1488 #824 #1406]: #430
+#1184 := [not-or-elim #1182]: #1168
+#1596 := [unit-resolution #1184 #1595]: #1165
+#1315 := (or #1164 #1161)
+#1316 := [def-axiom]: #1315
+#1597 := [unit-resolution #1316 #1596]: #1161
+#1533 := (or #288 #241)
+#1471 := (or #194 #288 #241)
+#1469 := (or #194 #288 #241 #1364 #1011)
+#1470 := [th-lemma]: #1469
+#1472 := [unit-resolution #1470 #1404 #750]: #1471
+#1473 := [unit-resolution #1472 #1405 #1406]: #194
+#1474 := [unit-resolution #993 #1473]: #963
+#1475 := [unit-resolution #1250 #1474]: #959
+#1476 := (or #147 #1373 #1364 #1011 #961 #241 #288)
+#1477 := [th-lemma]: #1476
+#1478 := [unit-resolution #1477 #1475 #724 #1406 #1404 #750 #1405]: #147
+#1479 := [unit-resolution #955 #1478]: #925
+#1480 := [unit-resolution #1238 #1479]: #921
+#1419 := (or #288 #241 #429)
+#1333 := [hypothesis]: #430
+#1408 := [unit-resolution #1280 #1407]: #1047
+#1410 := (or #335 #1049 #1409 #288 #241)
+#1411 := [th-lemma]: #1410
+#1412 := [unit-resolution #1411 #1405 #1408 #776 #1406]: #335
+#1413 := [unit-resolution #1107 #1412]: #1077
+#1414 := [unit-resolution #1286 #1413]: #1073
+#1415 := [unit-resolution #1352 #1414 #802 #1405 #828 #1333]: #1113
+#1416 := [unit-resolution #1298 #1415]: #1114
+#1417 := [unit-resolution #1145 #1416]: #383
+#1418 := [th-lemma #1414 #802 #1405 #1408 #776 #1406 #1417]: false
+#1420 := [lemma #1418]: #1419
+#1481 := [unit-resolution #1420 #1405 #1406]: #429
+#1482 := [unit-resolution #1183 #1481]: #1153
+#1483 := [unit-resolution #1308 #1482]: #1147
+#1490 := (or #477 #1150 #1489 #1365 #1112 #1049 #241 #1409 #288)
+#1491 := [th-lemma]: #1490
+#1492 := [unit-resolution #1491 #1405 #1468 #776 #1488 #824 #1483 #850 #1406]: #477
+#1493 := [unit-resolution #1222 #1492]: #1203
+#1494 := [unit-resolution #1326 #1493]: #1198
+#1495 := [unit-resolution #1310 #1482]: #1149
+#1500 := [unit-resolution #1499 #1475 #698 #724 #1497 #772 #1447 #802 #1496 #828 #1495 #854 #1494 #876 #1480]: #516
+#1501 := [unit-resolution #1236 #1479]: #919
+#1502 := [unit-resolution #1328 #1493]: #1199
+#1508 := [unit-resolution #1248 #1474]: #957
+#1514 := [unit-resolution #1513 #1508 #694 #720 #1468 #776 #1507 #798 #1488 #824 #1483 #850 #1502 #880 #1501]: #515
+#1517 := [unit-resolution #1516 #1480 #698 #1475 #724 #1406]: #100
+#1518 := [unit-resolution #917 #1517]: #887
+#1519 := [unit-resolution #1226 #1518]: #883
+#1526 := [unit-resolution #1525 #1480 #672 #698 #1521 #746 #1468 #776 #1488 #824 #1483 #850 #1519]: #509
+#1527 := [unit-resolution #1224 #1518]: #881
+#1531 := [unit-resolution #1530 #1501 #668 #694 #1404 #750 #1497 #772 #1496 #828 #1495 #854 #1527]: #506
+#1532 := [unit-resolution #646 #1531 #1526 #1514 #1500]: false
+#1534 := [lemma #1532]: #1533
+#1598 := [unit-resolution #1534 #1406]: #288
+#1599 := [unit-resolution #1069 #1598]: #1039
+#1271 := (or #1038 #1033)
+#1272 := [def-axiom]: #1271
+#1600 := [unit-resolution #1272 #1599]: #1033
+#1601 := [unit-resolution #1236 #1588]: #919
+#1602 := (or #506 #884 #1528 #1364 #1011 #1365 #1112 #1337 #1357 #922 #1510 #1036 #1163 #1074 #1358)
+#1603 := [th-lemma]: #1602
+#1604 := [unit-resolution #1603 #1601 #668 #694 #1404 #750 #1600 #772 #1507 #798 #1488 #824 #1597 #854 #1592]: #506
+#1605 := [unit-resolution #1226 #1591]: #883
+#1313 := (or #1164 #1160)
+#1314 := [def-axiom]: #1313
+#1606 := [unit-resolution #1314 #1596]: #1160
+#1607 := (or #509 #885 #1522 #1523 #1010 #1343 #1113 #1489 #1409 #923 #1371 #1037 #1162 #1075 #1350)
+#1608 := [th-lemma]: #1607
+#1609 := [unit-resolution #1608 #1589 #672 #698 #1521 #746 #1459 #776 #1447 #802 #1496 #828 #1606 #850 #1605]: #509
+#1610 := (or #476 #1036 #1337 #1365 #1112 #1357 #194 #1364 #1011 #1163 #1074 #1358)
+#1611 := [th-lemma]: #1610
+#1612 := [unit-resolution #1611 #1597 #750 #1600 #772 #1507 #798 #1488 #824 #1404 #854 #1535]: #476
+#1221 := [not-or-elim #1220]: #1194
+#1613 := [unit-resolution #1221 #1612]: #1191
+#1319 := (or #1190 #1185)
+#1320 := [def-axiom]: #1319
+#1614 := [unit-resolution #1320 #1613]: #1185
+#1615 := (or #516 #923 #1373 #1371 #1372 #1075 #1350 #1489 #1409 #1037 #973 #1162 #1188 #1343 #1113 #1523 #1010)
+#1616 := [th-lemma]: #1615
+#1617 := [unit-resolution #1616 #1606 #1440 #724 #1521 #746 #1459 #776 #1447 #802 #1496 #828 #698 #850 #1614 #876 #1589]: #516
+#1321 := (or #1190 #1187)
+#1322 := [def-axiom]: #1321
+#1618 := [unit-resolution #1322 #1613]: #1187
+#1619 := [unit-resolution #994 #1535]: #975
+#1253 := (or #974 #970)
+#1254 := [def-axiom]: #1253
+#1620 := [unit-resolution #1254 #1619]: #970
+#1621 := (or #515 #922 #1509 #1510 #1511 #1074 #1358 #1337 #1357 #1036 #972 #1163 #1189 #1365 #1112 #1364 #1011)
+#1622 := [th-lemma]: #1621
+#1623 := [unit-resolution #1622 #1620 #694 #720 #1404 #750 #1600 #772 #1507 #798 #1488 #824 #1597 #854 #1618 #880 #1601]: #515
+#1624 := [unit-resolution #646 #1623 #1617 #1609 #1604]: false
+#1626 := [lemma #1624]: #1625
+#1658 := [unit-resolution #1626 #1406]: #194
+#1659 := [unit-resolution #993 #1658]: #963
+#1660 := [unit-resolution #1250 #1659]: #959
+#1661 := [unit-resolution #1563 #1660 #1658 #724 #1406]: #147
+#1662 := [unit-resolution #955 #1661]: #925
+#1663 := [unit-resolution #1238 #1662]: #921
+#1664 := [unit-resolution #1516 #1663 #698 #1660 #724 #1406]: #100
+#1665 := [unit-resolution #917 #1664]: #887
+#1666 := [unit-resolution #1226 #1665]: #883
+#1667 := [unit-resolution #1224 #1665]: #881
+#1668 := [unit-resolution #1236 #1662]: #919
+#1669 := [unit-resolution #1248 #1659]: #957
+#1655 := (or #429 #1113 #1010 #960 #1036 #1074 #1112 #922 #923 #884 #885)
+#1632 := [hypothesis]: #919
+#1636 := [hypothesis]: #881
+#1638 := [hypothesis]: #1071
+#1639 := [hypothesis]: #1033
+#1334 := [unit-resolution #1184 #1333]: #1165
+#1335 := [unit-resolution #1316 #1334]: #1161
+#1640 := [unit-resolution #1603 #1335 #668 #694 #1404 #750 #1639 #772 #1638 #798 #1637 #824 #1632 #854 #1636]: #506
+#1641 := [hypothesis]: #883
+#1642 := [hypothesis]: #921
+#1643 := [hypothesis]: #1111
+#1644 := [hypothesis]: #1008
+#1631 := [unit-resolution #1314 #1334]: #1160
+#1645 := [unit-resolution #1608 #1631 #672 #698 #1644 #746 #1459 #776 #1447 #802 #1643 #828 #1642 #850 #1641]: #509
+#1634 := (or #1202 #922 #960 #632 #631 #429)
+#1627 := [hypothesis]: #506
+#1628 := [hypothesis]: #509
+#1384 := [hypothesis]: #1203
+#1396 := (or #1202 #516 #429)
+#1331 := [hypothesis]: #634
+#1385 := [unit-resolution #1326 #1384]: #1198
+#1382 := (or #1189 #1200 #516 #429)
+#1332 := [hypothesis]: #1198
+#1336 := [hypothesis]: #1187
+#1338 := (or #382 #1189 #1337 #429 #1163 #1200)
+#1339 := [th-lemma]: #1338
+#1340 := [unit-resolution #1339 #1336 #1335 #854 #1333 #1332]: #382
+#1341 := [unit-resolution #1145 #1340]: #1115
+#1342 := [unit-resolution #1298 #1341]: #1111
+#1344 := (or #335 #1113 #429 #1343 #1189 #1337 #1163 #1200)
+#1345 := [th-lemma]: #1344
+#1346 := [unit-resolution #1345 #1342 #828 #1333 #1335 #854 #1336 #1332]: #335
+#1347 := [unit-resolution #1107 #1346]: #1077
+#1348 := [unit-resolution #1284 #1347]: #1071
+#1349 := [unit-resolution #1286 #1347]: #1073
+#1353 := [unit-resolution #1352 #1349 #802 #1342 #828 #1333]: #288
+#1354 := [unit-resolution #1069 #1353]: #1039
+#1355 := [unit-resolution #1272 #1354]: #1033
+#1356 := [unit-resolution #1296 #1341]: #1109
+#1359 := (or #242 #1036 #1357 #429 #1189 #1337 #1163 #1200 #1074 #1358)
+#1360 := [th-lemma]: #1359
+#1361 := [unit-resolution #1360 #1355 #772 #1348 #798 #1333 #1335 #854 #1336 #1332]: #242
+#1362 := [unit-resolution #1032 #1361]: #1013
+#1363 := [unit-resolution #1268 #1362]: #1009
+#1366 := (or #194 #1011 #1364 #1074 #1358 #1112 #1365 #1036 #1357 #1189 #1337 #1163 #1200)
+#1367 := [th-lemma]: #1366
+#1368 := [unit-resolution #1367 #1363 #750 #1355 #772 #1348 #798 #1356 #824 #1335 #854 #1336 #1332]: #194
+#1369 := [unit-resolution #993 #1368]: #963
+#1370 := [unit-resolution #1250 #1369]: #959
+#1374 := (or #923 #1371 #516 #1372 #1200 #961 #1373 #1036 #1357 #1337 #1163 #1074 #1358)
+#1375 := [th-lemma]: #1374
+#1376 := [unit-resolution #1375 #1370 #698 #724 #1355 #772 #1348 #798 #1335 #854 #1332 #876 #1331]: #923
+#1377 := (or #147 #195 #961 #1373 #1036 #1357 #429 #1189 #1337 #1163 #1200 #1074 #1358)
+#1378 := [th-lemma]: #1377
+#1379 := [unit-resolution #1378 #1368 #1370 #724 #1355 #772 #1348 #798 #1333 #1335 #854 #1336 #1332]: #147
+#1380 := [unit-resolution #955 #1379]: #925
+#1381 := [unit-resolution #1238 #1380 #1376]: false
+#1383 := [lemma #1381]: #1382
+#1386 := [unit-resolution #1383 #1385 #1331 #1333]: #1189
+#1387 := [unit-resolution #1322 #1386]: #1190
+#1388 := [unit-resolution #1328 #1384]: #1199
+#1389 := (or #1187 #1185)
+#1390 := [th-lemma]: #1389
+#1391 := [unit-resolution #1390 #1386]: #1185
+#1392 := (or #476 #1188 #1201)
+#1393 := [th-lemma]: #1392
+#1394 := [unit-resolution #1393 #1391 #1388]: #476
+#1395 := [unit-resolution #1221 #1394 #1387]: false
+#1397 := [lemma #1395]: #1396
+#1629 := [unit-resolution #1397 #1384 #1333]: #516
+#1630 := [unit-resolution #646 #1629 #1628 #1627]: #633
+#1633 := [th-lemma #1632 #720 #694 #880 #1447 #802 #850 #776 #1459 #1631 #1536 #1388 #1630]: false
+#1635 := [lemma #1633]: #1634
+#1646 := [unit-resolution #1635 #1645 #1536 #1632 #1640 #1333]: #1202
+#1647 := [unit-resolution #1222 #1646]: #476
+#1648 := [unit-resolution #1221 #1647]: #1191
+#1649 := [unit-resolution #1322 #1648]: #1187
+#1650 := [unit-resolution #1320 #1648]: #1185
+#1651 := [unit-resolution #1616 #1650 #1440 #724 #1644 #746 #1459 #776 #1447 #802 #1643 #828 #698 #850 #1631 #876 #1642]: #516
+#1652 := [unit-resolution #646 #1651 #1645 #1640]: #633
+#1653 := [unit-resolution #1622 #1652 #694 #720 #1404 #750 #1639 #772 #1638 #798 #1637 #824 #1335 #854 #1649 #880 #1632]: #972
+#1654 := [th-lemma #1459 #1647 #850 #828 #1643 #776 #746 #1644 #1631 #1447 #802 #1536 #1653]: false
+#1656 := [lemma #1654]: #1655
+#1670 := [unit-resolution #1656 #1496 #1521 #1669 #1600 #1507 #1488 #1668 #1663 #1667 #1666]: #429
+#1671 := [th-lemma #1600 #1670 #824 #1507 #798 #1488 #1657]: false
+#1672 := [lemma #1671]: #241
+#1683 := [unit-resolution #1031 #1672]: #1001
+#1703 := [unit-resolution #1262 #1683]: #997
+#1920 := (or #194 #242 #1364 #999 #288)
+#1921 := [th-lemma]: #1920
+#1922 := [unit-resolution #1921 #1405 #1703 #750 #1672]: #194
+#1923 := [unit-resolution #993 #1922]: #963
+#1924 := [unit-resolution #1248 #1923]: #957
+#1925 := [unit-resolution #1250 #1923]: #959
+#1843 := (or #288 #961 #147)
+#1763 := [hypothesis]: #148
+#1828 := [hypothesis]: #959
+#1842 := [th-lemma #724 #750 #1703 #1828 #1405 #1763]: false
+#1844 := [lemma #1842]: #1843
+#1926 := [unit-resolution #1844 #1925 #1405]: #147
+#1927 := [unit-resolution #955 #1926]: #925
+#1928 := [unit-resolution #1236 #1927]: #919
+#2116 := [unit-resolution #1310 #2047]: #1149
+#2084 := (or #288 #516)
+#2050 := (or #288 #961 #516)
+#2037 := [hypothesis]: #1087
+#2038 := [unit-resolution #1292 #2037]: #1088
+#2041 := (or #1085 #336)
+#2039 := (or #1085 #1075 #336)
+#2040 := [th-lemma]: #2039
+#2042 := [unit-resolution #2040 #1447]: #2041
+#2043 := [unit-resolution #2042 #2037]: #336
+#2044 := [unit-resolution #1108 #2043 #2038]: false
+#2045 := [lemma #2044]: #1085
+#2035 := (or #1087 #1150 #961 #1048 #516)
+#1845 := [hypothesis]: #1085
+#1874 := [hypothesis]: #477
+#1901 := (or #335 #476)
+#1895 := [unit-resolution #1222 #1874]: #1203
+#1896 := [unit-resolution #1326 #1895]: #1198
+#1893 := (or #429 #1200)
+#1880 := (or #335 #1113 #429 #1163 #1200)
+#1857 := [hypothesis]: #1189
+#1858 := [unit-resolution #1322 #1857]: #1190
+#1859 := [unit-resolution #1221 #1858]: #477
+#1860 := [unit-resolution #1222 #1859]: #1203
+#1861 := [unit-resolution #1390 #1857]: #1185
+#1862 := [unit-resolution #1393 #1859 #1861]: #1201
+#1863 := [unit-resolution #1328 #1862 #1860]: false
+#1864 := [lemma #1863]: #1187
+#1878 := (or #335 #1113 #429 #1189 #1163 #1200)
+#1879 := [unit-resolution #1345 #828 #854]: #1878
+#1881 := [unit-resolution #1879 #1864]: #1880
+#1882 := [unit-resolution #1881 #1335 #1870 #1333 #1332]: #335
+#1883 := [unit-resolution #1107 #1882]: #1077
+#1884 := [unit-resolution #1689 #1333 #1870]: #288
+#1885 := [unit-resolution #1069 #1884]: #1039
+#1886 := [unit-resolution #1272 #1885]: #1033
+#1889 := (or #1036 #429 #1163 #1200 #1074)
+#1887 := (or #242 #1036 #429 #1189 #1163 #1200 #1074)
+#1888 := [unit-resolution #1360 #772 #798 #854]: #1887
+#1890 := [unit-resolution #1888 #1672 #1864]: #1889
+#1891 := [unit-resolution #1890 #1886 #1332 #1333 #1335]: #1074
+#1892 := [unit-resolution #1284 #1891 #1883]: false
+#1894 := [lemma #1892]: #1893
+#1897 := [unit-resolution #1894 #1896]: #429
+#1898 := [unit-resolution #1183 #1897]: #1153
+#1899 := [unit-resolution #1310 #1898]: #1149
+#1900 := [th-lemma #854 #1899 #1870 #828 #1422 #1874]: false
+#1902 := [lemma #1900]: #1901
+#1950 := [unit-resolution #1902 #1874]: #335
+#1951 := [unit-resolution #1107 #1950]: #1077
+#1952 := [unit-resolution #1284 #1951]: #1071
+#1953 := [unit-resolution #1328 #1895]: #1199
+#1876 := (or #1109 #476)
+#1673 := [hypothesis]: #1112
+#1760 := (or #429 #1109)
+#1674 := [unit-resolution #1296 #1673]: #1114
+#1675 := [unit-resolution #1145 #1674]: #383
+#1676 := [unit-resolution #1146 #1675]: #1127
+#1677 := [unit-resolution #1304 #1676]: #1123
+#1687 := [unit-resolution #1686 #1673]: #1111
+#1743 := [unit-resolution #1689 #1333 #1687]: #288
+#1744 := [unit-resolution #1069 #1743]: #1039
+#1745 := [unit-resolution #1272 #1744]: #1033
+#1678 := (or #335 #1343 #429 #382 #1125)
+#1679 := [th-lemma]: #1678
+#1746 := [unit-resolution #1679 #1333 #1675 #828 #1677]: #335
+#1747 := [unit-resolution #1107 #1746]: #1077
+#1748 := [unit-resolution #1284 #1747]: #1071
+#1259 := (or #1000 #995)
+#1260 := [def-axiom]: #1259
+#1684 := [unit-resolution #1260 #1683]: #995
+#1693 := (or #147 #1373 #1343 #1074 #1358 #1523 #429 #973 #998 #1036 #1357 #1125)
+#1694 := [th-lemma]: #1693
+#1749 := [unit-resolution #1694 #1745 #724 #1684 #746 #1440 #772 #1748 #798 #1677 #828 #1333]: #147
+#1750 := [unit-resolution #955 #1749]: #925
+#1751 := [unit-resolution #1238 #1750]: #921
+#1714 := (or #100 #923 #1373 #1371 #1343 #1523 #1074 #1358 #973 #429 #382 #1036 #1357 #998 #1125)
+#1715 := [th-lemma]: #1714
+#1752 := [unit-resolution #1715 #1751 #698 #1440 #724 #1684 #746 #1675 #772 #1748 #798 #1745 #1677 #828 #1333]: #100
+#1753 := [unit-resolution #1236 #1750]: #919
+#1727 := (or #1109 #429 #972)
+#1680 := [unit-resolution #1679 #1675 #1677 #828 #1333]: #335
+#1681 := [unit-resolution #1107 #1680]: #1077
+#1682 := [unit-resolution #1284 #1681]: #1071
+#1690 := [unit-resolution #1689 #1687 #1333]: #288
+#1691 := [unit-resolution #1069 #1690]: #1039
+#1692 := [unit-resolution #1272 #1691]: #1033
+#1695 := [unit-resolution #1694 #1692 #724 #1684 #746 #1440 #772 #1682 #798 #1677 #828 #1333]: #147
+#1696 := [unit-resolution #955 #1695]: #925
+#1697 := [unit-resolution #1236 #1696]: #919
+#1698 := (or #476 #429 #1337 #1163 #382)
+#1699 := [th-lemma]: #1698
+#1700 := [unit-resolution #1699 #1675 #1335 #854 #1333]: #476
+#1701 := [unit-resolution #1221 #1700]: #1191
+#1702 := [unit-resolution #1322 #1701]: #1187
+#1704 := [hypothesis]: #970
+#1301 := (or #1126 #1122)
+#1302 := [def-axiom]: #1301
+#1705 := [unit-resolution #1302 #1676]: #1122
+#1706 := (or #515 #922 #1509 #1510 #1511 #1075 #1350 #1337 #1409 #1037 #1163 #1365 #1364 #972 #999 #1124 #1189)
+#1707 := [th-lemma]: #1706
+#1708 := [unit-resolution #1707 #1705 #1704 #720 #1703 #750 #1459 #776 #1447 #802 #694 #824 #1335 #854 #1702 #880 #1697]: #515
+#1709 := [unit-resolution #1238 #1696]: #921
+#1710 := [unit-resolution #1320 #1701]: #1185
+#1711 := (or #516 #923 #1373 #1371 #1372 #1074 #1358 #1489 #1357 #1036 #1162 #1343 #1523 #973 #998 #1125 #1188)
+#1712 := [th-lemma]: #1711
+#1713 := [unit-resolution #1712 #1692 #1440 #724 #1684 #746 #698 #772 #1682 #798 #1677 #828 #1631 #850 #1710 #876 #1709]: #516
+#1716 := [unit-resolution #1715 #1709 #698 #1440 #724 #1684 #746 #1692 #772 #1682 #798 #1675 #1677 #828 #1333]: #100
+#1717 := [unit-resolution #917 #1716]: #887
+#1718 := [unit-resolution #1226 #1717]: #883
+#1719 := (or #509 #885 #1522 #1523 #1343 #1489 #1357 #923 #1371 #1036 #1162 #998 #1125)
+#1720 := [th-lemma]: #1719
+#1721 := [unit-resolution #1720 #1709 #672 #698 #1684 #746 #1692 #772 #1677 #828 #1631 #850 #1718]: #509
+#1722 := [unit-resolution #1224 #1717]: #881
+#1723 := (or #506 #884 #1528 #1364 #1365 #1337 #1409 #922 #1510 #1037 #1163 #999 #1124)
+#1724 := [th-lemma]: #1723
+#1725 := [unit-resolution #1724 #1697 #668 #694 #1703 #750 #1459 #776 #1705 #824 #1335 #854 #1722]: #506
+#1726 := [unit-resolution #646 #1725 #1721 #1713 #1708]: false
+#1728 := [lemma #1726]: #1727
+#1754 := [unit-resolution #1728 #1333 #1673]: #972
+#1755 := [unit-resolution #1254 #1754]: #974
+#1756 := [unit-resolution #994 #1755]: #194
+#1757 := [unit-resolution #993 #1756]: #963
+#1758 := [unit-resolution #1248 #1757]: #957
+#1759 := [th-lemma #1758 #1753 #720 #694 #1675 #1459 #776 #1447 #802 #1752]: false
+#1761 := [lemma #1759]: #1760
+#1871 := [unit-resolution #1761 #1673]: #429
+#1872 := [unit-resolution #1183 #1871]: #1153
+#1873 := [unit-resolution #1310 #1872]: #1149
+#1875 := [th-lemma #1675 #1874 #854 #1873 #1871]: false
+#1877 := [lemma #1875]: #1876
+#1954 := [unit-resolution #1877 #1874]: #1109
+#1948 := (or #288 #1112 #1200 #1201 #1074)
+#1917 := [unit-resolution #1894 #1332]: #429
+#1918 := [unit-resolution #1183 #1917]: #1153
+#1919 := [unit-resolution #1308 #1918]: #1147
+#1929 := [unit-resolution #1310 #1918]: #1149
+#1930 := [unit-resolution #1238 #1927]: #921
+#1931 := [hypothesis]: #1199
+#1932 := (or #515 #922 #1201 #1074 #1112 #960 #1150)
+#1933 := [unit-resolution #1513 #694 #720 #1468 #776 #798 #824 #850 #880]: #1932
+#1934 := [unit-resolution #1933 #1928 #1931 #1637 #1638 #1919 #1924]: #515
+#1935 := (or #516 #923 #1200 #1113 #961 #1151 #1048)
+#1936 := [unit-resolution #1499 #698 #724 #772 #1447 #802 #828 #854 #876]: #1935
+#1937 := [unit-resolution #1936 #1930 #1870 #1332 #1929 #1497 #1925]: #516
+#1915 := (or #898 #634 #633 #923 #961 #1048 #1151 #922 #960 #1112 #1150)
+#1903 := [hypothesis]: #515
+#1904 := [hypothesis]: #516
+#1905 := [hypothesis]: #899
+#1906 := [unit-resolution #1232 #1905]: #895
+#1907 := (or #509 #1522 #1523 #897 #998 #1489 #1150 #960 #1509 #1112 #1365 #1049 #922 #1510 #1409)
+#1908 := [th-lemma]: #1907
+#1909 := [unit-resolution #1908 #1906 #1632 #694 #1536 #720 #1684 #746 #1468 #776 #1637 #824 #1540 #850 #672]: #509
+#1774 := [hypothesis]: #1149
+#1229 := (or #898 #894)
+#1230 := [def-axiom]: #1229
+#1910 := [unit-resolution #1230 #1905]: #894
+#1911 := (or #506 #1528 #1364 #896 #999 #1337 #1151 #961 #1373 #1113 #1343 #1048 #923 #1371 #1357)
+#1912 := [th-lemma]: #1911
+#1913 := [unit-resolution #1912 #1910 #1642 #698 #1828 #724 #1703 #750 #1545 #772 #1870 #828 #1774 #854 #668]: #506
+#1914 := [unit-resolution #646 #1913 #1909 #1904 #1903]: false
+#1916 := [lemma #1914]: #1915
+#1938 := [unit-resolution #1916 #1937 #1934 #1930 #1925 #1497 #1929 #1928 #1924 #1637 #1919]: #898
+#1939 := [unit-resolution #918 #1938]: #100
+#1940 := [unit-resolution #917 #1939]: #887
+#1941 := [unit-resolution #1224 #1940]: #881
+#1942 := (or #506 #884 #1113 #1151 #1048 #922)
+#1943 := [unit-resolution #1530 #668 #694 #1404 #750 #772 #828 #854]: #1942
+#1944 := [unit-resolution #1943 #1941 #1497 #1870 #1929 #1928]: #506
+#1945 := [unit-resolution #646 #1944 #1937 #1934]: #632
+#1946 := [unit-resolution #1908 #1945 #1928 #694 #1924 #720 #1684 #746 #1468 #776 #1637 #824 #1919 #850 #672]: #897
+#1947 := [th-lemma #1946 #1939 #1742]: false
+#1949 := [lemma #1947]: #1948
+#1955 := [unit-resolution #1949 #1954 #1896 #1953 #1952]: #288
+#1956 := [unit-resolution #1069 #1955]: #1039
+#1957 := [unit-resolution #1272 #1956]: #1033
+#1958 := [unit-resolution #1735 #1954]: #382
+#1959 := (or #1123 #383 #1113)
+#1960 := [th-lemma]: #1959
+#1961 := [unit-resolution #1960 #1958 #1870]: #1123
+#1962 := [unit-resolution #1308 #1898]: #1147
+#1965 := (or #1160 #1112 #1074 #289 #1150)
+#1963 := (or #1160 #1365 #1112 #1074 #1358 #289 #1150)
+#1964 := [th-lemma]: #1963
+#1966 := [unit-resolution #1964 #798 #824]: #1965
+#1967 := [unit-resolution #1966 #1955 #1954 #1962 #1952]: #1160
+#1970 := (or #1162 #1151 #1036 #1125 #147 #1074)
+#1968 := (or #1162 #1151 #1343 #1523 #998 #1036 #1357 #1125 #973 #147 #1373 #1074 #1358)
+#1969 := [th-lemma]: #1968
+#1971 := [unit-resolution #1969 #724 #1684 #746 #1440 #772 #798 #828]: #1970
+#1972 := [unit-resolution #1971 #1967 #1952 #1961 #1899 #1957]: #147
+#1973 := [unit-resolution #955 #1972]: #925
+#1974 := [unit-resolution #1236 #1973]: #919
+#1975 := (or #1161 #1151 #430)
+#1976 := [th-lemma]: #1975
+#1977 := [unit-resolution #1976 #1899 #1897]: #1161
+#1978 := (or #476 #1036 #1112 #194 #1163 #1074)
+#1979 := [unit-resolution #1611 #750 #772 #798 #824 #1404 #854]: #1978
+#1980 := [unit-resolution #1979 #1957 #1874 #1954 #1952 #1977]: #194
+#1981 := [unit-resolution #993 #1980]: #963
+#1982 := [unit-resolution #1248 #1981]: #957
+#1983 := [unit-resolution #1933 #1974 #1953 #1954 #1952 #1962 #1982]: #515
+#1984 := [unit-resolution #1238 #1973]: #921
+#1985 := [unit-resolution #1250 #1981]: #959
+#1849 := (or #923 #516 #1200 #961 #1036 #1163 #1074)
+#1850 := [unit-resolution #1375 #698 #724 #772 #798 #854 #876]: #1849
+#1986 := [unit-resolution #1850 #1985 #1896 #1952 #1977 #1957 #1984]: #516
+#1987 := (or #509 #923 #1036 #1162 #1125)
+#1988 := [unit-resolution #1720 #672 #698 #1684 #746 #1742 #772 #828 #850]: #1987
+#1989 := [unit-resolution #1988 #1984 #1961 #1967 #1957]: #509
+#1990 := [unit-resolution #646 #1989 #1986 #1983]: #631
+#1991 := (or #506 #884 #1112 #922 #1036 #1163 #1074)
+#1992 := [unit-resolution #1603 #668 #694 #1404 #750 #772 #798 #824 #854]: #1991
+#1993 := [unit-resolution #1992 #1990 #1977 #1954 #1952 #1957 #1974]: #884
+#1994 := [unit-resolution #1224 #1993]: #886
+#1995 := [unit-resolution #917 #1994]: #101
+#1996 := [th-lemma #746 #1684 #1957 #1874 #854 #1899 #1870 #828 #1984 #1995 #698 #772 #1972]: false
+#1997 := [lemma #1996]: #476
+#2014 := [unit-resolution #1221 #1997]: #1191
+#2015 := [unit-resolution #1320 #2014]: #1185
+#2034 := [th-lemma #876 #850 #1540 #2015 #802 #2033 #698 #772 #1828 #724 #1545 #1845 #1331]: false
+#2036 := [lemma #2034]: #2035
+#2048 := [unit-resolution #2036 #1497 #2045 #1828 #1331]: #1150
+#2049 := [unit-resolution #1308 #2048 #2047]: false
+#2051 := [lemma #2049]: #2050
+#2082 := [unit-resolution #2051 #1405 #1331]: #961
+#2083 := [unit-resolution #1250 #1923 #2082]: false
+#2085 := [lemma #2083]: #2084
+#2089 := [unit-resolution #2085 #1331]: #288
+#2090 := [unit-resolution #1069 #2089]: #1039
+#2091 := [unit-resolution #1272 #2090]: #1033
+#2065 := [hypothesis]: #935
+#2066 := [unit-resolution #1244 #2065]: #936
+#2067 := [unit-resolution #956 #2066]: #147
+#2068 := [th-lemma #2065 #2033 #2067]: false
+#2069 := [lemma #2068]: #933
+#2100 := (or #429 #516)
+#2063 := (or #429 #1086 #516)
+#2052 := [unit-resolution #1761 #1333]: #1109
+#2053 := [unit-resolution #1735 #2052]: #382
+#2054 := [hypothesis]: #1084
+#2055 := (or #1200 #516 #429)
+#2056 := [unit-resolution #1383 #1864]: #2055
+#2057 := [unit-resolution #2056 #1333 #1331]: #1200
+#2060 := (or #1086 #383 #1113 #1188 #1162 #1198)
+#2058 := (or #1086 #383 #1113 #1343 #1188 #1489 #1162 #1198 #1075)
+#2059 := [th-lemma]: #2058
+#2061 := [unit-resolution #2059 #1447 #828 #850]: #2060
+#2062 := [unit-resolution #2061 #1631 #2057 #2015 #1870 #2054 #2053]: false
+#2064 := [lemma #2062]: #2063
+#2086 := [unit-resolution #2064 #1333 #1331]: #1086
+#2087 := [unit-resolution #1290 #2086]: #1088
+#2088 := [unit-resolution #1108 #2087]: #335
+#2080 := (or #1109 #516)
+#2070 := [unit-resolution #1308 #1872]: #1147
+#2020 := (or #194 #1150 #516 #1125 #1151 #1124)
+#1762 := [hypothesis]: #1122
+#1775 := [hypothesis]: #1123
+#1803 := (or #194 #1151 #1150 #1125 #147 #1124)
+#1764 := [unit-resolution #956 #1763]: #937
+#1765 := [unit-resolution #1244 #1764]: #933
+#1766 := (or #509 #885 #1522 #1364 #1365 #1489 #999 #1124 #1371 #1037 #1409 #935 #1150 #972 #1509 #1075 #1350)
+#1767 := [th-lemma]: #1766
+#1768 := [unit-resolution #1767 #1620 #1765 #698 #672 #720 #1703 #750 #1459 #776 #1447 #802 #1762 #824 #1540 #850 #1742]: #509
+#1769 := (or #100 #1371 #935 #194 #147)
+#1770 := [th-lemma]: #1769
+#1771 := [unit-resolution #1770 #1535 #1765 #698 #1763]: #100
+#1772 := [unit-resolution #917 #1771]: #887
+#1773 := [unit-resolution #1224 #1772]: #881
+#1776 := (or #335 #194 #1364 #1037 #1409 #999)
+#1777 := [th-lemma]: #1776
+#1778 := [unit-resolution #1777 #1535 #750 #1459 #776 #1703]: #335
+#1779 := [unit-resolution #1107 #1778]: #1077
+#1780 := [unit-resolution #1284 #1779]: #1071
+#1241 := (or #936 #932)
+#1242 := [def-axiom]: #1241
+#1781 := [unit-resolution #1242 #1764]: #932
+#1782 := (or #288 #1364 #999 #973 #147 #1373 #194)
+#1783 := [th-lemma]: #1782
+#1784 := [unit-resolution #1783 #1535 #1440 #724 #1703 #750 #1763]: #288
+#1785 := [unit-resolution #1069 #1784]: #1039
+#1786 := [unit-resolution #1272 #1785]: #1033
+#1787 := (or #506 #884 #1528 #1523 #1343 #1337 #998 #1125 #1510 #1036 #1357 #934 #1151 #973 #1373 #1074 #1358)
+#1788 := [th-lemma]: #1787
+#1789 := [unit-resolution #1788 #1786 #1781 #694 #1440 #724 #1684 #746 #668 #772 #1780 #798 #1775 #828 #1774 #854 #1773]: #506
+#1790 := (or #476 #1337 #1343 #1523 #1036 #1357 #998 #1125 #973 #147 #1373 #1074 #1358 #1151 #194)
+#1791 := [th-lemma]: #1790
+#1792 := [unit-resolution #1791 #1535 #1440 #724 #1684 #746 #1786 #772 #1780 #798 #1775 #828 #1774 #854 #1763]: #476
+#1793 := [unit-resolution #1221 #1792]: #1191
+#1794 := [unit-resolution #1320 #1793]: #1185
+#1795 := (or #516 #1372 #1489 #1409 #1037 #1188 #1371 #935 #972 #1509 #1075 #1350 #1150)
+#1796 := [th-lemma]: #1795
+#1797 := [unit-resolution #1796 #1620 #698 #720 #1459 #776 #1447 #802 #1540 #850 #1794 #876 #1765]: #516
+#1798 := [unit-resolution #1322 #1793]: #1187
+#1799 := (or #515 #1511 #1337 #1357 #1036 #1189 #1510 #934 #973 #1373 #1074 #1358 #1151)
+#1800 := [th-lemma]: #1799
+#1801 := [unit-resolution #1800 #1786 #1440 #724 #694 #772 #1780 #798 #1774 #854 #1798 #880 #1781]: #515
+#1802 := [unit-resolution #646 #1801 #1797 #1789 #1768]: false
+#1804 := [lemma #1802]: #1803
+#2011 := [unit-resolution #1804 #1535 #1540 #1775 #1774 #1762]: #147
+#2012 := [unit-resolution #955 #2011]: #925
+#2013 := [unit-resolution #1238 #2012]: #921
+#2016 := (or #516 #1188 #935 #972 #1150)
+#2017 := [unit-resolution #1796 #698 #720 #1459 #776 #1447 #802 #850 #876]: #2016
+#2018 := [unit-resolution #2017 #1620 #2015 #1540 #1331]: #935
+#2019 := [th-lemma #2018 #2013 #2011]: false
+#2021 := [lemma #2019]: #2020
+#2071 := [unit-resolution #2021 #2070 #1331 #1677 #1873 #1705]: #194
+#2072 := [unit-resolution #993 #2071]: #963
+#2073 := [unit-resolution #2010 #1675]: #288
+#2074 := [unit-resolution #1069 #2073]: #1039
+#2075 := [unit-resolution #1272 #2074]: #1033
+#2076 := (or #516 #1036 #1188 #935 #1150 #960 #1087)
+#1823 := (or #516 #1372 #1489 #1357 #1036 #1188 #1371 #935 #1509 #1350 #1150 #960 #1523 #998 #1087)
+#1824 := [th-lemma]: #1823
+#2077 := [unit-resolution #1824 #720 #1684 #746 #698 #772 #802 #850 #876]: #2076
+#2078 := [unit-resolution #2077 #2075 #2015 #2045 #2069 #1331 #2070]: #960
+#2079 := [unit-resolution #1248 #2078 #2072]: false
+#2081 := [lemma #2079]: #2080
+#2092 := [unit-resolution #2081 #1331]: #1109
+#2093 := [unit-resolution #1735 #2092]: #382
+#2094 := [unit-resolution #1960 #2093 #1870]: #1123
+#2095 := (or #516 #923 #1074 #1036 #1162 #1125 #1188)
+#2096 := [unit-resolution #1712 #1440 #724 #1684 #746 #698 #772 #798 #828 #850 #876]: #2095
+#2097 := [unit-resolution #2096 #1631 #2015 #2094 #1331 #2091 #2033]: #1074
+#2098 := [unit-resolution #1284 #2097]: #1076
+#2099 := [unit-resolution #1107 #2098 #2088]: false
+#2101 := [lemma #2099]: #2100
+#2102 := [unit-resolution #2101 #1331]: #429
+#2103 := [unit-resolution #1183 #2102]: #1153
+#2104 := [unit-resolution #1308 #2103]: #1147
+#2105 := [unit-resolution #2077 #2104 #2015 #2045 #2069 #1331 #2091]: #960
+#2106 := [unit-resolution #1248 #2105]: #962
+#2107 := [unit-resolution #2017 #2104 #2015 #2069 #1331]: #972
+#2108 := [unit-resolution #1254 #2107]: #974
+#2109 := [unit-resolution #994 #2108]: #194
+#2110 := [unit-resolution #993 #2109 #2106]: false
+#2111 := [lemma #2110]: #516
+#2127 := (or #1199 #1189 #477)
+#2128 := [th-lemma]: #2127
+#2129 := [unit-resolution #2128 #1864 #1997]: #1199
+#2125 := (or #335 #288)
+#1806 := [unit-resolution #1108 #1422]: #1089
+#1829 := [unit-resolution #1290 #1806]: #1084
+#2117 := (or #515 #1511 #1337 #1151 #1189 #1358 #922 #1510 #1409 #960 #1509 #1049 #1086)
+#2118 := [th-lemma]: #2117
+#2119 := [unit-resolution #2118 #1829 #1924 #720 #1468 #776 #694 #798 #2116 #854 #1864 #880 #1928]: #515
+#2120 := (or #101 #922 #1510 #1409 #960 #1509 #1049 #335 #288)
+#2121 := [th-lemma]: #2120
+#2122 := [unit-resolution #2121 #1422 #694 #1924 #720 #1405 #1468 #776 #1928]: #101
+#2123 := [unit-resolution #918 #2122]: #899
+#2124 := [unit-resolution #1916 #2123 #2119 #2111 #1497 #1925 #2033 #2116 #1928 #1924 #2115 #2112]: false
+#2126 := [lemma #2124]: #2125
+#2130 := [unit-resolution #2126 #1405]: #335
+#2131 := [unit-resolution #1107 #2130]: #1077
+#2132 := [unit-resolution #1284 #2131]: #1071
+#2133 := [unit-resolution #1933 #2132 #2129 #2115 #1928 #2112 #1924]: #515
+#2134 := [unit-resolution #1916 #2133 #2111 #1497 #1925 #2033 #2116 #1928 #1924 #2115 #2112]: #898
+#2135 := [unit-resolution #918 #2134]: #100
+#2136 := [unit-resolution #917 #2135]: #887
+#2137 := [unit-resolution #1224 #2136]: #881
+#2138 := [unit-resolution #1943 #2137 #1497 #1870 #2116 #1928]: #506
+#2139 := [unit-resolution #646 #2138 #2111 #2133]: #632
+#2140 := [unit-resolution #1908 #2139 #1928 #694 #1924 #720 #1684 #746 #1468 #776 #2115 #824 #2112 #850 #672]: #897
+#2141 := [th-lemma #2140 #2135 #1742]: false
+#2142 := [lemma #2141]: #288
+#2143 := [unit-resolution #1069 #2142]: #1039
+#2144 := [unit-resolution #1272 #2143]: #1033
+#2145 := [hypothesis]: #1150
+#2146 := [unit-resolution #1308 #2145]: #1152
+#2147 := [unit-resolution #1183 #2146]: #430
+#2148 := [unit-resolution #1184 #2147]: #1165
+#2149 := [unit-resolution #1314 #2148]: #1160
+#2150 := [unit-resolution #1761 #2147]: #1109
+#2151 := [unit-resolution #1735 #2150]: #382
+#2152 := [unit-resolution #1960 #2151 #1870]: #1123
+#2153 := [unit-resolution #1988 #2152 #2149 #2033 #2144]: #509
+#2154 := (or #1149 #1147)
+#2155 := [th-lemma]: #2154
+#2156 := [unit-resolution #2155 #2145]: #1149
+#2157 := [unit-resolution #1894 #2147]: #1200
+#2158 := [unit-resolution #2061 #2149 #2015 #1870 #2157 #2151]: #1086
+#2159 := [unit-resolution #1290 #2158]: #1088
+#2160 := [unit-resolution #1108 #2159]: #335
+#2161 := [unit-resolution #1107 #2160]: #1077
+#2162 := [unit-resolution #1284 #2161]: #1071
+#2163 := [unit-resolution #1971 #2162 #2149 #2152 #2156 #2144]: #147
+#2164 := [unit-resolution #955 #2163]: #925
+#2165 := [unit-resolution #1236 #2164]: #919
+#2166 := [unit-resolution #1316 #2148]: #1161
+#2167 := (or #100 #923 #1371 #1357 #1523 #998 #1036 #383 #429 #1343 #1113 #973 #1373 #1074 #1358)
+#2168 := [th-lemma]: #2167
+#2169 := [unit-resolution #2168 #2162 #698 #1440 #724 #1684 #746 #2144 #772 #2033 #798 #2151 #1870 #828 #2147]: #100
+#2170 := [unit-resolution #917 #2169]: #887
+#2171 := [unit-resolution #1224 #2170]: #881
+#2172 := [unit-resolution #1992 #2171 #2166 #2150 #2162 #2144 #2165]: #506
+#2173 := (or #195 #1357 #1523 #998 #1036 #383 #429 #1343 #1113)
+#2174 := [th-lemma]: #2173
+#2175 := [unit-resolution #2174 #2151 #746 #2144 #772 #1684 #1870 #828 #2147]: #195
+#2176 := [unit-resolution #994 #2175]: #975
+#2177 := [unit-resolution #1254 #2176]: #970
+#2178 := (or #515 #922 #1074 #1036 #972 #1163 #1112)
+#2179 := [unit-resolution #1622 #694 #720 #1404 #750 #772 #1864 #798 #824 #854 #880]: #2178
+#2180 := [unit-resolution #2179 #2177 #2150 #2162 #2166 #2144 #2165]: #515
+#2181 := [unit-resolution #646 #2180 #2172 #2111 #2153]: false
+#2182 := [lemma #2181]: #1147
+#1805 := [unit-resolution #1302 #1729]: #1122
+#2231 := (or #194 #382)
+#2183 := (or #1150 #429 #1163)
+#2184 := [th-lemma]: #2183
+#2185 := [unit-resolution #2184 #1333 #2182]: #1163
+#2186 := [unit-resolution #1316 #2185 #1334]: false
+#2187 := [lemma #2186]: #429
+#2196 := [unit-resolution #1183 #2187]: #1153
+#2197 := [unit-resolution #1310 #2196]: #1149
+#1817 := [unit-resolution #1304 #1729]: #1123
+#2217 := [unit-resolution #1804 #1535 #2182 #1817 #2197 #1805]: #147
+#2218 := [unit-resolution #955 #2217]: #925
+#2219 := [unit-resolution #1236 #2218]: #919
+#2210 := [unit-resolution #1976 #2197 #2187]: #1161
+#2220 := (or #509 #1124 #935 #1150 #972)
+#2221 := [unit-resolution #1767 #698 #672 #720 #1703 #750 #1459 #776 #1447 #802 #824 #850 #1742]: #2220
+#2222 := [unit-resolution #2221 #1620 #2069 #1805 #2182]: #509
+#2223 := (or #515 #922 #1163 #972 #1124)
+#2224 := [unit-resolution #1707 #720 #1703 #750 #1459 #776 #1447 #802 #694 #824 #854 #880 #1864]: #2223
+#2225 := [unit-resolution #2224 #2219 #1805 #1620 #2210]: #515
+#2226 := [unit-resolution #646 #2225 #2111 #2222]: #631
+#2211 := (or #506 #884 #922 #1163 #1124)
+#2212 := [unit-resolution #1724 #668 #694 #1703 #750 #1459 #776 #824 #854]: #2211
+#2227 := [unit-resolution #2212 #2226 #1805 #2210 #2219]: #884
+#2228 := [unit-resolution #1224 #2227]: #886
+#2229 := [unit-resolution #917 #2228]: #101
+#2230 := [th-lemma #1620 #720 #1459 #776 #1447 #802 #2033 #2229 #698 #1428 #2217]: false
+#2232 := [lemma #2230]: #2231
+#2242 := [unit-resolution #2232 #1428]: #194
+#2243 := [unit-resolution #993 #2242]: #963
+#2244 := [unit-resolution #1248 #2243]: #957
+#2193 := (or #509 #1124 #1036 #935 #1150 #960 #1087)
+#1814 := (or #509 #885 #1522 #1523 #1365 #1489 #998 #1124 #1371 #1036 #1357 #935 #1150 #1509 #1350 #960 #1087)
+#1815 := [th-lemma]: #1814
+#2194 := [unit-resolution #1815 #698 #720 #1684 #746 #672 #772 #802 #824 #850 #1742]: #2193
+#2245 := [unit-resolution #2194 #2244 #2069 #2144 #2045 #1805 #2182]: #509
+#2205 := (or #100 #935 #1036 #382 #960 #1087)
+#1834 := (or #100 #1371 #935 #1523 #1036 #1357 #998 #1509 #382 #1350 #960 #1087)
+#1835 := [th-lemma]: #1834
+#2206 := [unit-resolution #1835 #698 #720 #1684 #746 #772 #802]: #2205
+#2246 := [unit-resolution #2206 #2244 #2045 #2069 #2144 #1428]: #100
+#2247 := [unit-resolution #917 #2246]: #887
+#2248 := [unit-resolution #1224 #2247]: #881
+#2215 := (or #335 #382)
+#2188 := (or #335 #194)
+#2189 := [unit-resolution #1777 #750 #1459 #776 #1703]: #2188
+#2190 := [unit-resolution #2189 #1422]: #194
+#2191 := [unit-resolution #993 #2190]: #963
+#2192 := [unit-resolution #1248 #2191]: #957
+#2195 := [unit-resolution #2194 #2192 #2069 #2144 #2045 #1805 #2182]: #509
+#2198 := [unit-resolution #1250 #2191]: #959
+#1840 := (or #335 #934 #1151 #961 #935 #960 #1150 #382)
+#1807 := [unit-resolution #1292 #1806]: #1085
+#1808 := [hypothesis]: #933
+#1809 := (or #288 #382 #1350 #335 #1087)
+#1810 := [th-lemma]: #1809
+#1811 := [unit-resolution #1810 #1422 #1807 #802 #1428]: #288
+#1812 := [unit-resolution #1069 #1811]: #1039
+#1813 := [unit-resolution #1272 #1812]: #1033
+#1816 := [unit-resolution #1815 #1813 #1808 #698 #1536 #720 #1684 #746 #672 #772 #1807 #802 #1805 #824 #1540 #850 #1742]: #509
+#1818 := (or #476 #1337 #1343 #1125 #1151 #335 #382)
+#1819 := [th-lemma]: #1818
+#1820 := [unit-resolution #1819 #1422 #1817 #828 #1774 #854 #1428]: #476
+#1821 := [unit-resolution #1221 #1820]: #1191
+#1822 := [unit-resolution #1320 #1821]: #1185
+#1825 := [unit-resolution #1824 #1813 #1536 #720 #1684 #746 #698 #772 #1807 #802 #1540 #850 #1822 #876 #1808]: #516
+#1826 := [hypothesis]: #932
+#1827 := [unit-resolution #1322 #1821]: #1187
+#1830 := (or #515 #1511 #1337 #1409 #1037 #1189 #1510 #934 #1373 #1358 #1151 #961 #1364 #999 #1086)
+#1831 := [th-lemma]: #1830
+#1832 := [unit-resolution #1831 #1829 #1828 #724 #1703 #750 #1459 #776 #694 #798 #1774 #854 #1827 #880 #1826]: #515
+#1833 := [unit-resolution #646 #1832 #1825 #1816]: #631
+#1836 := [unit-resolution #1835 #1813 #698 #1536 #720 #1684 #746 #1808 #772 #1807 #802 #1428]: #100
+#1837 := [unit-resolution #917 #1836]: #887
+#1838 := [unit-resolution #1224 #1837]: #881
+#1839 := [th-lemma #1838 #668 #750 #828 #854 #1703 #1817 #694 #1459 #776 #1826 #1774 #724 #798 #1828 #1829 #1833]: false
+#1841 := [lemma #1839]: #1840
+#2199 := [unit-resolution #1841 #2198 #2069 #1422 #2197 #2192 #2182 #1428]: #934
+#2200 := [unit-resolution #1242 #2199]: #936
+#2201 := [unit-resolution #956 #2200]: #147
+#2202 := [unit-resolution #955 #2201]: #925
+#2203 := [unit-resolution #1236 #2202]: #919
+#2204 := [unit-resolution #2118 #2203 #1829 #720 #1468 #776 #694 #798 #2197 #854 #1864 #880 #2192]: #515
+#2207 := [unit-resolution #2206 #2192 #2045 #2069 #2144 #1428]: #100
+#2208 := [unit-resolution #917 #2207]: #887
+#2209 := [unit-resolution #1224 #2208]: #881
+#2213 := [unit-resolution #2212 #2203 #1805 #2210 #2209]: #506
+#2214 := [unit-resolution #646 #2213 #2204 #2111 #2195]: false
+#2216 := [lemma #2214]: #2215
+#2249 := [unit-resolution #2216 #1428]: #335
+#2250 := [unit-resolution #1107 #2249]: #1077
+#2251 := [unit-resolution #1284 #2250]: #1071
+#2252 := (or #1084 #1074 #1357 #1523 #998 #1036 #195)
+#2253 := [th-lemma]: #2252
+#2254 := [unit-resolution #2253 #2251 #746 #2144 #772 #1684 #2242]: #1084
+#2255 := [unit-resolution #1250 #2243]: #959
+#2240 := (or #934 #632 #884 #1074 #1125 #961 #1086)
+#2233 := (or #515 #934 #1151 #961 #1086)
+#2234 := [unit-resolution #1831 #1864 #724 #1703 #750 #1459 #776 #694 #798 #854 #880]: #2233
+#2235 := [unit-resolution #2234 #1826 #2197 #1828 #2054]: #515
+#2236 := (or #506 #884 #1125 #1036 #934 #1151 #1074)
+#2237 := [unit-resolution #1788 #694 #1440 #724 #1684 #746 #668 #772 #798 #828 #854]: #2236
+#2238 := [unit-resolution #2237 #1826 #1636 #1638 #1775 #2197 #2144]: #506
+#2239 := [unit-resolution #646 #2238 #2235 #2111 #1628]: false
+#2241 := [lemma #2239]: #2240
+#2256 := [unit-resolution #2241 #2245 #2248 #2251 #1817 #2255 #2254]: #934
+#2257 := [unit-resolution #1242 #2256]: #936
+#2258 := [unit-resolution #956 #2257]: #147
+#2259 := [unit-resolution #955 #2258]: #925
+#2260 := [unit-resolution #1236 #2259]: #919
+#2261 := [unit-resolution #2212 #2260 #1805 #2210 #2248]: #506
+#2262 := [unit-resolution #2118 #2260 #2254 #720 #1468 #776 #694 #798 #2197 #854 #1864 #880 #2244]: #515
+#2263 := [unit-resolution #646 #2262 #2261 #2111 #2245]: false
+#2264 := [lemma #2263]: #382
+#2265 := [unit-resolution #1145 #2264]: #1115
+#2266 := [unit-resolution #1296 #2265]: #1109
+#2267 := [unit-resolution #2189 #1535]: #335
+#2268 := [unit-resolution #1107 #2267]: #1077
+#2269 := [unit-resolution #1284 #2268]: #1071
+#2270 := [unit-resolution #1966 #2269 #2142 #2266 #2182]: #1160
+#2271 := (or #1008 #998 #1036 #1357 #1074 #1358 #383)
+#2272 := [th-lemma]: #2271
+#2273 := [unit-resolution #2272 #2269 #2144 #772 #1684 #798 #2264]: #1008
+#2274 := (or #509 #1010 #1113 #923 #1162)
+#2275 := [unit-resolution #1608 #672 #698 #1742 #746 #1459 #776 #1447 #802 #828 #850]: #2274
+#2276 := [unit-resolution #2275 #2273 #1870 #2270 #2033]: #509
+#2277 := [unit-resolution #1960 #2264 #1870]: #1123
+#2278 := [unit-resolution #1971 #2270 #2269 #2277 #2197 #2144]: #147
+#2279 := [unit-resolution #955 #2278]: #925
+#2280 := [unit-resolution #1236 #2279]: #919
+#2281 := (or #1010 #999 #923 #100 #1371 #961 #1373)
+#2282 := [th-lemma]: #2281
+#2283 := [unit-resolution #2282 #2273 #698 #1584 #724 #1703 #2033]: #100
+#2284 := [unit-resolution #917 #2283]: #887
+#2285 := [unit-resolution #1224 #2284]: #881
+#2286 := [unit-resolution #1992 #2285 #2210 #2266 #2269 #2144 #2280]: #506
+#2287 := [unit-resolution #2179 #2280 #2266 #1620 #2210 #2144 #2269]: #515
+#2288 := [unit-resolution #646 #2287 #2286 #2111 #2276]: false
+#2289 := [lemma #2288]: #194
+#2305 := [unit-resolution #2253 #2302 #746 #2144 #772 #1684 #2289]: #1074
+#2306 := [unit-resolution #1284 #2305]: #1076
+#2307 := [unit-resolution #1107 #2306 #2304]: false
+#2308 := [lemma #2307]: #1084
+#2300 := (or #1086 #515)
+#2290 := [hypothesis]: #633
+#2291 := [unit-resolution #993 #2289]: #963
+#2292 := [unit-resolution #1250 #2291]: #959
+#2293 := [unit-resolution #2234 #2054 #2197 #2292 #2290]: #934
+#2294 := [unit-resolution #1242 #2293]: #936
+#2295 := [unit-resolution #1248 #2291]: #957
+#2296 := [unit-resolution #2118 #2054 #2290 #720 #1468 #776 #694 #798 #2197 #854 #1864 #880 #2295]: #922
+#2297 := [unit-resolution #1236 #2296]: #924
+#2298 := [unit-resolution #955 #2297]: #148
+#2299 := [unit-resolution #956 #2298 #2294]: false
+#2301 := [lemma #2299]: #2300
+#1848 := [unit-resolution #2301 #2308]: #515
+#1851 := [hypothesis]: #632
+#1852 := (or #897 #1522 #509 #1523 #998 #1365 #1489 #1150 #1509 #1350 #633 #1372 #1188 #960 #1087 #1112)
+#1853 := [th-lemma]: #1852
+#1846 := [unit-resolution #1853 #1851 #2295 #720 #1684 #746 #2045 #802 #2266 #824 #2182 #850 #2015 #876 #672 #1848]: #897
+#1847 := [unit-resolution #1232 #1846]: #898
+#1854 := [unit-resolution #918 #1847]: #100
+#1855 := (or #509 #1124)
+#1856 := [unit-resolution #2194 #2069 #2144 #2045 #2295 #2182]: #1855
+#2309 := [unit-resolution #1856 #1851]: #1124
+#2310 := [th-lemma #1848 #876 #850 #2182 #2015 #2309 #2266 #1854]: false
+#2311 := [lemma #2310]: #509
+#2312 := (or #631 #632)
+#2313 := [unit-resolution #646 #2111 #1848]: #2312
+#2314 := [unit-resolution #2313 #2311]: #631
+#2315 := (or #884 #633 #1372 #1188 #1125 #1528 #506 #1364 #999 #1343 #1373 #1358 #961 #1086)
+#2316 := [th-lemma]: #2315
+#2317 := [unit-resolution #2316 #668 #2292 #724 #1703 #750 #2308 #798 #2277 #828 #2015 #876 #2314 #1848]: #884
+#2318 := [unit-resolution #1224 #2317]: #886
+#2319 := (or #896 #1528 #506 #1364 #999 #1343 #1337 #1151 #1373 #1358 #634 #1511 #1189 #961 #1086 #1113)
+#2320 := [th-lemma]: #2319
+#2321 := [unit-resolution #2320 #668 #2292 #724 #1703 #750 #2308 #798 #1870 #828 #2197 #854 #1864 #880 #2314 #2111]: #896
+#2322 := [unit-resolution #1230 #2321]: #898
+#2323 := [unit-resolution #918 #2322]: #100
+[unit-resolution #917 #2323 #2318]: false
+unsat
+
+SqgPFdiZeq8SOFuTISQN5g 1109
+#2 := false
+#8 := 1::real
+decl uf_1 :: real
+#4 := uf_1
+#6 := 2::real
+#7 := (* 2::real uf_1)
+#9 := (+ #7 1::real)
+#5 := (+ uf_1 uf_1)
+#10 := (< #5 #9)
+#11 := (or false #10)
+#12 := (or #10 #11)
+#13 := (not #12)
+#64 := (iff #13 false)
+#32 := (+ 1::real #7)
+#35 := (< #7 #32)
+#52 := (not #35)
+#62 := (iff #52 false)
+#1 := true
+#57 := (not true)
+#60 := (iff #57 false)
+#61 := [rewrite]: #60
+#58 := (iff #52 #57)
+#55 := (iff #35 true)
+#56 := [rewrite]: #55
+#59 := [monotonicity #56]: #58
+#63 := [trans #59 #61]: #62
+#53 := (iff #13 #52)
+#50 := (iff #12 #35)
+#45 := (or #35 #35)
+#48 := (iff #45 #35)
+#49 := [rewrite]: #48
+#46 := (iff #12 #45)
+#43 := (iff #11 #35)
+#38 := (or false #35)
+#41 := (iff #38 #35)
+#42 := [rewrite]: #41
+#39 := (iff #11 #38)
+#36 := (iff #10 #35)
+#33 := (= #9 #32)
+#34 := [rewrite]: #33
+#30 := (= #5 #7)
+#31 := [rewrite]: #30
+#37 := [monotonicity #31 #34]: #36
+#40 := [monotonicity #37]: #39
+#44 := [trans #40 #42]: #43
+#47 := [monotonicity #37 #44]: #46
+#51 := [trans #47 #49]: #50
+#54 := [monotonicity #51]: #53
+#65 := [trans #54 #63]: #64
+#29 := [asserted]: #13
+[mp #29 #65]: false
+unsat
+
+rOkYPs+Q++z/M3OPR/88JQ 1272
+#2 := false
+#55 := 0::int
+#7 := 2::int
+decl uf_1 :: int
+#4 := uf_1
+#8 := (mod uf_1 2::int)
+#58 := (>= #8 0::int)
+#61 := (not #58)
+#5 := 1::int
+#9 := (* 2::int #8)
+#10 := (+ #9 1::int)
+#11 := (+ uf_1 #10)
+#6 := (+ uf_1 1::int)
+#12 := (<= #6 #11)
+#13 := (not #12)
+#66 := (iff #13 #61)
+#39 := (+ uf_1 #9)
+#40 := (+ 1::int #39)
+#30 := (+ 1::int uf_1)
+#45 := (<= #30 #40)
+#48 := (not #45)
+#64 := (iff #48 #61)
+#56 := (>= #9 0::int)
+#51 := (not #56)
+#62 := (iff #51 #61)
+#59 := (iff #56 #58)
+#60 := [rewrite]: #59
+#63 := [monotonicity #60]: #62
+#52 := (iff #48 #51)
+#53 := (iff #45 #56)
+#54 := [rewrite]: #53
+#57 := [monotonicity #54]: #52
+#65 := [trans #57 #63]: #64
+#49 := (iff #13 #48)
+#46 := (iff #12 #45)
+#43 := (= #11 #40)
+#33 := (+ 1::int #9)
+#36 := (+ uf_1 #33)
+#41 := (= #36 #40)
+#42 := [rewrite]: #41
+#37 := (= #11 #36)
+#34 := (= #10 #33)
+#35 := [rewrite]: #34
+#38 := [monotonicity #35]: #37
+#44 := [trans #38 #42]: #43
+#31 := (= #6 #30)
+#32 := [rewrite]: #31
+#47 := [monotonicity #32 #44]: #46
+#50 := [monotonicity #47]: #49
+#67 := [trans #50 #65]: #66
+#29 := [asserted]: #13
+#68 := [mp #29 #67]: #61
+#1 := true
+#28 := [true-axiom]: true
+#142 := (or false #58)
+#143 := [th-lemma]: #142
+#144 := [unit-resolution #143 #28]: #58
+[unit-resolution #144 #68]: false
+unsat
+
+oSBTeiF6GKDeHPsMxXHXtQ 1161
+#2 := false
+#5 := 2::int
+decl uf_1 :: int
+#4 := uf_1
+#6 := (mod uf_1 2::int)
+#55 := (>= #6 2::int)
+#9 := 3::int
+#10 := (+ uf_1 3::int)
+#7 := (+ #6 #6)
+#8 := (+ uf_1 #7)
+#11 := (< #8 #10)
+#12 := (not #11)
+#60 := (iff #12 #55)
+#35 := (+ 3::int uf_1)
+#29 := (* 2::int #6)
+#32 := (+ uf_1 #29)
+#38 := (< #32 #35)
+#41 := (not #38)
+#58 := (iff #41 #55)
+#48 := (>= #29 3::int)
+#56 := (iff #48 #55)
+#57 := [rewrite]: #56
+#53 := (iff #41 #48)
+#46 := (not #48)
+#45 := (not #46)
+#51 := (iff #45 #48)
+#52 := [rewrite]: #51
+#49 := (iff #41 #45)
+#47 := (iff #38 #46)
+#44 := [rewrite]: #47
+#50 := [monotonicity #44]: #49
+#54 := [trans #50 #52]: #53
+#59 := [trans #54 #57]: #58
+#42 := (iff #12 #41)
+#39 := (iff #11 #38)
+#36 := (= #10 #35)
+#37 := [rewrite]: #36
+#33 := (= #8 #32)
+#30 := (= #7 #29)
+#31 := [rewrite]: #30
+#34 := [monotonicity #31]: #33
+#40 := [monotonicity #34 #37]: #39
+#43 := [monotonicity #40]: #42
+#61 := [trans #43 #59]: #60
+#28 := [asserted]: #12
+#62 := [mp #28 #61]: #55
+#127 := (not #55)
+#1 := true
+#27 := [true-axiom]: true
+#137 := (or false #127)
+#138 := [th-lemma]: #137
+#139 := [unit-resolution #138 #27]: #127
+[unit-resolution #139 #62]: false
+unsat
+
+hqH9yBHvmZgES3pBkvsqVQ 2715
+#2 := false
+#5 := 0::real
+decl uf_1 :: real
+#4 := uf_1
+#94 := (<= uf_1 0::real)
+#17 := 2::real
+#40 := (* 2::real uf_1)
+#102 := (<= #40 0::real)
+#103 := (>= #40 0::real)
+#105 := (not #103)
+#104 := (not #102)
+#106 := (or #104 #105)
+#107 := (not #106)
+#88 := (= #40 0::real)
+#108 := (iff #88 #107)
+#109 := [rewrite]: #108
+#16 := 4::real
+#11 := (- uf_1)
+#10 := (< uf_1 0::real)
+#12 := (ite #10 #11 uf_1)
+#9 := 1::real
+#13 := (< 1::real #12)
+#14 := (not #13)
+#15 := (or #13 #14)
+#18 := (ite #15 4::real 2::real)
+#19 := (* #18 uf_1)
+#8 := (+ uf_1 uf_1)
+#20 := (= #8 #19)
+#21 := (not #20)
+#22 := (not #21)
+#89 := (iff #22 #88)
+#70 := (* 4::real uf_1)
+#73 := (= #40 #70)
+#86 := (iff #73 #88)
+#87 := [rewrite]: #86
+#84 := (iff #22 #73)
+#76 := (not #73)
+#79 := (not #76)
+#82 := (iff #79 #73)
+#83 := [rewrite]: #82
+#80 := (iff #22 #79)
+#77 := (iff #21 #76)
+#74 := (iff #20 #73)
+#71 := (= #19 #70)
+#68 := (= #18 4::real)
+#1 := true
+#63 := (ite true 4::real 2::real)
+#66 := (= #63 4::real)
+#67 := [rewrite]: #66
+#64 := (= #18 #63)
+#61 := (iff #15 true)
+#43 := -1::real
+#44 := (* -1::real uf_1)
+#47 := (ite #10 #44 uf_1)
+#50 := (< 1::real #47)
+#53 := (not #50)
+#56 := (or #50 #53)
+#59 := (iff #56 true)
+#60 := [rewrite]: #59
+#57 := (iff #15 #56)
+#54 := (iff #14 #53)
+#51 := (iff #13 #50)
+#48 := (= #12 #47)
+#45 := (= #11 #44)
+#46 := [rewrite]: #45
+#49 := [monotonicity #46]: #48
+#52 := [monotonicity #49]: #51
+#55 := [monotonicity #52]: #54
+#58 := [monotonicity #52 #55]: #57
+#62 := [trans #58 #60]: #61
+#65 := [monotonicity #62]: #64
+#69 := [trans #65 #67]: #68
+#72 := [monotonicity #69]: #71
+#41 := (= #8 #40)
+#42 := [rewrite]: #41
+#75 := [monotonicity #42 #72]: #74
+#78 := [monotonicity #75]: #77
+#81 := [monotonicity #78]: #80
+#85 := [trans #81 #83]: #84
+#90 := [trans #85 #87]: #89
+#39 := [asserted]: #22
+#91 := [mp #39 #90]: #88
+#110 := [mp #91 #109]: #107
+#111 := [not-or-elim #110]: #102
+#127 := (or #94 #104)
+#128 := [th-lemma]: #127
+#129 := [unit-resolution #128 #111]: #94
+#92 := (>= uf_1 0::real)
+#112 := [not-or-elim #110]: #103
+#130 := (or #92 #105)
+#131 := [th-lemma]: #130
+#132 := [unit-resolution #131 #112]: #92
+#114 := (not #94)
+#113 := (not #92)
+#115 := (or #113 #114)
+#95 := (and #92 #94)
+#98 := (not #95)
+#124 := (iff #98 #115)
+#116 := (not #115)
+#119 := (not #116)
+#122 := (iff #119 #115)
+#123 := [rewrite]: #122
+#120 := (iff #98 #119)
+#117 := (iff #95 #116)
+#118 := [rewrite]: #117
+#121 := [monotonicity #118]: #120
+#125 := [trans #121 #123]: #124
+#6 := (= uf_1 0::real)
+#7 := (not #6)
+#99 := (iff #7 #98)
+#96 := (iff #6 #95)
+#97 := [rewrite]: #96
+#100 := [monotonicity #97]: #99
+#38 := [asserted]: #7
+#101 := [mp #38 #100]: #98
+#126 := [mp #101 #125]: #115
+[unit-resolution #126 #132 #129]: false
+unsat
+
+ar4IlNF9IylWgGSPOf9paw 5159
+#2 := false
+#9 := 0::int
+#11 := 4::int
+decl uf_1 :: int
+#4 := uf_1
+#189 := (div uf_1 4::int)
+#210 := -4::int
+#211 := (* -4::int #189)
+#12 := (mod uf_1 4::int)
+#134 := -1::int
+#209 := (* -1::int #12)
+#212 := (+ #209 #211)
+#213 := (+ uf_1 #212)
+#214 := (<= #213 0::int)
+#215 := (not #214)
+#208 := (>= #213 0::int)
+#207 := (not #208)
+#216 := (or #207 #215)
+#217 := (not #216)
+#1 := true
+#36 := [true-axiom]: true
+#393 := (or false #217)
+#394 := [th-lemma]: #393
+#395 := [unit-resolution #394 #36]: #217
+#224 := (or #216 #214)
+#225 := [def-axiom]: #224
+#396 := [unit-resolution #225 #395]: #214
+#222 := (or #216 #208)
+#223 := [def-axiom]: #222
+#397 := [unit-resolution #223 #395]: #208
+#250 := (>= #12 4::int)
+#251 := (not #250)
+#398 := (or false #251)
+#399 := [th-lemma]: #398
+#400 := [unit-resolution #399 #36]: #251
+#13 := 3::int
+#90 := (>= #12 3::int)
+#92 := (not #90)
+#89 := (<= #12 3::int)
+#91 := (not #89)
+#93 := (or #91 #92)
+#94 := (not #93)
+#14 := (= #12 3::int)
+#95 := (iff #14 #94)
+#96 := [rewrite]: #95
+#38 := [asserted]: #14
+#97 := [mp #38 #96]: #94
+#99 := [not-or-elim #97]: #90
+#7 := 2::int
+#261 := (div uf_1 2::int)
+#140 := -2::int
+#276 := (* -2::int #261)
+#15 := (mod uf_1 2::int)
+#275 := (* -1::int #15)
+#277 := (+ #275 #276)
+#278 := (+ uf_1 #277)
+#279 := (<= #278 0::int)
+#280 := (not #279)
+#274 := (>= #278 0::int)
+#273 := (not #274)
+#281 := (or #273 #280)
+#282 := (not #281)
+#408 := (or false #282)
+#409 := [th-lemma]: #408
+#410 := [unit-resolution #409 #36]: #282
+#289 := (or #281 #279)
+#290 := [def-axiom]: #289
+#411 := [unit-resolution #290 #410]: #279
+#287 := (or #281 #274)
+#288 := [def-axiom]: #287
+#412 := [unit-resolution #288 #410]: #274
+#16 := 1::int
+#55 := (>= #15 1::int)
+#100 := (not #55)
+decl uf_2 :: int
+#5 := uf_2
+#18 := (mod uf_2 2::int)
+#61 := (<= #18 1::int)
+#102 := (not #61)
+#375 := [hypothesis]: #102
+#358 := (>= #18 2::int)
+#359 := (not #358)
+#403 := (or false #359)
+#404 := [th-lemma]: #403
+#405 := [unit-resolution #404 #36]: #359
+#406 := [th-lemma #405 #375]: false
+#407 := [lemma #406]: #61
+#413 := (or #100 #102)
+#62 := (>= #18 1::int)
+#315 := (div uf_2 2::int)
+#330 := (* -2::int #315)
+#329 := (* -1::int #18)
+#331 := (+ #329 #330)
+#332 := (+ uf_2 #331)
+#333 := (<= #332 0::int)
+#334 := (not #333)
+#328 := (>= #332 0::int)
+#327 := (not #328)
+#335 := (or #327 #334)
+#336 := (not #335)
+#376 := (or false #336)
+#377 := [th-lemma]: #376
+#378 := [unit-resolution #377 #36]: #336
+#343 := (or #335 #333)
+#344 := [def-axiom]: #343
+#379 := [unit-resolution #344 #378]: #333
+#341 := (or #335 #328)
+#342 := [def-axiom]: #341
+#380 := [unit-resolution #342 #378]: #328
+#103 := (not #62)
+#381 := [hypothesis]: #103
+#352 := (>= #18 0::int)
+#382 := (or false #352)
+#383 := [th-lemma]: #382
+#384 := [unit-resolution #383 #36]: #352
+#6 := (+ uf_1 uf_2)
+#116 := (div #6 2::int)
+#141 := (* -2::int #116)
+#8 := (mod #6 2::int)
+#139 := (* -1::int #8)
+#142 := (+ #139 #141)
+#143 := (+ uf_2 #142)
+#144 := (+ uf_1 #143)
+#138 := (<= #144 0::int)
+#136 := (not #138)
+#137 := (>= #144 0::int)
+#135 := (not #137)
+#145 := (or #135 #136)
+#146 := (not #145)
+#385 := (or false #146)
+#386 := [th-lemma]: #385
+#387 := [unit-resolution #386 #36]: #146
+#153 := (or #145 #138)
+#154 := [def-axiom]: #153
+#388 := [unit-resolution #154 #387]: #138
+#151 := (or #145 #137)
+#152 := [def-axiom]: #151
+#389 := [unit-resolution #152 #387]: #137
+#78 := (<= #8 0::int)
+#79 := (>= #8 0::int)
+#81 := (not #79)
+#80 := (not #78)
+#82 := (or #80 #81)
+#83 := (not #82)
+#10 := (= #8 0::int)
+#84 := (iff #10 #83)
+#85 := [rewrite]: #84
+#37 := [asserted]: #10
+#86 := [mp #37 #85]: #83
+#87 := [not-or-elim #86]: #78
+#390 := (or false #79)
+#391 := [th-lemma]: #390
+#392 := [unit-resolution #391 #36]: #79
+#401 := [th-lemma #99 #400 #397 #396 #392 #87 #389 #388 #384 #381 #380 #379]: false
+#402 := [lemma #401]: #62
+#57 := (<= #15 1::int)
+#101 := (not #57)
+#369 := [hypothesis]: #101
+#304 := (>= #15 2::int)
+#305 := (not #304)
+#370 := (or false #305)
+#371 := [th-lemma]: #370
+#372 := [unit-resolution #371 #36]: #305
+#373 := [th-lemma #372 #369]: false
+#374 := [lemma #373]: #57
+#104 := (or #100 #101 #102 #103)
+#69 := (and #55 #57 #61 #62)
+#74 := (not #69)
+#113 := (iff #74 #104)
+#105 := (not #104)
+#108 := (not #105)
+#111 := (iff #108 #104)
+#112 := [rewrite]: #111
+#109 := (iff #74 #108)
+#106 := (iff #69 #105)
+#107 := [rewrite]: #106
+#110 := [monotonicity #107]: #109
+#114 := [trans #110 #112]: #113
+#19 := (= #18 1::int)
+#17 := (= #15 1::int)
+#20 := (and #17 #19)
+#21 := (not #20)
+#75 := (iff #21 #74)
+#72 := (iff #20 #69)
+#63 := (and #61 #62)
+#58 := (and #55 #57)
+#66 := (and #58 #63)
+#70 := (iff #66 #69)
+#71 := [rewrite]: #70
+#67 := (iff #20 #66)
+#64 := (iff #19 #63)
+#65 := [rewrite]: #64
+#59 := (iff #17 #58)
+#60 := [rewrite]: #59
+#68 := [monotonicity #60 #65]: #67
+#73 := [trans #68 #71]: #72
+#76 := [monotonicity #73]: #75
+#39 := [asserted]: #21
+#77 := [mp #39 #76]: #74
+#115 := [mp #77 #114]: #104
+#414 := [unit-resolution #115 #374 #402]: #413
+#415 := [unit-resolution #414 #407]: #100
+#298 := (>= #15 0::int)
+#416 := (or false #298)
+#417 := [th-lemma]: #416
+#418 := [unit-resolution #417 #36]: #298
+[th-lemma #418 #415 #412 #411 #99 #400 #397 #396]: false
+unsat
+
+o9WM91Nq0O5f08PEA0qA6w 515
+#2 := false
+#4 := (exists (vars (?x1 int)) false)
+#5 := (not #4)
+#6 := (not #5)
+#37 := (iff #6 false)
+#1 := true
+#32 := (not true)
+#35 := (iff #32 false)
+#36 := [rewrite]: #35
+#33 := (iff #6 #32)
+#30 := (iff #5 true)
+#25 := (not false)
+#28 := (iff #25 true)
+#29 := [rewrite]: #28
+#26 := (iff #5 #25)
+#23 := (iff #4 false)
+#24 := [elim-unused]: #23
+#27 := [monotonicity #24]: #26
+#31 := [trans #27 #29]: #30
+#34 := [monotonicity #31]: #33
+#38 := [trans #34 #36]: #37
+#22 := [asserted]: #6
+[mp #22 #38]: false
+unsat
+
+SJxvvXYE4z1G4iLRBCPerg 516
+#2 := false
+#4 := (exists (vars (?x1 real)) false)
+#5 := (not #4)
+#6 := (not #5)
+#37 := (iff #6 false)
+#1 := true
+#32 := (not true)
+#35 := (iff #32 false)
+#36 := [rewrite]: #35
+#33 := (iff #6 #32)
+#30 := (iff #5 true)
+#25 := (not false)
+#28 := (iff #25 true)
+#29 := [rewrite]: #28
+#26 := (iff #5 #25)
+#23 := (iff #4 false)
+#24 := [elim-unused]: #23
+#27 := [monotonicity #24]: #26
+#31 := [trans #27 #29]: #30
+#34 := [monotonicity #31]: #33
+#38 := [trans #34 #36]: #37
+#22 := [asserted]: #6
+[mp #22 #38]: false
+unsat
+
+Fr3hfDrqfMuGDpDYbXAHwg 7
+unsat
+
+bFuFCRUoQSRWyp0iCwo+PA 7
+unsat
+
+NJEgv3p5NO4/yEATNBBDNw 7
+unsat
+
+RC1LWjyqEEsh1xhocCgPmQ 1633
+#2 := false
+#5 := 0::int
+#8 := 1::int
+#143 := (= 1::int 0::int)
+#145 := (iff #143 false)
+#146 := [rewrite]: #145
+decl ?x1!1 :: int
+#47 := ?x1!1
+#51 := (= ?x1!1 0::int)
+decl ?x2!0 :: int
+#46 := ?x2!0
+#50 := (= ?x2!0 1::int)
+#63 := (and #50 #51)
+#69 := (= ?x2!0 ?x1!1)
+#72 := (not #69)
+#66 := (not #63)
+#75 := (or #66 #72)
+#78 := (not #75)
+#48 := (= ?x1!1 ?x2!0)
+#49 := (not #48)
+#52 := (and #51 #50)
+#53 := (not #52)
+#54 := (or #53 #49)
+#55 := (not #54)
+#79 := (iff #55 #78)
+#76 := (iff #54 #75)
+#73 := (iff #49 #72)
+#70 := (iff #48 #69)
+#71 := [rewrite]: #70
+#74 := [monotonicity #71]: #73
+#67 := (iff #53 #66)
+#64 := (iff #52 #63)
+#65 := [rewrite]: #64
+#68 := [monotonicity #65]: #67
+#77 := [monotonicity #68 #74]: #76
+#80 := [monotonicity #77]: #79
+#7 := (:var 0 int)
+#4 := (:var 1 int)
+#11 := (= #4 #7)
+#12 := (not #11)
+#9 := (= #7 1::int)
+#6 := (= #4 0::int)
+#10 := (and #6 #9)
+#32 := (not #10)
+#33 := (or #32 #12)
+#36 := (forall (vars (?x1 int) (?x2 int)) #33)
+#39 := (not #36)
+#56 := (~ #39 #55)
+#57 := [sk]: #56
+#13 := (implies #10 #12)
+#14 := (forall (vars (?x1 int) (?x2 int)) #13)
+#15 := (not #14)
+#40 := (iff #15 #39)
+#37 := (iff #14 #36)
+#34 := (iff #13 #33)
+#35 := [rewrite]: #34
+#38 := [quant-intro #35]: #37
+#41 := [monotonicity #38]: #40
+#31 := [asserted]: #15
+#44 := [mp #31 #41]: #39
+#60 := [mp~ #44 #57]: #55
+#61 := [mp #60 #80]: #78
+#62 := [not-or-elim #61]: #63
+#82 := [and-elim #62]: #51
+#141 := (= 1::int ?x1!1)
+#83 := [not-or-elim #61]: #69
+#139 := (= 1::int ?x2!0)
+#81 := [and-elim #62]: #50
+#140 := [symm #81]: #139
+#142 := [trans #140 #83]: #141
+#144 := [trans #142 #82]: #143
+[mp #144 #146]: false
+unsat
+
+6pnpFuE9SN6Jr5Upml9T0A 1896
+#2 := false
+#5 := (:var 0 int)
+#7 := 0::int
+#9 := (<= 0::int #5)
+#8 := (< #5 0::int)
+#10 := (or #8 #9)
+#4 := (:var 1 int)
+#6 := (< #4 #5)
+#11 := (implies #6 #10)
+#12 := (forall (vars (?x2 int)) #11)
+#13 := (exists (vars (?x1 int)) #12)
+#14 := (not #13)
+#95 := (iff #14 false)
+#31 := (not #6)
+#32 := (or #31 #10)
+#35 := (forall (vars (?x2 int)) #32)
+#38 := (exists (vars (?x1 int)) #35)
+#41 := (not #38)
+#93 := (iff #41 false)
+#1 := true
+#88 := (not true)
+#91 := (iff #88 false)
+#92 := [rewrite]: #91
+#89 := (iff #41 #88)
+#86 := (iff #38 true)
+#81 := (exists (vars (?x1 int)) true)
+#84 := (iff #81 true)
+#85 := [elim-unused]: #84
+#82 := (iff #38 #81)
+#79 := (iff #35 true)
+#74 := (forall (vars (?x2 int)) true)
+#77 := (iff #74 true)
+#78 := [elim-unused]: #77
+#75 := (iff #35 #74)
+#72 := (iff #32 true)
+#46 := (>= #5 0::int)
+#44 := (not #46)
+#64 := (or #44 #46)
+#50 := -1::int
+#53 := (* -1::int #5)
+#54 := (+ #4 #53)
+#52 := (>= #54 0::int)
+#67 := (or #52 #64)
+#70 := (iff #67 true)
+#71 := [rewrite]: #70
+#68 := (iff #32 #67)
+#65 := (iff #10 #64)
+#48 := (iff #9 #46)
+#49 := [rewrite]: #48
+#45 := (iff #8 #44)
+#47 := [rewrite]: #45
+#66 := [monotonicity #47 #49]: #65
+#62 := (iff #31 #52)
+#51 := (not #52)
+#57 := (not #51)
+#60 := (iff #57 #52)
+#61 := [rewrite]: #60
+#58 := (iff #31 #57)
+#55 := (iff #6 #51)
+#56 := [rewrite]: #55
+#59 := [monotonicity #56]: #58
+#63 := [trans #59 #61]: #62
+#69 := [monotonicity #63 #66]: #68
+#73 := [trans #69 #71]: #72
+#76 := [quant-intro #73]: #75
+#80 := [trans #76 #78]: #79
+#83 := [quant-intro #80]: #82
+#87 := [trans #83 #85]: #86
+#90 := [monotonicity #87]: #89
+#94 := [trans #90 #92]: #93
+#42 := (iff #14 #41)
+#39 := (iff #13 #38)
+#36 := (iff #12 #35)
+#33 := (iff #11 #32)
+#34 := [rewrite]: #33
+#37 := [quant-intro #34]: #36
+#40 := [quant-intro #37]: #39
+#43 := [monotonicity #40]: #42
+#96 := [trans #43 #94]: #95
+#30 := [asserted]: #14
+[mp #30 #96]: false
+unsat
+
+sHpY0IFBgZUNhP56aRB+/w 1765
+#2 := false
+#5 := (:var 0 int)
+#7 := 2::int
+#11 := (* 2::int #5)
+#9 := 1::int
+#4 := (:var 1 int)
+#8 := (* 2::int #4)
+#10 := (+ #8 1::int)
+#12 := (< #10 #11)
+#6 := (< #4 #5)
+#13 := (implies #6 #12)
+#14 := (forall (vars (?x1 int) (?x2 int)) #13)
+#15 := (not #14)
+#91 := (iff #15 false)
+#32 := (+ 1::int #8)
+#35 := (< #32 #11)
+#41 := (not #6)
+#42 := (or #41 #35)
+#47 := (forall (vars (?x1 int) (?x2 int)) #42)
+#50 := (not #47)
+#89 := (iff #50 false)
+#1 := true
+#84 := (not true)
+#87 := (iff #84 false)
+#88 := [rewrite]: #87
+#85 := (iff #50 #84)
+#82 := (iff #47 true)
+#77 := (forall (vars (?x1 int) (?x2 int)) true)
+#80 := (iff #77 true)
+#81 := [elim-unused]: #80
+#78 := (iff #47 #77)
+#75 := (iff #42 true)
+#55 := 0::int
+#53 := -1::int
+#57 := (* -1::int #5)
+#58 := (+ #4 #57)
+#56 := (>= #58 0::int)
+#54 := (not #56)
+#69 := (or #56 #54)
+#73 := (iff #69 true)
+#74 := [rewrite]: #73
+#71 := (iff #42 #69)
+#70 := (iff #35 #54)
+#68 := [rewrite]: #70
+#66 := (iff #41 #56)
+#61 := (not #54)
+#64 := (iff #61 #56)
+#65 := [rewrite]: #64
+#62 := (iff #41 #61)
+#59 := (iff #6 #54)
+#60 := [rewrite]: #59
+#63 := [monotonicity #60]: #62
+#67 := [trans #63 #65]: #66
+#72 := [monotonicity #67 #68]: #71
+#76 := [trans #72 #74]: #75
+#79 := [quant-intro #76]: #78
+#83 := [trans #79 #81]: #82
+#86 := [monotonicity #83]: #85
+#90 := [trans #86 #88]: #89
+#51 := (iff #15 #50)
+#48 := (iff #14 #47)
+#45 := (iff #13 #42)
+#38 := (implies #6 #35)
+#43 := (iff #38 #42)
+#44 := [rewrite]: #43
+#39 := (iff #13 #38)
+#36 := (iff #12 #35)
+#33 := (= #10 #32)
+#34 := [rewrite]: #33
+#37 := [monotonicity #34]: #36
+#40 := [monotonicity #37]: #39
+#46 := [trans #40 #44]: #45
+#49 := [quant-intro #46]: #48
+#52 := [monotonicity #49]: #51
+#92 := [trans #52 #90]: #91
+#31 := [asserted]: #15
+[mp #31 #92]: false
+unsat
+
+7GSX+dyn9XwHWNcjJ4X1ww 1400
+#2 := false
+#9 := (:var 0 int)
+#4 := 2::int
+#10 := (* 2::int #9)
+#7 := 1::int
+#5 := (:var 1 int)
+#6 := (* 2::int #5)
+#8 := (+ #6 1::int)
+#11 := (= #8 #10)
+#12 := (not #11)
+#13 := (forall (vars (?x1 int) (?x2 int)) #12)
+#14 := (not #13)
+#74 := (iff #14 false)
+#31 := (+ 1::int #6)
+#37 := (= #10 #31)
+#42 := (not #37)
+#45 := (forall (vars (?x1 int) (?x2 int)) #42)
+#48 := (not #45)
+#72 := (iff #48 false)
+#1 := true
+#67 := (not true)
+#70 := (iff #67 false)
+#71 := [rewrite]: #70
+#68 := (iff #48 #67)
+#65 := (iff #45 true)
+#60 := (forall (vars (?x1 int) (?x2 int)) true)
+#63 := (iff #60 true)
+#64 := [elim-unused]: #63
+#61 := (iff #45 #60)
+#58 := (iff #42 true)
+#51 := (not false)
+#56 := (iff #51 true)
+#57 := [rewrite]: #56
+#52 := (iff #42 #51)
+#53 := (iff #37 false)
+#54 := [rewrite]: #53
+#55 := [monotonicity #54]: #52
+#59 := [trans #55 #57]: #58
+#62 := [quant-intro #59]: #61
+#66 := [trans #62 #64]: #65
+#69 := [monotonicity #66]: #68
+#73 := [trans #69 #71]: #72
+#49 := (iff #14 #48)
+#46 := (iff #13 #45)
+#43 := (iff #12 #42)
+#40 := (iff #11 #37)
+#34 := (= #31 #10)
+#38 := (iff #34 #37)
+#39 := [rewrite]: #38
+#35 := (iff #11 #34)
+#32 := (= #8 #31)
+#33 := [rewrite]: #32
+#36 := [monotonicity #33]: #35
+#41 := [trans #36 #39]: #40
+#44 := [monotonicity #41]: #43
+#47 := [quant-intro #44]: #46
+#50 := [monotonicity #47]: #49
+#75 := [trans #50 #73]: #74
+#30 := [asserted]: #14
+[mp #30 #75]: false
+unsat
+
+oieid3+1h5s04LTQ9d796Q 2636
+#2 := false
+#4 := 2::int
+decl ?x1!1 :: int
+#85 := ?x1!1
+decl ?x2!0 :: int
+#84 := ?x2!0
+#101 := (+ ?x2!0 ?x1!1)
+#107 := (>= #101 2::int)
+#113 := (<= #101 2::int)
+#116 := (not #113)
+#110 := (not #107)
+#104 := (= #101 2::int)
+#119 := (or #104 #110 #116)
+#122 := (not #119)
+#86 := (+ ?x1!1 ?x2!0)
+#87 := (<= #86 2::int)
+#88 := (not #87)
+#89 := (>= #86 2::int)
+#90 := (not #89)
+#91 := (= #86 2::int)
+#92 := (or #91 #90 #88)
+#93 := (not #92)
+#123 := (iff #93 #122)
+#120 := (iff #92 #119)
+#117 := (iff #88 #116)
+#114 := (iff #87 #113)
+#102 := (= #86 #101)
+#103 := [rewrite]: #102
+#115 := [monotonicity #103]: #114
+#118 := [monotonicity #115]: #117
+#111 := (iff #90 #110)
+#108 := (iff #89 #107)
+#109 := [monotonicity #103]: #108
+#112 := [monotonicity #109]: #111
+#105 := (iff #91 #104)
+#106 := [monotonicity #103]: #105
+#121 := [monotonicity #106 #112 #118]: #120
+#124 := [monotonicity #121]: #123
+#6 := (:var 0 int)
+#5 := (:var 1 int)
+#7 := (+ #5 #6)
+#56 := (<= #7 2::int)
+#58 := (not #56)
+#54 := (>= #7 2::int)
+#51 := (not #54)
+#9 := (= #7 2::int)
+#67 := (or #9 #51 #58)
+#72 := (forall (vars (?x1 int) (?x2 int)) #67)
+#75 := (not #72)
+#94 := (~ #75 #93)
+#95 := [sk]: #94
+#10 := (< #7 2::int)
+#11 := (or #9 #10)
+#8 := (< 2::int #7)
+#12 := (or #8 #11)
+#13 := (forall (vars (?x1 int) (?x2 int)) #12)
+#14 := (not #13)
+#78 := (iff #14 #75)
+#31 := (= 2::int #7)
+#37 := (or #10 #31)
+#42 := (or #8 #37)
+#45 := (forall (vars (?x1 int) (?x2 int)) #42)
+#48 := (not #45)
+#76 := (iff #48 #75)
+#73 := (iff #45 #72)
+#70 := (iff #42 #67)
+#61 := (or #51 #9)
+#64 := (or #58 #61)
+#68 := (iff #64 #67)
+#69 := [rewrite]: #68
+#65 := (iff #42 #64)
+#62 := (iff #37 #61)
+#55 := (iff #31 #9)
+#57 := [rewrite]: #55
+#53 := (iff #10 #51)
+#52 := [rewrite]: #53
+#63 := [monotonicity #52 #57]: #62
+#59 := (iff #8 #58)
+#60 := [rewrite]: #59
+#66 := [monotonicity #60 #63]: #65
+#71 := [trans #66 #69]: #70
+#74 := [quant-intro #71]: #73
+#77 := [monotonicity #74]: #76
+#49 := (iff #14 #48)
+#46 := (iff #13 #45)
+#43 := (iff #12 #42)
+#40 := (iff #11 #37)
+#34 := (or #31 #10)
+#38 := (iff #34 #37)
+#39 := [rewrite]: #38
+#35 := (iff #11 #34)
+#32 := (iff #9 #31)
+#33 := [rewrite]: #32
+#36 := [monotonicity #33]: #35
+#41 := [trans #36 #39]: #40
+#44 := [monotonicity #41]: #43
+#47 := [quant-intro #44]: #46
+#50 := [monotonicity #47]: #49
+#79 := [trans #50 #77]: #78
+#30 := [asserted]: #14
+#80 := [mp #30 #79]: #75
+#98 := [mp~ #80 #95]: #93
+#99 := [mp #98 #124]: #122
+#126 := [not-or-elim #99]: #107
+#100 := (not #104)
+#125 := [not-or-elim #99]: #100
+#127 := [not-or-elim #99]: #113
+#183 := (or #104 #116 #110)
+#184 := [th-lemma]: #183
+[unit-resolution #184 #127 #125 #126]: false
+unsat
+
++RiWXCcHPvuSeYUjZ4Ky/g 2113
+#2 := false
+#4 := 0::int
+decl ?x1!0 :: int
+#78 := ?x1!0
+#83 := (<= ?x1!0 0::int)
+#146 := (not #83)
+#155 := [hypothesis]: #83
+#7 := 1::int
+#81 := (>= ?x1!0 1::int)
+#82 := (not #81)
+#156 := (or #82 #146)
+#157 := [th-lemma]: #156
+#158 := [unit-resolution #157 #155]: #82
+#159 := (or #146 #81)
+#49 := -1::int
+#79 := (<= ?x1!0 -1::int)
+#80 := (not #79)
+#84 := (ite #83 #82 #80)
+#85 := (not #84)
+#5 := (:var 0 int)
+#50 := (<= #5 -1::int)
+#51 := (not #50)
+#55 := (>= #5 1::int)
+#54 := (not #55)
+#45 := (<= #5 0::int)
+#61 := (ite #45 #54 #51)
+#66 := (forall (vars (?x1 int)) #61)
+#69 := (not #66)
+#86 := (~ #69 #85)
+#87 := [sk]: #86
+#10 := (< #5 1::int)
+#8 := (+ #5 1::int)
+#9 := (< 0::int #8)
+#6 := (< 0::int #5)
+#11 := (ite #6 #9 #10)
+#12 := (forall (vars (?x1 int)) #11)
+#13 := (not #12)
+#72 := (iff #13 #69)
+#30 := (+ 1::int #5)
+#33 := (< 0::int #30)
+#36 := (ite #6 #33 #10)
+#39 := (forall (vars (?x1 int)) #36)
+#42 := (not #39)
+#70 := (iff #42 #69)
+#67 := (iff #39 #66)
+#64 := (iff #36 #61)
+#46 := (not #45)
+#58 := (ite #46 #51 #54)
+#62 := (iff #58 #61)
+#63 := [rewrite]: #62
+#59 := (iff #36 #58)
+#56 := (iff #10 #54)
+#57 := [rewrite]: #56
+#52 := (iff #33 #51)
+#53 := [rewrite]: #52
+#47 := (iff #6 #46)
+#48 := [rewrite]: #47
+#60 := [monotonicity #48 #53 #57]: #59
+#65 := [trans #60 #63]: #64
+#68 := [quant-intro #65]: #67
+#71 := [monotonicity #68]: #70
+#43 := (iff #13 #42)
+#40 := (iff #12 #39)
+#37 := (iff #11 #36)
+#34 := (iff #9 #33)
+#31 := (= #8 #30)
+#32 := [rewrite]: #31
+#35 := [monotonicity #32]: #34
+#38 := [monotonicity #35]: #37
+#41 := [quant-intro #38]: #40
+#44 := [monotonicity #41]: #43
+#73 := [trans #44 #71]: #72
+#29 := [asserted]: #13
+#74 := [mp #29 #73]: #69
+#90 := [mp~ #74 #87]: #85
+#151 := (or #84 #146 #81)
+#152 := [def-axiom]: #151
+#160 := [unit-resolution #152 #90]: #159
+#161 := [unit-resolution #160 #158 #155]: false
+#162 := [lemma #161]: #146
+#163 := (or #80 #83)
+#164 := [th-lemma]: #163
+#165 := [unit-resolution #164 #162]: #80
+#166 := (or #83 #79)
+#153 := (or #84 #83 #79)
+#154 := [def-axiom]: #153
+#167 := [unit-resolution #154 #90]: #166
+[unit-resolution #167 #165 #162]: false
+unsat
+
+hlG1vHDJCcXbyvxKYDWifg 2036
+WARNING: failed to find a pattern for quantifier (quantifier id: k!1)
+#2 := false
+#5 := 0::int
+#4 := (:var 0 int)
+#42 := (<= #4 0::int)
+#43 := (not #42)
+#40 := (>= #4 0::int)
+#38 := (not #40)
+#46 := (or #38 #43)
+#49 := (forall (vars (?x1 int)) #46)
+#182 := (not #49)
+#118 := (<= 0::int 0::int)
+#204 := (not #118)
+#119 := (>= 0::int 0::int)
+#205 := (not #119)
+#206 := (or #205 #204)
+#187 := (or #182 #206)
+#172 := (iff #187 #182)
+#183 := (or #182 false)
+#509 := (iff #183 #182)
+#171 := [rewrite]: #509
+#525 := (iff #187 #183)
+#533 := (iff #206 false)
+#529 := (or false false)
+#532 := (iff #529 false)
+#527 := [rewrite]: #532
+#530 := (iff #206 #529)
+#195 := (iff #204 false)
+#1 := true
+#209 := (not true)
+#207 := (iff #209 false)
+#211 := [rewrite]: #207
+#315 := (iff #204 #209)
+#528 := (iff #118 true)
+#184 := [rewrite]: #528
+#522 := [monotonicity #184]: #315
+#196 := [trans #522 #211]: #195
+#190 := (iff #205 false)
+#137 := (iff #205 #209)
+#197 := (iff #119 true)
+#208 := [rewrite]: #197
+#210 := [monotonicity #208]: #137
+#526 := [trans #210 #211]: #190
+#531 := [monotonicity #526 #196]: #530
+#523 := [trans #531 #527]: #533
+#167 := [monotonicity #523]: #525
+#173 := [trans #167 #171]: #172
+#524 := [quant-inst]: #187
+#174 := [mp #524 #173]: #182
+#60 := (~ #49 #49)
+#58 := (~ #46 #46)
+#59 := [refl]: #58
+#61 := [nnf-pos #59]: #60
+#7 := (< 0::int #4)
+#6 := (< #4 0::int)
+#8 := (or #6 #7)
+#9 := (forall (vars (?x1 int)) #8)
+#10 := (ite #9 false true)
+#11 := (not #10)
+#52 := (iff #11 #49)
+#50 := (iff #9 #49)
+#47 := (iff #8 #46)
+#44 := (iff #7 #43)
+#45 := [rewrite]: #44
+#39 := (iff #6 #38)
+#41 := [rewrite]: #39
+#48 := [monotonicity #41 #45]: #47
+#51 := [quant-intro #48]: #50
+#36 := (iff #11 #9)
+#28 := (not #9)
+#31 := (not #28)
+#34 := (iff #31 #9)
+#35 := [rewrite]: #34
+#32 := (iff #11 #31)
+#29 := (iff #10 #28)
+#30 := [rewrite]: #29
+#33 := [monotonicity #30]: #32
+#37 := [trans #33 #35]: #36
+#53 := [trans #37 #51]: #52
+#27 := [asserted]: #11
+#54 := [mp #27 #53]: #49
+#62 := [mp~ #54 #61]: #49
+[unit-resolution #62 #174]: false
+unsat
+
+oOC8ghGUYboMezGio2exAg 4464
+WARNING: failed to find a pattern for quantifier (quantifier id: k!1)
+#2 := false
+#4 := 0::int
+#5 := (:var 0 int)
+#48 := (<= #5 0::int)
+#49 := (not #48)
+#45 := (>= #5 0::int)
+#44 := (not #45)
+#52 := (or #44 #49)
+#55 := (forall (vars (?x1 int)) #52)
+#86 := (not #55)
+#263 := (<= 0::int 0::int)
+#268 := (not #263)
+#604 := (>= 0::int 0::int)
+#264 := (not #604)
+#605 := (or #264 #268)
+#588 := (or #86 #605)
+#584 := (iff #588 #86)
+#311 := (or #86 false)
+#207 := (iff #311 #86)
+#583 := [rewrite]: #207
+#312 := (iff #588 #311)
+#601 := (iff #605 false)
+#599 := (or false false)
+#600 := (iff #599 false)
+#598 := [rewrite]: #600
+#239 := (iff #605 #599)
+#234 := (iff #268 false)
+#1 := true
+#252 := (not true)
+#255 := (iff #252 false)
+#591 := [rewrite]: #255
+#590 := (iff #268 #252)
+#594 := (iff #263 true)
+#595 := [rewrite]: #594
+#596 := [monotonicity #595]: #590
+#597 := [trans #596 #591]: #234
+#592 := (iff #264 false)
+#253 := (iff #264 #252)
+#248 := (iff #604 true)
+#589 := [rewrite]: #248
+#254 := [monotonicity #589]: #253
+#593 := [trans #254 #591]: #592
+#240 := [monotonicity #593 #597]: #239
+#587 := [trans #240 #598]: #601
+#313 := [monotonicity #587]: #312
+#306 := [trans #313 #583]: #584
+#310 := [quant-inst]: #588
+#307 := [mp #310 #306]: #86
+decl z3name!0 :: bool
+#83 := z3name!0
+#12 := 3::int
+#32 := -1::int
+#92 := (ite z3name!0 -1::int 3::int)
+#290 := (= #92 3::int)
+#610 := (not #290)
+#608 := (>= #92 3::int)
+#265 := (not #608)
+#95 := (<= #92 0::int)
+#58 := (ite #55 -1::int 3::int)
+#64 := (<= #58 0::int)
+#96 := (~ #64 #95)
+#93 := (= #58 #92)
+#90 := (~ #55 z3name!0)
+#87 := (or z3name!0 #86)
+#84 := (not z3name!0)
+#85 := (or #84 #55)
+#88 := (and #85 #87)
+#89 := [intro-def]: #88
+#91 := [apply-def #89]: #90
+#94 := [monotonicity #91]: #93
+#97 := [monotonicity #94]: #96
+#10 := 1::int
+#11 := (- 1::int)
+#7 := (< 0::int #5)
+#6 := (< #5 0::int)
+#8 := (or #6 #7)
+#9 := (forall (vars (?x1 int)) #8)
+#13 := (ite #9 #11 3::int)
+#14 := (< 0::int #13)
+#15 := (not #14)
+#77 := (iff #15 #64)
+#35 := (ite #9 -1::int 3::int)
+#38 := (< 0::int #35)
+#41 := (not #38)
+#75 := (iff #41 #64)
+#65 := (not #64)
+#70 := (not #65)
+#73 := (iff #70 #64)
+#74 := [rewrite]: #73
+#71 := (iff #41 #70)
+#68 := (iff #38 #65)
+#61 := (< 0::int #58)
+#66 := (iff #61 #65)
+#67 := [rewrite]: #66
+#62 := (iff #38 #61)
+#59 := (= #35 #58)
+#56 := (iff #9 #55)
+#53 := (iff #8 #52)
+#50 := (iff #7 #49)
+#51 := [rewrite]: #50
+#46 := (iff #6 #44)
+#47 := [rewrite]: #46
+#54 := [monotonicity #47 #51]: #53
+#57 := [quant-intro #54]: #56
+#60 := [monotonicity #57]: #59
+#63 := [monotonicity #60]: #62
+#69 := [trans #63 #67]: #68
+#72 := [monotonicity #69]: #71
+#76 := [trans #72 #74]: #75
+#42 := (iff #15 #41)
+#39 := (iff #14 #38)
+#36 := (= #13 #35)
+#33 := (= #11 -1::int)
+#34 := [rewrite]: #33
+#37 := [monotonicity #34]: #36
+#40 := [monotonicity #37]: #39
+#43 := [monotonicity #40]: #42
+#78 := [trans #43 #76]: #77
+#31 := [asserted]: #15
+#79 := [mp #31 #78]: #64
+#126 := [mp~ #79 #97]: #95
+#395 := (not #95)
+#602 := (or #265 #395)
+#276 := [th-lemma]: #602
+#277 := [unit-resolution #276 #126]: #265
+#609 := [hypothesis]: #290
+#611 := (or #610 #608)
+#612 := [th-lemma]: #611
+#607 := [unit-resolution #612 #609 #277]: false
+#613 := [lemma #607]: #610
+#292 := (or z3name!0 #290)
+#271 := [def-axiom]: #292
+#581 := [unit-resolution #271 #613]: z3name!0
+#129 := (or #55 #84)
+decl ?x1!1 :: int
+#108 := ?x1!1
+#111 := (>= ?x1!1 0::int)
+#112 := (not #111)
+#109 := (<= ?x1!1 0::int)
+#110 := (not #109)
+#132 := (or #110 #112)
+#135 := (not #132)
+#138 := (or z3name!0 #135)
+#141 := (and #129 #138)
+#113 := (or #112 #110)
+#114 := (not #113)
+#119 := (or z3name!0 #114)
+#122 := (and #85 #119)
+#142 := (iff #122 #141)
+#139 := (iff #119 #138)
+#136 := (iff #114 #135)
+#133 := (iff #113 #132)
+#134 := [rewrite]: #133
+#137 := [monotonicity #134]: #136
+#140 := [monotonicity #137]: #139
+#130 := (iff #85 #129)
+#131 := [rewrite]: #130
+#143 := [monotonicity #131 #140]: #142
+#123 := (~ #88 #122)
+#120 := (~ #87 #119)
+#115 := (~ #86 #114)
+#116 := [sk]: #115
+#106 := (~ z3name!0 z3name!0)
+#107 := [refl]: #106
+#121 := [monotonicity #107 #116]: #120
+#104 := (~ #85 #85)
+#102 := (~ #55 #55)
+#100 := (~ #52 #52)
+#101 := [refl]: #100
+#103 := [nnf-pos #101]: #102
+#98 := (~ #84 #84)
+#99 := [refl]: #98
+#105 := [monotonicity #99 #103]: #104
+#124 := [monotonicity #105 #121]: #123
+#125 := [mp~ #89 #124]: #122
+#127 := [mp #125 #143]: #141
+#128 := [and-elim #127]: #129
+#585 := [unit-resolution #128 #581]: #55
+[unit-resolution #585 #307]: false
+unsat
+
+4Dtb5Y1TTRPWZcHG9Griog 1594
+#2 := false
+#12 := 1::int
+#9 := (:var 1 int)
+#7 := 6::int
+#8 := (- 6::int)
+#10 := (* #8 #9)
+#5 := (:var 2 int)
+#4 := 4::int
+#6 := (* 4::int #5)
+#11 := (+ #6 #10)
+#13 := (= #11 1::int)
+#14 := (exists (vars (?x1 int) (?x2 int) (?x3 int)) #13)
+#15 := (not #14)
+#16 := (not #15)
+#82 := (iff #16 false)
+#53 := (:var 0 int)
+#33 := -6::int
+#54 := (* -6::int #53)
+#55 := (* 4::int #9)
+#56 := (+ #55 #54)
+#57 := (= 1::int #56)
+#58 := (exists (vars (?x1 int) (?x2 int)) #57)
+#80 := (iff #58 false)
+#76 := (exists (vars (?x1 int) (?x2 int)) false)
+#78 := (iff #76 false)
+#79 := [elim-unused]: #78
+#77 := (iff #58 #76)
+#73 := (iff #57 false)
+#74 := [rewrite]: #73
+#75 := [quant-intro #74]: #77
+#81 := [trans #75 #79]: #80
+#71 := (iff #16 #58)
+#63 := (not #58)
+#66 := (not #63)
+#69 := (iff #66 #58)
+#70 := [rewrite]: #69
+#67 := (iff #16 #66)
+#64 := (iff #15 #63)
+#61 := (iff #14 #58)
+#36 := (* -6::int #9)
+#39 := (+ #6 #36)
+#45 := (= 1::int #39)
+#50 := (exists (vars (?x1 int) (?x2 int) (?x3 int)) #45)
+#59 := (iff #50 #58)
+#60 := [elim-unused]: #59
+#51 := (iff #14 #50)
+#48 := (iff #13 #45)
+#42 := (= #39 1::int)
+#46 := (iff #42 #45)
+#47 := [rewrite]: #46
+#43 := (iff #13 #42)
+#40 := (= #11 #39)
+#37 := (= #10 #36)
+#34 := (= #8 -6::int)
+#35 := [rewrite]: #34
+#38 := [monotonicity #35]: #37
+#41 := [monotonicity #38]: #40
+#44 := [monotonicity #41]: #43
+#49 := [trans #44 #47]: #48
+#52 := [quant-intro #49]: #51
+#62 := [trans #52 #60]: #61
+#65 := [monotonicity #62]: #64
+#68 := [monotonicity #65]: #67
+#72 := [trans #68 #70]: #71
+#83 := [trans #72 #81]: #82
+#32 := [asserted]: #16
+[mp #32 #83]: false
+unsat
+
+dbOre63OdVavsqL3lG2ttw 2516
+#2 := false
+#4 := 0::int
+decl ?x2!1 :: int
+#83 := ?x2!1
+decl ?x3!0 :: int
+#82 := ?x3!0
+#108 := (+ ?x3!0 ?x2!1)
+#111 := (<= #108 0::int)
+#114 := (not #111)
+#89 := (<= ?x2!1 0::int)
+#90 := (not #89)
+#87 := (<= ?x3!0 0::int)
+#88 := (not #87)
+#102 := (and #88 #90)
+#105 := (not #102)
+#117 := (or #105 #114)
+#120 := (not #117)
+#84 := (+ ?x2!1 ?x3!0)
+#85 := (<= #84 0::int)
+#86 := (not #85)
+#91 := (and #90 #88)
+#92 := (not #91)
+#93 := (or #92 #86)
+#94 := (not #93)
+#121 := (iff #94 #120)
+#118 := (iff #93 #117)
+#115 := (iff #86 #114)
+#112 := (iff #85 #111)
+#109 := (= #84 #108)
+#110 := [rewrite]: #109
+#113 := [monotonicity #110]: #112
+#116 := [monotonicity #113]: #115
+#106 := (iff #92 #105)
+#103 := (iff #91 #102)
+#104 := [rewrite]: #103
+#107 := [monotonicity #104]: #106
+#119 := [monotonicity #107 #116]: #118
+#122 := [monotonicity #119]: #121
+#7 := (:var 0 int)
+#5 := (:var 1 int)
+#10 := (+ #5 #7)
+#63 := (<= #10 0::int)
+#64 := (not #63)
+#53 := (<= #7 0::int)
+#54 := (not #53)
+#49 := (<= #5 0::int)
+#50 := (not #49)
+#57 := (and #50 #54)
+#60 := (not #57)
+#67 := (or #60 #64)
+#70 := (forall (vars (?x2 int) (?x3 int)) #67)
+#73 := (not #70)
+#95 := (~ #73 #94)
+#96 := [sk]: #95
+#11 := (< 0::int #10)
+#8 := (< 0::int #7)
+#6 := (< 0::int #5)
+#9 := (and #6 #8)
+#12 := (implies #9 #11)
+#13 := (forall (vars (?x2 int) (?x3 int)) #12)
+#14 := (exists (vars (?x1 int)) #13)
+#15 := (not #14)
+#76 := (iff #15 #73)
+#32 := (not #9)
+#33 := (or #32 #11)
+#36 := (forall (vars (?x2 int) (?x3 int)) #33)
+#46 := (not #36)
+#74 := (iff #46 #73)
+#71 := (iff #36 #70)
+#68 := (iff #33 #67)
+#65 := (iff #11 #64)
+#66 := [rewrite]: #65
+#61 := (iff #32 #60)
+#58 := (iff #9 #57)
+#55 := (iff #8 #54)
+#56 := [rewrite]: #55
+#51 := (iff #6 #50)
+#52 := [rewrite]: #51
+#59 := [monotonicity #52 #56]: #58
+#62 := [monotonicity #59]: #61
+#69 := [monotonicity #62 #66]: #68
+#72 := [quant-intro #69]: #71
+#75 := [monotonicity #72]: #74
+#47 := (iff #15 #46)
+#44 := (iff #14 #36)
+#39 := (exists (vars (?x1 int)) #36)
+#42 := (iff #39 #36)
+#43 := [elim-unused]: #42
+#40 := (iff #14 #39)
+#37 := (iff #13 #36)
+#34 := (iff #12 #33)
+#35 := [rewrite]: #34
+#38 := [quant-intro #35]: #37
+#41 := [quant-intro #38]: #40
+#45 := [trans #41 #43]: #44
+#48 := [monotonicity #45]: #47
+#77 := [trans #48 #75]: #76
+#31 := [asserted]: #15
+#78 := [mp #31 #77]: #73
+#99 := [mp~ #78 #96]: #94
+#100 := [mp #99 #122]: #120
+#125 := [not-or-elim #100]: #111
+#101 := [not-or-elim #100]: #102
+#124 := [and-elim #101]: #90
+#123 := [and-elim #101]: #88
+[th-lemma #123 #124 #125]: false
+unsat
+
+LtM5szEGH9QAF1TwsVtH4w 2764
+#2 := false
+#4 := 0::int
+decl ?x2!1 :: int
+#91 := ?x2!1
+#98 := (<= ?x2!1 0::int)
+#99 := (not #98)
+#7 := 0::real
+decl ?x3!0 :: real
+#93 := ?x3!0
+#96 := (<= ?x3!0 0::real)
+#97 := (not #96)
+#111 := (and #97 #99)
+#114 := (not #111)
+#33 := -1::int
+#94 := (<= ?x2!1 -1::int)
+#95 := (not #94)
+#120 := (or #95 #114)
+#125 := (not #120)
+#100 := (and #99 #97)
+#101 := (not #100)
+#102 := (or #101 #95)
+#103 := (not #102)
+#126 := (iff #103 #125)
+#123 := (iff #102 #120)
+#117 := (or #114 #95)
+#121 := (iff #117 #120)
+#122 := [rewrite]: #121
+#118 := (iff #102 #117)
+#115 := (iff #101 #114)
+#112 := (iff #100 #111)
+#113 := [rewrite]: #112
+#116 := [monotonicity #113]: #115
+#119 := [monotonicity #116]: #118
+#124 := [trans #119 #122]: #123
+#127 := [monotonicity #124]: #126
+#5 := (:var 1 int)
+#75 := (<= #5 -1::int)
+#76 := (not #75)
+#8 := (:var 0 real)
+#65 := (<= #8 0::real)
+#66 := (not #65)
+#61 := (<= #5 0::int)
+#62 := (not #61)
+#69 := (and #62 #66)
+#72 := (not #69)
+#79 := (or #72 #76)
+#82 := (forall (vars (?x2 int) (?x3 real)) #79)
+#85 := (not #82)
+#104 := (~ #85 #103)
+#105 := [sk]: #104
+#11 := 1::int
+#12 := (- 1::int)
+#13 := (< #12 #5)
+#9 := (< 0::real #8)
+#6 := (< 0::int #5)
+#10 := (and #6 #9)
+#14 := (implies #10 #13)
+#15 := (forall (vars (?x2 int) (?x3 real)) #14)
+#16 := (exists (vars (?x1 int)) #15)
+#17 := (not #16)
+#88 := (iff #17 #85)
+#36 := (< -1::int #5)
+#42 := (not #10)
+#43 := (or #42 #36)
+#48 := (forall (vars (?x2 int) (?x3 real)) #43)
+#58 := (not #48)
+#86 := (iff #58 #85)
+#83 := (iff #48 #82)
+#80 := (iff #43 #79)
+#77 := (iff #36 #76)
+#78 := [rewrite]: #77
+#73 := (iff #42 #72)
+#70 := (iff #10 #69)
+#67 := (iff #9 #66)
+#68 := [rewrite]: #67
+#63 := (iff #6 #62)
+#64 := [rewrite]: #63
+#71 := [monotonicity #64 #68]: #70
+#74 := [monotonicity #71]: #73
+#81 := [monotonicity #74 #78]: #80
+#84 := [quant-intro #81]: #83
+#87 := [monotonicity #84]: #86
+#59 := (iff #17 #58)
+#56 := (iff #16 #48)
+#51 := (exists (vars (?x1 int)) #48)
+#54 := (iff #51 #48)
+#55 := [elim-unused]: #54
+#52 := (iff #16 #51)
+#49 := (iff #15 #48)
+#46 := (iff #14 #43)
+#39 := (implies #10 #36)
+#44 := (iff #39 #43)
+#45 := [rewrite]: #44
+#40 := (iff #14 #39)
+#37 := (iff #13 #36)
+#34 := (= #12 -1::int)
+#35 := [rewrite]: #34
+#38 := [monotonicity #35]: #37
+#41 := [monotonicity #38]: #40
+#47 := [trans #41 #45]: #46
+#50 := [quant-intro #47]: #49
+#53 := [quant-intro #50]: #52
+#57 := [trans #53 #55]: #56
+#60 := [monotonicity #57]: #59
+#89 := [trans #60 #87]: #88
+#32 := [asserted]: #17
+#90 := [mp #32 #89]: #85
+#108 := [mp~ #90 #105]: #103
+#109 := [mp #108 #127]: #125
+#128 := [not-or-elim #109]: #111
+#130 := [and-elim #128]: #99
+#110 := [not-or-elim #109]: #94
+#186 := (or #95 #98)
+#187 := [th-lemma]: #186
+#188 := [unit-resolution #187 #110]: #98
+[unit-resolution #188 #130]: false
+unsat
+
+ibIqbnIUB+oyERADdbFW6w 3631
+#2 := false
+#144 := (not false)
+#7 := 0::int
+#5 := (:var 0 int)
+#52 := (<= #5 0::int)
+#53 := (not #52)
+#147 := (or #53 #144)
+#150 := (not #147)
+#153 := (forall (vars (?x1 int)) #150)
+#180 := (iff #153 false)
+#175 := (forall (vars (?x1 int)) false)
+#178 := (iff #175 false)
+#179 := [elim-unused]: #178
+#176 := (iff #153 #175)
+#173 := (iff #150 false)
+#1 := true
+#168 := (not true)
+#171 := (iff #168 false)
+#172 := [rewrite]: #171
+#169 := (iff #150 #168)
+#166 := (iff #147 true)
+#161 := (or #53 true)
+#164 := (iff #161 true)
+#165 := [rewrite]: #164
+#162 := (iff #147 #161)
+#159 := (iff #144 true)
+#160 := [rewrite]: #159
+#163 := [monotonicity #160]: #162
+#167 := [trans #163 #165]: #166
+#170 := [monotonicity #167]: #169
+#174 := [trans #170 #172]: #173
+#177 := [quant-intro #174]: #176
+#181 := [trans #177 #179]: #180
+#56 := -1::int
+#57 := (* -1::int #5)
+#4 := (:var 1 int)
+#58 := (+ #4 #57)
+#59 := (<= #58 0::int)
+#62 := (not #59)
+#68 := (or #53 #62)
+#73 := (forall (vars (?x2 int)) #68)
+#76 := (not #73)
+#79 := (or #53 #76)
+#105 := (not #79)
+#123 := (forall (vars (?x1 int)) #105)
+#156 := (iff #123 #153)
+#127 := (forall (vars (?x2 int)) #53)
+#130 := (not #127)
+#133 := (or #53 #130)
+#136 := (not #133)
+#139 := (forall (vars (?x1 int)) #136)
+#154 := (iff #139 #153)
+#155 := [rewrite]: #154
+#140 := (iff #123 #139)
+#141 := [rewrite]: #140
+#157 := [trans #141 #155]: #156
+#116 := (and #52 #73)
+#119 := (forall (vars (?x1 int)) #116)
+#124 := (iff #119 #123)
+#113 := (iff #116 #105)
+#122 := [rewrite]: #113
+#125 := [quant-intro #122]: #124
+#94 := (not #53)
+#104 := (and #94 #73)
+#108 := (forall (vars (?x1 int)) #104)
+#120 := (iff #108 #119)
+#117 := (iff #104 #116)
+#114 := (iff #94 #52)
+#115 := [rewrite]: #114
+#118 := [monotonicity #115]: #117
+#121 := [quant-intro #118]: #120
+#82 := (exists (vars (?x1 int)) #79)
+#85 := (not #82)
+#109 := (~ #85 #108)
+#106 := (~ #105 #104)
+#101 := (not #76)
+#102 := (~ #101 #73)
+#99 := (~ #73 #73)
+#97 := (~ #68 #68)
+#98 := [refl]: #97
+#100 := [nnf-pos #98]: #99
+#103 := [nnf-neg #100]: #102
+#95 := (~ #94 #94)
+#96 := [refl]: #95
+#107 := [nnf-neg #96 #103]: #106
+#110 := [nnf-neg #107]: #109
+#8 := (< 0::int #5)
+#6 := (<= #4 #5)
+#9 := (implies #6 #8)
+#10 := (forall (vars (?x2 int)) #9)
+#11 := (implies #10 #8)
+#12 := (exists (vars (?x1 int)) #11)
+#13 := (not #12)
+#88 := (iff #13 #85)
+#30 := (not #6)
+#31 := (or #30 #8)
+#34 := (forall (vars (?x2 int)) #31)
+#40 := (not #34)
+#41 := (or #8 #40)
+#46 := (exists (vars (?x1 int)) #41)
+#49 := (not #46)
+#86 := (iff #49 #85)
+#83 := (iff #46 #82)
+#80 := (iff #41 #79)
+#77 := (iff #40 #76)
+#74 := (iff #34 #73)
+#71 := (iff #31 #68)
+#65 := (or #62 #53)
+#69 := (iff #65 #68)
+#70 := [rewrite]: #69
+#66 := (iff #31 #65)
+#54 := (iff #8 #53)
+#55 := [rewrite]: #54
+#63 := (iff #30 #62)
+#60 := (iff #6 #59)
+#61 := [rewrite]: #60
+#64 := [monotonicity #61]: #63
+#67 := [monotonicity #64 #55]: #66
+#72 := [trans #67 #70]: #71
+#75 := [quant-intro #72]: #74
+#78 := [monotonicity #75]: #77
+#81 := [monotonicity #55 #78]: #80
+#84 := [quant-intro #81]: #83
+#87 := [monotonicity #84]: #86
+#50 := (iff #13 #49)
+#47 := (iff #12 #46)
+#44 := (iff #11 #41)
+#37 := (implies #34 #8)
+#42 := (iff #37 #41)
+#43 := [rewrite]: #42
+#38 := (iff #11 #37)
+#35 := (iff #10 #34)
+#32 := (iff #9 #31)
+#33 := [rewrite]: #32
+#36 := [quant-intro #33]: #35
+#39 := [monotonicity #36]: #38
+#45 := [trans #39 #43]: #44
+#48 := [quant-intro #45]: #47
+#51 := [monotonicity #48]: #50
+#89 := [trans #51 #87]: #88
+#29 := [asserted]: #13
+#90 := [mp #29 #89]: #85
+#111 := [mp~ #90 #110]: #108
+#112 := [mp #111 #121]: #119
+#126 := [mp #112 #125]: #123
+#158 := [mp #126 #157]: #153
+[mp #158 #181]: false
+unsat
+
+1HbJvLWS/aG8IZEVLDIWyA 1506
+#2 := false
+#4 := (:var 0 int)
+#5 := (pattern #4)
+decl uf_1 :: int
+#6 := uf_1
+#8 := 2::int
+#10 := (* 2::int uf_1)
+#9 := (* 2::int #4)
+#11 := (< #9 #10)
+#7 := (< #4 uf_1)
+#12 := (implies #7 #11)
+#13 := (forall (vars (?x1 int)) (:pat #5) #12)
+#14 := (not #13)
+#79 := (iff #14 false)
+#31 := (not #7)
+#32 := (or #31 #11)
+#35 := (forall (vars (?x1 int)) (:pat #5) #32)
+#38 := (not #35)
+#77 := (iff #38 false)
+#1 := true
+#72 := (not true)
+#75 := (iff #72 false)
+#76 := [rewrite]: #75
+#73 := (iff #38 #72)
+#70 := (iff #35 true)
+#65 := (forall (vars (?x1 int)) (:pat #5) true)
+#68 := (iff #65 true)
+#69 := [elim-unused]: #68
+#66 := (iff #35 #65)
+#63 := (iff #32 true)
+#43 := 0::int
+#41 := -1::int
+#45 := (* -1::int uf_1)
+#46 := (+ #4 #45)
+#44 := (>= #46 0::int)
+#42 := (not #44)
+#57 := (or #44 #42)
+#61 := (iff #57 true)
+#62 := [rewrite]: #61
+#59 := (iff #32 #57)
+#58 := (iff #11 #42)
+#56 := [rewrite]: #58
+#54 := (iff #31 #44)
+#49 := (not #42)
+#52 := (iff #49 #44)
+#53 := [rewrite]: #52
+#50 := (iff #31 #49)
+#47 := (iff #7 #42)
+#48 := [rewrite]: #47
+#51 := [monotonicity #48]: #50
+#55 := [trans #51 #53]: #54
+#60 := [monotonicity #55 #56]: #59
+#64 := [trans #60 #62]: #63
+#67 := [quant-intro #64]: #66
+#71 := [trans #67 #69]: #70
+#74 := [monotonicity #71]: #73
+#78 := [trans #74 #76]: #77
+#39 := (iff #14 #38)
+#36 := (iff #13 #35)
+#33 := (iff #12 #32)
+#34 := [rewrite]: #33
+#37 := [quant-intro #34]: #36
+#40 := [monotonicity #37]: #39
+#80 := [trans #40 #78]: #79
+#30 := [asserted]: #14
+[mp #30 #80]: false
+unsat
+
+K7kWge9RT44bPFRy+hxaqg 7
+unsat
+
++rwKUm5bOzD9paEkmogLyw 1562
+#2 := false
+#6 := 1::int
+decl uf_3 :: int
+#8 := uf_3
+#12 := (+ uf_3 1::int)
+decl uf_1 :: int
+#4 := uf_1
+#13 := (* uf_1 #12)
+decl uf_2 :: int
+#5 := uf_2
+#11 := (* uf_1 uf_2)
+#14 := (+ #11 #13)
+#7 := (+ uf_2 1::int)
+#9 := (+ #7 uf_3)
+#10 := (* uf_1 #9)
+#15 := (= #10 #14)
+#16 := (not #15)
+#85 := (iff #16 false)
+#1 := true
+#80 := (not true)
+#83 := (iff #80 false)
+#84 := [rewrite]: #83
+#81 := (iff #16 #80)
+#78 := (iff #15 true)
+#48 := (* uf_1 uf_3)
+#49 := (+ #11 #48)
+#50 := (+ uf_1 #49)
+#73 := (= #50 #50)
+#76 := (iff #73 true)
+#77 := [rewrite]: #76
+#74 := (iff #15 #73)
+#71 := (= #14 #50)
+#61 := (+ uf_1 #48)
+#66 := (+ #11 #61)
+#69 := (= #66 #50)
+#70 := [rewrite]: #69
+#67 := (= #14 #66)
+#64 := (= #13 #61)
+#55 := (+ 1::int uf_3)
+#58 := (* uf_1 #55)
+#62 := (= #58 #61)
+#63 := [rewrite]: #62
+#59 := (= #13 #58)
+#56 := (= #12 #55)
+#57 := [rewrite]: #56
+#60 := [monotonicity #57]: #59
+#65 := [trans #60 #63]: #64
+#68 := [monotonicity #65]: #67
+#72 := [trans #68 #70]: #71
+#53 := (= #10 #50)
+#39 := (+ uf_2 uf_3)
+#40 := (+ 1::int #39)
+#45 := (* uf_1 #40)
+#51 := (= #45 #50)
+#52 := [rewrite]: #51
+#46 := (= #10 #45)
+#43 := (= #9 #40)
+#33 := (+ 1::int uf_2)
+#36 := (+ #33 uf_3)
+#41 := (= #36 #40)
+#42 := [rewrite]: #41
+#37 := (= #9 #36)
+#34 := (= #7 #33)
+#35 := [rewrite]: #34
+#38 := [monotonicity #35]: #37
+#44 := [trans #38 #42]: #43
+#47 := [monotonicity #44]: #46
+#54 := [trans #47 #52]: #53
+#75 := [monotonicity #54 #72]: #74
+#79 := [trans #75 #77]: #78
+#82 := [monotonicity #79]: #81
+#86 := [trans #82 #84]: #85
+#32 := [asserted]: #16
+[mp #32 #86]: false
+unsat
+
+iRJ30NP1Enylq9tZfpCPTA 1288
+#2 := false
+decl uf_2 :: real
+#6 := uf_2
+decl uf_1 :: real
+#4 := uf_1
+#12 := 2::real
+#13 := (* 2::real uf_1)
+#14 := (* #13 uf_2)
+#5 := 1::real
+#9 := (- 1::real uf_2)
+#10 := (* uf_1 #9)
+#7 := (+ 1::real uf_2)
+#8 := (* uf_1 #7)
+#11 := (- #8 #10)
+#15 := (= #11 #14)
+#16 := (not #15)
+#73 := (iff #16 false)
+#1 := true
+#68 := (not true)
+#71 := (iff #68 false)
+#72 := [rewrite]: #71
+#69 := (iff #16 #68)
+#66 := (iff #15 true)
+#33 := (* uf_1 uf_2)
+#55 := (* 2::real #33)
+#61 := (= #55 #55)
+#64 := (iff #61 true)
+#65 := [rewrite]: #64
+#62 := (iff #15 #61)
+#59 := (= #14 #55)
+#60 := [rewrite]: #59
+#57 := (= #11 #55)
+#37 := -1::real
+#45 := (* -1::real #33)
+#46 := (+ uf_1 #45)
+#34 := (+ uf_1 #33)
+#51 := (- #34 #46)
+#54 := (= #51 #55)
+#56 := [rewrite]: #54
+#52 := (= #11 #51)
+#49 := (= #10 #46)
+#38 := (* -1::real uf_2)
+#39 := (+ 1::real #38)
+#42 := (* uf_1 #39)
+#47 := (= #42 #46)
+#48 := [rewrite]: #47
+#43 := (= #10 #42)
+#40 := (= #9 #39)
+#41 := [rewrite]: #40
+#44 := [monotonicity #41]: #43
+#50 := [trans #44 #48]: #49
+#35 := (= #8 #34)
+#36 := [rewrite]: #35
+#53 := [monotonicity #36 #50]: #52
+#58 := [trans #53 #56]: #57
+#63 := [monotonicity #58 #60]: #62
+#67 := [trans #63 #65]: #66
+#70 := [monotonicity #67]: #69
+#74 := [trans #70 #72]: #73
+#32 := [asserted]: #16
+[mp #32 #74]: false
+unsat
+
+Ff1vqDiuUnet19j/x+mXkA 3029
+#2 := false
+decl uf_4 :: int
+#9 := uf_4
+decl uf_5 :: int
+#13 := uf_5
+decl uf_3 :: int
+#8 := uf_3
+#24 := (+ uf_3 uf_5)
+#25 := (+ #24 uf_4)
+decl uf_2 :: int
+#6 := uf_2
+#5 := 1::int
+#7 := (+ 1::int uf_2)
+#26 := (* #7 #25)
+#21 := (* uf_5 uf_2)
+#19 := (* #7 uf_5)
+#10 := (+ uf_3 uf_4)
+#16 := 2::int
+#17 := (* 2::int #7)
+#18 := (* #17 #10)
+#20 := (+ #18 #19)
+#22 := (+ #20 #21)
+decl uf_1 :: int
+#4 := uf_1
+#23 := (+ uf_1 #22)
+#27 := (- #23 #26)
+#14 := (* uf_2 uf_5)
+#11 := (* #7 #10)
+#12 := (+ uf_1 #11)
+#15 := (+ #12 #14)
+#28 := (= #15 #27)
+#29 := (not #28)
+#149 := (iff #29 false)
+#1 := true
+#144 := (not true)
+#147 := (iff #144 false)
+#148 := [rewrite]: #147
+#145 := (iff #29 #144)
+#142 := (iff #28 true)
+#47 := (* uf_2 uf_4)
+#46 := (* uf_2 uf_3)
+#48 := (+ #46 #47)
+#59 := (+ #14 #48)
+#60 := (+ uf_4 #59)
+#61 := (+ uf_3 #60)
+#62 := (+ uf_1 #61)
+#136 := (= #62 #62)
+#140 := (iff #136 true)
+#141 := [rewrite]: #140
+#135 := (iff #28 #136)
+#138 := (= #27 #62)
+#123 := (+ uf_5 #59)
+#124 := (+ uf_4 #123)
+#125 := (+ uf_3 #124)
+#77 := (* 2::int #47)
+#75 := (* 2::int #46)
+#78 := (+ #75 #77)
+#104 := (* 2::int #14)
+#105 := (+ #104 #78)
+#106 := (+ uf_5 #105)
+#76 := (* 2::int uf_4)
+#107 := (+ #76 #106)
+#74 := (* 2::int uf_3)
+#108 := (+ #74 #107)
+#113 := (+ uf_1 #108)
+#130 := (- #113 #125)
+#133 := (= #130 #62)
+#139 := [rewrite]: #133
+#131 := (= #27 #130)
+#128 := (= #26 #125)
+#116 := (+ uf_4 uf_5)
+#117 := (+ uf_3 #116)
+#120 := (* #7 #117)
+#126 := (= #120 #125)
+#127 := [rewrite]: #126
+#121 := (= #26 #120)
+#118 := (= #25 #117)
+#119 := [rewrite]: #118
+#122 := [monotonicity #119]: #121
+#129 := [trans #122 #127]: #128
+#114 := (= #23 #113)
+#111 := (= #22 #108)
+#91 := (+ #14 #78)
+#92 := (+ uf_5 #91)
+#93 := (+ #76 #92)
+#94 := (+ #74 #93)
+#101 := (+ #94 #14)
+#109 := (= #101 #108)
+#110 := [rewrite]: #109
+#102 := (= #22 #101)
+#99 := (= #21 #14)
+#100 := [rewrite]: #99
+#97 := (= #20 #94)
+#85 := (+ uf_5 #14)
+#79 := (+ #76 #78)
+#80 := (+ #74 #79)
+#88 := (+ #80 #85)
+#95 := (= #88 #94)
+#96 := [rewrite]: #95
+#89 := (= #20 #88)
+#86 := (= #19 #85)
+#87 := [rewrite]: #86
+#83 := (= #18 #80)
+#67 := (* 2::int uf_2)
+#68 := (+ 2::int #67)
+#71 := (* #68 #10)
+#81 := (= #71 #80)
+#82 := [rewrite]: #81
+#72 := (= #18 #71)
+#69 := (= #17 #68)
+#70 := [rewrite]: #69
+#73 := [monotonicity #70]: #72
+#84 := [trans #73 #82]: #83
+#90 := [monotonicity #84 #87]: #89
+#98 := [trans #90 #96]: #97
+#103 := [monotonicity #98 #100]: #102
+#112 := [trans #103 #110]: #111
+#115 := [monotonicity #112]: #114
+#132 := [monotonicity #115 #129]: #131
+#137 := [trans #132 #139]: #138
+#65 := (= #15 #62)
+#49 := (+ uf_4 #48)
+#50 := (+ uf_3 #49)
+#53 := (+ uf_1 #50)
+#56 := (+ #53 #14)
+#63 := (= #56 #62)
+#64 := [rewrite]: #63
+#57 := (= #15 #56)
+#54 := (= #12 #53)
+#51 := (= #11 #50)
+#52 := [rewrite]: #51
+#55 := [monotonicity #52]: #54
+#58 := [monotonicity #55]: #57
+#66 := [trans #58 #64]: #65
+#134 := [monotonicity #66 #137]: #135
+#143 := [trans #134 #141]: #142
+#146 := [monotonicity #143]: #145
+#150 := [trans #146 #148]: #149
+#45 := [asserted]: #29
+[mp #45 #150]: false
+unsat
+
+iPZ87GNdi7uQw4ehEpbTPQ 6383
+#2 := false
+#9 := 0::int
+decl uf_2 :: (-> T1 int)
+decl uf_1 :: (-> int T1)
+decl uf_3 :: T1
+#22 := uf_3
+#23 := (uf_2 uf_3)
+#21 := 2::int
+#24 := (* 2::int #23)
+#25 := (uf_1 #24)
+#293 := (uf_2 #25)
+#295 := -1::int
+#274 := (* -1::int #293)
+#610 := (+ #24 #274)
+#594 := (<= #610 0::int)
+#612 := (= #610 0::int)
+#606 := (>= #23 0::int)
+#237 := (= #293 0::int)
+#549 := (not #237)
+#588 := (<= #293 0::int)
+#457 := (not #588)
+#26 := 1::int
+#558 := (>= #293 1::int)
+#555 := (= #293 1::int)
+#27 := (uf_1 1::int)
+#589 := (uf_2 #27)
+#301 := (= #589 1::int)
+#10 := (:var 0 int)
+#12 := (uf_1 #10)
+#626 := (pattern #12)
+#70 := (>= #10 0::int)
+#71 := (not #70)
+#13 := (uf_2 #12)
+#52 := (= #10 #13)
+#77 := (or #52 #71)
+#627 := (forall (vars (?x2 int)) (:pat #626) #77)
+#82 := (forall (vars (?x2 int)) #77)
+#630 := (iff #82 #627)
+#628 := (iff #77 #77)
+#629 := [refl]: #628
+#631 := [quant-intro #629]: #630
+#132 := (~ #82 #82)
+#144 := (~ #77 #77)
+#145 := [refl]: #144
+#130 := [nnf-pos #145]: #132
+#14 := (= #13 #10)
+#11 := (<= 0::int #10)
+#15 := (implies #11 #14)
+#16 := (forall (vars (?x2 int)) #15)
+#85 := (iff #16 #82)
+#59 := (not #11)
+#60 := (or #59 #52)
+#65 := (forall (vars (?x2 int)) #60)
+#83 := (iff #65 #82)
+#80 := (iff #60 #77)
+#74 := (or #71 #52)
+#78 := (iff #74 #77)
+#79 := [rewrite]: #78
+#75 := (iff #60 #74)
+#72 := (iff #59 #71)
+#68 := (iff #11 #70)
+#69 := [rewrite]: #68
+#73 := [monotonicity #69]: #72
+#76 := [monotonicity #73]: #75
+#81 := [trans #76 #79]: #80
+#84 := [quant-intro #81]: #83
+#66 := (iff #16 #65)
+#63 := (iff #15 #60)
+#56 := (implies #11 #52)
+#61 := (iff #56 #60)
+#62 := [rewrite]: #61
+#57 := (iff #15 #56)
+#54 := (iff #14 #52)
+#55 := [rewrite]: #54
+#58 := [monotonicity #55]: #57
+#64 := [trans #58 #62]: #63
+#67 := [quant-intro #64]: #66
+#86 := [trans #67 #84]: #85
+#51 := [asserted]: #16
+#87 := [mp #51 #86]: #82
+#146 := [mp~ #87 #130]: #82
+#632 := [mp #146 #631]: #627
+#609 := (not #627)
+#578 := (or #609 #301)
+#311 := (>= 1::int 0::int)
+#585 := (not #311)
+#586 := (= 1::int #589)
+#590 := (or #586 #585)
+#419 := (or #609 #590)
+#421 := (iff #419 #578)
+#564 := (iff #578 #578)
+#565 := [rewrite]: #564
+#577 := (iff #590 #301)
+#574 := (or #301 false)
+#571 := (iff #574 #301)
+#576 := [rewrite]: #571
+#575 := (iff #590 #574)
+#584 := (iff #585 false)
+#1 := true
+#582 := (not true)
+#583 := (iff #582 false)
+#580 := [rewrite]: #583
+#296 := (iff #585 #582)
+#303 := (iff #311 true)
+#581 := [rewrite]: #303
+#579 := [monotonicity #581]: #296
+#573 := [trans #579 #580]: #584
+#300 := (iff #586 #301)
+#302 := [rewrite]: #300
+#570 := [monotonicity #302 #573]: #575
+#572 := [trans #570 #576]: #577
+#563 := [monotonicity #572]: #421
+#566 := [trans #563 #565]: #421
+#420 := [quant-inst]: #419
+#560 := [mp #420 #566]: #578
+#442 := [unit-resolution #560 #632]: #301
+#443 := (= #293 #589)
+#28 := (= #25 #27)
+#129 := [asserted]: #28
+#436 := [monotonicity #129]: #443
+#451 := [trans #436 #442]: #555
+#453 := (not #555)
+#454 := (or #453 #558)
+#447 := [th-lemma]: #454
+#455 := [unit-resolution #447 #451]: #558
+#456 := (not #558)
+#458 := (or #456 #457)
+#459 := [th-lemma]: #458
+#552 := [unit-resolution #459 #455]: #457
+#553 := (or #549 #588)
+#540 := [th-lemma]: #553
+#542 := [unit-resolution #540 #552]: #549
+#603 := (or #237 #606)
+#18 := (= #13 0::int)
+#118 := (or #18 #70)
+#633 := (forall (vars (?x3 int)) (:pat #626) #118)
+#123 := (forall (vars (?x3 int)) #118)
+#636 := (iff #123 #633)
+#634 := (iff #118 #118)
+#635 := [refl]: #634
+#637 := [quant-intro #635]: #636
+#133 := (~ #123 #123)
+#147 := (~ #118 #118)
+#148 := [refl]: #147
+#134 := [nnf-pos #148]: #133
+#17 := (< #10 0::int)
+#19 := (implies #17 #18)
+#20 := (forall (vars (?x3 int)) #19)
+#126 := (iff #20 #123)
+#89 := (= 0::int #13)
+#95 := (not #17)
+#96 := (or #95 #89)
+#101 := (forall (vars (?x3 int)) #96)
+#124 := (iff #101 #123)
+#121 := (iff #96 #118)
+#115 := (or #70 #18)
+#119 := (iff #115 #118)
+#120 := [rewrite]: #119
+#116 := (iff #96 #115)
+#113 := (iff #89 #18)
+#114 := [rewrite]: #113
+#111 := (iff #95 #70)
+#106 := (not #71)
+#109 := (iff #106 #70)
+#110 := [rewrite]: #109
+#107 := (iff #95 #106)
+#104 := (iff #17 #71)
+#105 := [rewrite]: #104
+#108 := [monotonicity #105]: #107
+#112 := [trans #108 #110]: #111
+#117 := [monotonicity #112 #114]: #116
+#122 := [trans #117 #120]: #121
+#125 := [quant-intro #122]: #124
+#102 := (iff #20 #101)
+#99 := (iff #19 #96)
+#92 := (implies #17 #89)
+#97 := (iff #92 #96)
+#98 := [rewrite]: #97
+#93 := (iff #19 #92)
+#90 := (iff #18 #89)
+#91 := [rewrite]: #90
+#94 := [monotonicity #91]: #93
+#100 := [trans #94 #98]: #99
+#103 := [quant-intro #100]: #102
+#127 := [trans #103 #125]: #126
+#88 := [asserted]: #20
+#128 := [mp #88 #127]: #123
+#149 := [mp~ #128 #134]: #123
+#638 := [mp #149 #637]: #633
+#604 := (not #633)
+#602 := (or #604 #237 #606)
+#204 := (>= #24 0::int)
+#601 := (or #237 #204)
+#605 := (or #604 #601)
+#317 := (iff #605 #602)
+#592 := (or #604 #603)
+#315 := (iff #592 #602)
+#316 := [rewrite]: #315
+#299 := (iff #605 #592)
+#242 := (iff #601 #603)
+#279 := (iff #204 #606)
+#280 := [rewrite]: #279
+#243 := [monotonicity #280]: #242
+#314 := [monotonicity #243]: #299
+#210 := [trans #314 #316]: #317
+#591 := [quant-inst]: #605
+#587 := [mp #591 #210]: #602
+#534 := [unit-resolution #587 #638]: #603
+#531 := [unit-resolution #534 #542]: #606
+#613 := (not #606)
+#607 := (or #613 #612)
+#251 := (or #609 #613 #612)
+#289 := (not #204)
+#294 := (= #24 #293)
+#291 := (or #294 #289)
+#593 := (or #609 #291)
+#597 := (iff #593 #251)
+#256 := (or #609 #607)
+#595 := (iff #256 #251)
+#596 := [rewrite]: #595
+#257 := (iff #593 #256)
+#608 := (iff #291 #607)
+#616 := (or #612 #613)
+#266 := (iff #616 #607)
+#271 := [rewrite]: #266
+#611 := (iff #291 #616)
+#614 := (iff #289 #613)
+#615 := [monotonicity #280]: #614
+#268 := (iff #294 #612)
+#399 := [rewrite]: #268
+#617 := [monotonicity #399 #615]: #611
+#267 := [trans #617 #271]: #608
+#258 := [monotonicity #267]: #257
+#598 := [trans #258 #596]: #597
+#255 := [quant-inst]: #593
+#599 := [mp #255 #598]: #251
+#533 := [unit-resolution #599 #632]: #607
+#543 := [unit-resolution #533 #531]: #612
+#544 := (not #612)
+#545 := (or #544 #594)
+#541 := [th-lemma]: #545
+#546 := [unit-resolution #541 #543]: #594
+#600 := (>= #610 0::int)
+#535 := (or #544 #600)
+#536 := [th-lemma]: #535
+#537 := [unit-resolution #536 #543]: #600
+#557 := (<= #293 1::int)
+#538 := (or #453 #557)
+#532 := [th-lemma]: #538
+#539 := [unit-resolution #532 #451]: #557
+[th-lemma #455 #539 #537 #546]: false
+unsat
+
+kDuOn7kAggfP4Um8ghhv5A 5551
+#2 := false
+#23 := 3::int
+decl uf_2 :: (-> T1 int)
+decl uf_3 :: T1
+#21 := uf_3
+#22 := (uf_2 uf_3)
+#137 := (>= #22 3::int)
+#135 := (not #137)
+#24 := (< #22 3::int)
+#136 := (iff #24 #135)
+#138 := [rewrite]: #136
+#132 := [asserted]: #24
+#139 := [mp #132 #138]: #135
+#9 := 0::int
+decl uf_1 :: (-> int T1)
+#25 := 2::int
+#26 := (* 2::int #22)
+#27 := (uf_1 #26)
+#28 := (uf_2 #27)
+#632 := -1::int
+#634 := (* -1::int #28)
+#290 := (+ #26 #634)
+#623 := (>= #290 0::int)
+#421 := (= #290 0::int)
+#302 := (>= #22 0::int)
+#625 := (= #28 0::int)
+#318 := (not #625)
+#322 := (<= #28 0::int)
+#324 := (not #322)
+#29 := 7::int
+#143 := (>= #28 7::int)
+#30 := (< #28 7::int)
+#31 := (not #30)
+#150 := (iff #31 #143)
+#141 := (not #143)
+#145 := (not #141)
+#148 := (iff #145 #143)
+#149 := [rewrite]: #148
+#146 := (iff #31 #145)
+#142 := (iff #30 #141)
+#144 := [rewrite]: #142
+#147 := [monotonicity #144]: #146
+#151 := [trans #147 #149]: #150
+#133 := [asserted]: #31
+#152 := [mp #133 #151]: #143
+#325 := (or #324 #141)
+#603 := [th-lemma]: #325
+#604 := [unit-resolution #603 #152]: #324
+#601 := (or #318 #322)
+#605 := [th-lemma]: #601
+#602 := [unit-resolution #605 #604]: #318
+#10 := (:var 0 int)
+#12 := (uf_1 #10)
+#648 := (pattern #12)
+#73 := (>= #10 0::int)
+#13 := (uf_2 #12)
+#18 := (= #13 0::int)
+#121 := (or #18 #73)
+#655 := (forall (vars (?x3 int)) (:pat #648) #121)
+#126 := (forall (vars (?x3 int)) #121)
+#658 := (iff #126 #655)
+#656 := (iff #121 #121)
+#657 := [refl]: #656
+#659 := [quant-intro #657]: #658
+#154 := (~ #126 #126)
+#170 := (~ #121 #121)
+#171 := [refl]: #170
+#155 := [nnf-pos #171]: #154
+#17 := (< #10 0::int)
+#19 := (implies #17 #18)
+#20 := (forall (vars (?x3 int)) #19)
+#129 := (iff #20 #126)
+#92 := (= 0::int #13)
+#98 := (not #17)
+#99 := (or #98 #92)
+#104 := (forall (vars (?x3 int)) #99)
+#127 := (iff #104 #126)
+#124 := (iff #99 #121)
+#118 := (or #73 #18)
+#122 := (iff #118 #121)
+#123 := [rewrite]: #122
+#119 := (iff #99 #118)
+#116 := (iff #92 #18)
+#117 := [rewrite]: #116
+#114 := (iff #98 #73)
+#74 := (not #73)
+#109 := (not #74)
+#112 := (iff #109 #73)
+#113 := [rewrite]: #112
+#110 := (iff #98 #109)
+#107 := (iff #17 #74)
+#108 := [rewrite]: #107
+#111 := [monotonicity #108]: #110
+#115 := [trans #111 #113]: #114
+#120 := [monotonicity #115 #117]: #119
+#125 := [trans #120 #123]: #124
+#128 := [quant-intro #125]: #127
+#105 := (iff #20 #104)
+#102 := (iff #19 #99)
+#95 := (implies #17 #92)
+#100 := (iff #95 #99)
+#101 := [rewrite]: #100
+#96 := (iff #19 #95)
+#93 := (iff #18 #92)
+#94 := [rewrite]: #93
+#97 := [monotonicity #94]: #96
+#103 := [trans #97 #101]: #102
+#106 := [quant-intro #103]: #105
+#130 := [trans #106 #128]: #129
+#91 := [asserted]: #20
+#131 := [mp #91 #130]: #126
+#172 := [mp~ #131 #155]: #126
+#660 := [mp #172 #659]: #655
+#337 := (not #655)
+#338 := (or #337 #302 #625)
+#315 := (>= #26 0::int)
+#264 := (or #625 #315)
+#339 := (or #337 #264)
+#611 := (iff #339 #338)
+#627 := (or #302 #625)
+#609 := (or #337 #627)
+#333 := (iff #609 #338)
+#607 := [rewrite]: #333
+#610 := (iff #339 #609)
+#321 := (iff #264 #627)
+#265 := (or #625 #302)
+#613 := (iff #265 #627)
+#614 := [rewrite]: #613
+#626 := (iff #264 #265)
+#635 := (iff #315 #302)
+#636 := [rewrite]: #635
+#624 := [monotonicity #636]: #626
+#336 := [trans #624 #614]: #321
+#332 := [monotonicity #336]: #610
+#608 := [trans #332 #607]: #611
+#231 := [quant-inst]: #339
+#612 := [mp #231 #608]: #338
+#606 := [unit-resolution #612 #660 #602]: #302
+#637 := (not #302)
+#293 := (or #637 #421)
+#55 := (= #10 #13)
+#80 := (or #55 #74)
+#649 := (forall (vars (?x2 int)) (:pat #648) #80)
+#85 := (forall (vars (?x2 int)) #80)
+#652 := (iff #85 #649)
+#650 := (iff #80 #80)
+#651 := [refl]: #650
+#653 := [quant-intro #651]: #652
+#153 := (~ #85 #85)
+#167 := (~ #80 #80)
+#168 := [refl]: #167
+#134 := [nnf-pos #168]: #153
+#14 := (= #13 #10)
+#11 := (<= 0::int #10)
+#15 := (implies #11 #14)
+#16 := (forall (vars (?x2 int)) #15)
+#88 := (iff #16 #85)
+#62 := (not #11)
+#63 := (or #62 #55)
+#68 := (forall (vars (?x2 int)) #63)
+#86 := (iff #68 #85)
+#83 := (iff #63 #80)
+#77 := (or #74 #55)
+#81 := (iff #77 #80)
+#82 := [rewrite]: #81
+#78 := (iff #63 #77)
+#75 := (iff #62 #74)
+#71 := (iff #11 #73)
+#72 := [rewrite]: #71
+#76 := [monotonicity #72]: #75
+#79 := [monotonicity #76]: #78
+#84 := [trans #79 #82]: #83
+#87 := [quant-intro #84]: #86
+#69 := (iff #16 #68)
+#66 := (iff #15 #63)
+#59 := (implies #11 #55)
+#64 := (iff #59 #63)
+#65 := [rewrite]: #64
+#60 := (iff #15 #59)
+#57 := (iff #14 #55)
+#58 := [rewrite]: #57
+#61 := [monotonicity #58]: #60
+#67 := [trans #61 #65]: #66
+#70 := [quant-intro #67]: #69
+#89 := [trans #70 #87]: #88
+#54 := [asserted]: #16
+#90 := [mp #54 #89]: #85
+#169 := [mp~ #90 #134]: #85
+#654 := [mp #169 #653]: #649
+#615 := (not #649)
+#277 := (or #615 #637 #421)
+#243 := (not #315)
+#317 := (= #26 #28)
+#296 := (or #317 #243)
+#278 := (or #615 #296)
+#621 := (iff #278 #277)
+#280 := (or #615 #293)
+#619 := (iff #280 #277)
+#620 := [rewrite]: #619
+#617 := (iff #278 #280)
+#631 := (iff #296 #293)
+#639 := (or #421 #637)
+#630 := (iff #639 #293)
+#289 := [rewrite]: #630
+#629 := (iff #296 #639)
+#638 := (iff #243 #637)
+#633 := [monotonicity #636]: #638
+#628 := (iff #317 #421)
+#301 := [rewrite]: #628
+#288 := [monotonicity #301 #633]: #629
+#273 := [trans #288 #289]: #631
+#618 := [monotonicity #273]: #617
+#616 := [trans #618 #620]: #621
+#279 := [quant-inst]: #278
+#622 := [mp #279 #616]: #277
+#595 := [unit-resolution #622 #654]: #293
+#596 := [unit-resolution #595 #606]: #421
+#597 := (not #421)
+#592 := (or #597 #623)
+#593 := [th-lemma]: #592
+#598 := [unit-resolution #593 #596]: #623
+[th-lemma #152 #598 #139]: false
+unsat
+
+aiB004AWADNjynNrqCDsxw 9284
+#2 := false
+#9 := 0::int
+decl uf_2 :: (-> T1 int)
+decl uf_3 :: T1
+#22 := uf_3
+#23 := (uf_2 uf_3)
+#469 := (= #23 0::int)
+decl uf_1 :: (-> int T1)
+#251 := (uf_1 #23)
+#557 := (uf_2 #251)
+#558 := (= #557 0::int)
+#556 := (>= #23 0::int)
+#477 := (not #556)
+#144 := -1::int
+#348 := (>= #23 -1::int)
+#628 := (not #348)
+#21 := 1::int
+#24 := (+ 1::int #23)
+#25 := (uf_1 #24)
+#26 := (uf_2 #25)
+#635 := (* -1::int #26)
+#632 := (+ #23 #635)
+#636 := (= #632 -1::int)
+#471 := (not #636)
+#606 := (<= #632 -1::int)
+#527 := (not #606)
+#145 := (* -1::int #23)
+#146 := (+ #145 #26)
+#149 := (uf_1 #146)
+#152 := (uf_2 #149)
+#504 := (+ #635 #152)
+#505 := (+ #23 #504)
+#573 := (>= #505 0::int)
+#502 := (= #505 0::int)
+#595 := (<= #632 0::int)
+#526 := [hypothesis]: #606
+#514 := (or #527 #595)
+#515 := [th-lemma]: #514
+#510 := [unit-resolution #515 #526]: #595
+#588 := (not #595)
+#579 := (or #502 #588)
+#10 := (:var 0 int)
+#12 := (uf_1 #10)
+#672 := (pattern #12)
+#76 := (>= #10 0::int)
+#77 := (not #76)
+#13 := (uf_2 #12)
+#58 := (= #10 #13)
+#83 := (or #58 #77)
+#673 := (forall (vars (?x2 int)) (:pat #672) #83)
+#88 := (forall (vars (?x2 int)) #83)
+#676 := (iff #88 #673)
+#674 := (iff #83 #83)
+#675 := [refl]: #674
+#677 := [quant-intro #675]: #676
+#179 := (~ #88 #88)
+#191 := (~ #83 #83)
+#192 := [refl]: #191
+#177 := [nnf-pos #192]: #179
+#14 := (= #13 #10)
+#11 := (<= 0::int #10)
+#15 := (implies #11 #14)
+#16 := (forall (vars (?x2 int)) #15)
+#91 := (iff #16 #88)
+#65 := (not #11)
+#66 := (or #65 #58)
+#71 := (forall (vars (?x2 int)) #66)
+#89 := (iff #71 #88)
+#86 := (iff #66 #83)
+#80 := (or #77 #58)
+#84 := (iff #80 #83)
+#85 := [rewrite]: #84
+#81 := (iff #66 #80)
+#78 := (iff #65 #77)
+#74 := (iff #11 #76)
+#75 := [rewrite]: #74
+#79 := [monotonicity #75]: #78
+#82 := [monotonicity #79]: #81
+#87 := [trans #82 #85]: #86
+#90 := [quant-intro #87]: #89
+#72 := (iff #16 #71)
+#69 := (iff #15 #66)
+#62 := (implies #11 #58)
+#67 := (iff #62 #66)
+#68 := [rewrite]: #67
+#63 := (iff #15 #62)
+#60 := (iff #14 #58)
+#61 := [rewrite]: #60
+#64 := [monotonicity #61]: #63
+#70 := [trans #64 #68]: #69
+#73 := [quant-intro #70]: #72
+#92 := [trans #73 #90]: #91
+#57 := [asserted]: #16
+#93 := [mp #57 #92]: #88
+#193 := [mp~ #93 #177]: #88
+#678 := [mp #193 #677]: #673
+#644 := (not #673)
+#591 := (or #644 #502 #588)
+#499 := (>= #146 0::int)
+#500 := (not #499)
+#493 := (= #146 #152)
+#501 := (or #493 #500)
+#587 := (or #644 #501)
+#585 := (iff #587 #591)
+#581 := (or #644 #579)
+#584 := (iff #581 #591)
+#578 := [rewrite]: #584
+#582 := (iff #587 #581)
+#589 := (iff #501 #579)
+#580 := (iff #500 #588)
+#599 := (iff #499 #595)
+#586 := [rewrite]: #599
+#577 := [monotonicity #586]: #580
+#503 := (iff #493 #502)
+#598 := [rewrite]: #503
+#590 := [monotonicity #598 #577]: #589
+#583 := [monotonicity #590]: #582
+#569 := [trans #583 #578]: #585
+#592 := [quant-inst]: #587
+#570 := [mp #592 #569]: #591
+#516 := [unit-resolution #570 #678]: #579
+#484 := [unit-resolution #516 #510]: #502
+#491 := (not #502)
+#450 := (or #491 #573)
+#481 := [th-lemma]: #450
+#483 := [unit-resolution #481 #484]: #573
+#554 := (<= #152 0::int)
+#163 := (* -1::int #152)
+#138 := (uf_1 0::int)
+#141 := (uf_2 #138)
+#164 := (+ #141 #163)
+#162 := (>= #164 0::int)
+#30 := (- #26 #23)
+#31 := (uf_1 #30)
+#32 := (uf_2 #31)
+#27 := (* 0::int #26)
+#28 := (uf_1 #27)
+#29 := (uf_2 #28)
+#33 := (< #29 #32)
+#34 := (not #33)
+#174 := (iff #34 #162)
+#155 := (< #141 #152)
+#158 := (not #155)
+#172 := (iff #158 #162)
+#161 := (not #162)
+#167 := (not #161)
+#170 := (iff #167 #162)
+#171 := [rewrite]: #170
+#168 := (iff #158 #167)
+#165 := (iff #155 #161)
+#166 := [rewrite]: #165
+#169 := [monotonicity #166]: #168
+#173 := [trans #169 #171]: #172
+#159 := (iff #34 #158)
+#156 := (iff #33 #155)
+#153 := (= #32 #152)
+#150 := (= #31 #149)
+#147 := (= #30 #146)
+#148 := [rewrite]: #147
+#151 := [monotonicity #148]: #150
+#154 := [monotonicity #151]: #153
+#142 := (= #29 #141)
+#139 := (= #28 #138)
+#136 := (= #27 0::int)
+#137 := [rewrite]: #136
+#140 := [monotonicity #137]: #139
+#143 := [monotonicity #140]: #142
+#157 := [monotonicity #143 #154]: #156
+#160 := [monotonicity #157]: #159
+#175 := [trans #160 #173]: #174
+#135 := [asserted]: #34
+#176 := [mp #135 #175]: #162
+#648 := (<= #141 0::int)
+#662 := (= #141 0::int)
+#645 := (or #644 #662)
+#445 := (>= 0::int 0::int)
+#652 := (not #445)
+#659 := (= 0::int #141)
+#660 := (or #659 #652)
+#640 := (or #644 #660)
+#284 := (iff #640 #645)
+#649 := (iff #645 #645)
+#289 := [rewrite]: #649
+#642 := (iff #660 #662)
+#302 := (or #662 false)
+#305 := (iff #302 #662)
+#641 := [rewrite]: #305
+#303 := (iff #660 #302)
+#298 := (iff #652 false)
+#1 := true
+#313 := (not true)
+#314 := (iff #313 false)
+#655 := [rewrite]: #314
+#318 := (iff #652 #313)
+#663 := (iff #445 true)
+#653 := [rewrite]: #663
+#654 := [monotonicity #653]: #318
+#639 := [trans #654 #655]: #298
+#661 := (iff #659 #662)
+#657 := [rewrite]: #661
+#304 := [monotonicity #657 #639]: #303
+#643 := [trans #304 #641]: #642
+#647 := [monotonicity #643]: #284
+#290 := [trans #647 #289]: #284
+#646 := [quant-inst]: #640
+#650 := [mp #646 #290]: #645
+#485 := [unit-resolution #650 #678]: #662
+#492 := (not #662)
+#494 := (or #492 #648)
+#495 := [th-lemma]: #494
+#496 := [unit-resolution #495 #485]: #648
+#506 := (not #648)
+#486 := (or #554 #506 #161)
+#507 := [th-lemma]: #486
+#462 := [unit-resolution #507 #496 #176]: #554
+#463 := [th-lemma #462 #526 #483]: false
+#468 := [lemma #463]: #527
+#472 := (or #471 #606)
+#473 := [th-lemma]: #472
+#474 := [unit-resolution #473 #468]: #471
+#619 := (or #628 #636)
+#622 := (or #644 #628 #636)
+#634 := (>= #24 0::int)
+#356 := (not #634)
+#357 := (= #24 #26)
+#631 := (or #357 #356)
+#623 := (or #644 #631)
+#610 := (iff #623 #622)
+#624 := (or #644 #619)
+#467 := (iff #624 #622)
+#609 := [rewrite]: #467
+#465 := (iff #623 #624)
+#616 := (iff #631 #619)
+#629 := (or #636 #628)
+#620 := (iff #629 #619)
+#621 := [rewrite]: #620
+#626 := (iff #631 #629)
+#343 := (iff #356 #628)
+#349 := (iff #634 #348)
+#627 := [rewrite]: #349
+#625 := [monotonicity #627]: #343
+#346 := (iff #357 #636)
+#347 := [rewrite]: #346
+#630 := [monotonicity #347 #625]: #626
+#617 := [trans #630 #621]: #616
+#466 := [monotonicity #617]: #465
+#611 := [trans #466 #609]: #610
+#618 := [quant-inst]: #623
+#612 := [mp #618 #611]: #622
+#475 := [unit-resolution #612 #678]: #619
+#476 := [unit-resolution #475 #474]: #628
+#478 := (or #477 #348)
+#479 := [th-lemma]: #478
+#464 := [unit-resolution #479 #476]: #477
+#560 := (or #556 #558)
+#18 := (= #13 0::int)
+#124 := (or #18 #76)
+#679 := (forall (vars (?x3 int)) (:pat #672) #124)
+#129 := (forall (vars (?x3 int)) #124)
+#682 := (iff #129 #679)
+#680 := (iff #124 #124)
+#681 := [refl]: #680
+#683 := [quant-intro #681]: #682
+#180 := (~ #129 #129)
+#194 := (~ #124 #124)
+#195 := [refl]: #194
+#181 := [nnf-pos #195]: #180
+#17 := (< #10 0::int)
+#19 := (implies #17 #18)
+#20 := (forall (vars (?x3 int)) #19)
+#132 := (iff #20 #129)
+#95 := (= 0::int #13)
+#101 := (not #17)
+#102 := (or #101 #95)
+#107 := (forall (vars (?x3 int)) #102)
+#130 := (iff #107 #129)
+#127 := (iff #102 #124)
+#121 := (or #76 #18)
+#125 := (iff #121 #124)
+#126 := [rewrite]: #125
+#122 := (iff #102 #121)
+#119 := (iff #95 #18)
+#120 := [rewrite]: #119
+#117 := (iff #101 #76)
+#112 := (not #77)
+#115 := (iff #112 #76)
+#116 := [rewrite]: #115
+#113 := (iff #101 #112)
+#110 := (iff #17 #77)
+#111 := [rewrite]: #110
+#114 := [monotonicity #111]: #113
+#118 := [trans #114 #116]: #117
+#123 := [monotonicity #118 #120]: #122
+#128 := [trans #123 #126]: #127
+#131 := [quant-intro #128]: #130
+#108 := (iff #20 #107)
+#105 := (iff #19 #102)
+#98 := (implies #17 #95)
+#103 := (iff #98 #102)
+#104 := [rewrite]: #103
+#99 := (iff #19 #98)
+#96 := (iff #18 #95)
+#97 := [rewrite]: #96
+#100 := [monotonicity #97]: #99
+#106 := [trans #100 #104]: #105
+#109 := [quant-intro #106]: #108
+#133 := [trans #109 #131]: #132
+#94 := [asserted]: #20
+#134 := [mp #94 #133]: #129
+#196 := [mp~ #134 #181]: #129
+#684 := [mp #196 #683]: #679
+#604 := (not #679)
+#539 := (or #604 #556 #558)
+#559 := (or #558 #556)
+#540 := (or #604 #559)
+#547 := (iff #540 #539)
+#543 := (or #604 #560)
+#546 := (iff #543 #539)
+#541 := [rewrite]: #546
+#544 := (iff #540 #543)
+#550 := (iff #559 #560)
+#561 := [rewrite]: #550
+#545 := [monotonicity #561]: #544
+#533 := [trans #545 #541]: #547
+#542 := [quant-inst]: #540
+#529 := [mp #542 #533]: #539
+#480 := [unit-resolution #529 #684]: #560
+#441 := [unit-resolution #480 #464]: #558
+#449 := (= #23 #557)
+#336 := (= uf_3 #251)
+#4 := (:var 0 T1)
+#5 := (uf_2 #4)
+#664 := (pattern #5)
+#6 := (uf_1 #5)
+#51 := (= #4 #6)
+#665 := (forall (vars (?x1 T1)) (:pat #664) #51)
+#54 := (forall (vars (?x1 T1)) #51)
+#666 := (iff #54 #665)
+#668 := (iff #665 #665)
+#669 := [rewrite]: #668
+#667 := [rewrite]: #666
+#670 := [trans #667 #669]: #666
+#188 := (~ #54 #54)
+#186 := (~ #51 #51)
+#187 := [refl]: #186
+#189 := [nnf-pos #187]: #188
+#7 := (= #6 #4)
+#8 := (forall (vars (?x1 T1)) #7)
+#55 := (iff #8 #54)
+#52 := (iff #7 #51)
+#53 := [rewrite]: #52
+#56 := [quant-intro #53]: #55
+#50 := [asserted]: #8
+#59 := [mp #50 #56]: #54
+#190 := [mp~ #59 #189]: #54
+#671 := [mp #190 #670]: #665
+#337 := (not #665)
+#338 := (or #337 #336)
+#342 := [quant-inst]: #338
+#442 := [unit-resolution #342 #671]: #336
+#451 := [monotonicity #442]: #449
+#452 := [trans #451 #441]: #469
+#453 := (not #469)
+#455 := (or #453 #556)
+#456 := [th-lemma]: #455
+[unit-resolution #456 #464 #452]: false
+unsat
+
+twoPNF2RBLeff4yYfubpfg 7634
+#2 := false
+#9 := 0::int
+decl uf_2 :: (-> T1 int)
+decl uf_1 :: (-> int T1)
+decl uf_3 :: T1
+#22 := uf_3
+#23 := (uf_2 uf_3)
+#21 := 1::int
+#24 := (+ 1::int #23)
+#25 := (uf_1 #24)
+#26 := (uf_2 #25)
+#138 := -1::int
+#139 := (+ -1::int #26)
+#142 := (uf_1 #139)
+#289 := (uf_2 #142)
+#674 := (* -1::int #289)
+#538 := (+ #23 #674)
+#532 := (>= #538 0::int)
+#536 := (= #23 #289)
+#148 := (= uf_3 #142)
+#167 := (<= #26 0::int)
+#168 := (not #167)
+#174 := (iff #148 #168)
+#189 := (not #174)
+#220 := (iff #189 #148)
+#210 := (not #148)
+#215 := (not #210)
+#218 := (iff #215 #148)
+#219 := [rewrite]: #218
+#216 := (iff #189 #215)
+#213 := (iff #174 #210)
+#207 := (iff #148 false)
+#211 := (iff #207 #210)
+#212 := [rewrite]: #211
+#208 := (iff #174 #207)
+#205 := (iff #168 false)
+#1 := true
+#200 := (not true)
+#203 := (iff #200 false)
+#204 := [rewrite]: #203
+#201 := (iff #168 #200)
+#198 := (iff #167 true)
+#179 := (or #168 #174)
+#182 := (not #179)
+#27 := (< 0::int #26)
+#28 := (ite #27 true false)
+#29 := (- #26 1::int)
+#30 := (uf_1 #29)
+#31 := (= #30 uf_3)
+#32 := (iff #28 #31)
+#33 := (or #32 #28)
+#34 := (not #33)
+#185 := (iff #34 #182)
+#153 := (iff #27 #148)
+#159 := (or #27 #153)
+#164 := (not #159)
+#183 := (iff #164 #182)
+#180 := (iff #159 #179)
+#177 := (iff #153 #174)
+#171 := (iff #168 #148)
+#175 := (iff #171 #174)
+#176 := [rewrite]: #175
+#172 := (iff #153 #171)
+#169 := (iff #27 #168)
+#170 := [rewrite]: #169
+#173 := [monotonicity #170]: #172
+#178 := [trans #173 #176]: #177
+#181 := [monotonicity #170 #178]: #180
+#184 := [monotonicity #181]: #183
+#165 := (iff #34 #164)
+#162 := (iff #33 #159)
+#156 := (or #153 #27)
+#160 := (iff #156 #159)
+#161 := [rewrite]: #160
+#157 := (iff #33 #156)
+#136 := (iff #28 #27)
+#137 := [rewrite]: #136
+#154 := (iff #32 #153)
+#151 := (iff #31 #148)
+#145 := (= #142 uf_3)
+#149 := (iff #145 #148)
+#150 := [rewrite]: #149
+#146 := (iff #31 #145)
+#143 := (= #30 #142)
+#140 := (= #29 #139)
+#141 := [rewrite]: #140
+#144 := [monotonicity #141]: #143
+#147 := [monotonicity #144]: #146
+#152 := [trans #147 #150]: #151
+#155 := [monotonicity #137 #152]: #154
+#158 := [monotonicity #155 #137]: #157
+#163 := [trans #158 #161]: #162
+#166 := [monotonicity #163]: #165
+#186 := [trans #166 #184]: #185
+#135 := [asserted]: #34
+#187 := [mp #135 #186]: #182
+#188 := [not-or-elim #187]: #167
+#199 := [iff-true #188]: #198
+#202 := [monotonicity #199]: #201
+#206 := [trans #202 #204]: #205
+#209 := [monotonicity #206]: #208
+#214 := [trans #209 #212]: #213
+#217 := [monotonicity #214]: #216
+#221 := [trans #217 #219]: #220
+#190 := [not-or-elim #187]: #189
+#222 := [mp #190 #221]: #148
+#543 := [monotonicity #222]: #536
+#544 := (not #536)
+#616 := (or #544 #532)
+#618 := [th-lemma]: #616
+#628 := [unit-resolution #618 #543]: #532
+#354 := (* -1::int #26)
+#484 := (+ #23 #354)
+#683 := (<= #484 -1::int)
+#691 := (= #484 -1::int)
+#698 := (>= #23 -1::int)
+#521 := (>= #289 0::int)
+#652 := (= #289 0::int)
+#387 := (>= #26 1::int)
+#667 := (not #387)
+#629 := (or #667 #168)
+#630 := [th-lemma]: #629
+#626 := [unit-resolution #630 #188]: #667
+#10 := (:var 0 int)
+#12 := (uf_1 #10)
+#711 := (pattern #12)
+#76 := (>= #10 0::int)
+#13 := (uf_2 #12)
+#18 := (= #13 0::int)
+#124 := (or #18 #76)
+#718 := (forall (vars (?x3 int)) (:pat #711) #124)
+#129 := (forall (vars (?x3 int)) #124)
+#721 := (iff #129 #718)
+#719 := (iff #124 #124)
+#720 := [refl]: #719
+#722 := [quant-intro #720]: #721
+#229 := (~ #129 #129)
+#227 := (~ #124 #124)
+#228 := [refl]: #227
+#230 := [nnf-pos #228]: #229
+#17 := (< #10 0::int)
+#19 := (implies #17 #18)
+#20 := (forall (vars (?x3 int)) #19)
+#132 := (iff #20 #129)
+#95 := (= 0::int #13)
+#101 := (not #17)
+#102 := (or #101 #95)
+#107 := (forall (vars (?x3 int)) #102)
+#130 := (iff #107 #129)
+#127 := (iff #102 #124)
+#121 := (or #76 #18)
+#125 := (iff #121 #124)
+#126 := [rewrite]: #125
+#122 := (iff #102 #121)
+#119 := (iff #95 #18)
+#120 := [rewrite]: #119
+#117 := (iff #101 #76)
+#77 := (not #76)
+#112 := (not #77)
+#115 := (iff #112 #76)
+#116 := [rewrite]: #115
+#113 := (iff #101 #112)
+#110 := (iff #17 #77)
+#111 := [rewrite]: #110
+#114 := [monotonicity #111]: #113
+#118 := [trans #114 #116]: #117
+#123 := [monotonicity #118 #120]: #122
+#128 := [trans #123 #126]: #127
+#131 := [quant-intro #128]: #130
+#108 := (iff #20 #107)
+#105 := (iff #19 #102)
+#98 := (implies #17 #95)
+#103 := (iff #98 #102)
+#104 := [rewrite]: #103
+#99 := (iff #19 #98)
+#96 := (iff #18 #95)
+#97 := [rewrite]: #96
+#100 := [monotonicity #97]: #99
+#106 := [trans #100 #104]: #105
+#109 := [quant-intro #106]: #108
+#133 := [trans #109 #131]: #132
+#94 := [asserted]: #20
+#134 := [mp #94 #133]: #129
+#231 := [mp~ #134 #230]: #129
+#723 := [mp #231 #722]: #718
+#328 := (not #718)
+#643 := (or #328 #387 #652)
+#673 := (>= #139 0::int)
+#653 := (or #652 #673)
+#641 := (or #328 #653)
+#537 := (iff #641 #643)
+#485 := (or #387 #652)
+#526 := (or #328 #485)
+#487 := (iff #526 #643)
+#635 := [rewrite]: #487
+#527 := (iff #641 #526)
+#640 := (iff #653 #485)
+#647 := (or #652 #387)
+#486 := (iff #647 #485)
+#639 := [rewrite]: #486
+#654 := (iff #653 #647)
+#388 := (iff #673 #387)
+#666 := [rewrite]: #388
+#483 := [monotonicity #666]: #654
+#642 := [trans #483 #639]: #640
+#528 := [monotonicity #642]: #527
+#632 := [trans #528 #635]: #537
+#644 := [quant-inst]: #641
+#633 := [mp #644 #632]: #643
+#631 := [unit-resolution #633 #723 #626]: #652
+#620 := (not #652)
+#621 := (or #620 #521)
+#622 := [th-lemma]: #621
+#623 := [unit-resolution #622 #631]: #521
+#624 := (not #532)
+#617 := (not #521)
+#608 := (or #698 #617 #624)
+#609 := [th-lemma]: #608
+#611 := [unit-resolution #609 #623 #628]: #698
+#701 := (not #698)
+#692 := (or #691 #701)
+#58 := (= #10 #13)
+#83 := (or #58 #77)
+#712 := (forall (vars (?x2 int)) (:pat #711) #83)
+#88 := (forall (vars (?x2 int)) #83)
+#715 := (iff #88 #712)
+#713 := (iff #83 #83)
+#714 := [refl]: #713
+#716 := [quant-intro #714]: #715
+#191 := (~ #88 #88)
+#195 := (~ #83 #83)
+#193 := [refl]: #195
+#225 := [nnf-pos #193]: #191
+#14 := (= #13 #10)
+#11 := (<= 0::int #10)
+#15 := (implies #11 #14)
+#16 := (forall (vars (?x2 int)) #15)
+#91 := (iff #16 #88)
+#65 := (not #11)
+#66 := (or #65 #58)
+#71 := (forall (vars (?x2 int)) #66)
+#89 := (iff #71 #88)
+#86 := (iff #66 #83)
+#80 := (or #77 #58)
+#84 := (iff #80 #83)
+#85 := [rewrite]: #84
+#81 := (iff #66 #80)
+#78 := (iff #65 #77)
+#74 := (iff #11 #76)
+#75 := [rewrite]: #74
+#79 := [monotonicity #75]: #78
+#82 := [monotonicity #79]: #81
+#87 := [trans #82 #85]: #86
+#90 := [quant-intro #87]: #89
+#72 := (iff #16 #71)
+#69 := (iff #15 #66)
+#62 := (implies #11 #58)
+#67 := (iff #62 #66)
+#68 := [rewrite]: #67
+#63 := (iff #15 #62)
+#60 := (iff #14 #58)
+#61 := [rewrite]: #60
+#64 := [monotonicity #61]: #63
+#70 := [trans #64 #68]: #69
+#73 := [quant-intro #70]: #72
+#92 := [trans #73 #90]: #91
+#57 := [asserted]: #16
+#93 := [mp #57 #92]: #88
+#226 := [mp~ #93 #225]: #88
+#717 := [mp #226 #716]: #712
+#693 := (not #712)
+#353 := (or #693 #691 #701)
+#380 := (>= #24 0::int)
+#377 := (not #380)
+#695 := (= #24 #26)
+#697 := (or #695 #377)
+#694 := (or #693 #697)
+#680 := (iff #694 #353)
+#678 := (or #693 #692)
+#343 := (iff #678 #353)
+#344 := [rewrite]: #343
+#341 := (iff #694 #678)
+#352 := (iff #697 #692)
+#696 := (iff #377 #701)
+#699 := (iff #380 #698)
+#700 := [rewrite]: #699
+#702 := [monotonicity #700]: #696
+#365 := (iff #695 #691)
+#366 := [rewrite]: #365
+#357 := [monotonicity #366 #702]: #352
+#342 := [monotonicity #357]: #341
+#681 := [trans #342 #344]: #680
+#337 := [quant-inst]: #694
+#682 := [mp #337 #681]: #353
+#612 := [unit-resolution #682 #717]: #692
+#613 := [unit-resolution #612 #611]: #691
+#614 := (not #691)
+#610 := (or #614 #683)
+#615 := [th-lemma]: #610
+#601 := [unit-resolution #615 #613]: #683
+[th-lemma #623 #188 #601 #628]: false
+unsat
+
+ZcLxnpFewGGQ0H47MfRXGw 11816
+#2 := false
+#9 := 0::int
+decl uf_2 :: (-> T1 int)
+decl uf_1 :: (-> int T1)
+decl uf_3 :: T1
+#21 := uf_3
+#22 := (uf_2 uf_3)
+#23 := 1::int
+#138 := (+ 1::int #22)
+#141 := (uf_1 #138)
+#297 := (uf_2 #141)
+#357 := (= #297 0::int)
+#166 := (uf_1 0::int)
+#531 := (uf_2 #166)
+#537 := (= #531 0::int)
+#10 := (:var 0 int)
+#12 := (uf_1 #10)
+#672 := (pattern #12)
+#78 := (>= #10 0::int)
+#79 := (not #78)
+#13 := (uf_2 #12)
+#60 := (= #10 #13)
+#85 := (or #60 #79)
+#673 := (forall (vars (?x2 int)) (:pat #672) #85)
+#90 := (forall (vars (?x2 int)) #85)
+#676 := (iff #90 #673)
+#674 := (iff #85 #85)
+#675 := [refl]: #674
+#677 := [quant-intro #675]: #676
+#178 := (~ #90 #90)
+#190 := (~ #85 #85)
+#191 := [refl]: #190
+#175 := [nnf-pos #191]: #178
+#14 := (= #13 #10)
+#11 := (<= 0::int #10)
+#15 := (implies #11 #14)
+#16 := (forall (vars (?x2 int)) #15)
+#93 := (iff #16 #90)
+#67 := (not #11)
+#68 := (or #67 #60)
+#73 := (forall (vars (?x2 int)) #68)
+#91 := (iff #73 #90)
+#88 := (iff #68 #85)
+#82 := (or #79 #60)
+#86 := (iff #82 #85)
+#87 := [rewrite]: #86
+#83 := (iff #68 #82)
+#80 := (iff #67 #79)
+#76 := (iff #11 #78)
+#77 := [rewrite]: #76
+#81 := [monotonicity #77]: #80
+#84 := [monotonicity #81]: #83
+#89 := [trans #84 #87]: #88
+#92 := [quant-intro #89]: #91
+#74 := (iff #16 #73)
+#71 := (iff #15 #68)
+#64 := (implies #11 #60)
+#69 := (iff #64 #68)
+#70 := [rewrite]: #69
+#65 := (iff #15 #64)
+#62 := (iff #14 #60)
+#63 := [rewrite]: #62
+#66 := [monotonicity #63]: #65
+#72 := [trans #66 #70]: #71
+#75 := [quant-intro #72]: #74
+#94 := [trans #75 #92]: #93
+#59 := [asserted]: #16
+#95 := [mp #59 #94]: #90
+#192 := [mp~ #95 #175]: #90
+#678 := [mp #192 #677]: #673
+#650 := (not #673)
+#528 := (or #650 #537)
+#529 := (>= 0::int 0::int)
+#530 := (not #529)
+#534 := (= 0::int #531)
+#535 := (or #534 #530)
+#508 := (or #650 #535)
+#509 := (iff #508 #528)
+#514 := (iff #528 #528)
+#515 := [rewrite]: #514
+#527 := (iff #535 #537)
+#520 := (or #537 false)
+#525 := (iff #520 #537)
+#526 := [rewrite]: #525
+#521 := (iff #535 #520)
+#519 := (iff #530 false)
+#1 := true
+#512 := (not true)
+#517 := (iff #512 false)
+#518 := [rewrite]: #517
+#513 := (iff #530 #512)
+#538 := (iff #529 true)
+#511 := [rewrite]: #538
+#406 := [monotonicity #511]: #513
+#524 := [trans #406 #518]: #519
+#536 := (iff #534 #537)
+#532 := [rewrite]: #536
+#522 := [monotonicity #532 #524]: #521
+#523 := [trans #522 #526]: #527
+#490 := [monotonicity #523]: #509
+#510 := [trans #490 #515]: #509
+#454 := [quant-inst]: #508
+#516 := [mp #454 #510]: #528
+#394 := [unit-resolution #516 #678]: #537
+#355 := (= #297 #531)
+#250 := (= #141 #166)
+#26 := 2::int
+#144 := (* 2::int #22)
+#147 := (uf_1 #144)
+#150 := (uf_2 #147)
+#30 := 3::int
+#156 := (+ 3::int #150)
+#161 := (uf_1 #156)
+#336 := (= #161 #166)
+#327 := (not #336)
+#588 := (uf_2 #161)
+#555 := (= #588 0::int)
+#398 := (= #588 #531)
+#395 := [hypothesis]: #336
+#387 := [monotonicity #395]: #398
+#399 := [trans #387 #394]: #555
+#390 := (not #555)
+#547 := (<= #588 0::int)
+#403 := (not #547)
+#595 := (>= #150 0::int)
+#302 := -1::int
+#618 := (* -1::int #150)
+#624 := (+ #144 #618)
+#488 := (<= #624 0::int)
+#465 := (= #624 0::int)
+#609 := (>= #22 0::int)
+#442 := (= #22 0::int)
+#660 := (uf_1 #22)
+#495 := (uf_2 #660)
+#496 := (= #495 0::int)
+#612 := (not #609)
+#451 := [hypothesis]: #612
+#506 := (or #496 #609)
+#18 := (= #13 0::int)
+#126 := (or #18 #78)
+#679 := (forall (vars (?x3 int)) (:pat #672) #126)
+#131 := (forall (vars (?x3 int)) #126)
+#682 := (iff #131 #679)
+#680 := (iff #126 #126)
+#681 := [refl]: #680
+#683 := [quant-intro #681]: #682
+#179 := (~ #131 #131)
+#193 := (~ #126 #126)
+#194 := [refl]: #193
+#180 := [nnf-pos #194]: #179
+#17 := (< #10 0::int)
+#19 := (implies #17 #18)
+#20 := (forall (vars (?x3 int)) #19)
+#134 := (iff #20 #131)
+#97 := (= 0::int #13)
+#103 := (not #17)
+#104 := (or #103 #97)
+#109 := (forall (vars (?x3 int)) #104)
+#132 := (iff #109 #131)
+#129 := (iff #104 #126)
+#123 := (or #78 #18)
+#127 := (iff #123 #126)
+#128 := [rewrite]: #127
+#124 := (iff #104 #123)
+#121 := (iff #97 #18)
+#122 := [rewrite]: #121
+#119 := (iff #103 #78)
+#114 := (not #79)
+#117 := (iff #114 #78)
+#118 := [rewrite]: #117
+#115 := (iff #103 #114)
+#112 := (iff #17 #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 #20 #109)
+#107 := (iff #19 #104)
+#100 := (implies #17 #97)
+#105 := (iff #100 #104)
+#106 := [rewrite]: #105
+#101 := (iff #19 #100)
+#98 := (iff #18 #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]: #20
+#136 := [mp #96 #135]: #131
+#195 := [mp~ #136 #180]: #131
+#684 := [mp #195 #683]: #679
+#346 := (not #679)
+#462 := (or #346 #496 #609)
+#463 := (or #346 #506)
+#469 := (iff #463 #462)
+#470 := [rewrite]: #469
+#468 := [quant-inst]: #463
+#471 := [mp #468 #470]: #462
+#452 := [unit-resolution #471 #684]: #506
+#453 := [unit-resolution #452 #451]: #496
+#456 := (= #22 #495)
+#661 := (= uf_3 #660)
+#4 := (:var 0 T1)
+#5 := (uf_2 #4)
+#664 := (pattern #5)
+#6 := (uf_1 #5)
+#53 := (= #4 #6)
+#665 := (forall (vars (?x1 T1)) (:pat #664) #53)
+#56 := (forall (vars (?x1 T1)) #53)
+#666 := (iff #56 #665)
+#668 := (iff #665 #665)
+#669 := [rewrite]: #668
+#667 := [rewrite]: #666
+#670 := [trans #667 #669]: #666
+#187 := (~ #56 #56)
+#185 := (~ #53 #53)
+#186 := [refl]: #185
+#188 := [nnf-pos #186]: #187
+#7 := (= #6 #4)
+#8 := (forall (vars (?x1 T1)) #7)
+#57 := (iff #8 #56)
+#54 := (iff #7 #53)
+#55 := [rewrite]: #54
+#58 := [quant-intro #55]: #57
+#52 := [asserted]: #8
+#61 := [mp #52 #58]: #56
+#189 := [mp~ #61 #188]: #56
+#671 := [mp #189 #670]: #665
+#663 := (not #665)
+#653 := (or #663 #661)
+#312 := [quant-inst]: #653
+#455 := [unit-resolution #312 #671]: #661
+#457 := [monotonicity #455]: #456
+#458 := [trans #457 #453]: #442
+#459 := (not #442)
+#460 := (or #459 #609)
+#443 := [th-lemma]: #460
+#461 := [unit-resolution #443 #451 #458]: false
+#431 := [lemma #461]: #609
+#613 := (or #465 #612)
+#615 := (or #650 #465 #612)
+#616 := (>= #144 0::int)
+#617 := (not #616)
+#622 := (= #144 #150)
+#623 := (or #622 #617)
+#444 := (or #650 #623)
+#602 := (iff #444 #615)
+#447 := (or #650 #613)
+#603 := (iff #447 #615)
+#604 := [rewrite]: #603
+#600 := (iff #444 #447)
+#614 := (iff #623 #613)
+#606 := (iff #617 #612)
+#610 := (iff #616 #609)
+#611 := [rewrite]: #610
+#607 := [monotonicity #611]: #606
+#466 := (iff #622 #465)
+#467 := [rewrite]: #466
+#608 := [monotonicity #467 #607]: #614
+#601 := [monotonicity #608]: #600
+#605 := [trans #601 #604]: #602
+#446 := [quant-inst]: #444
+#487 := [mp #446 #605]: #615
+#439 := [unit-resolution #487 #678]: #613
+#435 := [unit-resolution #439 #431]: #465
+#440 := (not #465)
+#419 := (or #440 #488)
+#422 := [th-lemma]: #419
+#426 := [unit-resolution #422 #435]: #488
+#430 := (not #488)
+#433 := (or #595 #612 #430)
+#438 := [th-lemma]: #433
+#402 := [unit-resolution #438 #431 #426]: #595
+#590 := -3::int
+#579 := (* -1::int #588)
+#589 := (+ #150 #579)
+#553 := (<= #589 -3::int)
+#591 := (= #589 -3::int)
+#581 := (>= #150 -3::int)
+#644 := (>= #22 -1::int)
+#428 := (or #612 #644)
+#429 := [th-lemma]: #428
+#427 := [unit-resolution #429 #431]: #644
+#646 := (not #644)
+#418 := (or #581 #646 #430)
+#421 := [th-lemma]: #418
+#423 := [unit-resolution #421 #426 #427]: #581
+#584 := (not #581)
+#573 := (or #584 #591)
+#562 := (or #650 #584 #591)
+#599 := (>= #156 0::int)
+#586 := (not #599)
+#580 := (= #156 #588)
+#577 := (or #580 #586)
+#563 := (or #650 #577)
+#549 := (iff #563 #562)
+#566 := (or #650 #573)
+#568 := (iff #566 #562)
+#548 := [rewrite]: #568
+#567 := (iff #563 #566)
+#571 := (iff #577 #573)
+#569 := (or #591 #584)
+#574 := (iff #569 #573)
+#575 := [rewrite]: #574
+#570 := (iff #577 #569)
+#578 := (iff #586 #584)
+#582 := (iff #599 #581)
+#583 := [rewrite]: #582
+#585 := [monotonicity #583]: #578
+#587 := (iff #580 #591)
+#592 := [rewrite]: #587
+#572 := [monotonicity #592 #585]: #570
+#576 := [trans #572 #575]: #571
+#564 := [monotonicity #576]: #567
+#551 := [trans #564 #548]: #549
+#565 := [quant-inst]: #563
+#552 := [mp #565 #551]: #562
+#424 := [unit-resolution #552 #678]: #573
+#420 := [unit-resolution #424 #423]: #591
+#425 := (not #591)
+#415 := (or #425 #553)
+#405 := [th-lemma]: #415
+#407 := [unit-resolution #405 #420]: #553
+#404 := (not #553)
+#401 := (not #595)
+#386 := (or #403 #401 #404)
+#388 := [th-lemma]: #386
+#389 := [unit-resolution #388 #407 #402]: #403
+#391 := (or #390 #547)
+#392 := [th-lemma]: #391
+#393 := [unit-resolution #392 #389]: #390
+#376 := [unit-resolution #393 #399]: false
+#378 := [lemma #376]: #327
+#249 := (= #141 #161)
+#334 := (not #249)
+#396 := (= #297 #588)
+#385 := [hypothesis]: #249
+#370 := [monotonicity #385]: #396
+#380 := (not #396)
+#434 := (+ #297 #579)
+#280 := (>= #434 0::int)
+#414 := (not #280)
+#303 := (* -1::int #297)
+#304 := (+ #22 #303)
+#356 := (>= #304 -1::int)
+#641 := (= #304 -1::int)
+#649 := (or #641 #646)
+#648 := (or #650 #641 #646)
+#317 := (>= #138 0::int)
+#654 := (not #317)
+#639 := (= #138 #297)
+#301 := (or #639 #654)
+#651 := (or #650 #301)
+#363 := (iff #651 #648)
+#638 := (or #650 #649)
+#361 := (iff #638 #648)
+#362 := [rewrite]: #361
+#345 := (iff #651 #638)
+#288 := (iff #301 #649)
+#283 := (iff #654 #646)
+#645 := (iff #317 #644)
+#640 := [rewrite]: #645
+#647 := [monotonicity #640]: #283
+#642 := (iff #639 #641)
+#643 := [rewrite]: #642
+#289 := [monotonicity #643 #647]: #288
+#360 := [monotonicity #289]: #345
+#256 := [trans #360 #362]: #363
+#637 := [quant-inst]: #651
+#633 := [mp #637 #256]: #648
+#408 := [unit-resolution #633 #678]: #649
+#411 := [unit-resolution #408 #427]: #641
+#412 := (not #641)
+#416 := (or #412 #356)
+#409 := [th-lemma]: #416
+#417 := [unit-resolution #409 #411]: #356
+#410 := [hypothesis]: #280
+#413 := [th-lemma #423 #410 #417 #407 #426]: false
+#400 := [lemma #413]: #414
+#381 := (or #380 #280)
+#382 := [th-lemma]: #381
+#377 := [unit-resolution #382 #400]: #380
+#371 := [unit-resolution #377 #370]: false
+#372 := [lemma #371]: #334
+#352 := (or #249 #250 #336)
+#335 := (not #250)
+#338 := (and #334 #335 #327)
+#339 := (not #338)
+#169 := (distinct #141 #161 #166)
+#172 := (not #169)
+#33 := (- #22 #22)
+#34 := (uf_1 #33)
+#27 := (* #22 2::int)
+#28 := (uf_1 #27)
+#29 := (uf_2 #28)
+#31 := (+ #29 3::int)
+#32 := (uf_1 #31)
+#24 := (+ #22 1::int)
+#25 := (uf_1 #24)
+#35 := (distinct #25 #32 #34)
+#36 := (not #35)
+#173 := (iff #36 #172)
+#170 := (iff #35 #169)
+#167 := (= #34 #166)
+#164 := (= #33 0::int)
+#165 := [rewrite]: #164
+#168 := [monotonicity #165]: #167
+#162 := (= #32 #161)
+#159 := (= #31 #156)
+#153 := (+ #150 3::int)
+#157 := (= #153 #156)
+#158 := [rewrite]: #157
+#154 := (= #31 #153)
+#151 := (= #29 #150)
+#148 := (= #28 #147)
+#145 := (= #27 #144)
+#146 := [rewrite]: #145
+#149 := [monotonicity #146]: #148
+#152 := [monotonicity #149]: #151
+#155 := [monotonicity #152]: #154
+#160 := [trans #155 #158]: #159
+#163 := [monotonicity #160]: #162
+#142 := (= #25 #141)
+#139 := (= #24 #138)
+#140 := [rewrite]: #139
+#143 := [monotonicity #140]: #142
+#171 := [monotonicity #143 #163 #168]: #170
+#174 := [monotonicity #171]: #173
+#137 := [asserted]: #36
+#177 := [mp #137 #174]: #172
+#326 := (or #169 #339)
+#659 := [def-axiom]: #326
+#351 := [unit-resolution #659 #177]: #339
+#314 := (or #338 #249 #250 #336)
+#445 := [def-axiom]: #314
+#343 := [unit-resolution #445 #351]: #352
+#353 := [unit-resolution #343 #372 #378]: #250
+#321 := [monotonicity #353]: #355
+#323 := [trans #321 #394]: #357
+#368 := (not #357)
+#620 := (<= #297 0::int)
+#364 := (not #620)
+#634 := (<= #304 -1::int)
+#374 := (or #412 #634)
+#373 := [th-lemma]: #374
+#375 := [unit-resolution #373 #411]: #634
+#365 := (not #634)
+#366 := (or #364 #612 #365)
+#358 := [th-lemma]: #366
+#367 := [unit-resolution #358 #375 #431]: #364
+#359 := (or #368 #620)
+#369 := [th-lemma]: #359
+#350 := [unit-resolution #369 #367]: #368
+[unit-resolution #350 #323]: false
+unsat
+
+ipe8HUL/t33OoeNl/z0smQ 4011
+#2 := false
+#9 := 0::int
+decl uf_3 :: int
+#21 := uf_3
+#130 := -1::int
+#131 := (* -1::int uf_3)
+#154 := (>= uf_3 0::int)
+#161 := (ite #154 uf_3 #131)
+#648 := (* -1::int #161)
+#645 := (+ #131 #648)
+#642 := (<= #645 0::int)
+#340 := (= #131 #161)
+#155 := (not #154)
+#649 := (+ uf_3 #648)
+#650 := (<= #649 0::int)
+#254 := (= uf_3 #161)
+#651 := [hypothesis]: #154
+#255 := (or #155 #254)
+#341 := [def-axiom]: #255
+#289 := [unit-resolution #341 #651]: #254
+#652 := (not #254)
+#654 := (or #652 #650)
+#294 := [th-lemma]: #654
+#295 := [unit-resolution #294 #289]: #650
+#273 := (>= #161 0::int)
+#346 := (not #273)
+decl uf_2 :: (-> T1 int)
+decl uf_1 :: (-> int T1)
+#166 := (uf_1 #161)
+#169 := (uf_2 #166)
+#172 := (= #161 #169)
+#175 := (not #172)
+#23 := (- uf_3)
+#22 := (< uf_3 0::int)
+#24 := (ite #22 #23 uf_3)
+#25 := (uf_1 #24)
+#26 := (uf_2 #25)
+#27 := (= #26 #24)
+#28 := (not #27)
+#178 := (iff #28 #175)
+#134 := (ite #22 #131 uf_3)
+#137 := (uf_1 #134)
+#140 := (uf_2 #137)
+#146 := (= #134 #140)
+#151 := (not #146)
+#176 := (iff #151 #175)
+#173 := (iff #146 #172)
+#170 := (= #140 #169)
+#167 := (= #137 #166)
+#164 := (= #134 #161)
+#158 := (ite #155 #131 uf_3)
+#162 := (= #158 #161)
+#163 := [rewrite]: #162
+#159 := (= #134 #158)
+#156 := (iff #22 #155)
+#157 := [rewrite]: #156
+#160 := [monotonicity #157]: #159
+#165 := [trans #160 #163]: #164
+#168 := [monotonicity #165]: #167
+#171 := [monotonicity #168]: #170
+#174 := [monotonicity #165 #171]: #173
+#177 := [monotonicity #174]: #176
+#152 := (iff #28 #151)
+#149 := (iff #27 #146)
+#143 := (= #140 #134)
+#147 := (iff #143 #146)
+#148 := [rewrite]: #147
+#144 := (iff #27 #143)
+#135 := (= #24 #134)
+#132 := (= #23 #131)
+#133 := [rewrite]: #132
+#136 := [monotonicity #133]: #135
+#141 := (= #26 #140)
+#138 := (= #25 #137)
+#139 := [monotonicity #136]: #138
+#142 := [monotonicity #139]: #141
+#145 := [monotonicity #142 #136]: #144
+#150 := [trans #145 #148]: #149
+#153 := [monotonicity #150]: #152
+#179 := [trans #153 #177]: #178
+#129 := [asserted]: #28
+#180 := [mp #129 #179]: #175
+#10 := (:var 0 int)
+#12 := (uf_1 #10)
+#677 := (pattern #12)
+#70 := (>= #10 0::int)
+#71 := (not #70)
+#13 := (uf_2 #12)
+#52 := (= #10 #13)
+#77 := (or #52 #71)
+#678 := (forall (vars (?x2 int)) (:pat #677) #77)
+#82 := (forall (vars (?x2 int)) #77)
+#681 := (iff #82 #678)
+#679 := (iff #77 #77)
+#680 := [refl]: #679
+#682 := [quant-intro #680]: #681
+#183 := (~ #82 #82)
+#195 := (~ #77 #77)
+#196 := [refl]: #195
+#181 := [nnf-pos #196]: #183
+#14 := (= #13 #10)
+#11 := (<= 0::int #10)
+#15 := (implies #11 #14)
+#16 := (forall (vars (?x2 int)) #15)
+#85 := (iff #16 #82)
+#59 := (not #11)
+#60 := (or #59 #52)
+#65 := (forall (vars (?x2 int)) #60)
+#83 := (iff #65 #82)
+#80 := (iff #60 #77)
+#74 := (or #71 #52)
+#78 := (iff #74 #77)
+#79 := [rewrite]: #78
+#75 := (iff #60 #74)
+#72 := (iff #59 #71)
+#68 := (iff #11 #70)
+#69 := [rewrite]: #68
+#73 := [monotonicity #69]: #72
+#76 := [monotonicity #73]: #75
+#81 := [trans #76 #79]: #80
+#84 := [quant-intro #81]: #83
+#66 := (iff #16 #65)
+#63 := (iff #15 #60)
+#56 := (implies #11 #52)
+#61 := (iff #56 #60)
+#62 := [rewrite]: #61
+#57 := (iff #15 #56)
+#54 := (iff #14 #52)
+#55 := [rewrite]: #54
+#58 := [monotonicity #55]: #57
+#64 := [trans #58 #62]: #63
+#67 := [quant-intro #64]: #66
+#86 := [trans #67 #84]: #85
+#51 := [asserted]: #16
+#87 := [mp #51 #86]: #82
+#197 := [mp~ #87 #181]: #82
+#683 := [mp #197 #682]: #678
+#450 := (not #678)
+#657 := (or #450 #172 #346)
+#661 := (or #172 #346)
+#331 := (or #450 #661)
+#664 := (iff #331 #657)
+#665 := [rewrite]: #664
+#332 := [quant-inst]: #331
+#666 := [mp #332 #665]: #657
+#655 := [unit-resolution #666 #683 #180]: #346
+#653 := [th-lemma #651 #655 #295]: false
+#656 := [lemma #653]: #155
+#342 := (or #154 #340)
+#333 := [def-axiom]: #342
+#365 := [unit-resolution #333 #656]: #340
+#366 := (not #340)
+#367 := (or #366 #642)
+#368 := [th-lemma]: #367
+#261 := [unit-resolution #368 #365]: #642
+#647 := (<= #161 0::int)
+#638 := (or #647 #273)
+#639 := [th-lemma]: #638
+#361 := [unit-resolution #639 #655]: #647
+[th-lemma #656 #361 #261]: false
+unsat
+
+eRjXXrQSzpzyc8Ro409d3Q 14366
+#2 := false
+#9 := 0::int
+decl uf_2 :: (-> T1 int)
+decl uf_1 :: (-> int T1)
+decl uf_5 :: T1
+#36 := uf_5
+#37 := (uf_2 uf_5)
+#35 := 4::int
+#38 := (* 4::int #37)
+#39 := (uf_1 #38)
+#40 := (uf_2 #39)
+#527 := (= #40 0::int)
+#976 := (not #527)
+#502 := (<= #40 0::int)
+#971 := (not #502)
+#22 := 1::int
+#186 := (+ 1::int #40)
+#189 := (uf_1 #186)
+#506 := (uf_2 #189)
+#407 := (<= #506 1::int)
+#876 := (not #407)
+decl up_4 :: (-> T1 T1 bool)
+#4 := (:var 0 T1)
+#408 := (up_4 #4 #189)
+#393 := (pattern #408)
+#413 := (= #4 #189)
+#414 := (not #408)
+#26 := (uf_1 1::int)
+#27 := (= #4 #26)
+#392 := (or #27 #414 #413)
+#397 := (forall (vars (?x5 T1)) (:pat #393) #392)
+#383 := (not #397)
+#382 := (or #383 #407)
+#375 := (not #382)
+decl up_3 :: (-> T1 bool)
+#192 := (up_3 #189)
+#404 := (not #192)
+#841 := (or #404 #375)
+decl ?x5!0 :: (-> T1 T1)
+#422 := (?x5!0 #189)
+#434 := (= #189 #422)
+#417 := (up_4 #422 #189)
+#418 := (not #417)
+#415 := (= #26 #422)
+#847 := (or #415 #418 #434)
+#850 := (not #847)
+#853 := (or #192 #407 #850)
+#856 := (not #853)
+#844 := (not #841)
+#859 := (or #844 #856)
+#862 := (not #859)
+#5 := (uf_2 #4)
+#787 := (pattern #5)
+#21 := (up_3 #4)
+#835 := (pattern #21)
+#210 := (?x5!0 #4)
+#274 := (= #4 #210)
+#271 := (= #26 #210)
+#232 := (up_4 #210 #4)
+#233 := (not #232)
+#277 := (or #233 #271 #274)
+#280 := (not #277)
+#163 := (<= #5 1::int)
+#289 := (or #21 #163 #280)
+#304 := (not #289)
+#24 := (:var 1 T1)
+#25 := (up_4 #4 #24)
+#808 := (pattern #25)
+#28 := (= #4 #24)
+#147 := (not #25)
+#167 := (or #147 #27 #28)
+#809 := (forall (vars (?x5 T1)) (:pat #808) #167)
+#814 := (not #809)
+#817 := (or #163 #814)
+#820 := (not #817)
+#253 := (not #21)
+#823 := (or #253 #820)
+#826 := (not #823)
+#829 := (or #826 #304)
+#832 := (not #829)
+#836 := (forall (vars (?x4 T1)) (:pat #835 #787) #832)
+#170 := (forall (vars (?x5 T1)) #167)
+#236 := (not #170)
+#239 := (or #163 #236)
+#240 := (not #239)
+#215 := (or #253 #240)
+#303 := (not #215)
+#305 := (or #303 #304)
+#306 := (not #305)
+#311 := (forall (vars (?x4 T1)) #306)
+#837 := (iff #311 #836)
+#833 := (iff #306 #832)
+#830 := (iff #305 #829)
+#827 := (iff #303 #826)
+#824 := (iff #215 #823)
+#821 := (iff #240 #820)
+#818 := (iff #239 #817)
+#815 := (iff #236 #814)
+#812 := (iff #170 #809)
+#810 := (iff #167 #167)
+#811 := [refl]: #810
+#813 := [quant-intro #811]: #812
+#816 := [monotonicity #813]: #815
+#819 := [monotonicity #816]: #818
+#822 := [monotonicity #819]: #821
+#825 := [monotonicity #822]: #824
+#828 := [monotonicity #825]: #827
+#831 := [monotonicity #828]: #830
+#834 := [monotonicity #831]: #833
+#838 := [quant-intro #834]: #837
+#164 := (not #163)
+#173 := (and #164 #170)
+#259 := (or #253 #173)
+#294 := (and #259 #289)
+#297 := (forall (vars (?x4 T1)) #294)
+#312 := (iff #297 #311)
+#309 := (iff #294 #306)
+#214 := (and #215 #289)
+#307 := (iff #214 #306)
+#308 := [rewrite]: #307
+#301 := (iff #294 #214)
+#216 := (iff #259 #215)
+#268 := (iff #173 #240)
+#300 := [rewrite]: #268
+#213 := [monotonicity #300]: #216
+#302 := [monotonicity #213]: #301
+#310 := [trans #302 #308]: #309
+#313 := [quant-intro #310]: #312
+#230 := (= #210 #4)
+#231 := (= #210 #26)
+#234 := (or #233 #231 #230)
+#235 := (not #234)
+#228 := (not #164)
+#241 := (or #228 #235)
+#258 := (or #21 #241)
+#260 := (and #259 #258)
+#263 := (forall (vars (?x4 T1)) #260)
+#298 := (iff #263 #297)
+#295 := (iff #260 #294)
+#292 := (iff #258 #289)
+#283 := (or #163 #280)
+#286 := (or #21 #283)
+#290 := (iff #286 #289)
+#291 := [rewrite]: #290
+#287 := (iff #258 #286)
+#284 := (iff #241 #283)
+#281 := (iff #235 #280)
+#278 := (iff #234 #277)
+#275 := (iff #230 #274)
+#276 := [rewrite]: #275
+#272 := (iff #231 #271)
+#273 := [rewrite]: #272
+#279 := [monotonicity #273 #276]: #278
+#282 := [monotonicity #279]: #281
+#269 := (iff #228 #163)
+#270 := [rewrite]: #269
+#285 := [monotonicity #270 #282]: #284
+#288 := [monotonicity #285]: #287
+#293 := [trans #288 #291]: #292
+#296 := [monotonicity #293]: #295
+#299 := [quant-intro #296]: #298
+#176 := (iff #21 #173)
+#179 := (forall (vars (?x4 T1)) #176)
+#264 := (~ #179 #263)
+#261 := (~ #176 #260)
+#251 := (~ #173 #173)
+#249 := (~ #170 #170)
+#247 := (~ #167 #167)
+#248 := [refl]: #247
+#250 := [nnf-pos #248]: #249
+#245 := (~ #164 #164)
+#246 := [refl]: #245
+#252 := [monotonicity #246 #250]: #251
+#242 := (not #173)
+#243 := (~ #242 #241)
+#237 := (~ #236 #235)
+#238 := [sk]: #237
+#229 := (~ #228 #228)
+#209 := [refl]: #229
+#244 := [nnf-neg #209 #238]: #243
+#256 := (~ #21 #21)
+#257 := [refl]: #256
+#254 := (~ #253 #253)
+#255 := [refl]: #254
+#262 := [nnf-pos #255 #257 #244 #252]: #261
+#265 := [nnf-pos #262]: #264
+#29 := (or #27 #28)
+#30 := (implies #25 #29)
+#31 := (forall (vars (?x5 T1)) #30)
+#23 := (< 1::int #5)
+#32 := (and #23 #31)
+#33 := (iff #21 #32)
+#34 := (forall (vars (?x4 T1)) #33)
+#182 := (iff #34 #179)
+#148 := (or #147 #29)
+#151 := (forall (vars (?x5 T1)) #148)
+#154 := (and #23 #151)
+#157 := (iff #21 #154)
+#160 := (forall (vars (?x4 T1)) #157)
+#180 := (iff #160 #179)
+#177 := (iff #157 #176)
+#174 := (iff #154 #173)
+#171 := (iff #151 #170)
+#168 := (iff #148 #167)
+#169 := [rewrite]: #168
+#172 := [quant-intro #169]: #171
+#165 := (iff #23 #164)
+#166 := [rewrite]: #165
+#175 := [monotonicity #166 #172]: #174
+#178 := [monotonicity #175]: #177
+#181 := [quant-intro #178]: #180
+#161 := (iff #34 #160)
+#158 := (iff #33 #157)
+#155 := (iff #32 #154)
+#152 := (iff #31 #151)
+#149 := (iff #30 #148)
+#150 := [rewrite]: #149
+#153 := [quant-intro #150]: #152
+#156 := [monotonicity #153]: #155
+#159 := [monotonicity #156]: #158
+#162 := [quant-intro #159]: #161
+#183 := [trans #162 #181]: #182
+#146 := [asserted]: #34
+#184 := [mp #146 #183]: #179
+#266 := [mp~ #184 #265]: #263
+#267 := [mp #266 #299]: #297
+#314 := [mp #267 #313]: #311
+#839 := [mp #314 #838]: #836
+#589 := (not #836)
+#865 := (or #589 #862)
+#416 := (or #418 #415 #434)
+#419 := (not #416)
+#409 := (or #192 #407 #419)
+#410 := (not #409)
+#389 := (or #414 #27 #413)
+#394 := (forall (vars (?x5 T1)) (:pat #393) #389)
+#399 := (not #394)
+#401 := (or #407 #399)
+#402 := (not #401)
+#400 := (or #404 #402)
+#405 := (not #400)
+#388 := (or #405 #410)
+#391 := (not #388)
+#866 := (or #589 #391)
+#868 := (iff #866 #865)
+#870 := (iff #865 #865)
+#871 := [rewrite]: #870
+#863 := (iff #391 #862)
+#860 := (iff #388 #859)
+#857 := (iff #410 #856)
+#854 := (iff #409 #853)
+#851 := (iff #419 #850)
+#848 := (iff #416 #847)
+#849 := [rewrite]: #848
+#852 := [monotonicity #849]: #851
+#855 := [monotonicity #852]: #854
+#858 := [monotonicity #855]: #857
+#845 := (iff #405 #844)
+#842 := (iff #400 #841)
+#378 := (iff #402 #375)
+#376 := (iff #401 #382)
+#384 := (or #407 #383)
+#387 := (iff #384 #382)
+#374 := [rewrite]: #387
+#385 := (iff #401 #384)
+#380 := (iff #399 #383)
+#390 := (iff #394 #397)
+#395 := (iff #389 #392)
+#396 := [rewrite]: #395
+#398 := [quant-intro #396]: #390
+#381 := [monotonicity #398]: #380
+#386 := [monotonicity #381]: #385
+#377 := [trans #386 #374]: #376
+#840 := [monotonicity #377]: #378
+#843 := [monotonicity #840]: #842
+#846 := [monotonicity #843]: #845
+#861 := [monotonicity #846 #858]: #860
+#864 := [monotonicity #861]: #863
+#869 := [monotonicity #864]: #868
+#872 := [trans #869 #871]: #868
+#867 := [quant-inst]: #866
+#873 := [mp #867 #872]: #865
+#947 := [unit-resolution #873 #839]: #862
+#905 := (or #859 #841)
+#906 := [def-axiom]: #905
+#948 := [unit-resolution #906 #947]: #841
+#951 := (or #844 #375)
+#41 := (+ #40 1::int)
+#42 := (uf_1 #41)
+#43 := (up_3 #42)
+#193 := (iff #43 #192)
+#190 := (= #42 #189)
+#187 := (= #41 #186)
+#188 := [rewrite]: #187
+#191 := [monotonicity #188]: #190
+#194 := [monotonicity #191]: #193
+#185 := [asserted]: #43
+#197 := [mp #185 #194]: #192
+#885 := (or #844 #404 #375)
+#886 := [def-axiom]: #885
+#952 := [unit-resolution #886 #197]: #951
+#953 := [unit-resolution #952 #948]: #375
+#877 := (or #382 #876)
+#878 := [def-axiom]: #877
+#954 := [unit-resolution #878 #953]: #876
+#542 := -1::int
+#508 := (* -1::int #506)
+#493 := (+ #40 #508)
+#438 := (>= #493 -1::int)
+#494 := (= #493 -1::int)
+#496 := (>= #40 -1::int)
+#451 := (= #506 0::int)
+#959 := (not #451)
+#432 := (<= #506 0::int)
+#955 := (not #432)
+#956 := (or #955 #407)
+#957 := [th-lemma]: #956
+#958 := [unit-resolution #957 #954]: #955
+#960 := (or #959 #432)
+#961 := [th-lemma]: #960
+#962 := [unit-resolution #961 #958]: #959
+#453 := (or #451 #496)
+#10 := (:var 0 int)
+#12 := (uf_1 #10)
+#795 := (pattern #12)
+#87 := (>= #10 0::int)
+#13 := (uf_2 #12)
+#18 := (= #13 0::int)
+#135 := (or #18 #87)
+#802 := (forall (vars (?x3 int)) (:pat #795) #135)
+#140 := (forall (vars (?x3 int)) #135)
+#805 := (iff #140 #802)
+#803 := (iff #135 #135)
+#804 := [refl]: #803
+#806 := [quant-intro #804]: #805
+#207 := (~ #140 #140)
+#225 := (~ #135 #135)
+#226 := [refl]: #225
+#208 := [nnf-pos #226]: #207
+#17 := (< #10 0::int)
+#19 := (implies #17 #18)
+#20 := (forall (vars (?x3 int)) #19)
+#143 := (iff #20 #140)
+#106 := (= 0::int #13)
+#112 := (not #17)
+#113 := (or #112 #106)
+#118 := (forall (vars (?x3 int)) #113)
+#141 := (iff #118 #140)
+#138 := (iff #113 #135)
+#132 := (or #87 #18)
+#136 := (iff #132 #135)
+#137 := [rewrite]: #136
+#133 := (iff #113 #132)
+#130 := (iff #106 #18)
+#131 := [rewrite]: #130
+#128 := (iff #112 #87)
+#88 := (not #87)
+#123 := (not #88)
+#126 := (iff #123 #87)
+#127 := [rewrite]: #126
+#124 := (iff #112 #123)
+#121 := (iff #17 #88)
+#122 := [rewrite]: #121
+#125 := [monotonicity #122]: #124
+#129 := [trans #125 #127]: #128
+#134 := [monotonicity #129 #131]: #133
+#139 := [trans #134 #137]: #138
+#142 := [quant-intro #139]: #141
+#119 := (iff #20 #118)
+#116 := (iff #19 #113)
+#109 := (implies #17 #106)
+#114 := (iff #109 #113)
+#115 := [rewrite]: #114
+#110 := (iff #19 #109)
+#107 := (iff #18 #106)
+#108 := [rewrite]: #107
+#111 := [monotonicity #108]: #110
+#117 := [trans #111 #115]: #116
+#120 := [quant-intro #117]: #119
+#144 := [trans #120 #142]: #143
+#105 := [asserted]: #20
+#145 := [mp #105 #144]: #140
+#227 := [mp~ #145 #208]: #140
+#807 := [mp #227 #806]: #802
+#514 := (not #802)
+#445 := (or #514 #451 #496)
+#504 := (>= #186 0::int)
+#452 := (or #451 #504)
+#456 := (or #514 #452)
+#429 := (iff #456 #445)
+#441 := (or #514 #453)
+#423 := (iff #441 #445)
+#428 := [rewrite]: #423
+#442 := (iff #456 #441)
+#454 := (iff #452 #453)
+#498 := (iff #504 #496)
+#487 := [rewrite]: #498
+#455 := [monotonicity #487]: #454
+#421 := [monotonicity #455]: #442
+#430 := [trans #421 #428]: #429
+#439 := [quant-inst]: #456
+#431 := [mp #439 #430]: #445
+#963 := [unit-resolution #431 #807]: #453
+#964 := [unit-resolution #963 #962]: #496
+#488 := (not #496)
+#490 := (or #494 #488)
+#69 := (= #10 #13)
+#94 := (or #69 #88)
+#796 := (forall (vars (?x2 int)) (:pat #795) #94)
+#99 := (forall (vars (?x2 int)) #94)
+#799 := (iff #99 #796)
+#797 := (iff #94 #94)
+#798 := [refl]: #797
+#800 := [quant-intro #798]: #799
+#206 := (~ #99 #99)
+#222 := (~ #94 #94)
+#223 := [refl]: #222
+#196 := [nnf-pos #223]: #206
+#14 := (= #13 #10)
+#11 := (<= 0::int #10)
+#15 := (implies #11 #14)
+#16 := (forall (vars (?x2 int)) #15)
+#102 := (iff #16 #99)
+#76 := (not #11)
+#77 := (or #76 #69)
+#82 := (forall (vars (?x2 int)) #77)
+#100 := (iff #82 #99)
+#97 := (iff #77 #94)
+#91 := (or #88 #69)
+#95 := (iff #91 #94)
+#96 := [rewrite]: #95
+#92 := (iff #77 #91)
+#89 := (iff #76 #88)
+#85 := (iff #11 #87)
+#86 := [rewrite]: #85
+#90 := [monotonicity #86]: #89
+#93 := [monotonicity #90]: #92
+#98 := [trans #93 #96]: #97
+#101 := [quant-intro #98]: #100
+#83 := (iff #16 #82)
+#80 := (iff #15 #77)
+#73 := (implies #11 #69)
+#78 := (iff #73 #77)
+#79 := [rewrite]: #78
+#74 := (iff #15 #73)
+#71 := (iff #14 #69)
+#72 := [rewrite]: #71
+#75 := [monotonicity #72]: #74
+#81 := [trans #75 #79]: #80
+#84 := [quant-intro #81]: #83
+#103 := [trans #84 #101]: #102
+#68 := [asserted]: #16
+#104 := [mp #68 #103]: #99
+#224 := [mp~ #104 #196]: #99
+#801 := [mp #224 #800]: #796
+#530 := (not #796)
+#492 := (or #530 #494 #488)
+#505 := (not #504)
+#507 := (= #186 #506)
+#500 := (or #507 #505)
+#473 := (or #530 #500)
+#478 := (iff #473 #492)
+#475 := (or #530 #490)
+#477 := (iff #475 #492)
+#467 := [rewrite]: #477
+#466 := (iff #473 #475)
+#491 := (iff #500 #490)
+#489 := (iff #505 #488)
+#481 := [monotonicity #487]: #489
+#495 := (iff #507 #494)
+#497 := [rewrite]: #495
+#482 := [monotonicity #497 #481]: #491
+#476 := [monotonicity #482]: #466
+#444 := [trans #476 #467]: #478
+#474 := [quant-inst]: #473
+#446 := [mp #474 #444]: #492
+#965 := [unit-resolution #446 #801]: #490
+#966 := [unit-resolution #965 #964]: #494
+#967 := (not #494)
+#968 := (or #967 #438)
+#969 := [th-lemma]: #968
+#970 := [unit-resolution #969 #966]: #438
+#972 := (not #438)
+#973 := (or #971 #407 #972)
+#974 := [th-lemma]: #973
+#975 := [unit-resolution #974 #970 #954]: #971
+#977 := (or #976 #502)
+#978 := [th-lemma]: #977
+#979 := [unit-resolution #978 #975]: #976
+#553 := (>= #37 0::int)
+#546 := (not #553)
+#545 := (* -1::int #40)
+#549 := (+ #38 #545)
+#551 := (= #549 0::int)
+#984 := (not #551)
+#524 := (>= #549 0::int)
+#980 := (not #524)
+#201 := (>= #37 1::int)
+#202 := (not #201)
+#44 := (<= 1::int #37)
+#45 := (not #44)
+#203 := (iff #45 #202)
+#199 := (iff #44 #201)
+#200 := [rewrite]: #199
+#204 := [monotonicity #200]: #203
+#195 := [asserted]: #45
+#205 := [mp #195 #204]: #202
+#981 := (or #980 #201 #407 #972)
+#982 := [th-lemma]: #981
+#983 := [unit-resolution #982 #205 #970 #954]: #980
+#985 := (or #984 #524)
+#986 := [th-lemma]: #985
+#987 := [unit-resolution #986 #983]: #984
+#548 := (or #551 #546)
+#531 := (or #530 #551 #546)
+#403 := (>= #38 0::int)
+#562 := (not #403)
+#558 := (= #38 #40)
+#563 := (or #558 #562)
+#534 := (or #530 #563)
+#537 := (iff #534 #531)
+#539 := (or #530 #548)
+#533 := (iff #539 #531)
+#536 := [rewrite]: #533
+#532 := (iff #534 #539)
+#538 := (iff #563 #548)
+#547 := (iff #562 #546)
+#541 := (iff #403 #553)
+#544 := [rewrite]: #541
+#543 := [monotonicity #544]: #547
+#552 := (iff #558 #551)
+#550 := [rewrite]: #552
+#528 := [monotonicity #550 #543]: #538
+#540 := [monotonicity #528]: #532
+#523 := [trans #540 #536]: #537
+#535 := [quant-inst]: #534
+#525 := [mp #535 #523]: #531
+#988 := [unit-resolution #525 #801]: #548
+#989 := [unit-resolution #988 #987]: #546
+#511 := (or #527 #553)
+#515 := (or #514 #527 #553)
+#509 := (or #527 #403)
+#516 := (or #514 #509)
+#522 := (iff #516 #515)
+#518 := (or #514 #511)
+#521 := (iff #518 #515)
+#510 := [rewrite]: #521
+#519 := (iff #516 #518)
+#512 := (iff #509 #511)
+#513 := [monotonicity #544]: #512
+#520 := [monotonicity #513]: #519
+#499 := [trans #520 #510]: #522
+#517 := [quant-inst]: #516
+#501 := [mp #517 #499]: #515
+#990 := [unit-resolution #501 #807]: #511
+[unit-resolution #990 #989 #979]: false
+unsat
+
+uq2MbDTgTG1nxWdwzZtWew 7
+unsat
+
+E5BydeDaPocMMvChMGY+og 7
+unsat
+
+p81EQzqwJrGunGO7GuNt3g 7
+unsat
+
+KpYfvnTcz2WncWNg3dJDCA 7
+unsat
+
+ybGRm230DLO0wH6aROKBBw 7
+unsat
+
+goFtZfJ6kkxA8sqBVpZutw 7
+unsat
+
+0+nmgsMqioeTuwam1ScP7g 7
+unsat
+
+nI63LP/VCL//bjsS1gNB2A 7
+unsat
+
+9+2QHvrRgbKz97Zg0kfybw 7
+unsat
+
+6kquszLXeBUhTwzaw6gd2Q 7
+unsat
+
+j5Z04lpza+N5I1cpno5mtw 7
+unsat
+
+mapbfUM6Ils30x5nEBJmaw 7
+unsat
+
+e8P++0FczY3zhNhEVclACw 7
+unsat
+
+yXMQNOyCylhI+EH8hNYxHA 7
+unsat
+
+GkYN9j7cjrR2KR/lb04/qw 7
+unsat
+
+PajzgNjLWHwVHpjoje+gnA 7
+unsat
+
+URpJYU7D8PO0VRnciRgE5A 7
+unsat
+
+D9ZGhymoV3L6zbWsJlwG2A 7
+unsat
+
+0QLuovrnnANWnCkUY3l10g 7
+unsat
+
+CYprps2G0Au5F3Z7n3KTRg 7
+unsat
+
+iyIMuJd6zijfEao8zKnx2w 7
+unsat
+
+49jzsAwAEfR6NSFBhBEisQ 7
+unsat
+
+T0j6xFgrghxif91jL+2yAw 7
+unsat
+
+md/M3rVve0+8sQ6oqIoL2w 7
+unsat
+
+pY7C8PCf5lVVaim6q7PJcQ 7
+unsat
+
+4zCFLQf4Jrov/gmEvsBm4Q 1036
+#2 := false
+#6 := 0::int
+decl uf_1 :: (-> bv[2] int)
+#4 := bv[0:2]
+#5 := (uf_1 bv[0:2])
+#225 := (<= #5 0::int)
+#309 := (not #225)
+#20 := (:var 0 bv[2])
+#21 := (uf_1 #20)
+#638 := (pattern #21)
+#54 := (<= #21 0::int)
+#55 := (not #54)
+#639 := (forall (vars (?x1 bv[2])) (:pat #638) #55)
+#58 := (forall (vars (?x1 bv[2])) #55)
+#642 := (iff #58 #639)
+#640 := (iff #55 #55)
+#641 := [refl]: #640
+#643 := [quant-intro #641]: #642
+#113 := (~ #58 #58)
+#115 := (~ #55 #55)
+#116 := [refl]: #115
+#114 := [nnf-pos #116]: #113
+#22 := (< 0::int #21)
+#23 := (forall (vars (?x1 bv[2])) #22)
+#59 := (iff #23 #58)
+#56 := (iff #22 #55)
+#57 := [rewrite]: #56
+#60 := [quant-intro #57]: #59
+#51 := [asserted]: #23
+#61 := [mp #51 #60]: #58
+#111 := [mp~ #61 #114]: #58
+#644 := [mp #111 #643]: #639
+#302 := (not #639)
+#313 := (or #302 #309)
+#314 := [quant-inst]: #313
+#635 := [unit-resolution #314 #644]: #309
+#7 := (= #5 0::int)
+#47 := [asserted]: #7
+#637 := (not #7)
+#627 := (or #637 #225)
+#287 := [th-lemma]: #627
+[unit-resolution #287 #47 #635]: false
+unsat
+
+czvSLyjMowmFNi82us4N2Q 7
+unsat
+
+aU+7kcyE8oAPbs5RjfuwIw 1160
+#2 := false
+decl uf_6 :: T2
+#23 := uf_6
+decl uf_4 :: T2
+#19 := uf_4
+#25 := (= uf_4 uf_6)
+decl uf_2 :: (-> T1 T2)
+decl uf_1 :: (-> T2 T3 T1)
+decl uf_5 :: T3
+#20 := uf_5
+#21 := (uf_1 uf_4 uf_5)
+#22 := (uf_2 #21)
+#24 := (= #22 uf_6)
+#65 := [asserted]: #24
+#143 := (= uf_4 #22)
+#11 := (:var 0 T3)
+#10 := (:var 1 T2)
+#12 := (uf_1 #10 #11)
+#567 := (pattern #12)
+#16 := (uf_2 #12)
+#58 := (= #10 #16)
+#574 := (forall (vars (?x4 T2) (?x5 T3)) (:pat #567) #58)
+#62 := (forall (vars (?x4 T2) (?x5 T3)) #58)
+#577 := (iff #62 #574)
+#575 := (iff #58 #58)
+#576 := [refl]: #575
+#578 := [quant-intro #576]: #577
+#71 := (~ #62 #62)
+#87 := (~ #58 #58)
+#88 := [refl]: #87
+#72 := [nnf-pos #88]: #71
+#17 := (= #16 #10)
+#18 := (forall (vars (?x4 T2) (?x5 T3)) #17)
+#63 := (iff #18 #62)
+#60 := (iff #17 #58)
+#61 := [rewrite]: #60
+#64 := [quant-intro #61]: #63
+#57 := [asserted]: #18
+#67 := [mp #57 #64]: #62
+#89 := [mp~ #67 #72]: #62
+#579 := [mp #89 #578]: #574
+#214 := (not #574)
+#551 := (or #214 #143)
+#553 := [quant-inst]: #551
+#233 := [unit-resolution #553 #579]: #143
+#235 := [trans #233 #65]: #25
+#26 := (not #25)
+#66 := [asserted]: #26
+[unit-resolution #66 #235]: false
+unsat
+
+dXfueqZAXkudfE6G0VKkwg 2559
+#2 := false
+decl uf_6 :: (-> T4 T2)
+decl uf_10 :: T4
+#39 := uf_10
+#44 := (uf_6 uf_10)
+decl uf_2 :: (-> T1 T2)
+decl uf_7 :: T1
+#34 := uf_7
+#43 := (uf_2 uf_7)
+#45 := (= #43 #44)
+decl uf_4 :: (-> T3 T2 T4)
+decl uf_8 :: T2
+#35 := uf_8
+decl uf_9 :: T3
+#36 := uf_9
+#40 := (uf_4 uf_9 uf_8)
+#204 := (uf_6 #40)
+#598 := (= #204 #44)
+#595 := (= #44 #204)
+#41 := (= uf_10 #40)
+decl uf_1 :: (-> T2 T3 T1)
+#37 := (uf_1 uf_8 uf_9)
+#38 := (= uf_7 #37)
+#42 := (and #38 #41)
+#109 := [asserted]: #42
+#114 := [and-elim #109]: #41
+#256 := [monotonicity #114]: #595
+#599 := [symm #256]: #598
+#596 := (= #43 #204)
+#269 := (= uf_8 #204)
+#23 := (:var 0 T2)
+#22 := (:var 1 T3)
+#24 := (uf_4 #22 #23)
+#643 := (pattern #24)
+#25 := (uf_6 #24)
+#86 := (= #23 #25)
+#644 := (forall (vars (?x5 T3) (?x6 T2)) (:pat #643) #86)
+#90 := (forall (vars (?x5 T3) (?x6 T2)) #86)
+#647 := (iff #90 #644)
+#645 := (iff #86 #86)
+#646 := [refl]: #645
+#648 := [quant-intro #646]: #647
+#119 := (~ #90 #90)
+#144 := (~ #86 #86)
+#145 := [refl]: #144
+#120 := [nnf-pos #145]: #119
+#26 := (= #25 #23)
+#27 := (forall (vars (?x5 T3) (?x6 T2)) #26)
+#91 := (iff #27 #90)
+#88 := (iff #26 #86)
+#89 := [rewrite]: #88
+#92 := [quant-intro #89]: #91
+#85 := [asserted]: #27
+#95 := [mp #85 #92]: #90
+#146 := [mp~ #95 #120]: #90
+#649 := [mp #146 #648]: #644
+#613 := (not #644)
+#619 := (or #613 #269)
+#609 := [quant-inst]: #619
+#267 := [unit-resolution #609 #649]: #269
+#600 := (= #43 uf_8)
+#289 := (uf_2 #37)
+#259 := (= #289 uf_8)
+#296 := (= uf_8 #289)
+#17 := (:var 0 T3)
+#16 := (:var 1 T2)
+#18 := (uf_1 #16 #17)
+#636 := (pattern #18)
+#28 := (uf_2 #18)
+#94 := (= #16 #28)
+#650 := (forall (vars (?x7 T2) (?x8 T3)) (:pat #636) #94)
+#98 := (forall (vars (?x7 T2) (?x8 T3)) #94)
+#653 := (iff #98 #650)
+#651 := (iff #94 #94)
+#652 := [refl]: #651
+#654 := [quant-intro #652]: #653
+#121 := (~ #98 #98)
+#147 := (~ #94 #94)
+#148 := [refl]: #147
+#122 := [nnf-pos #148]: #121
+#29 := (= #28 #16)
+#30 := (forall (vars (?x7 T2) (?x8 T3)) #29)
+#99 := (iff #30 #98)
+#96 := (iff #29 #94)
+#97 := [rewrite]: #96
+#100 := [quant-intro #97]: #99
+#93 := [asserted]: #30
+#103 := [mp #93 #100]: #98
+#149 := [mp~ #103 #122]: #98
+#655 := [mp #149 #654]: #650
+#615 := (not #650)
+#616 := (or #615 #296)
+#617 := [quant-inst]: #616
+#618 := [unit-resolution #617 #655]: #296
+#597 := [symm #618]: #259
+#611 := (= #43 #289)
+#113 := [and-elim #109]: #38
+#252 := [monotonicity #113]: #611
+#601 := [trans #252 #597]: #600
+#602 := [trans #601 #267]: #596
+#238 := [trans #602 #599]: #45
+#46 := (not #45)
+#110 := [asserted]: #46
+[unit-resolution #110 #238]: false
+unsat
+
+Dc/6bNJffjYYplvoitScJQ 4578
+#2 := false
+decl uf_1 :: (-> T1 T2 T3)
+decl uf_3 :: T2
+#22 := uf_3
+decl uf_6 :: T1
+#30 := uf_6
+#36 := (uf_1 uf_6 uf_3)
+decl uf_2 :: (-> T1 T2 T3 T1)
+decl uf_8 :: T3
+#33 := uf_8
+decl uf_5 :: T2
+#26 := uf_5
+decl uf_7 :: T3
+#31 := uf_7
+decl uf_4 :: T2
+#23 := uf_4
+#32 := (uf_2 uf_6 uf_4 uf_7)
+#34 := (uf_2 #32 uf_5 uf_8)
+#35 := (uf_1 #34 uf_3)
+#37 := (= #35 #36)
+#223 := (uf_1 #32 uf_4)
+#214 := (uf_2 uf_6 uf_4 #223)
+#552 := (uf_1 #214 uf_3)
+#555 := (= #552 #36)
+#560 := (= #36 #552)
+#556 := (= #223 #552)
+#24 := (= uf_3 uf_4)
+#561 := (ite #24 #556 #560)
+#8 := (:var 0 T2)
+#6 := (:var 1 T3)
+#5 := (:var 2 T2)
+#4 := (:var 3 T1)
+#7 := (uf_2 #4 #5 #6)
+#9 := (uf_1 #7 #8)
+#575 := (pattern #9)
+#11 := (uf_1 #4 #8)
+#100 := (= #9 #11)
+#99 := (= #6 #9)
+#55 := (= #5 #8)
+#83 := (ite #55 #99 #100)
+#576 := (forall (vars (?x1 T1) (?x2 T2) (?x3 T3) (?x4 T2)) (:pat #575) #83)
+#90 := (forall (vars (?x1 T1) (?x2 T2) (?x3 T3) (?x4 T2)) #83)
+#579 := (iff #90 #576)
+#577 := (iff #83 #83)
+#578 := [refl]: #577
+#580 := [quant-intro #578]: #579
+#58 := (ite #55 #6 #11)
+#61 := (= #9 #58)
+#64 := (forall (vars (?x1 T1) (?x2 T2) (?x3 T3) (?x4 T2)) #61)
+#87 := (iff #64 #90)
+#84 := (iff #61 #83)
+#89 := [rewrite]: #84
+#88 := [quant-intro #89]: #87
+#93 := (~ #64 #64)
+#91 := (~ #61 #61)
+#92 := [refl]: #91
+#94 := [nnf-pos #92]: #93
+#10 := (= #8 #5)
+#12 := (ite #10 #6 #11)
+#13 := (= #9 #12)
+#14 := (forall (vars (?x1 T1) (?x2 T2) (?x3 T3) (?x4 T2)) #13)
+#65 := (iff #14 #64)
+#62 := (iff #13 #61)
+#59 := (= #12 #58)
+#56 := (iff #10 #55)
+#57 := [rewrite]: #56
+#60 := [monotonicity #57]: #59
+#63 := [monotonicity #60]: #62
+#66 := [quant-intro #63]: #65
+#54 := [asserted]: #14
+#69 := [mp #54 #66]: #64
+#95 := [mp~ #69 #94]: #64
+#85 := [mp #95 #88]: #90
+#581 := [mp #85 #580]: #576
+#250 := (not #576)
+#548 := (or #250 #561)
+#551 := (= uf_4 uf_3)
+#557 := (ite #551 #556 #555)
+#549 := (or #250 #557)
+#271 := (iff #549 #548)
+#273 := (iff #548 #548)
+#259 := [rewrite]: #273
+#559 := (iff #557 #561)
+#198 := (iff #555 #560)
+#199 := [rewrite]: #198
+#193 := (iff #551 #24)
+#558 := [rewrite]: #193
+#562 := [monotonicity #558 #199]: #559
+#272 := [monotonicity #562]: #271
+#274 := [trans #272 #259]: #271
+#255 := [quant-inst]: #549
+#165 := [mp #255 #274]: #548
+#510 := [unit-resolution #165 #581]: #561
+#544 := (not #561)
+#497 := (or #544 #560)
+#25 := (not #24)
+#27 := (= uf_3 uf_5)
+#28 := (not #27)
+#29 := (and #25 #28)
+#75 := [asserted]: #29
+#79 := [and-elim #75]: #25
+#268 := (or #544 #24 #560)
+#542 := [def-axiom]: #268
+#499 := [unit-resolution #542 #79]: #497
+#491 := [unit-resolution #499 #510]: #560
+#493 := [symm #491]: #555
+#494 := (= #35 #552)
+#157 := (uf_1 #32 uf_3)
+#503 := (= #157 #552)
+#502 := (= #552 #157)
+#509 := (= #214 #32)
+#415 := (= #223 uf_7)
+#566 := (= uf_7 #223)
+#17 := (:var 0 T3)
+#16 := (:var 1 T2)
+#15 := (:var 2 T1)
+#18 := (uf_2 #15 #16 #17)
+#582 := (pattern #18)
+#19 := (uf_1 #18 #16)
+#68 := (= #17 #19)
+#584 := (forall (vars (?x5 T1) (?x6 T2) (?x7 T3)) (:pat #582) #68)
+#72 := (forall (vars (?x5 T1) (?x6 T2) (?x7 T3)) #68)
+#583 := (iff #72 #584)
+#586 := (iff #584 #584)
+#587 := [rewrite]: #586
+#585 := [rewrite]: #583
+#588 := [trans #585 #587]: #583
+#82 := (~ #72 #72)
+#96 := (~ #68 #68)
+#97 := [refl]: #96
+#78 := [nnf-pos #97]: #82
+#20 := (= #19 #17)
+#21 := (forall (vars (?x5 T1) (?x6 T2) (?x7 T3)) #20)
+#73 := (iff #21 #72)
+#70 := (iff #20 #68)
+#71 := [rewrite]: #70
+#74 := [quant-intro #71]: #73
+#67 := [asserted]: #21
+#77 := [mp #67 #74]: #72
+#98 := [mp~ #77 #78]: #72
+#589 := [mp #98 #588]: #584
+#211 := (not #584)
+#212 := (or #211 #566)
+#213 := [quant-inst]: #212
+#414 := [unit-resolution #213 #589]: #566
+#416 := [symm #414]: #415
+#506 := [monotonicity #416]: #509
+#498 := [monotonicity #506]: #502
+#492 := [symm #498]: #503
+#244 := (= #35 #157)
+#158 := (= uf_8 #35)
+#248 := (ite #27 #158 #244)
+#247 := (or #250 #248)
+#245 := (= uf_5 uf_3)
+#159 := (ite #245 #158 #244)
+#251 := (or #250 #159)
+#567 := (iff #251 #247)
+#224 := (iff #247 #247)
+#356 := [rewrite]: #224
+#249 := (iff #159 #248)
+#246 := (iff #245 #27)
+#237 := [rewrite]: #246
+#177 := [monotonicity #237]: #249
+#569 := [monotonicity #177]: #567
+#563 := [trans #569 #356]: #567
+#230 := [quant-inst]: #251
+#235 := [mp #230 #563]: #247
+#488 := [unit-resolution #235 #581]: #248
+#236 := (not #248)
+#490 := (or #236 #244)
+#80 := [and-elim #75]: #28
+#572 := (or #236 #27 #244)
+#573 := [def-axiom]: #572
+#500 := [unit-resolution #573 #80]: #490
+#501 := [unit-resolution #500 #488]: #244
+#495 := [trans #501 #492]: #494
+#489 := [trans #495 #493]: #37
+#38 := (not #37)
+#76 := [asserted]: #38
+[unit-resolution #76 #489]: false
+unsat
+
+jdmsd1j41Osn+WzTtqTUSQ 1352
+#2 := false
+decl up_4 :: (-> T1 T2 bool)
+decl uf_3 :: T2
+#5 := uf_3
+decl uf_2 :: T1
+#4 := uf_2
+#7 := (up_4 uf_2 uf_3)
+#60 := (not #7)
+decl up_1 :: (-> T1 T2 bool)
+#6 := (up_1 uf_2 uf_3)
+#33 := (iff #6 #7)
+#49 := (or #6 #7 #33)
+#52 := (not #49)
+#1 := true
+#11 := (iff #7 true)
+#10 := (iff #6 true)
+#12 := (or #10 #11)
+#8 := (and #7 true)
+#9 := (iff #6 #8)
+#13 := (or #9 #12)
+#14 := (not #13)
+#55 := (iff #14 #52)
+#40 := (or #6 #7)
+#43 := (or #33 #40)
+#46 := (not #43)
+#53 := (iff #46 #52)
+#50 := (iff #43 #49)
+#51 := [rewrite]: #50
+#54 := [monotonicity #51]: #53
+#47 := (iff #14 #46)
+#44 := (iff #13 #43)
+#41 := (iff #12 #40)
+#38 := (iff #11 #7)
+#39 := [rewrite]: #38
+#36 := (iff #10 #6)
+#37 := [rewrite]: #36
+#42 := [monotonicity #37 #39]: #41
+#34 := (iff #9 #33)
+#31 := (iff #8 #7)
+#32 := [rewrite]: #31
+#35 := [monotonicity #32]: #34
+#45 := [monotonicity #35 #42]: #44
+#48 := [monotonicity #45]: #47
+#56 := [trans #48 #54]: #55
+#30 := [asserted]: #14
+#57 := [mp #30 #56]: #52
+#61 := [not-or-elim #57]: #60
+#58 := (not #6)
+#59 := [not-or-elim #57]: #58
+#72 := (or #7 #6)
+#66 := (iff #7 #58)
+#62 := (not #33)
+#64 := (iff #62 #66)
+#67 := [rewrite]: #64
+#63 := [not-or-elim #57]: #62
+#68 := [mp #63 #67]: #66
+#69 := (not #66)
+#70 := (or #7 #6 #69)
+#71 := [def-axiom]: #70
+#73 := [unit-resolution #71 #68]: #72
+[unit-resolution #73 #59 #61]: false
+unsat
+
+EA8ecQ7czWL46/C3k7D7tg 2697
+#2 := false
+decl up_2 :: (-> T2 bool)
+decl uf_3 :: T2
+#10 := uf_3
+#17 := (up_2 uf_3)
+#78 := (not #17)
+decl uf_1 :: (-> T1 T1)
+decl uf_4 :: T1
+#14 := uf_4
+#15 := (uf_1 uf_4)
+#46 := (= uf_4 #15)
+#79 := (not #46)
+#145 := [hypothesis]: #79
+#4 := (:var 0 T1)
+#5 := (uf_1 #4)
+#563 := (pattern #5)
+#37 := (= #4 #5)
+#564 := (forall (vars (?x1 T1)) (:pat #563) #37)
+#40 := (forall (vars (?x1 T1)) #37)
+#567 := (iff #40 #564)
+#565 := (iff #37 #37)
+#566 := [refl]: #565
+#568 := [quant-intro #566]: #567
+#72 := (~ #40 #40)
+#70 := (~ #37 #37)
+#71 := [refl]: #70
+#73 := [nnf-pos #71]: #72
+#6 := (= #5 #4)
+#7 := (forall (vars (?x1 T1)) #6)
+#41 := (iff #7 #40)
+#38 := (iff #6 #37)
+#39 := [rewrite]: #38
+#42 := [quant-intro #39]: #41
+#36 := [asserted]: #7
+#45 := [mp #36 #42]: #40
+#74 := [mp~ #45 #73]: #40
+#569 := [mp #74 #568]: #564
+#146 := (not #564)
+#233 := (or #146 #46)
+#147 := [quant-inst]: #233
+#232 := [unit-resolution #147 #569 #145]: false
+#234 := [lemma #232]: #46
+#66 := (or #78 #79)
+#54 := (and #17 #46)
+#59 := (not #54)
+#85 := (iff #59 #66)
+#67 := (not #66)
+#80 := (not #67)
+#83 := (iff #80 #66)
+#84 := [rewrite]: #83
+#81 := (iff #59 #80)
+#68 := (iff #54 #67)
+#69 := [rewrite]: #68
+#82 := [monotonicity #69]: #81
+#86 := [trans #82 #84]: #85
+#1 := true
+#18 := (iff #17 true)
+#16 := (= #15 uf_4)
+#19 := (and #16 #18)
+#20 := (not #19)
+#60 := (iff #20 #59)
+#57 := (iff #19 #54)
+#51 := (and #46 #17)
+#55 := (iff #51 #54)
+#56 := [rewrite]: #55
+#52 := (iff #19 #51)
+#49 := (iff #18 #17)
+#50 := [rewrite]: #49
+#47 := (iff #16 #46)
+#48 := [rewrite]: #47
+#53 := [monotonicity #48 #50]: #52
+#58 := [trans #53 #56]: #57
+#61 := [monotonicity #58]: #60
+#44 := [asserted]: #20
+#64 := [mp #44 #61]: #59
+#87 := [mp #64 #86]: #66
+#561 := [unit-resolution #87 #234]: #78
+#8 := (:var 0 T2)
+#9 := (up_2 #8)
+#570 := (pattern #9)
+#11 := (= #8 uf_3)
+#12 := (iff #9 #11)
+#571 := (forall (vars (?x2 T2)) (:pat #570) #12)
+#13 := (forall (vars (?x2 T2)) #12)
+#574 := (iff #13 #571)
+#572 := (iff #12 #12)
+#573 := [refl]: #572
+#575 := [quant-intro #573]: #574
+#65 := (~ #13 #13)
+#75 := (~ #12 #12)
+#76 := [refl]: #75
+#62 := [nnf-pos #76]: #65
+#43 := [asserted]: #13
+#77 := [mp~ #43 #62]: #13
+#576 := [mp #77 #575]: #571
+#555 := (not #571)
+#557 := (or #555 #17)
+#225 := (= uf_3 uf_3)
+#236 := (iff #17 #225)
+#212 := (or #555 #236)
+#551 := (iff #212 #557)
+#224 := (iff #557 #557)
+#558 := [rewrite]: #224
+#239 := (iff #236 #17)
+#238 := (iff #236 #18)
+#237 := (iff #225 true)
+#165 := [rewrite]: #237
+#235 := [monotonicity #165]: #238
+#218 := [trans #235 #50]: #239
+#223 := [monotonicity #218]: #551
+#559 := [trans #223 #558]: #551
+#344 := [quant-inst]: #212
+#560 := [mp #344 #559]: #557
+[unit-resolution #560 #576 #561]: false
+unsat
+
+mNfbN3NQCB4ik2xJmLE1UQ 11936
+#2 := false
+decl uf_2 :: (-> T2 T3 T3)
+decl uf_4 :: T3
+#15 := uf_4
+decl uf_6 :: (-> int T2)
+#48 := 2::int
+#49 := (uf_6 2::int)
+#50 := (uf_2 #49 uf_4)
+#23 := 1::int
+#44 := (uf_6 1::int)
+#51 := (uf_2 #44 #50)
+decl uf_1 :: (-> T1 T3 T3)
+#45 := (uf_2 #44 uf_4)
+#31 := 0::int
+#43 := (uf_6 0::int)
+#46 := (uf_2 #43 #45)
+decl uf_5 :: T1
+#19 := uf_5
+#47 := (uf_1 uf_5 #46)
+#52 := (= #47 #51)
+#266 := (uf_1 uf_5 #45)
+decl uf_3 :: (-> T1 T2 T2)
+#351 := (uf_3 uf_5 #43)
+#267 := (uf_2 #351 #266)
+#791 := (= #267 #51)
+#789 := (= #51 #267)
+#752 := (= #50 #266)
+#521 := (uf_1 uf_5 uf_4)
+#522 := (uf_3 uf_5 #44)
+#615 := (uf_2 #522 #521)
+#750 := (= #615 #266)
+#612 := (= #266 #615)
+#6 := (:var 0 T3)
+#4 := (:var 2 T1)
+#10 := (uf_1 #4 #6)
+#5 := (:var 1 T2)
+#9 := (uf_3 #4 #5)
+#11 := (uf_2 #9 #10)
+#682 := (pattern #11)
+#7 := (uf_2 #5 #6)
+#8 := (uf_1 #4 #7)
+#681 := (pattern #8)
+#12 := (= #8 #11)
+#683 := (forall (vars (?x1 T1) (?x2 T2) (?x3 T3)) (:pat #681 #682) #12)
+#13 := (forall (vars (?x1 T1) (?x2 T2) (?x3 T3)) #12)
+#686 := (iff #13 #683)
+#684 := (iff #12 #12)
+#685 := [refl]: #684
+#687 := [quant-intro #685]: #686
+#195 := (~ #13 #13)
+#193 := (~ #12 #12)
+#194 := [refl]: #193
+#196 := [nnf-pos #194]: #195
+#69 := [asserted]: #13
+#197 := [mp~ #69 #196]: #13
+#688 := [mp #197 #687]: #683
+#355 := (not #683)
+#605 := (or #355 #612)
+#597 := [quant-inst]: #605
+#274 := [unit-resolution #597 #688]: #612
+#751 := [symm #274]: #750
+#748 := (= #50 #615)
+#579 := (= uf_4 #521)
+#14 := (:var 0 T1)
+#16 := (uf_1 #14 uf_4)
+#689 := (pattern #16)
+#71 := (= uf_4 #16)
+#690 := (forall (vars (?x4 T1)) (:pat #689) #71)
+#74 := (forall (vars (?x4 T1)) #71)
+#693 := (iff #74 #690)
+#691 := (iff #71 #71)
+#692 := [refl]: #691
+#694 := [quant-intro #692]: #693
+#180 := (~ #74 #74)
+#198 := (~ #71 #71)
+#199 := [refl]: #198
+#178 := [nnf-pos #199]: #180
+#17 := (= #16 uf_4)
+#18 := (forall (vars (?x4 T1)) #17)
+#75 := (iff #18 #74)
+#72 := (iff #17 #71)
+#73 := [rewrite]: #72
+#76 := [quant-intro #73]: #75
+#70 := [asserted]: #18
+#79 := [mp #70 #76]: #74
+#200 := [mp~ #79 #178]: #74
+#695 := [mp #200 #694]: #690
+#583 := (not #690)
+#584 := (or #583 #579)
+#581 := [quant-inst]: #584
+#275 := [unit-resolution #581 #695]: #579
+#746 := (= #49 #522)
+decl uf_7 :: (-> T2 int)
+#668 := (uf_7 #44)
+#596 := (+ 1::int #668)
+#606 := (uf_6 #596)
+#742 := (= #606 #522)
+#609 := (= #522 #606)
+#20 := (:var 0 T2)
+#22 := (uf_7 #20)
+#697 := (pattern #22)
+#21 := (uf_3 uf_5 #20)
+#696 := (pattern #21)
+#78 := (+ 1::int #22)
+#82 := (uf_6 #78)
+#85 := (= #21 #82)
+#698 := (forall (vars (?x5 T2)) (:pat #696 #697) #85)
+#88 := (forall (vars (?x5 T2)) #85)
+#701 := (iff #88 #698)
+#699 := (iff #85 #85)
+#700 := [refl]: #699
+#702 := [quant-intro #700]: #701
+#181 := (~ #88 #88)
+#201 := (~ #85 #85)
+#202 := [refl]: #201
+#182 := [nnf-pos #202]: #181
+#24 := (+ #22 1::int)
+#25 := (uf_6 #24)
+#26 := (= #21 #25)
+#27 := (forall (vars (?x5 T2)) #26)
+#89 := (iff #27 #88)
+#86 := (iff #26 #85)
+#83 := (= #25 #82)
+#80 := (= #24 #78)
+#81 := [rewrite]: #80
+#84 := [monotonicity #81]: #83
+#87 := [monotonicity #84]: #86
+#90 := [quant-intro #87]: #89
+#77 := [asserted]: #27
+#93 := [mp #77 #90]: #88
+#203 := [mp~ #93 #182]: #88
+#703 := [mp #203 #702]: #698
+#607 := (not #698)
+#600 := (or #607 #609)
+#601 := [quant-inst]: #600
+#278 := [unit-resolution #601 #703]: #609
+#743 := [symm #278]: #742
+#744 := (= #49 #606)
+#526 := (uf_7 #606)
+#325 := (uf_6 #526)
+#327 := (= #325 #606)
+#28 := (uf_6 #22)
+#92 := (= #20 #28)
+#704 := (forall (vars (?x6 T2)) (:pat #697) #92)
+#96 := (forall (vars (?x6 T2)) #92)
+#705 := (iff #96 #704)
+#707 := (iff #704 #704)
+#708 := [rewrite]: #707
+#706 := [rewrite]: #705
+#709 := [trans #706 #708]: #705
+#183 := (~ #96 #96)
+#204 := (~ #92 #92)
+#205 := [refl]: #204
+#184 := [nnf-pos #205]: #183
+#29 := (= #28 #20)
+#30 := (forall (vars (?x6 T2)) #29)
+#97 := (iff #30 #96)
+#94 := (iff #29 #92)
+#95 := [rewrite]: #94
+#98 := [quant-intro #95]: #97
+#91 := [asserted]: #30
+#101 := [mp #91 #98]: #96
+#206 := [mp~ #101 #184]: #96
+#710 := [mp #206 #709]: #704
+#368 := (not #704)
+#309 := (or #368 #327)
+#326 := (= #606 #325)
+#311 := (or #368 #326)
+#310 := (iff #311 #309)
+#301 := (iff #309 #309)
+#303 := [rewrite]: #301
+#316 := (iff #326 #327)
+#328 := [rewrite]: #316
+#313 := [monotonicity #328]: #310
+#304 := [trans #313 #303]: #310
+#312 := [quant-inst]: #311
+#307 := [mp #312 #304]: #309
+#279 := [unit-resolution #307 #710]: #327
+#740 := (= #49 #325)
+#738 := (= 2::int #526)
+#736 := (= #526 2::int)
+#568 := -1::int
+#533 := (* -1::int #668)
+#501 := (+ #526 #533)
+#481 := (<= #501 1::int)
+#527 := (= #501 1::int)
+#467 := (>= #668 -1::int)
+#592 := (>= #668 1::int)
+#378 := (= #668 1::int)
+#32 := (:var 0 int)
+#34 := (uf_6 #32)
+#711 := (pattern #34)
+#118 := (>= #32 0::int)
+#119 := (not #118)
+#35 := (uf_7 #34)
+#100 := (= #32 #35)
+#125 := (or #100 #119)
+#712 := (forall (vars (?x7 int)) (:pat #711) #125)
+#130 := (forall (vars (?x7 int)) #125)
+#715 := (iff #130 #712)
+#713 := (iff #125 #125)
+#714 := [refl]: #713
+#716 := [quant-intro #714]: #715
+#185 := (~ #130 #130)
+#207 := (~ #125 #125)
+#208 := [refl]: #207
+#186 := [nnf-pos #208]: #185
+#36 := (= #35 #32)
+#33 := (<= 0::int #32)
+#37 := (implies #33 #36)
+#38 := (forall (vars (?x7 int)) #37)
+#133 := (iff #38 #130)
+#107 := (not #33)
+#108 := (or #107 #100)
+#113 := (forall (vars (?x7 int)) #108)
+#131 := (iff #113 #130)
+#128 := (iff #108 #125)
+#122 := (or #119 #100)
+#126 := (iff #122 #125)
+#127 := [rewrite]: #126
+#123 := (iff #108 #122)
+#120 := (iff #107 #119)
+#116 := (iff #33 #118)
+#117 := [rewrite]: #116
+#121 := [monotonicity #117]: #120
+#124 := [monotonicity #121]: #123
+#129 := [trans #124 #127]: #128
+#132 := [quant-intro #129]: #131
+#114 := (iff #38 #113)
+#111 := (iff #37 #108)
+#104 := (implies #33 #100)
+#109 := (iff #104 #108)
+#110 := [rewrite]: #109
+#105 := (iff #37 #104)
+#102 := (iff #36 #100)
+#103 := [rewrite]: #102
+#106 := [monotonicity #103]: #105
+#112 := [trans #106 #110]: #111
+#115 := [quant-intro #112]: #114
+#134 := [trans #115 #132]: #133
+#99 := [asserted]: #38
+#135 := [mp #99 #134]: #130
+#209 := [mp~ #135 #186]: #130
+#717 := [mp #209 #716]: #712
+#314 := (not #712)
+#365 := (or #314 #378)
+#667 := (>= 1::int 0::int)
+#665 := (not #667)
+#654 := (= 1::int #668)
+#655 := (or #654 #665)
+#366 := (or #314 #655)
+#645 := (iff #366 #365)
+#642 := (iff #365 #365)
+#646 := [rewrite]: #642
+#363 := (iff #655 #378)
+#374 := (or #378 false)
+#649 := (iff #374 #378)
+#653 := [rewrite]: #649
+#648 := (iff #655 #374)
+#651 := (iff #665 false)
+#1 := true
+#342 := (not true)
+#677 := (iff #342 false)
+#678 := [rewrite]: #677
+#273 := (iff #665 #342)
+#379 := (iff #667 true)
+#380 := [rewrite]: #379
+#650 := [monotonicity #380]: #273
+#373 := [trans #650 #678]: #651
+#362 := (iff #654 #378)
+#377 := [rewrite]: #362
+#652 := [monotonicity #377 #373]: #648
+#364 := [trans #652 #653]: #363
+#359 := [monotonicity #364]: #645
+#643 := [trans #359 #646]: #645
+#644 := [quant-inst]: #366
+#647 := [mp #644 #643]: #365
+#280 := [unit-resolution #647 #717]: #378
+#276 := (not #378)
+#281 := (or #276 #592)
+#268 := [th-lemma]: #281
+#270 := [unit-resolution #268 #280]: #592
+#271 := (not #592)
+#269 := (or #271 #467)
+#272 := [th-lemma]: #269
+#724 := [unit-resolution #272 #270]: #467
+#502 := (not #467)
+#486 := (or #314 #502 #527)
+#525 := (>= #596 0::int)
+#471 := (not #525)
+#507 := (= #596 #526)
+#531 := (or #507 #471)
+#487 := (or #314 #531)
+#494 := (iff #487 #486)
+#503 := (or #502 #527)
+#489 := (or #314 #503)
+#492 := (iff #489 #486)
+#493 := [rewrite]: #492
+#490 := (iff #487 #489)
+#480 := (iff #531 #503)
+#512 := (or #527 #502)
+#524 := (iff #512 #503)
+#479 := [rewrite]: #524
+#513 := (iff #531 #512)
+#509 := (iff #471 #502)
+#498 := (iff #525 #467)
+#500 := [rewrite]: #498
+#511 := [monotonicity #500]: #509
+#532 := (iff #507 #527)
+#508 := [rewrite]: #532
+#523 := [monotonicity #508 #511]: #513
+#485 := [trans #523 #479]: #480
+#491 := [monotonicity #485]: #490
+#495 := [trans #491 #493]: #494
+#488 := [quant-inst]: #487
+#496 := [mp #488 #495]: #486
+#725 := [unit-resolution #496 #717 #724]: #527
+#726 := (not #527)
+#727 := (or #726 #481)
+#728 := [th-lemma]: #727
+#729 := [unit-resolution #728 #725]: #481
+#497 := (>= #501 1::int)
+#730 := (or #726 #497)
+#731 := [th-lemma]: #730
+#732 := [unit-resolution #731 #725]: #497
+#591 := (<= #668 1::int)
+#733 := (or #276 #591)
+#734 := [th-lemma]: #733
+#735 := [unit-resolution #734 #280]: #591
+#737 := [th-lemma #270 #735 #732 #729]: #736
+#739 := [symm #737]: #738
+#741 := [monotonicity #739]: #740
+#745 := [trans #741 #279]: #744
+#747 := [trans #745 #743]: #746
+#749 := [monotonicity #747 #275]: #748
+#753 := [trans #749 #751]: #752
+#786 := (= #44 #351)
+#354 := (uf_7 #43)
+#616 := (+ 1::int #354)
+#603 := (uf_6 #616)
+#782 := (= #603 #351)
+#594 := (= #351 #603)
+#608 := (or #607 #594)
+#604 := [quant-inst]: #608
+#754 := [unit-resolution #604 #703]: #594
+#783 := [symm #754]: #782
+#784 := (= #44 #603)
+#585 := (uf_7 #603)
+#384 := (uf_6 #585)
+#376 := (= #384 #603)
+#369 := (or #368 #376)
+#385 := (= #603 #384)
+#360 := (or #368 #385)
+#371 := (iff #360 #369)
+#372 := (iff #369 #369)
+#338 := [rewrite]: #372
+#386 := (iff #385 #376)
+#367 := [rewrite]: #386
+#361 := [monotonicity #367]: #371
+#340 := [trans #361 #338]: #371
+#370 := [quant-inst]: #360
+#341 := [mp #370 #340]: #369
+#755 := [unit-resolution #341 #710]: #376
+#780 := (= #44 #384)
+#778 := (= 1::int #585)
+#776 := (= #585 1::int)
+#569 := (* -1::int #585)
+#570 := (+ #354 #569)
+#552 := (<= #570 -1::int)
+#571 := (= #570 -1::int)
+#574 := (>= #354 -1::int)
+#587 := (>= #354 0::int)
+#331 := (= #354 0::int)
+#656 := (or #314 #331)
+#353 := (>= 0::int 0::int)
+#344 := (not #353)
+#358 := (= 0::int #354)
+#337 := (or #358 #344)
+#318 := (or #314 #337)
+#320 := (iff #318 #656)
+#658 := (iff #656 #656)
+#659 := [rewrite]: #658
+#330 := (iff #337 #331)
+#680 := (or #331 false)
+#334 := (iff #680 #331)
+#671 := [rewrite]: #334
+#670 := (iff #337 #680)
+#679 := (iff #344 false)
+#343 := (iff #344 #342)
+#462 := (iff #353 true)
+#669 := [rewrite]: #462
+#676 := [monotonicity #669]: #343
+#674 := [trans #676 #678]: #679
+#673 := (iff #358 #331)
+#675 := [rewrite]: #673
+#329 := [monotonicity #675 #674]: #670
+#672 := [trans #329 #671]: #330
+#321 := [monotonicity #672]: #320
+#660 := [trans #321 #659]: #320
+#319 := [quant-inst]: #318
+#661 := [mp #319 #660]: #656
+#756 := [unit-resolution #661 #717]: #331
+#757 := (not #331)
+#758 := (or #757 #587)
+#759 := [th-lemma]: #758
+#760 := [unit-resolution #759 #756]: #587
+#761 := (not #587)
+#762 := (or #761 #574)
+#763 := [th-lemma]: #762
+#764 := [unit-resolution #763 #760]: #574
+#577 := (not #574)
+#560 := (or #314 #571 #577)
+#580 := (>= #616 0::int)
+#582 := (not #580)
+#565 := (= #616 #585)
+#566 := (or #565 #582)
+#561 := (or #314 #566)
+#547 := (iff #561 #560)
+#556 := (or #571 #577)
+#563 := (or #314 #556)
+#550 := (iff #563 #560)
+#546 := [rewrite]: #550
+#558 := (iff #561 #563)
+#557 := (iff #566 #556)
+#567 := (iff #582 #577)
+#575 := (iff #580 #574)
+#576 := [rewrite]: #575
+#578 := [monotonicity #576]: #567
+#572 := (iff #565 #571)
+#573 := [rewrite]: #572
+#559 := [monotonicity #573 #578]: #557
+#564 := [monotonicity #559]: #558
+#548 := [trans #564 #546]: #547
+#562 := [quant-inst]: #561
+#551 := [mp #562 #548]: #560
+#765 := [unit-resolution #551 #717 #764]: #571
+#766 := (not #571)
+#767 := (or #766 #552)
+#768 := [th-lemma]: #767
+#769 := [unit-resolution #768 #765]: #552
+#553 := (>= #570 -1::int)
+#770 := (or #766 #553)
+#771 := [th-lemma]: #770
+#772 := [unit-resolution #771 #765]: #553
+#586 := (<= #354 0::int)
+#773 := (or #757 #586)
+#774 := [th-lemma]: #773
+#775 := [unit-resolution #774 #756]: #586
+#777 := [th-lemma #760 #775 #772 #769]: #776
+#779 := [symm #777]: #778
+#781 := [monotonicity #779]: #780
+#785 := [trans #781 #755]: #784
+#787 := [trans #785 #783]: #786
+#790 := [monotonicity #787 #753]: #789
+#792 := [symm #790]: #791
+#352 := (= #47 #267)
+#356 := (or #355 #352)
+#357 := [quant-inst]: #356
+#788 := [unit-resolution #357 #688]: #352
+#793 := [trans #788 #792]: #52
+#53 := (not #52)
+#177 := [asserted]: #53
+[unit-resolution #177 #793]: false
+unsat
+
+Jtmz+P173L9nRQkQk52h+Q 420
+#2 := false
+decl up_1 :: (-> T1 bool)
+#4 := (:var 0 T1)
+#5 := (up_1 #4)
+#6 := (forall (vars (?x1 T1)) #5)
+#7 := (not #6)
+#8 := (or #6 #7)
+#9 := (not #8)
+#33 := (iff #9 false)
+#1 := true
+#28 := (not true)
+#31 := (iff #28 false)
+#32 := [rewrite]: #31
+#29 := (iff #9 #28)
+#26 := (iff #8 true)
+#27 := [rewrite]: #26
+#30 := [monotonicity #27]: #29
+#34 := [trans #30 #32]: #33
+#25 := [asserted]: #9
+[mp #25 #34]: false
+unsat
+
+YG20f6Uf93koN6rVg/alpA 9362
+#2 := false
+decl uf_1 :: (-> int T1)
+#37 := 6::int
+#38 := (uf_1 6::int)
+decl uf_3 :: (-> T1 T1)
+decl uf_2 :: (-> T1 int)
+#30 := 4::int
+#31 := (uf_1 4::int)
+#32 := (uf_3 #31)
+#33 := (uf_2 #32)
+#34 := (* 4::int #33)
+#35 := (uf_1 #34)
+#36 := (uf_3 #35)
+#39 := (= #36 #38)
+#476 := (uf_3 #38)
+#403 := (= #476 #38)
+#531 := (= #38 #476)
+#620 := (uf_2 #38)
+#142 := -10::int
+#513 := (+ -10::int #620)
+#472 := (uf_1 #513)
+#503 := (uf_3 #472)
+#505 := (= #476 #503)
+#22 := 10::int
+#507 := (>= #620 10::int)
+#514 := (ite #507 #505 #531)
+#4 := (:var 0 T1)
+#21 := (uf_3 #4)
+#707 := (pattern #21)
+#5 := (uf_2 #4)
+#686 := (pattern #5)
+#209 := (= #4 #21)
+#143 := (+ -10::int #5)
+#146 := (uf_1 #143)
+#149 := (uf_3 #146)
+#208 := (= #21 #149)
+#163 := (>= #5 10::int)
+#190 := (ite #163 #208 #209)
+#708 := (forall (vars (?x4 T1)) (:pat #686 #707) #190)
+#193 := (forall (vars (?x4 T1)) #190)
+#711 := (iff #193 #708)
+#709 := (iff #190 #190)
+#710 := [refl]: #709
+#712 := [quant-intro #710]: #711
+#168 := (ite #163 #149 #4)
+#173 := (= #21 #168)
+#176 := (forall (vars (?x4 T1)) #173)
+#210 := (iff #176 #193)
+#191 := (iff #173 #190)
+#192 := [rewrite]: #191
+#211 := [quant-intro #192]: #210
+#188 := (~ #176 #176)
+#205 := (~ #173 #173)
+#206 := [refl]: #205
+#189 := [nnf-pos #206]: #188
+#24 := (- #5 10::int)
+#25 := (uf_1 #24)
+#26 := (uf_3 #25)
+#23 := (< #5 10::int)
+#27 := (ite #23 #4 #26)
+#28 := (= #21 #27)
+#29 := (forall (vars (?x4 T1)) #28)
+#179 := (iff #29 #176)
+#152 := (ite #23 #4 #149)
+#155 := (= #21 #152)
+#158 := (forall (vars (?x4 T1)) #155)
+#177 := (iff #158 #176)
+#174 := (iff #155 #173)
+#171 := (= #152 #168)
+#161 := (not #163)
+#165 := (ite #161 #4 #149)
+#169 := (= #165 #168)
+#170 := [rewrite]: #169
+#166 := (= #152 #165)
+#162 := (iff #23 #161)
+#164 := [rewrite]: #162
+#167 := [monotonicity #164]: #166
+#172 := [trans #167 #170]: #171
+#175 := [monotonicity #172]: #174
+#178 := [quant-intro #175]: #177
+#159 := (iff #29 #158)
+#156 := (iff #28 #155)
+#153 := (= #27 #152)
+#150 := (= #26 #149)
+#147 := (= #25 #146)
+#144 := (= #24 #143)
+#145 := [rewrite]: #144
+#148 := [monotonicity #145]: #147
+#151 := [monotonicity #148]: #150
+#154 := [monotonicity #151]: #153
+#157 := [monotonicity #154]: #156
+#160 := [quant-intro #157]: #159
+#180 := [trans #160 #178]: #179
+#141 := [asserted]: #29
+#181 := [mp #141 #180]: #176
+#207 := [mp~ #181 #189]: #176
+#212 := [mp #207 #211]: #193
+#713 := [mp #212 #712]: #708
+#336 := (not #708)
+#518 := (or #336 #514)
+#528 := [quant-inst]: #518
+#477 := [unit-resolution #528 #713]: #514
+#529 := (not #507)
+#498 := (<= #620 6::int)
+#610 := (= #620 6::int)
+#10 := (:var 0 int)
+#12 := (uf_1 #10)
+#694 := (pattern #12)
+#9 := 0::int
+#82 := (>= #10 0::int)
+#83 := (not #82)
+#13 := (uf_2 #12)
+#64 := (= #10 #13)
+#89 := (or #64 #83)
+#695 := (forall (vars (?x2 int)) (:pat #694) #89)
+#94 := (forall (vars (?x2 int)) #89)
+#698 := (iff #94 #695)
+#696 := (iff #89 #89)
+#697 := [refl]: #696
+#699 := [quant-intro #697]: #698
+#185 := (~ #94 #94)
+#199 := (~ #89 #89)
+#200 := [refl]: #199
+#183 := [nnf-pos #200]: #185
+#14 := (= #13 #10)
+#11 := (<= 0::int #10)
+#15 := (implies #11 #14)
+#16 := (forall (vars (?x2 int)) #15)
+#97 := (iff #16 #94)
+#71 := (not #11)
+#72 := (or #71 #64)
+#77 := (forall (vars (?x2 int)) #72)
+#95 := (iff #77 #94)
+#92 := (iff #72 #89)
+#86 := (or #83 #64)
+#90 := (iff #86 #89)
+#91 := [rewrite]: #90
+#87 := (iff #72 #86)
+#84 := (iff #71 #83)
+#80 := (iff #11 #82)
+#81 := [rewrite]: #80
+#85 := [monotonicity #81]: #84
+#88 := [monotonicity #85]: #87
+#93 := [trans #88 #91]: #92
+#96 := [quant-intro #93]: #95
+#78 := (iff #16 #77)
+#75 := (iff #15 #72)
+#68 := (implies #11 #64)
+#73 := (iff #68 #72)
+#74 := [rewrite]: #73
+#69 := (iff #15 #68)
+#66 := (iff #14 #64)
+#67 := [rewrite]: #66
+#70 := [monotonicity #67]: #69
+#76 := [trans #70 #74]: #75
+#79 := [quant-intro #76]: #78
+#98 := [trans #79 #96]: #97
+#63 := [asserted]: #16
+#99 := [mp #63 #98]: #94
+#201 := [mp~ #99 #183]: #94
+#700 := [mp #201 #699]: #695
+#673 := (not #695)
+#591 := (or #673 #610)
+#526 := (>= 6::int 0::int)
+#527 := (not #526)
+#617 := (= 6::int #620)
+#621 := (or #617 #527)
+#592 := (or #673 #621)
+#595 := (iff #592 #591)
+#597 := (iff #591 #591)
+#593 := [rewrite]: #597
+#600 := (iff #621 #610)
+#614 := (or #610 false)
+#605 := (iff #614 #610)
+#606 := [rewrite]: #605
+#603 := (iff #621 #614)
+#613 := (iff #527 false)
+#1 := true
+#663 := (not true)
+#666 := (iff #663 false)
+#667 := [rewrite]: #666
+#611 := (iff #527 #663)
+#599 := (iff #526 true)
+#601 := [rewrite]: #599
+#612 := [monotonicity #601]: #611
+#609 := [trans #612 #667]: #613
+#608 := (iff #617 #610)
+#602 := [rewrite]: #608
+#604 := [monotonicity #602 #609]: #603
+#607 := [trans #604 #606]: #600
+#596 := [monotonicity #607]: #595
+#598 := [trans #596 #593]: #595
+#594 := [quant-inst]: #592
+#584 := [mp #594 #598]: #591
+#478 := [unit-resolution #584 #700]: #610
+#453 := (not #610)
+#454 := (or #453 #498)
+#455 := [th-lemma]: #454
+#456 := [unit-resolution #455 #478]: #498
+#458 := (not #498)
+#459 := (or #458 #529)
+#460 := [th-lemma]: #459
+#302 := [unit-resolution #460 #456]: #529
+#508 := (not #514)
+#490 := (or #508 #507 #531)
+#491 := [def-axiom]: #490
+#461 := [unit-resolution #491 #302 #477]: #531
+#404 := [symm #461]: #403
+#405 := (= #36 #476)
+#649 := (uf_2 #35)
+#582 := (+ -10::int #649)
+#553 := (uf_1 #582)
+#556 := (uf_3 #553)
+#401 := (= #556 #476)
+#417 := (= #553 #38)
+#415 := (= #582 6::int)
+#335 := (uf_2 #31)
+#647 := -1::int
+#502 := (* -1::int #335)
+#463 := (+ #33 #502)
+#464 := (<= #463 0::int)
+#486 := (= #33 #335)
+#445 := (= #32 #31)
+#574 := (= #31 #32)
+#575 := (+ -10::int #335)
+#576 := (uf_1 #575)
+#577 := (uf_3 #576)
+#578 := (= #32 #577)
+#579 := (>= #335 10::int)
+#580 := (ite #579 #578 #574)
+#572 := (or #336 #580)
+#583 := [quant-inst]: #572
+#457 := [unit-resolution #583 #713]: #580
+#562 := (not #579)
+#554 := (<= #335 4::int)
+#324 := (= #335 4::int)
+#659 := (or #673 #324)
+#678 := (>= 4::int 0::int)
+#680 := (not #678)
+#677 := (= 4::int #335)
+#319 := (or #677 #680)
+#660 := (or #673 #319)
+#382 := (iff #660 #659)
+#384 := (iff #659 #659)
+#385 := [rewrite]: #384
+#672 := (iff #319 #324)
+#305 := (or #324 false)
+#310 := (iff #305 #324)
+#311 := [rewrite]: #310
+#669 := (iff #319 #305)
+#662 := (iff #680 false)
+#664 := (iff #680 #663)
+#325 := (iff #678 true)
+#326 := [rewrite]: #325
+#665 := [monotonicity #326]: #664
+#668 := [trans #665 #667]: #662
+#661 := (iff #677 #324)
+#323 := [rewrite]: #661
+#671 := [monotonicity #323 #668]: #669
+#670 := [trans #671 #311]: #672
+#383 := [monotonicity #670]: #382
+#277 := [trans #383 #385]: #382
+#367 := [quant-inst]: #660
+#655 := [mp #367 #277]: #659
+#462 := [unit-resolution #655 #700]: #324
+#441 := (not #324)
+#444 := (or #441 #554)
+#448 := [th-lemma]: #444
+#450 := [unit-resolution #448 #462]: #554
+#451 := (not #554)
+#449 := (or #451 #562)
+#452 := [th-lemma]: #449
+#440 := [unit-resolution #452 #450]: #562
+#561 := (not #580)
+#566 := (or #561 #579 #574)
+#567 := [def-axiom]: #566
+#443 := [unit-resolution #567 #440 #457]: #574
+#446 := [symm #443]: #445
+#442 := [monotonicity #446]: #486
+#447 := (not #486)
+#437 := (or #447 #464)
+#427 := [th-lemma]: #437
+#429 := [unit-resolution #427 #442]: #464
+#471 := (>= #463 0::int)
+#430 := (or #447 #471)
+#433 := [th-lemma]: #430
+#434 := [unit-resolution #433 #442]: #471
+#560 := (>= #335 4::int)
+#438 := (or #441 #560)
+#431 := [th-lemma]: #438
+#439 := [unit-resolution #431 #462]: #560
+#651 := (* -1::int #649)
+#648 := (+ #34 #651)
+#625 := (<= #648 0::int)
+#652 := (= #648 0::int)
+#643 := (>= #33 0::int)
+#435 := (not #471)
+#432 := (not #560)
+#436 := (or #643 #432 #435)
+#422 := [th-lemma]: #436
+#424 := [unit-resolution #422 #439 #434]: #643
+#644 := (not #643)
+#489 := (or #644 #652)
+#628 := (or #673 #644 #652)
+#370 := (>= #34 0::int)
+#371 := (not #370)
+#650 := (= #34 #649)
+#364 := (or #650 #371)
+#629 := (or #673 #364)
+#469 := (iff #629 #628)
+#636 := (or #673 #489)
+#466 := (iff #636 #628)
+#468 := [rewrite]: #466
+#630 := (iff #629 #636)
+#633 := (iff #364 #489)
+#646 := (or #652 #644)
+#631 := (iff #646 #489)
+#632 := [rewrite]: #631
+#487 := (iff #364 #646)
+#645 := (iff #371 #644)
+#638 := (iff #370 #643)
+#639 := [rewrite]: #638
+#640 := [monotonicity #639]: #645
+#641 := (iff #650 #652)
+#642 := [rewrite]: #641
+#488 := [monotonicity #642 #640]: #487
+#634 := [trans #488 #632]: #633
+#637 := [monotonicity #634]: #630
+#622 := [trans #637 #468]: #469
+#635 := [quant-inst]: #629
+#623 := [mp #635 #622]: #628
+#425 := [unit-resolution #623 #700]: #489
+#423 := [unit-resolution #425 #424]: #652
+#426 := (not #652)
+#408 := (or #426 #625)
+#410 := [th-lemma]: #408
+#411 := [unit-resolution #410 #423]: #625
+#626 := (>= #648 0::int)
+#412 := (or #426 #626)
+#413 := [th-lemma]: #412
+#414 := [unit-resolution #413 #423]: #626
+#416 := [th-lemma #414 #411 #439 #450 #434 #429]: #415
+#418 := [monotonicity #416]: #417
+#402 := [monotonicity #418]: #401
+#557 := (= #36 #556)
+#581 := (= #35 #36)
+#558 := (>= #649 10::int)
+#559 := (ite #558 #557 #581)
+#533 := (or #336 #559)
+#534 := [quant-inst]: #533
+#419 := [unit-resolution #534 #713]: #559
+#420 := (not #625)
+#409 := (or #558 #420 #432 #435)
+#421 := [th-lemma]: #409
+#398 := [unit-resolution #421 #411 #439 #434]: #558
+#428 := (not #558)
+#535 := (not #559)
+#539 := (or #535 #428 #557)
+#540 := [def-axiom]: #539
+#400 := [unit-resolution #540 #398 #419]: #557
+#406 := [trans #400 #402]: #405
+#399 := [trans #406 #404]: #39
+#40 := (not #39)
+#182 := [asserted]: #40
+[unit-resolution #182 #399]: false
+unsat
+
+/fwo5o8vhLVHyS4oYEs4QA 10833
+#2 := false
+#22 := 0::int
+#8 := 2::int
+decl uf_6 :: (-> T3 T4 int)
+decl uf_9 :: T4
+#50 := uf_9
+decl uf_8 :: T3
+#49 := uf_8
+#51 := (uf_6 uf_8 uf_9)
+#624 := (mod #51 2::int)
+#172 := -1::int
+#640 := (* -1::int #624)
+decl uf_7 :: (-> T2 T4 T4)
+decl uf_5 :: T2
+#13 := uf_5
+#54 := (uf_7 uf_5 uf_9)
+#55 := (uf_6 uf_8 #54)
+#56 := (mod #55 2::int)
+#620 := (+ #56 #640)
+#608 := (>= #620 0::int)
+#566 := (= #620 0::int)
+#35 := (:var 0 T4)
+#38 := (uf_7 uf_5 #35)
+#34 := (:var 1 T3)
+#39 := (uf_6 #34 #38)
+#811 := (pattern #39)
+#40 := (mod #39 2::int)
+#173 := (* -1::int #40)
+#36 := (uf_6 #34 #35)
+#37 := (mod #36 2::int)
+#174 := (+ #37 #173)
+#175 := (= #174 0::int)
+#812 := (forall (vars (?x6 T3) (?x7 T4)) (:pat #811) #175)
+#178 := (forall (vars (?x6 T3) (?x7 T4)) #175)
+#815 := (iff #178 #812)
+#813 := (iff #175 #175)
+#814 := [refl]: #813
+#816 := [quant-intro #814]: #815
+#277 := (~ #178 #178)
+#302 := (~ #175 #175)
+#303 := [refl]: #302
+#278 := [nnf-pos #303]: #277
+#41 := (= #37 #40)
+#42 := (forall (vars (?x6 T3) (?x7 T4)) #41)
+#179 := (iff #42 #178)
+#176 := (iff #41 #175)
+#177 := [rewrite]: #176
+#180 := [quant-intro #177]: #179
+#169 := [asserted]: #42
+#181 := [mp #169 #180]: #178
+#304 := [mp~ #181 #278]: #178
+#817 := [mp #304 #816]: #812
+#622 := (not #812)
+#628 := (or #622 #566)
+#756 := (* -1::int #56)
+#625 := (+ #624 #756)
+#632 := (= #625 0::int)
+#596 := (or #622 #632)
+#562 := (iff #596 #628)
+#595 := (iff #628 #628)
+#597 := [rewrite]: #595
+#626 := (iff #632 #566)
+#633 := (+ #756 #624)
+#638 := (= #633 0::int)
+#621 := (iff #638 #566)
+#602 := [rewrite]: #621
+#639 := (iff #632 #638)
+#634 := (= #625 #633)
+#637 := [rewrite]: #634
+#635 := [monotonicity #637]: #639
+#627 := [trans #635 #602]: #626
+#593 := [monotonicity #627]: #562
+#604 := [trans #593 #597]: #562
+#603 := [quant-inst]: #596
+#606 := [mp #603 #604]: #628
+#528 := [unit-resolution #606 #817]: #566
+#521 := (not #566)
+#464 := (or #521 #608)
+#456 := [th-lemma]: #464
+#465 := [unit-resolution #456 #528]: #608
+decl uf_10 :: int
+#52 := uf_10
+#57 := (mod uf_10 2::int)
+#243 := (* -1::int #57)
+#244 := (+ #56 #243)
+#447 := (>= #244 0::int)
+#387 := (not #447)
+#245 := (= #244 0::int)
+#248 := (not #245)
+#218 := (* -1::int #55)
+#219 := (+ uf_10 #218)
+#222 := (div #219 2::int)
+#251 := (* -1::int #222)
+decl uf_2 :: T2
+#4 := uf_2
+#59 := (uf_7 uf_2 uf_9)
+#60 := (uf_6 uf_8 #59)
+#252 := (+ #60 #251)
+#253 := (= #252 0::int)
+#448 := (<= #252 0::int)
+#605 := (+ uf_10 #55)
+#613 := (mod #605 2::int)
+#672 := (>= #613 2::int)
+#662 := (not #672)
+#1 := true
+#81 := [true-axiom]: true
+#520 := (or false #662)
+#523 := [th-lemma]: #520
+#524 := [unit-resolution #523 #81]: #662
+#701 := (* -1::int #613)
+#204 := -2::int
+#691 := (* -2::int #222)
+#702 := (+ #691 #701)
+#703 := (+ #218 #702)
+#699 := (+ uf_10 #703)
+#694 := (<= #699 0::int)
+#692 := (= #699 0::int)
+#545 := (or false #692)
+#546 := [th-lemma]: #545
+#548 := [unit-resolution #546 #81]: #692
+#549 := (not #692)
+#497 := (or #549 #694)
+#482 := [th-lemma]: #497
+#483 := [unit-resolution #482 #548]: #694
+#536 := (not #448)
+#395 := [hypothesis]: #536
+#555 := (* -1::int uf_10)
+#573 := (+ #51 #555)
+#543 := (<= #573 0::int)
+#53 := (= #51 uf_10)
+#256 := (not #253)
+#259 := (or #248 #256)
+#502 := 1::int
+#731 := (div uf_10 2::int)
+#515 := (* -1::int #731)
+#513 := (+ #640 #515)
+#618 := (div #51 2::int)
+#514 := (* -1::int #618)
+#516 := (+ #514 #513)
+#498 := (+ #243 #516)
+#500 := (+ #56 #498)
+#501 := (+ uf_10 #500)
+#503 := (>= #501 1::int)
+#486 := (not #503)
+#361 := (<= #244 0::int)
+#453 := (not #259)
+#517 := [hypothesis]: #453
+#440 := (or #259 #245)
+#451 := [def-axiom]: #440
+#519 := [unit-resolution #451 #517]: #245
+#478 := (or #248 #361)
+#470 := [th-lemma]: #478
+#479 := [unit-resolution #470 #519]: #361
+#449 := (>= #252 0::int)
+#452 := (or #259 #253)
+#380 := [def-axiom]: #452
+#480 := [unit-resolution #380 #517]: #253
+#471 := (or #256 #449)
+#481 := [th-lemma]: #471
+#462 := [unit-resolution #481 #480]: #449
+#487 := (not #361)
+#485 := (not #449)
+#476 := (or #486 #485 #487)
+#607 := (<= #620 0::int)
+#529 := (or #521 #607)
+#522 := [th-lemma]: #529
+#525 := [unit-resolution #522 #528]: #607
+#723 := (* -2::int #731)
+#724 := (+ #243 #723)
+#718 := (+ uf_10 #724)
+#720 := (<= #718 0::int)
+#722 := (= #718 0::int)
+#526 := (or false #722)
+#512 := [th-lemma]: #526
+#504 := [unit-resolution #512 #81]: #722
+#505 := (not #722)
+#506 := (or #505 #720)
+#507 := [th-lemma]: #506
+#508 := [unit-resolution #507 #504]: #720
+#509 := [hypothesis]: #361
+#583 := (* -2::int #618)
+#584 := (+ #583 #640)
+#585 := (+ #51 #584)
+#587 := (<= #585 0::int)
+#582 := (= #585 0::int)
+#510 := (or false #582)
+#499 := [th-lemma]: #510
+#511 := [unit-resolution #499 #81]: #582
+#488 := (not #582)
+#490 := (or #488 #587)
+#491 := [th-lemma]: #490
+#492 := [unit-resolution #491 #511]: #587
+#493 := [hypothesis]: #503
+#649 := (* -2::int #60)
+#644 := (+ #218 #649)
+#650 := (+ #51 #644)
+#636 := (>= #650 0::int)
+#623 := (= #650 0::int)
+#43 := (uf_7 uf_2 #35)
+#44 := (uf_6 #34 #43)
+#818 := (pattern #44)
+#205 := (* -2::int #44)
+#203 := (* -1::int #39)
+#206 := (+ #203 #205)
+#207 := (+ #36 #206)
+#208 := (= #207 0::int)
+#819 := (forall (vars (?x8 T3) (?x9 T4)) (:pat #811 #818) #208)
+#211 := (forall (vars (?x8 T3) (?x9 T4)) #208)
+#822 := (iff #211 #819)
+#820 := (iff #208 #208)
+#821 := [refl]: #820
+#823 := [quant-intro #821]: #822
+#279 := (~ #211 #211)
+#305 := (~ #208 #208)
+#306 := [refl]: #305
+#280 := [nnf-pos #306]: #279
+#45 := (* #44 2::int)
+#46 := (+ #45 #39)
+#47 := (= #46 #36)
+#48 := (forall (vars (?x8 T3) (?x9 T4)) #47)
+#214 := (iff #48 #211)
+#171 := (* 2::int #44)
+#187 := (+ #39 #171)
+#195 := (= #36 #187)
+#200 := (forall (vars (?x8 T3) (?x9 T4)) #195)
+#212 := (iff #200 #211)
+#209 := (iff #195 #208)
+#210 := [rewrite]: #209
+#213 := [quant-intro #210]: #212
+#201 := (iff #48 #200)
+#198 := (iff #47 #195)
+#192 := (= #187 #36)
+#196 := (iff #192 #195)
+#197 := [rewrite]: #196
+#193 := (iff #47 #192)
+#190 := (= #46 #187)
+#184 := (+ #171 #39)
+#188 := (= #184 #187)
+#189 := [rewrite]: #188
+#185 := (= #46 #184)
+#182 := (= #45 #171)
+#183 := [rewrite]: #182
+#186 := [monotonicity #183]: #185
+#191 := [trans #186 #189]: #190
+#194 := [monotonicity #191]: #193
+#199 := [trans #194 #197]: #198
+#202 := [quant-intro #199]: #201
+#215 := [trans #202 #213]: #214
+#170 := [asserted]: #48
+#216 := [mp #170 #215]: #211
+#307 := [mp~ #216 #280]: #211
+#824 := [mp #307 #823]: #819
+#518 := (not #819)
+#629 := (or #518 #623)
+#630 := [quant-inst]: #629
+#531 := [unit-resolution #630 #824]: #623
+#534 := (not #623)
+#494 := (or #534 #636)
+#495 := [th-lemma]: #494
+#496 := [unit-resolution #495 #531]: #636
+#489 := [hypothesis]: #449
+#484 := [th-lemma #483 #489 #496 #493 #492 #509 #508 #525 #524]: false
+#477 := [lemma #484]: #476
+#463 := [unit-resolution #477 #462 #479]: #486
+#727 := (>= #718 0::int)
+#466 := (or #505 #727)
+#457 := [th-lemma]: #466
+#467 := [unit-resolution #457 #504]: #727
+#434 := (or #248 #447)
+#436 := [th-lemma]: #434
+#437 := [unit-resolution #436 #519]: #447
+#544 := (>= #573 0::int)
+#445 := (not #544)
+#428 := (or #256 #448)
+#441 := [th-lemma]: #428
+#442 := [unit-resolution #441 #480]: #448
+#532 := (or #543 #536)
+#695 := (>= #699 0::int)
+#550 := (or #549 #695)
+#393 := [th-lemma]: #550
+#551 := [unit-resolution #393 #548]: #695
+#547 := (not #543)
+#552 := [hypothesis]: #547
+#631 := (<= #650 0::int)
+#538 := (or #534 #631)
+#540 := [th-lemma]: #538
+#541 := [unit-resolution #540 #531]: #631
+#539 := [hypothesis]: #448
+#666 := (>= #613 0::int)
+#542 := (or false #666)
+#530 := [th-lemma]: #542
+#533 := [unit-resolution #530 #81]: #666
+#535 := [th-lemma #533 #539 #541 #552 #551]: false
+#537 := [lemma #535]: #532
+#443 := [unit-resolution #537 #442]: #543
+#429 := (or #547 #445)
+#764 := (not #53)
+#771 := (or #764 #259)
+#262 := (iff #53 #259)
+#61 := (- uf_10 #55)
+#62 := (div #61 2::int)
+#63 := (= #60 #62)
+#64 := (not #63)
+#58 := (= #56 #57)
+#65 := (implies #58 #64)
+#66 := (iff #53 #65)
+#265 := (iff #66 #262)
+#225 := (= #60 #222)
+#228 := (not #225)
+#234 := (not #58)
+#235 := (or #234 #228)
+#240 := (iff #53 #235)
+#263 := (iff #240 #262)
+#260 := (iff #235 #259)
+#257 := (iff #228 #256)
+#254 := (iff #225 #253)
+#255 := [rewrite]: #254
+#258 := [monotonicity #255]: #257
+#249 := (iff #234 #248)
+#246 := (iff #58 #245)
+#247 := [rewrite]: #246
+#250 := [monotonicity #247]: #249
+#261 := [monotonicity #250 #258]: #260
+#264 := [monotonicity #261]: #263
+#241 := (iff #66 #240)
+#238 := (iff #65 #235)
+#231 := (implies #58 #228)
+#236 := (iff #231 #235)
+#237 := [rewrite]: #236
+#232 := (iff #65 #231)
+#229 := (iff #64 #228)
+#226 := (iff #63 #225)
+#223 := (= #62 #222)
+#220 := (= #61 #219)
+#221 := [rewrite]: #220
+#224 := [monotonicity #221]: #223
+#227 := [monotonicity #224]: #226
+#230 := [monotonicity #227]: #229
+#233 := [monotonicity #230]: #232
+#239 := [trans #233 #237]: #238
+#242 := [monotonicity #239]: #241
+#266 := [trans #242 #264]: #265
+#217 := [asserted]: #66
+#267 := [mp #217 #266]: #262
+#433 := (not #262)
+#438 := (or #764 #259 #433)
+#439 := [def-axiom]: #438
+#772 := [unit-resolution #439 #267]: #771
+#444 := [unit-resolution #772 #517]: #764
+#435 := (or #53 #547 #445)
+#446 := [th-lemma]: #435
+#431 := [unit-resolution #446 #444]: #429
+#432 := [unit-resolution #431 #443]: #445
+#588 := (>= #585 0::int)
+#411 := (or #488 #588)
+#413 := [th-lemma]: #411
+#418 := [unit-resolution #413 #511]: #588
+#419 := [th-lemma #418 #432 #437 #467 #465 #463]: false
+#420 := [lemma #419]: #259
+#427 := (or #53 #453)
+#768 := (or #53 #453 #433)
+#770 := [def-axiom]: #768
+#557 := [unit-resolution #770 #267]: #427
+#406 := [unit-resolution #557 #420]: #53
+#377 := (or #764 #543)
+#381 := [th-lemma]: #377
+#382 := [unit-resolution #381 #406]: #543
+#385 := [th-lemma #496 #382 #395 #483 #524]: false
+#386 := [lemma #385]: #448
+#390 := (or #253 #536)
+#408 := [hypothesis]: #485
+#409 := (or #764 #544)
+#397 := [th-lemma]: #409
+#399 := [unit-resolution #397 #406]: #544
+#400 := [th-lemma #399 #408 #533 #551 #541]: false
+#403 := [lemma #400]: #449
+#392 := (or #253 #536 #485)
+#394 := [th-lemma]: #392
+#657 := [unit-resolution #394 #403]: #390
+#658 := [unit-resolution #657 #386]: #253
+#450 := (or #453 #248 #256)
+#454 := [def-axiom]: #450
+#762 := [unit-resolution #454 #420]: #259
+#664 := [unit-resolution #762 #658]: #248
+#372 := (or #245 #387)
+#560 := (+ #57 #640)
+#610 := (>= #560 0::int)
+#742 := (= #57 #624)
+#424 := (= #624 #57)
+#405 := [monotonicity #406]: #424
+#407 := [symm #405]: #742
+#705 := (not #742)
+#706 := (or #705 #610)
+#568 := [th-lemma]: #706
+#569 := [unit-resolution #568 #407]: #610
+#398 := [hypothesis]: #487
+#404 := [th-lemma #525 #398 #569]: false
+#378 := [lemma #404]: #361
+#379 := (or #245 #487 #387)
+#388 := [th-lemma]: #379
+#369 := [unit-resolution #388 #378]: #372
+#370 := [unit-resolution #369 #664]: #387
+#708 := (<= #560 0::int)
+#373 := (or #705 #708)
+#374 := [th-lemma]: #373
+#375 := [unit-resolution #374 #407]: #708
+[th-lemma #375 #370 #465]: false
+unsat
+
+s8LL71+1HTN+eIuEYeKhUw 1251
+#2 := false
+decl up_35 :: (-> int bool)
+#112 := 1::int
+#113 := (up_35 1::int)
+#114 := (not #113)
+#297 := [asserted]: #114
+#103 := (:var 0 int)
+#104 := (up_35 #103)
+#910 := (pattern #104)
+#911 := (forall (vars (?x12 int)) (:pat #910) #104)
+#294 := (forall (vars (?x12 int)) #104)
+#914 := (iff #294 #911)
+#912 := (iff #104 #104)
+#913 := [refl]: #912
+#915 := [quant-intro #913]: #914
+#320 := (~ #294 #294)
+#361 := (~ #104 #104)
+#362 := [refl]: #361
+#321 := [nnf-pos #362]: #320
+decl up_32 :: (-> T13 bool)
+decl uf_36 :: (-> int T13 T13)
+decl uf_37 :: T13
+#105 := uf_37
+#106 := (uf_36 #103 uf_37)
+#107 := (up_32 #106)
+#108 := (not #107)
+#109 := (or #107 #108)
+#110 := (and #104 #109)
+#111 := (forall (vars (?x12 int)) #110)
+#295 := (iff #111 #294)
+#292 := (iff #110 #104)
+#1 := true
+#287 := (and #104 true)
+#290 := (iff #287 #104)
+#291 := [rewrite]: #290
+#288 := (iff #110 #287)
+#284 := (iff #109 true)
+#286 := [rewrite]: #284
+#289 := [monotonicity #286]: #288
+#293 := [trans #289 #291]: #292
+#296 := [quant-intro #293]: #295
+#283 := [asserted]: #111
+#299 := [mp #283 #296]: #294
+#363 := [mp~ #299 #321]: #294
+#916 := [mp #363 #915]: #911
+#418 := (not #911)
+#503 := (or #418 #113)
+#504 := [quant-inst]: #503
+[unit-resolution #504 #916 #297]: false
+unsat
+
+MZYsU5krlrOK4MkB71Ikeg 12985
+#2 := false
+decl uf_17 :: (-> T8 T3)
+decl uf_18 :: (-> T1 T8)
+decl uf_19 :: T1
+#104 := uf_19
+#105 := (uf_18 uf_19)
+#106 := (uf_17 #105)
+decl uf_15 :: (-> T7 T3)
+decl uf_16 :: (-> int T7)
+#101 := 3::int
+#102 := (uf_16 3::int)
+#103 := (uf_15 #102)
+#107 := (= #103 #106)
+decl uf_13 :: (-> T2 T3)
+decl uf_2 :: (-> T1 T2 T2)
+decl uf_7 :: T2
+#29 := uf_7
+#513 := (uf_2 uf_19 uf_7)
+#644 := (uf_13 #513)
+#564 := (= #644 #106)
+#858 := (= #106 #644)
+#79 := (:var 0 T1)
+#82 := (uf_2 #79 uf_7)
+#930 := (pattern #82)
+#80 := (uf_18 #79)
+#929 := (pattern #80)
+#83 := (uf_13 #82)
+#81 := (uf_17 #80)
+#84 := (= #81 #83)
+#931 := (forall (vars (?x16 T1)) (:pat #929 #930) #84)
+#85 := (forall (vars (?x16 T1)) #84)
+#934 := (iff #85 #931)
+#932 := (iff #84 #84)
+#933 := [refl]: #932
+#935 := [quant-intro #933]: #934
+#347 := (~ #85 #85)
+#384 := (~ #84 #84)
+#385 := [refl]: #384
+#348 := [nnf-pos #385]: #347
+#238 := [asserted]: #85
+#386 := [mp~ #238 #348]: #85
+#936 := [mp #386 #935]: #931
+#861 := (not #931)
+#856 := (or #861 #858)
+#862 := [quant-inst]: #856
+#579 := [unit-resolution #862 #936]: #858
+#565 := [symm #579]: #564
+#553 := (= #103 #644)
+decl uf_1 :: (-> T2 T3)
+#834 := (uf_1 #513)
+#831 := (= #834 #644)
+#835 := (= #644 #834)
+#5 := (:var 0 T2)
+#66 := (uf_13 #5)
+#906 := (pattern #66)
+#8 := (uf_1 #5)
+#905 := (pattern #8)
+#222 := (= #8 #66)
+#907 := (forall (vars (?x13 T2)) (:pat #905 #906) #222)
+#226 := (forall (vars (?x13 T2)) #222)
+#910 := (iff #226 #907)
+#908 := (iff #222 #222)
+#909 := [refl]: #908
+#911 := [quant-intro #909]: #910
+#341 := (~ #226 #226)
+#375 := (~ #222 #222)
+#376 := [refl]: #375
+#342 := [nnf-pos #376]: #341
+#67 := (= #66 #8)
+#68 := (forall (vars (?x13 T2)) #67)
+#227 := (iff #68 #226)
+#224 := (iff #67 #222)
+#225 := [rewrite]: #224
+#228 := [quant-intro #225]: #227
+#221 := [asserted]: #68
+#231 := [mp #221 #228]: #226
+#377 := [mp~ #231 #342]: #226
+#912 := [mp #377 #911]: #907
+#526 := (not #907)
+#547 := (or #526 #835)
+#548 := (or #526 #831)
+#827 := (iff #548 #547)
+#824 := (iff #547 #547)
+#828 := [rewrite]: #824
+#545 := (iff #831 #835)
+#546 := [rewrite]: #545
+#541 := [monotonicity #546]: #827
+#825 := [trans #541 #828]: #827
+#826 := [quant-inst]: #548
+#829 := [mp #826 #825]: #547
+#578 := [unit-resolution #829 #912]: #835
+#563 := [symm #578]: #831
+#542 := (= #103 #834)
+decl uf_3 :: (-> int T3)
+decl uf_4 :: (-> T3 int)
+#30 := (uf_1 uf_7)
+#698 := (uf_4 #30)
+#11 := 1::int
+#127 := (uf_3 1::int)
+#130 := (uf_4 #127)
+#701 := (+ #130 #698)
+#704 := (uf_3 #701)
+#779 := (= #704 #834)
+#4 := (:var 1 T1)
+#6 := (uf_2 #4 #5)
+#863 := (pattern #6)
+#9 := (uf_4 #8)
+#133 := (+ #9 #130)
+#136 := (uf_3 #133)
+#7 := (uf_1 #6)
+#139 := (= #7 #136)
+#864 := (forall (vars (?x1 T1) (?x2 T2)) (:pat #863) #139)
+#142 := (forall (vars (?x1 T1) (?x2 T2)) #139)
+#867 := (iff #142 #864)
+#865 := (iff #139 #139)
+#866 := [refl]: #865
+#868 := [quant-intro #866]: #867
+#361 := (~ #142 #142)
+#359 := (~ #139 #139)
+#360 := [refl]: #359
+#362 := [nnf-pos #360]: #361
+#10 := 0::int
+#12 := (+ 0::int 1::int)
+#13 := (uf_3 #12)
+#14 := (uf_4 #13)
+#15 := (+ #9 #14)
+#16 := (uf_3 #15)
+#17 := (= #7 #16)
+#18 := (forall (vars (?x1 T1) (?x2 T2)) #17)
+#143 := (iff #18 #142)
+#140 := (iff #17 #139)
+#137 := (= #16 #136)
+#134 := (= #15 #133)
+#131 := (= #14 #130)
+#128 := (= #13 #127)
+#125 := (= #12 1::int)
+#126 := [rewrite]: #125
+#129 := [monotonicity #126]: #128
+#132 := [monotonicity #129]: #131
+#135 := [monotonicity #132]: #134
+#138 := [monotonicity #135]: #137
+#141 := [monotonicity #138]: #140
+#144 := [quant-intro #141]: #143
+#124 := [asserted]: #18
+#147 := [mp #124 #144]: #142
+#363 := [mp~ #147 #362]: #142
+#869 := [mp #363 #868]: #864
+#790 := (not #864)
+#786 := (or #790 #779)
+#699 := (+ #698 #130)
+#692 := (uf_3 #699)
+#700 := (= #834 #692)
+#791 := (or #790 #700)
+#781 := (iff #791 #786)
+#783 := (iff #786 #786)
+#777 := [rewrite]: #783
+#788 := (iff #700 #779)
+#798 := (= #834 #704)
+#776 := (iff #798 #779)
+#778 := [rewrite]: #776
+#785 := (iff #700 #798)
+#797 := (= #692 #704)
+#702 := (= #699 #701)
+#703 := [rewrite]: #702
+#794 := [monotonicity #703]: #797
+#787 := [monotonicity #794]: #785
+#789 := [trans #787 #778]: #788
+#782 := [monotonicity #789]: #781
+#784 := [trans #782 #777]: #781
+#780 := [quant-inst]: #791
+#768 := [mp #780 #784]: #786
+#577 := [unit-resolution #768 #869]: #779
+#550 := (= #103 #704)
+#572 := (= #127 #704)
+#582 := (= #704 #127)
+#598 := (= #701 1::int)
+#774 := (<= #698 0::int)
+#773 := (= #698 0::int)
+#31 := (uf_3 0::int)
+#852 := (uf_4 #31)
+#854 := (= #852 0::int)
+#72 := (:var 0 int)
+#92 := (uf_3 #72)
+#945 := (pattern #92)
+#266 := (>= #72 0::int)
+#267 := (not #266)
+#93 := (uf_4 #92)
+#248 := (= #72 #93)
+#273 := (or #248 #267)
+#946 := (forall (vars (?x18 int)) (:pat #945) #273)
+#278 := (forall (vars (?x18 int)) #273)
+#949 := (iff #278 #946)
+#947 := (iff #273 #273)
+#948 := [refl]: #947
+#950 := [quant-intro #948]: #949
+#351 := (~ #278 #278)
+#390 := (~ #273 #273)
+#391 := [refl]: #390
+#352 := [nnf-pos #391]: #351
+#94 := (= #93 #72)
+#91 := (<= 0::int #72)
+#95 := (implies #91 #94)
+#96 := (forall (vars (?x18 int)) #95)
+#281 := (iff #96 #278)
+#255 := (not #91)
+#256 := (or #255 #248)
+#261 := (forall (vars (?x18 int)) #256)
+#279 := (iff #261 #278)
+#276 := (iff #256 #273)
+#270 := (or #267 #248)
+#274 := (iff #270 #273)
+#275 := [rewrite]: #274
+#271 := (iff #256 #270)
+#268 := (iff #255 #267)
+#264 := (iff #91 #266)
+#265 := [rewrite]: #264
+#269 := [monotonicity #265]: #268
+#272 := [monotonicity #269]: #271
+#277 := [trans #272 #275]: #276
+#280 := [quant-intro #277]: #279
+#262 := (iff #96 #261)
+#259 := (iff #95 #256)
+#252 := (implies #91 #248)
+#257 := (iff #252 #256)
+#258 := [rewrite]: #257
+#253 := (iff #95 #252)
+#250 := (iff #94 #248)
+#251 := [rewrite]: #250
+#254 := [monotonicity #251]: #253
+#260 := [trans #254 #258]: #259
+#263 := [quant-intro #260]: #262
+#282 := [trans #263 #280]: #281
+#247 := [asserted]: #96
+#283 := [mp #247 #282]: #278
+#392 := [mp~ #283 #352]: #278
+#951 := [mp #392 #950]: #946
+#487 := (not #946)
+#488 := (or #487 #854)
+#859 := (>= 0::int 0::int)
+#860 := (not #859)
+#511 := (= 0::int #852)
+#516 := (or #511 #860)
+#849 := (or #487 #516)
+#850 := (iff #849 #488)
+#837 := (iff #488 #488)
+#544 := [rewrite]: #837
+#846 := (iff #516 #854)
+#843 := (or #854 false)
+#845 := (iff #843 #854)
+#482 := [rewrite]: #845
+#844 := (iff #516 #843)
+#841 := (iff #860 false)
+#1 := true
+#500 := (not true)
+#503 := (iff #500 false)
+#840 := [rewrite]: #503
+#501 := (iff #860 #500)
+#496 := (iff #859 true)
+#838 := [rewrite]: #496
+#502 := [monotonicity #838]: #501
+#842 := [trans #502 #840]: #841
+#853 := (iff #511 #854)
+#512 := [rewrite]: #853
+#839 := [monotonicity #512 #842]: #844
+#848 := [trans #839 #482]: #846
+#836 := [monotonicity #848]: #850
+#559 := [trans #836 #544]: #850
+#847 := [quant-inst]: #849
+#560 := [mp #847 #559]: #488
+#622 := [unit-resolution #560 #951]: #854
+#589 := (= #698 #852)
+#32 := (= #30 #31)
+#159 := [asserted]: #32
+#590 := [monotonicity #159]: #589
+#591 := [trans #590 #622]: #773
+#592 := (not #773)
+#593 := (or #592 #774)
+#594 := [th-lemma]: #593
+#595 := [unit-resolution #594 #591]: #774
+#770 := (>= #698 0::int)
+#596 := (or #592 #770)
+#597 := [th-lemma]: #596
+#586 := [unit-resolution #597 #591]: #770
+#680 := (<= #130 1::int)
+#605 := (= #130 1::int)
+#708 := (or #487 #605)
+#746 := (>= 1::int 0::int)
+#732 := (not #746)
+#710 := (= 1::int #130)
+#711 := (or #710 #732)
+#689 := (or #487 #711)
+#714 := (iff #689 #708)
+#715 := (iff #708 #708)
+#683 := [rewrite]: #715
+#707 := (iff #711 #605)
+#724 := (or #605 false)
+#722 := (iff #724 #605)
+#727 := [rewrite]: #722
+#725 := (iff #711 #724)
+#720 := (iff #732 false)
+#723 := (iff #732 #500)
+#717 := (iff #746 true)
+#718 := [rewrite]: #717
+#719 := [monotonicity #718]: #723
+#721 := [trans #719 #840]: #720
+#712 := (iff #710 #605)
+#716 := [rewrite]: #712
+#726 := [monotonicity #716 #721]: #725
+#653 := [trans #726 #727]: #707
+#709 := [monotonicity #653]: #714
+#690 := [trans #709 #683]: #714
+#713 := [quant-inst]: #689
+#649 := [mp #713 #690]: #708
+#616 := [unit-resolution #649 #951]: #605
+#609 := (not #605)
+#612 := (or #609 #680)
+#613 := [th-lemma]: #612
+#599 := [unit-resolution #613 #616]: #680
+#682 := (>= #130 1::int)
+#601 := (or #609 #682)
+#602 := [th-lemma]: #601
+#600 := [unit-resolution #602 #616]: #682
+#575 := [th-lemma #600 #599 #586 #595]: #598
+#583 := [monotonicity #575]: #582
+#574 := [symm #583]: #572
+#568 := (= #103 #127)
+decl uf_5 :: (-> T4 T3)
+decl uf_8 :: T4
+#33 := uf_8
+#34 := (uf_5 uf_8)
+#810 := (uf_4 #34)
+#812 := (+ #130 #810)
+#814 := (uf_3 #812)
+#571 := (= #814 #127)
+#576 := (= #127 #814)
+#587 := (= 1::int #812)
+#603 := (= #812 1::int)
+#771 := (<= #810 0::int)
+#769 := (= #810 0::int)
+#619 := (= #810 #852)
+#35 := (= #34 #31)
+#162 := (= #31 #34)
+#163 := (iff #35 #162)
+#164 := [rewrite]: #163
+#160 := [asserted]: #35
+#167 := [mp #160 #164]: #162
+#623 := [symm #167]: #35
+#624 := [monotonicity #623]: #619
+#614 := [trans #624 #622]: #769
+#604 := (not #769)
+#606 := (or #604 #771)
+#607 := [th-lemma]: #606
+#610 := [unit-resolution #607 #614]: #771
+#772 := (>= #810 0::int)
+#611 := (or #604 #772)
+#615 := [th-lemma]: #611
+#608 := [unit-resolution #615 #614]: #772
+#585 := [th-lemma #600 #599 #608 #610]: #603
+#588 := [symm #585]: #587
+#584 := [monotonicity #588]: #576
+#573 := [symm #584]: #571
+#567 := (= #103 #814)
+decl uf_6 :: (-> int T4 T4)
+#539 := (uf_6 3::int uf_8)
+#818 := (uf_5 #539)
+#646 := (= #818 #814)
+#802 := (= #814 #818)
+#20 := (:var 0 T4)
+#19 := (:var 1 int)
+#21 := (uf_6 #19 #20)
+#870 := (pattern #21)
+#23 := (uf_5 #20)
+#24 := (uf_4 #23)
+#146 := (+ #24 #130)
+#150 := (uf_3 #146)
+#22 := (uf_5 #21)
+#153 := (= #22 #150)
+#871 := (forall (vars (?x3 int) (?x4 T4)) (:pat #870) #153)
+#156 := (forall (vars (?x3 int) (?x4 T4)) #153)
+#874 := (iff #156 #871)
+#872 := (iff #153 #153)
+#873 := [refl]: #872
+#875 := [quant-intro #873]: #874
+#328 := (~ #156 #156)
+#364 := (~ #153 #153)
+#365 := [refl]: #364
+#326 := [nnf-pos #365]: #328
+#25 := (+ #24 #14)
+#26 := (uf_3 #25)
+#27 := (= #22 #26)
+#28 := (forall (vars (?x3 int) (?x4 T4)) #27)
+#157 := (iff #28 #156)
+#154 := (iff #27 #153)
+#151 := (= #26 #150)
+#148 := (= #25 #146)
+#149 := [monotonicity #132]: #148
+#152 := [monotonicity #149]: #151
+#155 := [monotonicity #152]: #154
+#158 := [quant-intro #155]: #157
+#145 := [asserted]: #28
+#161 := [mp #145 #158]: #156
+#366 := [mp~ #161 #326]: #156
+#876 := [mp #366 #875]: #871
+#687 := (not #871)
+#688 := (or #687 #802)
+#811 := (+ #810 #130)
+#805 := (uf_3 #811)
+#806 := (= #818 #805)
+#647 := (or #687 #806)
+#697 := (iff #647 #688)
+#793 := (iff #688 #688)
+#796 := [rewrite]: #793
+#804 := (iff #806 #802)
+#803 := (iff #646 #802)
+#801 := [rewrite]: #803
+#799 := (iff #806 #646)
+#643 := (= #805 #814)
+#813 := (= #811 #812)
+#807 := [rewrite]: #813
+#645 := [monotonicity #807]: #643
+#800 := [monotonicity #645]: #799
+#686 := [trans #800 #801]: #804
+#792 := [monotonicity #686]: #697
+#681 := [trans #792 #796]: #697
+#795 := [quant-inst]: #647
+#696 := [mp #795 #681]: #688
+#626 := [unit-resolution #696 #876]: #802
+#570 := [symm #626]: #646
+#557 := (= #103 #818)
+decl uf_14 :: (-> T4 T3)
+#536 := (uf_14 #539)
+#820 := (= #536 #818)
+#69 := (uf_14 #20)
+#914 := (pattern #69)
+#913 := (pattern #23)
+#230 := (= #23 #69)
+#915 := (forall (vars (?x14 T4)) (:pat #913 #914) #230)
+#234 := (forall (vars (?x14 T4)) #230)
+#918 := (iff #234 #915)
+#916 := (iff #230 #230)
+#917 := [refl]: #916
+#919 := [quant-intro #917]: #918
+#343 := (~ #234 #234)
+#378 := (~ #230 #230)
+#379 := [refl]: #378
+#344 := [nnf-pos #379]: #343
+#70 := (= #69 #23)
+#71 := (forall (vars (?x14 T4)) #70)
+#235 := (iff #71 #234)
+#232 := (iff #70 #230)
+#233 := [rewrite]: #232
+#236 := [quant-intro #233]: #235
+#229 := [asserted]: #71
+#239 := [mp #229 #236]: #234
+#380 := [mp~ #239 #344]: #234
+#920 := [mp #380 #919]: #915
+#540 := (not #915)
+#821 := (or #540 #820)
+#819 := (= #818 #536)
+#822 := (or #540 #819)
+#823 := (iff #822 #821)
+#665 := (iff #821 #821)
+#666 := [rewrite]: #665
+#815 := (iff #819 #820)
+#816 := [rewrite]: #815
+#664 := [monotonicity #816]: #823
+#808 := [trans #664 #666]: #823
+#817 := [quant-inst]: #822
+#809 := [mp #817 #808]: #821
+#628 := [unit-resolution #809 #920]: #820
+#857 := (= #103 #536)
+#75 := (uf_6 #72 uf_8)
+#922 := (pattern #75)
+#73 := (uf_16 #72)
+#921 := (pattern #73)
+#76 := (uf_14 #75)
+#74 := (uf_15 #73)
+#77 := (= #74 #76)
+#923 := (forall (vars (?x15 int)) (:pat #921 #922) #77)
+#78 := (forall (vars (?x15 int)) #77)
+#926 := (iff #78 #923)
+#924 := (iff #77 #77)
+#925 := [refl]: #924
+#927 := [quant-intro #925]: #926
+#345 := (~ #78 #78)
+#381 := (~ #77 #77)
+#382 := [refl]: #381
+#346 := [nnf-pos #382]: #345
+#237 := [asserted]: #78
+#383 := [mp~ #237 #346]: #78
+#928 := [mp #383 #927]: #923
+#851 := (not #923)
+#524 := (or #851 #857)
+#525 := [quant-inst]: #524
+#580 := [unit-resolution #525 #928]: #857
+#566 := [trans #580 #628]: #557
+#558 := [trans #566 #570]: #567
+#549 := [trans #558 #573]: #568
+#551 := [trans #549 #574]: #550
+#552 := [trans #551 #577]: #542
+#543 := [trans #552 #563]: #553
+#554 := [trans #543 #565]: #107
+#108 := (not #107)
+#325 := [asserted]: #108
+[unit-resolution #325 #554]: false
+unsat
+
--- a/src/HOL/SMT/Examples/SMT_Examples.thy Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/SMT/Examples/SMT_Examples.thy Tue Feb 02 18:16:48 2010 +0100
@@ -10,42 +10,34 @@
declare [[smt_solver=z3, z3_proofs=true]]
+declare [[smt_certificates="$ISABELLE_SMT/Examples/SMT_Examples.certs"]]
+
text {*
-To re-generate the certificates, replace the option 'smt_cert' with 'smt_keep'
-(while keeping the paths as they are) and let Isabelle process this theory.
+To avoid re-generation of certificates,
+the following option is set to "false":
*}
+declare [[smt_record=false]]
+
+
section {* Propositional and first-order logic *}
-lemma "True"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_prop_01"]]
- by smt
+lemma "True" by smt
-lemma "p \<or> \<not>p"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_prop_02"]]
- by smt
+lemma "p \<or> \<not>p" by smt
-lemma "(p \<and> True) = p"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_prop_03"]]
- by smt
+lemma "(p \<and> True) = p" by smt
-lemma "(p \<or> q) \<and> \<not>p \<Longrightarrow> q"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_prop_04"]]
- by smt
+lemma "(p \<or> q) \<and> \<not>p \<Longrightarrow> q" by smt
lemma "(a \<and> b) \<or> (c \<and> d) \<Longrightarrow> (a \<and> b) \<or> (c \<and> d)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_prop_05"]]
using [[z3_proofs=false]] (* no Z3 proof *)
by smt
-lemma "(p1 \<and> p2) \<or> p3 \<longrightarrow> (p1 \<longrightarrow> (p3 \<and> p2) \<or> (p1 \<and> p3)) \<or> p1"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_prop_06"]]
- by smt
+lemma "(p1 \<and> p2) \<or> p3 \<longrightarrow> (p1 \<longrightarrow> (p3 \<and> p2) \<or> (p1 \<and> p3)) \<or> p1" by smt
-lemma "P=P=P=P=P=P=P=P=P=P"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_prop_07"]]
- by smt
+lemma "P=P=P=P=P=P=P=P=P=P" by smt
lemma
assumes "a | b | c | d"
@@ -55,14 +47,11 @@
and "~(d | False) | c"
and "~(c | (~p & (p | (q & ~q))))"
shows False
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_prop_08"]]
using assms by smt
axiomatization symm_f :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" where
symm_f: "symm_f x y = symm_f y x"
-lemma "a = a \<and> symm_f a b = symm_f b a"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_prop_09"]]
- by (smt symm_f)
+lemma "a = a \<and> symm_f a b = symm_f b a" by (smt symm_f)
(*
Taken from ~~/src/HOL/ex/SAT_Examples.thy.
@@ -254,106 +243,69 @@
and "~x29 | ~x58"
and "~x28 | ~x58"
shows False
- using assms
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_prop_10"]]
- by smt
+ using assms by smt
lemma "\<forall>x::int. P x \<longrightarrow> (\<forall>y::int. P x \<or> P y)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_fol_01"]]
by smt
lemma
assumes "(\<forall>x y. P x y = x)"
shows "(\<exists>y. P x y) = P x c"
- using assms
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_fol_02"]]
- by smt
+ using assms by smt
lemma
assumes "(\<forall>x y. P x y = x)"
and "(\<forall>x. \<exists>y. P x y) = (\<forall>x. P x c)"
shows "(EX y. P x y) = P x c"
- using assms
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_fol_03"]]
- by smt
+ using assms by smt
lemma
assumes "if P x then \<not>(\<exists>y. P y) else (\<forall>y. \<not>P y)"
shows "P x \<longrightarrow> P y"
- using assms
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_fol_04"]]
- by smt
+ using assms by smt
section {* Arithmetic *}
subsection {* Linear arithmetic over integers and reals *}
-lemma "(3::int) = 3"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_01"]]
- by smt
+lemma "(3::int) = 3" by smt
-lemma "(3::real) = 3"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_02"]]
- by smt
+lemma "(3::real) = 3" by smt
-lemma "(3 :: int) + 1 = 4"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_03"]]
- by smt
+lemma "(3 :: int) + 1 = 4" by smt
-lemma "x + (y + z) = y + (z + (x::int))"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_04"]]
- by smt
+lemma "x + (y + z) = y + (z + (x::int))" by smt
-lemma "max (3::int) 8 > 5"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_05"]]
- by smt
+lemma "max (3::int) 8 > 5" by smt
-lemma "abs (x :: real) + abs y \<ge> abs (x + y)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_06"]]
- by smt
+lemma "abs (x :: real) + abs y \<ge> abs (x + y)" by smt
-lemma "P ((2::int) < 3) = P True"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_07"]]
- by smt
+lemma "P ((2::int) < 3) = P True" by smt
-lemma "x + 3 \<ge> 4 \<or> x < (1::int)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_08"]]
- by smt
+lemma "x + 3 \<ge> 4 \<or> x < (1::int)" by smt
lemma
assumes "x \<ge> (3::int)" and "y = x + 4"
shows "y - x > 0"
- using assms
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_09"]]
- by smt
+ using assms by smt
-lemma "let x = (2 :: int) in x + x \<noteq> 5"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_10"]]
- by smt
+lemma "let x = (2 :: int) in x + x \<noteq> 5" by smt
lemma
fixes x :: real
assumes "3 * x + 7 * a < 4" and "3 < 2 * x"
shows "a < 0"
- using assms
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_11"]]
- by smt
+ using assms by smt
-lemma "(0 \<le> y + -1 * x \<or> \<not> 0 \<le> x \<or> 0 \<le> (x::int)) = (\<not> False)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_12"]]
- by smt
+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]"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_13"]]
- by smt
+lemma "distinct [x < (3::int), 3 \<le> x]" by smt
lemma
assumes "a > (0::int)"
shows "distinct [a, a * 2, a - a]"
- using assms
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_14"]]
- by smt
+ using assms by smt
lemma "
(n < m & m < n') | (n < m & m = n') | (n < n' & n' < m) |
@@ -363,7 +315,6 @@
(m < n & n < n') | (m < n & n' = n) | (m < n' & n' < n) |
(m = n & n < n') | (m = n' & n' < n) |
(n' = m & m = (n::int))"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_15"]]
by smt
text{*
@@ -386,172 +337,109 @@
x6 = abs x5 - x4; x7 = abs x6 - x5; x8 = abs x7 - x6;
x9 = abs x8 - x7; x10 = abs x9 - x8; x11 = abs x10 - x9 \<rbrakk>
\<Longrightarrow> x1 = x10 & x2 = (x11::int)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_16"]]
by smt
-lemma "let P = 2 * x + 1 > x + (x::real) in P \<or> False \<or> P"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_17"]]
- by smt
+lemma "let P = 2 * x + 1 > x + (x::real) in P \<or> False \<or> P" by smt
+
+lemma "x + (let y = x mod 2 in 2 * y + 1) \<ge> x + (1::int)" by smt
-lemma "x + (let y = x mod 2 in 2 * y + 1) \<ge> x + (1::int)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_18"]]
- by smt
+lemma "x + (let y = x mod 2 in y + y) < x + (3::int)" by smt
-lemma "x + (let y = x mod 2 in y + y) < x + (3::int)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_19"]]
- by smt
-
-lemma
+lemma
assumes "x \<noteq> (0::real)"
shows "x + x \<noteq> (let P = (abs x > 1) in if P \<or> \<not>P then 4 else 2) * x"
- using assms
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_20"]]
- by smt
+ using assms by smt
lemma
assumes "(n + m) mod 2 = 0" and "n mod 4 = 3"
shows "n mod 2 = 1 & m mod 2 = (1::int)"
- using assms
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_linarith_21"]]
- by smt
+ using assms by smt
subsection {* Linear arithmetic with quantifiers *}
-lemma "~ (\<exists>x::int. False)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_01"]]
- by smt
+lemma "~ (\<exists>x::int. False)" by smt
-lemma "~ (\<exists>x::real. False)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_02"]]
- by smt
+lemma "~ (\<exists>x::real. False)" by smt
lemma "\<exists>x::int. 0 < x"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_03"]]
using [[z3_proofs=false]] (* no Z3 proof *)
by smt
lemma "\<exists>x::real. 0 < x"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_04"]]
using [[z3_proofs=false]] (* no Z3 proof *)
by smt
lemma "\<forall>x::int. \<exists>y. y > x"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_05"]]
using [[z3_proofs=false]] (* no Z3 proof *)
by smt
-lemma "\<forall>x y::int. (x = 0 \<and> y = 1) \<longrightarrow> x \<noteq> y"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_06"]]
- by smt
+lemma "\<forall>x y::int. (x = 0 \<and> y = 1) \<longrightarrow> x \<noteq> y" by smt
-lemma "\<exists>x::int. \<forall>y. x < y \<longrightarrow> y < 0 \<or> y >= 0"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_07"]]
- by smt
+lemma "\<exists>x::int. \<forall>y. x < y \<longrightarrow> y < 0 \<or> y >= 0" by smt
-lemma "\<forall>x y::int. x < y \<longrightarrow> (2 * x + 1) < (2 * y)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_08"]]
- by smt
+lemma "\<forall>x y::int. x < y \<longrightarrow> (2 * x + 1) < (2 * y)" by smt
-lemma "\<forall>x y::int. (2 * x + 1) \<noteq> (2 * y)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_09"]]
- by smt
+lemma "\<forall>x y::int. (2 * x + 1) \<noteq> (2 * y)" by smt
-lemma "\<forall>x y::int. x + y > 2 \<or> x + y = 2 \<or> x + y < 2"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_10"]]
- by smt
+lemma "\<forall>x y::int. x + y > 2 \<or> x + y = 2 \<or> x + y < 2" by smt
-lemma "\<forall>x::int. if x > 0 then x + 1 > 0 else 1 > x"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_11"]]
- by smt
+lemma "\<forall>x::int. if x > 0 then x + 1 > 0 else 1 > x" by smt
-lemma "if (ALL x::int. x < 0 \<or> x > 0) then False else True"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_12"]]
- by smt
+lemma "if (ALL x::int. x < 0 \<or> x > 0) then False else True" by smt
-lemma "(if (ALL x::int. x < 0 \<or> x > 0) then -1 else 3) > (0::int)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_13"]]
- by smt
+lemma "(if (ALL x::int. x < 0 \<or> x > 0) then -1 else 3) > (0::int)" by smt
-lemma "~ (\<exists>x y z::int. 4 * x + -6 * y = (1::int))"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_14"]]
- by smt
+lemma "~ (\<exists>x y z::int. 4 * x + -6 * y = (1::int))" by smt
-lemma "\<exists>x::int. \<forall>x y. 0 < x \<and> 0 < y \<longrightarrow> (0::int) < x + y"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_15"]]
- by smt
+lemma "\<exists>x::int. \<forall>x y. 0 < x \<and> 0 < y \<longrightarrow> (0::int) < x + y" by smt
-lemma "\<exists>u::int. \<forall>(x::int) y::real. 0 < x \<and> 0 < y \<longrightarrow> -1 < x"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_16"]]
- by smt
+lemma "\<exists>u::int. \<forall>(x::int) y::real. 0 < x \<and> 0 < y \<longrightarrow> -1 < x" by smt
-lemma "\<exists>x::int. (\<forall>y. y \<ge> x \<longrightarrow> y > 0) \<longrightarrow> x > 0"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_17"]]
- by smt
+lemma "\<exists>x::int. (\<forall>y. y \<ge> x \<longrightarrow> y > 0) \<longrightarrow> x > 0" by smt
-lemma "\<forall>x::int. trigger [pat x] (x < a \<longrightarrow> 2 * x < 2 * a)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_arith_quant_18"]]
- by smt
+lemma "\<forall>x::int. trigger [pat x] (x < a \<longrightarrow> 2 * x < 2 * a)" by smt
subsection {* Non-linear arithmetic over integers and reals *}
lemma "a > (0::int) \<Longrightarrow> a*b > 0 \<Longrightarrow> b > 0"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_nlarith_01"]]
using [[z3_proofs=false]] -- {* Isabelle's arithmetic decision procedures
are too weak to automatically prove @{thm zero_less_mult_pos}. *}
by smt
-lemma "(a::int) * (x + 1 + y) = a * x + a * (y + 1)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_nlarith_02"]]
- by smt
+lemma "(a::int) * (x + 1 + y) = a * x + a * (y + 1)" by smt
-lemma "((x::real) * (1 + y) - x * (1 - y)) = (2 * x * y)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_nlarith_03"]]
- by smt
+lemma "((x::real) * (1 + y) - x * (1 - y)) = (2 * x * y)" by smt
lemma
"(U::int) + (1 + p) * (b + e) + p * d =
U + (2 * (1 + p) * (b + e) + (1 + p) * d + d * p) - (1 + p) * (b + d + e)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_nlarith_04"]]
by smt
subsection {* Linear arithmetic for natural numbers *}
-lemma "2 * (x::nat) ~= 1"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_nat_arith_01"]]
- by smt
+lemma "2 * (x::nat) ~= 1" by smt
-lemma "a < 3 \<Longrightarrow> (7::nat) > 2 * a"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_nat_arith_02"]]
- by smt
+lemma "a < 3 \<Longrightarrow> (7::nat) > 2 * a" by smt
-lemma "let x = (1::nat) + y in x - y > 0 * x"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_nat_arith_03"]]
- by smt
+lemma "let x = (1::nat) + y in x - y > 0 * x" by smt
lemma
"let x = (1::nat) + y in
let P = (if x > 0 then True else False) in
False \<or> P = (x - 1 = y) \<or> (\<not>P \<longrightarrow> False)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_nat_arith_04"]]
by smt
-lemma "distinct [a + (1::nat), a * 2 + 3, a - a]"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_nat_arith_05"]]
- by smt
+lemma "distinct [a + (1::nat), a * 2 + 3, a - a]" by smt
-lemma "int (nat \<bar>x::int\<bar>) = \<bar>x\<bar>"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_nat_arith_06"]]
- by smt
+lemma "int (nat \<bar>x::int\<bar>) = \<bar>x\<bar>" by smt
definition prime_nat :: "nat \<Rightarrow> bool" where
"prime_nat p = (1 < p \<and> (\<forall>m. m dvd p --> m = 1 \<or> m = p))"
-lemma "prime_nat (4*m + 1) \<Longrightarrow> m \<ge> (1::nat)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_nat_arith_07"]]
- by (smt prime_nat_def)
+lemma "prime_nat (4*m + 1) \<Longrightarrow> m \<ge> (1::nat)" by (smt prime_nat_def)
section {* Bitvectors *}
@@ -568,107 +456,59 @@
subsection {* Bitvector arithmetic *}
-lemma "(27 :: 4 word) = -5"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_arith_01"]]
- by smt
+lemma "(27 :: 4 word) = -5" by smt
-lemma "(27 :: 4 word) = 11"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_arith_02"]]
- by smt
+lemma "(27 :: 4 word) = 11" by smt
-lemma "23 < (27::8 word)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_arith_03"]]
- by smt
+lemma "23 < (27::8 word)" by smt
-lemma "27 + 11 = (6::5 word)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_arith_04"]]
- by smt
+lemma "27 + 11 = (6::5 word)" by smt
+
+lemma "7 * 3 = (21::8 word)" by smt
-lemma "7 * 3 = (21::8 word)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_arith_05"]]
- by smt
-lemma "11 - 27 = (-16::8 word)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_arith_06"]]
- by smt
+lemma "11 - 27 = (-16::8 word)" by smt
-lemma "- -11 = (11::5 word)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_arith_07"]]
- by smt
+lemma "- -11 = (11::5 word)" by smt
-lemma "-40 + 1 = (-39::7 word)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_arith_08"]]
- by smt
+lemma "-40 + 1 = (-39::7 word)" by smt
-lemma "a + 2 * b + c - b = (b + c) + (a :: 32 word)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_arith_09"]]
- by smt
+lemma "a + 2 * b + c - b = (b + c) + (a :: 32 word)" by smt
-lemma "x = (5 :: 4 word) \<Longrightarrow> 4 * x = 4"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_arith_10"]]
- by smt
+lemma "x = (5 :: 4 word) \<Longrightarrow> 4 * x = 4" by smt
subsection {* Bit-level logic *}
-lemma "0b110 AND 0b101 = (0b100 :: 32 word)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_bit_01"]]
- by smt
+lemma "0b110 AND 0b101 = (0b100 :: 32 word)" by smt
-lemma "0b110 OR 0b011 = (0b111 :: 8 word)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_bit_02"]]
- by smt
+lemma "0b110 OR 0b011 = (0b111 :: 8 word)" by smt
-lemma "0xF0 XOR 0xFF = (0x0F :: 8 word)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_bit_03"]]
- by smt
+lemma "0xF0 XOR 0xFF = (0x0F :: 8 word)" by smt
-lemma "NOT (0xF0 :: 16 word) = 0xFF0F"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_bit_04"]]
- by smt
+lemma "NOT (0xF0 :: 16 word) = 0xFF0F" by smt
-lemma "word_cat (27::4 word) (27::8 word) = (2843::12 word)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_bit_05"]]
- by smt
+lemma "word_cat (27::4 word) (27::8 word) = (2843::12 word)" by smt
lemma "word_cat (0b0011::4 word) (0b1111::6word) = (0b0011001111 :: 10 word)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_bit_06"]]
- by smt
-
-lemma "slice 1 (0b10110 :: 4 word) = (0b11 :: 2 word)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_bit_07"]]
by smt
-lemma "ucast (0b1010 :: 4 word) = (0b1010 :: 10 word)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_bit_08"]]
- by smt
+lemma "slice 1 (0b10110 :: 4 word) = (0b11 :: 2 word)" by smt
-lemma "scast (0b1010 :: 4 word) = (0b111010 :: 6 word)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_bit_09"]]
- by smt
+lemma "ucast (0b1010 :: 4 word) = (0b1010 :: 10 word)" by smt
-lemma "bv_lshr 0b10011 2 = (0b100::8 word)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_bit_10"]]
- by smt
+lemma "scast (0b1010 :: 4 word) = (0b111010 :: 6 word)" by smt
-lemma "bv_ashr 0b10011 2 = (0b100::8 word)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_bit_11"]]
- by smt
+lemma "bv_lshr 0b10011 2 = (0b100::8 word)" by smt
-lemma "word_rotr 2 0b0110 = (0b1001::4 word)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_bit_12"]]
- by smt
+lemma "bv_ashr 0b10011 2 = (0b100::8 word)" by smt
-lemma "word_rotl 1 0b1110 = (0b1101::4 word)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_bit_13"]]
- by smt
+lemma "word_rotr 2 0b0110 = (0b1001::4 word)" by smt
-lemma "(x AND 0xff00) OR (x AND 0x00ff) = (x::16 word)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_bit_14"]]
- by smt
+lemma "word_rotl 1 0b1110 = (0b1101::4 word)" by smt
-lemma "w < 256 \<Longrightarrow> (w :: 16 word) AND 0x00FF = w"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_bit_15"]]
- by smt
+lemma "(x AND 0xff00) OR (x AND 0x00ff) = (x::16 word)" by smt
+
+lemma "w < 256 \<Longrightarrow> (w :: 16 word) AND 0x00FF = w" by smt
end
@@ -681,57 +521,37 @@
shows "\<forall>i::int. i < 0 \<longrightarrow> (\<forall>x::2 word. bv2int x > i)"
using assms
using [[smt_solver=z3]]
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_01"]]
by smt
lemma "P (0 \<le> (a :: 4 word)) = P True"
using [[smt_solver=z3, z3_proofs=false]]
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_bv_02"]]
by smt
section {* Pairs *}
-lemma "fst (x, y) = a \<Longrightarrow> x = a"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_pair_01"]]
- by smt
+lemma "fst (x, y) = a \<Longrightarrow> x = a" by smt
-lemma "p1 = (x, y) \<and> p2 = (y, x) \<Longrightarrow> fst p1 = snd p2"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_pair_02"]]
- by smt
+lemma "p1 = (x, y) \<and> p2 = (y, x) \<Longrightarrow> fst p1 = snd p2" by smt
section {* Higher-order problems and recursion *}
-lemma "i \<noteq> i1 \<and> i \<noteq> i2 \<Longrightarrow> (f (i1 := v1, i2 := v2)) i = f i"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_hol_01"]]
- by smt
+lemma "i \<noteq> i1 \<and> i \<noteq> i2 \<Longrightarrow> (f (i1 := v1, i2 := v2)) i = f i" by smt
-lemma "(f g x = (g x \<and> True)) \<or> (f g x = True) \<or> (g x = True)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_hol_02"]]
- by smt
+lemma "(f g x = (g x \<and> True)) \<or> (f g x = True) \<or> (g x = True)" by smt
-lemma "id 3 = 3 \<and> id True = True"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_hol_03"]]
- by (smt id_def)
+lemma "id 3 = 3 \<and> id True = True" by (smt id_def)
-lemma "i \<noteq> i1 \<and> i \<noteq> i2 \<Longrightarrow> ((f (i1 := v1)) (i2 := v2)) i = f i"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_hol_04"]]
- by smt
+lemma "i \<noteq> i1 \<and> i \<noteq> i2 \<Longrightarrow> ((f (i1 := v1)) (i2 := v2)) i = f i" by smt
-lemma "map (\<lambda>i::nat. i + 1) [0, 1] = [1, 2]"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_hol_05"]]
- by (smt map.simps)
+lemma "map (\<lambda>i::nat. i + 1) [0, 1] = [1, 2]" by (smt map.simps)
-lemma "(ALL x. P x) | ~ All P"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_hol_06"]]
- by smt
+lemma "(ALL x. P x) | ~ All P" by smt
fun dec_10 :: "nat \<Rightarrow> nat" where
"dec_10 n = (if n < 10 then n else dec_10 (n - 10))"
-lemma "dec_10 (4 * dec_10 4) = 6"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_hol_07"]]
- by (smt dec_10.simps)
+lemma "dec_10 (4 * dec_10 4) = 6" by (smt dec_10.simps)
axiomatization
eval_dioph :: "int list \<Rightarrow> nat list \<Rightarrow> int"
@@ -747,7 +567,6 @@
(eval_dioph ks (map (\<lambda>x. x mod 2) xs) mod 2 = l mod 2 \<and>
eval_dioph ks (map (\<lambda>x. x div 2) xs) =
(l - eval_dioph ks (map (\<lambda>x. x mod 2) xs)) div 2)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_hol_08"]]
by (smt eval_dioph_mod[where n=2] eval_dioph_div_mult[where n=2])
@@ -755,17 +574,13 @@
definition P :: "'a \<Rightarrow> bool" where "P x = True"
lemma poly_P: "P x \<and> (P [x] \<or> \<not>P[x])" by (simp add: P_def)
-lemma "P (1::int)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_mono_01"]]
- by (smt poly_P)
+lemma "P (1::int)" by (smt poly_P)
consts g :: "'a \<Rightarrow> nat"
axioms
g1: "g (Some x) = g [x]"
g2: "g None = g []"
g3: "g xs = length xs"
-lemma "g (Some (3::int)) = g (Some True)"
- using [[smt_cert="$ISABELLE_SMT/Examples/cert/z3_mono_02"]]
- by (smt g1 g2 g3 list.size)
+lemma "g (Some (3::int)) = g (Some True)" by (smt g1 g2 g3 list.size)
end
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_01 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (not (exists (?x1 Int) false)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_01.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,24 +0,0 @@
-#2 := false
-#4 := (exists (vars (?x1 int)) false)
-#5 := (not #4)
-#6 := (not #5)
-#37 := (iff #6 false)
-#1 := true
-#32 := (not true)
-#35 := (iff #32 false)
-#36 := [rewrite]: #35
-#33 := (iff #6 #32)
-#30 := (iff #5 true)
-#25 := (not false)
-#28 := (iff #25 true)
-#29 := [rewrite]: #28
-#26 := (iff #5 #25)
-#23 := (iff #4 false)
-#24 := [elim-unused]: #23
-#27 := [monotonicity #24]: #26
-#31 := [trans #27 #29]: #30
-#34 := [monotonicity #31]: #33
-#38 := [trans #34 #36]: #37
-#22 := [asserted]: #6
-[mp #22 #38]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_02 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (not (exists (?x1 Real) false)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_02.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,24 +0,0 @@
-#2 := false
-#4 := (exists (vars (?x1 real)) false)
-#5 := (not #4)
-#6 := (not #5)
-#37 := (iff #6 false)
-#1 := true
-#32 := (not true)
-#35 := (iff #32 false)
-#36 := [rewrite]: #35
-#33 := (iff #6 #32)
-#30 := (iff #5 true)
-#25 := (not false)
-#28 := (iff #25 true)
-#29 := [rewrite]: #28
-#26 := (iff #5 #25)
-#23 := (iff #4 false)
-#24 := [elim-unused]: #23
-#27 := [monotonicity #24]: #26
-#31 := [trans #27 #29]: #30
-#34 := [monotonicity #31]: #33
-#38 := [trans #34 #36]: #37
-#22 := [asserted]: #6
-[mp #22 #38]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_03 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (exists (?x1 Int) (< 0 ?x1)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_03.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_04 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (exists (?x1 Real) (< 0.0 ?x1)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_04.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_05 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (forall (?x1 Int) (exists (?x2 Int) (< ?x1 ?x2))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_05.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_06 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (forall (?x1 Int) (?x2 Int) (implies (and (= ?x1 0) (= ?x2 1)) (not (= ?x1 ?x2)))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_06.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,73 +0,0 @@
-#2 := false
-#5 := 0::int
-#8 := 1::int
-#143 := (= 1::int 0::int)
-#145 := (iff #143 false)
-#146 := [rewrite]: #145
-decl ?x1!1 :: int
-#47 := ?x1!1
-#51 := (= ?x1!1 0::int)
-decl ?x2!0 :: int
-#46 := ?x2!0
-#50 := (= ?x2!0 1::int)
-#63 := (and #50 #51)
-#69 := (= ?x2!0 ?x1!1)
-#72 := (not #69)
-#66 := (not #63)
-#75 := (or #66 #72)
-#78 := (not #75)
-#48 := (= ?x1!1 ?x2!0)
-#49 := (not #48)
-#52 := (and #51 #50)
-#53 := (not #52)
-#54 := (or #53 #49)
-#55 := (not #54)
-#79 := (iff #55 #78)
-#76 := (iff #54 #75)
-#73 := (iff #49 #72)
-#70 := (iff #48 #69)
-#71 := [rewrite]: #70
-#74 := [monotonicity #71]: #73
-#67 := (iff #53 #66)
-#64 := (iff #52 #63)
-#65 := [rewrite]: #64
-#68 := [monotonicity #65]: #67
-#77 := [monotonicity #68 #74]: #76
-#80 := [monotonicity #77]: #79
-#7 := (:var 0 int)
-#4 := (:var 1 int)
-#11 := (= #4 #7)
-#12 := (not #11)
-#9 := (= #7 1::int)
-#6 := (= #4 0::int)
-#10 := (and #6 #9)
-#32 := (not #10)
-#33 := (or #32 #12)
-#36 := (forall (vars (?x1 int) (?x2 int)) #33)
-#39 := (not #36)
-#56 := (~ #39 #55)
-#57 := [sk]: #56
-#13 := (implies #10 #12)
-#14 := (forall (vars (?x1 int) (?x2 int)) #13)
-#15 := (not #14)
-#40 := (iff #15 #39)
-#37 := (iff #14 #36)
-#34 := (iff #13 #33)
-#35 := [rewrite]: #34
-#38 := [quant-intro #35]: #37
-#41 := [monotonicity #38]: #40
-#31 := [asserted]: #15
-#44 := [mp #31 #41]: #39
-#60 := [mp~ #44 #57]: #55
-#61 := [mp #60 #80]: #78
-#62 := [not-or-elim #61]: #63
-#82 := [and-elim #62]: #51
-#141 := (= 1::int ?x1!1)
-#83 := [not-or-elim #61]: #69
-#139 := (= 1::int ?x2!0)
-#81 := [and-elim #62]: #50
-#140 := [symm #81]: #139
-#142 := [trans #140 #83]: #141
-#144 := [trans #142 #82]: #143
-[mp #144 #146]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_07 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (exists (?x1 Int) (forall (?x2 Int) (implies (< ?x1 ?x2) (or (< ?x2 0) (<= 0 ?x2))))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_07.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,82 +0,0 @@
-#2 := false
-#5 := (:var 0 int)
-#7 := 0::int
-#9 := (<= 0::int #5)
-#8 := (< #5 0::int)
-#10 := (or #8 #9)
-#4 := (:var 1 int)
-#6 := (< #4 #5)
-#11 := (implies #6 #10)
-#12 := (forall (vars (?x2 int)) #11)
-#13 := (exists (vars (?x1 int)) #12)
-#14 := (not #13)
-#95 := (iff #14 false)
-#31 := (not #6)
-#32 := (or #31 #10)
-#35 := (forall (vars (?x2 int)) #32)
-#38 := (exists (vars (?x1 int)) #35)
-#41 := (not #38)
-#93 := (iff #41 false)
-#1 := true
-#88 := (not true)
-#91 := (iff #88 false)
-#92 := [rewrite]: #91
-#89 := (iff #41 #88)
-#86 := (iff #38 true)
-#81 := (exists (vars (?x1 int)) true)
-#84 := (iff #81 true)
-#85 := [elim-unused]: #84
-#82 := (iff #38 #81)
-#79 := (iff #35 true)
-#74 := (forall (vars (?x2 int)) true)
-#77 := (iff #74 true)
-#78 := [elim-unused]: #77
-#75 := (iff #35 #74)
-#72 := (iff #32 true)
-#46 := (>= #5 0::int)
-#44 := (not #46)
-#64 := (or #44 #46)
-#50 := -1::int
-#53 := (* -1::int #5)
-#54 := (+ #4 #53)
-#52 := (>= #54 0::int)
-#67 := (or #52 #64)
-#70 := (iff #67 true)
-#71 := [rewrite]: #70
-#68 := (iff #32 #67)
-#65 := (iff #10 #64)
-#48 := (iff #9 #46)
-#49 := [rewrite]: #48
-#45 := (iff #8 #44)
-#47 := [rewrite]: #45
-#66 := [monotonicity #47 #49]: #65
-#62 := (iff #31 #52)
-#51 := (not #52)
-#57 := (not #51)
-#60 := (iff #57 #52)
-#61 := [rewrite]: #60
-#58 := (iff #31 #57)
-#55 := (iff #6 #51)
-#56 := [rewrite]: #55
-#59 := [monotonicity #56]: #58
-#63 := [trans #59 #61]: #62
-#69 := [monotonicity #63 #66]: #68
-#73 := [trans #69 #71]: #72
-#76 := [quant-intro #73]: #75
-#80 := [trans #76 #78]: #79
-#83 := [quant-intro #80]: #82
-#87 := [trans #83 #85]: #86
-#90 := [monotonicity #87]: #89
-#94 := [trans #90 #92]: #93
-#42 := (iff #14 #41)
-#39 := (iff #13 #38)
-#36 := (iff #12 #35)
-#33 := (iff #11 #32)
-#34 := [rewrite]: #33
-#37 := [quant-intro #34]: #36
-#40 := [quant-intro #37]: #39
-#43 := [monotonicity #40]: #42
-#96 := [trans #43 #94]: #95
-#30 := [asserted]: #14
-[mp #30 #96]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_08 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (forall (?x1 Int) (?x2 Int) (implies (< ?x1 ?x2) (< (+ (* 2 ?x1) 1) (* 2 ?x2)))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_08.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,128 +0,0 @@
-#2 := false
-#9 := 1::int
-decl ?x1!1 :: int
-#91 := ?x1!1
-#68 := -2::int
-#129 := (* -2::int ?x1!1)
-decl ?x2!0 :: int
-#90 := ?x2!0
-#7 := 2::int
-#128 := (* 2::int ?x2!0)
-#130 := (+ #128 #129)
-#131 := (<= #130 1::int)
-#136 := (not #131)
-#55 := 0::int
-#53 := -1::int
-#115 := (* -1::int ?x1!1)
-#116 := (+ ?x2!0 #115)
-#117 := (<= #116 0::int)
-#139 := (or #117 #136)
-#142 := (not #139)
-#92 := (* -2::int ?x2!0)
-#93 := (* 2::int ?x1!1)
-#94 := (+ #93 #92)
-#95 := (>= #94 -1::int)
-#96 := (not #95)
-#97 := (* -1::int ?x2!0)
-#98 := (+ ?x1!1 #97)
-#99 := (>= #98 0::int)
-#100 := (or #99 #96)
-#101 := (not #100)
-#143 := (iff #101 #142)
-#140 := (iff #100 #139)
-#137 := (iff #96 #136)
-#134 := (iff #95 #131)
-#122 := (+ #92 #93)
-#125 := (>= #122 -1::int)
-#132 := (iff #125 #131)
-#133 := [rewrite]: #132
-#126 := (iff #95 #125)
-#123 := (= #94 #122)
-#124 := [rewrite]: #123
-#127 := [monotonicity #124]: #126
-#135 := [trans #127 #133]: #134
-#138 := [monotonicity #135]: #137
-#120 := (iff #99 #117)
-#109 := (+ #97 ?x1!1)
-#112 := (>= #109 0::int)
-#118 := (iff #112 #117)
-#119 := [rewrite]: #118
-#113 := (iff #99 #112)
-#110 := (= #98 #109)
-#111 := [rewrite]: #110
-#114 := [monotonicity #111]: #113
-#121 := [trans #114 #119]: #120
-#141 := [monotonicity #121 #138]: #140
-#144 := [monotonicity #141]: #143
-#5 := (:var 0 int)
-#71 := (* -2::int #5)
-#4 := (:var 1 int)
-#8 := (* 2::int #4)
-#72 := (+ #8 #71)
-#70 := (>= #72 -1::int)
-#69 := (not #70)
-#57 := (* -1::int #5)
-#58 := (+ #4 #57)
-#56 := (>= #58 0::int)
-#75 := (or #56 #69)
-#78 := (forall (vars (?x1 int) (?x2 int)) #75)
-#81 := (not #78)
-#102 := (~ #81 #101)
-#103 := [sk]: #102
-#11 := (* 2::int #5)
-#10 := (+ #8 1::int)
-#12 := (< #10 #11)
-#6 := (< #4 #5)
-#13 := (implies #6 #12)
-#14 := (forall (vars (?x1 int) (?x2 int)) #13)
-#15 := (not #14)
-#84 := (iff #15 #81)
-#32 := (+ 1::int #8)
-#35 := (< #32 #11)
-#41 := (not #6)
-#42 := (or #41 #35)
-#47 := (forall (vars (?x1 int) (?x2 int)) #42)
-#50 := (not #47)
-#82 := (iff #50 #81)
-#79 := (iff #47 #78)
-#76 := (iff #42 #75)
-#73 := (iff #35 #69)
-#74 := [rewrite]: #73
-#66 := (iff #41 #56)
-#54 := (not #56)
-#61 := (not #54)
-#64 := (iff #61 #56)
-#65 := [rewrite]: #64
-#62 := (iff #41 #61)
-#59 := (iff #6 #54)
-#60 := [rewrite]: #59
-#63 := [monotonicity #60]: #62
-#67 := [trans #63 #65]: #66
-#77 := [monotonicity #67 #74]: #76
-#80 := [quant-intro #77]: #79
-#83 := [monotonicity #80]: #82
-#51 := (iff #15 #50)
-#48 := (iff #14 #47)
-#45 := (iff #13 #42)
-#38 := (implies #6 #35)
-#43 := (iff #38 #42)
-#44 := [rewrite]: #43
-#39 := (iff #13 #38)
-#36 := (iff #12 #35)
-#33 := (= #10 #32)
-#34 := [rewrite]: #33
-#37 := [monotonicity #34]: #36
-#40 := [monotonicity #37]: #39
-#46 := [trans #40 #44]: #45
-#49 := [quant-intro #46]: #48
-#52 := [monotonicity #49]: #51
-#85 := [trans #52 #83]: #84
-#31 := [asserted]: #15
-#86 := [mp #31 #85]: #81
-#106 := [mp~ #86 #103]: #101
-#107 := [mp #106 #144]: #142
-#146 := [not-or-elim #107]: #131
-#108 := (not #117)
-#145 := [not-or-elim #107]: #108
-[th-lemma #145 #146]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_09 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (forall (?x1 Int) (?x2 Int) (not (= (+ (* 2 ?x1) 1) (* 2 ?x2)))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_09.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,100 +0,0 @@
-#2 := false
-#7 := 1::int
-decl ?x1!1 :: int
-#74 := ?x1!1
-#51 := -2::int
-#96 := (* -2::int ?x1!1)
-decl ?x2!0 :: int
-#73 := ?x2!0
-#4 := 2::int
-#95 := (* 2::int ?x2!0)
-#97 := (+ #95 #96)
-#166 := (<= #97 1::int)
-#94 := (= #97 1::int)
-#53 := -1::int
-#75 := (* -2::int ?x2!0)
-#76 := (* 2::int ?x1!1)
-#77 := (+ #76 #75)
-#78 := (= #77 -1::int)
-#79 := (not #78)
-#80 := (not #79)
-#110 := (iff #80 #94)
-#102 := (not #94)
-#105 := (not #102)
-#108 := (iff #105 #94)
-#109 := [rewrite]: #108
-#106 := (iff #80 #105)
-#103 := (iff #79 #102)
-#100 := (iff #78 #94)
-#88 := (+ #75 #76)
-#91 := (= #88 -1::int)
-#98 := (iff #91 #94)
-#99 := [rewrite]: #98
-#92 := (iff #78 #91)
-#89 := (= #77 #88)
-#90 := [rewrite]: #89
-#93 := [monotonicity #90]: #92
-#101 := [trans #93 #99]: #100
-#104 := [monotonicity #101]: #103
-#107 := [monotonicity #104]: #106
-#111 := [trans #107 #109]: #110
-#9 := (:var 0 int)
-#55 := (* -2::int #9)
-#5 := (:var 1 int)
-#6 := (* 2::int #5)
-#56 := (+ #6 #55)
-#54 := (= #56 -1::int)
-#58 := (not #54)
-#61 := (forall (vars (?x1 int) (?x2 int)) #58)
-#64 := (not #61)
-#81 := (~ #64 #80)
-#82 := [sk]: #81
-#10 := (* 2::int #9)
-#8 := (+ #6 1::int)
-#11 := (= #8 #10)
-#12 := (not #11)
-#13 := (forall (vars (?x1 int) (?x2 int)) #12)
-#14 := (not #13)
-#67 := (iff #14 #64)
-#31 := (+ 1::int #6)
-#37 := (= #10 #31)
-#42 := (not #37)
-#45 := (forall (vars (?x1 int) (?x2 int)) #42)
-#48 := (not #45)
-#65 := (iff #48 #64)
-#62 := (iff #45 #61)
-#59 := (iff #42 #58)
-#52 := (iff #37 #54)
-#57 := [rewrite]: #52
-#60 := [monotonicity #57]: #59
-#63 := [quant-intro #60]: #62
-#66 := [monotonicity #63]: #65
-#49 := (iff #14 #48)
-#46 := (iff #13 #45)
-#43 := (iff #12 #42)
-#40 := (iff #11 #37)
-#34 := (= #31 #10)
-#38 := (iff #34 #37)
-#39 := [rewrite]: #38
-#35 := (iff #11 #34)
-#32 := (= #8 #31)
-#33 := [rewrite]: #32
-#36 := [monotonicity #33]: #35
-#41 := [trans #36 #39]: #40
-#44 := [monotonicity #41]: #43
-#47 := [quant-intro #44]: #46
-#50 := [monotonicity #47]: #49
-#68 := [trans #50 #66]: #67
-#30 := [asserted]: #14
-#69 := [mp #30 #68]: #64
-#85 := [mp~ #69 #82]: #80
-#86 := [mp #85 #111]: #94
-#168 := (or #102 #166)
-#169 := [th-lemma]: #168
-#170 := [unit-resolution #169 #86]: #166
-#167 := (>= #97 1::int)
-#171 := (or #102 #167)
-#172 := [th-lemma]: #171
-#173 := [unit-resolution #172 #86]: #167
-[th-lemma #173 #170]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_10 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (forall (?x1 Int) (?x2 Int) (or (< 2 (+ ?x1 ?x2)) (or (= (+ ?x1 ?x2) 2) (< (+ ?x1 ?x2) 2)))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_10.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,111 +0,0 @@
-#2 := false
-#4 := 2::int
-decl ?x1!1 :: int
-#85 := ?x1!1
-decl ?x2!0 :: int
-#84 := ?x2!0
-#101 := (+ ?x2!0 ?x1!1)
-#107 := (>= #101 2::int)
-#113 := (<= #101 2::int)
-#116 := (not #113)
-#110 := (not #107)
-#104 := (= #101 2::int)
-#119 := (or #104 #110 #116)
-#122 := (not #119)
-#86 := (+ ?x1!1 ?x2!0)
-#87 := (<= #86 2::int)
-#88 := (not #87)
-#89 := (>= #86 2::int)
-#90 := (not #89)
-#91 := (= #86 2::int)
-#92 := (or #91 #90 #88)
-#93 := (not #92)
-#123 := (iff #93 #122)
-#120 := (iff #92 #119)
-#117 := (iff #88 #116)
-#114 := (iff #87 #113)
-#102 := (= #86 #101)
-#103 := [rewrite]: #102
-#115 := [monotonicity #103]: #114
-#118 := [monotonicity #115]: #117
-#111 := (iff #90 #110)
-#108 := (iff #89 #107)
-#109 := [monotonicity #103]: #108
-#112 := [monotonicity #109]: #111
-#105 := (iff #91 #104)
-#106 := [monotonicity #103]: #105
-#121 := [monotonicity #106 #112 #118]: #120
-#124 := [monotonicity #121]: #123
-#6 := (:var 0 int)
-#5 := (:var 1 int)
-#7 := (+ #5 #6)
-#56 := (<= #7 2::int)
-#58 := (not #56)
-#54 := (>= #7 2::int)
-#51 := (not #54)
-#9 := (= #7 2::int)
-#67 := (or #9 #51 #58)
-#72 := (forall (vars (?x1 int) (?x2 int)) #67)
-#75 := (not #72)
-#94 := (~ #75 #93)
-#95 := [sk]: #94
-#10 := (< #7 2::int)
-#11 := (or #9 #10)
-#8 := (< 2::int #7)
-#12 := (or #8 #11)
-#13 := (forall (vars (?x1 int) (?x2 int)) #12)
-#14 := (not #13)
-#78 := (iff #14 #75)
-#31 := (= 2::int #7)
-#37 := (or #10 #31)
-#42 := (or #8 #37)
-#45 := (forall (vars (?x1 int) (?x2 int)) #42)
-#48 := (not #45)
-#76 := (iff #48 #75)
-#73 := (iff #45 #72)
-#70 := (iff #42 #67)
-#61 := (or #51 #9)
-#64 := (or #58 #61)
-#68 := (iff #64 #67)
-#69 := [rewrite]: #68
-#65 := (iff #42 #64)
-#62 := (iff #37 #61)
-#55 := (iff #31 #9)
-#57 := [rewrite]: #55
-#53 := (iff #10 #51)
-#52 := [rewrite]: #53
-#63 := [monotonicity #52 #57]: #62
-#59 := (iff #8 #58)
-#60 := [rewrite]: #59
-#66 := [monotonicity #60 #63]: #65
-#71 := [trans #66 #69]: #70
-#74 := [quant-intro #71]: #73
-#77 := [monotonicity #74]: #76
-#49 := (iff #14 #48)
-#46 := (iff #13 #45)
-#43 := (iff #12 #42)
-#40 := (iff #11 #37)
-#34 := (or #31 #10)
-#38 := (iff #34 #37)
-#39 := [rewrite]: #38
-#35 := (iff #11 #34)
-#32 := (iff #9 #31)
-#33 := [rewrite]: #32
-#36 := [monotonicity #33]: #35
-#41 := [trans #36 #39]: #40
-#44 := [monotonicity #41]: #43
-#47 := [quant-intro #44]: #46
-#50 := [monotonicity #47]: #49
-#79 := [trans #50 #77]: #78
-#30 := [asserted]: #14
-#80 := [mp #30 #79]: #75
-#98 := [mp~ #80 #95]: #93
-#99 := [mp #98 #124]: #122
-#126 := [not-or-elim #99]: #107
-#100 := (not #104)
-#125 := [not-or-elim #99]: #100
-#127 := [not-or-elim #99]: #113
-#183 := (or #104 #116 #110)
-#184 := [th-lemma]: #183
-[unit-resolution #184 #127 #125 #126]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_11 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (forall (?x1 Int) (if_then_else (< 0 ?x1) (< 0 (+ ?x1 1)) (< ?x1 1))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_11.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,89 +0,0 @@
-#2 := false
-#4 := 0::int
-decl ?x1!0 :: int
-#78 := ?x1!0
-#83 := (<= ?x1!0 0::int)
-#146 := (not #83)
-#155 := [hypothesis]: #83
-#7 := 1::int
-#81 := (>= ?x1!0 1::int)
-#82 := (not #81)
-#156 := (or #82 #146)
-#157 := [th-lemma]: #156
-#158 := [unit-resolution #157 #155]: #82
-#159 := (or #146 #81)
-#49 := -1::int
-#79 := (<= ?x1!0 -1::int)
-#80 := (not #79)
-#84 := (ite #83 #82 #80)
-#85 := (not #84)
-#5 := (:var 0 int)
-#50 := (<= #5 -1::int)
-#51 := (not #50)
-#55 := (>= #5 1::int)
-#54 := (not #55)
-#45 := (<= #5 0::int)
-#61 := (ite #45 #54 #51)
-#66 := (forall (vars (?x1 int)) #61)
-#69 := (not #66)
-#86 := (~ #69 #85)
-#87 := [sk]: #86
-#10 := (< #5 1::int)
-#8 := (+ #5 1::int)
-#9 := (< 0::int #8)
-#6 := (< 0::int #5)
-#11 := (ite #6 #9 #10)
-#12 := (forall (vars (?x1 int)) #11)
-#13 := (not #12)
-#72 := (iff #13 #69)
-#30 := (+ 1::int #5)
-#33 := (< 0::int #30)
-#36 := (ite #6 #33 #10)
-#39 := (forall (vars (?x1 int)) #36)
-#42 := (not #39)
-#70 := (iff #42 #69)
-#67 := (iff #39 #66)
-#64 := (iff #36 #61)
-#46 := (not #45)
-#58 := (ite #46 #51 #54)
-#62 := (iff #58 #61)
-#63 := [rewrite]: #62
-#59 := (iff #36 #58)
-#56 := (iff #10 #54)
-#57 := [rewrite]: #56
-#52 := (iff #33 #51)
-#53 := [rewrite]: #52
-#47 := (iff #6 #46)
-#48 := [rewrite]: #47
-#60 := [monotonicity #48 #53 #57]: #59
-#65 := [trans #60 #63]: #64
-#68 := [quant-intro #65]: #67
-#71 := [monotonicity #68]: #70
-#43 := (iff #13 #42)
-#40 := (iff #12 #39)
-#37 := (iff #11 #36)
-#34 := (iff #9 #33)
-#31 := (= #8 #30)
-#32 := [rewrite]: #31
-#35 := [monotonicity #32]: #34
-#38 := [monotonicity #35]: #37
-#41 := [quant-intro #38]: #40
-#44 := [monotonicity #41]: #43
-#73 := [trans #44 #71]: #72
-#29 := [asserted]: #13
-#74 := [mp #29 #73]: #69
-#90 := [mp~ #74 #87]: #85
-#151 := (or #84 #146 #81)
-#152 := [def-axiom]: #151
-#160 := [unit-resolution #152 #90]: #159
-#161 := [unit-resolution #160 #158 #155]: false
-#162 := [lemma #161]: #146
-#163 := (or #80 #83)
-#164 := [th-lemma]: #163
-#165 := [unit-resolution #164 #162]: #80
-#166 := (or #83 #79)
-#153 := (or #84 #83 #79)
-#154 := [def-axiom]: #153
-#167 := [unit-resolution #154 #90]: #166
-[unit-resolution #167 #165 #162]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_12 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (if_then_else (forall (?x1 Int) (or (< ?x1 0) (< 0 ?x1))) false true))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_12.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,83 +0,0 @@
-#2 := false
-#5 := 0::int
-#4 := (:var 0 int)
-#42 := (<= #4 0::int)
-#43 := (not #42)
-#40 := (>= #4 0::int)
-#38 := (not #40)
-#46 := (or #38 #43)
-#49 := (forall (vars (?x1 int)) #46)
-#524 := (not #49)
-#118 := (<= 0::int 0::int)
-#205 := (not #118)
-#119 := (>= 0::int 0::int)
-#206 := (not #119)
-#120 := (or #206 #205)
-#183 := (or #524 #120)
-#172 := (iff #183 #524)
-#525 := (or #524 false)
-#168 := (iff #525 #524)
-#510 := [rewrite]: #168
-#184 := (iff #183 #525)
-#528 := (iff #120 false)
-#197 := (or false false)
-#532 := (iff #197 false)
-#533 := [rewrite]: #532
-#530 := (iff #120 #197)
-#523 := (iff #205 false)
-#1 := true
-#209 := (not true)
-#211 := (iff #209 false)
-#208 := [rewrite]: #211
-#185 := (iff #205 #209)
-#527 := (iff #118 true)
-#529 := [rewrite]: #527
-#316 := [monotonicity #529]: #185
-#196 := [trans #316 #208]: #523
-#212 := (iff #206 false)
-#210 := (iff #206 #209)
-#207 := (iff #119 true)
-#198 := [rewrite]: #207
-#138 := [monotonicity #198]: #210
-#191 := [trans #138 #208]: #212
-#531 := [monotonicity #191 #196]: #530
-#534 := [trans #531 #533]: #528
-#526 := [monotonicity #534]: #184
-#173 := [trans #526 #510]: #172
-#188 := [quant-inst]: #183
-#174 := [mp #188 #173]: #524
-#60 := (~ #49 #49)
-#58 := (~ #46 #46)
-#59 := [refl]: #58
-#61 := [nnf-pos #59]: #60
-#7 := (< 0::int #4)
-#6 := (< #4 0::int)
-#8 := (or #6 #7)
-#9 := (forall (vars (?x1 int)) #8)
-#10 := (ite #9 false true)
-#11 := (not #10)
-#52 := (iff #11 #49)
-#50 := (iff #9 #49)
-#47 := (iff #8 #46)
-#44 := (iff #7 #43)
-#45 := [rewrite]: #44
-#39 := (iff #6 #38)
-#41 := [rewrite]: #39
-#48 := [monotonicity #41 #45]: #47
-#51 := [quant-intro #48]: #50
-#36 := (iff #11 #9)
-#28 := (not #9)
-#31 := (not #28)
-#34 := (iff #31 #9)
-#35 := [rewrite]: #34
-#32 := (iff #11 #31)
-#29 := (iff #10 #28)
-#30 := [rewrite]: #29
-#33 := [monotonicity #30]: #32
-#37 := [trans #33 #35]: #36
-#53 := [trans #37 #51]: #52
-#27 := [asserted]: #11
-#54 := [mp #27 #53]: #49
-#62 := [mp~ #54 #61]: #49
-[unit-resolution #62 #174]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_13 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (< 0 (ite (forall (?x1 Int) (or (< ?x1 0) (< 0 ?x1))) (~ 1) 3)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_13.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,180 +0,0 @@
-#2 := false
-#4 := 0::int
-#5 := (:var 0 int)
-#48 := (<= #5 0::int)
-#49 := (not #48)
-#45 := (>= #5 0::int)
-#44 := (not #45)
-#52 := (or #44 #49)
-#55 := (forall (vars (?x1 int)) #52)
-#86 := (not #55)
-#604 := (<= 0::int 0::int)
-#264 := (not #604)
-#269 := (>= 0::int 0::int)
-#605 := (not #269)
-#265 := (or #605 #264)
-#588 := (or #86 #265)
-#584 := (iff #588 #86)
-#311 := (or #86 false)
-#314 := (iff #311 #86)
-#208 := [rewrite]: #314
-#312 := (iff #588 #311)
-#599 := (iff #265 false)
-#598 := (or false false)
-#241 := (iff #598 false)
-#601 := [rewrite]: #241
-#600 := (iff #265 #598)
-#597 := (iff #264 false)
-#1 := true
-#590 := (not true)
-#255 := (iff #590 false)
-#256 := [rewrite]: #255
-#596 := (iff #264 #590)
-#594 := (iff #604 true)
-#595 := [rewrite]: #594
-#591 := [monotonicity #595]: #596
-#235 := [trans #591 #256]: #597
-#592 := (iff #605 false)
-#253 := (iff #605 #590)
-#606 := (iff #269 true)
-#249 := [rewrite]: #606
-#254 := [monotonicity #249]: #253
-#593 := [trans #254 #256]: #592
-#240 := [monotonicity #593 #235]: #600
-#602 := [trans #240 #601]: #599
-#313 := [monotonicity #602]: #312
-#585 := [trans #313 #208]: #584
-#589 := [quant-inst]: #588
-#307 := [mp #589 #585]: #86
-decl z3name!0 :: bool
-#83 := z3name!0
-#12 := 3::int
-#32 := -1::int
-#92 := (ite z3name!0 -1::int 3::int)
-#290 := (= #92 3::int)
-#610 := (not #290)
-#607 := (>= #92 3::int)
-#609 := (not #607)
-#95 := (<= #92 0::int)
-#58 := (ite #55 -1::int 3::int)
-#64 := (<= #58 0::int)
-#96 := (~ #64 #95)
-#93 := (= #58 #92)
-#90 := (~ #55 z3name!0)
-#87 := (or z3name!0 #86)
-#84 := (not z3name!0)
-#85 := (or #84 #55)
-#88 := (and #85 #87)
-#89 := [intro-def]: #88
-#91 := [apply-def #89]: #90
-#94 := [monotonicity #91]: #93
-#97 := [monotonicity #94]: #96
-#10 := 1::int
-#11 := (- 1::int)
-#7 := (< 0::int #5)
-#6 := (< #5 0::int)
-#8 := (or #6 #7)
-#9 := (forall (vars (?x1 int)) #8)
-#13 := (ite #9 #11 3::int)
-#14 := (< 0::int #13)
-#15 := (not #14)
-#77 := (iff #15 #64)
-#35 := (ite #9 -1::int 3::int)
-#38 := (< 0::int #35)
-#41 := (not #38)
-#75 := (iff #41 #64)
-#65 := (not #64)
-#70 := (not #65)
-#73 := (iff #70 #64)
-#74 := [rewrite]: #73
-#71 := (iff #41 #70)
-#68 := (iff #38 #65)
-#61 := (< 0::int #58)
-#66 := (iff #61 #65)
-#67 := [rewrite]: #66
-#62 := (iff #38 #61)
-#59 := (= #35 #58)
-#56 := (iff #9 #55)
-#53 := (iff #8 #52)
-#50 := (iff #7 #49)
-#51 := [rewrite]: #50
-#46 := (iff #6 #44)
-#47 := [rewrite]: #46
-#54 := [monotonicity #47 #51]: #53
-#57 := [quant-intro #54]: #56
-#60 := [monotonicity #57]: #59
-#63 := [monotonicity #60]: #62
-#69 := [trans #63 #67]: #68
-#72 := [monotonicity #69]: #71
-#76 := [trans #72 #74]: #75
-#42 := (iff #15 #41)
-#39 := (iff #14 #38)
-#36 := (= #13 #35)
-#33 := (= #11 -1::int)
-#34 := [rewrite]: #33
-#37 := [monotonicity #34]: #36
-#40 := [monotonicity #37]: #39
-#43 := [monotonicity #40]: #42
-#78 := [trans #43 #76]: #77
-#31 := [asserted]: #15
-#79 := [mp #31 #78]: #64
-#126 := [mp~ #79 #97]: #95
-#266 := (not #95)
-#396 := (or #609 #266)
-#603 := [th-lemma]: #396
-#277 := [unit-resolution #603 #126]: #609
-#278 := [hypothesis]: #290
-#611 := (or #610 #607)
-#612 := [th-lemma]: #611
-#613 := [unit-resolution #612 #278 #277]: false
-#608 := [lemma #613]: #610
-#289 := (or z3name!0 #290)
-#293 := [def-axiom]: #289
-#308 := [unit-resolution #293 #608]: z3name!0
-#129 := (or #55 #84)
-decl ?x1!1 :: int
-#108 := ?x1!1
-#111 := (>= ?x1!1 0::int)
-#112 := (not #111)
-#109 := (<= ?x1!1 0::int)
-#110 := (not #109)
-#132 := (or #110 #112)
-#135 := (not #132)
-#138 := (or z3name!0 #135)
-#141 := (and #129 #138)
-#113 := (or #112 #110)
-#114 := (not #113)
-#119 := (or z3name!0 #114)
-#122 := (and #85 #119)
-#142 := (iff #122 #141)
-#139 := (iff #119 #138)
-#136 := (iff #114 #135)
-#133 := (iff #113 #132)
-#134 := [rewrite]: #133
-#137 := [monotonicity #134]: #136
-#140 := [monotonicity #137]: #139
-#130 := (iff #85 #129)
-#131 := [rewrite]: #130
-#143 := [monotonicity #131 #140]: #142
-#123 := (~ #88 #122)
-#120 := (~ #87 #119)
-#115 := (~ #86 #114)
-#116 := [sk]: #115
-#106 := (~ z3name!0 z3name!0)
-#107 := [refl]: #106
-#121 := [monotonicity #107 #116]: #120
-#104 := (~ #85 #85)
-#102 := (~ #55 #55)
-#100 := (~ #52 #52)
-#101 := [refl]: #100
-#103 := [nnf-pos #101]: #102
-#98 := (~ #84 #84)
-#99 := [refl]: #98
-#105 := [monotonicity #99 #103]: #104
-#124 := [monotonicity #105 #121]: #123
-#125 := [mp~ #89 #124]: #122
-#127 := [mp #125 #143]: #141
-#128 := [and-elim #127]: #129
-#582 := [unit-resolution #128 #308]: #55
-[unit-resolution #582 #307]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_14 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (not (exists (?x1 Int) (?x2 Int) (?x3 Int) (= (+ (* 4 ?x1) (* (~ 6) ?x2)) 1))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_14.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,103 +0,0 @@
-#2 := false
-#104 := -1::int
-decl ?x1!1 :: int
-#86 := ?x1!1
-#106 := -4::int
-#107 := (* -4::int ?x1!1)
-decl ?x2!0 :: int
-#85 := ?x2!0
-#7 := 6::int
-#105 := (* 6::int ?x2!0)
-#108 := (+ #105 #107)
-#168 := (<= #108 -1::int)
-#109 := (= #108 -1::int)
-#12 := 1::int
-#33 := -6::int
-#87 := (* -6::int ?x2!0)
-#4 := 4::int
-#88 := (* 4::int ?x1!1)
-#89 := (+ #88 #87)
-#90 := (= #89 1::int)
-#112 := (iff #90 #109)
-#98 := (+ #87 #88)
-#101 := (= #98 1::int)
-#110 := (iff #101 #109)
-#111 := [rewrite]: #110
-#102 := (iff #90 #101)
-#99 := (= #89 #98)
-#100 := [rewrite]: #99
-#103 := [monotonicity #100]: #102
-#113 := [trans #103 #111]: #112
-#53 := (:var 0 int)
-#54 := (* -6::int #53)
-#9 := (:var 1 int)
-#55 := (* 4::int #9)
-#56 := (+ #55 #54)
-#76 := (= #56 1::int)
-#74 := (exists (vars (?x1 int) (?x2 int)) #76)
-#91 := (~ #74 #90)
-#92 := [sk]: #91
-#8 := (- 6::int)
-#10 := (* #8 #9)
-#5 := (:var 2 int)
-#6 := (* 4::int #5)
-#11 := (+ #6 #10)
-#13 := (= #11 1::int)
-#14 := (exists (vars (?x1 int) (?x2 int) (?x3 int)) #13)
-#15 := (not #14)
-#16 := (not #15)
-#79 := (iff #16 #74)
-#57 := (= 1::int #56)
-#58 := (exists (vars (?x1 int) (?x2 int)) #57)
-#77 := (iff #58 #74)
-#75 := (iff #57 #76)
-#73 := [rewrite]: #75
-#78 := [quant-intro #73]: #77
-#71 := (iff #16 #58)
-#63 := (not #58)
-#66 := (not #63)
-#69 := (iff #66 #58)
-#70 := [rewrite]: #69
-#67 := (iff #16 #66)
-#64 := (iff #15 #63)
-#61 := (iff #14 #58)
-#36 := (* -6::int #9)
-#39 := (+ #6 #36)
-#45 := (= 1::int #39)
-#50 := (exists (vars (?x1 int) (?x2 int) (?x3 int)) #45)
-#59 := (iff #50 #58)
-#60 := [elim-unused]: #59
-#51 := (iff #14 #50)
-#48 := (iff #13 #45)
-#42 := (= #39 1::int)
-#46 := (iff #42 #45)
-#47 := [rewrite]: #46
-#43 := (iff #13 #42)
-#40 := (= #11 #39)
-#37 := (= #10 #36)
-#34 := (= #8 -6::int)
-#35 := [rewrite]: #34
-#38 := [monotonicity #35]: #37
-#41 := [monotonicity #38]: #40
-#44 := [monotonicity #41]: #43
-#49 := [trans #44 #47]: #48
-#52 := [quant-intro #49]: #51
-#62 := [trans #52 #60]: #61
-#65 := [monotonicity #62]: #64
-#68 := [monotonicity #65]: #67
-#72 := [trans #68 #70]: #71
-#80 := [trans #72 #78]: #79
-#32 := [asserted]: #16
-#81 := [mp #32 #80]: #74
-#95 := [mp~ #81 #92]: #90
-#96 := [mp #95 #113]: #109
-#170 := (not #109)
-#171 := (or #170 #168)
-#172 := [th-lemma]: #171
-#173 := [unit-resolution #172 #96]: #168
-#169 := (>= #108 -1::int)
-#174 := (or #170 #169)
-#175 := [th-lemma]: #174
-#176 := [unit-resolution #175 #96]: #169
-[th-lemma #176 #173]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_15 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (exists (?x1 Int) (forall (?x2 Int) (?x3 Int) (implies (and (< 0 ?x2) (< 0 ?x3)) (< 0 (+ ?x2 ?x3))))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_15.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,107 +0,0 @@
-#2 := false
-#4 := 0::int
-decl ?x2!1 :: int
-#83 := ?x2!1
-decl ?x3!0 :: int
-#82 := ?x3!0
-#108 := (+ ?x3!0 ?x2!1)
-#111 := (<= #108 0::int)
-#114 := (not #111)
-#89 := (<= ?x2!1 0::int)
-#90 := (not #89)
-#87 := (<= ?x3!0 0::int)
-#88 := (not #87)
-#102 := (and #88 #90)
-#105 := (not #102)
-#117 := (or #105 #114)
-#120 := (not #117)
-#84 := (+ ?x2!1 ?x3!0)
-#85 := (<= #84 0::int)
-#86 := (not #85)
-#91 := (and #90 #88)
-#92 := (not #91)
-#93 := (or #92 #86)
-#94 := (not #93)
-#121 := (iff #94 #120)
-#118 := (iff #93 #117)
-#115 := (iff #86 #114)
-#112 := (iff #85 #111)
-#109 := (= #84 #108)
-#110 := [rewrite]: #109
-#113 := [monotonicity #110]: #112
-#116 := [monotonicity #113]: #115
-#106 := (iff #92 #105)
-#103 := (iff #91 #102)
-#104 := [rewrite]: #103
-#107 := [monotonicity #104]: #106
-#119 := [monotonicity #107 #116]: #118
-#122 := [monotonicity #119]: #121
-#7 := (:var 0 int)
-#5 := (:var 1 int)
-#10 := (+ #5 #7)
-#63 := (<= #10 0::int)
-#64 := (not #63)
-#53 := (<= #7 0::int)
-#54 := (not #53)
-#49 := (<= #5 0::int)
-#50 := (not #49)
-#57 := (and #50 #54)
-#60 := (not #57)
-#67 := (or #60 #64)
-#70 := (forall (vars (?x2 int) (?x3 int)) #67)
-#73 := (not #70)
-#95 := (~ #73 #94)
-#96 := [sk]: #95
-#11 := (< 0::int #10)
-#8 := (< 0::int #7)
-#6 := (< 0::int #5)
-#9 := (and #6 #8)
-#12 := (implies #9 #11)
-#13 := (forall (vars (?x2 int) (?x3 int)) #12)
-#14 := (exists (vars (?x1 int)) #13)
-#15 := (not #14)
-#76 := (iff #15 #73)
-#32 := (not #9)
-#33 := (or #32 #11)
-#36 := (forall (vars (?x2 int) (?x3 int)) #33)
-#46 := (not #36)
-#74 := (iff #46 #73)
-#71 := (iff #36 #70)
-#68 := (iff #33 #67)
-#65 := (iff #11 #64)
-#66 := [rewrite]: #65
-#61 := (iff #32 #60)
-#58 := (iff #9 #57)
-#55 := (iff #8 #54)
-#56 := [rewrite]: #55
-#51 := (iff #6 #50)
-#52 := [rewrite]: #51
-#59 := [monotonicity #52 #56]: #58
-#62 := [monotonicity #59]: #61
-#69 := [monotonicity #62 #66]: #68
-#72 := [quant-intro #69]: #71
-#75 := [monotonicity #72]: #74
-#47 := (iff #15 #46)
-#44 := (iff #14 #36)
-#39 := (exists (vars (?x1 int)) #36)
-#42 := (iff #39 #36)
-#43 := [elim-unused]: #42
-#40 := (iff #14 #39)
-#37 := (iff #13 #36)
-#34 := (iff #12 #33)
-#35 := [rewrite]: #34
-#38 := [quant-intro #35]: #37
-#41 := [quant-intro #38]: #40
-#45 := [trans #41 #43]: #44
-#48 := [monotonicity #45]: #47
-#77 := [trans #48 #75]: #76
-#31 := [asserted]: #15
-#78 := [mp #31 #77]: #73
-#99 := [mp~ #78 #96]: #94
-#100 := [mp #99 #122]: #120
-#125 := [not-or-elim #100]: #111
-#101 := [not-or-elim #100]: #102
-#124 := [and-elim #101]: #90
-#123 := [and-elim #101]: #88
-[th-lemma #123 #124 #125]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_16 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (exists (?x1 Int) (forall (?x2 Int) (?x3 Real) (implies (and (< 0 ?x2) (< 0.0 ?x3)) (< (~ 1) ?x2)))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_16.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,117 +0,0 @@
-#2 := false
-#4 := 0::int
-decl ?x2!1 :: int
-#91 := ?x2!1
-#98 := (<= ?x2!1 0::int)
-#99 := (not #98)
-#7 := 0::real
-decl ?x3!0 :: real
-#93 := ?x3!0
-#96 := (<= ?x3!0 0::real)
-#97 := (not #96)
-#111 := (and #97 #99)
-#114 := (not #111)
-#33 := -1::int
-#94 := (<= ?x2!1 -1::int)
-#95 := (not #94)
-#120 := (or #95 #114)
-#125 := (not #120)
-#100 := (and #99 #97)
-#101 := (not #100)
-#102 := (or #101 #95)
-#103 := (not #102)
-#126 := (iff #103 #125)
-#123 := (iff #102 #120)
-#117 := (or #114 #95)
-#121 := (iff #117 #120)
-#122 := [rewrite]: #121
-#118 := (iff #102 #117)
-#115 := (iff #101 #114)
-#112 := (iff #100 #111)
-#113 := [rewrite]: #112
-#116 := [monotonicity #113]: #115
-#119 := [monotonicity #116]: #118
-#124 := [trans #119 #122]: #123
-#127 := [monotonicity #124]: #126
-#5 := (:var 1 int)
-#75 := (<= #5 -1::int)
-#76 := (not #75)
-#8 := (:var 0 real)
-#65 := (<= #8 0::real)
-#66 := (not #65)
-#61 := (<= #5 0::int)
-#62 := (not #61)
-#69 := (and #62 #66)
-#72 := (not #69)
-#79 := (or #72 #76)
-#82 := (forall (vars (?x2 int) (?x3 real)) #79)
-#85 := (not #82)
-#104 := (~ #85 #103)
-#105 := [sk]: #104
-#11 := 1::int
-#12 := (- 1::int)
-#13 := (< #12 #5)
-#9 := (< 0::real #8)
-#6 := (< 0::int #5)
-#10 := (and #6 #9)
-#14 := (implies #10 #13)
-#15 := (forall (vars (?x2 int) (?x3 real)) #14)
-#16 := (exists (vars (?x1 int)) #15)
-#17 := (not #16)
-#88 := (iff #17 #85)
-#36 := (< -1::int #5)
-#42 := (not #10)
-#43 := (or #42 #36)
-#48 := (forall (vars (?x2 int) (?x3 real)) #43)
-#58 := (not #48)
-#86 := (iff #58 #85)
-#83 := (iff #48 #82)
-#80 := (iff #43 #79)
-#77 := (iff #36 #76)
-#78 := [rewrite]: #77
-#73 := (iff #42 #72)
-#70 := (iff #10 #69)
-#67 := (iff #9 #66)
-#68 := [rewrite]: #67
-#63 := (iff #6 #62)
-#64 := [rewrite]: #63
-#71 := [monotonicity #64 #68]: #70
-#74 := [monotonicity #71]: #73
-#81 := [monotonicity #74 #78]: #80
-#84 := [quant-intro #81]: #83
-#87 := [monotonicity #84]: #86
-#59 := (iff #17 #58)
-#56 := (iff #16 #48)
-#51 := (exists (vars (?x1 int)) #48)
-#54 := (iff #51 #48)
-#55 := [elim-unused]: #54
-#52 := (iff #16 #51)
-#49 := (iff #15 #48)
-#46 := (iff #14 #43)
-#39 := (implies #10 #36)
-#44 := (iff #39 #43)
-#45 := [rewrite]: #44
-#40 := (iff #14 #39)
-#37 := (iff #13 #36)
-#34 := (= #12 -1::int)
-#35 := [rewrite]: #34
-#38 := [monotonicity #35]: #37
-#41 := [monotonicity #38]: #40
-#47 := [trans #41 #45]: #46
-#50 := [quant-intro #47]: #49
-#53 := [quant-intro #50]: #52
-#57 := [trans #53 #55]: #56
-#60 := [monotonicity #57]: #59
-#89 := [trans #60 #87]: #88
-#32 := [asserted]: #17
-#90 := [mp #32 #89]: #85
-#108 := [mp~ #90 #105]: #103
-#109 := [mp #108 #127]: #125
-#128 := [not-or-elim #109]: #111
-#130 := [and-elim #128]: #99
-#110 := [not-or-elim #109]: #94
-#186 := (or #95 #98)
-#187 := [th-lemma]: #186
-#188 := [unit-resolution #187 #110]: #98
-[unit-resolution #188 #130]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_17 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (exists (?x1 Int) (implies (forall (?x2 Int) (implies (<= ?x1 ?x2) (< 0 ?x2))) (< 0 ?x1))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_17.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,148 +0,0 @@
-#2 := false
-#144 := (not false)
-#7 := 0::int
-#5 := (:var 0 int)
-#52 := (<= #5 0::int)
-#53 := (not #52)
-#147 := (or #53 #144)
-#150 := (not #147)
-#153 := (forall (vars (?x1 int)) #150)
-#180 := (iff #153 false)
-#175 := (forall (vars (?x1 int)) false)
-#178 := (iff #175 false)
-#179 := [elim-unused]: #178
-#176 := (iff #153 #175)
-#173 := (iff #150 false)
-#1 := true
-#168 := (not true)
-#171 := (iff #168 false)
-#172 := [rewrite]: #171
-#169 := (iff #150 #168)
-#166 := (iff #147 true)
-#161 := (or #53 true)
-#164 := (iff #161 true)
-#165 := [rewrite]: #164
-#162 := (iff #147 #161)
-#159 := (iff #144 true)
-#160 := [rewrite]: #159
-#163 := [monotonicity #160]: #162
-#167 := [trans #163 #165]: #166
-#170 := [monotonicity #167]: #169
-#174 := [trans #170 #172]: #173
-#177 := [quant-intro #174]: #176
-#181 := [trans #177 #179]: #180
-#56 := -1::int
-#57 := (* -1::int #5)
-#4 := (:var 1 int)
-#58 := (+ #4 #57)
-#59 := (<= #58 0::int)
-#62 := (not #59)
-#68 := (or #53 #62)
-#73 := (forall (vars (?x2 int)) #68)
-#76 := (not #73)
-#79 := (or #53 #76)
-#105 := (not #79)
-#123 := (forall (vars (?x1 int)) #105)
-#156 := (iff #123 #153)
-#127 := (forall (vars (?x2 int)) #53)
-#130 := (not #127)
-#133 := (or #53 #130)
-#136 := (not #133)
-#139 := (forall (vars (?x1 int)) #136)
-#154 := (iff #139 #153)
-#155 := [rewrite]: #154
-#140 := (iff #123 #139)
-#141 := [rewrite]: #140
-#157 := [trans #141 #155]: #156
-#116 := (and #52 #73)
-#119 := (forall (vars (?x1 int)) #116)
-#124 := (iff #119 #123)
-#113 := (iff #116 #105)
-#122 := [rewrite]: #113
-#125 := [quant-intro #122]: #124
-#94 := (not #53)
-#104 := (and #94 #73)
-#108 := (forall (vars (?x1 int)) #104)
-#120 := (iff #108 #119)
-#117 := (iff #104 #116)
-#114 := (iff #94 #52)
-#115 := [rewrite]: #114
-#118 := [monotonicity #115]: #117
-#121 := [quant-intro #118]: #120
-#82 := (exists (vars (?x1 int)) #79)
-#85 := (not #82)
-#109 := (~ #85 #108)
-#106 := (~ #105 #104)
-#101 := (not #76)
-#102 := (~ #101 #73)
-#99 := (~ #73 #73)
-#97 := (~ #68 #68)
-#98 := [refl]: #97
-#100 := [nnf-pos #98]: #99
-#103 := [nnf-neg #100]: #102
-#95 := (~ #94 #94)
-#96 := [refl]: #95
-#107 := [nnf-neg #96 #103]: #106
-#110 := [nnf-neg #107]: #109
-#8 := (< 0::int #5)
-#6 := (<= #4 #5)
-#9 := (implies #6 #8)
-#10 := (forall (vars (?x2 int)) #9)
-#11 := (implies #10 #8)
-#12 := (exists (vars (?x1 int)) #11)
-#13 := (not #12)
-#88 := (iff #13 #85)
-#30 := (not #6)
-#31 := (or #30 #8)
-#34 := (forall (vars (?x2 int)) #31)
-#40 := (not #34)
-#41 := (or #8 #40)
-#46 := (exists (vars (?x1 int)) #41)
-#49 := (not #46)
-#86 := (iff #49 #85)
-#83 := (iff #46 #82)
-#80 := (iff #41 #79)
-#77 := (iff #40 #76)
-#74 := (iff #34 #73)
-#71 := (iff #31 #68)
-#65 := (or #62 #53)
-#69 := (iff #65 #68)
-#70 := [rewrite]: #69
-#66 := (iff #31 #65)
-#54 := (iff #8 #53)
-#55 := [rewrite]: #54
-#63 := (iff #30 #62)
-#60 := (iff #6 #59)
-#61 := [rewrite]: #60
-#64 := [monotonicity #61]: #63
-#67 := [monotonicity #64 #55]: #66
-#72 := [trans #67 #70]: #71
-#75 := [quant-intro #72]: #74
-#78 := [monotonicity #75]: #77
-#81 := [monotonicity #55 #78]: #80
-#84 := [quant-intro #81]: #83
-#87 := [monotonicity #84]: #86
-#50 := (iff #13 #49)
-#47 := (iff #12 #46)
-#44 := (iff #11 #41)
-#37 := (implies #34 #8)
-#42 := (iff #37 #41)
-#43 := [rewrite]: #42
-#38 := (iff #11 #37)
-#35 := (iff #10 #34)
-#32 := (iff #9 #31)
-#33 := [rewrite]: #32
-#36 := [quant-intro #33]: #35
-#39 := [monotonicity #36]: #38
-#45 := [trans #39 #43]: #44
-#48 := [quant-intro #45]: #47
-#51 := [monotonicity #48]: #50
-#89 := [trans #51 #87]: #88
-#29 := [asserted]: #13
-#90 := [mp #29 #89]: #85
-#111 := [mp~ #90 #110]: #108
-#112 := [mp #111 #121]: #119
-#126 := [mp #112 #125]: #123
-#158 := [mp #126 #157]: #153
-[mp #158 #181]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_18 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,7 +0,0 @@
-(benchmark Isabelle
-:extrafuns (
- (uf_1 Int)
- )
-:assumption (not (forall (?x1 Int) (implies (< ?x1 uf_1) (< (* 2 ?x1) (* 2 uf_1))) :pat { ?x1 }))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_arith_quant_18.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,112 +0,0 @@
-#2 := false
-#43 := 0::int
-decl ?x1!0 :: int
-#78 := ?x1!0
-#56 := -2::int
-#113 := (* -2::int ?x1!0)
-decl uf_1 :: int
-#6 := uf_1
-#8 := 2::int
-#10 := (* 2::int uf_1)
-#114 := (+ #10 #113)
-#115 := (<= #114 0::int)
-#120 := (not #115)
-#41 := -1::int
-#100 := (* -1::int ?x1!0)
-#101 := (+ uf_1 #100)
-#102 := (<= #101 0::int)
-#123 := (or #102 #120)
-#126 := (not #123)
-#59 := (* -2::int uf_1)
-#79 := (* 2::int ?x1!0)
-#80 := (+ #79 #59)
-#81 := (>= #80 0::int)
-#82 := (not #81)
-#45 := (* -1::int uf_1)
-#83 := (+ ?x1!0 #45)
-#84 := (>= #83 0::int)
-#85 := (or #84 #82)
-#86 := (not #85)
-#127 := (iff #86 #126)
-#124 := (iff #85 #123)
-#121 := (iff #82 #120)
-#118 := (iff #81 #115)
-#107 := (+ #59 #79)
-#110 := (>= #107 0::int)
-#116 := (iff #110 #115)
-#117 := [rewrite]: #116
-#111 := (iff #81 #110)
-#108 := (= #80 #107)
-#109 := [rewrite]: #108
-#112 := [monotonicity #109]: #111
-#119 := [trans #112 #117]: #118
-#122 := [monotonicity #119]: #121
-#105 := (iff #84 #102)
-#94 := (+ #45 ?x1!0)
-#97 := (>= #94 0::int)
-#103 := (iff #97 #102)
-#104 := [rewrite]: #103
-#98 := (iff #84 #97)
-#95 := (= #83 #94)
-#96 := [rewrite]: #95
-#99 := [monotonicity #96]: #98
-#106 := [trans #99 #104]: #105
-#125 := [monotonicity #106 #122]: #124
-#128 := [monotonicity #125]: #127
-#4 := (:var 0 int)
-#5 := (pattern #4)
-#9 := (* 2::int #4)
-#60 := (+ #9 #59)
-#58 := (>= #60 0::int)
-#57 := (not #58)
-#46 := (+ #4 #45)
-#44 := (>= #46 0::int)
-#63 := (or #44 #57)
-#66 := (forall (vars (?x1 int)) (:pat #5) #63)
-#69 := (not #66)
-#87 := (~ #69 #86)
-#88 := [sk]: #87
-#11 := (< #9 #10)
-#7 := (< #4 uf_1)
-#12 := (implies #7 #11)
-#13 := (forall (vars (?x1 int)) (:pat #5) #12)
-#14 := (not #13)
-#72 := (iff #14 #69)
-#31 := (not #7)
-#32 := (or #31 #11)
-#35 := (forall (vars (?x1 int)) (:pat #5) #32)
-#38 := (not #35)
-#70 := (iff #38 #69)
-#67 := (iff #35 #66)
-#64 := (iff #32 #63)
-#61 := (iff #11 #57)
-#62 := [rewrite]: #61
-#54 := (iff #31 #44)
-#42 := (not #44)
-#49 := (not #42)
-#52 := (iff #49 #44)
-#53 := [rewrite]: #52
-#50 := (iff #31 #49)
-#47 := (iff #7 #42)
-#48 := [rewrite]: #47
-#51 := [monotonicity #48]: #50
-#55 := [trans #51 #53]: #54
-#65 := [monotonicity #55 #62]: #64
-#68 := [quant-intro #65]: #67
-#71 := [monotonicity #68]: #70
-#39 := (iff #14 #38)
-#36 := (iff #13 #35)
-#33 := (iff #12 #32)
-#34 := [rewrite]: #33
-#37 := [quant-intro #34]: #36
-#40 := [monotonicity #37]: #39
-#73 := [trans #40 #71]: #72
-#30 := [asserted]: #14
-#74 := [mp #30 #73]: #69
-#91 := [mp~ #74 #88]: #86
-#92 := [mp #91 #128]: #126
-#130 := [not-or-elim #92]: #115
-#93 := (not #102)
-#129 := [not-or-elim #92]: #93
-[th-lemma #129 #130]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_bv_01 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,12 +0,0 @@
-(benchmark Isabelle
-:extrafuns (
- (uf_1 BitVec[2] Int)
- )
-:assumption (= (uf_1 bv0[2]) 0)
-:assumption (= (uf_1 bv1[2]) 1)
-:assumption (= (uf_1 bv2[2]) 2)
-:assumption (= (uf_1 bv3[2]) 3)
-:assumption (forall (?x1 BitVec[2]) (< 0 (uf_1 ?x1)))
-:assumption (not (forall (?x2 Int) (implies (< ?x2 0) (forall (?x3 BitVec[2]) (< ?x2 (uf_1 ?x3))))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_bv_01.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,43 +0,0 @@
-#2 := false
-#6 := 0::int
-decl uf_1 :: (-> bv[2] int)
-#4 := bv[0:2]
-#5 := (uf_1 bv[0:2])
-#225 := (<= #5 0::int)
-#311 := (not #225)
-#20 := (:var 0 bv[2])
-#21 := (uf_1 #20)
-#640 := (pattern #21)
-#54 := (<= #21 0::int)
-#55 := (not #54)
-#641 := (forall (vars (?x1 bv[2])) (:pat #640) #55)
-#58 := (forall (vars (?x1 bv[2])) #55)
-#644 := (iff #58 #641)
-#642 := (iff #55 #55)
-#643 := [refl]: #642
-#645 := [quant-intro #643]: #644
-#113 := (~ #58 #58)
-#115 := (~ #55 #55)
-#116 := [refl]: #115
-#114 := [nnf-pos #116]: #113
-#22 := (< 0::int #21)
-#23 := (forall (vars (?x1 bv[2])) #22)
-#59 := (iff #23 #58)
-#56 := (iff #22 #55)
-#57 := [rewrite]: #56
-#60 := [quant-intro #57]: #59
-#51 := [asserted]: #23
-#61 := [mp #51 #60]: #58
-#111 := [mp~ #61 #114]: #58
-#646 := [mp #111 #645]: #641
-#227 := (not #641)
-#313 := (or #227 #311)
-#304 := [quant-inst]: #313
-#635 := [unit-resolution #304 #646]: #311
-#7 := (= #5 0::int)
-#47 := [asserted]: #7
-#638 := (not #7)
-#633 := (or #638 #225)
-#639 := [th-lemma]: #633
-[unit-resolution #639 #47 #635]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_bv_02 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,12 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T2 T1)
-:extrafuns (
- (uf_4 T1)
- (uf_2 BitVec[4] BitVec[4] T1)
- (uf_1 T1 T2)
- (uf_3 BitVec[4])
- )
-:assumption (not (= (uf_1 (uf_2 bv0[4] uf_3)) (uf_1 uf_4)))
-:assumption (forall (?x1 BitVec[4]) (?x2 BitVec[4]) (iff (= (uf_2 ?x1 ?x2) uf_4) (bvule ?x1 ?x2)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_bv_02.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_bv_arith_01 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (= bv27[4] (bvneg bv5[4])))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_bv_arith_01.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_bv_arith_02 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (= bv27[4] bv11[4]))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_bv_arith_02.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_bv_arith_03 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (bvult bv23[8] bv27[8]))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_bv_arith_03.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_bv_arith_04 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (= (bvadd bv27[5] bv11[5]) bv6[5]))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_bv_arith_04.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_bv_arith_05 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (= (bvmul bv7[8] bv3[8]) bv21[8]))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_bv_arith_05.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_bv_arith_06 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (= (bvsub bv11[8] bv27[8]) (bvneg bv16[8])))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_bv_arith_06.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_bv_arith_07 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (= (bvneg (bvneg bv11[5])) bv11[5]))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_bv_arith_07.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_bv_arith_08 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (= (bvadd (bvneg bv40[7]) bv1[7]) (bvneg bv39[7])))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_bv_arith_08.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_bv_arith_09 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,9 +0,0 @@
-(benchmark Isabelle
-:extrafuns (
- (uf_1 BitVec[32])
- (uf_2 BitVec[32])
- (uf_3 BitVec[32])
- )
-:assumption (not (= (bvsub (bvadd (bvadd uf_1 (bvmul bv2[32] uf_2)) uf_3) uf_2) (bvadd (bvadd uf_2 uf_3) uf_1)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_bv_arith_09.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_bv_arith_10 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,8 +0,0 @@
-(benchmark Isabelle
-:extrafuns (
- (uf_1 BitVec[4])
- )
-:assumption (= uf_1 bv5[4])
-:assumption (not (= (bvmul bv4[4] uf_1) bv4[4]))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_bv_arith_10.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_bv_bit_01 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (= (bvand bv6[32] bv5[32]) bv4[32]))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_bv_bit_01.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_bv_bit_02 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (= (bvor bv6[8] bv3[8]) bv7[8]))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_bv_bit_02.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_bv_bit_03 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (= (bvxor bv240[8] bv255[8]) bv15[8]))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_bv_bit_03.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_bv_bit_04 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (= (bvnot bv240[16]) bv65295[16]))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_bv_bit_04.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_bv_bit_05 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (= (concat bv27[4] bv27[8]) bv2843[12]))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_bv_bit_05.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_bv_bit_06 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (= (concat bv3[4] bv15[6]) bv207[10]))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_bv_bit_06.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_bv_bit_07 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,12 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T1)
-:extrafuns (
- (uf_1 Int T1)
- (uf_2 T1 Int)
- )
-:assumption (forall (?x1 T1) (= (uf_1 (uf_2 ?x1)) ?x1))
-:assumption (forall (?x2 Int) (implies (<= 0 ?x2) (= (uf_2 (uf_1 ?x2)) ?x2)))
-:assumption (forall (?x3 Int) (implies (< ?x3 0) (= (uf_2 (uf_1 ?x3)) 0)))
-:assumption (not (= (extract[2:1] bv22[4]) bv3[2]))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_bv_bit_07.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_bv_bit_08 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (= (zero_extend[6] bv10[4]) bv10[10]))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_bv_bit_08.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_bv_bit_09 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (= (sign_extend[2] bv10[4]) bv58[6]))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_bv_bit_09.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_bv_bit_10 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (= (bvlshr bv19[8] bv2[8]) bv4[8]))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_bv_bit_10.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_bv_bit_11 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (= (bvashr bv19[8] bv2[8]) bv4[8]))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_bv_bit_11.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_bv_bit_12 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,12 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T1)
-:extrafuns (
- (uf_1 Int T1)
- (uf_2 T1 Int)
- )
-:assumption (forall (?x1 T1) (= (uf_1 (uf_2 ?x1)) ?x1))
-:assumption (forall (?x2 Int) (implies (<= 0 ?x2) (= (uf_2 (uf_1 ?x2)) ?x2)))
-:assumption (forall (?x3 Int) (implies (< ?x3 0) (= (uf_2 (uf_1 ?x3)) 0)))
-:assumption (not (= (rotate_right[2] bv6[4]) bv9[4]))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_bv_bit_12.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_bv_bit_13 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,12 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T1)
-:extrafuns (
- (uf_1 Int T1)
- (uf_2 T1 Int)
- )
-:assumption (forall (?x1 T1) (= (uf_1 (uf_2 ?x1)) ?x1))
-:assumption (forall (?x2 Int) (implies (<= 0 ?x2) (= (uf_2 (uf_1 ?x2)) ?x2)))
-:assumption (forall (?x3 Int) (implies (< ?x3 0) (= (uf_2 (uf_1 ?x3)) 0)))
-:assumption (not (= (rotate_left[1] bv14[4]) bv13[4]))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_bv_bit_13.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_bv_bit_14 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,7 +0,0 @@
-(benchmark Isabelle
-:extrafuns (
- (uf_1 BitVec[16])
- )
-:assumption (not (= (bvor (bvand uf_1 bv65280[16]) (bvand uf_1 bv255[16])) uf_1))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_bv_bit_14.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_bv_bit_15 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,8 +0,0 @@
-(benchmark Isabelle
-:extrafuns (
- (uf_1 BitVec[16])
- )
-:assumption (bvult uf_1 bv256[16])
-:assumption (not (= (bvand uf_1 bv255[16]) uf_1))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_bv_bit_15.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_fol_01 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,8 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T1)
-:extrapreds (
- (up_1 Int)
- )
-:assumption (not (forall (?x1 Int) (implies (up_1 ?x1) (forall (?x2 Int) (or (up_1 ?x1) (up_1 ?x2))))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_fol_01.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,78 +0,0 @@
-#2 := false
-decl up_1 :: (-> int bool)
-decl ?x1!0 :: int
-#54 := ?x1!0
-#55 := (up_1 ?x1!0)
-#58 := (not #55)
-decl ?x2!1 :: int
-#66 := ?x2!1
-#67 := (up_1 ?x2!1)
-#85 := (or #55 #67)
-#88 := (not #85)
-#91 := (and #55 #88)
-#68 := (or #67 #55)
-#69 := (not #68)
-#63 := (not #58)
-#75 := (and #63 #69)
-#92 := (iff #75 #91)
-#89 := (iff #69 #88)
-#86 := (iff #68 #85)
-#87 := [rewrite]: #86
-#90 := [monotonicity #87]: #89
-#83 := (iff #63 #55)
-#84 := [rewrite]: #83
-#93 := [monotonicity #84 #90]: #92
-#6 := (:var 1 int)
-#7 := (up_1 #6)
-#4 := (:var 0 int)
-#5 := (up_1 #4)
-#29 := (or #5 #7)
-#32 := (forall (vars (?x2 int)) #29)
-#38 := (not #5)
-#39 := (or #38 #32)
-#44 := (forall (vars (?x1 int)) #39)
-#47 := (not #44)
-#78 := (~ #47 #75)
-#56 := (or #5 #55)
-#57 := (forall (vars (?x2 int)) #56)
-#59 := (or #58 #57)
-#60 := (not #59)
-#76 := (~ #60 #75)
-#70 := (not #57)
-#71 := (~ #70 #69)
-#72 := [sk]: #71
-#64 := (~ #63 #63)
-#65 := [refl]: #64
-#77 := [nnf-neg #65 #72]: #76
-#61 := (~ #47 #60)
-#62 := [sk]: #61
-#79 := [trans #62 #77]: #78
-#8 := (or #7 #5)
-#9 := (forall (vars (?x2 int)) #8)
-#10 := (implies #5 #9)
-#11 := (forall (vars (?x1 int)) #10)
-#12 := (not #11)
-#48 := (iff #12 #47)
-#45 := (iff #11 #44)
-#42 := (iff #10 #39)
-#35 := (implies #5 #32)
-#40 := (iff #35 #39)
-#41 := [rewrite]: #40
-#36 := (iff #10 #35)
-#33 := (iff #9 #32)
-#30 := (iff #8 #29)
-#31 := [rewrite]: #30
-#34 := [quant-intro #31]: #33
-#37 := [monotonicity #34]: #36
-#43 := [trans #37 #41]: #42
-#46 := [quant-intro #43]: #45
-#49 := [monotonicity #46]: #48
-#28 := [asserted]: #12
-#52 := [mp #28 #49]: #47
-#80 := [mp~ #52 #79]: #75
-#81 := [mp #80 #93]: #91
-#94 := [and-elim #81]: #88
-#95 := [not-or-elim #94]: #58
-#82 := [and-elim #81]: #55
-[unit-resolution #82 #95]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_fol_02 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,14 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T2 T1)
-:extrafuns (
- (uf_2 T1)
- (uf_4 T2)
- (uf_3 T1)
- )
-:extrapreds (
- (up_1 T1 T2)
- )
-:assumption (forall (?x1 T1) (?x2 T2) (iff (up_1 ?x1 ?x2) (= ?x1 uf_2)))
-:assumption (not (iff (exists (?x3 T2) (up_1 uf_3 ?x3)) (up_1 uf_3 uf_4)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_fol_02.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,135 +0,0 @@
-#2 := false
-decl up_1 :: (-> T1 T2 bool)
-#5 := (:var 0 T2)
-decl uf_3 :: T1
-#11 := uf_3
-#12 := (up_1 uf_3 #5)
-#560 := (pattern #12)
-#57 := (not #12)
-#561 := (forall (vars (?x3 T2)) (:pat #560) #57)
-decl uf_4 :: T2
-#14 := uf_4
-#15 := (up_1 uf_3 uf_4)
-decl uf_2 :: T1
-#7 := uf_2
-#136 := (= uf_2 uf_3)
-#543 := (iff #15 #136)
-#4 := (:var 1 T1)
-#6 := (up_1 #4 #5)
-#553 := (pattern #6)
-#8 := (= #4 uf_2)
-#9 := (iff #6 #8)
-#554 := (forall (vars (?x1 T1) (?x2 T2)) (:pat #553) #9)
-#10 := (forall (vars (?x1 T1) (?x2 T2)) #9)
-#557 := (iff #10 #554)
-#555 := (iff #9 #9)
-#556 := [refl]: #555
-#558 := [quant-intro #556]: #557
-#47 := (~ #10 #10)
-#45 := (~ #9 #9)
-#46 := [refl]: #45
-#48 := [nnf-pos #46]: #47
-#33 := [asserted]: #10
-#49 := [mp~ #33 #48]: #10
-#559 := [mp #49 #558]: #554
-#227 := (not #554)
-#185 := (or #227 #543)
-#135 := (= uf_3 uf_2)
-#205 := (iff #15 #135)
-#528 := (or #227 #205)
-#190 := (iff #528 #185)
-#192 := (iff #185 #185)
-#530 := [rewrite]: #192
-#201 := (iff #205 #543)
-#223 := (iff #135 #136)
-#137 := [rewrite]: #223
-#544 := [monotonicity #137]: #201
-#191 := [monotonicity #544]: #190
-#531 := [trans #191 #530]: #190
-#189 := [quant-inst]: #528
-#532 := [mp #189 #531]: #185
-#539 := [unit-resolution #532 #559]: #543
-decl ?x3!0 :: T2
-#50 := ?x3!0
-#51 := (up_1 uf_3 ?x3!0)
-#224 := (iff #51 #136)
-#155 := (or #227 #224)
-#222 := (iff #51 #135)
-#228 := (or #227 #222)
-#229 := (iff #228 #155)
-#545 := (iff #155 #155)
-#547 := [rewrite]: #545
-#215 := (iff #222 #224)
-#226 := [monotonicity #137]: #215
-#208 := [monotonicity #226]: #229
-#202 := [trans #208 #547]: #229
-#225 := [quant-inst]: #228
-#334 := [mp #225 #202]: #155
-#537 := [unit-resolution #334 #559]: #224
-#541 := (not #224)
-#527 := (or #541 #136)
-#63 := (not #15)
-#540 := [hypothesis]: #63
-#68 := (or #15 #51)
-#60 := (forall (vars (?x3 T2)) #57)
-#69 := (or #63 #60)
-#76 := (and #68 #69)
-#70 := (and #69 #68)
-#77 := (iff #70 #76)
-#78 := [rewrite]: #77
-#13 := (exists (vars (?x3 T2)) #12)
-#35 := (not #13)
-#36 := (iff #15 #35)
-#71 := (~ #36 #70)
-#61 := (~ #35 #60)
-#58 := (~ #57 #57)
-#59 := [refl]: #58
-#62 := [nnf-neg #59]: #61
-#54 := (not #35)
-#55 := (~ #54 #51)
-#42 := (~ #13 #51)
-#39 := [sk]: #42
-#56 := [nnf-neg #39]: #55
-#66 := (~ #15 #15)
-#67 := [refl]: #66
-#64 := (~ #63 #63)
-#65 := [refl]: #64
-#72 := [nnf-pos #65 #67 #56 #62]: #71
-#16 := (iff #13 #15)
-#17 := (not #16)
-#37 := (iff #17 #36)
-#38 := [rewrite]: #37
-#34 := [asserted]: #17
-#41 := [mp #34 #38]: #36
-#73 := [mp~ #41 #72]: #70
-#74 := [mp #73 #78]: #76
-#75 := [and-elim #74]: #68
-#526 := [unit-resolution #75 #540]: #51
-#549 := (not #51)
-#550 := (or #541 #549 #136)
-#551 := [def-axiom]: #550
-#233 := [unit-resolution #551 #526]: #527
-#249 := [unit-resolution #233 #537]: #136
-#213 := (not #136)
-#533 := (not #543)
-#250 := (or #533 #213)
-#534 := (or #533 #15 #213)
-#529 := [def-axiom]: #534
-#251 := [unit-resolution #529 #540]: #250
-#237 := [unit-resolution #251 #249 #539]: false
-#252 := [lemma #237]: #15
-#566 := (or #63 #561)
-#567 := (iff #69 #566)
-#564 := (iff #60 #561)
-#562 := (iff #57 #57)
-#563 := [refl]: #562
-#565 := [quant-intro #563]: #564
-#568 := [monotonicity #565]: #567
-#79 := [and-elim #74]: #69
-#569 := [mp #79 #568]: #566
-#535 := [unit-resolution #569 #252]: #561
-#536 := (not #561)
-#538 := (or #536 #63)
-#176 := [quant-inst]: #538
-[unit-resolution #176 #252 #535]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_fol_03 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,15 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T2 T1)
-:extrafuns (
- (uf_2 T1)
- (uf_3 T2)
- (uf_4 T1)
- )
-:extrapreds (
- (up_1 T1 T2)
- )
-:assumption (forall (?x1 T1) (?x2 T2) (iff (up_1 ?x1 ?x2) (= ?x1 uf_2)))
-:assumption (iff (forall (?x3 T1) (exists (?x4 T2) (up_1 ?x3 ?x4))) (forall (?x5 T1) (up_1 ?x5 uf_3)))
-:assumption (not (iff (exists (?x6 T2) (up_1 uf_4 ?x6)) (up_1 uf_4 uf_3)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_fol_03.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,135 +0,0 @@
-#2 := false
-decl up_1 :: (-> T1 T2 bool)
-#5 := (:var 0 T2)
-decl uf_4 :: T1
-#18 := uf_4
-#19 := (up_1 uf_4 #5)
-#635 := (pattern #19)
-#116 := (not #19)
-#636 := (forall (vars (?x6 T2)) (:pat #635) #116)
-decl uf_3 :: T2
-#14 := uf_3
-#21 := (up_1 uf_4 uf_3)
-decl uf_2 :: T1
-#7 := uf_2
-#195 := (= uf_2 uf_4)
-#602 := (iff #21 #195)
-#4 := (:var 1 T1)
-#6 := (up_1 #4 #5)
-#612 := (pattern #6)
-#8 := (= #4 uf_2)
-#9 := (iff #6 #8)
-#613 := (forall (vars (?x1 T1) (?x2 T2)) (:pat #612) #9)
-#10 := (forall (vars (?x1 T1) (?x2 T2)) #9)
-#616 := (iff #10 #613)
-#614 := (iff #9 #9)
-#615 := [refl]: #614
-#617 := [quant-intro #615]: #616
-#56 := (~ #10 #10)
-#54 := (~ #9 #9)
-#55 := [refl]: #54
-#57 := [nnf-pos #55]: #56
-#39 := [asserted]: #10
-#58 := [mp~ #39 #57]: #10
-#618 := [mp #58 #617]: #613
-#286 := (not #613)
-#244 := (or #286 #602)
-#194 := (= uf_4 uf_2)
-#264 := (iff #21 #194)
-#587 := (or #286 #264)
-#249 := (iff #587 #244)
-#251 := (iff #244 #244)
-#589 := [rewrite]: #251
-#260 := (iff #264 #602)
-#282 := (iff #194 #195)
-#196 := [rewrite]: #282
-#603 := [monotonicity #196]: #260
-#250 := [monotonicity #603]: #249
-#590 := [trans #250 #589]: #249
-#248 := [quant-inst]: #587
-#591 := [mp #248 #590]: #244
-#598 := [unit-resolution #591 #618]: #602
-decl ?x6!3 :: T2
-#63 := ?x6!3
-#64 := (up_1 uf_4 ?x6!3)
-#283 := (iff #64 #195)
-#214 := (or #286 #283)
-#281 := (iff #64 #194)
-#287 := (or #286 #281)
-#288 := (iff #287 #214)
-#604 := (iff #214 #214)
-#606 := [rewrite]: #604
-#274 := (iff #281 #283)
-#285 := [monotonicity #196]: #274
-#267 := [monotonicity #285]: #288
-#261 := [trans #267 #606]: #288
-#284 := [quant-inst]: #287
-#393 := [mp #284 #261]: #214
-#596 := [unit-resolution #393 #618]: #283
-#600 := (not #283)
-#586 := (or #600 #195)
-#122 := (not #21)
-#599 := [hypothesis]: #122
-#127 := (or #21 #64)
-#119 := (forall (vars (?x6 T2)) #116)
-#128 := (or #122 #119)
-#135 := (and #127 #128)
-#129 := (and #128 #127)
-#136 := (iff #129 #135)
-#137 := [rewrite]: #136
-#20 := (exists (vars (?x6 T2)) #19)
-#42 := (not #20)
-#43 := (iff #21 #42)
-#130 := (~ #43 #129)
-#120 := (~ #42 #119)
-#117 := (~ #116 #116)
-#118 := [refl]: #117
-#121 := [nnf-neg #118]: #120
-#113 := (not #42)
-#114 := (~ #113 #64)
-#88 := (~ #20 #64)
-#89 := [sk]: #88
-#115 := [nnf-neg #89]: #114
-#125 := (~ #21 #21)
-#126 := [refl]: #125
-#123 := (~ #122 #122)
-#124 := [refl]: #123
-#131 := [nnf-pos #124 #126 #115 #121]: #130
-#22 := (iff #20 #21)
-#23 := (not #22)
-#44 := (iff #23 #43)
-#45 := [rewrite]: #44
-#41 := [asserted]: #23
-#48 := [mp #41 #45]: #43
-#132 := [mp~ #48 #131]: #129
-#133 := [mp #132 #137]: #135
-#134 := [and-elim #133]: #127
-#585 := [unit-resolution #134 #599]: #64
-#608 := (not #64)
-#609 := (or #600 #608 #195)
-#610 := [def-axiom]: #609
-#292 := [unit-resolution #610 #585]: #586
-#308 := [unit-resolution #292 #596]: #195
-#272 := (not #195)
-#592 := (not #602)
-#309 := (or #592 #272)
-#593 := (or #592 #21 #272)
-#588 := [def-axiom]: #593
-#310 := [unit-resolution #588 #599]: #309
-#296 := [unit-resolution #310 #308 #598]: false
-#311 := [lemma #296]: #21
-#641 := (or #122 #636)
-#642 := (iff #128 #641)
-#639 := (iff #119 #636)
-#637 := (iff #116 #116)
-#638 := [refl]: #637
-#640 := [quant-intro #638]: #639
-#643 := [monotonicity #640]: #642
-#138 := [and-elim #133]: #128
-#644 := [mp #138 #643]: #641
-#594 := [unit-resolution #644 #311]: #636
-#595 := (not #636)
-#597 := (or #595 #122)
-#235 := [quant-inst]: #597
-[unit-resolution #235 #311 #594]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_fol_04 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,13 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T1 T2)
-:extrafuns (
- (uf_2 T1)
- (uf_3 T1)
- )
-:extrapreds (
- (up_1 T1)
- )
-:assumption (if_then_else (up_1 uf_2) (not (exists (?x1 T1) (up_1 ?x1))) (forall (?x2 T1) (not (up_1 ?x2))))
-:assumption (not (implies (up_1 uf_2) (up_1 uf_3)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_fol_04.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,56 +0,0 @@
-#2 := false
-decl up_1 :: (-> T1 bool)
-decl uf_2 :: T1
-#4 := uf_2
-#5 := (up_1 uf_2)
-decl uf_3 :: T1
-#13 := uf_3
-#14 := (up_1 uf_3)
-#34 := (not #5)
-#35 := (or #34 #14)
-#38 := (not #35)
-#15 := (implies #5 #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]: #5
-#6 := (:var 0 T1)
-#7 := (up_1 #6)
-#536 := (pattern #7)
-#10 := (not #7)
-#537 := (forall (vars (?x2 T1)) (:pat #536) #10)
-#11 := (forall (vars (?x2 T1)) #10)
-#540 := (iff #11 #537)
-#538 := (iff #10 #10)
-#539 := [refl]: #538
-#541 := [quant-intro #539]: #540
-#8 := (exists (vars (?x1 T1)) #7)
-#9 := (not #8)
-#45 := (~ #9 #11)
-#50 := (~ #10 #10)
-#51 := [refl]: #50
-#59 := [nnf-neg #51]: #45
-#12 := (ite #5 #9 #11)
-#57 := (iff #12 #9)
-#1 := true
-#52 := (ite true #9 #11)
-#55 := (iff #52 #9)
-#56 := [rewrite]: #55
-#53 := (iff #12 #52)
-#48 := (iff #5 true)
-#49 := [iff-true #41]: #48
-#54 := [monotonicity #49]: #53
-#58 := [trans #54 #56]: #57
-#32 := [asserted]: #12
-#47 := [mp #32 #58]: #9
-#60 := [mp~ #47 #59]: #11
-#542 := [mp #60 #541]: #537
-#119 := (not #537)
-#206 := (or #119 #34)
-#120 := [quant-inst]: #206
-[unit-resolution #120 #542 #41]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_hol_01 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,18 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T2 T3 T1)
-:extrafuns (
- (uf_2 T1 T2 T3 T1)
- (uf_1 T1 T2 T3)
- (uf_6 T1)
- (uf_3 T2)
- (uf_4 T2)
- (uf_5 T2)
- (uf_7 T3)
- (uf_8 T3)
- )
-:assumption (forall (?x1 T1) (?x2 T2) (?x3 T3) (?x4 T2) (= (uf_1 (uf_2 ?x1 ?x2 ?x3) ?x4) (ite (= ?x4 ?x2) ?x3 (uf_1 ?x1 ?x4))))
-:assumption (forall (?x5 T1) (?x6 T2) (?x7 T3) (= (uf_1 (uf_2 ?x5 ?x6 ?x7) ?x6) ?x7))
-:assumption (and (not (= uf_3 uf_4)) (not (= uf_3 uf_5)))
-:assumption (not (= (uf_1 (uf_2 (uf_2 uf_6 uf_4 uf_7) uf_5 uf_8) uf_3) (uf_1 uf_6 uf_3)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_hol_01.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,181 +0,0 @@
-#2 := false
-decl uf_1 :: (-> T1 T2 T3)
-decl uf_3 :: T2
-#22 := uf_3
-decl uf_6 :: T1
-#30 := uf_6
-#36 := (uf_1 uf_6 uf_3)
-decl uf_2 :: (-> T1 T2 T3 T1)
-decl uf_8 :: T3
-#33 := uf_8
-decl uf_5 :: T2
-#26 := uf_5
-decl uf_7 :: T3
-#31 := uf_7
-decl uf_4 :: T2
-#23 := uf_4
-#32 := (uf_2 uf_6 uf_4 uf_7)
-#34 := (uf_2 #32 uf_5 uf_8)
-#35 := (uf_1 #34 uf_3)
-#37 := (= #35 #36)
-#223 := (uf_1 #32 uf_4)
-#214 := (uf_2 uf_6 uf_4 #223)
-#552 := (uf_1 #214 uf_3)
-#555 := (= #552 #36)
-#560 := (= #36 #552)
-#556 := (= #223 #552)
-#24 := (= uf_3 uf_4)
-#561 := (ite #24 #556 #560)
-#8 := (:var 0 T2)
-#6 := (:var 1 T3)
-#5 := (:var 2 T2)
-#4 := (:var 3 T1)
-#7 := (uf_2 #4 #5 #6)
-#9 := (uf_1 #7 #8)
-#575 := (pattern #9)
-#11 := (uf_1 #4 #8)
-#100 := (= #9 #11)
-#99 := (= #6 #9)
-#55 := (= #5 #8)
-#83 := (ite #55 #99 #100)
-#576 := (forall (vars (?x1 T1) (?x2 T2) (?x3 T3) (?x4 T2)) (:pat #575) #83)
-#90 := (forall (vars (?x1 T1) (?x2 T2) (?x3 T3) (?x4 T2)) #83)
-#579 := (iff #90 #576)
-#577 := (iff #83 #83)
-#578 := [refl]: #577
-#580 := [quant-intro #578]: #579
-#58 := (ite #55 #6 #11)
-#61 := (= #9 #58)
-#64 := (forall (vars (?x1 T1) (?x2 T2) (?x3 T3) (?x4 T2)) #61)
-#87 := (iff #64 #90)
-#84 := (iff #61 #83)
-#89 := [rewrite]: #84
-#88 := [quant-intro #89]: #87
-#93 := (~ #64 #64)
-#91 := (~ #61 #61)
-#92 := [refl]: #91
-#94 := [nnf-pos #92]: #93
-#10 := (= #8 #5)
-#12 := (ite #10 #6 #11)
-#13 := (= #9 #12)
-#14 := (forall (vars (?x1 T1) (?x2 T2) (?x3 T3) (?x4 T2)) #13)
-#65 := (iff #14 #64)
-#62 := (iff #13 #61)
-#59 := (= #12 #58)
-#56 := (iff #10 #55)
-#57 := [rewrite]: #56
-#60 := [monotonicity #57]: #59
-#63 := [monotonicity #60]: #62
-#66 := [quant-intro #63]: #65
-#54 := [asserted]: #14
-#69 := [mp #54 #66]: #64
-#95 := [mp~ #69 #94]: #64
-#85 := [mp #95 #88]: #90
-#581 := [mp #85 #580]: #576
-#250 := (not #576)
-#548 := (or #250 #561)
-#551 := (= uf_4 uf_3)
-#557 := (ite #551 #556 #555)
-#549 := (or #250 #557)
-#271 := (iff #549 #548)
-#273 := (iff #548 #548)
-#259 := [rewrite]: #273
-#559 := (iff #557 #561)
-#198 := (iff #555 #560)
-#199 := [rewrite]: #198
-#193 := (iff #551 #24)
-#558 := [rewrite]: #193
-#562 := [monotonicity #558 #199]: #559
-#272 := [monotonicity #562]: #271
-#274 := [trans #272 #259]: #271
-#255 := [quant-inst]: #549
-#165 := [mp #255 #274]: #548
-#510 := [unit-resolution #165 #581]: #561
-#544 := (not #561)
-#497 := (or #544 #560)
-#25 := (not #24)
-#27 := (= uf_3 uf_5)
-#28 := (not #27)
-#29 := (and #25 #28)
-#75 := [asserted]: #29
-#79 := [and-elim #75]: #25
-#268 := (or #544 #24 #560)
-#542 := [def-axiom]: #268
-#499 := [unit-resolution #542 #79]: #497
-#491 := [unit-resolution #499 #510]: #560
-#493 := [symm #491]: #555
-#494 := (= #35 #552)
-#157 := (uf_1 #32 uf_3)
-#503 := (= #157 #552)
-#502 := (= #552 #157)
-#509 := (= #214 #32)
-#415 := (= #223 uf_7)
-#566 := (= uf_7 #223)
-#17 := (:var 0 T3)
-#16 := (:var 1 T2)
-#15 := (:var 2 T1)
-#18 := (uf_2 #15 #16 #17)
-#582 := (pattern #18)
-#19 := (uf_1 #18 #16)
-#68 := (= #17 #19)
-#584 := (forall (vars (?x5 T1) (?x6 T2) (?x7 T3)) (:pat #582) #68)
-#72 := (forall (vars (?x5 T1) (?x6 T2) (?x7 T3)) #68)
-#583 := (iff #72 #584)
-#586 := (iff #584 #584)
-#587 := [rewrite]: #586
-#585 := [rewrite]: #583
-#588 := [trans #585 #587]: #583
-#82 := (~ #72 #72)
-#96 := (~ #68 #68)
-#97 := [refl]: #96
-#78 := [nnf-pos #97]: #82
-#20 := (= #19 #17)
-#21 := (forall (vars (?x5 T1) (?x6 T2) (?x7 T3)) #20)
-#73 := (iff #21 #72)
-#70 := (iff #20 #68)
-#71 := [rewrite]: #70
-#74 := [quant-intro #71]: #73
-#67 := [asserted]: #21
-#77 := [mp #67 #74]: #72
-#98 := [mp~ #77 #78]: #72
-#589 := [mp #98 #588]: #584
-#211 := (not #584)
-#212 := (or #211 #566)
-#213 := [quant-inst]: #212
-#414 := [unit-resolution #213 #589]: #566
-#416 := [symm #414]: #415
-#506 := [monotonicity #416]: #509
-#498 := [monotonicity #506]: #502
-#492 := [symm #498]: #503
-#244 := (= #35 #157)
-#158 := (= uf_8 #35)
-#248 := (ite #27 #158 #244)
-#247 := (or #250 #248)
-#245 := (= uf_5 uf_3)
-#159 := (ite #245 #158 #244)
-#251 := (or #250 #159)
-#567 := (iff #251 #247)
-#224 := (iff #247 #247)
-#356 := [rewrite]: #224
-#249 := (iff #159 #248)
-#246 := (iff #245 #27)
-#237 := [rewrite]: #246
-#177 := [monotonicity #237]: #249
-#569 := [monotonicity #177]: #567
-#563 := [trans #569 #356]: #567
-#230 := [quant-inst]: #251
-#235 := [mp #230 #563]: #247
-#488 := [unit-resolution #235 #581]: #248
-#236 := (not #248)
-#490 := (or #236 #244)
-#80 := [and-elim #75]: #28
-#572 := (or #236 #27 #244)
-#573 := [def-axiom]: #572
-#500 := [unit-resolution #573 #80]: #490
-#501 := [unit-resolution #500 #488]: #244
-#495 := [trans #501 #492]: #494
-#489 := [trans #495 #493]: #37
-#38 := (not #37)
-#76 := [asserted]: #38
-[unit-resolution #76 #489]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_hol_02 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,13 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T2 T1 T3)
-:extrafuns (
- (uf_2 T1)
- (uf_3 T2)
- )
-:extrapreds (
- (up_4 T1 T2)
- (up_1 T1 T2)
- )
-:assumption (not (or (iff (up_1 uf_2 uf_3) (and (up_4 uf_2 uf_3) true)) (or (iff (up_1 uf_2 uf_3) true) (iff (up_4 uf_2 uf_3) true))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_hol_02.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,62 +0,0 @@
-#2 := false
-decl up_4 :: (-> T1 T2 bool)
-decl uf_3 :: T2
-#5 := uf_3
-decl uf_2 :: T1
-#4 := uf_2
-#7 := (up_4 uf_2 uf_3)
-#60 := (not #7)
-decl up_1 :: (-> T1 T2 bool)
-#6 := (up_1 uf_2 uf_3)
-#33 := (iff #6 #7)
-#49 := (or #6 #7 #33)
-#52 := (not #49)
-#1 := true
-#11 := (iff #7 true)
-#10 := (iff #6 true)
-#12 := (or #10 #11)
-#8 := (and #7 true)
-#9 := (iff #6 #8)
-#13 := (or #9 #12)
-#14 := (not #13)
-#55 := (iff #14 #52)
-#40 := (or #6 #7)
-#43 := (or #33 #40)
-#46 := (not #43)
-#53 := (iff #46 #52)
-#50 := (iff #43 #49)
-#51 := [rewrite]: #50
-#54 := [monotonicity #51]: #53
-#47 := (iff #14 #46)
-#44 := (iff #13 #43)
-#41 := (iff #12 #40)
-#38 := (iff #11 #7)
-#39 := [rewrite]: #38
-#36 := (iff #10 #6)
-#37 := [rewrite]: #36
-#42 := [monotonicity #37 #39]: #41
-#34 := (iff #9 #33)
-#31 := (iff #8 #7)
-#32 := [rewrite]: #31
-#35 := [monotonicity #32]: #34
-#45 := [monotonicity #35 #42]: #44
-#48 := [monotonicity #45]: #47
-#56 := [trans #48 #54]: #55
-#30 := [asserted]: #14
-#57 := [mp #30 #56]: #52
-#61 := [not-or-elim #57]: #60
-#58 := (not #6)
-#59 := [not-or-elim #57]: #58
-#72 := (or #7 #6)
-#66 := (iff #7 #58)
-#62 := (not #33)
-#64 := (iff #62 #66)
-#67 := [rewrite]: #64
-#63 := [not-or-elim #57]: #62
-#68 := [mp #63 #67]: #66
-#69 := (not #66)
-#70 := (or #7 #6 #69)
-#71 := [def-axiom]: #70
-#73 := [unit-resolution #71 #68]: #72
-[unit-resolution #73 #59 #61]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_hol_03 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,15 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T1 T2)
-:extrafuns (
- (uf_3 T2)
- (uf_1 T1 T1)
- (uf_4 T1)
- )
-:extrapreds (
- (up_2 T2)
- )
-:assumption (forall (?x1 T1) (= (uf_1 ?x1) ?x1))
-:assumption (forall (?x2 T2) (iff (up_2 ?x2) (= ?x2 uf_3)))
-:assumption (not (and (= (uf_1 uf_4) uf_4) (iff (up_2 uf_3) true)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_hol_03.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,115 +0,0 @@
-#2 := false
-decl up_2 :: (-> T2 bool)
-decl uf_3 :: T2
-#10 := uf_3
-#17 := (up_2 uf_3)
-#78 := (not #17)
-decl uf_1 :: (-> T1 T1)
-decl uf_4 :: T1
-#14 := uf_4
-#15 := (uf_1 uf_4)
-#46 := (= uf_4 #15)
-#79 := (not #46)
-#145 := [hypothesis]: #79
-#4 := (:var 0 T1)
-#5 := (uf_1 #4)
-#563 := (pattern #5)
-#37 := (= #4 #5)
-#564 := (forall (vars (?x1 T1)) (:pat #563) #37)
-#40 := (forall (vars (?x1 T1)) #37)
-#567 := (iff #40 #564)
-#565 := (iff #37 #37)
-#566 := [refl]: #565
-#568 := [quant-intro #566]: #567
-#72 := (~ #40 #40)
-#70 := (~ #37 #37)
-#71 := [refl]: #70
-#73 := [nnf-pos #71]: #72
-#6 := (= #5 #4)
-#7 := (forall (vars (?x1 T1)) #6)
-#41 := (iff #7 #40)
-#38 := (iff #6 #37)
-#39 := [rewrite]: #38
-#42 := [quant-intro #39]: #41
-#36 := [asserted]: #7
-#45 := [mp #36 #42]: #40
-#74 := [mp~ #45 #73]: #40
-#569 := [mp #74 #568]: #564
-#146 := (not #564)
-#233 := (or #146 #46)
-#147 := [quant-inst]: #233
-#232 := [unit-resolution #147 #569 #145]: false
-#234 := [lemma #232]: #46
-#66 := (or #78 #79)
-#54 := (and #17 #46)
-#59 := (not #54)
-#85 := (iff #59 #66)
-#67 := (not #66)
-#80 := (not #67)
-#83 := (iff #80 #66)
-#84 := [rewrite]: #83
-#81 := (iff #59 #80)
-#68 := (iff #54 #67)
-#69 := [rewrite]: #68
-#82 := [monotonicity #69]: #81
-#86 := [trans #82 #84]: #85
-#1 := true
-#18 := (iff #17 true)
-#16 := (= #15 uf_4)
-#19 := (and #16 #18)
-#20 := (not #19)
-#60 := (iff #20 #59)
-#57 := (iff #19 #54)
-#51 := (and #46 #17)
-#55 := (iff #51 #54)
-#56 := [rewrite]: #55
-#52 := (iff #19 #51)
-#49 := (iff #18 #17)
-#50 := [rewrite]: #49
-#47 := (iff #16 #46)
-#48 := [rewrite]: #47
-#53 := [monotonicity #48 #50]: #52
-#58 := [trans #53 #56]: #57
-#61 := [monotonicity #58]: #60
-#44 := [asserted]: #20
-#64 := [mp #44 #61]: #59
-#87 := [mp #64 #86]: #66
-#561 := [unit-resolution #87 #234]: #78
-#8 := (:var 0 T2)
-#9 := (up_2 #8)
-#570 := (pattern #9)
-#11 := (= #8 uf_3)
-#12 := (iff #9 #11)
-#571 := (forall (vars (?x2 T2)) (:pat #570) #12)
-#13 := (forall (vars (?x2 T2)) #12)
-#574 := (iff #13 #571)
-#572 := (iff #12 #12)
-#573 := [refl]: #572
-#575 := [quant-intro #573]: #574
-#65 := (~ #13 #13)
-#75 := (~ #12 #12)
-#76 := [refl]: #75
-#62 := [nnf-pos #76]: #65
-#43 := [asserted]: #13
-#77 := [mp~ #43 #62]: #13
-#576 := [mp #77 #575]: #571
-#555 := (not #571)
-#557 := (or #555 #17)
-#225 := (= uf_3 uf_3)
-#236 := (iff #17 #225)
-#212 := (or #555 #236)
-#551 := (iff #212 #557)
-#224 := (iff #557 #557)
-#558 := [rewrite]: #224
-#239 := (iff #236 #17)
-#238 := (iff #236 #18)
-#237 := (iff #225 true)
-#165 := [rewrite]: #237
-#235 := [monotonicity #165]: #238
-#218 := [trans #235 #50]: #239
-#223 := [monotonicity #218]: #551
-#559 := [trans #223 #558]: #551
-#344 := [quant-inst]: #212
-#560 := [mp #344 #559]: #557
-[unit-resolution #560 #576 #561]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_hol_04 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,18 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T2 T3 T1)
-:extrafuns (
- (uf_2 T1 T2 T3 T1)
- (uf_1 T1 T2 T3)
- (uf_6 T1)
- (uf_3 T2)
- (uf_4 T2)
- (uf_5 T2)
- (uf_7 T3)
- (uf_8 T3)
- )
-:assumption (forall (?x1 T1) (?x2 T2) (?x3 T3) (?x4 T2) (= (uf_1 (uf_2 ?x1 ?x2 ?x3) ?x4) (ite (= ?x4 ?x2) ?x3 (uf_1 ?x1 ?x4))))
-:assumption (forall (?x5 T1) (?x6 T2) (?x7 T3) (= (uf_1 (uf_2 ?x5 ?x6 ?x7) ?x6) ?x7))
-:assumption (and (not (= uf_3 uf_4)) (not (= uf_3 uf_5)))
-:assumption (not (= (uf_1 (uf_2 (uf_2 uf_6 uf_4 uf_7) uf_5 uf_8) uf_3) (uf_1 uf_6 uf_3)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_hol_04.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,181 +0,0 @@
-#2 := false
-decl uf_1 :: (-> T1 T2 T3)
-decl uf_3 :: T2
-#22 := uf_3
-decl uf_6 :: T1
-#30 := uf_6
-#36 := (uf_1 uf_6 uf_3)
-decl uf_2 :: (-> T1 T2 T3 T1)
-decl uf_8 :: T3
-#33 := uf_8
-decl uf_5 :: T2
-#26 := uf_5
-decl uf_7 :: T3
-#31 := uf_7
-decl uf_4 :: T2
-#23 := uf_4
-#32 := (uf_2 uf_6 uf_4 uf_7)
-#34 := (uf_2 #32 uf_5 uf_8)
-#35 := (uf_1 #34 uf_3)
-#37 := (= #35 #36)
-#223 := (uf_1 #32 uf_4)
-#214 := (uf_2 uf_6 uf_4 #223)
-#552 := (uf_1 #214 uf_3)
-#555 := (= #552 #36)
-#560 := (= #36 #552)
-#556 := (= #223 #552)
-#24 := (= uf_3 uf_4)
-#561 := (ite #24 #556 #560)
-#8 := (:var 0 T2)
-#6 := (:var 1 T3)
-#5 := (:var 2 T2)
-#4 := (:var 3 T1)
-#7 := (uf_2 #4 #5 #6)
-#9 := (uf_1 #7 #8)
-#575 := (pattern #9)
-#11 := (uf_1 #4 #8)
-#100 := (= #9 #11)
-#99 := (= #6 #9)
-#55 := (= #5 #8)
-#83 := (ite #55 #99 #100)
-#576 := (forall (vars (?x1 T1) (?x2 T2) (?x3 T3) (?x4 T2)) (:pat #575) #83)
-#90 := (forall (vars (?x1 T1) (?x2 T2) (?x3 T3) (?x4 T2)) #83)
-#579 := (iff #90 #576)
-#577 := (iff #83 #83)
-#578 := [refl]: #577
-#580 := [quant-intro #578]: #579
-#58 := (ite #55 #6 #11)
-#61 := (= #9 #58)
-#64 := (forall (vars (?x1 T1) (?x2 T2) (?x3 T3) (?x4 T2)) #61)
-#87 := (iff #64 #90)
-#84 := (iff #61 #83)
-#89 := [rewrite]: #84
-#88 := [quant-intro #89]: #87
-#93 := (~ #64 #64)
-#91 := (~ #61 #61)
-#92 := [refl]: #91
-#94 := [nnf-pos #92]: #93
-#10 := (= #8 #5)
-#12 := (ite #10 #6 #11)
-#13 := (= #9 #12)
-#14 := (forall (vars (?x1 T1) (?x2 T2) (?x3 T3) (?x4 T2)) #13)
-#65 := (iff #14 #64)
-#62 := (iff #13 #61)
-#59 := (= #12 #58)
-#56 := (iff #10 #55)
-#57 := [rewrite]: #56
-#60 := [monotonicity #57]: #59
-#63 := [monotonicity #60]: #62
-#66 := [quant-intro #63]: #65
-#54 := [asserted]: #14
-#69 := [mp #54 #66]: #64
-#95 := [mp~ #69 #94]: #64
-#85 := [mp #95 #88]: #90
-#581 := [mp #85 #580]: #576
-#250 := (not #576)
-#548 := (or #250 #561)
-#551 := (= uf_4 uf_3)
-#557 := (ite #551 #556 #555)
-#549 := (or #250 #557)
-#271 := (iff #549 #548)
-#273 := (iff #548 #548)
-#259 := [rewrite]: #273
-#559 := (iff #557 #561)
-#198 := (iff #555 #560)
-#199 := [rewrite]: #198
-#193 := (iff #551 #24)
-#558 := [rewrite]: #193
-#562 := [monotonicity #558 #199]: #559
-#272 := [monotonicity #562]: #271
-#274 := [trans #272 #259]: #271
-#255 := [quant-inst]: #549
-#165 := [mp #255 #274]: #548
-#510 := [unit-resolution #165 #581]: #561
-#544 := (not #561)
-#497 := (or #544 #560)
-#25 := (not #24)
-#27 := (= uf_3 uf_5)
-#28 := (not #27)
-#29 := (and #25 #28)
-#75 := [asserted]: #29
-#79 := [and-elim #75]: #25
-#268 := (or #544 #24 #560)
-#542 := [def-axiom]: #268
-#499 := [unit-resolution #542 #79]: #497
-#491 := [unit-resolution #499 #510]: #560
-#493 := [symm #491]: #555
-#494 := (= #35 #552)
-#157 := (uf_1 #32 uf_3)
-#503 := (= #157 #552)
-#502 := (= #552 #157)
-#509 := (= #214 #32)
-#415 := (= #223 uf_7)
-#566 := (= uf_7 #223)
-#17 := (:var 0 T3)
-#16 := (:var 1 T2)
-#15 := (:var 2 T1)
-#18 := (uf_2 #15 #16 #17)
-#582 := (pattern #18)
-#19 := (uf_1 #18 #16)
-#68 := (= #17 #19)
-#584 := (forall (vars (?x5 T1) (?x6 T2) (?x7 T3)) (:pat #582) #68)
-#72 := (forall (vars (?x5 T1) (?x6 T2) (?x7 T3)) #68)
-#583 := (iff #72 #584)
-#586 := (iff #584 #584)
-#587 := [rewrite]: #586
-#585 := [rewrite]: #583
-#588 := [trans #585 #587]: #583
-#82 := (~ #72 #72)
-#96 := (~ #68 #68)
-#97 := [refl]: #96
-#78 := [nnf-pos #97]: #82
-#20 := (= #19 #17)
-#21 := (forall (vars (?x5 T1) (?x6 T2) (?x7 T3)) #20)
-#73 := (iff #21 #72)
-#70 := (iff #20 #68)
-#71 := [rewrite]: #70
-#74 := [quant-intro #71]: #73
-#67 := [asserted]: #21
-#77 := [mp #67 #74]: #72
-#98 := [mp~ #77 #78]: #72
-#589 := [mp #98 #588]: #584
-#211 := (not #584)
-#212 := (or #211 #566)
-#213 := [quant-inst]: #212
-#414 := [unit-resolution #213 #589]: #566
-#416 := [symm #414]: #415
-#506 := [monotonicity #416]: #509
-#498 := [monotonicity #506]: #502
-#492 := [symm #498]: #503
-#244 := (= #35 #157)
-#158 := (= uf_8 #35)
-#248 := (ite #27 #158 #244)
-#247 := (or #250 #248)
-#245 := (= uf_5 uf_3)
-#159 := (ite #245 #158 #244)
-#251 := (or #250 #159)
-#567 := (iff #251 #247)
-#224 := (iff #247 #247)
-#356 := [rewrite]: #224
-#249 := (iff #159 #248)
-#246 := (iff #245 #27)
-#237 := [rewrite]: #246
-#177 := [monotonicity #237]: #249
-#569 := [monotonicity #177]: #567
-#563 := [trans #569 #356]: #567
-#230 := [quant-inst]: #251
-#235 := [mp #230 #563]: #247
-#488 := [unit-resolution #235 #581]: #248
-#236 := (not #248)
-#490 := (or #236 #244)
-#80 := [and-elim #75]: #28
-#572 := (or #236 #27 #244)
-#573 := [def-axiom]: #572
-#500 := [unit-resolution #573 #80]: #490
-#501 := [unit-resolution #500 #488]: #244
-#495 := [trans #501 #492]: #494
-#489 := [trans #495 #493]: #37
-#38 := (not #37)
-#76 := [asserted]: #38
-[unit-resolution #76 #489]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_hol_05 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,20 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T1 T2 T3)
-:extrafuns (
- (uf_6 Int T2)
- (uf_1 T1 T3 T3)
- (uf_4 T3)
- (uf_2 T2 T3 T3)
- (uf_3 T1 T2 T2)
- (uf_7 T2 Int)
- (uf_5 T1)
- )
-:assumption (forall (?x1 T1) (?x2 T2) (?x3 T3) (= (uf_1 ?x1 (uf_2 ?x2 ?x3)) (uf_2 (uf_3 ?x1 ?x2) (uf_1 ?x1 ?x3))))
-:assumption (forall (?x4 T1) (= (uf_1 ?x4 uf_4) uf_4))
-:assumption (forall (?x5 T2) (= (uf_3 uf_5 ?x5) (uf_6 (+ (uf_7 ?x5) 1))))
-:assumption (forall (?x6 T2) (= (uf_6 (uf_7 ?x6)) ?x6))
-:assumption (forall (?x7 Int) (implies (<= 0 ?x7) (= (uf_7 (uf_6 ?x7)) ?x7)))
-:assumption (forall (?x8 Int) (implies (< ?x8 0) (= (uf_7 (uf_6 ?x8)) 0)))
-:assumption (not (= (uf_1 uf_5 (uf_2 (uf_6 0) (uf_2 (uf_6 1) uf_4))) (uf_2 (uf_6 1) (uf_2 (uf_6 2) uf_4))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_hol_05.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,464 +0,0 @@
-#2 := false
-decl uf_2 :: (-> T2 T3 T3)
-decl uf_4 :: T3
-#15 := uf_4
-decl uf_6 :: (-> int T2)
-#48 := 2::int
-#49 := (uf_6 2::int)
-#50 := (uf_2 #49 uf_4)
-#23 := 1::int
-#44 := (uf_6 1::int)
-#51 := (uf_2 #44 #50)
-decl uf_1 :: (-> T1 T3 T3)
-#45 := (uf_2 #44 uf_4)
-#31 := 0::int
-#43 := (uf_6 0::int)
-#46 := (uf_2 #43 #45)
-decl uf_5 :: T1
-#19 := uf_5
-#47 := (uf_1 uf_5 #46)
-#52 := (= #47 #51)
-#266 := (uf_1 uf_5 #45)
-decl uf_3 :: (-> T1 T2 T2)
-#352 := (uf_3 uf_5 #43)
-#267 := (uf_2 #352 #266)
-#797 := (= #267 #51)
-#795 := (= #51 #267)
-#758 := (= #50 #266)
-#521 := (uf_1 uf_5 uf_4)
-#522 := (uf_3 uf_5 #44)
-#523 := (uf_2 #522 #521)
-#756 := (= #523 #266)
-#616 := (= #266 #523)
-#6 := (:var 0 T3)
-#4 := (:var 2 T1)
-#10 := (uf_1 #4 #6)
-#5 := (:var 1 T2)
-#9 := (uf_3 #4 #5)
-#11 := (uf_2 #9 #10)
-#683 := (pattern #11)
-#7 := (uf_2 #5 #6)
-#8 := (uf_1 #4 #7)
-#682 := (pattern #8)
-#12 := (= #8 #11)
-#684 := (forall (vars (?x1 T1) (?x2 T2) (?x3 T3)) (:pat #682 #683) #12)
-#13 := (forall (vars (?x1 T1) (?x2 T2) (?x3 T3)) #12)
-#687 := (iff #13 #684)
-#685 := (iff #12 #12)
-#686 := [refl]: #685
-#688 := [quant-intro #686]: #687
-#195 := (~ #13 #13)
-#193 := (~ #12 #12)
-#194 := [refl]: #193
-#196 := [nnf-pos #194]: #195
-#69 := [asserted]: #13
-#197 := [mp~ #69 #196]: #13
-#689 := [mp #197 #688]: #684
-#345 := (not #684)
-#604 := (or #345 #616)
-#606 := [quant-inst]: #604
-#277 := [unit-resolution #606 #689]: #616
-#757 := [symm #277]: #756
-#754 := (= #50 #523)
-#569 := (= uf_4 #521)
-#14 := (:var 0 T1)
-#16 := (uf_1 #14 uf_4)
-#690 := (pattern #16)
-#71 := (= uf_4 #16)
-#691 := (forall (vars (?x4 T1)) (:pat #690) #71)
-#74 := (forall (vars (?x4 T1)) #71)
-#694 := (iff #74 #691)
-#692 := (iff #71 #71)
-#693 := [refl]: #692
-#695 := [quant-intro #693]: #694
-#180 := (~ #74 #74)
-#198 := (~ #71 #71)
-#199 := [refl]: #198
-#178 := [nnf-pos #199]: #180
-#17 := (= #16 uf_4)
-#18 := (forall (vars (?x4 T1)) #17)
-#75 := (iff #18 #74)
-#72 := (iff #17 #71)
-#73 := [rewrite]: #72
-#76 := [quant-intro #73]: #75
-#70 := [asserted]: #18
-#79 := [mp #70 #76]: #74
-#200 := [mp~ #79 #178]: #74
-#696 := [mp #200 #695]: #691
-#572 := (not #691)
-#573 := (or #572 #569)
-#574 := [quant-inst]: #573
-#282 := [unit-resolution #574 #696]: #569
-#752 := (= #49 #522)
-decl uf_7 :: (-> T2 int)
-#666 := (uf_7 #44)
-#595 := (+ 1::int #666)
-#597 := (uf_6 #595)
-#748 := (= #597 #522)
-#605 := (= #522 #597)
-#20 := (:var 0 T2)
-#22 := (uf_7 #20)
-#698 := (pattern #22)
-#21 := (uf_3 uf_5 #20)
-#697 := (pattern #21)
-#78 := (+ 1::int #22)
-#82 := (uf_6 #78)
-#85 := (= #21 #82)
-#699 := (forall (vars (?x5 T2)) (:pat #697 #698) #85)
-#88 := (forall (vars (?x5 T2)) #85)
-#702 := (iff #88 #699)
-#700 := (iff #85 #85)
-#701 := [refl]: #700
-#703 := [quant-intro #701]: #702
-#181 := (~ #88 #88)
-#201 := (~ #85 #85)
-#202 := [refl]: #201
-#182 := [nnf-pos #202]: #181
-#24 := (+ #22 1::int)
-#25 := (uf_6 #24)
-#26 := (= #21 #25)
-#27 := (forall (vars (?x5 T2)) #26)
-#89 := (iff #27 #88)
-#86 := (iff #26 #85)
-#83 := (= #25 #82)
-#80 := (= #24 #78)
-#81 := [rewrite]: #80
-#84 := [monotonicity #81]: #83
-#87 := [monotonicity #84]: #86
-#90 := [quant-intro #87]: #89
-#77 := [asserted]: #27
-#93 := [mp #77 #90]: #88
-#203 := [mp~ #93 #182]: #88
-#704 := [mp #203 #703]: #699
-#607 := (not #699)
-#600 := (or #607 #605)
-#601 := [quant-inst]: #600
-#269 := [unit-resolution #601 #704]: #605
-#749 := [symm #269]: #748
-#750 := (= #49 #597)
-#499 := (uf_7 #597)
-#337 := (uf_6 #499)
-#318 := (= #337 #597)
-#28 := (uf_6 #22)
-#92 := (= #20 #28)
-#705 := (forall (vars (?x6 T2)) (:pat #698) #92)
-#96 := (forall (vars (?x6 T2)) #92)
-#706 := (iff #96 #705)
-#708 := (iff #705 #705)
-#709 := [rewrite]: #708
-#707 := [rewrite]: #706
-#710 := [trans #707 #709]: #706
-#183 := (~ #96 #96)
-#204 := (~ #92 #92)
-#205 := [refl]: #204
-#184 := [nnf-pos #205]: #183
-#29 := (= #28 #20)
-#30 := (forall (vars (?x6 T2)) #29)
-#97 := (iff #30 #96)
-#94 := (iff #29 #92)
-#95 := [rewrite]: #94
-#98 := [quant-intro #95]: #97
-#91 := [asserted]: #30
-#101 := [mp #91 #98]: #96
-#206 := [mp~ #101 #184]: #96
-#711 := [mp #206 #710]: #705
-#376 := (not #705)
-#325 := (or #376 #318)
-#316 := (= #597 #337)
-#326 := (or #376 #316)
-#328 := (iff #326 #325)
-#329 := (iff #325 #325)
-#310 := [rewrite]: #329
-#323 := (iff #316 #318)
-#324 := [rewrite]: #323
-#317 := [monotonicity #324]: #328
-#312 := [trans #317 #310]: #328
-#327 := [quant-inst]: #326
-#313 := [mp #327 #312]: #325
-#271 := [unit-resolution #313 #711]: #318
-#746 := (= #49 #337)
-#744 := (= 2::int #499)
-#742 := (= #499 2::int)
-#578 := -1::int
-#513 := (* -1::int #666)
-#514 := (+ #499 #513)
-#474 := (<= #514 1::int)
-#512 := (= #514 1::int)
-#504 := (>= #666 -1::int)
-#586 := (>= #666 1::int)
-#378 := (= #666 1::int)
-#32 := (:var 0 int)
-#34 := (uf_6 #32)
-#712 := (pattern #34)
-#118 := (>= #32 0::int)
-#119 := (not #118)
-#35 := (uf_7 #34)
-#100 := (= #32 #35)
-#125 := (or #100 #119)
-#713 := (forall (vars (?x7 int)) (:pat #712) #125)
-#130 := (forall (vars (?x7 int)) #125)
-#716 := (iff #130 #713)
-#714 := (iff #125 #125)
-#715 := [refl]: #714
-#717 := [quant-intro #715]: #716
-#185 := (~ #130 #130)
-#207 := (~ #125 #125)
-#208 := [refl]: #207
-#186 := [nnf-pos #208]: #185
-#36 := (= #35 #32)
-#33 := (<= 0::int #32)
-#37 := (implies #33 #36)
-#38 := (forall (vars (?x7 int)) #37)
-#133 := (iff #38 #130)
-#107 := (not #33)
-#108 := (or #107 #100)
-#113 := (forall (vars (?x7 int)) #108)
-#131 := (iff #113 #130)
-#128 := (iff #108 #125)
-#122 := (or #119 #100)
-#126 := (iff #122 #125)
-#127 := [rewrite]: #126
-#123 := (iff #108 #122)
-#120 := (iff #107 #119)
-#116 := (iff #33 #118)
-#117 := [rewrite]: #116
-#121 := [monotonicity #117]: #120
-#124 := [monotonicity #121]: #123
-#129 := [trans #124 #127]: #128
-#132 := [quant-intro #129]: #131
-#114 := (iff #38 #113)
-#111 := (iff #37 #108)
-#104 := (implies #33 #100)
-#109 := (iff #104 #108)
-#110 := [rewrite]: #109
-#105 := (iff #37 #104)
-#102 := (iff #36 #100)
-#103 := [rewrite]: #102
-#106 := [monotonicity #103]: #105
-#112 := [trans #106 #110]: #111
-#115 := [quant-intro #112]: #114
-#134 := [trans #115 #132]: #133
-#99 := [asserted]: #38
-#135 := [mp #99 #134]: #130
-#209 := [mp~ #135 #186]: #130
-#718 := [mp #209 #717]: #713
-#673 := (not #713)
-#365 := (or #673 #378)
-#307 := (>= 1::int 0::int)
-#668 := (not #307)
-#669 := (= 1::int #666)
-#655 := (or #669 #668)
-#366 := (or #673 #655)
-#645 := (iff #366 #365)
-#360 := (iff #365 #365)
-#643 := [rewrite]: #360
-#654 := (iff #655 #378)
-#374 := (or #378 false)
-#653 := (iff #374 #378)
-#650 := [rewrite]: #653
-#375 := (iff #655 #374)
-#651 := (iff #668 false)
-#1 := true
-#670 := (not true)
-#677 := (iff #670 false)
-#678 := [rewrite]: #677
-#381 := (iff #668 #670)
-#379 := (iff #307 true)
-#380 := [rewrite]: #379
-#274 := [monotonicity #380]: #381
-#652 := [trans #274 #678]: #651
-#656 := (iff #669 #378)
-#363 := [rewrite]: #656
-#649 := [monotonicity #363 #652]: #375
-#364 := [trans #649 #650]: #654
-#646 := [monotonicity #364]: #645
-#647 := [trans #646 #643]: #645
-#367 := [quant-inst]: #366
-#644 := [mp #367 #647]: #365
-#272 := [unit-resolution #644 #718]: #378
-#270 := (not #378)
-#273 := (or #270 #586)
-#725 := [th-lemma]: #273
-#726 := [unit-resolution #725 #272]: #586
-#727 := (not #586)
-#728 := (or #727 #504)
-#729 := [th-lemma]: #728
-#730 := [unit-resolution #729 #726]: #504
-#481 := (not #504)
-#496 := (or #673 #481 #512)
-#509 := (>= #595 0::int)
-#468 := (not #509)
-#501 := (= #595 #499)
-#503 := (or #501 #468)
-#497 := (or #673 #503)
-#470 := (iff #497 #496)
-#491 := (or #481 #512)
-#498 := (or #673 #491)
-#467 := (iff #498 #496)
-#469 := [rewrite]: #467
-#459 := (iff #497 #498)
-#494 := (iff #503 #491)
-#488 := (or #512 #481)
-#492 := (iff #488 #491)
-#493 := [rewrite]: #492
-#489 := (iff #503 #488)
-#486 := (iff #468 #481)
-#525 := (iff #509 #504)
-#480 := [rewrite]: #525
-#487 := [monotonicity #480]: #486
-#510 := (iff #501 #512)
-#524 := [rewrite]: #510
-#490 := [monotonicity #524 #487]: #489
-#495 := [trans #490 #493]: #494
-#460 := [monotonicity #495]: #459
-#471 := [trans #460 #469]: #470
-#482 := [quant-inst]: #497
-#473 := [mp #482 #471]: #496
-#731 := [unit-resolution #473 #718 #730]: #512
-#732 := (not #512)
-#733 := (or #732 #474)
-#734 := [th-lemma]: #733
-#735 := [unit-resolution #734 #731]: #474
-#475 := (>= #514 1::int)
-#736 := (or #732 #475)
-#737 := [th-lemma]: #736
-#738 := [unit-resolution #737 #731]: #475
-#582 := (<= #666 1::int)
-#739 := (or #270 #582)
-#740 := [th-lemma]: #739
-#741 := [unit-resolution #740 #272]: #582
-#743 := [th-lemma #726 #741 #738 #735]: #742
-#745 := [symm #743]: #744
-#747 := [monotonicity #745]: #746
-#751 := [trans #747 #271]: #750
-#753 := [trans #751 #749]: #752
-#755 := [monotonicity #753 #282]: #754
-#759 := [trans #755 #757]: #758
-#792 := (= #44 #352)
-#358 := (uf_7 #43)
-#613 := (+ 1::int #358)
-#617 := (uf_6 #613)
-#788 := (= #617 #352)
-#598 := (= #352 #617)
-#608 := (or #607 #598)
-#609 := [quant-inst]: #608
-#760 := [unit-resolution #609 #704]: #598
-#789 := [symm #760]: #788
-#790 := (= #44 #617)
-#575 := (uf_7 #617)
-#390 := (uf_6 #575)
-#382 := (= #390 #617)
-#385 := (or #376 #382)
-#392 := (= #617 #390)
-#386 := (or #376 #392)
-#387 := (iff #386 #385)
-#369 := (iff #385 #385)
-#370 := [rewrite]: #369
-#383 := (iff #392 #382)
-#384 := [rewrite]: #383
-#368 := [monotonicity #384]: #387
-#361 := [trans #368 #370]: #387
-#377 := [quant-inst]: #386
-#371 := [mp #377 #361]: #385
-#761 := [unit-resolution #371 #711]: #382
-#786 := (= #44 #390)
-#784 := (= 1::int #575)
-#782 := (= #575 1::int)
-#568 := (* -1::int #575)
-#579 := (+ #358 #568)
-#535 := (<= #579 -1::int)
-#557 := (= #579 -1::int)
-#561 := (>= #358 -1::int)
-#585 := (>= #358 0::int)
-#676 := (= #358 0::int)
-#315 := (or #673 #676)
-#268 := (>= 0::int 0::int)
-#354 := (not #268)
-#355 := (= 0::int #358)
-#359 := (or #355 #354)
-#657 := (or #673 #359)
-#320 := (iff #657 #315)
-#322 := (iff #315 #315)
-#659 := [rewrite]: #322
-#672 := (iff #359 #676)
-#675 := (or #676 false)
-#330 := (iff #675 #676)
-#335 := [rewrite]: #330
-#681 := (iff #359 #675)
-#679 := (iff #354 false)
-#343 := (iff #354 #670)
-#332 := (iff #268 true)
-#463 := [rewrite]: #332
-#344 := [monotonicity #463]: #343
-#680 := [trans #344 #678]: #679
-#338 := (iff #355 #676)
-#674 := [rewrite]: #338
-#671 := [monotonicity #674 #680]: #681
-#331 := [trans #671 #335]: #672
-#321 := [monotonicity #331]: #320
-#660 := [trans #321 #659]: #320
-#319 := [quant-inst]: #657
-#661 := [mp #319 #660]: #315
-#762 := [unit-resolution #661 #718]: #676
-#763 := (not #676)
-#764 := (or #763 #585)
-#765 := [th-lemma]: #764
-#766 := [unit-resolution #765 #762]: #585
-#767 := (not #585)
-#768 := (or #767 #561)
-#769 := [th-lemma]: #768
-#770 := [unit-resolution #769 #766]: #561
-#564 := (not #561)
-#549 := (or #673 #557 #564)
-#570 := (>= #613 0::int)
-#571 := (not #570)
-#576 := (= #613 #575)
-#577 := (or #576 #571)
-#552 := (or #673 #577)
-#530 := (iff #552 #549)
-#551 := (or #557 #564)
-#554 := (or #673 #551)
-#556 := (iff #554 #549)
-#529 := [rewrite]: #556
-#555 := (iff #552 #554)
-#547 := (iff #577 #551)
-#559 := (iff #571 #564)
-#562 := (iff #570 #561)
-#563 := [rewrite]: #562
-#565 := [monotonicity #563]: #559
-#558 := (iff #576 #557)
-#560 := [rewrite]: #558
-#548 := [monotonicity #560 #565]: #547
-#550 := [monotonicity #548]: #555
-#531 := [trans #550 #529]: #530
-#553 := [quant-inst]: #552
-#424 := [mp #553 #531]: #549
-#771 := [unit-resolution #424 #718 #770]: #557
-#772 := (not #557)
-#773 := (or #772 #535)
-#774 := [th-lemma]: #773
-#775 := [unit-resolution #774 #771]: #535
-#536 := (>= #579 -1::int)
-#776 := (or #772 #536)
-#777 := [th-lemma]: #776
-#778 := [unit-resolution #777 #771]: #536
-#584 := (<= #358 0::int)
-#779 := (or #763 #584)
-#780 := [th-lemma]: #779
-#781 := [unit-resolution #780 #762]: #584
-#783 := [th-lemma #766 #781 #778 #775]: #782
-#785 := [symm #783]: #784
-#787 := [monotonicity #785]: #786
-#791 := [trans #787 #761]: #790
-#793 := [trans #791 #789]: #792
-#796 := [monotonicity #793 #759]: #795
-#798 := [symm #796]: #797
-#353 := (= #47 #267)
-#356 := (or #345 #353)
-#357 := [quant-inst]: #356
-#794 := [unit-resolution #357 #689]: #353
-#799 := [trans #794 #798]: #52
-#53 := (not #52)
-#177 := [asserted]: #53
-[unit-resolution #177 #799]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_hol_06 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,8 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T1 T2)
-:extrapreds (
- (up_1 T1)
- )
-:assumption (not (or (forall (?x1 T1) (up_1 ?x1)) (not (forall (?x2 T1) (up_1 ?x2)))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_hol_06.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,21 +0,0 @@
-#2 := false
-decl up_1 :: (-> T1 bool)
-#4 := (:var 0 T1)
-#5 := (up_1 #4)
-#6 := (forall (vars (?x1 T1)) #5)
-#7 := (not #6)
-#8 := (or #6 #7)
-#9 := (not #8)
-#33 := (iff #9 false)
-#1 := true
-#28 := (not true)
-#31 := (iff #28 false)
-#32 := [rewrite]: #31
-#29 := (iff #9 #28)
-#26 := (iff #8 true)
-#27 := [rewrite]: #26
-#30 := [monotonicity #27]: #29
-#34 := [trans #30 #32]: #33
-#25 := [asserted]: #9
-[mp #25 #34]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_hol_07 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,14 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T1)
-:extrafuns (
- (uf_1 Int T1)
- (uf_3 T1 T1)
- (uf_2 T1 Int)
- )
-:assumption (forall (?x1 T1) (= (uf_1 (uf_2 ?x1)) ?x1))
-:assumption (forall (?x2 Int) (implies (<= 0 ?x2) (= (uf_2 (uf_1 ?x2)) ?x2)))
-:assumption (forall (?x3 Int) (implies (< ?x3 0) (= (uf_2 (uf_1 ?x3)) 0)))
-:assumption (forall (?x4 T1) (= (uf_3 ?x4) (ite (< (uf_2 ?x4) 10) ?x4 (uf_3 (uf_1 (- (uf_2 ?x4) 10))))))
-:assumption (not (= (uf_3 (uf_1 (* 4 (uf_2 (uf_3 (uf_1 4)))))) (uf_1 6)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_hol_07.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,373 +0,0 @@
-#2 := false
-decl uf_1 :: (-> int T1)
-#37 := 6::int
-#38 := (uf_1 6::int)
-decl uf_3 :: (-> T1 T1)
-decl uf_2 :: (-> T1 int)
-#30 := 4::int
-#31 := (uf_1 4::int)
-#32 := (uf_3 #31)
-#33 := (uf_2 #32)
-#34 := (* 4::int #33)
-#35 := (uf_1 #34)
-#36 := (uf_3 #35)
-#39 := (= #36 #38)
-#548 := (uf_3 #38)
-#394 := (= #548 #38)
-#549 := (= #38 #548)
-#523 := (uf_2 #38)
-#142 := -10::int
-#513 := (+ -10::int #523)
-#537 := (uf_1 #513)
-#538 := (uf_3 #537)
-#514 := (= #538 #548)
-#22 := 10::int
-#539 := (>= #523 10::int)
-#506 := (ite #539 #514 #549)
-#4 := (:var 0 T1)
-#21 := (uf_3 #4)
-#708 := (pattern #21)
-#5 := (uf_2 #4)
-#687 := (pattern #5)
-#209 := (= #4 #21)
-#143 := (+ -10::int #5)
-#146 := (uf_1 #143)
-#149 := (uf_3 #146)
-#208 := (= #21 #149)
-#163 := (>= #5 10::int)
-#190 := (ite #163 #208 #209)
-#709 := (forall (vars (?x4 T1)) (:pat #687 #708) #190)
-#193 := (forall (vars (?x4 T1)) #190)
-#712 := (iff #193 #709)
-#710 := (iff #190 #190)
-#711 := [refl]: #710
-#713 := [quant-intro #711]: #712
-#168 := (ite #163 #149 #4)
-#173 := (= #21 #168)
-#176 := (forall (vars (?x4 T1)) #173)
-#210 := (iff #176 #193)
-#191 := (iff #173 #190)
-#192 := [rewrite]: #191
-#211 := [quant-intro #192]: #210
-#188 := (~ #176 #176)
-#205 := (~ #173 #173)
-#206 := [refl]: #205
-#189 := [nnf-pos #206]: #188
-#24 := (- #5 10::int)
-#25 := (uf_1 #24)
-#26 := (uf_3 #25)
-#23 := (< #5 10::int)
-#27 := (ite #23 #4 #26)
-#28 := (= #21 #27)
-#29 := (forall (vars (?x4 T1)) #28)
-#179 := (iff #29 #176)
-#152 := (ite #23 #4 #149)
-#155 := (= #21 #152)
-#158 := (forall (vars (?x4 T1)) #155)
-#177 := (iff #158 #176)
-#174 := (iff #155 #173)
-#171 := (= #152 #168)
-#161 := (not #163)
-#165 := (ite #161 #4 #149)
-#169 := (= #165 #168)
-#170 := [rewrite]: #169
-#166 := (= #152 #165)
-#162 := (iff #23 #161)
-#164 := [rewrite]: #162
-#167 := [monotonicity #164]: #166
-#172 := [trans #167 #170]: #171
-#175 := [monotonicity #172]: #174
-#178 := [quant-intro #175]: #177
-#159 := (iff #29 #158)
-#156 := (iff #28 #155)
-#153 := (= #27 #152)
-#150 := (= #26 #149)
-#147 := (= #25 #146)
-#144 := (= #24 #143)
-#145 := [rewrite]: #144
-#148 := [monotonicity #145]: #147
-#151 := [monotonicity #148]: #150
-#154 := [monotonicity #151]: #153
-#157 := [monotonicity #154]: #156
-#160 := [quant-intro #157]: #159
-#180 := [trans #160 #178]: #179
-#141 := [asserted]: #29
-#181 := [mp #141 #180]: #176
-#207 := [mp~ #181 #189]: #176
-#212 := [mp #207 #211]: #193
-#714 := [mp #212 #713]: #709
-#681 := (not #709)
-#517 := (or #681 #506)
-#533 := (= #548 #538)
-#507 := (ite #539 #533 #549)
-#518 := (or #681 #507)
-#529 := (iff #518 #517)
-#530 := (iff #517 #517)
-#485 := [rewrite]: #530
-#508 := (iff #507 #506)
-#473 := (iff #533 #514)
-#504 := [rewrite]: #473
-#515 := [monotonicity #504]: #508
-#509 := [monotonicity #515]: #529
-#486 := [trans #509 #485]: #529
-#519 := [quant-inst]: #518
-#491 := [mp #519 #486]: #517
-#484 := [unit-resolution #491 #714]: #506
-#493 := (not #539)
-#465 := (<= #523 6::int)
-#526 := (= #523 6::int)
-#10 := (:var 0 int)
-#12 := (uf_1 #10)
-#695 := (pattern #12)
-#9 := 0::int
-#82 := (>= #10 0::int)
-#83 := (not #82)
-#13 := (uf_2 #12)
-#64 := (= #10 #13)
-#89 := (or #64 #83)
-#696 := (forall (vars (?x2 int)) (:pat #695) #89)
-#94 := (forall (vars (?x2 int)) #89)
-#699 := (iff #94 #696)
-#697 := (iff #89 #89)
-#698 := [refl]: #697
-#700 := [quant-intro #698]: #699
-#185 := (~ #94 #94)
-#199 := (~ #89 #89)
-#200 := [refl]: #199
-#183 := [nnf-pos #200]: #185
-#14 := (= #13 #10)
-#11 := (<= 0::int #10)
-#15 := (implies #11 #14)
-#16 := (forall (vars (?x2 int)) #15)
-#97 := (iff #16 #94)
-#71 := (not #11)
-#72 := (or #71 #64)
-#77 := (forall (vars (?x2 int)) #72)
-#95 := (iff #77 #94)
-#92 := (iff #72 #89)
-#86 := (or #83 #64)
-#90 := (iff #86 #89)
-#91 := [rewrite]: #90
-#87 := (iff #72 #86)
-#84 := (iff #71 #83)
-#80 := (iff #11 #82)
-#81 := [rewrite]: #80
-#85 := [monotonicity #81]: #84
-#88 := [monotonicity #85]: #87
-#93 := [trans #88 #91]: #92
-#96 := [quant-intro #93]: #95
-#78 := (iff #16 #77)
-#75 := (iff #15 #72)
-#68 := (implies #11 #64)
-#73 := (iff #68 #72)
-#74 := [rewrite]: #73
-#69 := (iff #15 #68)
-#66 := (iff #14 #64)
-#67 := [rewrite]: #66
-#70 := [monotonicity #67]: #69
-#76 := [trans #70 #74]: #75
-#79 := [quant-intro #76]: #78
-#98 := [trans #79 #96]: #97
-#63 := [asserted]: #16
-#99 := [mp #63 #98]: #94
-#201 := [mp~ #99 #183]: #94
-#701 := [mp #201 #700]: #696
-#671 := (not #696)
-#615 := (or #671 #526)
-#520 := (>= 6::int 0::int)
-#522 := (not #520)
-#516 := (= 6::int #523)
-#524 := (or #516 #522)
-#604 := (or #671 #524)
-#606 := (iff #604 #615)
-#601 := (iff #615 #615)
-#608 := [rewrite]: #601
-#614 := (iff #524 #526)
-#603 := (or #526 false)
-#612 := (iff #603 #526)
-#613 := [rewrite]: #612
-#600 := (iff #524 #603)
-#609 := (iff #522 false)
-#1 := true
-#327 := (not true)
-#666 := (iff #327 false)
-#667 := [rewrite]: #666
-#618 := (iff #522 #327)
-#528 := (iff #520 true)
-#621 := [rewrite]: #528
-#622 := [monotonicity #621]: #618
-#611 := [trans #622 #667]: #609
-#525 := (iff #516 #526)
-#527 := [rewrite]: #525
-#602 := [monotonicity #527 #611]: #600
-#610 := [trans #602 #613]: #614
-#607 := [monotonicity #610]: #606
-#592 := [trans #607 #608]: #606
-#605 := [quant-inst]: #604
-#593 := [mp #605 #592]: #615
-#454 := [unit-resolution #593 #701]: #526
-#303 := (not #526)
-#462 := (or #303 #465)
-#458 := [th-lemma]: #462
-#463 := [unit-resolution #458 #454]: #465
-#442 := (not #465)
-#445 := (or #442 #493)
-#449 := [th-lemma]: #445
-#451 := [unit-resolution #449 #463]: #493
-#492 := (not #506)
-#496 := (or #492 #539 #549)
-#497 := [def-axiom]: #496
-#452 := [unit-resolution #497 #451 #484]: #549
-#395 := [symm #452]: #394
-#397 := (= #36 #548)
-#372 := (uf_2 #35)
-#576 := (+ -10::int #372)
-#568 := (uf_1 #576)
-#569 := (uf_3 #568)
-#408 := (= #569 #548)
-#401 := (= #568 #38)
-#422 := (= #576 6::int)
-#677 := (uf_2 #31)
-#365 := -1::int
-#478 := (* -1::int #677)
-#479 := (+ #33 #478)
-#480 := (<= #479 0::int)
-#476 := (= #33 #677)
-#431 := (= #32 #31)
-#589 := (= #31 #32)
-#590 := (+ -10::int #677)
-#587 := (uf_1 #590)
-#591 := (uf_3 #587)
-#571 := (= #32 #591)
-#572 := (>= #677 10::int)
-#574 := (ite #572 #571 #589)
-#577 := (or #681 #574)
-#578 := [quant-inst]: #577
-#450 := [unit-resolution #578 #714]: #574
-#580 := (not #572)
-#552 := (<= #677 4::int)
-#324 := (= #677 4::int)
-#674 := (or #671 #324)
-#343 := (>= 4::int 0::int)
-#679 := (not #343)
-#336 := (= 4::int #677)
-#678 := (or #336 #679)
-#660 := (or #671 #678)
-#368 := (iff #660 #674)
-#384 := (iff #674 #674)
-#385 := [rewrite]: #384
-#312 := (iff #678 #324)
-#669 := (or #324 false)
-#672 := (iff #669 #324)
-#311 := [rewrite]: #672
-#306 := (iff #678 #669)
-#668 := (iff #679 false)
-#664 := (iff #679 #327)
-#325 := (iff #343 true)
-#326 := [rewrite]: #325
-#665 := [monotonicity #326]: #664
-#663 := [trans #665 #667]: #668
-#320 := (iff #336 #324)
-#662 := [rewrite]: #320
-#670 := [monotonicity #662 #663]: #306
-#673 := [trans #670 #311]: #312
-#383 := [monotonicity #673]: #368
-#386 := [trans #383 #385]: #368
-#661 := [quant-inst]: #660
-#278 := [mp #661 #386]: #674
-#453 := [unit-resolution #278 #701]: #324
-#441 := (not #324)
-#444 := (or #441 #552)
-#446 := [th-lemma]: #444
-#447 := [unit-resolution #446 #453]: #552
-#443 := (not #552)
-#448 := (or #443 #580)
-#438 := [th-lemma]: #448
-#428 := [unit-resolution #438 #447]: #580
-#579 := (not #574)
-#583 := (or #579 #572 #589)
-#573 := [def-axiom]: #583
-#430 := [unit-resolution #573 #428 #450]: #589
-#434 := [symm #430]: #431
-#435 := [monotonicity #434]: #476
-#439 := (not #476)
-#432 := (or #439 #480)
-#440 := [th-lemma]: #432
-#433 := [unit-resolution #440 #435]: #480
-#481 := (>= #479 0::int)
-#436 := (or #439 #481)
-#437 := [th-lemma]: #436
-#423 := [unit-resolution #437 #435]: #481
-#553 := (>= #677 4::int)
-#425 := (or #441 #553)
-#426 := [th-lemma]: #425
-#424 := [unit-resolution #426 #453]: #553
-#648 := (* -1::int #372)
-#652 := (+ #34 #648)
-#631 := (<= #652 0::int)
-#649 := (= #652 0::int)
-#370 := (>= #34 0::int)
-#409 := (not #481)
-#427 := (not #553)
-#411 := (or #370 #427 #409)
-#412 := [th-lemma]: #411
-#413 := [unit-resolution #412 #424 #423]: #370
-#371 := (not #370)
-#640 := (or #371 #649)
-#488 := (or #671 #371 #649)
-#650 := (= #34 #372)
-#651 := (or #650 #371)
-#489 := (or #671 #651)
-#630 := (iff #489 #488)
-#632 := (or #671 #640)
-#635 := (iff #632 #488)
-#629 := [rewrite]: #635
-#633 := (iff #489 #632)
-#641 := (iff #651 #640)
-#643 := (or #649 #371)
-#645 := (iff #643 #640)
-#646 := [rewrite]: #645
-#644 := (iff #651 #643)
-#653 := (iff #650 #649)
-#642 := [rewrite]: #653
-#639 := [monotonicity #642]: #644
-#647 := [trans #639 #646]: #641
-#634 := [monotonicity #647]: #633
-#636 := [trans #634 #629]: #630
-#490 := [quant-inst]: #489
-#637 := [mp #490 #636]: #488
-#414 := [unit-resolution #637 #701]: #640
-#415 := [unit-resolution #414 #413]: #649
-#416 := (not #649)
-#417 := (or #416 #631)
-#418 := [th-lemma]: #417
-#419 := [unit-resolution #418 #415]: #631
-#638 := (>= #652 0::int)
-#420 := (or #416 #638)
-#421 := [th-lemma]: #420
-#410 := [unit-resolution #421 #415]: #638
-#399 := [th-lemma #410 #419 #424 #447 #423 #433]: #422
-#402 := [monotonicity #399]: #401
-#393 := [monotonicity #402]: #408
-#564 := (= #36 #569)
-#575 := (= #35 #36)
-#570 := (>= #372 10::int)
-#556 := (ite #570 #564 #575)
-#554 := (or #681 #556)
-#557 := [quant-inst]: #554
-#403 := [unit-resolution #557 #714]: #556
-#404 := (not #631)
-#405 := (or #570 #404 #427 #409)
-#406 := [th-lemma]: #405
-#407 := [unit-resolution #406 #419 #424 #423]: #570
-#559 := (not #570)
-#558 := (not #556)
-#560 := (or #558 #559 #564)
-#555 := [def-axiom]: #560
-#400 := [unit-resolution #555 #407 #403]: #564
-#396 := [trans #400 #393]: #397
-#398 := [trans #396 #395]: #39
-#40 := (not #39)
-#182 := [asserted]: #40
-[unit-resolution #182 #398]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_hol_08 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,24 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T2 T1 T4 T3)
-:extrafuns (
- (uf_3 Int T1)
- (uf_7 T2 T4 T4)
- (uf_1 T2 T1 T1)
- (uf_6 T3 T4 Int)
- (uf_4 T1 Int)
- (uf_5 T2)
- (uf_2 T2)
- (uf_10 Int)
- (uf_8 T3)
- (uf_9 T4)
- )
-:assumption (forall (?x1 T1) (= (uf_1 uf_2 ?x1) (uf_3 (div (uf_4 ?x1) 2))))
-:assumption (forall (?x2 T1) (= (uf_1 uf_5 ?x2) (uf_3 (mod (uf_4 ?x2) 2))))
-:assumption (forall (?x3 T1) (= (uf_3 (uf_4 ?x3)) ?x3))
-:assumption (forall (?x4 Int) (implies (<= 0 ?x4) (= (uf_4 (uf_3 ?x4)) ?x4)))
-:assumption (forall (?x5 Int) (implies (< ?x5 0) (= (uf_4 (uf_3 ?x5)) 0)))
-:assumption (forall (?x6 T3) (?x7 T4) (= (mod (uf_6 ?x6 ?x7) 2) (mod (uf_6 ?x6 (uf_7 uf_5 ?x7)) 2)))
-:assumption (forall (?x8 T3) (?x9 T4) (= (+ (* (uf_6 ?x8 (uf_7 uf_2 ?x9)) 2) (uf_6 ?x8 (uf_7 uf_5 ?x9))) (uf_6 ?x8 ?x9)))
-:assumption (iff (= (uf_6 uf_8 uf_9) uf_10) (implies (= (mod (uf_6 uf_8 (uf_7 uf_5 uf_9)) 2) (mod uf_10 2)) (not (= (uf_6 uf_8 (uf_7 uf_2 uf_9)) (div (- uf_10 (uf_6 uf_8 (uf_7 uf_5 uf_9))) 2)))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_hol_08.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,410 +0,0 @@
-#2 := false
-#22 := 0::int
-decl uf_6 :: (-> T3 T4 int)
-decl uf_7 :: (-> T2 T4 T4)
-decl uf_9 :: T4
-#50 := uf_9
-decl uf_2 :: T2
-#4 := uf_2
-#59 := (uf_7 uf_2 uf_9)
-decl uf_8 :: T3
-#49 := uf_8
-#60 := (uf_6 uf_8 #59)
-#204 := -2::int
-#683 := (* -2::int #60)
-decl uf_5 :: T2
-#13 := uf_5
-#54 := (uf_7 uf_5 uf_9)
-#55 := (uf_6 uf_8 #54)
-#172 := -1::int
-#218 := (* -1::int #55)
-#685 := (+ #218 #683)
-#51 := (uf_6 uf_8 uf_9)
-#686 := (+ #51 #685)
-#679 := (>= #686 0::int)
-#687 := (= #686 0::int)
-#35 := (:var 0 T4)
-#43 := (uf_7 uf_2 #35)
-#34 := (:var 1 T3)
-#44 := (uf_6 #34 #43)
-#819 := (pattern #44)
-#38 := (uf_7 uf_5 #35)
-#39 := (uf_6 #34 #38)
-#812 := (pattern #39)
-#205 := (* -2::int #44)
-#203 := (* -1::int #39)
-#206 := (+ #203 #205)
-#36 := (uf_6 #34 #35)
-#207 := (+ #36 #206)
-#208 := (= #207 0::int)
-#820 := (forall (vars (?x8 T3) (?x9 T4)) (:pat #812 #819) #208)
-#211 := (forall (vars (?x8 T3) (?x9 T4)) #208)
-#823 := (iff #211 #820)
-#821 := (iff #208 #208)
-#822 := [refl]: #821
-#824 := [quant-intro #822]: #823
-#279 := (~ #211 #211)
-#305 := (~ #208 #208)
-#306 := [refl]: #305
-#280 := [nnf-pos #306]: #279
-#8 := 2::int
-#45 := (* #44 2::int)
-#46 := (+ #45 #39)
-#47 := (= #46 #36)
-#48 := (forall (vars (?x8 T3) (?x9 T4)) #47)
-#214 := (iff #48 #211)
-#171 := (* 2::int #44)
-#187 := (+ #39 #171)
-#195 := (= #36 #187)
-#200 := (forall (vars (?x8 T3) (?x9 T4)) #195)
-#212 := (iff #200 #211)
-#209 := (iff #195 #208)
-#210 := [rewrite]: #209
-#213 := [quant-intro #210]: #212
-#201 := (iff #48 #200)
-#198 := (iff #47 #195)
-#192 := (= #187 #36)
-#196 := (iff #192 #195)
-#197 := [rewrite]: #196
-#193 := (iff #47 #192)
-#190 := (= #46 #187)
-#184 := (+ #171 #39)
-#188 := (= #184 #187)
-#189 := [rewrite]: #188
-#185 := (= #46 #184)
-#182 := (= #45 #171)
-#183 := [rewrite]: #182
-#186 := [monotonicity #183]: #185
-#191 := [trans #186 #189]: #190
-#194 := [monotonicity #191]: #193
-#199 := [trans #194 #197]: #198
-#202 := [quant-intro #199]: #201
-#215 := [trans #202 #213]: #214
-#170 := [asserted]: #48
-#216 := [mp #170 #215]: #211
-#307 := [mp~ #216 #280]: #211
-#825 := [mp #307 #824]: #820
-#689 := (not #820)
-#675 := (or #689 #687)
-#676 := [quant-inst]: #675
-#536 := [unit-resolution #676 #825]: #687
-#537 := (not #687)
-#533 := (or #537 #679)
-#538 := [th-lemma]: #533
-#528 := [unit-resolution #538 #536]: #679
-decl uf_10 :: int
-#52 := uf_10
-#219 := (+ uf_10 #218)
-#222 := (div #219 2::int)
-#251 := (* -1::int #222)
-#252 := (+ #60 #251)
-#449 := (<= #252 0::int)
-#399 := (not #449)
-#253 := (= #252 0::int)
-#256 := (not #253)
-#57 := (mod uf_10 2::int)
-#243 := (* -1::int #57)
-#56 := (mod #55 2::int)
-#244 := (+ #56 #243)
-#245 := (= #244 0::int)
-#448 := (>= #244 0::int)
-#688 := (mod #51 2::int)
-#666 := (* -1::int #688)
-#667 := (+ #56 #666)
-#660 := (>= #667 0::int)
-#668 := (= #667 0::int)
-#40 := (mod #39 2::int)
-#173 := (* -1::int #40)
-#37 := (mod #36 2::int)
-#174 := (+ #37 #173)
-#175 := (= #174 0::int)
-#813 := (forall (vars (?x6 T3) (?x7 T4)) (:pat #812) #175)
-#178 := (forall (vars (?x6 T3) (?x7 T4)) #175)
-#816 := (iff #178 #813)
-#814 := (iff #175 #175)
-#815 := [refl]: #814
-#817 := [quant-intro #815]: #816
-#277 := (~ #178 #178)
-#302 := (~ #175 #175)
-#303 := [refl]: #302
-#278 := [nnf-pos #303]: #277
-#41 := (= #37 #40)
-#42 := (forall (vars (?x6 T3) (?x7 T4)) #41)
-#179 := (iff #42 #178)
-#176 := (iff #41 #175)
-#177 := [rewrite]: #176
-#180 := [quant-intro #177]: #179
-#169 := [asserted]: #42
-#181 := [mp #169 #180]: #178
-#304 := [mp~ #181 #278]: #178
-#818 := [mp #304 #817]: #813
-#673 := (not #813)
-#663 := (or #673 #668)
-#756 := (* -1::int #56)
-#684 := (+ #688 #756)
-#680 := (= #684 0::int)
-#674 := (or #673 #680)
-#653 := (iff #674 #663)
-#656 := (iff #663 #663)
-#657 := [rewrite]: #656
-#671 := (iff #680 #668)
-#677 := (+ #756 #688)
-#662 := (= #677 0::int)
-#669 := (iff #662 #668)
-#670 := [rewrite]: #669
-#664 := (iff #680 #662)
-#681 := (= #684 #677)
-#661 := [rewrite]: #681
-#665 := [monotonicity #661]: #664
-#672 := [trans #665 #670]: #671
-#655 := [monotonicity #672]: #653
-#658 := [trans #655 #657]: #653
-#652 := [quant-inst]: #674
-#659 := [mp #652 #658]: #663
-#394 := [unit-resolution #659 #818]: #668
-#552 := (not #668)
-#514 := (or #552 #660)
-#517 := [th-lemma]: #514
-#499 := [unit-resolution #517 #394]: #660
-#503 := (not #448)
-#414 := [hypothesis]: #503
-#561 := (+ #57 #666)
-#709 := (<= #561 0::int)
-#602 := (= #57 #688)
-#468 := (= #688 #57)
-#53 := (= #51 uf_10)
-#248 := (not #245)
-#259 := (or #248 #256)
-#362 := (mod #219 2::int)
-#699 := (>= #362 0::int)
-#1 := true
-#81 := [true-axiom]: true
-#604 := (or false #699)
-#506 := [th-lemma]: #604
-#507 := [unit-resolution #506 #81]: #699
-#628 := (* -1::int uf_10)
-#623 := (+ #51 #628)
-#629 := (<= #623 0::int)
-#498 := (not #629)
-#597 := (>= #623 0::int)
-#381 := (not #259)
-#508 := [hypothesis]: #381
-#450 := (or #259 #245)
-#441 := [def-axiom]: #450
-#509 := [unit-resolution #441 #508]: #245
-#510 := (or #248 #448)
-#511 := [th-lemma]: #510
-#500 := [unit-resolution #511 #509]: #448
-#743 := (div uf_10 2::int)
-#723 := (* -2::int #743)
-#545 := (* -2::int #688)
-#546 := (+ #545 #723)
-#646 := (div #51 2::int)
-#645 := (* -2::int #646)
-#547 := (+ #645 #546)
-#605 := (* -2::int #57)
-#549 := (+ #605 #547)
-#594 := (* 2::int #56)
-#550 := (+ #594 #549)
-#598 := (* 2::int uf_10)
-#551 := (+ #598 #550)
-#563 := (>= #551 2::int)
-#520 := (not #563)
-#361 := (<= #244 0::int)
-#512 := (or #248 #361)
-#489 := [th-lemma]: #512
-#491 := [unit-resolution #489 #509]: #361
-#363 := (>= #252 0::int)
-#452 := (or #259 #253)
-#453 := [def-axiom]: #452
-#492 := [unit-resolution #453 #508]: #253
-#493 := (or #256 #363)
-#494 := [th-lemma]: #493
-#495 := [unit-resolution #494 #492]: #363
-#556 := (not #361)
-#573 := (not #363)
-#521 := (or #520 #573 #556)
-#703 := (>= #362 2::int)
-#704 := (not #703)
-#599 := (or false #704)
-#620 := [th-lemma]: #599
-#575 := [unit-resolution #620 #81]: #704
-#654 := (<= #667 0::int)
-#548 := (or #552 #654)
-#553 := [th-lemma]: #548
-#532 := [unit-resolution #553 #394]: #654
-#651 := (+ #645 #666)
-#624 := (+ #51 #651)
-#626 := (<= #624 0::int)
-#650 := (= #624 0::int)
-#535 := (or false #650)
-#539 := [th-lemma]: #535
-#541 := [unit-resolution #539 #81]: #650
-#542 := (not #650)
-#540 := (or #542 #626)
-#543 := [th-lemma]: #540
-#531 := [unit-resolution #543 #541]: #626
-#587 := [hypothesis]: #361
-#724 := (+ #243 #723)
-#725 := (+ uf_10 #724)
-#727 := (<= #725 0::int)
-#722 := (= #725 0::int)
-#576 := (or false #722)
-#581 := [th-lemma]: #576
-#582 := [unit-resolution #581 #81]: #722
-#583 := (not #722)
-#584 := (or #583 #727)
-#585 := [th-lemma]: #584
-#586 := [unit-resolution #585 #582]: #727
-#534 := [hypothesis]: #563
-#555 := [hypothesis]: #363
-#616 := (* -1::int #362)
-#615 := (* -2::int #222)
-#617 := (+ #615 #616)
-#618 := (+ #218 #617)
-#711 := (+ uf_10 #618)
-#708 := (<= #711 0::int)
-#606 := (= #711 0::int)
-#562 := (or false #606)
-#564 := [th-lemma]: #562
-#565 := [unit-resolution #564 #81]: #606
-#566 := (not #606)
-#568 := (or #566 #708)
-#569 := [th-lemma]: #568
-#570 := [unit-resolution #569 #565]: #708
-#518 := [th-lemma #570 #555 #528 #534 #586 #587 #531 #532 #575]: false
-#524 := [lemma #518]: #521
-#496 := [unit-resolution #524 #495 #491]: #520
-#504 := (or #597 #563 #503)
-#529 := (not #597)
-#522 := [hypothesis]: #529
-#519 := (>= #624 0::int)
-#530 := (or #542 #519)
-#523 := [th-lemma]: #530
-#526 := [unit-resolution #523 #541]: #519
-#527 := [hypothesis]: #448
-#721 := (>= #725 0::int)
-#513 := (or #583 #721)
-#515 := [th-lemma]: #513
-#516 := [unit-resolution #515 #582]: #721
-#501 := [th-lemma #499 #516 #527 #526 #522]: #563
-#525 := [hypothesis]: #520
-#502 := [unit-resolution #525 #501]: false
-#505 := [lemma #502]: #504
-#497 := [unit-resolution #505 #496 #500]: #597
-#485 := (or #498 #529)
-#558 := (not #53)
-#440 := (or #558 #259)
-#262 := (iff #53 #259)
-#61 := (- uf_10 #55)
-#62 := (div #61 2::int)
-#63 := (= #60 #62)
-#64 := (not #63)
-#58 := (= #56 #57)
-#65 := (implies #58 #64)
-#66 := (iff #53 #65)
-#265 := (iff #66 #262)
-#225 := (= #60 #222)
-#228 := (not #225)
-#234 := (not #58)
-#235 := (or #234 #228)
-#240 := (iff #53 #235)
-#263 := (iff #240 #262)
-#260 := (iff #235 #259)
-#257 := (iff #228 #256)
-#254 := (iff #225 #253)
-#255 := [rewrite]: #254
-#258 := [monotonicity #255]: #257
-#249 := (iff #234 #248)
-#246 := (iff #58 #245)
-#247 := [rewrite]: #246
-#250 := [monotonicity #247]: #249
-#261 := [monotonicity #250 #258]: #260
-#264 := [monotonicity #261]: #263
-#241 := (iff #66 #240)
-#238 := (iff #65 #235)
-#231 := (implies #58 #228)
-#236 := (iff #231 #235)
-#237 := [rewrite]: #236
-#232 := (iff #65 #231)
-#229 := (iff #64 #228)
-#226 := (iff #63 #225)
-#223 := (= #62 #222)
-#220 := (= #61 #219)
-#221 := [rewrite]: #220
-#224 := [monotonicity #221]: #223
-#227 := [monotonicity #224]: #226
-#230 := [monotonicity #227]: #229
-#233 := [monotonicity #230]: #232
-#239 := [trans #233 #237]: #238
-#242 := [monotonicity #239]: #241
-#266 := [trans #242 #264]: #265
-#217 := [asserted]: #66
-#267 := [mp #217 #266]: #262
-#455 := (not #262)
-#765 := (or #558 #259 #455)
-#439 := [def-axiom]: #765
-#772 := [unit-resolution #439 #267]: #440
-#490 := [unit-resolution #772 #508]: #558
-#483 := (or #53 #498 #529)
-#484 := [th-lemma]: #483
-#487 := [unit-resolution #484 #490]: #485
-#486 := [unit-resolution #487 #497]: #498
-#678 := (<= #686 0::int)
-#488 := (or #537 #678)
-#477 := [th-lemma]: #488
-#478 := [unit-resolution #477 #536]: #678
-#479 := (or #256 #449)
-#471 := [th-lemma]: #479
-#480 := [unit-resolution #471 #492]: #449
-#712 := (>= #711 0::int)
-#481 := (or #566 #712)
-#472 := [th-lemma]: #481
-#482 := [unit-resolution #472 #565]: #712
-#463 := [th-lemma #482 #480 #478 #486 #507]: false
-#464 := [lemma #463]: #259
-#771 := (or #53 #381)
-#434 := (or #53 #381 #455)
-#769 := [def-axiom]: #434
-#428 := [unit-resolution #769 #267]: #771
-#442 := [unit-resolution #428 #464]: #53
-#435 := [monotonicity #442]: #468
-#437 := [symm #435]: #602
-#438 := (not #602)
-#419 := (or #438 #709)
-#420 := [th-lemma]: #419
-#421 := [unit-resolution #420 #437]: #709
-#422 := [th-lemma #421 #414 #499]: false
-#423 := [lemma #422]: #448
-#410 := (or #245 #503)
-#611 := (>= #561 0::int)
-#682 := (or #438 #611)
-#447 := [th-lemma]: #682
-#430 := [unit-resolution #447 #437]: #611
-#432 := [hypothesis]: #556
-#433 := [th-lemma #532 #432 #430]: false
-#412 := [lemma #433]: #361
-#409 := (or #245 #556 #503)
-#407 := [th-lemma]: #409
-#398 := [unit-resolution #407 #412]: #410
-#400 := [unit-resolution #398 #423]: #245
-#454 := (or #381 #248 #256)
-#451 := [def-axiom]: #454
-#401 := [unit-resolution #451 #464]: #259
-#404 := [unit-resolution #401 #400]: #256
-#384 := (or #253 #399)
-#429 := [hypothesis]: #573
-#443 := (or #558 #597)
-#444 := [th-lemma]: #443
-#445 := [unit-resolution #444 #442]: #597
-#446 := [th-lemma #445 #507 #482 #429 #478]: false
-#436 := [lemma #446]: #363
-#405 := (or #253 #399 #573)
-#379 := [th-lemma]: #405
-#385 := [unit-resolution #379 #436]: #384
-#390 := [unit-resolution #385 #404]: #399
-#392 := (or #558 #629)
-#393 := [th-lemma]: #392
-#395 := [unit-resolution #393 #442]: #629
-[th-lemma #395 #575 #570 #390 #528]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_linarith_01 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (= 3 3))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_linarith_01.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,17 +0,0 @@
-#2 := false
-#4 := 3::int
-#5 := (= 3::int 3::int)
-#6 := (not #5)
-#30 := (iff #6 false)
-#1 := true
-#25 := (not true)
-#28 := (iff #25 false)
-#29 := [rewrite]: #28
-#26 := (iff #6 #25)
-#23 := (iff #5 true)
-#24 := [rewrite]: #23
-#27 := [monotonicity #24]: #26
-#31 := [trans #27 #29]: #30
-#22 := [asserted]: #6
-[mp #22 #31]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_linarith_02 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (= 3.0 3.0))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_linarith_02.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,17 +0,0 @@
-#2 := false
-#4 := 3::real
-#5 := (= 3::real 3::real)
-#6 := (not #5)
-#30 := (iff #6 false)
-#1 := true
-#25 := (not true)
-#28 := (iff #25 false)
-#29 := [rewrite]: #28
-#26 := (iff #6 #25)
-#23 := (iff #5 true)
-#24 := [rewrite]: #23
-#27 := [monotonicity #24]: #26
-#31 := [trans #27 #29]: #30
-#22 := [asserted]: #6
-[mp #22 #31]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_linarith_03 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (= (+ 3 1) 4))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_linarith_03.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,26 +0,0 @@
-#2 := false
-#7 := 4::int
-#5 := 1::int
-#4 := 3::int
-#6 := (+ 3::int 1::int)
-#8 := (= #6 4::int)
-#9 := (not #8)
-#39 := (iff #9 false)
-#1 := true
-#34 := (not true)
-#37 := (iff #34 false)
-#38 := [rewrite]: #37
-#35 := (iff #9 #34)
-#32 := (iff #8 true)
-#27 := (= 4::int 4::int)
-#30 := (iff #27 true)
-#31 := [rewrite]: #30
-#28 := (iff #8 #27)
-#26 := [rewrite]: #8
-#29 := [monotonicity #26]: #28
-#33 := [trans #29 #31]: #32
-#36 := [monotonicity #33]: #35
-#40 := [trans #36 #38]: #39
-#25 := [asserted]: #9
-[mp #25 #40]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_linarith_04 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,9 +0,0 @@
-(benchmark Isabelle
-:extrafuns (
- (uf_1 Int)
- (uf_2 Int)
- (uf_3 Int)
- )
-:assumption (not (= (+ uf_1 (+ uf_2 uf_3)) (+ uf_2 (+ uf_3 uf_1))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_linarith_04.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,41 +0,0 @@
-#2 := false
-decl uf_1 :: int
-#4 := uf_1
-decl uf_3 :: int
-#6 := uf_3
-#9 := (+ uf_3 uf_1)
-decl uf_2 :: int
-#5 := uf_2
-#10 := (+ uf_2 #9)
-#7 := (+ uf_2 uf_3)
-#8 := (+ uf_1 #7)
-#11 := (= #8 #10)
-#12 := (not #11)
-#51 := (iff #12 false)
-#1 := true
-#46 := (not true)
-#49 := (iff #46 false)
-#50 := [rewrite]: #49
-#47 := (iff #12 #46)
-#44 := (iff #11 true)
-#39 := (= #8 #8)
-#42 := (iff #39 true)
-#43 := [rewrite]: #42
-#40 := (iff #11 #39)
-#37 := (= #10 #8)
-#29 := (+ uf_1 uf_3)
-#32 := (+ uf_2 #29)
-#35 := (= #32 #8)
-#36 := [rewrite]: #35
-#33 := (= #10 #32)
-#30 := (= #9 #29)
-#31 := [rewrite]: #30
-#34 := [monotonicity #31]: #33
-#38 := [trans #34 #36]: #37
-#41 := [monotonicity #38]: #40
-#45 := [trans #41 #43]: #44
-#48 := [monotonicity #45]: #47
-#52 := [trans #48 #50]: #51
-#28 := [asserted]: #12
-[mp #28 #52]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_linarith_05 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (< 5 (ite (<= 3 8) 8 3)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_linarith_05.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,35 +0,0 @@
-#2 := false
-#5 := 3::int
-#6 := 8::int
-#7 := (<= 3::int 8::int)
-#8 := (ite #7 8::int 3::int)
-#4 := 5::int
-#9 := (< 5::int #8)
-#10 := (not #9)
-#50 := (iff #10 false)
-#1 := true
-#45 := (not true)
-#48 := (iff #45 false)
-#49 := [rewrite]: #48
-#46 := (iff #10 #45)
-#43 := (iff #9 true)
-#38 := (< 5::int 8::int)
-#41 := (iff #38 true)
-#42 := [rewrite]: #41
-#39 := (iff #9 #38)
-#36 := (= #8 8::int)
-#31 := (ite true 8::int 3::int)
-#34 := (= #31 8::int)
-#35 := [rewrite]: #34
-#32 := (= #8 #31)
-#29 := (iff #7 true)
-#30 := [rewrite]: #29
-#33 := [monotonicity #30]: #32
-#37 := [trans #33 #35]: #36
-#40 := [monotonicity #37]: #39
-#44 := [trans #40 #42]: #43
-#47 := [monotonicity #44]: #46
-#51 := [trans #47 #49]: #50
-#26 := [asserted]: #10
-[mp #26 #51]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_linarith_06 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,8 +0,0 @@
-(benchmark Isabelle
-:extrafuns (
- (uf_1 Real)
- (uf_2 Real)
- )
-:assumption (not (<= (ite (< (+ uf_1 uf_2) 0.0) (~ (+ uf_1 uf_2)) (+ uf_1 uf_2)) (+ (ite (< uf_1 0.0) (~ uf_1) uf_1) (ite (< uf_2 0.0) (~ uf_2) uf_2))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_linarith_06.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,250 +0,0 @@
-#2 := false
-#7 := 0::real
-decl uf_2 :: real
-#5 := uf_2
-#143 := 2::real
-#144 := (* 2::real uf_2)
-#165 := (<= #144 0::real)
-#188 := (not #165)
-#88 := (>= uf_2 0::real)
-#166 := (or #88 #165)
-#191 := (not #166)
-decl uf_1 :: real
-#4 := uf_1
-#76 := (>= uf_1 0::real)
-#89 := (not #88)
-#146 := (* 2::real uf_1)
-#167 := (<= #146 0::real)
-#199 := (not #167)
-#263 := [hypothesis]: #88
-#147 := (+ #146 #144)
-#168 := (<= #147 0::real)
-#169 := (ite #88 #167 #168)
-#194 := (not #169)
-#186 := (or #166 #89)
-#187 := [def-axiom]: #186
-#271 := [unit-resolution #187 #263]: #166
-#170 := (ite #76 #166 #169)
-#205 := (not #170)
-#6 := (+ uf_1 uf_2)
-#64 := (>= #6 0::real)
-#269 := (or #64 #89)
-#65 := (not #64)
-#262 := [hypothesis]: #65
-#174 := (>= #144 0::real)
-#175 := (or #89 #174)
-#230 := (not #175)
-#257 := [hypothesis]: #230
-#225 := (or #175 #88)
-#226 := [def-axiom]: #225
-#258 := [unit-resolution #226 #257]: #88
-#227 := (not #174)
-#228 := (or #175 #227)
-#229 := [def-axiom]: #228
-#259 := [unit-resolution #229 #257]: #227
-#260 := [th-lemma #259 #258]: false
-#261 := [lemma #260]: #175
-#172 := (>= #146 0::real)
-#171 := (>= #147 0::real)
-#173 := (ite #88 #171 #172)
-#176 := (ite #76 #173 #175)
-#233 := (not #176)
-#264 := (or #64 #233)
-#177 := (ite #64 #170 #176)
-#182 := (not #177)
-#36 := -1::real
-#38 := (* -1::real uf_2)
-#95 := (ite #88 uf_2 #38)
-#107 := (* -1::real #95)
-#37 := (* -1::real uf_1)
-#83 := (ite #76 uf_1 #37)
-#106 := (* -1::real #83)
-#108 := (+ #106 #107)
-#39 := (+ #37 #38)
-#71 := (ite #64 #6 #39)
-#109 := (+ #71 #108)
-#110 := (<= #109 0::real)
-#115 := (not #110)
-#183 := (iff #115 #182)
-#180 := (iff #110 #177)
-#150 := -2::real
-#152 := (* -2::real uf_2)
-#155 := (ite #88 #152 0::real)
-#151 := (* -2::real uf_1)
-#153 := (+ #151 #152)
-#154 := (ite #88 #153 #151)
-#156 := (ite #76 #154 #155)
-#148 := (ite #88 #146 #147)
-#145 := (ite #88 0::real #144)
-#149 := (ite #76 #145 #148)
-#157 := (ite #64 #149 #156)
-#162 := (<= #157 0::real)
-#178 := (iff #162 #177)
-#179 := [rewrite]: #178
-#163 := (iff #110 #162)
-#160 := (= #109 #157)
-#133 := (+ uf_1 #38)
-#134 := (ite #88 #133 #6)
-#131 := (+ #37 uf_2)
-#132 := (ite #88 #39 #131)
-#135 := (ite #76 #132 #134)
-#140 := (+ #71 #135)
-#158 := (= #140 #157)
-#159 := [rewrite]: #158
-#141 := (= #109 #140)
-#138 := (= #108 #135)
-#125 := (ite #88 #38 uf_2)
-#123 := (ite #76 #37 uf_1)
-#128 := (+ #123 #125)
-#136 := (= #128 #135)
-#137 := [rewrite]: #136
-#129 := (= #108 #128)
-#126 := (= #107 #125)
-#127 := [rewrite]: #126
-#121 := (= #106 #123)
-#124 := [rewrite]: #121
-#130 := [monotonicity #124 #127]: #129
-#139 := [trans #130 #137]: #138
-#142 := [monotonicity #139]: #141
-#161 := [trans #142 #159]: #160
-#164 := [monotonicity #161]: #163
-#181 := [trans #164 #179]: #180
-#184 := [monotonicity #181]: #183
-#15 := (- uf_2)
-#14 := (< uf_2 0::real)
-#16 := (ite #14 #15 uf_2)
-#12 := (- uf_1)
-#11 := (< uf_1 0::real)
-#13 := (ite #11 #12 uf_1)
-#17 := (+ #13 #16)
-#9 := (- #6)
-#8 := (< #6 0::real)
-#10 := (ite #8 #9 #6)
-#18 := (<= #10 #17)
-#19 := (not #18)
-#118 := (iff #19 #115)
-#52 := (ite #14 #38 uf_2)
-#47 := (ite #11 #37 uf_1)
-#55 := (+ #47 #52)
-#42 := (ite #8 #39 #6)
-#58 := (<= #42 #55)
-#61 := (not #58)
-#116 := (iff #61 #115)
-#113 := (iff #58 #110)
-#100 := (+ #83 #95)
-#103 := (<= #71 #100)
-#111 := (iff #103 #110)
-#112 := [rewrite]: #111
-#104 := (iff #58 #103)
-#101 := (= #55 #100)
-#98 := (= #52 #95)
-#92 := (ite #89 #38 uf_2)
-#96 := (= #92 #95)
-#97 := [rewrite]: #96
-#93 := (= #52 #92)
-#90 := (iff #14 #89)
-#91 := [rewrite]: #90
-#94 := [monotonicity #91]: #93
-#99 := [trans #94 #97]: #98
-#86 := (= #47 #83)
-#77 := (not #76)
-#80 := (ite #77 #37 uf_1)
-#84 := (= #80 #83)
-#85 := [rewrite]: #84
-#81 := (= #47 #80)
-#78 := (iff #11 #77)
-#79 := [rewrite]: #78
-#82 := [monotonicity #79]: #81
-#87 := [trans #82 #85]: #86
-#102 := [monotonicity #87 #99]: #101
-#74 := (= #42 #71)
-#68 := (ite #65 #39 #6)
-#72 := (= #68 #71)
-#73 := [rewrite]: #72
-#69 := (= #42 #68)
-#66 := (iff #8 #65)
-#67 := [rewrite]: #66
-#70 := [monotonicity #67]: #69
-#75 := [trans #70 #73]: #74
-#105 := [monotonicity #75 #102]: #104
-#114 := [trans #105 #112]: #113
-#117 := [monotonicity #114]: #116
-#62 := (iff #19 #61)
-#59 := (iff #18 #58)
-#56 := (= #17 #55)
-#53 := (= #16 #52)
-#50 := (= #15 #38)
-#51 := [rewrite]: #50
-#54 := [monotonicity #51]: #53
-#48 := (= #13 #47)
-#45 := (= #12 #37)
-#46 := [rewrite]: #45
-#49 := [monotonicity #46]: #48
-#57 := [monotonicity #49 #54]: #56
-#43 := (= #10 #42)
-#40 := (= #9 #39)
-#41 := [rewrite]: #40
-#44 := [monotonicity #41]: #43
-#60 := [monotonicity #44 #57]: #59
-#63 := [monotonicity #60]: #62
-#119 := [trans #63 #117]: #118
-#35 := [asserted]: #19
-#120 := [mp #35 #119]: #115
-#185 := [mp #120 #184]: #182
-#248 := (or #177 #64 #233)
-#249 := [def-axiom]: #248
-#265 := [unit-resolution #249 #185]: #264
-#266 := [unit-resolution #265 #262]: #233
-#240 := (or #176 #76 #230)
-#241 := [def-axiom]: #240
-#267 := [unit-resolution #241 #266 #261]: #76
-#268 := [th-lemma #267 #263 #262]: false
-#270 := [lemma #268]: #269
-#272 := [unit-resolution #270 #263]: #64
-#273 := (or #65 #205)
-#246 := (or #177 #65 #205)
-#247 := [def-axiom]: #246
-#274 := [unit-resolution #247 #185]: #273
-#275 := [unit-resolution #274 #272]: #205
-#255 := (or #170 #194 #191)
-#250 := [hypothesis]: #169
-#251 := [hypothesis]: #205
-#252 := [hypothesis]: #166
-#210 := (or #170 #77 #191)
-#211 := [def-axiom]: #210
-#253 := [unit-resolution #211 #251 #252]: #77
-#212 := (or #170 #76 #194)
-#213 := [def-axiom]: #212
-#254 := [unit-resolution #213 #253 #251 #250]: false
-#256 := [lemma #254]: #255
-#276 := [unit-resolution #256 #275 #271]: #194
-#200 := (or #169 #89 #199)
-#201 := [def-axiom]: #200
-#277 := [unit-resolution #201 #276 #263]: #199
-#278 := [unit-resolution #211 #275 #271]: #77
-#279 := [th-lemma #278 #277]: false
-#280 := [lemma #279]: #89
-#281 := [hypothesis]: #77
-#282 := [unit-resolution #241 #281 #261]: #176
-#283 := [unit-resolution #265 #282]: #64
-#284 := [th-lemma #281 #283 #280]: false
-#285 := [lemma #284]: #76
-#222 := (not #172)
-#286 := [hypothesis]: #222
-#287 := [th-lemma #285 #286]: false
-#288 := [lemma #287]: #172
-#223 := (or #173 #88 #222)
-#224 := [def-axiom]: #223
-#289 := [unit-resolution #224 #288 #280]: #173
-#214 := (not #173)
-#238 := (or #176 #77 #214)
-#239 := [def-axiom]: #238
-#290 := [unit-resolution #239 #289 #285]: #176
-#291 := [unit-resolution #265 #290]: #64
-#292 := [unit-resolution #274 #291]: #205
-#293 := [unit-resolution #211 #292 #285]: #191
-#189 := (or #166 #188)
-#190 := [def-axiom]: #189
-#294 := [unit-resolution #190 #293]: #188
-[th-lemma #280 #294]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_linarith_07 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,11 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T2 T1)
-:extrafuns (
- (uf_3 T1)
- (uf_2 Int Int T1)
- (uf_1 T1 T2)
- )
-:assumption (not (= (uf_1 (uf_2 2 3)) (uf_1 uf_3)))
-:assumption (forall (?x1 Int) (?x2 Int) (iff (= (uf_2 ?x1 ?x2) uf_3) (< ?x1 ?x2)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_linarith_07.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,124 +0,0 @@
-#2 := false
-decl uf_1 :: (-> T1 T2)
-decl uf_3 :: T1
-#8 := uf_3
-#9 := (uf_1 uf_3)
-decl uf_2 :: (-> int int T1)
-#5 := 3::int
-#4 := 2::int
-#6 := (uf_2 2::int 3::int)
-#7 := (uf_1 #6)
-#10 := (= #7 #9)
-#225 := (= #6 uf_3)
-#13 := (:var 0 int)
-#12 := (:var 1 int)
-#14 := (uf_2 #12 #13)
-#549 := (pattern #14)
-#52 := 0::int
-#50 := -1::int
-#54 := (* -1::int #13)
-#55 := (+ #12 #54)
-#53 := (>= #55 0::int)
-#51 := (not #53)
-#36 := (= uf_3 #14)
-#61 := (iff #36 #51)
-#550 := (forall (vars (?x1 int) (?x2 int)) (:pat #549) #61)
-#66 := (forall (vars (?x1 int) (?x2 int)) #61)
-#553 := (iff #66 #550)
-#551 := (iff #61 #61)
-#552 := [refl]: #551
-#554 := [quant-intro #552]: #553
-#79 := (~ #66 #66)
-#77 := (~ #61 #61)
-#78 := [refl]: #77
-#80 := [nnf-pos #78]: #79
-#16 := (< #12 #13)
-#15 := (= #14 uf_3)
-#17 := (iff #15 #16)
-#18 := (forall (vars (?x1 int) (?x2 int)) #17)
-#69 := (iff #18 #66)
-#42 := (iff #16 #36)
-#47 := (forall (vars (?x1 int) (?x2 int)) #42)
-#67 := (iff #47 #66)
-#64 := (iff #42 #61)
-#58 := (iff #51 #36)
-#62 := (iff #58 #61)
-#63 := [rewrite]: #62
-#59 := (iff #42 #58)
-#56 := (iff #16 #51)
-#57 := [rewrite]: #56
-#60 := [monotonicity #57]: #59
-#65 := [trans #60 #63]: #64
-#68 := [quant-intro #65]: #67
-#48 := (iff #18 #47)
-#45 := (iff #17 #42)
-#39 := (iff #36 #16)
-#43 := (iff #39 #42)
-#44 := [rewrite]: #43
-#40 := (iff #17 #39)
-#37 := (iff #15 #36)
-#38 := [rewrite]: #37
-#41 := [monotonicity #38]: #40
-#46 := [trans #41 #44]: #45
-#49 := [quant-intro #46]: #48
-#70 := [trans #49 #68]: #69
-#35 := [asserted]: #18
-#71 := [mp #35 #70]: #66
-#74 := [mp~ #71 #80]: #66
-#555 := [mp #74 #554]: #550
-#529 := (not #550)
-#530 := (or #529 #225)
-#220 := (* -1::int 3::int)
-#221 := (+ 2::int #220)
-#222 := (>= #221 0::int)
-#213 := (not #222)
-#135 := (= uf_3 #6)
-#224 := (iff #135 #213)
-#525 := (or #529 #224)
-#169 := (iff #525 #530)
-#534 := (iff #530 #530)
-#174 := [rewrite]: #534
-#527 := (iff #224 #225)
-#1 := true
-#187 := (iff #225 true)
-#190 := (iff #187 #225)
-#526 := [rewrite]: #190
-#188 := (iff #224 #187)
-#183 := (iff #213 true)
-#198 := (not false)
-#199 := (iff #198 true)
-#540 := [rewrite]: #199
-#203 := (iff #213 #198)
-#548 := (iff #222 false)
-#544 := (>= -1::int 0::int)
-#547 := (iff #544 false)
-#542 := [rewrite]: #547
-#545 := (iff #222 #544)
-#211 := (= #221 -1::int)
-#223 := -3::int
-#541 := (+ 2::int -3::int)
-#330 := (= #541 -1::int)
-#537 := [rewrite]: #330
-#543 := (= #221 #541)
-#227 := (= #220 -3::int)
-#206 := [rewrite]: #227
-#200 := [monotonicity #206]: #543
-#212 := [trans #200 #537]: #211
-#546 := [monotonicity #212]: #545
-#538 := [trans #546 #542]: #548
-#539 := [monotonicity #538]: #203
-#524 := [trans #539 #540]: #183
-#153 := (iff #135 #225)
-#226 := [rewrite]: #153
-#189 := [monotonicity #226 #524]: #188
-#528 := [trans #189 #526]: #527
-#532 := [monotonicity #528]: #169
-#175 := [trans #532 #174]: #169
-#531 := [quant-inst]: #525
-#535 := [mp #531 #175]: #530
-#533 := [unit-resolution #535 #555]: #225
-#536 := [monotonicity #533]: #10
-#11 := (not #10)
-#34 := [asserted]: #11
-[unit-resolution #34 #536]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_linarith_08 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,7 +0,0 @@
-(benchmark Isabelle
-:extrafuns (
- (uf_1 Int)
- )
-:assumption (not (or (<= 4 (+ uf_1 3)) (< uf_1 1)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_linarith_08.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,54 +0,0 @@
-#2 := false
-#9 := 1::int
-decl uf_1 :: int
-#5 := uf_1
-#10 := (< uf_1 1::int)
-#6 := 3::int
-#7 := (+ uf_1 3::int)
-#4 := 4::int
-#8 := (<= 4::int #7)
-#11 := (or #8 #10)
-#12 := (not #11)
-#66 := (iff #12 false)
-#29 := (+ 3::int uf_1)
-#32 := (<= 4::int #29)
-#38 := (or #10 #32)
-#43 := (not #38)
-#64 := (iff #43 false)
-#1 := true
-#59 := (not true)
-#62 := (iff #59 false)
-#63 := [rewrite]: #62
-#60 := (iff #43 #59)
-#57 := (iff #38 true)
-#48 := (>= uf_1 1::int)
-#46 := (not #48)
-#52 := (or #46 #48)
-#55 := (iff #52 true)
-#56 := [rewrite]: #55
-#53 := (iff #38 #52)
-#50 := (iff #32 #48)
-#51 := [rewrite]: #50
-#47 := (iff #10 #46)
-#49 := [rewrite]: #47
-#54 := [monotonicity #49 #51]: #53
-#58 := [trans #54 #56]: #57
-#61 := [monotonicity #58]: #60
-#65 := [trans #61 #63]: #64
-#44 := (iff #12 #43)
-#41 := (iff #11 #38)
-#35 := (or #32 #10)
-#39 := (iff #35 #38)
-#40 := [rewrite]: #39
-#36 := (iff #11 #35)
-#33 := (iff #8 #32)
-#30 := (= #7 #29)
-#31 := [rewrite]: #30
-#34 := [monotonicity #31]: #33
-#37 := [monotonicity #34]: #36
-#42 := [trans #37 #40]: #41
-#45 := [monotonicity #42]: #44
-#67 := [trans #45 #65]: #66
-#28 := [asserted]: #12
-[mp #28 #67]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_linarith_09 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,10 +0,0 @@
-(benchmark Isabelle
-:extrafuns (
- (uf_1 Int)
- (uf_2 Int)
- )
-:assumption (<= 3 uf_1)
-:assumption (= uf_2 (+ uf_1 4))
-:assumption (not (< 0 (- uf_2 uf_1)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_linarith_09.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,63 +0,0 @@
-#2 := false
-#11 := 0::int
-decl uf_2 :: int
-#7 := uf_2
-#42 := -1::int
-#45 := (* -1::int uf_2)
-decl uf_1 :: int
-#5 := uf_1
-#46 := (+ uf_1 #45)
-#63 := (>= #46 0::int)
-#83 := (iff #63 false)
-#44 := -4::int
-#79 := (>= -4::int 0::int)
-#81 := (iff #79 false)
-#82 := [rewrite]: #81
-#77 := (iff #63 #79)
-#47 := (= #46 -4::int)
-#8 := 4::int
-#9 := (+ uf_1 4::int)
-#10 := (= uf_2 #9)
-#49 := (iff #10 #47)
-#32 := (+ 4::int uf_1)
-#39 := (= uf_2 #32)
-#43 := (iff #39 #47)
-#48 := [rewrite]: #43
-#40 := (iff #10 #39)
-#37 := (= #9 #32)
-#38 := [rewrite]: #37
-#41 := [monotonicity #38]: #40
-#50 := [trans #41 #48]: #49
-#31 := [asserted]: #10
-#51 := [mp #31 #50]: #47
-#80 := [monotonicity #51]: #77
-#84 := [trans #80 #82]: #83
-#12 := (- uf_2 uf_1)
-#13 := (< 0::int #12)
-#14 := (not #13)
-#74 := (iff #14 #63)
-#53 := (* -1::int uf_1)
-#54 := (+ #53 uf_2)
-#57 := (< 0::int #54)
-#60 := (not #57)
-#72 := (iff #60 #63)
-#64 := (not #63)
-#67 := (not #64)
-#70 := (iff #67 #63)
-#71 := [rewrite]: #70
-#68 := (iff #60 #67)
-#65 := (iff #57 #64)
-#66 := [rewrite]: #65
-#69 := [monotonicity #66]: #68
-#73 := [trans #69 #71]: #72
-#61 := (iff #14 #60)
-#58 := (iff #13 #57)
-#55 := (= #12 #54)
-#56 := [rewrite]: #55
-#59 := [monotonicity #56]: #58
-#62 := [monotonicity #59]: #61
-#75 := [trans #62 #73]: #74
-#52 := [asserted]: #14
-#76 := [mp #52 #75]: #63
-[mp #76 #84]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_linarith_10 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not (let (?x1 2) (not (= (+ ?x1 ?x1) 5))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_linarith_10.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,35 +0,0 @@
-#2 := false
-#6 := 5::int
-#4 := 2::int
-#5 := (+ 2::int 2::int)
-#7 := (= #5 5::int)
-#8 := (not #7)
-#9 := (not #8)
-#48 := (iff #9 false)
-#1 := true
-#43 := (not true)
-#46 := (iff #43 false)
-#47 := [rewrite]: #46
-#44 := (iff #9 #43)
-#41 := (iff #8 true)
-#36 := (not false)
-#39 := (iff #36 true)
-#40 := [rewrite]: #39
-#37 := (iff #8 #36)
-#34 := (iff #7 false)
-#26 := 4::int
-#29 := (= 4::int 5::int)
-#32 := (iff #29 false)
-#33 := [rewrite]: #32
-#30 := (iff #7 #29)
-#27 := (= #5 4::int)
-#28 := [rewrite]: #27
-#31 := [monotonicity #28]: #30
-#35 := [trans #31 #33]: #34
-#38 := [monotonicity #35]: #37
-#42 := [trans #38 #40]: #41
-#45 := [monotonicity #42]: #44
-#49 := [trans #45 #47]: #48
-#25 := [asserted]: #9
-[mp #25 #49]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_linarith_11 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,10 +0,0 @@
-(benchmark Isabelle
-:extrafuns (
- (uf_2 Real)
- (uf_1 Real)
- )
-:assumption (< (+ (* 3.0 uf_1) (* 7.0 uf_2)) 4.0)
-:assumption (< 3.0 (* 2.0 uf_1))
-:assumption (not (< uf_2 0.0))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_linarith_11.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,45 +0,0 @@
-#2 := false
-#11 := 4::real
-decl uf_2 :: real
-#8 := uf_2
-#7 := 7::real
-#9 := (* 7::real uf_2)
-decl uf_1 :: real
-#5 := uf_1
-#4 := 3::real
-#6 := (* 3::real uf_1)
-#10 := (+ #6 #9)
-#41 := (>= #10 4::real)
-#39 := (not #41)
-#12 := (< #10 4::real)
-#40 := (iff #12 #39)
-#37 := [rewrite]: #40
-#34 := [asserted]: #12
-#38 := [mp #34 #37]: #39
-#13 := 2::real
-#14 := (* 2::real uf_1)
-#43 := (<= #14 3::real)
-#44 := (not #43)
-#15 := (< 3::real #14)
-#45 := (iff #15 #44)
-#46 := [rewrite]: #45
-#35 := [asserted]: #15
-#47 := [mp #35 #46]: #44
-#16 := 0::real
-#51 := (>= uf_2 0::real)
-#17 := (< uf_2 0::real)
-#18 := (not #17)
-#58 := (iff #18 #51)
-#49 := (not #51)
-#53 := (not #49)
-#56 := (iff #53 #51)
-#57 := [rewrite]: #56
-#54 := (iff #18 #53)
-#50 := (iff #17 #49)
-#52 := [rewrite]: #50
-#55 := [monotonicity #52]: #54
-#59 := [trans #55 #57]: #58
-#36 := [asserted]: #18
-#60 := [mp #36 #59]: #51
-[th-lemma #60 #47 #38]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_linarith_12 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,8 +0,0 @@
-(benchmark Isabelle
-:extrafuns (
- (uf_2 Int)
- (uf_1 Int)
- )
-:assumption (not (iff (or (<= 0 (+ uf_1 (* (~ 1) uf_2))) (or (not (<= 0 uf_2)) (<= 0 uf_2))) (not false)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_linarith_12.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,59 +0,0 @@
-#2 := false
-#16 := (not false)
-decl uf_2 :: int
-#8 := uf_2
-#4 := 0::int
-#12 := (<= 0::int uf_2)
-#13 := (not #12)
-#14 := (or #13 #12)
-#6 := 1::int
-#7 := (- 1::int)
-#9 := (* #7 uf_2)
-decl uf_1 :: int
-#5 := uf_1
-#10 := (+ uf_1 #9)
-#11 := (<= 0::int #10)
-#15 := (or #11 #14)
-#17 := (iff #15 #16)
-#18 := (not #17)
-#70 := (iff #18 false)
-#1 := true
-#65 := (not true)
-#68 := (iff #65 false)
-#69 := [rewrite]: #68
-#66 := (iff #18 #65)
-#63 := (iff #17 true)
-#58 := (iff true true)
-#61 := (iff #58 true)
-#62 := [rewrite]: #61
-#59 := (iff #17 #58)
-#56 := (iff #16 true)
-#57 := [rewrite]: #56
-#54 := (iff #15 true)
-#35 := -1::int
-#38 := (* -1::int uf_2)
-#41 := (+ uf_1 #38)
-#44 := (<= 0::int #41)
-#49 := (or #44 true)
-#52 := (iff #49 true)
-#53 := [rewrite]: #52
-#50 := (iff #15 #49)
-#47 := (iff #14 true)
-#48 := [rewrite]: #47
-#45 := (iff #11 #44)
-#42 := (= #10 #41)
-#39 := (= #9 #38)
-#36 := (= #7 -1::int)
-#37 := [rewrite]: #36
-#40 := [monotonicity #37]: #39
-#43 := [monotonicity #40]: #42
-#46 := [monotonicity #43]: #45
-#51 := [monotonicity #46 #48]: #50
-#55 := [trans #51 #53]: #54
-#60 := [monotonicity #55 #57]: #59
-#64 := [trans #60 #62]: #63
-#67 := [monotonicity #64]: #66
-#71 := [trans #67 #69]: #70
-#34 := [asserted]: #18
-[mp #34 #71]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_linarith_13 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,13 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T1)
-:extrafuns (
- (uf_4 T1)
- (uf_1 Int Int T1)
- (uf_3 Int Int T1)
- (uf_2 Int)
- )
-:assumption (not (distinct (uf_1 uf_2 3) (uf_3 3 uf_2)))
-:assumption (forall (?x1 Int) (?x2 Int) (iff (= (uf_3 ?x1 ?x2) uf_4) (<= ?x1 ?x2)))
-:assumption (forall (?x3 Int) (?x4 Int) (iff (= (uf_1 ?x3 ?x4) uf_4) (< ?x3 ?x4)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_linarith_13.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,212 +0,0 @@
-#2 := false
-decl uf_4 :: T1
-#13 := uf_4
-decl uf_1 :: (-> int int T1)
-#5 := 3::int
-decl uf_2 :: int
-#4 := uf_2
-#6 := (uf_1 uf_2 3::int)
-#559 := (= #6 uf_4)
-decl uf_3 :: (-> int int T1)
-#7 := (uf_3 3::int uf_2)
-#254 := (= #7 uf_4)
-#524 := (iff #254 #559)
-#529 := (iff #559 #254)
-#39 := (= #6 #7)
-#8 := (distinct #6 #7)
-#9 := (not #8)
-#48 := (iff #9 #39)
-#40 := (not #39)
-#43 := (not #40)
-#46 := (iff #43 #39)
-#47 := [rewrite]: #46
-#44 := (iff #9 #43)
-#41 := (iff #8 #40)
-#42 := [rewrite]: #41
-#45 := [monotonicity #42]: #44
-#49 := [trans #45 #47]: #48
-#38 := [asserted]: #9
-#52 := [mp #38 #49]: #39
-#523 := [monotonicity #52]: #529
-#530 := [symm #523]: #524
-#547 := (not #559)
-#570 := (not #254)
-#531 := (iff #570 #547)
-#525 := [monotonicity #530]: #531
-#540 := [hypothesis]: #570
-#532 := [mp #540 #525]: #547
-#256 := (>= uf_2 3::int)
-#579 := (not #256)
-#541 := (or #254 #579)
-#258 := (iff #254 #256)
-#11 := (:var 0 int)
-#10 := (:var 1 int)
-#12 := (uf_3 #10 #11)
-#581 := (pattern #12)
-#57 := 0::int
-#54 := -1::int
-#55 := (* -1::int #11)
-#56 := (+ #10 #55)
-#58 := (<= #56 0::int)
-#14 := (= #12 uf_4)
-#61 := (iff #14 #58)
-#582 := (forall (vars (?x1 int) (?x2 int)) (:pat #581) #61)
-#64 := (forall (vars (?x1 int) (?x2 int)) #61)
-#585 := (iff #64 #582)
-#583 := (iff #61 #61)
-#584 := [refl]: #583
-#586 := [quant-intro #584]: #585
-#108 := (~ #64 #64)
-#106 := (~ #61 #61)
-#107 := [refl]: #106
-#109 := [nnf-pos #107]: #108
-#15 := (<= #10 #11)
-#16 := (iff #14 #15)
-#17 := (forall (vars (?x1 int) (?x2 int)) #16)
-#65 := (iff #17 #64)
-#62 := (iff #16 #61)
-#59 := (iff #15 #58)
-#60 := [rewrite]: #59
-#63 := [monotonicity #60]: #62
-#66 := [quant-intro #63]: #65
-#50 := [asserted]: #17
-#67 := [mp #50 #66]: #64
-#101 := [mp~ #67 #109]: #64
-#587 := [mp #101 #586]: #582
-#238 := (not #582)
-#573 := (or #238 #258)
-#167 := (* -1::int uf_2)
-#252 := (+ 3::int #167)
-#253 := (<= #252 0::int)
-#245 := (iff #254 #253)
-#575 := (or #238 #245)
-#362 := (iff #575 #573)
-#243 := (iff #573 #573)
-#244 := [rewrite]: #243
-#255 := (iff #245 #258)
-#257 := (iff #253 #256)
-#185 := [rewrite]: #257
-#259 := [monotonicity #185]: #255
-#569 := [monotonicity #259]: #362
-#576 := [trans #569 #244]: #362
-#232 := [quant-inst]: #575
-#577 := [mp #232 #576]: #573
-#535 := [unit-resolution #577 #587]: #258
-#578 := (not #258)
-#574 := (or #578 #254 #579)
-#580 := [def-axiom]: #574
-#382 := [unit-resolution #580 #535]: #541
-#383 := [unit-resolution #382 #540]: #579
-#526 := (or #559 #256)
-#273 := (iff #559 #579)
-#18 := (uf_1 #10 #11)
-#588 := (pattern #18)
-#82 := (>= #56 0::int)
-#81 := (not #82)
-#53 := (= uf_4 #18)
-#88 := (iff #53 #81)
-#589 := (forall (vars (?x3 int) (?x4 int)) (:pat #588) #88)
-#93 := (forall (vars (?x3 int) (?x4 int)) #88)
-#592 := (iff #93 #589)
-#590 := (iff #88 #88)
-#591 := [refl]: #590
-#593 := [quant-intro #591]: #592
-#102 := (~ #93 #93)
-#99 := (~ #88 #88)
-#110 := [refl]: #99
-#103 := [nnf-pos #110]: #102
-#20 := (< #10 #11)
-#19 := (= #18 uf_4)
-#21 := (iff #19 #20)
-#22 := (forall (vars (?x3 int) (?x4 int)) #21)
-#96 := (iff #22 #93)
-#73 := (iff #20 #53)
-#78 := (forall (vars (?x3 int) (?x4 int)) #73)
-#94 := (iff #78 #93)
-#91 := (iff #73 #88)
-#85 := (iff #81 #53)
-#89 := (iff #85 #88)
-#90 := [rewrite]: #89
-#86 := (iff #73 #85)
-#83 := (iff #20 #81)
-#84 := [rewrite]: #83
-#87 := [monotonicity #84]: #86
-#92 := [trans #87 #90]: #91
-#95 := [quant-intro #92]: #94
-#79 := (iff #22 #78)
-#76 := (iff #21 #73)
-#70 := (iff #53 #20)
-#74 := (iff #70 #73)
-#75 := [rewrite]: #74
-#71 := (iff #21 #70)
-#68 := (iff #19 #53)
-#69 := [rewrite]: #68
-#72 := [monotonicity #69]: #71
-#77 := [trans #72 #75]: #76
-#80 := [quant-intro #77]: #79
-#97 := [trans #80 #95]: #96
-#51 := [asserted]: #22
-#98 := [mp #51 #97]: #93
-#111 := [mp~ #98 #103]: #93
-#594 := [mp #111 #593]: #589
-#552 := (not #589)
-#549 := (or #552 #273)
-#219 := (* -1::int 3::int)
-#220 := (+ uf_2 #219)
-#221 := (>= #220 0::int)
-#222 := (not #221)
-#556 := (= uf_4 #6)
-#558 := (iff #556 #222)
-#553 := (or #552 #558)
-#264 := (iff #553 #549)
-#266 := (iff #549 #549)
-#544 := [rewrite]: #266
-#274 := (iff #558 #273)
-#550 := (iff #222 #579)
-#280 := (iff #221 #256)
-#562 := -3::int
-#206 := (+ -3::int uf_2)
-#554 := (>= #206 0::int)
-#278 := (iff #554 #256)
-#279 := [rewrite]: #278
-#555 := (iff #221 #554)
-#565 := (= #220 #206)
-#201 := (+ uf_2 -3::int)
-#207 := (= #201 #206)
-#567 := [rewrite]: #207
-#564 := (= #220 #201)
-#557 := (= #219 -3::int)
-#563 := [rewrite]: #557
-#566 := [monotonicity #563]: #564
-#568 := [trans #566 #567]: #565
-#277 := [monotonicity #568]: #555
-#173 := [trans #277 #279]: #280
-#551 := [monotonicity #173]: #550
-#560 := (iff #556 #559)
-#561 := [rewrite]: #560
-#548 := [monotonicity #561 #551]: #274
-#265 := [monotonicity #548]: #264
-#545 := [trans #265 #544]: #264
-#263 := [quant-inst]: #553
-#260 := [mp #263 #545]: #549
-#384 := [unit-resolution #260 #594]: #273
-#542 := (not #273)
-#546 := (or #542 #559 #256)
-#543 := [def-axiom]: #546
-#527 := [unit-resolution #543 #384]: #526
-#528 := [unit-resolution #527 #383]: #559
-#361 := [unit-resolution #528 #532]: false
-#363 := [lemma #361]: #254
-#522 := [mp #363 #530]: #559
-#364 := (or #570 #256)
-#230 := (or #578 #570 #256)
-#235 := [def-axiom]: #230
-#517 := [unit-resolution #235 #535]: #364
-#518 := [unit-resolution #517 #363]: #256
-#520 := (or #547 #579)
-#536 := (or #542 #547 #579)
-#537 := [def-axiom]: #536
-#521 := [unit-resolution #537 #384]: #520
-#519 := [unit-resolution #521 #518]: #547
-[unit-resolution #519 #522]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_linarith_14 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,8 +0,0 @@
-(benchmark Isabelle
-:extrafuns (
- (uf_1 Int)
- )
-:assumption (< 0 uf_1)
-:assumption (not (distinct uf_1 (* uf_1 2) (- uf_1 uf_1)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_linarith_14.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,86 +0,0 @@
-#2 := false
-decl uf_1 :: int
-#5 := uf_1
-#7 := 2::int
-#29 := (* 2::int uf_1)
-#4 := 0::int
-#54 := (= 0::int #29)
-#55 := (not #54)
-#61 := (= #29 0::int)
-#104 := (not #61)
-#110 := (iff #104 #55)
-#108 := (iff #61 #54)
-#109 := [commutativity]: #108
-#111 := [monotonicity #109]: #110
-#62 := (<= #29 0::int)
-#100 := (not #62)
-#30 := (<= uf_1 0::int)
-#31 := (not #30)
-#6 := (< 0::int uf_1)
-#32 := (iff #6 #31)
-#33 := [rewrite]: #32
-#27 := [asserted]: #6
-#34 := [mp #27 #33]: #31
-#101 := (or #100 #30)
-#102 := [th-lemma]: #101
-#103 := [unit-resolution #102 #34]: #100
-#105 := (or #104 #62)
-#106 := [th-lemma]: #105
-#107 := [unit-resolution #106 #103]: #104
-#112 := [mp #107 #111]: #55
-#56 := (= uf_1 #29)
-#57 := (not #56)
-#53 := (= 0::int uf_1)
-#50 := (not #53)
-#58 := (and #50 #55 #57)
-#69 := (not #58)
-#42 := (distinct 0::int uf_1 #29)
-#47 := (not #42)
-#9 := (- uf_1 uf_1)
-#8 := (* uf_1 2::int)
-#10 := (distinct uf_1 #8 #9)
-#11 := (not #10)
-#48 := (iff #11 #47)
-#45 := (iff #10 #42)
-#39 := (distinct uf_1 #29 0::int)
-#43 := (iff #39 #42)
-#44 := [rewrite]: #43
-#40 := (iff #10 #39)
-#37 := (= #9 0::int)
-#38 := [rewrite]: #37
-#35 := (= #8 #29)
-#36 := [rewrite]: #35
-#41 := [monotonicity #36 #38]: #40
-#46 := [trans #41 #44]: #45
-#49 := [monotonicity #46]: #48
-#28 := [asserted]: #11
-#52 := [mp #28 #49]: #47
-#80 := (or #42 #69)
-#81 := [def-axiom]: #80
-#82 := [unit-resolution #81 #52]: #69
-#59 := (= uf_1 0::int)
-#83 := (not #59)
-#89 := (iff #83 #50)
-#87 := (iff #59 #53)
-#88 := [commutativity]: #87
-#90 := [monotonicity #88]: #89
-#84 := (or #83 #30)
-#85 := [th-lemma]: #84
-#86 := [unit-resolution #85 #34]: #83
-#91 := [mp #86 #90]: #50
-#64 := -1::int
-#65 := (* -1::int #29)
-#66 := (+ uf_1 #65)
-#68 := (>= #66 0::int)
-#92 := (not #68)
-#93 := (or #92 #30)
-#94 := [th-lemma]: #93
-#95 := [unit-resolution #94 #34]: #92
-#96 := (or #57 #68)
-#97 := [th-lemma]: #96
-#98 := [unit-resolution #97 #95]: #57
-#76 := (or #58 #53 #54 #56)
-#77 := [def-axiom]: #76
-#99 := [unit-resolution #77 #98 #91 #82]: #54
-[unit-resolution #99 #112]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_linarith_15 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,9 +0,0 @@
-(benchmark Isabelle
-:extrafuns (
- (uf_2 Int)
- (uf_1 Int)
- (uf_3 Int)
- )
-:assumption (not (or (and (< uf_1 uf_2) (< uf_2 uf_3)) (or (and (< uf_1 uf_2) (= uf_2 uf_3)) (or (and (< uf_1 uf_3) (< uf_3 uf_2)) (or (and (= uf_1 uf_3) (< uf_3 uf_2)) (or (and (= uf_1 uf_2) (< uf_2 uf_3)) (or (and (< uf_3 uf_2) (< uf_2 uf_1)) (or (and (< uf_3 uf_2) (= uf_2 uf_1)) (or (and (< uf_3 uf_1) (< uf_1 uf_2)) (or (and (= uf_3 uf_1) (< uf_1 uf_2)) (or (and (= uf_3 uf_2) (< uf_2 uf_1)) (or (and (< uf_2 uf_1) (< uf_1 uf_3)) (or (and (< uf_2 uf_1) (= uf_3 uf_1)) (or (and (< uf_2 uf_3) (< uf_3 uf_1)) (or (and (= uf_2 uf_1) (< uf_1 uf_3)) (or (and (= uf_2 uf_3) (< uf_3 uf_1)) (and (= uf_3 uf_2) (= uf_2 uf_1))))))))))))))))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_linarith_15.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,673 +0,0 @@
-#2 := false
-#169 := 0::int
-decl uf_2 :: int
-#5 := uf_2
-#166 := -1::int
-#202 := (* -1::int uf_2)
-decl uf_1 :: int
-#4 := uf_1
-#203 := (+ uf_1 #202)
-#218 := (>= #203 0::int)
-decl uf_3 :: int
-#7 := uf_3
-#167 := (* -1::int uf_3)
-#168 := (+ uf_1 #167)
-#178 := (>= #168 0::int)
-#217 := (not #218)
-#204 := (<= #203 0::int)
-#205 := (not #204)
-#692 := [hypothesis]: #205
-#177 := (not #178)
-#693 := (or #177 #204)
-#170 := (<= #168 0::int)
-#191 := (+ uf_2 #167)
-#237 := (<= #191 0::int)
-#238 := (not #237)
-#171 := (not #170)
-#685 := [hypothesis]: #171
-#190 := (>= #191 0::int)
-#455 := (or #170 #190)
-#189 := (not #190)
-#197 := (and #171 #189)
-#354 := (not #197)
-#464 := (iff #354 #455)
-#456 := (not #455)
-#459 := (not #456)
-#462 := (iff #459 #455)
-#463 := [rewrite]: #462
-#460 := (iff #354 #459)
-#457 := (iff #197 #456)
-#458 := [rewrite]: #457
-#461 := [monotonicity #458]: #460
-#465 := [trans #461 #463]: #464
-#287 := (and #189 #217)
-#10 := (= uf_2 uf_3)
-#279 := (and #10 #217)
-#273 := (and #177 #238)
-#15 := (= uf_1 uf_3)
-#268 := (and #15 #238)
-#17 := (= uf_1 uf_2)
-#260 := (and #17 #189)
-#252 := (and #205 #238)
-#244 := (and #17 #238)
-#232 := (and #171 #217)
-#224 := (and #15 #217)
-#214 := (and #10 #205)
-#211 := (and #177 #205)
-#208 := (and #15 #205)
-#184 := (and #17 #177)
-#174 := (and #10 #171)
-#115 := (and #10 #17)
-#337 := (or #115 #174 #184 #197 #208 #211 #214 #224 #232 #244 #252 #260 #268 #273 #279 #287)
-#342 := (not #337)
-#21 := (= uf_2 uf_1)
-#27 := (= uf_3 uf_2)
-#34 := (and #27 #21)
-#23 := (< uf_3 uf_1)
-#33 := (and #10 #23)
-#35 := (or #33 #34)
-#12 := (< uf_1 uf_3)
-#32 := (and #21 #12)
-#36 := (or #32 #35)
-#8 := (< uf_2 uf_3)
-#31 := (and #8 #23)
-#37 := (or #31 #36)
-#25 := (= uf_3 uf_1)
-#19 := (< uf_2 uf_1)
-#30 := (and #19 #25)
-#38 := (or #30 #37)
-#29 := (and #19 #12)
-#39 := (or #29 #38)
-#28 := (and #27 #19)
-#40 := (or #28 #39)
-#6 := (< uf_1 uf_2)
-#26 := (and #25 #6)
-#41 := (or #26 #40)
-#24 := (and #23 #6)
-#42 := (or #24 #41)
-#13 := (< uf_3 uf_2)
-#22 := (and #13 #21)
-#43 := (or #22 #42)
-#20 := (and #13 #19)
-#44 := (or #20 #43)
-#18 := (and #17 #8)
-#45 := (or #18 #44)
-#16 := (and #15 #13)
-#46 := (or #16 #45)
-#14 := (and #12 #13)
-#47 := (or #14 #46)
-#11 := (and #6 #10)
-#48 := (or #11 #47)
-#9 := (and #6 #8)
-#49 := (or #9 #48)
-#50 := (not #49)
-#345 := (iff #50 #342)
-#118 := (or #33 #115)
-#110 := (and #12 #17)
-#121 := (or #110 #118)
-#124 := (or #31 #121)
-#102 := (and #15 #19)
-#127 := (or #102 #124)
-#96 := (and #12 #19)
-#130 := (or #96 #127)
-#93 := (and #10 #19)
-#133 := (or #93 #130)
-#86 := (and #6 #15)
-#136 := (or #86 #133)
-#78 := (and #6 #23)
-#139 := (or #78 #136)
-#75 := (and #13 #17)
-#142 := (or #75 #139)
-#145 := (or #20 #142)
-#70 := (and #8 #17)
-#148 := (or #70 #145)
-#67 := (and #13 #15)
-#151 := (or #67 #148)
-#154 := (or #14 #151)
-#157 := (or #11 #154)
-#160 := (or #9 #157)
-#163 := (not #160)
-#343 := (iff #163 #342)
-#340 := (iff #160 #337)
-#292 := (or #174 #115)
-#295 := (or #184 #292)
-#298 := (or #197 #295)
-#301 := (or #208 #298)
-#304 := (or #211 #301)
-#307 := (or #214 #304)
-#310 := (or #224 #307)
-#313 := (or #232 #310)
-#316 := (or #244 #313)
-#319 := (or #252 #316)
-#322 := (or #260 #319)
-#325 := (or #268 #322)
-#328 := (or #273 #325)
-#331 := (or #279 #328)
-#334 := (or #287 #331)
-#338 := (iff #334 #337)
-#339 := [rewrite]: #338
-#335 := (iff #160 #334)
-#332 := (iff #157 #331)
-#329 := (iff #154 #328)
-#326 := (iff #151 #325)
-#323 := (iff #148 #322)
-#320 := (iff #145 #319)
-#317 := (iff #142 #316)
-#314 := (iff #139 #313)
-#311 := (iff #136 #310)
-#308 := (iff #133 #307)
-#305 := (iff #130 #304)
-#302 := (iff #127 #301)
-#299 := (iff #124 #298)
-#296 := (iff #121 #295)
-#293 := (iff #118 #292)
-#175 := (iff #33 #174)
-#172 := (iff #23 #171)
-#173 := [rewrite]: #172
-#176 := [monotonicity #173]: #175
-#294 := [monotonicity #176]: #293
-#187 := (iff #110 #184)
-#181 := (and #177 #17)
-#185 := (iff #181 #184)
-#186 := [rewrite]: #185
-#182 := (iff #110 #181)
-#179 := (iff #12 #177)
-#180 := [rewrite]: #179
-#183 := [monotonicity #180]: #182
-#188 := [trans #183 #186]: #187
-#297 := [monotonicity #188 #294]: #296
-#200 := (iff #31 #197)
-#194 := (and #189 #171)
-#198 := (iff #194 #197)
-#199 := [rewrite]: #198
-#195 := (iff #31 #194)
-#192 := (iff #8 #189)
-#193 := [rewrite]: #192
-#196 := [monotonicity #193 #173]: #195
-#201 := [trans #196 #199]: #200
-#300 := [monotonicity #201 #297]: #299
-#209 := (iff #102 #208)
-#206 := (iff #19 #205)
-#207 := [rewrite]: #206
-#210 := [monotonicity #207]: #209
-#303 := [monotonicity #210 #300]: #302
-#212 := (iff #96 #211)
-#213 := [monotonicity #180 #207]: #212
-#306 := [monotonicity #213 #303]: #305
-#215 := (iff #93 #214)
-#216 := [monotonicity #207]: #215
-#309 := [monotonicity #216 #306]: #308
-#227 := (iff #86 #224)
-#221 := (and #217 #15)
-#225 := (iff #221 #224)
-#226 := [rewrite]: #225
-#222 := (iff #86 #221)
-#219 := (iff #6 #217)
-#220 := [rewrite]: #219
-#223 := [monotonicity #220]: #222
-#228 := [trans #223 #226]: #227
-#312 := [monotonicity #228 #309]: #311
-#235 := (iff #78 #232)
-#229 := (and #217 #171)
-#233 := (iff #229 #232)
-#234 := [rewrite]: #233
-#230 := (iff #78 #229)
-#231 := [monotonicity #220 #173]: #230
-#236 := [trans #231 #234]: #235
-#315 := [monotonicity #236 #312]: #314
-#247 := (iff #75 #244)
-#241 := (and #238 #17)
-#245 := (iff #241 #244)
-#246 := [rewrite]: #245
-#242 := (iff #75 #241)
-#239 := (iff #13 #238)
-#240 := [rewrite]: #239
-#243 := [monotonicity #240]: #242
-#248 := [trans #243 #246]: #247
-#318 := [monotonicity #248 #315]: #317
-#255 := (iff #20 #252)
-#249 := (and #238 #205)
-#253 := (iff #249 #252)
-#254 := [rewrite]: #253
-#250 := (iff #20 #249)
-#251 := [monotonicity #240 #207]: #250
-#256 := [trans #251 #254]: #255
-#321 := [monotonicity #256 #318]: #320
-#263 := (iff #70 #260)
-#257 := (and #189 #17)
-#261 := (iff #257 #260)
-#262 := [rewrite]: #261
-#258 := (iff #70 #257)
-#259 := [monotonicity #193]: #258
-#264 := [trans #259 #262]: #263
-#324 := [monotonicity #264 #321]: #323
-#271 := (iff #67 #268)
-#265 := (and #238 #15)
-#269 := (iff #265 #268)
-#270 := [rewrite]: #269
-#266 := (iff #67 #265)
-#267 := [monotonicity #240]: #266
-#272 := [trans #267 #270]: #271
-#327 := [monotonicity #272 #324]: #326
-#274 := (iff #14 #273)
-#275 := [monotonicity #180 #240]: #274
-#330 := [monotonicity #275 #327]: #329
-#282 := (iff #11 #279)
-#276 := (and #217 #10)
-#280 := (iff #276 #279)
-#281 := [rewrite]: #280
-#277 := (iff #11 #276)
-#278 := [monotonicity #220]: #277
-#283 := [trans #278 #281]: #282
-#333 := [monotonicity #283 #330]: #332
-#290 := (iff #9 #287)
-#284 := (and #217 #189)
-#288 := (iff #284 #287)
-#289 := [rewrite]: #288
-#285 := (iff #9 #284)
-#286 := [monotonicity #220 #193]: #285
-#291 := [trans #286 #289]: #290
-#336 := [monotonicity #291 #333]: #335
-#341 := [trans #336 #339]: #340
-#344 := [monotonicity #341]: #343
-#164 := (iff #50 #163)
-#161 := (iff #49 #160)
-#158 := (iff #48 #157)
-#155 := (iff #47 #154)
-#152 := (iff #46 #151)
-#149 := (iff #45 #148)
-#146 := (iff #44 #145)
-#143 := (iff #43 #142)
-#140 := (iff #42 #139)
-#137 := (iff #41 #136)
-#134 := (iff #40 #133)
-#131 := (iff #39 #130)
-#128 := (iff #38 #127)
-#125 := (iff #37 #124)
-#122 := (iff #36 #121)
-#119 := (iff #35 #118)
-#116 := (iff #34 #115)
-#73 := (iff #21 #17)
-#74 := [rewrite]: #73
-#91 := (iff #27 #10)
-#92 := [rewrite]: #91
-#117 := [monotonicity #92 #74]: #116
-#120 := [monotonicity #117]: #119
-#113 := (iff #32 #110)
-#107 := (and #17 #12)
-#111 := (iff #107 #110)
-#112 := [rewrite]: #111
-#108 := (iff #32 #107)
-#109 := [monotonicity #74]: #108
-#114 := [trans #109 #112]: #113
-#123 := [monotonicity #114 #120]: #122
-#126 := [monotonicity #123]: #125
-#105 := (iff #30 #102)
-#99 := (and #19 #15)
-#103 := (iff #99 #102)
-#104 := [rewrite]: #103
-#100 := (iff #30 #99)
-#81 := (iff #25 #15)
-#82 := [rewrite]: #81
-#101 := [monotonicity #82]: #100
-#106 := [trans #101 #104]: #105
-#129 := [monotonicity #106 #126]: #128
-#97 := (iff #29 #96)
-#98 := [rewrite]: #97
-#132 := [monotonicity #98 #129]: #131
-#94 := (iff #28 #93)
-#95 := [monotonicity #92]: #94
-#135 := [monotonicity #95 #132]: #134
-#89 := (iff #26 #86)
-#83 := (and #15 #6)
-#87 := (iff #83 #86)
-#88 := [rewrite]: #87
-#84 := (iff #26 #83)
-#85 := [monotonicity #82]: #84
-#90 := [trans #85 #88]: #89
-#138 := [monotonicity #90 #135]: #137
-#79 := (iff #24 #78)
-#80 := [rewrite]: #79
-#141 := [monotonicity #80 #138]: #140
-#76 := (iff #22 #75)
-#77 := [monotonicity #74]: #76
-#144 := [monotonicity #77 #141]: #143
-#147 := [monotonicity #144]: #146
-#71 := (iff #18 #70)
-#72 := [rewrite]: #71
-#150 := [monotonicity #72 #147]: #149
-#68 := (iff #16 #67)
-#69 := [rewrite]: #68
-#153 := [monotonicity #69 #150]: #152
-#156 := [monotonicity #153]: #155
-#159 := [monotonicity #156]: #158
-#162 := [monotonicity #159]: #161
-#165 := [monotonicity #162]: #164
-#346 := [trans #165 #344]: #345
-#66 := [asserted]: #50
-#347 := [mp #66 #346]: #342
-#355 := [not-or-elim #347]: #354
-#466 := [mp #355 #465]: #455
-#686 := [unit-resolution #466 #685]: #190
-#427 := (or #170 #189 #238)
-#350 := (not #174)
-#430 := (iff #350 #427)
-#382 := (or #189 #238)
-#414 := (or #170 #382)
-#428 := (iff #414 #427)
-#429 := [rewrite]: #428
-#425 := (iff #350 #414)
-#415 := (not #414)
-#420 := (not #415)
-#423 := (iff #420 #414)
-#424 := [rewrite]: #423
-#421 := (iff #350 #420)
-#418 := (iff #174 #415)
-#380 := (not #382)
-#411 := (and #380 #171)
-#416 := (iff #411 #415)
-#417 := [rewrite]: #416
-#412 := (iff #174 #411)
-#383 := (iff #10 #380)
-#384 := [rewrite]: #383
-#413 := [monotonicity #384]: #412
-#419 := [trans #413 #417]: #418
-#422 := [monotonicity #419]: #421
-#426 := [trans #422 #424]: #425
-#431 := [trans #426 #429]: #430
-#351 := [not-or-elim #347]: #350
-#432 := [mp #351 #431]: #427
-#687 := [unit-resolution #432 #686 #685]: #238
-#549 := (or #170 #218)
-#364 := (not #232)
-#558 := (iff #364 #549)
-#550 := (not #549)
-#553 := (not #550)
-#556 := (iff #553 #549)
-#557 := [rewrite]: #556
-#554 := (iff #364 #553)
-#551 := (iff #232 #550)
-#552 := [rewrite]: #551
-#555 := [monotonicity #552]: #554
-#559 := [trans #555 #557]: #558
-#365 := [not-or-elim #347]: #364
-#560 := [mp #365 #559]: #549
-#688 := [unit-resolution #560 #685]: #218
-#577 := (or #205 #217 #237)
-#366 := (not #244)
-#580 := (iff #366 #577)
-#385 := (or #205 #217)
-#564 := (or #237 #385)
-#578 := (iff #564 #577)
-#579 := [rewrite]: #578
-#575 := (iff #366 #564)
-#565 := (not #564)
-#570 := (not #565)
-#573 := (iff #570 #564)
-#574 := [rewrite]: #573
-#571 := (iff #366 #570)
-#568 := (iff #244 #565)
-#386 := (not #385)
-#561 := (and #386 #238)
-#566 := (iff #561 #565)
-#567 := [rewrite]: #566
-#562 := (iff #244 #561)
-#387 := (iff #17 #386)
-#388 := [rewrite]: #387
-#563 := [monotonicity #388]: #562
-#569 := [trans #563 #567]: #568
-#572 := [monotonicity #569]: #571
-#576 := [trans #572 #574]: #575
-#581 := [trans #576 #579]: #580
-#367 := [not-or-elim #347]: #366
-#582 := [mp #367 #581]: #577
-#689 := [unit-resolution #582 #688 #687]: #205
-#583 := (or #204 #237)
-#368 := (not #252)
-#592 := (iff #368 #583)
-#584 := (not #583)
-#587 := (not #584)
-#590 := (iff #587 #583)
-#591 := [rewrite]: #590
-#588 := (iff #368 #587)
-#585 := (iff #252 #584)
-#586 := [rewrite]: #585
-#589 := [monotonicity #586]: #588
-#593 := [trans #589 #591]: #592
-#369 := [not-or-elim #347]: #368
-#594 := [mp #369 #593]: #583
-#690 := [unit-resolution #594 #689 #687]: false
-#691 := [lemma #690]: #170
-#487 := (or #171 #177 #204)
-#356 := (not #208)
-#490 := (iff #356 #487)
-#467 := (or #171 #177)
-#474 := (or #204 #467)
-#488 := (iff #474 #487)
-#489 := [rewrite]: #488
-#485 := (iff #356 #474)
-#475 := (not #474)
-#480 := (not #475)
-#483 := (iff #480 #474)
-#484 := [rewrite]: #483
-#481 := (iff #356 #480)
-#478 := (iff #208 #475)
-#468 := (not #467)
-#471 := (and #468 #205)
-#476 := (iff #471 #475)
-#477 := [rewrite]: #476
-#472 := (iff #208 #471)
-#469 := (iff #15 #468)
-#470 := [rewrite]: #469
-#473 := [monotonicity #470]: #472
-#479 := [trans #473 #477]: #478
-#482 := [monotonicity #479]: #481
-#486 := [trans #482 #484]: #485
-#491 := [trans #486 #489]: #490
-#357 := [not-or-elim #347]: #356
-#492 := [mp #357 #491]: #487
-#694 := [unit-resolution #492 #691]: #693
-#695 := [unit-resolution #694 #692]: #177
-#493 := (or #178 #204)
-#358 := (not #211)
-#502 := (iff #358 #493)
-#494 := (not #493)
-#497 := (not #494)
-#500 := (iff #497 #493)
-#501 := [rewrite]: #500
-#498 := (iff #358 #497)
-#495 := (iff #211 #494)
-#496 := [rewrite]: #495
-#499 := [monotonicity #496]: #498
-#503 := [trans #499 #501]: #502
-#359 := [not-or-elim #347]: #358
-#504 := [mp #359 #503]: #493
-#696 := [unit-resolution #504 #695 #692]: false
-#697 := [lemma #696]: #204
-#698 := [hypothesis]: #177
-#449 := (or #178 #205 #217)
-#352 := (not #184)
-#452 := (iff #352 #449)
-#436 := (or #178 #385)
-#450 := (iff #436 #449)
-#451 := [rewrite]: #450
-#447 := (iff #352 #436)
-#437 := (not #436)
-#442 := (not #437)
-#445 := (iff #442 #436)
-#446 := [rewrite]: #445
-#443 := (iff #352 #442)
-#440 := (iff #184 #437)
-#433 := (and #386 #177)
-#438 := (iff #433 #437)
-#439 := [rewrite]: #438
-#434 := (iff #184 #433)
-#435 := [monotonicity #388]: #434
-#441 := [trans #435 #439]: #440
-#444 := [monotonicity #441]: #443
-#448 := [trans #444 #446]: #447
-#453 := [trans #448 #451]: #452
-#353 := [not-or-elim #347]: #352
-#454 := [mp #353 #453]: #449
-#699 := [unit-resolution #454 #698 #697]: #217
-#639 := (or #178 #237)
-#374 := (not #273)
-#648 := (iff #374 #639)
-#640 := (not #639)
-#643 := (not #640)
-#646 := (iff #643 #639)
-#647 := [rewrite]: #646
-#644 := (iff #374 #643)
-#641 := (iff #273 #640)
-#642 := [rewrite]: #641
-#645 := [monotonicity #642]: #644
-#649 := [trans #645 #647]: #648
-#375 := [not-or-elim #347]: #374
-#650 := [mp #375 #649]: #639
-#700 := [unit-resolution #650 #698]: #237
-#667 := (or #189 #218 #238)
-#376 := (not #279)
-#670 := (iff #376 #667)
-#654 := (or #218 #382)
-#668 := (iff #654 #667)
-#669 := [rewrite]: #668
-#665 := (iff #376 #654)
-#655 := (not #654)
-#660 := (not #655)
-#663 := (iff #660 #654)
-#664 := [rewrite]: #663
-#661 := (iff #376 #660)
-#658 := (iff #279 #655)
-#651 := (and #380 #217)
-#656 := (iff #651 #655)
-#657 := [rewrite]: #656
-#652 := (iff #279 #651)
-#653 := [monotonicity #384]: #652
-#659 := [trans #653 #657]: #658
-#662 := [monotonicity #659]: #661
-#666 := [trans #662 #664]: #665
-#671 := [trans #666 #669]: #670
-#377 := [not-or-elim #347]: #376
-#672 := [mp #377 #671]: #667
-#701 := [unit-resolution #672 #699 #700]: #189
-#673 := (or #190 #218)
-#378 := (not #287)
-#682 := (iff #378 #673)
-#674 := (not #673)
-#677 := (not #674)
-#680 := (iff #677 #673)
-#681 := [rewrite]: #680
-#678 := (iff #378 #677)
-#675 := (iff #287 #674)
-#676 := [rewrite]: #675
-#679 := [monotonicity #676]: #678
-#683 := [trans #679 #681]: #682
-#379 := [not-or-elim #347]: #378
-#684 := [mp #379 #683]: #673
-#702 := [unit-resolution #684 #701 #699]: false
-#703 := [lemma #702]: #178
-#704 := (or #177 #218)
-#543 := (or #171 #177 #218)
-#362 := (not #224)
-#546 := (iff #362 #543)
-#530 := (or #218 #467)
-#544 := (iff #530 #543)
-#545 := [rewrite]: #544
-#541 := (iff #362 #530)
-#531 := (not #530)
-#536 := (not #531)
-#539 := (iff #536 #530)
-#540 := [rewrite]: #539
-#537 := (iff #362 #536)
-#534 := (iff #224 #531)
-#527 := (and #468 #217)
-#532 := (iff #527 #531)
-#533 := [rewrite]: #532
-#528 := (iff #224 #527)
-#529 := [monotonicity #470]: #528
-#535 := [trans #529 #533]: #534
-#538 := [monotonicity #535]: #537
-#542 := [trans #538 #540]: #541
-#547 := [trans #542 #545]: #546
-#363 := [not-or-elim #347]: #362
-#548 := [mp #363 #547]: #543
-#705 := [unit-resolution #548 #691]: #704
-#706 := [unit-resolution #705 #703]: #218
-#707 := (or #177 #237)
-#633 := (or #171 #177 #237)
-#372 := (not #268)
-#636 := (iff #372 #633)
-#620 := (or #237 #467)
-#634 := (iff #620 #633)
-#635 := [rewrite]: #634
-#631 := (iff #372 #620)
-#621 := (not #620)
-#626 := (not #621)
-#629 := (iff #626 #620)
-#630 := [rewrite]: #629
-#627 := (iff #372 #626)
-#624 := (iff #268 #621)
-#617 := (and #468 #238)
-#622 := (iff #617 #621)
-#623 := [rewrite]: #622
-#618 := (iff #268 #617)
-#619 := [monotonicity #470]: #618
-#625 := [trans #619 #623]: #624
-#628 := [monotonicity #625]: #627
-#632 := [trans #628 #630]: #631
-#637 := [trans #632 #635]: #636
-#373 := [not-or-elim #347]: #372
-#638 := [mp #373 #637]: #633
-#708 := [unit-resolution #638 #691]: #707
-#709 := [unit-resolution #708 #703]: #237
-#611 := (or #190 #205 #217)
-#370 := (not #260)
-#614 := (iff #370 #611)
-#598 := (or #190 #385)
-#612 := (iff #598 #611)
-#613 := [rewrite]: #612
-#609 := (iff #370 #598)
-#599 := (not #598)
-#604 := (not #599)
-#607 := (iff #604 #598)
-#608 := [rewrite]: #607
-#605 := (iff #370 #604)
-#602 := (iff #260 #599)
-#595 := (and #386 #189)
-#600 := (iff #595 #599)
-#601 := [rewrite]: #600
-#596 := (iff #260 #595)
-#597 := [monotonicity #388]: #596
-#603 := [trans #597 #601]: #602
-#606 := [monotonicity #603]: #605
-#610 := [trans #606 #608]: #609
-#615 := [trans #610 #613]: #614
-#371 := [not-or-elim #347]: #370
-#616 := [mp #371 #615]: #611
-#710 := [unit-resolution #616 #706 #697]: #190
-#405 := (or #189 #205 #217 #238)
-#348 := (not #115)
-#408 := (iff #348 #405)
-#392 := (or #382 #385)
-#406 := (iff #392 #405)
-#407 := [rewrite]: #406
-#403 := (iff #348 #392)
-#393 := (not #392)
-#398 := (not #393)
-#401 := (iff #398 #392)
-#402 := [rewrite]: #401
-#399 := (iff #348 #398)
-#396 := (iff #115 #393)
-#389 := (and #380 #386)
-#394 := (iff #389 #393)
-#395 := [rewrite]: #394
-#390 := (iff #115 #389)
-#391 := [monotonicity #384 #388]: #390
-#397 := [trans #391 #395]: #396
-#400 := [monotonicity #397]: #399
-#404 := [trans #400 #402]: #403
-#409 := [trans #404 #407]: #408
-#349 := [not-or-elim #347]: #348
-#410 := [mp #349 #409]: #405
-[unit-resolution #410 #710 #709 #697 #706]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_linarith_16 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,26 +0,0 @@
-(benchmark Isabelle
-:extrafuns (
- (uf_3 Int)
- (uf_2 Int)
- (uf_1 Int)
- (uf_4 Int)
- (uf_5 Int)
- (uf_6 Int)
- (uf_7 Int)
- (uf_8 Int)
- (uf_9 Int)
- (uf_10 Int)
- (uf_11 Int)
- )
-:assumption (= uf_1 (- (ite (< uf_2 0) (~ uf_2) uf_2) uf_3))
-:assumption (= uf_4 (- (ite (< uf_1 0) (~ uf_1) uf_1) uf_2))
-:assumption (= uf_5 (- (ite (< uf_4 0) (~ uf_4) uf_4) uf_1))
-:assumption (= uf_6 (- (ite (< uf_5 0) (~ uf_5) uf_5) uf_4))
-:assumption (= uf_7 (- (ite (< uf_6 0) (~ uf_6) uf_6) uf_5))
-:assumption (= uf_8 (- (ite (< uf_7 0) (~ uf_7) uf_7) uf_6))
-:assumption (= uf_9 (- (ite (< uf_8 0) (~ uf_8) uf_8) uf_7))
-:assumption (= uf_10 (- (ite (< uf_9 0) (~ uf_9) uf_9) uf_8))
-:assumption (= uf_11 (- (ite (< uf_10 0) (~ uf_10) uf_10) uf_9))
-:assumption (not (and (= uf_3 uf_10) (= uf_2 uf_11)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_linarith_16.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,2291 +0,0 @@
-#2 := false
-#6 := 0::int
-decl z3name!0 :: int
-#647 := z3name!0
-#81 := -1::int
-#656 := (* -1::int z3name!0)
-decl uf_2 :: int
-#5 := uf_2
-#882 := (+ uf_2 #656)
-#883 := (<= #882 0::int)
-#885 := (not #883)
-#881 := (>= #882 0::int)
-#884 := (not #881)
-#886 := (or #884 #885)
-decl uf_11 :: int
-#55 := uf_11
-#513 := (* -1::int uf_11)
-#514 := (+ uf_2 #513)
-#515 := (<= #514 0::int)
-decl z3name!5 :: int
-#777 := z3name!5
-decl uf_7 :: int
-#31 := uf_7
-#1083 := (+ uf_7 z3name!5)
-#1084 := (<= #1083 0::int)
-#335 := (>= uf_7 0::int)
-#1085 := (>= #1083 0::int)
-#1087 := (not #1085)
-#1086 := (not #1084)
-#1088 := (or #1086 #1087)
-#2302 := [hypothesis]: #1086
-#1289 := (or #1088 #1084)
-#1290 := [def-axiom]: #1289
-#2303 := [unit-resolution #1290 #2302]: #1088
-#1089 := (not #1088)
-#1092 := (or #335 #1089)
-#1099 := (not #1092)
-#786 := (* -1::int z3name!5)
-#1072 := (+ uf_7 #786)
-#1073 := (<= #1072 0::int)
-#1075 := (not #1073)
-#1071 := (>= #1072 0::int)
-#1074 := (not #1071)
-#1076 := (or #1074 #1075)
-#1077 := (not #1076)
-#336 := (not #335)
-#1080 := (or #336 #1077)
-#1098 := (not #1080)
-#1100 := (or #1098 #1099)
-#1101 := (not #1100)
-#318 := (* -1::int uf_7)
-#780 := (= z3name!5 #318)
-#781 := (or #335 #780)
-#778 := (= z3name!5 uf_7)
-#779 := (or #336 #778)
-#782 := (and #779 #781)
-#1104 := (iff #782 #1101)
-#1095 := (and #1080 #1092)
-#1102 := (iff #1095 #1101)
-#1103 := [rewrite]: #1102
-#1096 := (iff #782 #1095)
-#1093 := (iff #781 #1092)
-#1090 := (iff #780 #1089)
-#1091 := [rewrite]: #1090
-#1094 := [monotonicity #1091]: #1093
-#1081 := (iff #779 #1080)
-#1078 := (iff #778 #1077)
-#1079 := [rewrite]: #1078
-#1082 := [monotonicity #1079]: #1081
-#1097 := [monotonicity #1082 #1094]: #1096
-#1105 := [trans #1097 #1103]: #1104
-#783 := [intro-def]: #782
-#1106 := [mp #783 #1105]: #1101
-#1108 := [not-or-elim #1106]: #1092
-#2304 := [unit-resolution #1108 #2303]: #335
-decl uf_4 :: int
-#13 := uf_4
-#194 := (>= uf_4 0::int)
-decl uf_10 :: int
-#49 := uf_10
-#459 := (* -1::int uf_10)
-decl uf_3 :: int
-#10 := uf_3
-#508 := (+ uf_3 #459)
-#509 := (>= #508 0::int)
-decl z3name!1 :: int
-#673 := z3name!1
-#682 := (* -1::int z3name!1)
-decl uf_1 :: int
-#4 := uf_1
-#920 := (+ uf_1 #682)
-#921 := (<= #920 0::int)
-#931 := (+ uf_1 z3name!1)
-#933 := (>= #931 0::int)
-#935 := (not #933)
-#932 := (<= #931 0::int)
-#934 := (not #932)
-#936 := (or #934 #935)
-#937 := (not #936)
-#147 := (>= uf_1 0::int)
-#148 := (not #147)
-#923 := (not #921)
-#919 := (>= #920 0::int)
-#922 := (not #919)
-#924 := (or #922 #923)
-#2022 := [hypothesis]: #923
-#1237 := (or #924 #921)
-#1238 := [def-axiom]: #1237
-#2023 := [unit-resolution #1238 #2022]: #924
-#925 := (not #924)
-#928 := (or #148 #925)
-#940 := (or #147 #937)
-#947 := (not #940)
-#946 := (not #928)
-#948 := (or #946 #947)
-#949 := (not #948)
-#130 := (* -1::int uf_1)
-#676 := (= z3name!1 #130)
-#677 := (or #147 #676)
-#674 := (= z3name!1 uf_1)
-#675 := (or #148 #674)
-#678 := (and #675 #677)
-#952 := (iff #678 #949)
-#943 := (and #928 #940)
-#950 := (iff #943 #949)
-#951 := [rewrite]: #950
-#944 := (iff #678 #943)
-#941 := (iff #677 #940)
-#938 := (iff #676 #937)
-#939 := [rewrite]: #938
-#942 := [monotonicity #939]: #941
-#929 := (iff #675 #928)
-#926 := (iff #674 #925)
-#927 := [rewrite]: #926
-#930 := [monotonicity #927]: #929
-#945 := [monotonicity #930 #942]: #944
-#953 := [trans #945 #951]: #952
-#679 := [intro-def]: #678
-#954 := [mp #679 #953]: #949
-#955 := [not-or-elim #954]: #928
-#2024 := [unit-resolution #955 #2023]: #148
-#956 := [not-or-elim #954]: #940
-#2025 := [unit-resolution #956 #2024]: #937
-#2026 := (or #921 #919)
-#2027 := [th-lemma]: #2026
-#2028 := [unit-resolution #2027 #2022]: #919
-#2029 := (or #922 #147 #935)
-#2030 := [th-lemma]: #2029
-#2031 := [unit-resolution #2030 #2024 #2028]: #935
-#1243 := (or #936 #933)
-#1244 := [def-axiom]: #1243
-#2032 := [unit-resolution #1244 #2031 #2025]: false
-#2033 := [lemma #2032]: #921
-decl z3name!7 :: int
-#829 := z3name!7
-decl uf_9 :: int
-#43 := uf_9
-#1159 := (+ uf_9 z3name!7)
-#1160 := (<= #1159 0::int)
-#838 := (* -1::int z3name!7)
-#1148 := (+ uf_9 #838)
-#1147 := (>= #1148 0::int)
-decl z3name!4 :: int
-#751 := z3name!4
-#760 := (* -1::int z3name!4)
-decl uf_6 :: int
-#25 := uf_6
-#1034 := (+ uf_6 #760)
-#1033 := (>= #1034 0::int)
-#1035 := (<= #1034 0::int)
-#1037 := (not #1035)
-#1036 := (not #1033)
-#1038 := (or #1036 #1037)
-#1039 := (not #1038)
-#288 := (>= uf_6 0::int)
-#893 := (+ uf_2 z3name!0)
-#895 := (>= #893 0::int)
-#897 := (not #895)
-#894 := (<= #893 0::int)
-#896 := (not #894)
-#898 := (or #896 #897)
-#899 := (not #898)
-#100 := (>= uf_2 0::int)
-#101 := (not #100)
-#1736 := [hypothesis]: #885
-#1225 := (or #886 #883)
-#1226 := [def-axiom]: #1225
-#1737 := [unit-resolution #1226 #1736]: #886
-#887 := (not #886)
-#890 := (or #101 #887)
-#902 := (or #100 #899)
-#909 := (not #902)
-#908 := (not #890)
-#910 := (or #908 #909)
-#911 := (not #910)
-#82 := (* -1::int uf_2)
-#650 := (= z3name!0 #82)
-#651 := (or #100 #650)
-#648 := (= z3name!0 uf_2)
-#649 := (or #101 #648)
-#652 := (and #649 #651)
-#914 := (iff #652 #911)
-#905 := (and #890 #902)
-#912 := (iff #905 #911)
-#913 := [rewrite]: #912
-#906 := (iff #652 #905)
-#903 := (iff #651 #902)
-#900 := (iff #650 #899)
-#901 := [rewrite]: #900
-#904 := [monotonicity #901]: #903
-#891 := (iff #649 #890)
-#888 := (iff #648 #887)
-#889 := [rewrite]: #888
-#892 := [monotonicity #889]: #891
-#907 := [monotonicity #892 #904]: #906
-#915 := [trans #907 #913]: #914
-#653 := [intro-def]: #652
-#916 := [mp #653 #915]: #911
-#917 := [not-or-elim #916]: #890
-#1738 := [unit-resolution #917 #1737]: #101
-#918 := [not-or-elim #916]: #902
-#1739 := [unit-resolution #918 #1738]: #899
-#1231 := (or #898 #895)
-#1232 := [def-axiom]: #1231
-#1740 := [unit-resolution #1232 #1739]: #895
-#1741 := [th-lemma #1736 #1738 #1740]: false
-#1742 := [lemma #1741]: #883
-#1149 := (<= #1148 0::int)
-#1151 := (not #1149)
-#1150 := (not #1147)
-#1152 := (or #1150 #1151)
-#1153 := (not #1152)
-#429 := (>= uf_9 0::int)
-decl z3name!6 :: int
-#803 := z3name!6
-#812 := (* -1::int z3name!6)
-decl uf_8 :: int
-#37 := uf_8
-#1110 := (+ uf_8 #812)
-#1111 := (<= #1110 0::int)
-#1113 := (not #1111)
-#1109 := (>= #1110 0::int)
-#1112 := (not #1109)
-#1114 := (or #1112 #1113)
-#1865 := [hypothesis]: #1113
-#1297 := (or #1114 #1111)
-#1298 := [def-axiom]: #1297
-#1866 := [unit-resolution #1298 #1865]: #1114
-#382 := (>= uf_8 0::int)
-#1685 := (or #1111 #1109)
-#1686 := [th-lemma]: #1685
-#1867 := [unit-resolution #1686 #1865]: #1109
-#1734 := (or #382 #1112)
-#1121 := (+ uf_8 z3name!6)
-#1123 := (>= #1121 0::int)
-#1125 := (not #1123)
-#1122 := (<= #1121 0::int)
-#1124 := (not #1122)
-#1126 := (or #1124 #1125)
-#1127 := (not #1126)
-#383 := (not #382)
-#1428 := [hypothesis]: #383
-#1130 := (or #382 #1127)
-#1137 := (not #1130)
-#1115 := (not #1114)
-#1118 := (or #383 #1115)
-#1136 := (not #1118)
-#1138 := (or #1136 #1137)
-#1139 := (not #1138)
-#365 := (* -1::int uf_8)
-#806 := (= z3name!6 #365)
-#807 := (or #382 #806)
-#804 := (= z3name!6 uf_8)
-#805 := (or #383 #804)
-#808 := (and #805 #807)
-#1142 := (iff #808 #1139)
-#1133 := (and #1118 #1130)
-#1140 := (iff #1133 #1139)
-#1141 := [rewrite]: #1140
-#1134 := (iff #808 #1133)
-#1131 := (iff #807 #1130)
-#1128 := (iff #806 #1127)
-#1129 := [rewrite]: #1128
-#1132 := [monotonicity #1129]: #1131
-#1119 := (iff #805 #1118)
-#1116 := (iff #804 #1115)
-#1117 := [rewrite]: #1116
-#1120 := [monotonicity #1117]: #1119
-#1135 := [monotonicity #1120 #1132]: #1134
-#1143 := [trans #1135 #1141]: #1142
-#809 := [intro-def]: #808
-#1144 := [mp #809 #1143]: #1139
-#1146 := [not-or-elim #1144]: #1130
-#1729 := [unit-resolution #1146 #1428]: #1127
-#1637 := [hypothesis]: #1109
-#1730 := (or #1112 #1125 #382)
-#1731 := [th-lemma]: #1730
-#1732 := [unit-resolution #1731 #1428 #1637]: #1125
-#1303 := (or #1126 #1123)
-#1304 := [def-axiom]: #1303
-#1733 := [unit-resolution #1304 #1732 #1729]: false
-#1735 := [lemma #1733]: #1734
-#1868 := [unit-resolution #1735 #1867]: #382
-#1145 := [not-or-elim #1144]: #1118
-#1869 := [unit-resolution #1145 #1868 #1866]: false
-#1870 := [lemma #1869]: #1111
-#289 := (not #288)
-#1405 := [hypothesis]: #289
-#1688 := (or #288 #429 #1113)
-#815 := (+ uf_9 #812)
-#818 := (+ uf_7 #815)
-#825 := (>= #818 0::int)
-#389 := (ite #382 uf_8 #365)
-#400 := (* -1::int #389)
-#401 := (+ uf_9 #400)
-#402 := (+ uf_7 #401)
-#599 := (>= #402 0::int)
-#826 := (= #599 #825)
-#819 := (~ #402 #818)
-#816 := (~ #401 #815)
-#813 := (~ #400 #812)
-#810 := (~ #389 z3name!6)
-#811 := [apply-def #809]: #810
-#814 := [monotonicity #811]: #813
-#817 := [monotonicity #814]: #816
-#820 := [monotonicity #817]: #819
-#827 := [monotonicity #820]: #826
-#601 := (not #599)
-#598 := (<= #402 0::int)
-#600 := (not #598)
-#602 := (or #600 #601)
-#603 := (not #602)
-#403 := (= #402 0::int)
-#604 := (iff #403 #603)
-#605 := [rewrite]: #604
-#45 := (- uf_8)
-#44 := (< uf_8 0::int)
-#46 := (ite #44 #45 uf_8)
-#47 := (- #46 uf_7)
-#48 := (= uf_9 #47)
-#408 := (iff #48 #403)
-#368 := (ite #44 #365 uf_8)
-#374 := (+ #318 #368)
-#379 := (= uf_9 #374)
-#406 := (iff #379 #403)
-#394 := (+ #318 #389)
-#397 := (= uf_9 #394)
-#404 := (iff #397 #403)
-#405 := [rewrite]: #404
-#398 := (iff #379 #397)
-#395 := (= #374 #394)
-#392 := (= #368 #389)
-#386 := (ite #383 #365 uf_8)
-#390 := (= #386 #389)
-#391 := [rewrite]: #390
-#387 := (= #368 #386)
-#384 := (iff #44 #383)
-#385 := [rewrite]: #384
-#388 := [monotonicity #385]: #387
-#393 := [trans #388 #391]: #392
-#396 := [monotonicity #393]: #395
-#399 := [monotonicity #396]: #398
-#407 := [trans #399 #405]: #406
-#380 := (iff #48 #379)
-#377 := (= #47 #374)
-#371 := (- #368 uf_7)
-#375 := (= #371 #374)
-#376 := [rewrite]: #375
-#372 := (= #47 #371)
-#369 := (= #46 #368)
-#366 := (= #45 #365)
-#367 := [rewrite]: #366
-#370 := [monotonicity #367]: #369
-#373 := [monotonicity #370]: #372
-#378 := [trans #373 #376]: #377
-#381 := [monotonicity #378]: #380
-#409 := [trans #381 #407]: #408
-#364 := [asserted]: #48
-#410 := [mp #364 #409]: #403
-#606 := [mp #410 #605]: #603
-#608 := [not-or-elim #606]: #599
-#828 := [mp~ #608 #827]: #825
-#1441 := [hypothesis]: #1075
-#1285 := (or #1076 #1073)
-#1286 := [def-axiom]: #1285
-#1442 := [unit-resolution #1286 #1441]: #1076
-#1107 := [not-or-elim #1106]: #1080
-#1443 := [unit-resolution #1107 #1442]: #336
-#1444 := [unit-resolution #1108 #1443]: #1089
-#1291 := (or #1088 #1085)
-#1292 := [def-axiom]: #1291
-#1445 := [unit-resolution #1292 #1444]: #1085
-#1446 := [th-lemma #1441 #1445 #1443]: false
-#1447 := [lemma #1446]: #1073
-#789 := (+ uf_8 #786)
-#792 := (+ uf_6 #789)
-#799 := (>= #792 0::int)
-#342 := (ite #335 uf_7 #318)
-#353 := (* -1::int #342)
-#354 := (+ uf_8 #353)
-#355 := (+ uf_6 #354)
-#588 := (>= #355 0::int)
-#800 := (= #588 #799)
-#793 := (~ #355 #792)
-#790 := (~ #354 #789)
-#787 := (~ #353 #786)
-#784 := (~ #342 z3name!5)
-#785 := [apply-def #783]: #784
-#788 := [monotonicity #785]: #787
-#791 := [monotonicity #788]: #790
-#794 := [monotonicity #791]: #793
-#801 := [monotonicity #794]: #800
-#590 := (not #588)
-#587 := (<= #355 0::int)
-#589 := (not #587)
-#591 := (or #589 #590)
-#592 := (not #591)
-#356 := (= #355 0::int)
-#593 := (iff #356 #592)
-#594 := [rewrite]: #593
-#39 := (- uf_7)
-#38 := (< uf_7 0::int)
-#40 := (ite #38 #39 uf_7)
-#41 := (- #40 uf_6)
-#42 := (= uf_8 #41)
-#361 := (iff #42 #356)
-#321 := (ite #38 #318 uf_7)
-#271 := (* -1::int uf_6)
-#327 := (+ #271 #321)
-#332 := (= uf_8 #327)
-#359 := (iff #332 #356)
-#347 := (+ #271 #342)
-#350 := (= uf_8 #347)
-#357 := (iff #350 #356)
-#358 := [rewrite]: #357
-#351 := (iff #332 #350)
-#348 := (= #327 #347)
-#345 := (= #321 #342)
-#339 := (ite #336 #318 uf_7)
-#343 := (= #339 #342)
-#344 := [rewrite]: #343
-#340 := (= #321 #339)
-#337 := (iff #38 #336)
-#338 := [rewrite]: #337
-#341 := [monotonicity #338]: #340
-#346 := [trans #341 #344]: #345
-#349 := [monotonicity #346]: #348
-#352 := [monotonicity #349]: #351
-#360 := [trans #352 #358]: #359
-#333 := (iff #42 #332)
-#330 := (= #41 #327)
-#324 := (- #321 uf_6)
-#328 := (= #324 #327)
-#329 := [rewrite]: #328
-#325 := (= #41 #324)
-#322 := (= #40 #321)
-#319 := (= #39 #318)
-#320 := [rewrite]: #319
-#323 := [monotonicity #320]: #322
-#326 := [monotonicity #323]: #325
-#331 := [trans #326 #329]: #330
-#334 := [monotonicity #331]: #333
-#362 := [trans #334 #360]: #361
-#317 := [asserted]: #42
-#363 := [mp #317 #362]: #356
-#595 := [mp #363 #594]: #592
-#597 := [not-or-elim #595]: #588
-#802 := [mp~ #597 #801]: #799
-#1343 := (not #825)
-#1350 := (not #799)
-#1351 := (or #288 #1075 #1350 #429 #1113 #1343)
-#1352 := [th-lemma]: #1351
-#1689 := [unit-resolution #1352 #802 #1447 #828]: #1688
-#2046 := [unit-resolution #1689 #1405 #1870]: #429
-#430 := (not #429)
-#1156 := (or #430 #1153)
-#1161 := (>= #1159 0::int)
-#1163 := (not #1161)
-#1162 := (not #1160)
-#1164 := (or #1162 #1163)
-#1165 := (not #1164)
-#1168 := (or #429 #1165)
-#1175 := (not #1168)
-#1174 := (not #1156)
-#1176 := (or #1174 #1175)
-#1177 := (not #1176)
-#412 := (* -1::int uf_9)
-#832 := (= z3name!7 #412)
-#833 := (or #429 #832)
-#830 := (= z3name!7 uf_9)
-#831 := (or #430 #830)
-#834 := (and #831 #833)
-#1180 := (iff #834 #1177)
-#1171 := (and #1156 #1168)
-#1178 := (iff #1171 #1177)
-#1179 := [rewrite]: #1178
-#1172 := (iff #834 #1171)
-#1169 := (iff #833 #1168)
-#1166 := (iff #832 #1165)
-#1167 := [rewrite]: #1166
-#1170 := [monotonicity #1167]: #1169
-#1157 := (iff #831 #1156)
-#1154 := (iff #830 #1153)
-#1155 := [rewrite]: #1154
-#1158 := [monotonicity #1155]: #1157
-#1173 := [monotonicity #1158 #1170]: #1172
-#1181 := [trans #1173 #1179]: #1180
-#835 := [intro-def]: #834
-#1182 := [mp #835 #1181]: #1177
-#1183 := [not-or-elim #1182]: #1156
-#2047 := [unit-resolution #1183 #2046]: #1153
-#1307 := (or #1152 #1147)
-#1308 := [def-axiom]: #1307
-#2112 := [unit-resolution #1308 #2047]: #1147
-#2009 := (or #288 #382)
-#1998 := (or #1036 #288)
-#1045 := (+ uf_6 z3name!4)
-#1047 := (>= #1045 0::int)
-#1049 := (not #1047)
-#1046 := (<= #1045 0::int)
-#1048 := (not #1046)
-#1050 := (or #1048 #1049)
-#1460 := [hypothesis]: #1049
-#1279 := (or #1050 #1047)
-#1280 := [def-axiom]: #1279
-#1461 := [unit-resolution #1280 #1460]: #1050
-#1464 := (or #1047 #289)
-#1051 := (not #1050)
-#1448 := [hypothesis]: #1037
-#1273 := (or #1038 #1035)
-#1274 := [def-axiom]: #1273
-#1449 := [unit-resolution #1274 #1448]: #1038
-#1042 := (or #289 #1039)
-#1054 := (or #288 #1051)
-#1061 := (not #1054)
-#1060 := (not #1042)
-#1062 := (or #1060 #1061)
-#1063 := (not #1062)
-#754 := (= z3name!4 #271)
-#755 := (or #288 #754)
-#752 := (= z3name!4 uf_6)
-#753 := (or #289 #752)
-#756 := (and #753 #755)
-#1066 := (iff #756 #1063)
-#1057 := (and #1042 #1054)
-#1064 := (iff #1057 #1063)
-#1065 := [rewrite]: #1064
-#1058 := (iff #756 #1057)
-#1055 := (iff #755 #1054)
-#1052 := (iff #754 #1051)
-#1053 := [rewrite]: #1052
-#1056 := [monotonicity #1053]: #1055
-#1043 := (iff #753 #1042)
-#1040 := (iff #752 #1039)
-#1041 := [rewrite]: #1040
-#1044 := [monotonicity #1041]: #1043
-#1059 := [monotonicity #1044 #1056]: #1058
-#1067 := [trans #1059 #1065]: #1066
-#757 := [intro-def]: #756
-#1068 := [mp #757 #1067]: #1063
-#1069 := [not-or-elim #1068]: #1042
-#1450 := [unit-resolution #1069 #1449]: #289
-#1070 := [not-or-elim #1068]: #1054
-#1451 := [unit-resolution #1070 #1450]: #1051
-#1452 := (or #1035 #1033)
-#1453 := [th-lemma]: #1452
-#1454 := [unit-resolution #1453 #1448]: #1033
-#1455 := (or #1036 #288 #1049)
-#1456 := [th-lemma]: #1455
-#1457 := [unit-resolution #1456 #1450 #1454]: #1049
-#1458 := [unit-resolution #1280 #1457 #1451]: false
-#1459 := [lemma #1458]: #1035
-#1462 := (or #1047 #1037 #289)
-#1463 := [th-lemma]: #1462
-#1465 := [unit-resolution #1463 #1459]: #1464
-#1466 := [unit-resolution #1465 #1460]: #289
-#1467 := [unit-resolution #1070 #1466 #1461]: false
-#1468 := [lemma #1467]: #1047
-#1999 := [unit-resolution #1456 #1468]: #1998
-#2000 := [unit-resolution #1999 #1405]: #1036
-#1407 := [unit-resolution #1070 #1405]: #1051
-#1277 := (or #1050 #1046)
-#1278 := [def-axiom]: #1277
-#1497 := [unit-resolution #1278 #1407]: #1046
-#2001 := (or #336 #1048 #1033 #382 #1350 #1075)
-#2002 := [th-lemma]: #2001
-#2003 := [unit-resolution #2002 #1497 #2000 #1447 #802 #1428]: #336
-#2004 := (or #1087 #1075 #1048 #1033 #382 #1350)
-#2005 := [th-lemma]: #2004
-#2006 := [unit-resolution #2005 #1497 #1447 #2000 #802 #1428]: #1087
-#2007 := [unit-resolution #1292 #2006]: #1088
-#2008 := [unit-resolution #1108 #2007 #2003]: false
-#2010 := [lemma #2008]: #2009
-#2113 := [unit-resolution #2010 #1405]: #382
-#2114 := [unit-resolution #1145 #2113]: #1115
-#1295 := (or #1114 #1109)
-#1296 := [def-axiom]: #1295
-#2115 := [unit-resolution #1296 #2114]: #1109
-decl z3name!2 :: int
-#699 := z3name!2
-#708 := (* -1::int z3name!2)
-#958 := (+ uf_4 #708)
-#957 := (>= #958 0::int)
-#959 := (<= #958 0::int)
-#961 := (not #959)
-#960 := (not #957)
-#962 := (or #960 #961)
-#963 := (not #962)
-decl uf_5 :: int
-#19 := uf_5
-#241 := (>= uf_5 0::int)
-#242 := (not #241)
-#1406 := [hypothesis]: #242
-#1579 := (or #1048 #241)
-#516 := (>= #514 0::int)
-#476 := (>= uf_10 0::int)
-#477 := (not #476)
-#1484 := (or #382 #241)
-#1430 := (or #382 #241 #1075 #1037)
-#1421 := [hypothesis]: #1035
-#1427 := [hypothesis]: #1073
-#763 := (+ uf_7 #760)
-#766 := (+ uf_5 #763)
-#773 := (>= #766 0::int)
-#295 := (ite #288 uf_6 #271)
-#306 := (* -1::int #295)
-#307 := (+ uf_7 #306)
-#308 := (+ uf_5 #307)
-#577 := (>= #308 0::int)
-#774 := (= #577 #773)
-#767 := (~ #308 #766)
-#764 := (~ #307 #763)
-#761 := (~ #306 #760)
-#758 := (~ #295 z3name!4)
-#759 := [apply-def #757]: #758
-#762 := [monotonicity #759]: #761
-#765 := [monotonicity #762]: #764
-#768 := [monotonicity #765]: #767
-#775 := [monotonicity #768]: #774
-#579 := (not #577)
-#576 := (<= #308 0::int)
-#578 := (not #576)
-#580 := (or #578 #579)
-#581 := (not #580)
-#309 := (= #308 0::int)
-#582 := (iff #309 #581)
-#583 := [rewrite]: #582
-#33 := (- uf_6)
-#32 := (< uf_6 0::int)
-#34 := (ite #32 #33 uf_6)
-#35 := (- #34 uf_5)
-#36 := (= uf_7 #35)
-#314 := (iff #36 #309)
-#274 := (ite #32 #271 uf_6)
-#224 := (* -1::int uf_5)
-#280 := (+ #224 #274)
-#285 := (= uf_7 #280)
-#312 := (iff #285 #309)
-#300 := (+ #224 #295)
-#303 := (= uf_7 #300)
-#310 := (iff #303 #309)
-#311 := [rewrite]: #310
-#304 := (iff #285 #303)
-#301 := (= #280 #300)
-#298 := (= #274 #295)
-#292 := (ite #289 #271 uf_6)
-#296 := (= #292 #295)
-#297 := [rewrite]: #296
-#293 := (= #274 #292)
-#290 := (iff #32 #289)
-#291 := [rewrite]: #290
-#294 := [monotonicity #291]: #293
-#299 := [trans #294 #297]: #298
-#302 := [monotonicity #299]: #301
-#305 := [monotonicity #302]: #304
-#313 := [trans #305 #311]: #312
-#286 := (iff #36 #285)
-#283 := (= #35 #280)
-#277 := (- #274 uf_5)
-#281 := (= #277 #280)
-#282 := [rewrite]: #281
-#278 := (= #35 #277)
-#275 := (= #34 #274)
-#272 := (= #33 #271)
-#273 := [rewrite]: #272
-#276 := [monotonicity #273]: #275
-#279 := [monotonicity #276]: #278
-#284 := [trans #279 #282]: #283
-#287 := [monotonicity #284]: #286
-#315 := [trans #287 #313]: #314
-#270 := [asserted]: #36
-#316 := [mp #270 #315]: #309
-#584 := [mp #316 #583]: #581
-#586 := [not-or-elim #584]: #577
-#776 := [mp~ #586 #775]: #773
-#1429 := [th-lemma #776 #1406 #1428 #1427 #802 #1421]: false
-#1431 := [lemma #1429]: #1430
-#1485 := [unit-resolution #1431 #1447 #1459]: #1484
-#1486 := [unit-resolution #1485 #1406]: #382
-#1487 := [unit-resolution #1145 #1486]: #1115
-#1496 := [unit-resolution #1298 #1487]: #1111
-#1545 := [hypothesis]: #1046
-#1548 := (or #1048 #1113 #429)
-#1546 := (or #1048 #1113 #429 #1343 #1075 #1350 #1037)
-#1547 := [th-lemma]: #1546
-#1549 := [unit-resolution #1547 #1447 #802 #1459 #828]: #1548
-#1550 := [unit-resolution #1549 #1545 #1496]: #429
-#1551 := [unit-resolution #1183 #1550]: #1153
-#1552 := [unit-resolution #1308 #1551]: #1147
-#1543 := (or #477 #241 #1150)
-#1488 := [unit-resolution #1296 #1487]: #1109
-#821 := (<= #818 0::int)
-#822 := (= #598 #821)
-#823 := [monotonicity #820]: #822
-#607 := [not-or-elim #606]: #598
-#824 := [mp~ #607 #823]: #821
-#841 := (+ uf_10 #838)
-#844 := (+ uf_8 #841)
-#847 := (<= #844 0::int)
-#436 := (ite #429 uf_9 #412)
-#447 := (* -1::int #436)
-#448 := (+ uf_10 #447)
-#449 := (+ uf_8 #448)
-#609 := (<= #449 0::int)
-#848 := (= #609 #847)
-#845 := (~ #449 #844)
-#842 := (~ #448 #841)
-#839 := (~ #447 #838)
-#836 := (~ #436 z3name!7)
-#837 := [apply-def #835]: #836
-#840 := [monotonicity #837]: #839
-#843 := [monotonicity #840]: #842
-#846 := [monotonicity #843]: #845
-#849 := [monotonicity #846]: #848
-#610 := (>= #449 0::int)
-#612 := (not #610)
-#611 := (not #609)
-#613 := (or #611 #612)
-#614 := (not #613)
-#450 := (= #449 0::int)
-#615 := (iff #450 #614)
-#616 := [rewrite]: #615
-#51 := (- uf_9)
-#50 := (< uf_9 0::int)
-#52 := (ite #50 #51 uf_9)
-#53 := (- #52 uf_8)
-#54 := (= uf_10 #53)
-#455 := (iff #54 #450)
-#415 := (ite #50 #412 uf_9)
-#421 := (+ #365 #415)
-#426 := (= uf_10 #421)
-#453 := (iff #426 #450)
-#441 := (+ #365 #436)
-#444 := (= uf_10 #441)
-#451 := (iff #444 #450)
-#452 := [rewrite]: #451
-#445 := (iff #426 #444)
-#442 := (= #421 #441)
-#439 := (= #415 #436)
-#433 := (ite #430 #412 uf_9)
-#437 := (= #433 #436)
-#438 := [rewrite]: #437
-#434 := (= #415 #433)
-#431 := (iff #50 #430)
-#432 := [rewrite]: #431
-#435 := [monotonicity #432]: #434
-#440 := [trans #435 #438]: #439
-#443 := [monotonicity #440]: #442
-#446 := [monotonicity #443]: #445
-#454 := [trans #446 #452]: #453
-#427 := (iff #54 #426)
-#424 := (= #53 #421)
-#418 := (- #415 uf_8)
-#422 := (= #418 #421)
-#423 := [rewrite]: #422
-#419 := (= #53 #418)
-#416 := (= #52 #415)
-#413 := (= #51 #412)
-#414 := [rewrite]: #413
-#417 := [monotonicity #414]: #416
-#420 := [monotonicity #417]: #419
-#425 := [trans #420 #423]: #424
-#428 := [monotonicity #425]: #427
-#456 := [trans #428 #454]: #455
-#411 := [asserted]: #54
-#457 := [mp #411 #456]: #450
-#617 := [mp #457 #616]: #614
-#618 := [not-or-elim #617]: #609
-#850 := [mp~ #618 #849]: #847
-#1540 := [hypothesis]: #1147
-#1541 := [hypothesis]: #476
-#1542 := [th-lemma #1468 #1406 #1541 #1540 #850 #824 #1488 #776 #1459]: false
-#1544 := [lemma #1542]: #1543
-#1553 := [unit-resolution #1544 #1552 #1406]: #477
-#851 := (>= #844 0::int)
-#852 := (= #610 #851)
-#853 := [monotonicity #846]: #852
-#619 := [not-or-elim #617]: #610
-#854 := [mp~ #619 #853]: #851
-#1309 := (or #1152 #1149)
-#1310 := [def-axiom]: #1309
-#1554 := [unit-resolution #1310 #1551]: #1149
-#769 := (<= #766 0::int)
-#770 := (= #576 #769)
-#771 := [monotonicity #768]: #770
-#585 := [not-or-elim #584]: #576
-#772 := [mp~ #585 #771]: #769
-decl z3name!3 :: int
-#725 := z3name!3
-#1007 := (+ uf_5 z3name!3)
-#1009 := (>= #1007 0::int)
-#1011 := (not #1009)
-#1398 := [hypothesis]: #1011
-#734 := (* -1::int z3name!3)
-#996 := (+ uf_5 #734)
-#997 := (<= #996 0::int)
-#999 := (not #997)
-#995 := (>= #996 0::int)
-#998 := (not #995)
-#1000 := (or #998 #999)
-#1001 := (not #1000)
-#1008 := (<= #1007 0::int)
-#1010 := (not #1008)
-#1012 := (or #1010 #1011)
-#1267 := (or #1012 #1009)
-#1268 := [def-axiom]: #1267
-#1399 := [unit-resolution #1268 #1398]: #1012
-#1013 := (not #1012)
-#1016 := (or #241 #1013)
-#1023 := (not #1016)
-#1004 := (or #242 #1001)
-#1022 := (not #1004)
-#1024 := (or #1022 #1023)
-#1025 := (not #1024)
-#728 := (= z3name!3 #224)
-#729 := (or #241 #728)
-#726 := (= z3name!3 uf_5)
-#727 := (or #242 #726)
-#730 := (and #727 #729)
-#1028 := (iff #730 #1025)
-#1019 := (and #1004 #1016)
-#1026 := (iff #1019 #1025)
-#1027 := [rewrite]: #1026
-#1020 := (iff #730 #1019)
-#1017 := (iff #729 #1016)
-#1014 := (iff #728 #1013)
-#1015 := [rewrite]: #1014
-#1018 := [monotonicity #1015]: #1017
-#1005 := (iff #727 #1004)
-#1002 := (iff #726 #1001)
-#1003 := [rewrite]: #1002
-#1006 := [monotonicity #1003]: #1005
-#1021 := [monotonicity #1006 #1018]: #1020
-#1029 := [trans #1021 #1027]: #1028
-#731 := [intro-def]: #730
-#1030 := [mp #731 #1029]: #1025
-#1032 := [not-or-elim #1030]: #1016
-#1400 := [unit-resolution #1032 #1399]: #241
-#1031 := [not-or-elim #1030]: #1004
-#1401 := [unit-resolution #1031 #1400]: #1001
-#1261 := (or #1000 #997)
-#1262 := [def-axiom]: #1261
-#1402 := [unit-resolution #1262 #1401]: #997
-#1403 := [th-lemma #1400 #1402 #1398]: false
-#1404 := [lemma #1403]: #1009
-#737 := (+ uf_6 #734)
-#740 := (+ uf_4 #737)
-#747 := (>= #740 0::int)
-#248 := (ite #241 uf_5 #224)
-#259 := (* -1::int #248)
-#260 := (+ uf_6 #259)
-#261 := (+ uf_4 #260)
-#566 := (>= #261 0::int)
-#748 := (= #566 #747)
-#741 := (~ #261 #740)
-#738 := (~ #260 #737)
-#735 := (~ #259 #734)
-#732 := (~ #248 z3name!3)
-#733 := [apply-def #731]: #732
-#736 := [monotonicity #733]: #735
-#739 := [monotonicity #736]: #738
-#742 := [monotonicity #739]: #741
-#749 := [monotonicity #742]: #748
-#568 := (not #566)
-#565 := (<= #261 0::int)
-#567 := (not #565)
-#569 := (or #567 #568)
-#570 := (not #569)
-#262 := (= #261 0::int)
-#571 := (iff #262 #570)
-#572 := [rewrite]: #571
-#27 := (- uf_5)
-#26 := (< uf_5 0::int)
-#28 := (ite #26 #27 uf_5)
-#29 := (- #28 uf_4)
-#30 := (= uf_6 #29)
-#267 := (iff #30 #262)
-#227 := (ite #26 #224 uf_5)
-#177 := (* -1::int uf_4)
-#233 := (+ #177 #227)
-#238 := (= uf_6 #233)
-#265 := (iff #238 #262)
-#253 := (+ #177 #248)
-#256 := (= uf_6 #253)
-#263 := (iff #256 #262)
-#264 := [rewrite]: #263
-#257 := (iff #238 #256)
-#254 := (= #233 #253)
-#251 := (= #227 #248)
-#245 := (ite #242 #224 uf_5)
-#249 := (= #245 #248)
-#250 := [rewrite]: #249
-#246 := (= #227 #245)
-#243 := (iff #26 #242)
-#244 := [rewrite]: #243
-#247 := [monotonicity #244]: #246
-#252 := [trans #247 #250]: #251
-#255 := [monotonicity #252]: #254
-#258 := [monotonicity #255]: #257
-#266 := [trans #258 #264]: #265
-#239 := (iff #30 #238)
-#236 := (= #29 #233)
-#230 := (- #227 uf_4)
-#234 := (= #230 #233)
-#235 := [rewrite]: #234
-#231 := (= #29 #230)
-#228 := (= #28 #227)
-#225 := (= #27 #224)
-#226 := [rewrite]: #225
-#229 := [monotonicity #226]: #228
-#232 := [monotonicity #229]: #231
-#237 := [trans #232 #235]: #236
-#240 := [monotonicity #237]: #239
-#268 := [trans #240 #266]: #267
-#223 := [asserted]: #30
-#269 := [mp #223 #268]: #262
-#573 := [mp #269 #572]: #570
-#575 := [not-or-elim #573]: #566
-#750 := [mp~ #575 #749]: #747
-#1364 := (not #747)
-#1357 := (not #769)
-#1337 := (not #851)
-#1555 := (or #194 #476 #1151 #1337 #1343 #1113 #1048 #1357 #1364 #1011)
-#1556 := [th-lemma]: #1555
-#1557 := [unit-resolution #1556 #1545 #750 #1404 #772 #1496 #828 #1554 #854 #1553]: #194
-#195 := (not #194)
-#966 := (or #195 #963)
-#969 := (+ uf_4 z3name!2)
-#971 := (>= #969 0::int)
-#973 := (not #971)
-#970 := (<= #969 0::int)
-#972 := (not #970)
-#974 := (or #972 #973)
-#975 := (not #974)
-#978 := (or #194 #975)
-#985 := (not #978)
-#984 := (not #966)
-#986 := (or #984 #985)
-#987 := (not #986)
-#702 := (= z3name!2 #177)
-#703 := (or #194 #702)
-#700 := (= z3name!2 uf_4)
-#701 := (or #195 #700)
-#704 := (and #701 #703)
-#990 := (iff #704 #987)
-#981 := (and #966 #978)
-#988 := (iff #981 #987)
-#989 := [rewrite]: #988
-#982 := (iff #704 #981)
-#979 := (iff #703 #978)
-#976 := (iff #702 #975)
-#977 := [rewrite]: #976
-#980 := [monotonicity #977]: #979
-#967 := (iff #701 #966)
-#964 := (iff #700 #963)
-#965 := [rewrite]: #964
-#968 := [monotonicity #965]: #967
-#983 := [monotonicity #968 #980]: #982
-#991 := [trans #983 #989]: #990
-#705 := [intro-def]: #704
-#992 := [mp #705 #991]: #987
-#993 := [not-or-elim #992]: #966
-#1558 := [unit-resolution #993 #1557]: #963
-#1249 := (or #962 #959)
-#1250 := [def-axiom]: #1249
-#1559 := [unit-resolution #1250 #1558]: #959
-decl z3name!8 :: int
-#855 := z3name!8
-#864 := (* -1::int z3name!8)
-#867 := (+ uf_11 #864)
-#870 := (+ uf_9 #867)
-#873 := (<= #870 0::int)
-#483 := (ite #476 uf_10 #459)
-#494 := (* -1::int #483)
-#495 := (+ uf_11 #494)
-#496 := (+ uf_9 #495)
-#620 := (<= #496 0::int)
-#874 := (= #620 #873)
-#871 := (~ #496 #870)
-#868 := (~ #495 #867)
-#865 := (~ #494 #864)
-#862 := (~ #483 z3name!8)
-#858 := (= z3name!8 #459)
-#859 := (or #476 #858)
-#856 := (= z3name!8 uf_10)
-#857 := (or #477 #856)
-#860 := (and #857 #859)
-#861 := [intro-def]: #860
-#863 := [apply-def #861]: #862
-#866 := [monotonicity #863]: #865
-#869 := [monotonicity #866]: #868
-#872 := [monotonicity #869]: #871
-#875 := [monotonicity #872]: #874
-#621 := (>= #496 0::int)
-#623 := (not #621)
-#622 := (not #620)
-#624 := (or #622 #623)
-#625 := (not #624)
-#497 := (= #496 0::int)
-#626 := (iff #497 #625)
-#627 := [rewrite]: #626
-#57 := (- uf_10)
-#56 := (< uf_10 0::int)
-#58 := (ite #56 #57 uf_10)
-#59 := (- #58 uf_9)
-#60 := (= uf_11 #59)
-#502 := (iff #60 #497)
-#462 := (ite #56 #459 uf_10)
-#468 := (+ #412 #462)
-#473 := (= uf_11 #468)
-#500 := (iff #473 #497)
-#488 := (+ #412 #483)
-#491 := (= uf_11 #488)
-#498 := (iff #491 #497)
-#499 := [rewrite]: #498
-#492 := (iff #473 #491)
-#489 := (= #468 #488)
-#486 := (= #462 #483)
-#480 := (ite #477 #459 uf_10)
-#484 := (= #480 #483)
-#485 := [rewrite]: #484
-#481 := (= #462 #480)
-#478 := (iff #56 #477)
-#479 := [rewrite]: #478
-#482 := [monotonicity #479]: #481
-#487 := [trans #482 #485]: #486
-#490 := [monotonicity #487]: #489
-#493 := [monotonicity #490]: #492
-#501 := [trans #493 #499]: #500
-#474 := (iff #60 #473)
-#471 := (= #59 #468)
-#465 := (- #462 uf_9)
-#469 := (= #465 #468)
-#470 := [rewrite]: #469
-#466 := (= #59 #465)
-#463 := (= #58 #462)
-#460 := (= #57 #459)
-#461 := [rewrite]: #460
-#464 := [monotonicity #461]: #463
-#467 := [monotonicity #464]: #466
-#472 := [trans #467 #470]: #471
-#475 := [monotonicity #472]: #474
-#503 := [trans #475 #501]: #502
-#458 := [asserted]: #60
-#504 := [mp #458 #503]: #497
-#628 := [mp #504 #627]: #625
-#629 := [not-or-elim #628]: #620
-#876 := [mp~ #629 #875]: #873
-#1197 := (+ uf_10 z3name!8)
-#1198 := (<= #1197 0::int)
-#1199 := (>= #1197 0::int)
-#1201 := (not #1199)
-#1200 := (not #1198)
-#1202 := (or #1200 #1201)
-#1203 := (not #1202)
-#1206 := (or #476 #1203)
-#1213 := (not #1206)
-#1186 := (+ uf_10 #864)
-#1187 := (<= #1186 0::int)
-#1189 := (not #1187)
-#1185 := (>= #1186 0::int)
-#1188 := (not #1185)
-#1190 := (or #1188 #1189)
-#1191 := (not #1190)
-#1194 := (or #477 #1191)
-#1212 := (not #1194)
-#1214 := (or #1212 #1213)
-#1215 := (not #1214)
-#1218 := (iff #860 #1215)
-#1209 := (and #1194 #1206)
-#1216 := (iff #1209 #1215)
-#1217 := [rewrite]: #1216
-#1210 := (iff #860 #1209)
-#1207 := (iff #859 #1206)
-#1204 := (iff #858 #1203)
-#1205 := [rewrite]: #1204
-#1208 := [monotonicity #1205]: #1207
-#1195 := (iff #857 #1194)
-#1192 := (iff #856 #1191)
-#1193 := [rewrite]: #1192
-#1196 := [monotonicity #1193]: #1195
-#1211 := [monotonicity #1196 #1208]: #1210
-#1219 := [trans #1211 #1217]: #1218
-#1220 := [mp #861 #1219]: #1215
-#1222 := [not-or-elim #1220]: #1206
-#1560 := [unit-resolution #1222 #1553]: #1203
-#1325 := (or #1202 #1198)
-#1326 := [def-axiom]: #1325
-#1561 := [unit-resolution #1326 #1560]: #1198
-#711 := (+ uf_5 #708)
-#714 := (+ uf_1 #711)
-#721 := (>= #714 0::int)
-#201 := (ite #194 uf_4 #177)
-#212 := (* -1::int #201)
-#213 := (+ uf_5 #212)
-#214 := (+ uf_1 #213)
-#555 := (>= #214 0::int)
-#722 := (= #555 #721)
-#715 := (~ #214 #714)
-#712 := (~ #213 #711)
-#709 := (~ #212 #708)
-#706 := (~ #201 z3name!2)
-#707 := [apply-def #705]: #706
-#710 := [monotonicity #707]: #709
-#713 := [monotonicity #710]: #712
-#716 := [monotonicity #713]: #715
-#723 := [monotonicity #716]: #722
-#557 := (not #555)
-#554 := (<= #214 0::int)
-#556 := (not #554)
-#558 := (or #556 #557)
-#559 := (not #558)
-#215 := (= #214 0::int)
-#560 := (iff #215 #559)
-#561 := [rewrite]: #560
-#21 := (- uf_4)
-#20 := (< uf_4 0::int)
-#22 := (ite #20 #21 uf_4)
-#23 := (- #22 uf_1)
-#24 := (= uf_5 #23)
-#220 := (iff #24 #215)
-#180 := (ite #20 #177 uf_4)
-#186 := (+ #130 #180)
-#191 := (= uf_5 #186)
-#218 := (iff #191 #215)
-#206 := (+ #130 #201)
-#209 := (= uf_5 #206)
-#216 := (iff #209 #215)
-#217 := [rewrite]: #216
-#210 := (iff #191 #209)
-#207 := (= #186 #206)
-#204 := (= #180 #201)
-#198 := (ite #195 #177 uf_4)
-#202 := (= #198 #201)
-#203 := [rewrite]: #202
-#199 := (= #180 #198)
-#196 := (iff #20 #195)
-#197 := [rewrite]: #196
-#200 := [monotonicity #197]: #199
-#205 := [trans #200 #203]: #204
-#208 := [monotonicity #205]: #207
-#211 := [monotonicity #208]: #210
-#219 := [trans #211 #217]: #218
-#192 := (iff #24 #191)
-#189 := (= #23 #186)
-#183 := (- #180 uf_1)
-#187 := (= #183 #186)
-#188 := [rewrite]: #187
-#184 := (= #23 #183)
-#181 := (= #22 #180)
-#178 := (= #21 #177)
-#179 := [rewrite]: #178
-#182 := [monotonicity #179]: #181
-#185 := [monotonicity #182]: #184
-#190 := [trans #185 #188]: #189
-#193 := [monotonicity #190]: #192
-#221 := [trans #193 #219]: #220
-#176 := [asserted]: #24
-#222 := [mp #176 #221]: #215
-#562 := [mp #222 #561]: #559
-#564 := [not-or-elim #562]: #555
-#724 := [mp~ #564 #723]: #721
-#685 := (+ uf_4 #682)
-#688 := (+ uf_2 #685)
-#695 := (>= #688 0::int)
-#154 := (ite #147 uf_1 #130)
-#165 := (* -1::int #154)
-#166 := (+ uf_4 #165)
-#167 := (+ uf_2 #166)
-#544 := (>= #167 0::int)
-#696 := (= #544 #695)
-#689 := (~ #167 #688)
-#686 := (~ #166 #685)
-#683 := (~ #165 #682)
-#680 := (~ #154 z3name!1)
-#681 := [apply-def #679]: #680
-#684 := [monotonicity #681]: #683
-#687 := [monotonicity #684]: #686
-#690 := [monotonicity #687]: #689
-#697 := [monotonicity #690]: #696
-#546 := (not #544)
-#543 := (<= #167 0::int)
-#545 := (not #543)
-#547 := (or #545 #546)
-#548 := (not #547)
-#168 := (= #167 0::int)
-#549 := (iff #168 #548)
-#550 := [rewrite]: #549
-#15 := (- uf_1)
-#14 := (< uf_1 0::int)
-#16 := (ite #14 #15 uf_1)
-#17 := (- #16 uf_2)
-#18 := (= uf_4 #17)
-#173 := (iff #18 #168)
-#133 := (ite #14 #130 uf_1)
-#139 := (+ #82 #133)
-#144 := (= uf_4 #139)
-#171 := (iff #144 #168)
-#159 := (+ #82 #154)
-#162 := (= uf_4 #159)
-#169 := (iff #162 #168)
-#170 := [rewrite]: #169
-#163 := (iff #144 #162)
-#160 := (= #139 #159)
-#157 := (= #133 #154)
-#151 := (ite #148 #130 uf_1)
-#155 := (= #151 #154)
-#156 := [rewrite]: #155
-#152 := (= #133 #151)
-#149 := (iff #14 #148)
-#150 := [rewrite]: #149
-#153 := [monotonicity #150]: #152
-#158 := [trans #153 #156]: #157
-#161 := [monotonicity #158]: #160
-#164 := [monotonicity #161]: #163
-#172 := [trans #164 #170]: #171
-#145 := (iff #18 #144)
-#142 := (= #17 #139)
-#136 := (- #133 uf_2)
-#140 := (= #136 #139)
-#141 := [rewrite]: #140
-#137 := (= #17 #136)
-#134 := (= #16 #133)
-#131 := (= #15 #130)
-#132 := [rewrite]: #131
-#135 := [monotonicity #132]: #134
-#138 := [monotonicity #135]: #137
-#143 := [trans #138 #141]: #142
-#146 := [monotonicity #143]: #145
-#174 := [trans #146 #172]: #173
-#129 := [asserted]: #18
-#175 := [mp #129 #174]: #168
-#551 := [mp #175 #550]: #548
-#553 := [not-or-elim #551]: #544
-#698 := [mp~ #553 #697]: #695
-#1373 := (not #721)
-#1562 := (or #147 #1373 #961 #241 #195)
-#1563 := [th-lemma]: #1562
-#1564 := [unit-resolution #1563 #1559 #1557 #724 #1406]: #147
-#1565 := [unit-resolution #955 #1564]: #925
-#1566 := [unit-resolution #1238 #1565]: #921
-#1372 := (not #873)
-#1371 := (not #695)
-#1498 := (or #516 #923 #1373 #1371 #1372 #1343 #1200 #1075 #1350 #1113 #961 #1151 #1337 #1048 #1357)
-#1499 := [th-lemma]: #1498
-#1567 := [unit-resolution #1499 #1566 #698 #724 #1545 #772 #1447 #802 #1496 #828 #1554 #854 #1561 #876 #1559]: #516
-#1247 := (or #962 #957)
-#1248 := [def-axiom]: #1247
-#1568 := [unit-resolution #1248 #1558]: #957
-#877 := (>= #870 0::int)
-#878 := (= #621 #877)
-#879 := [monotonicity #872]: #878
-#630 := [not-or-elim #628]: #621
-#880 := [mp~ #630 #879]: #877
-#1327 := (or #1202 #1199)
-#1328 := [def-axiom]: #1327
-#1569 := [unit-resolution #1328 #1560]: #1199
-#795 := (<= #792 0::int)
-#796 := (= #587 #795)
-#797 := [monotonicity #794]: #796
-#596 := [not-or-elim #595]: #587
-#798 := [mp~ #596 #797]: #795
-#1503 := (or #335 #1049 #241)
-#1425 := (or #335 #1049 #241 #1037)
-#1422 := [hypothesis]: #336
-#1423 := [hypothesis]: #1047
-#1424 := [th-lemma #1423 #1422 #776 #1406 #1421]: false
-#1426 := [lemma #1424]: #1425
-#1504 := [unit-resolution #1426 #1459]: #1503
-#1505 := [unit-resolution #1504 #1406 #1468]: #335
-#1506 := [unit-resolution #1107 #1505]: #1077
-#1283 := (or #1076 #1071)
-#1284 := [def-axiom]: #1283
-#1507 := [unit-resolution #1284 #1506]: #1071
-#717 := (<= #714 0::int)
-#718 := (= #554 #717)
-#719 := [monotonicity #716]: #718
-#563 := [not-or-elim #562]: #554
-#720 := [mp~ #563 #719]: #717
-#691 := (<= #688 0::int)
-#692 := (= #543 #691)
-#693 := [monotonicity #690]: #692
-#552 := [not-or-elim #551]: #543
-#694 := [mp~ #552 #693]: #691
-#1235 := (or #924 #919)
-#1236 := [def-axiom]: #1235
-#1570 := [unit-resolution #1236 #1565]: #919
-#1409 := (not #773)
-#1489 := (not #847)
-#1358 := (not #795)
-#1365 := (not #821)
-#1511 := (not #877)
-#1510 := (not #691)
-#1509 := (not #717)
-#1512 := (or #515 #922 #1509 #1510 #1511 #1365 #1201 #1074 #1358 #1112 #960 #1150 #1489 #1049 #1409)
-#1513 := [th-lemma]: #1512
-#1571 := [unit-resolution #1513 #1570 #694 #720 #1468 #776 #1507 #798 #1488 #824 #1552 #850 #1569 #880 #1568]: #515
-#506 := (<= #508 0::int)
-#659 := (+ uf_3 #656)
-#662 := (+ uf_1 #659)
-#665 := (<= #662 0::int)
-#107 := (ite #100 uf_2 #82)
-#118 := (* -1::int #107)
-#119 := (+ uf_3 #118)
-#120 := (+ uf_1 #119)
-#532 := (<= #120 0::int)
-#666 := (= #532 #665)
-#663 := (~ #120 #662)
-#660 := (~ #119 #659)
-#657 := (~ #118 #656)
-#654 := (~ #107 z3name!0)
-#655 := [apply-def #653]: #654
-#658 := [monotonicity #655]: #657
-#661 := [monotonicity #658]: #660
-#664 := [monotonicity #661]: #663
-#667 := [monotonicity #664]: #666
-#533 := (>= #120 0::int)
-#535 := (not #533)
-#534 := (not #532)
-#536 := (or #534 #535)
-#537 := (not #536)
-#121 := (= #120 0::int)
-#538 := (iff #121 #537)
-#539 := [rewrite]: #538
-#8 := (- uf_2)
-#7 := (< uf_2 0::int)
-#9 := (ite #7 #8 uf_2)
-#11 := (- #9 uf_3)
-#12 := (= uf_1 #11)
-#126 := (iff #12 #121)
-#85 := (ite #7 #82 uf_2)
-#91 := (* -1::int uf_3)
-#92 := (+ #91 #85)
-#97 := (= uf_1 #92)
-#124 := (iff #97 #121)
-#112 := (+ #91 #107)
-#115 := (= uf_1 #112)
-#122 := (iff #115 #121)
-#123 := [rewrite]: #122
-#116 := (iff #97 #115)
-#113 := (= #92 #112)
-#110 := (= #85 #107)
-#104 := (ite #101 #82 uf_2)
-#108 := (= #104 #107)
-#109 := [rewrite]: #108
-#105 := (= #85 #104)
-#102 := (iff #7 #101)
-#103 := [rewrite]: #102
-#106 := [monotonicity #103]: #105
-#111 := [trans #106 #109]: #110
-#114 := [monotonicity #111]: #113
-#117 := [monotonicity #114]: #116
-#125 := [trans #117 #123]: #124
-#98 := (iff #12 #97)
-#95 := (= #11 #92)
-#88 := (- #85 uf_3)
-#93 := (= #88 #92)
-#94 := [rewrite]: #93
-#89 := (= #11 #88)
-#86 := (= #9 #85)
-#83 := (= #8 #82)
-#84 := [rewrite]: #83
-#87 := [monotonicity #84]: #86
-#90 := [monotonicity #87]: #89
-#96 := [trans #90 #94]: #95
-#99 := [monotonicity #96]: #98
-#127 := [trans #99 #125]: #126
-#80 := [asserted]: #12
-#128 := [mp #80 #127]: #121
-#540 := [mp #128 #539]: #537
-#541 := [not-or-elim #540]: #532
-#668 := [mp~ #541 #667]: #665
-#1515 := (or #100 #241 #923 #1373 #1371 #961)
-#1516 := [th-lemma]: #1515
-#1572 := [unit-resolution #1516 #1566 #698 #1559 #724 #1406]: #100
-#1573 := [unit-resolution #917 #1572]: #887
-#1223 := (or #886 #881)
-#1224 := [def-axiom]: #1223
-#1574 := [unit-resolution #1224 #1573]: #881
-#1528 := (not #665)
-#1529 := (or #506 #884 #1528 #1364 #1011 #1343 #1113 #1151 #1337 #1048 #1357 #922 #1510)
-#1530 := [th-lemma]: #1529
-#1575 := [unit-resolution #1530 #1574 #668 #694 #1404 #750 #1545 #772 #1496 #828 #1554 #854 #1570]: #506
-#743 := (<= #740 0::int)
-#744 := (= #565 #743)
-#745 := [monotonicity #742]: #744
-#574 := [not-or-elim #573]: #565
-#746 := [mp~ #574 #745]: #743
-#1520 := [unit-resolution #1032 #1406]: #1013
-#1265 := (or #1012 #1008)
-#1266 := [def-axiom]: #1265
-#1521 := [unit-resolution #1266 #1520]: #1008
-#669 := (>= #662 0::int)
-#670 := (= #533 #669)
-#671 := [monotonicity #664]: #670
-#542 := [not-or-elim #540]: #533
-#672 := [mp~ #542 #671]: #669
-#1576 := [unit-resolution #1226 #1573]: #883
-#1523 := (not #743)
-#1522 := (not #669)
-#1524 := (or #509 #885 #1522 #1523 #1010 #1365 #1112 #1150 #1489 #1049 #1409 #923 #1371)
-#1525 := [th-lemma]: #1524
-#1577 := [unit-resolution #1525 #1576 #672 #698 #1521 #746 #1468 #776 #1488 #824 #1552 #850 #1566]: #509
-#634 := (not #516)
-#633 := (not #515)
-#632 := (not #509)
-#631 := (not #506)
-#635 := (or #631 #632 #633 #634)
-#523 := (and #506 #509 #515 #516)
-#528 := (not #523)
-#644 := (iff #528 #635)
-#636 := (not #635)
-#639 := (not #636)
-#642 := (iff #639 #635)
-#643 := [rewrite]: #642
-#640 := (iff #528 #639)
-#637 := (iff #523 #636)
-#638 := [rewrite]: #637
-#641 := [monotonicity #638]: #640
-#645 := [trans #641 #643]: #644
-#62 := (= uf_2 uf_11)
-#61 := (= uf_3 uf_10)
-#63 := (and #61 #62)
-#64 := (not #63)
-#529 := (iff #64 #528)
-#526 := (iff #63 #523)
-#517 := (and #515 #516)
-#510 := (and #506 #509)
-#520 := (and #510 #517)
-#524 := (iff #520 #523)
-#525 := [rewrite]: #524
-#521 := (iff #63 #520)
-#518 := (iff #62 #517)
-#519 := [rewrite]: #518
-#511 := (iff #61 #510)
-#512 := [rewrite]: #511
-#522 := [monotonicity #512 #519]: #521
-#527 := [trans #522 #525]: #526
-#530 := [monotonicity #527]: #529
-#505 := [asserted]: #64
-#531 := [mp #505 #530]: #528
-#646 := [mp #531 #645]: #635
-#1578 := [unit-resolution #646 #1577 #1575 #1571 #1567]: false
-#1580 := [lemma #1578]: #1579
-#1657 := [unit-resolution #1580 #1406]: #1048
-#1625 := (or #194 #241)
-#1535 := [hypothesis]: #195
-#1538 := (or #194 #960)
-#1432 := [hypothesis]: #973
-#1255 := (or #974 #971)
-#1256 := [def-axiom]: #1255
-#1433 := [unit-resolution #1256 #1432]: #974
-#994 := [not-or-elim #992]: #978
-#1434 := [unit-resolution #994 #1433]: #194
-#1435 := [unit-resolution #993 #1434]: #963
-#1436 := (or #971 #195 #961)
-#1437 := [th-lemma]: #1436
-#1438 := [unit-resolution #1437 #1434 #1432]: #961
-#1439 := [unit-resolution #1250 #1438 #1435]: false
-#1440 := [lemma #1439]: #971
-#1536 := [hypothesis]: #957
-#1537 := [th-lemma #1536 #1535 #1440]: false
-#1539 := [lemma #1537]: #1538
-#1581 := [unit-resolution #1539 #1535]: #960
-#1582 := (or #959 #957)
-#1583 := [th-lemma]: #1582
-#1584 := [unit-resolution #1583 #1581]: #959
-#1585 := (or #147 #1373 #241 #194 #973)
-#1586 := [th-lemma]: #1585
-#1587 := [unit-resolution #1586 #1535 #1440 #724 #1406]: #147
-#1588 := [unit-resolution #955 #1587]: #925
-#1589 := [unit-resolution #1238 #1588]: #921
-#1590 := [unit-resolution #1516 #1589 #698 #1584 #724 #1406]: #100
-#1591 := [unit-resolution #917 #1590]: #887
-#1592 := [unit-resolution #1224 #1591]: #881
-#1593 := (or #430 #1365 #1074 #1358 #1112 #194 #1364 #1011 #241)
-#1594 := [th-lemma]: #1593
-#1595 := [unit-resolution #1594 #1535 #1404 #750 #1507 #798 #1488 #824 #1406]: #430
-#1184 := [not-or-elim #1182]: #1168
-#1596 := [unit-resolution #1184 #1595]: #1165
-#1315 := (or #1164 #1161)
-#1316 := [def-axiom]: #1315
-#1597 := [unit-resolution #1316 #1596]: #1161
-#1533 := (or #288 #241)
-#1471 := (or #194 #288 #241)
-#1469 := (or #194 #288 #241 #1364 #1011)
-#1470 := [th-lemma]: #1469
-#1472 := [unit-resolution #1470 #1404 #750]: #1471
-#1473 := [unit-resolution #1472 #1405 #1406]: #194
-#1474 := [unit-resolution #993 #1473]: #963
-#1475 := [unit-resolution #1250 #1474]: #959
-#1476 := (or #147 #1373 #1364 #1011 #961 #241 #288)
-#1477 := [th-lemma]: #1476
-#1478 := [unit-resolution #1477 #1475 #724 #1406 #1404 #750 #1405]: #147
-#1479 := [unit-resolution #955 #1478]: #925
-#1480 := [unit-resolution #1238 #1479]: #921
-#1419 := (or #288 #241 #429)
-#1333 := [hypothesis]: #430
-#1408 := [unit-resolution #1280 #1407]: #1047
-#1410 := (or #335 #1049 #1409 #288 #241)
-#1411 := [th-lemma]: #1410
-#1412 := [unit-resolution #1411 #1405 #1408 #776 #1406]: #335
-#1413 := [unit-resolution #1107 #1412]: #1077
-#1414 := [unit-resolution #1286 #1413]: #1073
-#1415 := [unit-resolution #1352 #1414 #802 #1405 #828 #1333]: #1113
-#1416 := [unit-resolution #1298 #1415]: #1114
-#1417 := [unit-resolution #1145 #1416]: #383
-#1418 := [th-lemma #1414 #802 #1405 #1408 #776 #1406 #1417]: false
-#1420 := [lemma #1418]: #1419
-#1481 := [unit-resolution #1420 #1405 #1406]: #429
-#1482 := [unit-resolution #1183 #1481]: #1153
-#1483 := [unit-resolution #1308 #1482]: #1147
-#1490 := (or #477 #1150 #1489 #1365 #1112 #1049 #241 #1409 #288)
-#1491 := [th-lemma]: #1490
-#1492 := [unit-resolution #1491 #1405 #1468 #776 #1488 #824 #1483 #850 #1406]: #477
-#1493 := [unit-resolution #1222 #1492]: #1203
-#1494 := [unit-resolution #1326 #1493]: #1198
-#1495 := [unit-resolution #1310 #1482]: #1149
-#1500 := [unit-resolution #1499 #1475 #698 #724 #1497 #772 #1447 #802 #1496 #828 #1495 #854 #1494 #876 #1480]: #516
-#1501 := [unit-resolution #1236 #1479]: #919
-#1502 := [unit-resolution #1328 #1493]: #1199
-#1508 := [unit-resolution #1248 #1474]: #957
-#1514 := [unit-resolution #1513 #1508 #694 #720 #1468 #776 #1507 #798 #1488 #824 #1483 #850 #1502 #880 #1501]: #515
-#1517 := [unit-resolution #1516 #1480 #698 #1475 #724 #1406]: #100
-#1518 := [unit-resolution #917 #1517]: #887
-#1519 := [unit-resolution #1226 #1518]: #883
-#1526 := [unit-resolution #1525 #1480 #672 #698 #1521 #746 #1468 #776 #1488 #824 #1483 #850 #1519]: #509
-#1527 := [unit-resolution #1224 #1518]: #881
-#1531 := [unit-resolution #1530 #1501 #668 #694 #1404 #750 #1497 #772 #1496 #828 #1495 #854 #1527]: #506
-#1532 := [unit-resolution #646 #1531 #1526 #1514 #1500]: false
-#1534 := [lemma #1532]: #1533
-#1598 := [unit-resolution #1534 #1406]: #288
-#1599 := [unit-resolution #1069 #1598]: #1039
-#1271 := (or #1038 #1033)
-#1272 := [def-axiom]: #1271
-#1600 := [unit-resolution #1272 #1599]: #1033
-#1601 := [unit-resolution #1236 #1588]: #919
-#1602 := (or #506 #884 #1528 #1364 #1011 #1365 #1112 #1337 #1357 #922 #1510 #1036 #1163 #1074 #1358)
-#1603 := [th-lemma]: #1602
-#1604 := [unit-resolution #1603 #1601 #668 #694 #1404 #750 #1600 #772 #1507 #798 #1488 #824 #1597 #854 #1592]: #506
-#1605 := [unit-resolution #1226 #1591]: #883
-#1313 := (or #1164 #1160)
-#1314 := [def-axiom]: #1313
-#1606 := [unit-resolution #1314 #1596]: #1160
-#1607 := (or #509 #885 #1522 #1523 #1010 #1343 #1113 #1489 #1409 #923 #1371 #1037 #1162 #1075 #1350)
-#1608 := [th-lemma]: #1607
-#1609 := [unit-resolution #1608 #1589 #672 #698 #1521 #746 #1459 #776 #1447 #802 #1496 #828 #1606 #850 #1605]: #509
-#1610 := (or #476 #1036 #1337 #1365 #1112 #1357 #194 #1364 #1011 #1163 #1074 #1358)
-#1611 := [th-lemma]: #1610
-#1612 := [unit-resolution #1611 #1597 #750 #1600 #772 #1507 #798 #1488 #824 #1404 #854 #1535]: #476
-#1221 := [not-or-elim #1220]: #1194
-#1613 := [unit-resolution #1221 #1612]: #1191
-#1319 := (or #1190 #1185)
-#1320 := [def-axiom]: #1319
-#1614 := [unit-resolution #1320 #1613]: #1185
-#1615 := (or #516 #923 #1373 #1371 #1372 #1075 #1350 #1489 #1409 #1037 #973 #1162 #1188 #1343 #1113 #1523 #1010)
-#1616 := [th-lemma]: #1615
-#1617 := [unit-resolution #1616 #1606 #1440 #724 #1521 #746 #1459 #776 #1447 #802 #1496 #828 #698 #850 #1614 #876 #1589]: #516
-#1321 := (or #1190 #1187)
-#1322 := [def-axiom]: #1321
-#1618 := [unit-resolution #1322 #1613]: #1187
-#1619 := [unit-resolution #994 #1535]: #975
-#1253 := (or #974 #970)
-#1254 := [def-axiom]: #1253
-#1620 := [unit-resolution #1254 #1619]: #970
-#1621 := (or #515 #922 #1509 #1510 #1511 #1074 #1358 #1337 #1357 #1036 #972 #1163 #1189 #1365 #1112 #1364 #1011)
-#1622 := [th-lemma]: #1621
-#1623 := [unit-resolution #1622 #1620 #694 #720 #1404 #750 #1600 #772 #1507 #798 #1488 #824 #1597 #854 #1618 #880 #1601]: #515
-#1624 := [unit-resolution #646 #1623 #1617 #1609 #1604]: false
-#1626 := [lemma #1624]: #1625
-#1658 := [unit-resolution #1626 #1406]: #194
-#1659 := [unit-resolution #993 #1658]: #963
-#1660 := [unit-resolution #1250 #1659]: #959
-#1661 := [unit-resolution #1563 #1660 #1658 #724 #1406]: #147
-#1662 := [unit-resolution #955 #1661]: #925
-#1663 := [unit-resolution #1238 #1662]: #921
-#1664 := [unit-resolution #1516 #1663 #698 #1660 #724 #1406]: #100
-#1665 := [unit-resolution #917 #1664]: #887
-#1666 := [unit-resolution #1226 #1665]: #883
-#1667 := [unit-resolution #1224 #1665]: #881
-#1668 := [unit-resolution #1236 #1662]: #919
-#1669 := [unit-resolution #1248 #1659]: #957
-#1655 := (or #429 #1113 #1010 #960 #1036 #1074 #1112 #922 #923 #884 #885)
-#1632 := [hypothesis]: #919
-#1636 := [hypothesis]: #881
-#1638 := [hypothesis]: #1071
-#1639 := [hypothesis]: #1033
-#1334 := [unit-resolution #1184 #1333]: #1165
-#1335 := [unit-resolution #1316 #1334]: #1161
-#1640 := [unit-resolution #1603 #1335 #668 #694 #1404 #750 #1639 #772 #1638 #798 #1637 #824 #1632 #854 #1636]: #506
-#1641 := [hypothesis]: #883
-#1642 := [hypothesis]: #921
-#1643 := [hypothesis]: #1111
-#1644 := [hypothesis]: #1008
-#1631 := [unit-resolution #1314 #1334]: #1160
-#1645 := [unit-resolution #1608 #1631 #672 #698 #1644 #746 #1459 #776 #1447 #802 #1643 #828 #1642 #850 #1641]: #509
-#1634 := (or #1202 #922 #960 #632 #631 #429)
-#1627 := [hypothesis]: #506
-#1628 := [hypothesis]: #509
-#1384 := [hypothesis]: #1203
-#1396 := (or #1202 #516 #429)
-#1331 := [hypothesis]: #634
-#1385 := [unit-resolution #1326 #1384]: #1198
-#1382 := (or #1189 #1200 #516 #429)
-#1332 := [hypothesis]: #1198
-#1336 := [hypothesis]: #1187
-#1338 := (or #382 #1189 #1337 #429 #1163 #1200)
-#1339 := [th-lemma]: #1338
-#1340 := [unit-resolution #1339 #1336 #1335 #854 #1333 #1332]: #382
-#1341 := [unit-resolution #1145 #1340]: #1115
-#1342 := [unit-resolution #1298 #1341]: #1111
-#1344 := (or #335 #1113 #429 #1343 #1189 #1337 #1163 #1200)
-#1345 := [th-lemma]: #1344
-#1346 := [unit-resolution #1345 #1342 #828 #1333 #1335 #854 #1336 #1332]: #335
-#1347 := [unit-resolution #1107 #1346]: #1077
-#1348 := [unit-resolution #1284 #1347]: #1071
-#1349 := [unit-resolution #1286 #1347]: #1073
-#1353 := [unit-resolution #1352 #1349 #802 #1342 #828 #1333]: #288
-#1354 := [unit-resolution #1069 #1353]: #1039
-#1355 := [unit-resolution #1272 #1354]: #1033
-#1356 := [unit-resolution #1296 #1341]: #1109
-#1359 := (or #242 #1036 #1357 #429 #1189 #1337 #1163 #1200 #1074 #1358)
-#1360 := [th-lemma]: #1359
-#1361 := [unit-resolution #1360 #1355 #772 #1348 #798 #1333 #1335 #854 #1336 #1332]: #242
-#1362 := [unit-resolution #1032 #1361]: #1013
-#1363 := [unit-resolution #1268 #1362]: #1009
-#1366 := (or #194 #1011 #1364 #1074 #1358 #1112 #1365 #1036 #1357 #1189 #1337 #1163 #1200)
-#1367 := [th-lemma]: #1366
-#1368 := [unit-resolution #1367 #1363 #750 #1355 #772 #1348 #798 #1356 #824 #1335 #854 #1336 #1332]: #194
-#1369 := [unit-resolution #993 #1368]: #963
-#1370 := [unit-resolution #1250 #1369]: #959
-#1374 := (or #923 #1371 #516 #1372 #1200 #961 #1373 #1036 #1357 #1337 #1163 #1074 #1358)
-#1375 := [th-lemma]: #1374
-#1376 := [unit-resolution #1375 #1370 #698 #724 #1355 #772 #1348 #798 #1335 #854 #1332 #876 #1331]: #923
-#1377 := (or #147 #195 #961 #1373 #1036 #1357 #429 #1189 #1337 #1163 #1200 #1074 #1358)
-#1378 := [th-lemma]: #1377
-#1379 := [unit-resolution #1378 #1368 #1370 #724 #1355 #772 #1348 #798 #1333 #1335 #854 #1336 #1332]: #147
-#1380 := [unit-resolution #955 #1379]: #925
-#1381 := [unit-resolution #1238 #1380 #1376]: false
-#1383 := [lemma #1381]: #1382
-#1386 := [unit-resolution #1383 #1385 #1331 #1333]: #1189
-#1387 := [unit-resolution #1322 #1386]: #1190
-#1388 := [unit-resolution #1328 #1384]: #1199
-#1389 := (or #1187 #1185)
-#1390 := [th-lemma]: #1389
-#1391 := [unit-resolution #1390 #1386]: #1185
-#1392 := (or #476 #1188 #1201)
-#1393 := [th-lemma]: #1392
-#1394 := [unit-resolution #1393 #1391 #1388]: #476
-#1395 := [unit-resolution #1221 #1394 #1387]: false
-#1397 := [lemma #1395]: #1396
-#1629 := [unit-resolution #1397 #1384 #1333]: #516
-#1630 := [unit-resolution #646 #1629 #1628 #1627]: #633
-#1633 := [th-lemma #1632 #720 #694 #880 #1447 #802 #850 #776 #1459 #1631 #1536 #1388 #1630]: false
-#1635 := [lemma #1633]: #1634
-#1646 := [unit-resolution #1635 #1645 #1536 #1632 #1640 #1333]: #1202
-#1647 := [unit-resolution #1222 #1646]: #476
-#1648 := [unit-resolution #1221 #1647]: #1191
-#1649 := [unit-resolution #1322 #1648]: #1187
-#1650 := [unit-resolution #1320 #1648]: #1185
-#1651 := [unit-resolution #1616 #1650 #1440 #724 #1644 #746 #1459 #776 #1447 #802 #1643 #828 #698 #850 #1631 #876 #1642]: #516
-#1652 := [unit-resolution #646 #1651 #1645 #1640]: #633
-#1653 := [unit-resolution #1622 #1652 #694 #720 #1404 #750 #1639 #772 #1638 #798 #1637 #824 #1335 #854 #1649 #880 #1632]: #972
-#1654 := [th-lemma #1459 #1647 #850 #828 #1643 #776 #746 #1644 #1631 #1447 #802 #1536 #1653]: false
-#1656 := [lemma #1654]: #1655
-#1670 := [unit-resolution #1656 #1496 #1521 #1669 #1600 #1507 #1488 #1668 #1663 #1667 #1666]: #429
-#1671 := [th-lemma #1600 #1670 #824 #1507 #798 #1488 #1657]: false
-#1672 := [lemma #1671]: #241
-#1683 := [unit-resolution #1031 #1672]: #1001
-#1703 := [unit-resolution #1262 #1683]: #997
-#1920 := (or #194 #242 #1364 #999 #288)
-#1921 := [th-lemma]: #1920
-#1922 := [unit-resolution #1921 #1405 #1703 #750 #1672]: #194
-#1923 := [unit-resolution #993 #1922]: #963
-#1924 := [unit-resolution #1248 #1923]: #957
-#1925 := [unit-resolution #1250 #1923]: #959
-#1843 := (or #288 #961 #147)
-#1763 := [hypothesis]: #148
-#1828 := [hypothesis]: #959
-#1842 := [th-lemma #724 #750 #1703 #1828 #1405 #1763]: false
-#1844 := [lemma #1842]: #1843
-#1926 := [unit-resolution #1844 #1925 #1405]: #147
-#1927 := [unit-resolution #955 #1926]: #925
-#1928 := [unit-resolution #1236 #1927]: #919
-#2116 := [unit-resolution #1310 #2047]: #1149
-#2084 := (or #288 #516)
-#2050 := (or #288 #961 #516)
-#2037 := [hypothesis]: #1087
-#2038 := [unit-resolution #1292 #2037]: #1088
-#2041 := (or #1085 #336)
-#2039 := (or #1085 #1075 #336)
-#2040 := [th-lemma]: #2039
-#2042 := [unit-resolution #2040 #1447]: #2041
-#2043 := [unit-resolution #2042 #2037]: #336
-#2044 := [unit-resolution #1108 #2043 #2038]: false
-#2045 := [lemma #2044]: #1085
-#2035 := (or #1087 #1150 #961 #1048 #516)
-#1845 := [hypothesis]: #1085
-#1874 := [hypothesis]: #477
-#1901 := (or #335 #476)
-#1895 := [unit-resolution #1222 #1874]: #1203
-#1896 := [unit-resolution #1326 #1895]: #1198
-#1893 := (or #429 #1200)
-#1880 := (or #335 #1113 #429 #1163 #1200)
-#1857 := [hypothesis]: #1189
-#1858 := [unit-resolution #1322 #1857]: #1190
-#1859 := [unit-resolution #1221 #1858]: #477
-#1860 := [unit-resolution #1222 #1859]: #1203
-#1861 := [unit-resolution #1390 #1857]: #1185
-#1862 := [unit-resolution #1393 #1859 #1861]: #1201
-#1863 := [unit-resolution #1328 #1862 #1860]: false
-#1864 := [lemma #1863]: #1187
-#1878 := (or #335 #1113 #429 #1189 #1163 #1200)
-#1879 := [unit-resolution #1345 #828 #854]: #1878
-#1881 := [unit-resolution #1879 #1864]: #1880
-#1882 := [unit-resolution #1881 #1335 #1870 #1333 #1332]: #335
-#1883 := [unit-resolution #1107 #1882]: #1077
-#1884 := [unit-resolution #1689 #1333 #1870]: #288
-#1885 := [unit-resolution #1069 #1884]: #1039
-#1886 := [unit-resolution #1272 #1885]: #1033
-#1889 := (or #1036 #429 #1163 #1200 #1074)
-#1887 := (or #242 #1036 #429 #1189 #1163 #1200 #1074)
-#1888 := [unit-resolution #1360 #772 #798 #854]: #1887
-#1890 := [unit-resolution #1888 #1672 #1864]: #1889
-#1891 := [unit-resolution #1890 #1886 #1332 #1333 #1335]: #1074
-#1892 := [unit-resolution #1284 #1891 #1883]: false
-#1894 := [lemma #1892]: #1893
-#1897 := [unit-resolution #1894 #1896]: #429
-#1898 := [unit-resolution #1183 #1897]: #1153
-#1899 := [unit-resolution #1310 #1898]: #1149
-#1900 := [th-lemma #854 #1899 #1870 #828 #1422 #1874]: false
-#1902 := [lemma #1900]: #1901
-#1950 := [unit-resolution #1902 #1874]: #335
-#1951 := [unit-resolution #1107 #1950]: #1077
-#1952 := [unit-resolution #1284 #1951]: #1071
-#1953 := [unit-resolution #1328 #1895]: #1199
-#1876 := (or #1109 #476)
-#1673 := [hypothesis]: #1112
-#1760 := (or #429 #1109)
-#1674 := [unit-resolution #1296 #1673]: #1114
-#1675 := [unit-resolution #1145 #1674]: #383
-#1676 := [unit-resolution #1146 #1675]: #1127
-#1677 := [unit-resolution #1304 #1676]: #1123
-#1687 := [unit-resolution #1686 #1673]: #1111
-#1743 := [unit-resolution #1689 #1333 #1687]: #288
-#1744 := [unit-resolution #1069 #1743]: #1039
-#1745 := [unit-resolution #1272 #1744]: #1033
-#1678 := (or #335 #1343 #429 #382 #1125)
-#1679 := [th-lemma]: #1678
-#1746 := [unit-resolution #1679 #1333 #1675 #828 #1677]: #335
-#1747 := [unit-resolution #1107 #1746]: #1077
-#1748 := [unit-resolution #1284 #1747]: #1071
-#1259 := (or #1000 #995)
-#1260 := [def-axiom]: #1259
-#1684 := [unit-resolution #1260 #1683]: #995
-#1693 := (or #147 #1373 #1343 #1074 #1358 #1523 #429 #973 #998 #1036 #1357 #1125)
-#1694 := [th-lemma]: #1693
-#1749 := [unit-resolution #1694 #1745 #724 #1684 #746 #1440 #772 #1748 #798 #1677 #828 #1333]: #147
-#1750 := [unit-resolution #955 #1749]: #925
-#1751 := [unit-resolution #1238 #1750]: #921
-#1714 := (or #100 #923 #1373 #1371 #1343 #1523 #1074 #1358 #973 #429 #382 #1036 #1357 #998 #1125)
-#1715 := [th-lemma]: #1714
-#1752 := [unit-resolution #1715 #1751 #698 #1440 #724 #1684 #746 #1675 #772 #1748 #798 #1745 #1677 #828 #1333]: #100
-#1753 := [unit-resolution #1236 #1750]: #919
-#1727 := (or #1109 #429 #972)
-#1680 := [unit-resolution #1679 #1675 #1677 #828 #1333]: #335
-#1681 := [unit-resolution #1107 #1680]: #1077
-#1682 := [unit-resolution #1284 #1681]: #1071
-#1690 := [unit-resolution #1689 #1687 #1333]: #288
-#1691 := [unit-resolution #1069 #1690]: #1039
-#1692 := [unit-resolution #1272 #1691]: #1033
-#1695 := [unit-resolution #1694 #1692 #724 #1684 #746 #1440 #772 #1682 #798 #1677 #828 #1333]: #147
-#1696 := [unit-resolution #955 #1695]: #925
-#1697 := [unit-resolution #1236 #1696]: #919
-#1698 := (or #476 #429 #1337 #1163 #382)
-#1699 := [th-lemma]: #1698
-#1700 := [unit-resolution #1699 #1675 #1335 #854 #1333]: #476
-#1701 := [unit-resolution #1221 #1700]: #1191
-#1702 := [unit-resolution #1322 #1701]: #1187
-#1704 := [hypothesis]: #970
-#1301 := (or #1126 #1122)
-#1302 := [def-axiom]: #1301
-#1705 := [unit-resolution #1302 #1676]: #1122
-#1706 := (or #515 #922 #1509 #1510 #1511 #1075 #1350 #1337 #1409 #1037 #1163 #1365 #1364 #972 #999 #1124 #1189)
-#1707 := [th-lemma]: #1706
-#1708 := [unit-resolution #1707 #1705 #1704 #720 #1703 #750 #1459 #776 #1447 #802 #694 #824 #1335 #854 #1702 #880 #1697]: #515
-#1709 := [unit-resolution #1238 #1696]: #921
-#1710 := [unit-resolution #1320 #1701]: #1185
-#1711 := (or #516 #923 #1373 #1371 #1372 #1074 #1358 #1489 #1357 #1036 #1162 #1343 #1523 #973 #998 #1125 #1188)
-#1712 := [th-lemma]: #1711
-#1713 := [unit-resolution #1712 #1692 #1440 #724 #1684 #746 #698 #772 #1682 #798 #1677 #828 #1631 #850 #1710 #876 #1709]: #516
-#1716 := [unit-resolution #1715 #1709 #698 #1440 #724 #1684 #746 #1692 #772 #1682 #798 #1675 #1677 #828 #1333]: #100
-#1717 := [unit-resolution #917 #1716]: #887
-#1718 := [unit-resolution #1226 #1717]: #883
-#1719 := (or #509 #885 #1522 #1523 #1343 #1489 #1357 #923 #1371 #1036 #1162 #998 #1125)
-#1720 := [th-lemma]: #1719
-#1721 := [unit-resolution #1720 #1709 #672 #698 #1684 #746 #1692 #772 #1677 #828 #1631 #850 #1718]: #509
-#1722 := [unit-resolution #1224 #1717]: #881
-#1723 := (or #506 #884 #1528 #1364 #1365 #1337 #1409 #922 #1510 #1037 #1163 #999 #1124)
-#1724 := [th-lemma]: #1723
-#1725 := [unit-resolution #1724 #1697 #668 #694 #1703 #750 #1459 #776 #1705 #824 #1335 #854 #1722]: #506
-#1726 := [unit-resolution #646 #1725 #1721 #1713 #1708]: false
-#1728 := [lemma #1726]: #1727
-#1754 := [unit-resolution #1728 #1333 #1673]: #972
-#1755 := [unit-resolution #1254 #1754]: #974
-#1756 := [unit-resolution #994 #1755]: #194
-#1757 := [unit-resolution #993 #1756]: #963
-#1758 := [unit-resolution #1248 #1757]: #957
-#1759 := [th-lemma #1758 #1753 #720 #694 #1675 #1459 #776 #1447 #802 #1752]: false
-#1761 := [lemma #1759]: #1760
-#1871 := [unit-resolution #1761 #1673]: #429
-#1872 := [unit-resolution #1183 #1871]: #1153
-#1873 := [unit-resolution #1310 #1872]: #1149
-#1875 := [th-lemma #1675 #1874 #854 #1873 #1871]: false
-#1877 := [lemma #1875]: #1876
-#1954 := [unit-resolution #1877 #1874]: #1109
-#1948 := (or #288 #1112 #1200 #1201 #1074)
-#1917 := [unit-resolution #1894 #1332]: #429
-#1918 := [unit-resolution #1183 #1917]: #1153
-#1919 := [unit-resolution #1308 #1918]: #1147
-#1929 := [unit-resolution #1310 #1918]: #1149
-#1930 := [unit-resolution #1238 #1927]: #921
-#1931 := [hypothesis]: #1199
-#1932 := (or #515 #922 #1201 #1074 #1112 #960 #1150)
-#1933 := [unit-resolution #1513 #694 #720 #1468 #776 #798 #824 #850 #880]: #1932
-#1934 := [unit-resolution #1933 #1928 #1931 #1637 #1638 #1919 #1924]: #515
-#1935 := (or #516 #923 #1200 #1113 #961 #1151 #1048)
-#1936 := [unit-resolution #1499 #698 #724 #772 #1447 #802 #828 #854 #876]: #1935
-#1937 := [unit-resolution #1936 #1930 #1870 #1332 #1929 #1497 #1925]: #516
-#1915 := (or #898 #634 #633 #923 #961 #1048 #1151 #922 #960 #1112 #1150)
-#1903 := [hypothesis]: #515
-#1904 := [hypothesis]: #516
-#1905 := [hypothesis]: #899
-#1906 := [unit-resolution #1232 #1905]: #895
-#1907 := (or #509 #1522 #1523 #897 #998 #1489 #1150 #960 #1509 #1112 #1365 #1049 #922 #1510 #1409)
-#1908 := [th-lemma]: #1907
-#1909 := [unit-resolution #1908 #1906 #1632 #694 #1536 #720 #1684 #746 #1468 #776 #1637 #824 #1540 #850 #672]: #509
-#1774 := [hypothesis]: #1149
-#1229 := (or #898 #894)
-#1230 := [def-axiom]: #1229
-#1910 := [unit-resolution #1230 #1905]: #894
-#1911 := (or #506 #1528 #1364 #896 #999 #1337 #1151 #961 #1373 #1113 #1343 #1048 #923 #1371 #1357)
-#1912 := [th-lemma]: #1911
-#1913 := [unit-resolution #1912 #1910 #1642 #698 #1828 #724 #1703 #750 #1545 #772 #1870 #828 #1774 #854 #668]: #506
-#1914 := [unit-resolution #646 #1913 #1909 #1904 #1903]: false
-#1916 := [lemma #1914]: #1915
-#1938 := [unit-resolution #1916 #1937 #1934 #1930 #1925 #1497 #1929 #1928 #1924 #1637 #1919]: #898
-#1939 := [unit-resolution #918 #1938]: #100
-#1940 := [unit-resolution #917 #1939]: #887
-#1941 := [unit-resolution #1224 #1940]: #881
-#1942 := (or #506 #884 #1113 #1151 #1048 #922)
-#1943 := [unit-resolution #1530 #668 #694 #1404 #750 #772 #828 #854]: #1942
-#1944 := [unit-resolution #1943 #1941 #1497 #1870 #1929 #1928]: #506
-#1945 := [unit-resolution #646 #1944 #1937 #1934]: #632
-#1946 := [unit-resolution #1908 #1945 #1928 #694 #1924 #720 #1684 #746 #1468 #776 #1637 #824 #1919 #850 #672]: #897
-#1947 := [th-lemma #1946 #1939 #1742]: false
-#1949 := [lemma #1947]: #1948
-#1955 := [unit-resolution #1949 #1954 #1896 #1953 #1952]: #288
-#1956 := [unit-resolution #1069 #1955]: #1039
-#1957 := [unit-resolution #1272 #1956]: #1033
-#1958 := [unit-resolution #1735 #1954]: #382
-#1959 := (or #1123 #383 #1113)
-#1960 := [th-lemma]: #1959
-#1961 := [unit-resolution #1960 #1958 #1870]: #1123
-#1962 := [unit-resolution #1308 #1898]: #1147
-#1965 := (or #1160 #1112 #1074 #289 #1150)
-#1963 := (or #1160 #1365 #1112 #1074 #1358 #289 #1150)
-#1964 := [th-lemma]: #1963
-#1966 := [unit-resolution #1964 #798 #824]: #1965
-#1967 := [unit-resolution #1966 #1955 #1954 #1962 #1952]: #1160
-#1970 := (or #1162 #1151 #1036 #1125 #147 #1074)
-#1968 := (or #1162 #1151 #1343 #1523 #998 #1036 #1357 #1125 #973 #147 #1373 #1074 #1358)
-#1969 := [th-lemma]: #1968
-#1971 := [unit-resolution #1969 #724 #1684 #746 #1440 #772 #798 #828]: #1970
-#1972 := [unit-resolution #1971 #1967 #1952 #1961 #1899 #1957]: #147
-#1973 := [unit-resolution #955 #1972]: #925
-#1974 := [unit-resolution #1236 #1973]: #919
-#1975 := (or #1161 #1151 #430)
-#1976 := [th-lemma]: #1975
-#1977 := [unit-resolution #1976 #1899 #1897]: #1161
-#1978 := (or #476 #1036 #1112 #194 #1163 #1074)
-#1979 := [unit-resolution #1611 #750 #772 #798 #824 #1404 #854]: #1978
-#1980 := [unit-resolution #1979 #1957 #1874 #1954 #1952 #1977]: #194
-#1981 := [unit-resolution #993 #1980]: #963
-#1982 := [unit-resolution #1248 #1981]: #957
-#1983 := [unit-resolution #1933 #1974 #1953 #1954 #1952 #1962 #1982]: #515
-#1984 := [unit-resolution #1238 #1973]: #921
-#1985 := [unit-resolution #1250 #1981]: #959
-#1849 := (or #923 #516 #1200 #961 #1036 #1163 #1074)
-#1850 := [unit-resolution #1375 #698 #724 #772 #798 #854 #876]: #1849
-#1986 := [unit-resolution #1850 #1985 #1896 #1952 #1977 #1957 #1984]: #516
-#1987 := (or #509 #923 #1036 #1162 #1125)
-#1988 := [unit-resolution #1720 #672 #698 #1684 #746 #1742 #772 #828 #850]: #1987
-#1989 := [unit-resolution #1988 #1984 #1961 #1967 #1957]: #509
-#1990 := [unit-resolution #646 #1989 #1986 #1983]: #631
-#1991 := (or #506 #884 #1112 #922 #1036 #1163 #1074)
-#1992 := [unit-resolution #1603 #668 #694 #1404 #750 #772 #798 #824 #854]: #1991
-#1993 := [unit-resolution #1992 #1990 #1977 #1954 #1952 #1957 #1974]: #884
-#1994 := [unit-resolution #1224 #1993]: #886
-#1995 := [unit-resolution #917 #1994]: #101
-#1996 := [th-lemma #746 #1684 #1957 #1874 #854 #1899 #1870 #828 #1984 #1995 #698 #772 #1972]: false
-#1997 := [lemma #1996]: #476
-#2014 := [unit-resolution #1221 #1997]: #1191
-#2015 := [unit-resolution #1320 #2014]: #1185
-#2034 := [th-lemma #876 #850 #1540 #2015 #802 #2033 #698 #772 #1828 #724 #1545 #1845 #1331]: false
-#2036 := [lemma #2034]: #2035
-#2048 := [unit-resolution #2036 #1497 #2045 #1828 #1331]: #1150
-#2049 := [unit-resolution #1308 #2048 #2047]: false
-#2051 := [lemma #2049]: #2050
-#2082 := [unit-resolution #2051 #1405 #1331]: #961
-#2083 := [unit-resolution #1250 #1923 #2082]: false
-#2085 := [lemma #2083]: #2084
-#2089 := [unit-resolution #2085 #1331]: #288
-#2090 := [unit-resolution #1069 #2089]: #1039
-#2091 := [unit-resolution #1272 #2090]: #1033
-#2065 := [hypothesis]: #935
-#2066 := [unit-resolution #1244 #2065]: #936
-#2067 := [unit-resolution #956 #2066]: #147
-#2068 := [th-lemma #2065 #2033 #2067]: false
-#2069 := [lemma #2068]: #933
-#2100 := (or #429 #516)
-#2063 := (or #429 #1086 #516)
-#2052 := [unit-resolution #1761 #1333]: #1109
-#2053 := [unit-resolution #1735 #2052]: #382
-#2054 := [hypothesis]: #1084
-#2055 := (or #1200 #516 #429)
-#2056 := [unit-resolution #1383 #1864]: #2055
-#2057 := [unit-resolution #2056 #1333 #1331]: #1200
-#2060 := (or #1086 #383 #1113 #1188 #1162 #1198)
-#2058 := (or #1086 #383 #1113 #1343 #1188 #1489 #1162 #1198 #1075)
-#2059 := [th-lemma]: #2058
-#2061 := [unit-resolution #2059 #1447 #828 #850]: #2060
-#2062 := [unit-resolution #2061 #1631 #2057 #2015 #1870 #2054 #2053]: false
-#2064 := [lemma #2062]: #2063
-#2086 := [unit-resolution #2064 #1333 #1331]: #1086
-#2087 := [unit-resolution #1290 #2086]: #1088
-#2088 := [unit-resolution #1108 #2087]: #335
-#2080 := (or #1109 #516)
-#2070 := [unit-resolution #1308 #1872]: #1147
-#2020 := (or #194 #1150 #516 #1125 #1151 #1124)
-#1762 := [hypothesis]: #1122
-#1775 := [hypothesis]: #1123
-#1803 := (or #194 #1151 #1150 #1125 #147 #1124)
-#1764 := [unit-resolution #956 #1763]: #937
-#1765 := [unit-resolution #1244 #1764]: #933
-#1766 := (or #509 #885 #1522 #1364 #1365 #1489 #999 #1124 #1371 #1037 #1409 #935 #1150 #972 #1509 #1075 #1350)
-#1767 := [th-lemma]: #1766
-#1768 := [unit-resolution #1767 #1620 #1765 #698 #672 #720 #1703 #750 #1459 #776 #1447 #802 #1762 #824 #1540 #850 #1742]: #509
-#1769 := (or #100 #1371 #935 #194 #147)
-#1770 := [th-lemma]: #1769
-#1771 := [unit-resolution #1770 #1535 #1765 #698 #1763]: #100
-#1772 := [unit-resolution #917 #1771]: #887
-#1773 := [unit-resolution #1224 #1772]: #881
-#1776 := (or #335 #194 #1364 #1037 #1409 #999)
-#1777 := [th-lemma]: #1776
-#1778 := [unit-resolution #1777 #1535 #750 #1459 #776 #1703]: #335
-#1779 := [unit-resolution #1107 #1778]: #1077
-#1780 := [unit-resolution #1284 #1779]: #1071
-#1241 := (or #936 #932)
-#1242 := [def-axiom]: #1241
-#1781 := [unit-resolution #1242 #1764]: #932
-#1782 := (or #288 #1364 #999 #973 #147 #1373 #194)
-#1783 := [th-lemma]: #1782
-#1784 := [unit-resolution #1783 #1535 #1440 #724 #1703 #750 #1763]: #288
-#1785 := [unit-resolution #1069 #1784]: #1039
-#1786 := [unit-resolution #1272 #1785]: #1033
-#1787 := (or #506 #884 #1528 #1523 #1343 #1337 #998 #1125 #1510 #1036 #1357 #934 #1151 #973 #1373 #1074 #1358)
-#1788 := [th-lemma]: #1787
-#1789 := [unit-resolution #1788 #1786 #1781 #694 #1440 #724 #1684 #746 #668 #772 #1780 #798 #1775 #828 #1774 #854 #1773]: #506
-#1790 := (or #476 #1337 #1343 #1523 #1036 #1357 #998 #1125 #973 #147 #1373 #1074 #1358 #1151 #194)
-#1791 := [th-lemma]: #1790
-#1792 := [unit-resolution #1791 #1535 #1440 #724 #1684 #746 #1786 #772 #1780 #798 #1775 #828 #1774 #854 #1763]: #476
-#1793 := [unit-resolution #1221 #1792]: #1191
-#1794 := [unit-resolution #1320 #1793]: #1185
-#1795 := (or #516 #1372 #1489 #1409 #1037 #1188 #1371 #935 #972 #1509 #1075 #1350 #1150)
-#1796 := [th-lemma]: #1795
-#1797 := [unit-resolution #1796 #1620 #698 #720 #1459 #776 #1447 #802 #1540 #850 #1794 #876 #1765]: #516
-#1798 := [unit-resolution #1322 #1793]: #1187
-#1799 := (or #515 #1511 #1337 #1357 #1036 #1189 #1510 #934 #973 #1373 #1074 #1358 #1151)
-#1800 := [th-lemma]: #1799
-#1801 := [unit-resolution #1800 #1786 #1440 #724 #694 #772 #1780 #798 #1774 #854 #1798 #880 #1781]: #515
-#1802 := [unit-resolution #646 #1801 #1797 #1789 #1768]: false
-#1804 := [lemma #1802]: #1803
-#2011 := [unit-resolution #1804 #1535 #1540 #1775 #1774 #1762]: #147
-#2012 := [unit-resolution #955 #2011]: #925
-#2013 := [unit-resolution #1238 #2012]: #921
-#2016 := (or #516 #1188 #935 #972 #1150)
-#2017 := [unit-resolution #1796 #698 #720 #1459 #776 #1447 #802 #850 #876]: #2016
-#2018 := [unit-resolution #2017 #1620 #2015 #1540 #1331]: #935
-#2019 := [th-lemma #2018 #2013 #2011]: false
-#2021 := [lemma #2019]: #2020
-#2071 := [unit-resolution #2021 #2070 #1331 #1677 #1873 #1705]: #194
-#2072 := [unit-resolution #993 #2071]: #963
-#2073 := [unit-resolution #2010 #1675]: #288
-#2074 := [unit-resolution #1069 #2073]: #1039
-#2075 := [unit-resolution #1272 #2074]: #1033
-#2076 := (or #516 #1036 #1188 #935 #1150 #960 #1087)
-#1823 := (or #516 #1372 #1489 #1357 #1036 #1188 #1371 #935 #1509 #1350 #1150 #960 #1523 #998 #1087)
-#1824 := [th-lemma]: #1823
-#2077 := [unit-resolution #1824 #720 #1684 #746 #698 #772 #802 #850 #876]: #2076
-#2078 := [unit-resolution #2077 #2075 #2015 #2045 #2069 #1331 #2070]: #960
-#2079 := [unit-resolution #1248 #2078 #2072]: false
-#2081 := [lemma #2079]: #2080
-#2092 := [unit-resolution #2081 #1331]: #1109
-#2093 := [unit-resolution #1735 #2092]: #382
-#2094 := [unit-resolution #1960 #2093 #1870]: #1123
-#2095 := (or #516 #923 #1074 #1036 #1162 #1125 #1188)
-#2096 := [unit-resolution #1712 #1440 #724 #1684 #746 #698 #772 #798 #828 #850 #876]: #2095
-#2097 := [unit-resolution #2096 #1631 #2015 #2094 #1331 #2091 #2033]: #1074
-#2098 := [unit-resolution #1284 #2097]: #1076
-#2099 := [unit-resolution #1107 #2098 #2088]: false
-#2101 := [lemma #2099]: #2100
-#2102 := [unit-resolution #2101 #1331]: #429
-#2103 := [unit-resolution #1183 #2102]: #1153
-#2104 := [unit-resolution #1308 #2103]: #1147
-#2105 := [unit-resolution #2077 #2104 #2015 #2045 #2069 #1331 #2091]: #960
-#2106 := [unit-resolution #1248 #2105]: #962
-#2107 := [unit-resolution #2017 #2104 #2015 #2069 #1331]: #972
-#2108 := [unit-resolution #1254 #2107]: #974
-#2109 := [unit-resolution #994 #2108]: #194
-#2110 := [unit-resolution #993 #2109 #2106]: false
-#2111 := [lemma #2110]: #516
-#2127 := (or #1199 #1189 #477)
-#2128 := [th-lemma]: #2127
-#2129 := [unit-resolution #2128 #1864 #1997]: #1199
-#2125 := (or #335 #288)
-#1806 := [unit-resolution #1108 #1422]: #1089
-#1829 := [unit-resolution #1290 #1806]: #1084
-#2117 := (or #515 #1511 #1337 #1151 #1189 #1358 #922 #1510 #1409 #960 #1509 #1049 #1086)
-#2118 := [th-lemma]: #2117
-#2119 := [unit-resolution #2118 #1829 #1924 #720 #1468 #776 #694 #798 #2116 #854 #1864 #880 #1928]: #515
-#2120 := (or #101 #922 #1510 #1409 #960 #1509 #1049 #335 #288)
-#2121 := [th-lemma]: #2120
-#2122 := [unit-resolution #2121 #1422 #694 #1924 #720 #1405 #1468 #776 #1928]: #101
-#2123 := [unit-resolution #918 #2122]: #899
-#2124 := [unit-resolution #1916 #2123 #2119 #2111 #1497 #1925 #2033 #2116 #1928 #1924 #2115 #2112]: false
-#2126 := [lemma #2124]: #2125
-#2130 := [unit-resolution #2126 #1405]: #335
-#2131 := [unit-resolution #1107 #2130]: #1077
-#2132 := [unit-resolution #1284 #2131]: #1071
-#2133 := [unit-resolution #1933 #2132 #2129 #2115 #1928 #2112 #1924]: #515
-#2134 := [unit-resolution #1916 #2133 #2111 #1497 #1925 #2033 #2116 #1928 #1924 #2115 #2112]: #898
-#2135 := [unit-resolution #918 #2134]: #100
-#2136 := [unit-resolution #917 #2135]: #887
-#2137 := [unit-resolution #1224 #2136]: #881
-#2138 := [unit-resolution #1943 #2137 #1497 #1870 #2116 #1928]: #506
-#2139 := [unit-resolution #646 #2138 #2111 #2133]: #632
-#2140 := [unit-resolution #1908 #2139 #1928 #694 #1924 #720 #1684 #746 #1468 #776 #2115 #824 #2112 #850 #672]: #897
-#2141 := [th-lemma #2140 #2135 #1742]: false
-#2142 := [lemma #2141]: #288
-#2143 := [unit-resolution #1069 #2142]: #1039
-#2144 := [unit-resolution #1272 #2143]: #1033
-#2145 := [hypothesis]: #1150
-#2146 := [unit-resolution #1308 #2145]: #1152
-#2147 := [unit-resolution #1183 #2146]: #430
-#2148 := [unit-resolution #1184 #2147]: #1165
-#2149 := [unit-resolution #1314 #2148]: #1160
-#2150 := [unit-resolution #1761 #2147]: #1109
-#2151 := [unit-resolution #1735 #2150]: #382
-#2152 := [unit-resolution #1960 #2151 #1870]: #1123
-#2153 := [unit-resolution #1988 #2152 #2149 #2033 #2144]: #509
-#2154 := (or #1149 #1147)
-#2155 := [th-lemma]: #2154
-#2156 := [unit-resolution #2155 #2145]: #1149
-#2157 := [unit-resolution #1894 #2147]: #1200
-#2158 := [unit-resolution #2061 #2149 #2015 #1870 #2157 #2151]: #1086
-#2159 := [unit-resolution #1290 #2158]: #1088
-#2160 := [unit-resolution #1108 #2159]: #335
-#2161 := [unit-resolution #1107 #2160]: #1077
-#2162 := [unit-resolution #1284 #2161]: #1071
-#2163 := [unit-resolution #1971 #2162 #2149 #2152 #2156 #2144]: #147
-#2164 := [unit-resolution #955 #2163]: #925
-#2165 := [unit-resolution #1236 #2164]: #919
-#2166 := [unit-resolution #1316 #2148]: #1161
-#2167 := (or #100 #923 #1371 #1357 #1523 #998 #1036 #383 #429 #1343 #1113 #973 #1373 #1074 #1358)
-#2168 := [th-lemma]: #2167
-#2169 := [unit-resolution #2168 #2162 #698 #1440 #724 #1684 #746 #2144 #772 #2033 #798 #2151 #1870 #828 #2147]: #100
-#2170 := [unit-resolution #917 #2169]: #887
-#2171 := [unit-resolution #1224 #2170]: #881
-#2172 := [unit-resolution #1992 #2171 #2166 #2150 #2162 #2144 #2165]: #506
-#2173 := (or #195 #1357 #1523 #998 #1036 #383 #429 #1343 #1113)
-#2174 := [th-lemma]: #2173
-#2175 := [unit-resolution #2174 #2151 #746 #2144 #772 #1684 #1870 #828 #2147]: #195
-#2176 := [unit-resolution #994 #2175]: #975
-#2177 := [unit-resolution #1254 #2176]: #970
-#2178 := (or #515 #922 #1074 #1036 #972 #1163 #1112)
-#2179 := [unit-resolution #1622 #694 #720 #1404 #750 #772 #1864 #798 #824 #854 #880]: #2178
-#2180 := [unit-resolution #2179 #2177 #2150 #2162 #2166 #2144 #2165]: #515
-#2181 := [unit-resolution #646 #2180 #2172 #2111 #2153]: false
-#2182 := [lemma #2181]: #1147
-#1805 := [unit-resolution #1302 #1729]: #1122
-#2231 := (or #194 #382)
-#2183 := (or #1150 #429 #1163)
-#2184 := [th-lemma]: #2183
-#2185 := [unit-resolution #2184 #1333 #2182]: #1163
-#2186 := [unit-resolution #1316 #2185 #1334]: false
-#2187 := [lemma #2186]: #429
-#2196 := [unit-resolution #1183 #2187]: #1153
-#2197 := [unit-resolution #1310 #2196]: #1149
-#1817 := [unit-resolution #1304 #1729]: #1123
-#2217 := [unit-resolution #1804 #1535 #2182 #1817 #2197 #1805]: #147
-#2218 := [unit-resolution #955 #2217]: #925
-#2219 := [unit-resolution #1236 #2218]: #919
-#2210 := [unit-resolution #1976 #2197 #2187]: #1161
-#2220 := (or #509 #1124 #935 #1150 #972)
-#2221 := [unit-resolution #1767 #698 #672 #720 #1703 #750 #1459 #776 #1447 #802 #824 #850 #1742]: #2220
-#2222 := [unit-resolution #2221 #1620 #2069 #1805 #2182]: #509
-#2223 := (or #515 #922 #1163 #972 #1124)
-#2224 := [unit-resolution #1707 #720 #1703 #750 #1459 #776 #1447 #802 #694 #824 #854 #880 #1864]: #2223
-#2225 := [unit-resolution #2224 #2219 #1805 #1620 #2210]: #515
-#2226 := [unit-resolution #646 #2225 #2111 #2222]: #631
-#2211 := (or #506 #884 #922 #1163 #1124)
-#2212 := [unit-resolution #1724 #668 #694 #1703 #750 #1459 #776 #824 #854]: #2211
-#2227 := [unit-resolution #2212 #2226 #1805 #2210 #2219]: #884
-#2228 := [unit-resolution #1224 #2227]: #886
-#2229 := [unit-resolution #917 #2228]: #101
-#2230 := [th-lemma #1620 #720 #1459 #776 #1447 #802 #2033 #2229 #698 #1428 #2217]: false
-#2232 := [lemma #2230]: #2231
-#2242 := [unit-resolution #2232 #1428]: #194
-#2243 := [unit-resolution #993 #2242]: #963
-#2244 := [unit-resolution #1248 #2243]: #957
-#2193 := (or #509 #1124 #1036 #935 #1150 #960 #1087)
-#1814 := (or #509 #885 #1522 #1523 #1365 #1489 #998 #1124 #1371 #1036 #1357 #935 #1150 #1509 #1350 #960 #1087)
-#1815 := [th-lemma]: #1814
-#2194 := [unit-resolution #1815 #698 #720 #1684 #746 #672 #772 #802 #824 #850 #1742]: #2193
-#2245 := [unit-resolution #2194 #2244 #2069 #2144 #2045 #1805 #2182]: #509
-#2205 := (or #100 #935 #1036 #382 #960 #1087)
-#1834 := (or #100 #1371 #935 #1523 #1036 #1357 #998 #1509 #382 #1350 #960 #1087)
-#1835 := [th-lemma]: #1834
-#2206 := [unit-resolution #1835 #698 #720 #1684 #746 #772 #802]: #2205
-#2246 := [unit-resolution #2206 #2244 #2045 #2069 #2144 #1428]: #100
-#2247 := [unit-resolution #917 #2246]: #887
-#2248 := [unit-resolution #1224 #2247]: #881
-#2215 := (or #335 #382)
-#2188 := (or #335 #194)
-#2189 := [unit-resolution #1777 #750 #1459 #776 #1703]: #2188
-#2190 := [unit-resolution #2189 #1422]: #194
-#2191 := [unit-resolution #993 #2190]: #963
-#2192 := [unit-resolution #1248 #2191]: #957
-#2195 := [unit-resolution #2194 #2192 #2069 #2144 #2045 #1805 #2182]: #509
-#2198 := [unit-resolution #1250 #2191]: #959
-#1840 := (or #335 #934 #1151 #961 #935 #960 #1150 #382)
-#1807 := [unit-resolution #1292 #1806]: #1085
-#1808 := [hypothesis]: #933
-#1809 := (or #288 #382 #1350 #335 #1087)
-#1810 := [th-lemma]: #1809
-#1811 := [unit-resolution #1810 #1422 #1807 #802 #1428]: #288
-#1812 := [unit-resolution #1069 #1811]: #1039
-#1813 := [unit-resolution #1272 #1812]: #1033
-#1816 := [unit-resolution #1815 #1813 #1808 #698 #1536 #720 #1684 #746 #672 #772 #1807 #802 #1805 #824 #1540 #850 #1742]: #509
-#1818 := (or #476 #1337 #1343 #1125 #1151 #335 #382)
-#1819 := [th-lemma]: #1818
-#1820 := [unit-resolution #1819 #1422 #1817 #828 #1774 #854 #1428]: #476
-#1821 := [unit-resolution #1221 #1820]: #1191
-#1822 := [unit-resolution #1320 #1821]: #1185
-#1825 := [unit-resolution #1824 #1813 #1536 #720 #1684 #746 #698 #772 #1807 #802 #1540 #850 #1822 #876 #1808]: #516
-#1826 := [hypothesis]: #932
-#1827 := [unit-resolution #1322 #1821]: #1187
-#1830 := (or #515 #1511 #1337 #1409 #1037 #1189 #1510 #934 #1373 #1358 #1151 #961 #1364 #999 #1086)
-#1831 := [th-lemma]: #1830
-#1832 := [unit-resolution #1831 #1829 #1828 #724 #1703 #750 #1459 #776 #694 #798 #1774 #854 #1827 #880 #1826]: #515
-#1833 := [unit-resolution #646 #1832 #1825 #1816]: #631
-#1836 := [unit-resolution #1835 #1813 #698 #1536 #720 #1684 #746 #1808 #772 #1807 #802 #1428]: #100
-#1837 := [unit-resolution #917 #1836]: #887
-#1838 := [unit-resolution #1224 #1837]: #881
-#1839 := [th-lemma #1838 #668 #750 #828 #854 #1703 #1817 #694 #1459 #776 #1826 #1774 #724 #798 #1828 #1829 #1833]: false
-#1841 := [lemma #1839]: #1840
-#2199 := [unit-resolution #1841 #2198 #2069 #1422 #2197 #2192 #2182 #1428]: #934
-#2200 := [unit-resolution #1242 #2199]: #936
-#2201 := [unit-resolution #956 #2200]: #147
-#2202 := [unit-resolution #955 #2201]: #925
-#2203 := [unit-resolution #1236 #2202]: #919
-#2204 := [unit-resolution #2118 #2203 #1829 #720 #1468 #776 #694 #798 #2197 #854 #1864 #880 #2192]: #515
-#2207 := [unit-resolution #2206 #2192 #2045 #2069 #2144 #1428]: #100
-#2208 := [unit-resolution #917 #2207]: #887
-#2209 := [unit-resolution #1224 #2208]: #881
-#2213 := [unit-resolution #2212 #2203 #1805 #2210 #2209]: #506
-#2214 := [unit-resolution #646 #2213 #2204 #2111 #2195]: false
-#2216 := [lemma #2214]: #2215
-#2249 := [unit-resolution #2216 #1428]: #335
-#2250 := [unit-resolution #1107 #2249]: #1077
-#2251 := [unit-resolution #1284 #2250]: #1071
-#2252 := (or #1084 #1074 #1357 #1523 #998 #1036 #195)
-#2253 := [th-lemma]: #2252
-#2254 := [unit-resolution #2253 #2251 #746 #2144 #772 #1684 #2242]: #1084
-#2255 := [unit-resolution #1250 #2243]: #959
-#2240 := (or #934 #632 #884 #1074 #1125 #961 #1086)
-#2233 := (or #515 #934 #1151 #961 #1086)
-#2234 := [unit-resolution #1831 #1864 #724 #1703 #750 #1459 #776 #694 #798 #854 #880]: #2233
-#2235 := [unit-resolution #2234 #1826 #2197 #1828 #2054]: #515
-#2236 := (or #506 #884 #1125 #1036 #934 #1151 #1074)
-#2237 := [unit-resolution #1788 #694 #1440 #724 #1684 #746 #668 #772 #798 #828 #854]: #2236
-#2238 := [unit-resolution #2237 #1826 #1636 #1638 #1775 #2197 #2144]: #506
-#2239 := [unit-resolution #646 #2238 #2235 #2111 #1628]: false
-#2241 := [lemma #2239]: #2240
-#2256 := [unit-resolution #2241 #2245 #2248 #2251 #1817 #2255 #2254]: #934
-#2257 := [unit-resolution #1242 #2256]: #936
-#2258 := [unit-resolution #956 #2257]: #147
-#2259 := [unit-resolution #955 #2258]: #925
-#2260 := [unit-resolution #1236 #2259]: #919
-#2261 := [unit-resolution #2212 #2260 #1805 #2210 #2248]: #506
-#2262 := [unit-resolution #2118 #2260 #2254 #720 #1468 #776 #694 #798 #2197 #854 #1864 #880 #2244]: #515
-#2263 := [unit-resolution #646 #2262 #2261 #2111 #2245]: false
-#2264 := [lemma #2263]: #382
-#2265 := [unit-resolution #1145 #2264]: #1115
-#2266 := [unit-resolution #1296 #2265]: #1109
-#2267 := [unit-resolution #2189 #1535]: #335
-#2268 := [unit-resolution #1107 #2267]: #1077
-#2269 := [unit-resolution #1284 #2268]: #1071
-#2270 := [unit-resolution #1966 #2269 #2142 #2266 #2182]: #1160
-#2271 := (or #1008 #998 #1036 #1357 #1074 #1358 #383)
-#2272 := [th-lemma]: #2271
-#2273 := [unit-resolution #2272 #2269 #2144 #772 #1684 #798 #2264]: #1008
-#2274 := (or #509 #1010 #1113 #923 #1162)
-#2275 := [unit-resolution #1608 #672 #698 #1742 #746 #1459 #776 #1447 #802 #828 #850]: #2274
-#2276 := [unit-resolution #2275 #2273 #1870 #2270 #2033]: #509
-#2277 := [unit-resolution #1960 #2264 #1870]: #1123
-#2278 := [unit-resolution #1971 #2270 #2269 #2277 #2197 #2144]: #147
-#2279 := [unit-resolution #955 #2278]: #925
-#2280 := [unit-resolution #1236 #2279]: #919
-#2281 := (or #1010 #999 #923 #100 #1371 #961 #1373)
-#2282 := [th-lemma]: #2281
-#2283 := [unit-resolution #2282 #2273 #698 #1584 #724 #1703 #2033]: #100
-#2284 := [unit-resolution #917 #2283]: #887
-#2285 := [unit-resolution #1224 #2284]: #881
-#2286 := [unit-resolution #1992 #2285 #2210 #2266 #2269 #2144 #2280]: #506
-#2287 := [unit-resolution #2179 #2280 #2266 #1620 #2210 #2144 #2269]: #515
-#2288 := [unit-resolution #646 #2287 #2286 #2111 #2276]: false
-#2289 := [lemma #2288]: #194
-#2305 := [unit-resolution #2253 #2302 #746 #2144 #772 #1684 #2289]: #1074
-#2306 := [unit-resolution #1284 #2305]: #1076
-#2307 := [unit-resolution #1107 #2306 #2304]: false
-#2308 := [lemma #2307]: #1084
-#2300 := (or #1086 #515)
-#2290 := [hypothesis]: #633
-#2291 := [unit-resolution #993 #2289]: #963
-#2292 := [unit-resolution #1250 #2291]: #959
-#2293 := [unit-resolution #2234 #2054 #2197 #2292 #2290]: #934
-#2294 := [unit-resolution #1242 #2293]: #936
-#2295 := [unit-resolution #1248 #2291]: #957
-#2296 := [unit-resolution #2118 #2054 #2290 #720 #1468 #776 #694 #798 #2197 #854 #1864 #880 #2295]: #922
-#2297 := [unit-resolution #1236 #2296]: #924
-#2298 := [unit-resolution #955 #2297]: #148
-#2299 := [unit-resolution #956 #2298 #2294]: false
-#2301 := [lemma #2299]: #2300
-#1848 := [unit-resolution #2301 #2308]: #515
-#1851 := [hypothesis]: #632
-#1852 := (or #897 #1522 #509 #1523 #998 #1365 #1489 #1150 #1509 #1350 #633 #1372 #1188 #960 #1087 #1112)
-#1853 := [th-lemma]: #1852
-#1846 := [unit-resolution #1853 #1851 #2295 #720 #1684 #746 #2045 #802 #2266 #824 #2182 #850 #2015 #876 #672 #1848]: #897
-#1847 := [unit-resolution #1232 #1846]: #898
-#1854 := [unit-resolution #918 #1847]: #100
-#1855 := (or #509 #1124)
-#1856 := [unit-resolution #2194 #2069 #2144 #2045 #2295 #2182]: #1855
-#2309 := [unit-resolution #1856 #1851]: #1124
-#2310 := [th-lemma #1848 #876 #850 #2182 #2015 #2309 #2266 #1854]: false
-#2311 := [lemma #2310]: #509
-#2312 := (or #631 #632)
-#2313 := [unit-resolution #646 #2111 #1848]: #2312
-#2314 := [unit-resolution #2313 #2311]: #631
-#2315 := (or #884 #633 #1372 #1188 #1125 #1528 #506 #1364 #999 #1343 #1373 #1358 #961 #1086)
-#2316 := [th-lemma]: #2315
-#2317 := [unit-resolution #2316 #668 #2292 #724 #1703 #750 #2308 #798 #2277 #828 #2015 #876 #2314 #1848]: #884
-#2318 := [unit-resolution #1224 #2317]: #886
-#2319 := (or #896 #1528 #506 #1364 #999 #1343 #1337 #1151 #1373 #1358 #634 #1511 #1189 #961 #1086 #1113)
-#2320 := [th-lemma]: #2319
-#2321 := [unit-resolution #2320 #668 #2292 #724 #1703 #750 #2308 #798 #1870 #828 #2197 #854 #1864 #880 #2314 #2111]: #896
-#2322 := [unit-resolution #1230 #2321]: #898
-#2323 := [unit-resolution #918 #2322]: #100
-[unit-resolution #917 #2323 #2318]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_linarith_17 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,8 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T1)
-:extrafuns (
- (uf_1 Real)
- )
-:assumption (not (flet ($x1 (< (+ uf_1 uf_1) (+ (* 2.0 uf_1) 1.0))) (or $x1 (or false $x1))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_linarith_17.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,52 +0,0 @@
-#2 := false
-#8 := 1::real
-decl uf_1 :: real
-#4 := uf_1
-#6 := 2::real
-#7 := (* 2::real uf_1)
-#9 := (+ #7 1::real)
-#5 := (+ uf_1 uf_1)
-#10 := (< #5 #9)
-#11 := (or false #10)
-#12 := (or #10 #11)
-#13 := (not #12)
-#64 := (iff #13 false)
-#32 := (+ 1::real #7)
-#35 := (< #7 #32)
-#52 := (not #35)
-#62 := (iff #52 false)
-#1 := true
-#57 := (not true)
-#60 := (iff #57 false)
-#61 := [rewrite]: #60
-#58 := (iff #52 #57)
-#55 := (iff #35 true)
-#56 := [rewrite]: #55
-#59 := [monotonicity #56]: #58
-#63 := [trans #59 #61]: #62
-#53 := (iff #13 #52)
-#50 := (iff #12 #35)
-#45 := (or #35 #35)
-#48 := (iff #45 #35)
-#49 := [rewrite]: #48
-#46 := (iff #12 #45)
-#43 := (iff #11 #35)
-#38 := (or false #35)
-#41 := (iff #38 #35)
-#42 := [rewrite]: #41
-#39 := (iff #11 #38)
-#36 := (iff #10 #35)
-#33 := (= #9 #32)
-#34 := [rewrite]: #33
-#30 := (= #5 #7)
-#31 := [rewrite]: #30
-#37 := [monotonicity #31 #34]: #36
-#40 := [monotonicity #37]: #39
-#44 := [trans #40 #42]: #43
-#47 := [monotonicity #37 #44]: #46
-#51 := [trans #47 #49]: #50
-#54 := [monotonicity #51]: #53
-#65 := [trans #54 #63]: #64
-#29 := [asserted]: #13
-[mp #29 #65]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_linarith_18 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,7 +0,0 @@
-(benchmark Isabelle
-:extrafuns (
- (uf_1 Int)
- )
-:assumption (not (<= (+ uf_1 1) (+ uf_1 (+ (* 2 (mod uf_1 2)) 1))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_linarith_18.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,52 +0,0 @@
-#2 := false
-#55 := 0::int
-#7 := 2::int
-decl uf_1 :: int
-#4 := uf_1
-#8 := (mod uf_1 2::int)
-#9 := (* 2::int #8)
-#56 := (>= #9 0::int)
-#51 := (not #56)
-#5 := 1::int
-#10 := (+ #9 1::int)
-#11 := (+ uf_1 #10)
-#6 := (+ uf_1 1::int)
-#12 := (<= #6 #11)
-#13 := (not #12)
-#58 := (iff #13 #51)
-#39 := (+ uf_1 #9)
-#40 := (+ 1::int #39)
-#30 := (+ 1::int uf_1)
-#45 := (<= #30 #40)
-#48 := (not #45)
-#52 := (iff #48 #51)
-#53 := (iff #45 #56)
-#54 := [rewrite]: #53
-#57 := [monotonicity #54]: #52
-#49 := (iff #13 #48)
-#46 := (iff #12 #45)
-#43 := (= #11 #40)
-#33 := (+ 1::int #9)
-#36 := (+ uf_1 #33)
-#41 := (= #36 #40)
-#42 := [rewrite]: #41
-#37 := (= #11 #36)
-#34 := (= #10 #33)
-#35 := [rewrite]: #34
-#38 := [monotonicity #35]: #37
-#44 := [trans #38 #42]: #43
-#31 := (= #6 #30)
-#32 := [rewrite]: #31
-#47 := [monotonicity #32 #44]: #46
-#50 := [monotonicity #47]: #49
-#59 := [trans #50 #57]: #58
-#29 := [asserted]: #13
-#60 := [mp #29 #59]: #51
-#107 := (>= #8 0::int)
-#1 := true
-#28 := [true-axiom]: true
-#135 := (or false #107)
-#136 := [th-lemma]: #135
-#137 := [unit-resolution #136 #28]: #107
-[th-lemma #137 #60]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_linarith_19 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,7 +0,0 @@
-(benchmark Isabelle
-:extrafuns (
- (uf_1 Int)
- )
-:assumption (not (< (+ uf_1 (+ (mod uf_1 2) (mod uf_1 2))) (+ uf_1 3)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_linarith_19.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,50 +0,0 @@
-#2 := false
-#5 := 2::int
-decl uf_1 :: int
-#4 := uf_1
-#6 := (mod uf_1 2::int)
-#122 := (>= #6 2::int)
-#123 := (not #122)
-#1 := true
-#27 := [true-axiom]: true
-#133 := (or false #123)
-#134 := [th-lemma]: #133
-#135 := [unit-resolution #134 #27]: #123
-#9 := 3::int
-#29 := (* 2::int #6)
-#48 := (>= #29 3::int)
-#10 := (+ uf_1 3::int)
-#7 := (+ #6 #6)
-#8 := (+ uf_1 #7)
-#11 := (< #8 #10)
-#12 := (not #11)
-#55 := (iff #12 #48)
-#35 := (+ 3::int uf_1)
-#32 := (+ uf_1 #29)
-#38 := (< #32 #35)
-#41 := (not #38)
-#53 := (iff #41 #48)
-#46 := (not #48)
-#45 := (not #46)
-#51 := (iff #45 #48)
-#52 := [rewrite]: #51
-#49 := (iff #41 #45)
-#47 := (iff #38 #46)
-#44 := [rewrite]: #47
-#50 := [monotonicity #44]: #49
-#54 := [trans #50 #52]: #53
-#42 := (iff #12 #41)
-#39 := (iff #11 #38)
-#36 := (= #10 #35)
-#37 := [rewrite]: #36
-#33 := (= #8 #32)
-#30 := (= #7 #29)
-#31 := [rewrite]: #30
-#34 := [monotonicity #31]: #33
-#40 := [monotonicity #34 #37]: #39
-#43 := [monotonicity #40]: #42
-#56 := [trans #43 #54]: #55
-#28 := [asserted]: #12
-#57 := [mp #28 #56]: #48
-[th-lemma #57 #135]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_linarith_20 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,9 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T1)
-:extrafuns (
- (uf_1 Real)
- )
-:assumption (not (= uf_1 0.0))
-:assumption (not (not (= (+ uf_1 uf_1) (* (ite (or (< 1.0 (ite (< uf_1 0.0) (~ uf_1) uf_1)) (not (< 1.0 (ite (< uf_1 0.0) (~ uf_1) uf_1)))) 4.0 2.0) uf_1))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_linarith_20.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,118 +0,0 @@
-#2 := false
-#5 := 0::real
-decl uf_1 :: real
-#4 := uf_1
-#94 := (<= uf_1 0::real)
-#17 := 2::real
-#40 := (* 2::real uf_1)
-#102 := (<= #40 0::real)
-#103 := (>= #40 0::real)
-#105 := (not #103)
-#104 := (not #102)
-#106 := (or #104 #105)
-#107 := (not #106)
-#88 := (= #40 0::real)
-#108 := (iff #88 #107)
-#109 := [rewrite]: #108
-#16 := 4::real
-#11 := (- uf_1)
-#10 := (< uf_1 0::real)
-#12 := (ite #10 #11 uf_1)
-#9 := 1::real
-#13 := (< 1::real #12)
-#14 := (not #13)
-#15 := (or #13 #14)
-#18 := (ite #15 4::real 2::real)
-#19 := (* #18 uf_1)
-#8 := (+ uf_1 uf_1)
-#20 := (= #8 #19)
-#21 := (not #20)
-#22 := (not #21)
-#89 := (iff #22 #88)
-#70 := (* 4::real uf_1)
-#73 := (= #40 #70)
-#86 := (iff #73 #88)
-#87 := [rewrite]: #86
-#84 := (iff #22 #73)
-#76 := (not #73)
-#79 := (not #76)
-#82 := (iff #79 #73)
-#83 := [rewrite]: #82
-#80 := (iff #22 #79)
-#77 := (iff #21 #76)
-#74 := (iff #20 #73)
-#71 := (= #19 #70)
-#68 := (= #18 4::real)
-#1 := true
-#63 := (ite true 4::real 2::real)
-#66 := (= #63 4::real)
-#67 := [rewrite]: #66
-#64 := (= #18 #63)
-#61 := (iff #15 true)
-#43 := -1::real
-#44 := (* -1::real uf_1)
-#47 := (ite #10 #44 uf_1)
-#50 := (< 1::real #47)
-#53 := (not #50)
-#56 := (or #50 #53)
-#59 := (iff #56 true)
-#60 := [rewrite]: #59
-#57 := (iff #15 #56)
-#54 := (iff #14 #53)
-#51 := (iff #13 #50)
-#48 := (= #12 #47)
-#45 := (= #11 #44)
-#46 := [rewrite]: #45
-#49 := [monotonicity #46]: #48
-#52 := [monotonicity #49]: #51
-#55 := [monotonicity #52]: #54
-#58 := [monotonicity #52 #55]: #57
-#62 := [trans #58 #60]: #61
-#65 := [monotonicity #62]: #64
-#69 := [trans #65 #67]: #68
-#72 := [monotonicity #69]: #71
-#41 := (= #8 #40)
-#42 := [rewrite]: #41
-#75 := [monotonicity #42 #72]: #74
-#78 := [monotonicity #75]: #77
-#81 := [monotonicity #78]: #80
-#85 := [trans #81 #83]: #84
-#90 := [trans #85 #87]: #89
-#39 := [asserted]: #22
-#91 := [mp #39 #90]: #88
-#110 := [mp #91 #109]: #107
-#111 := [not-or-elim #110]: #102
-#127 := (or #94 #104)
-#128 := [th-lemma]: #127
-#129 := [unit-resolution #128 #111]: #94
-#92 := (>= uf_1 0::real)
-#112 := [not-or-elim #110]: #103
-#130 := (or #92 #105)
-#131 := [th-lemma]: #130
-#132 := [unit-resolution #131 #112]: #92
-#114 := (not #94)
-#113 := (not #92)
-#115 := (or #113 #114)
-#95 := (and #92 #94)
-#98 := (not #95)
-#124 := (iff #98 #115)
-#116 := (not #115)
-#119 := (not #116)
-#122 := (iff #119 #115)
-#123 := [rewrite]: #122
-#120 := (iff #98 #119)
-#117 := (iff #95 #116)
-#118 := [rewrite]: #117
-#121 := [monotonicity #118]: #120
-#125 := [trans #121 #123]: #124
-#6 := (= uf_1 0::real)
-#7 := (not #6)
-#99 := (iff #7 #98)
-#96 := (iff #6 #95)
-#97 := [rewrite]: #96
-#100 := [monotonicity #97]: #99
-#38 := [asserted]: #7
-#101 := [mp #38 #100]: #98
-#126 := [mp #101 #125]: #115
-[unit-resolution #126 #132 #129]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_linarith_21 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,10 +0,0 @@
-(benchmark Isabelle
-:extrafuns (
- (uf_2 Int)
- (uf_1 Int)
- )
-:assumption (= (mod (+ uf_1 uf_2) 2) 0)
-:assumption (= (mod uf_1 4) 3)
-:assumption (not (and (= (mod uf_1 2) 1) (= (mod uf_2 2) 1)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_linarith_21.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,208 +0,0 @@
-#2 := false
-#9 := 0::int
-#11 := 4::int
-decl uf_1 :: int
-#4 := uf_1
-#189 := (div uf_1 4::int)
-#210 := -4::int
-#211 := (* -4::int #189)
-#12 := (mod uf_1 4::int)
-#134 := -1::int
-#209 := (* -1::int #12)
-#212 := (+ #209 #211)
-#213 := (+ uf_1 #212)
-#214 := (<= #213 0::int)
-#215 := (not #214)
-#208 := (>= #213 0::int)
-#207 := (not #208)
-#216 := (or #207 #215)
-#217 := (not #216)
-#1 := true
-#36 := [true-axiom]: true
-#393 := (or false #217)
-#394 := [th-lemma]: #393
-#395 := [unit-resolution #394 #36]: #217
-#224 := (or #216 #214)
-#225 := [def-axiom]: #224
-#396 := [unit-resolution #225 #395]: #214
-#222 := (or #216 #208)
-#223 := [def-axiom]: #222
-#397 := [unit-resolution #223 #395]: #208
-#250 := (>= #12 4::int)
-#251 := (not #250)
-#398 := (or false #251)
-#399 := [th-lemma]: #398
-#400 := [unit-resolution #399 #36]: #251
-#13 := 3::int
-#90 := (>= #12 3::int)
-#92 := (not #90)
-#89 := (<= #12 3::int)
-#91 := (not #89)
-#93 := (or #91 #92)
-#94 := (not #93)
-#14 := (= #12 3::int)
-#95 := (iff #14 #94)
-#96 := [rewrite]: #95
-#38 := [asserted]: #14
-#97 := [mp #38 #96]: #94
-#99 := [not-or-elim #97]: #90
-#7 := 2::int
-#261 := (div uf_1 2::int)
-#140 := -2::int
-#276 := (* -2::int #261)
-#15 := (mod uf_1 2::int)
-#275 := (* -1::int #15)
-#277 := (+ #275 #276)
-#278 := (+ uf_1 #277)
-#279 := (<= #278 0::int)
-#280 := (not #279)
-#274 := (>= #278 0::int)
-#273 := (not #274)
-#281 := (or #273 #280)
-#282 := (not #281)
-#408 := (or false #282)
-#409 := [th-lemma]: #408
-#410 := [unit-resolution #409 #36]: #282
-#289 := (or #281 #279)
-#290 := [def-axiom]: #289
-#411 := [unit-resolution #290 #410]: #279
-#287 := (or #281 #274)
-#288 := [def-axiom]: #287
-#412 := [unit-resolution #288 #410]: #274
-#16 := 1::int
-#55 := (>= #15 1::int)
-#100 := (not #55)
-decl uf_2 :: int
-#5 := uf_2
-#18 := (mod uf_2 2::int)
-#61 := (<= #18 1::int)
-#102 := (not #61)
-#375 := [hypothesis]: #102
-#358 := (>= #18 2::int)
-#359 := (not #358)
-#403 := (or false #359)
-#404 := [th-lemma]: #403
-#405 := [unit-resolution #404 #36]: #359
-#406 := [th-lemma #405 #375]: false
-#407 := [lemma #406]: #61
-#413 := (or #100 #102)
-#62 := (>= #18 1::int)
-#315 := (div uf_2 2::int)
-#330 := (* -2::int #315)
-#329 := (* -1::int #18)
-#331 := (+ #329 #330)
-#332 := (+ uf_2 #331)
-#333 := (<= #332 0::int)
-#334 := (not #333)
-#328 := (>= #332 0::int)
-#327 := (not #328)
-#335 := (or #327 #334)
-#336 := (not #335)
-#376 := (or false #336)
-#377 := [th-lemma]: #376
-#378 := [unit-resolution #377 #36]: #336
-#343 := (or #335 #333)
-#344 := [def-axiom]: #343
-#379 := [unit-resolution #344 #378]: #333
-#341 := (or #335 #328)
-#342 := [def-axiom]: #341
-#380 := [unit-resolution #342 #378]: #328
-#103 := (not #62)
-#381 := [hypothesis]: #103
-#352 := (>= #18 0::int)
-#382 := (or false #352)
-#383 := [th-lemma]: #382
-#384 := [unit-resolution #383 #36]: #352
-#6 := (+ uf_1 uf_2)
-#116 := (div #6 2::int)
-#141 := (* -2::int #116)
-#8 := (mod #6 2::int)
-#139 := (* -1::int #8)
-#142 := (+ #139 #141)
-#143 := (+ uf_2 #142)
-#144 := (+ uf_1 #143)
-#138 := (<= #144 0::int)
-#136 := (not #138)
-#137 := (>= #144 0::int)
-#135 := (not #137)
-#145 := (or #135 #136)
-#146 := (not #145)
-#385 := (or false #146)
-#386 := [th-lemma]: #385
-#387 := [unit-resolution #386 #36]: #146
-#153 := (or #145 #138)
-#154 := [def-axiom]: #153
-#388 := [unit-resolution #154 #387]: #138
-#151 := (or #145 #137)
-#152 := [def-axiom]: #151
-#389 := [unit-resolution #152 #387]: #137
-#78 := (<= #8 0::int)
-#79 := (>= #8 0::int)
-#81 := (not #79)
-#80 := (not #78)
-#82 := (or #80 #81)
-#83 := (not #82)
-#10 := (= #8 0::int)
-#84 := (iff #10 #83)
-#85 := [rewrite]: #84
-#37 := [asserted]: #10
-#86 := [mp #37 #85]: #83
-#87 := [not-or-elim #86]: #78
-#390 := (or false #79)
-#391 := [th-lemma]: #390
-#392 := [unit-resolution #391 #36]: #79
-#401 := [th-lemma #99 #400 #397 #396 #392 #87 #389 #388 #384 #381 #380 #379]: false
-#402 := [lemma #401]: #62
-#57 := (<= #15 1::int)
-#101 := (not #57)
-#369 := [hypothesis]: #101
-#304 := (>= #15 2::int)
-#305 := (not #304)
-#370 := (or false #305)
-#371 := [th-lemma]: #370
-#372 := [unit-resolution #371 #36]: #305
-#373 := [th-lemma #372 #369]: false
-#374 := [lemma #373]: #57
-#104 := (or #100 #101 #102 #103)
-#69 := (and #55 #57 #61 #62)
-#74 := (not #69)
-#113 := (iff #74 #104)
-#105 := (not #104)
-#108 := (not #105)
-#111 := (iff #108 #104)
-#112 := [rewrite]: #111
-#109 := (iff #74 #108)
-#106 := (iff #69 #105)
-#107 := [rewrite]: #106
-#110 := [monotonicity #107]: #109
-#114 := [trans #110 #112]: #113
-#19 := (= #18 1::int)
-#17 := (= #15 1::int)
-#20 := (and #17 #19)
-#21 := (not #20)
-#75 := (iff #21 #74)
-#72 := (iff #20 #69)
-#63 := (and #61 #62)
-#58 := (and #55 #57)
-#66 := (and #58 #63)
-#70 := (iff #66 #69)
-#71 := [rewrite]: #70
-#67 := (iff #20 #66)
-#64 := (iff #19 #63)
-#65 := [rewrite]: #64
-#59 := (iff #17 #58)
-#60 := [rewrite]: #59
-#68 := [monotonicity #60 #65]: #67
-#73 := [trans #68 #71]: #72
-#76 := [monotonicity #73]: #75
-#39 := [asserted]: #21
-#77 := [mp #39 #76]: #74
-#115 := [mp #77 #114]: #104
-#414 := [unit-resolution #115 #374 #402]: #413
-#415 := [unit-resolution #414 #407]: #100
-#298 := (>= #15 0::int)
-#416 := (or false #298)
-#417 := [th-lemma]: #416
-#418 := [unit-resolution #417 #36]: #298
-[th-lemma #418 #415 #412 #411 #99 #400 #397 #396]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_mono_01 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,58 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T2 T13 T12 T11 T10 T9 T8 T7 T6 T5 T4 T1 T3)
-:extrafuns (
- (uf_37 T13)
- (uf_34 T12)
- (uf_31 T11)
- (uf_28 T10)
- (uf_25 T9)
- (uf_22 T8)
- (uf_19 T7)
- (uf_16 T6)
- (uf_13 T5)
- (uf_10 T4)
- (uf_7 T1)
- (uf_4 T3)
- (uf_36 Int T13 T13)
- (uf_33 T13 T12 T12)
- (uf_30 T12 T11 T11)
- (uf_27 T11 T10 T10)
- (uf_24 T10 T9 T9)
- (uf_21 T9 T8 T8)
- (uf_18 T8 T7 T7)
- (uf_15 T7 T6 T6)
- (uf_12 T6 T5 T5)
- (uf_9 T5 T4 T4)
- (uf_6 T4 T1 T1)
- (uf_3 T1 T3 T3)
- )
-:extrapreds (
- (up_35 Int)
- (up_32 T13)
- (up_29 T12)
- (up_26 T11)
- (up_23 T10)
- (up_20 T9)
- (up_17 T8)
- (up_14 T7)
- (up_11 T6)
- (up_8 T5)
- (up_5 T4)
- (up_1 T1)
- (up_2 T3)
- )
-:assumption (forall (?x1 T1) (and (up_1 ?x1) (or (up_2 (uf_3 ?x1 uf_4)) (not (up_2 (uf_3 ?x1 uf_4))))))
-:assumption (forall (?x2 T4) (and (up_5 ?x2) (or (up_1 (uf_6 ?x2 uf_7)) (not (up_1 (uf_6 ?x2 uf_7))))))
-:assumption (forall (?x3 T5) (and (up_8 ?x3) (or (up_5 (uf_9 ?x3 uf_10)) (not (up_5 (uf_9 ?x3 uf_10))))))
-:assumption (forall (?x4 T6) (and (up_11 ?x4) (or (up_8 (uf_12 ?x4 uf_13)) (not (up_8 (uf_12 ?x4 uf_13))))))
-:assumption (forall (?x5 T7) (and (up_14 ?x5) (or (up_11 (uf_15 ?x5 uf_16)) (not (up_11 (uf_15 ?x5 uf_16))))))
-:assumption (forall (?x6 T8) (and (up_17 ?x6) (or (up_14 (uf_18 ?x6 uf_19)) (not (up_14 (uf_18 ?x6 uf_19))))))
-:assumption (forall (?x7 T9) (and (up_20 ?x7) (or (up_17 (uf_21 ?x7 uf_22)) (not (up_17 (uf_21 ?x7 uf_22))))))
-:assumption (forall (?x8 T10) (and (up_23 ?x8) (or (up_20 (uf_24 ?x8 uf_25)) (not (up_20 (uf_24 ?x8 uf_25))))))
-:assumption (forall (?x9 T11) (and (up_26 ?x9) (or (up_23 (uf_27 ?x9 uf_28)) (not (up_23 (uf_27 ?x9 uf_28))))))
-:assumption (forall (?x10 T12) (and (up_29 ?x10) (or (up_26 (uf_30 ?x10 uf_31)) (not (up_26 (uf_30 ?x10 uf_31))))))
-:assumption (forall (?x11 T13) (and (up_32 ?x11) (or (up_29 (uf_33 ?x11 uf_34)) (not (up_29 (uf_33 ?x11 uf_34))))))
-:assumption (forall (?x12 Int) (and (up_35 ?x12) (or (up_32 (uf_36 ?x12 uf_37)) (not (up_32 (uf_36 ?x12 uf_37))))))
-:assumption (not (up_35 1))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_mono_01.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,50 +0,0 @@
-#2 := false
-decl up_35 :: (-> int bool)
-#112 := 1::int
-#113 := (up_35 1::int)
-#114 := (not #113)
-#297 := [asserted]: #114
-#103 := (:var 0 int)
-#104 := (up_35 #103)
-#911 := (pattern #104)
-#912 := (forall (vars (?x12 int)) (:pat #911) #104)
-#294 := (forall (vars (?x12 int)) #104)
-#915 := (iff #294 #912)
-#913 := (iff #104 #104)
-#914 := [refl]: #913
-#916 := [quant-intro #914]: #915
-#320 := (~ #294 #294)
-#361 := (~ #104 #104)
-#362 := [refl]: #361
-#321 := [nnf-pos #362]: #320
-decl up_32 :: (-> T13 bool)
-decl uf_36 :: (-> int T13 T13)
-decl uf_37 :: T13
-#105 := uf_37
-#106 := (uf_36 #103 uf_37)
-#107 := (up_32 #106)
-#108 := (not #107)
-#109 := (or #107 #108)
-#110 := (and #104 #109)
-#111 := (forall (vars (?x12 int)) #110)
-#295 := (iff #111 #294)
-#292 := (iff #110 #104)
-#1 := true
-#287 := (and #104 true)
-#290 := (iff #287 #104)
-#291 := [rewrite]: #290
-#288 := (iff #110 #287)
-#284 := (iff #109 true)
-#286 := [rewrite]: #284
-#289 := [monotonicity #286]: #288
-#293 := [trans #289 #291]: #292
-#296 := [quant-intro #293]: #295
-#283 := [asserted]: #111
-#299 := [mp #283 #296]: #294
-#363 := [mp~ #299 #321]: #294
-#917 := [mp #363 #916]: #912
-#418 := (not #912)
-#504 := (or #418 #113)
-#419 := [quant-inst]: #504
-[unit-resolution #419 #917 #297]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_mono_02 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,41 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T5 T6 T3 T1 T2 T4 T8 T7)
-:extrafuns (
- (uf_19 T1)
- (uf_3 Int T3)
- (uf_7 T2)
- (uf_8 T4)
- (uf_2 T1 T2 T2)
- (uf_6 Int T4 T4)
- (uf_10 T5 T1 T3)
- (uf_12 T6 Int T3)
- (uf_13 T2 T3)
- (uf_14 T4 T3)
- (uf_17 T8 T3)
- (uf_15 T7 T3)
- (uf_18 T1 T8)
- (uf_16 Int T7)
- (uf_9 T5 T2 T3)
- (uf_11 T6 T4 T3)
- (uf_1 T2 T3)
- (uf_5 T4 T3)
- (uf_4 T3 Int)
- )
-:assumption (forall (?x1 T1) (?x2 T2) (= (uf_1 (uf_2 ?x1 ?x2)) (uf_3 (+ (uf_4 (uf_1 ?x2)) (uf_4 (uf_3 (+ 0 1)))))))
-:assumption (forall (?x3 Int) (?x4 T4) (= (uf_5 (uf_6 ?x3 ?x4)) (uf_3 (+ (uf_4 (uf_5 ?x4)) (uf_4 (uf_3 (+ 0 1)))))))
-:assumption (= (uf_1 uf_7) (uf_3 0))
-:assumption (= (uf_5 uf_8) (uf_3 0))
-:assumption (forall (?x5 T5) (?x6 T1) (?x7 T2) (= (uf_9 ?x5 (uf_2 ?x6 ?x7)) (uf_3 (+ (+ (uf_4 (uf_10 ?x5 ?x6)) (uf_4 (uf_9 ?x5 ?x7))) (uf_4 (uf_3 (+ 0 1)))))))
-:assumption (forall (?x8 T6) (?x9 Int) (?x10 T4) (= (uf_11 ?x8 (uf_6 ?x9 ?x10)) (uf_3 (+ (+ (uf_4 (uf_12 ?x8 ?x9)) (uf_4 (uf_11 ?x8 ?x10))) (uf_4 (uf_3 (+ 0 1)))))))
-:assumption (forall (?x11 T5) (= (uf_9 ?x11 uf_7) (uf_3 0)))
-:assumption (forall (?x12 T6) (= (uf_11 ?x12 uf_8) (uf_3 0)))
-:assumption (forall (?x13 T2) (= (uf_13 ?x13) (uf_1 ?x13)))
-:assumption (forall (?x14 T4) (= (uf_14 ?x14) (uf_5 ?x14)))
-:assumption (forall (?x15 Int) (= (uf_15 (uf_16 ?x15)) (uf_14 (uf_6 ?x15 uf_8))))
-:assumption (forall (?x16 T1) (= (uf_17 (uf_18 ?x16)) (uf_13 (uf_2 ?x16 uf_7))))
-:assumption (forall (?x17 T3) (= (uf_3 (uf_4 ?x17)) ?x17))
-:assumption (forall (?x18 Int) (implies (<= 0 ?x18) (= (uf_4 (uf_3 ?x18)) ?x18)))
-:assumption (forall (?x19 Int) (implies (< ?x19 0) (= (uf_4 (uf_3 ?x19)) 0)))
-:assumption (not (= (uf_15 (uf_16 3)) (uf_17 (uf_18 uf_19))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_mono_02.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,492 +0,0 @@
-#2 := false
-decl uf_17 :: (-> T8 T3)
-decl uf_18 :: (-> T1 T8)
-decl uf_19 :: T1
-#104 := uf_19
-#105 := (uf_18 uf_19)
-#106 := (uf_17 #105)
-decl uf_15 :: (-> T7 T3)
-decl uf_16 :: (-> int T7)
-#101 := 3::int
-#102 := (uf_16 3::int)
-#103 := (uf_15 #102)
-#107 := (= #103 #106)
-decl uf_13 :: (-> T2 T3)
-decl uf_2 :: (-> T1 T2 T2)
-decl uf_7 :: T2
-#29 := uf_7
-#857 := (uf_2 uf_19 uf_7)
-#859 := (uf_13 #857)
-#599 := (= #859 #106)
-#526 := (= #106 #859)
-#79 := (:var 0 T1)
-#82 := (uf_2 #79 uf_7)
-#932 := (pattern #82)
-#80 := (uf_18 #79)
-#931 := (pattern #80)
-#83 := (uf_13 #82)
-#81 := (uf_17 #80)
-#84 := (= #81 #83)
-#933 := (forall (vars (?x16 T1)) (:pat #931 #932) #84)
-#85 := (forall (vars (?x16 T1)) #84)
-#936 := (iff #85 #933)
-#934 := (iff #84 #84)
-#935 := [refl]: #934
-#937 := [quant-intro #935]: #936
-#347 := (~ #85 #85)
-#384 := (~ #84 #84)
-#385 := [refl]: #384
-#348 := [nnf-pos #385]: #347
-#238 := [asserted]: #85
-#386 := [mp~ #238 #348]: #85
-#938 := [mp #386 #937]: #933
-#861 := (not #933)
-#862 := (or #861 #526)
-#863 := [quant-inst]: #862
-#601 := [unit-resolution #863 #938]: #526
-#588 := [symm #601]: #599
-#586 := (= #103 #859)
-decl uf_1 :: (-> T2 T3)
-#558 := (uf_1 #857)
-#832 := (= #558 #859)
-#5 := (:var 0 T2)
-#66 := (uf_13 #5)
-#908 := (pattern #66)
-#8 := (uf_1 #5)
-#907 := (pattern #8)
-#222 := (= #8 #66)
-#909 := (forall (vars (?x13 T2)) (:pat #907 #908) #222)
-#226 := (forall (vars (?x13 T2)) #222)
-#912 := (iff #226 #909)
-#910 := (iff #222 #222)
-#911 := [refl]: #910
-#913 := [quant-intro #911]: #912
-#341 := (~ #226 #226)
-#375 := (~ #222 #222)
-#376 := [refl]: #375
-#342 := [nnf-pos #376]: #341
-#67 := (= #66 #8)
-#68 := (forall (vars (?x13 T2)) #67)
-#227 := (iff #68 #226)
-#224 := (iff #67 #222)
-#225 := [rewrite]: #224
-#228 := [quant-intro #225]: #227
-#221 := [asserted]: #68
-#231 := [mp #221 #228]: #226
-#377 := [mp~ #231 #342]: #226
-#914 := [mp #377 #913]: #909
-#451 := (not #909)
-#837 := (or #451 #832)
-#547 := [quant-inst]: #837
-#615 := [unit-resolution #547 #914]: #832
-#585 := (= #103 #558)
-decl uf_3 :: (-> int T3)
-decl uf_4 :: (-> T3 int)
-#30 := (uf_1 uf_7)
-#806 := (uf_4 #30)
-#11 := 1::int
-#127 := (uf_3 1::int)
-#130 := (uf_4 #127)
-#649 := (+ #130 #806)
-#794 := (uf_3 #649)
-#597 := (= #794 #558)
-#683 := (= #558 #794)
-#4 := (:var 1 T1)
-#6 := (uf_2 #4 #5)
-#865 := (pattern #6)
-#9 := (uf_4 #8)
-#133 := (+ #9 #130)
-#136 := (uf_3 #133)
-#7 := (uf_1 #6)
-#139 := (= #7 #136)
-#866 := (forall (vars (?x1 T1) (?x2 T2)) (:pat #865) #139)
-#142 := (forall (vars (?x1 T1) (?x2 T2)) #139)
-#869 := (iff #142 #866)
-#867 := (iff #139 #139)
-#868 := [refl]: #867
-#870 := [quant-intro #868]: #869
-#361 := (~ #142 #142)
-#359 := (~ #139 #139)
-#360 := [refl]: #359
-#362 := [nnf-pos #360]: #361
-#10 := 0::int
-#12 := (+ 0::int 1::int)
-#13 := (uf_3 #12)
-#14 := (uf_4 #13)
-#15 := (+ #9 #14)
-#16 := (uf_3 #15)
-#17 := (= #7 #16)
-#18 := (forall (vars (?x1 T1) (?x2 T2)) #17)
-#143 := (iff #18 #142)
-#140 := (iff #17 #139)
-#137 := (= #16 #136)
-#134 := (= #15 #133)
-#131 := (= #14 #130)
-#128 := (= #13 #127)
-#125 := (= #12 1::int)
-#126 := [rewrite]: #125
-#129 := [monotonicity #126]: #128
-#132 := [monotonicity #129]: #131
-#135 := [monotonicity #132]: #134
-#138 := [monotonicity #135]: #137
-#141 := [monotonicity #138]: #140
-#144 := [quant-intro #141]: #143
-#124 := [asserted]: #18
-#147 := [mp #124 #144]: #142
-#363 := [mp~ #147 #362]: #142
-#871 := [mp #363 #870]: #866
-#701 := (not #866)
-#694 := (or #701 #683)
-#688 := (+ #806 #130)
-#689 := (uf_3 #688)
-#690 := (= #558 #689)
-#702 := (or #701 #690)
-#704 := (iff #702 #694)
-#706 := (iff #694 #694)
-#799 := [rewrite]: #706
-#698 := (iff #690 #683)
-#795 := (= #689 #794)
-#797 := (= #688 #649)
-#699 := [rewrite]: #797
-#798 := [monotonicity #699]: #795
-#700 := [monotonicity #798]: #698
-#705 := [monotonicity #700]: #704
-#796 := [trans #705 #799]: #704
-#703 := [quant-inst]: #702
-#800 := [mp #703 #796]: #694
-#614 := [unit-resolution #800 #871]: #683
-#598 := [symm #614]: #597
-#583 := (= #103 #794)
-#595 := (= #127 #794)
-#605 := (= #794 #127)
-#618 := (= #649 1::int)
-#780 := (<= #806 0::int)
-#778 := (= #806 0::int)
-#31 := (uf_3 0::int)
-#858 := (uf_4 #31)
-#855 := (= #858 0::int)
-#72 := (:var 0 int)
-#92 := (uf_3 #72)
-#947 := (pattern #92)
-#266 := (>= #72 0::int)
-#267 := (not #266)
-#93 := (uf_4 #92)
-#248 := (= #72 #93)
-#273 := (or #248 #267)
-#948 := (forall (vars (?x18 int)) (:pat #947) #273)
-#278 := (forall (vars (?x18 int)) #273)
-#951 := (iff #278 #948)
-#949 := (iff #273 #273)
-#950 := [refl]: #949
-#952 := [quant-intro #950]: #951
-#351 := (~ #278 #278)
-#390 := (~ #273 #273)
-#391 := [refl]: #390
-#352 := [nnf-pos #391]: #351
-#94 := (= #93 #72)
-#91 := (<= 0::int #72)
-#95 := (implies #91 #94)
-#96 := (forall (vars (?x18 int)) #95)
-#281 := (iff #96 #278)
-#255 := (not #91)
-#256 := (or #255 #248)
-#261 := (forall (vars (?x18 int)) #256)
-#279 := (iff #261 #278)
-#276 := (iff #256 #273)
-#270 := (or #267 #248)
-#274 := (iff #270 #273)
-#275 := [rewrite]: #274
-#271 := (iff #256 #270)
-#268 := (iff #255 #267)
-#264 := (iff #91 #266)
-#265 := [rewrite]: #264
-#269 := [monotonicity #265]: #268
-#272 := [monotonicity #269]: #271
-#277 := [trans #272 #275]: #276
-#280 := [quant-intro #277]: #279
-#262 := (iff #96 #261)
-#259 := (iff #95 #256)
-#252 := (implies #91 #248)
-#257 := (iff #252 #256)
-#258 := [rewrite]: #257
-#253 := (iff #95 #252)
-#250 := (iff #94 #248)
-#251 := [rewrite]: #250
-#254 := [monotonicity #251]: #253
-#260 := [trans #254 #258]: #259
-#263 := [quant-intro #260]: #262
-#282 := [trans #263 #280]: #281
-#247 := [asserted]: #96
-#283 := [mp #247 #282]: #278
-#392 := [mp~ #283 #352]: #278
-#953 := [mp #392 #952]: #948
-#848 := (not #948)
-#850 := (or #848 #855)
-#527 := (>= 0::int 0::int)
-#860 := (not #527)
-#864 := (= 0::int #858)
-#854 := (or #864 #860)
-#489 := (or #848 #854)
-#851 := (iff #489 #850)
-#852 := (iff #850 #850)
-#838 := [rewrite]: #852
-#847 := (iff #854 #855)
-#843 := (or #855 false)
-#846 := (iff #843 #855)
-#841 := [rewrite]: #846
-#844 := (iff #854 #843)
-#505 := (iff #860 false)
-#1 := true
-#498 := (not true)
-#503 := (iff #498 false)
-#504 := [rewrite]: #503
-#840 := (iff #860 #498)
-#514 := (iff #527 true)
-#856 := [rewrite]: #514
-#502 := [monotonicity #856]: #840
-#842 := [trans #502 #504]: #505
-#513 := (iff #864 #855)
-#518 := [rewrite]: #513
-#845 := [monotonicity #518 #842]: #844
-#484 := [trans #845 #841]: #847
-#849 := [monotonicity #484]: #851
-#839 := [trans #849 #838]: #851
-#490 := [quant-inst]: #489
-#546 := [mp #490 #839]: #850
-#644 := [unit-resolution #546 #953]: #855
-#621 := (= #806 #858)
-#32 := (= #30 #31)
-#159 := [asserted]: #32
-#626 := [monotonicity #159]: #621
-#616 := [trans #626 #644]: #778
-#606 := (not #778)
-#608 := (or #606 #780)
-#609 := [th-lemma]: #608
-#612 := [unit-resolution #609 #616]: #780
-#790 := (>= #806 0::int)
-#613 := (or #606 #790)
-#617 := [th-lemma]: #613
-#610 := [unit-resolution #617 #616]: #790
-#723 := (<= #130 1::int)
-#746 := (= #130 1::int)
-#713 := (or #848 #746)
-#755 := (>= 1::int 0::int)
-#756 := (not #755)
-#743 := (= 1::int #130)
-#744 := (or #743 #756)
-#714 := (or #848 #744)
-#718 := (iff #714 #713)
-#720 := (iff #713 #713)
-#725 := [rewrite]: #720
-#739 := (iff #744 #746)
-#735 := (or #746 false)
-#738 := (iff #735 #746)
-#733 := [rewrite]: #738
-#736 := (iff #744 #735)
-#731 := (iff #756 false)
-#734 := (iff #756 #498)
-#742 := (iff #755 true)
-#748 := [rewrite]: #742
-#730 := [monotonicity #748]: #734
-#732 := [trans #730 #504]: #731
-#745 := (iff #743 #746)
-#747 := [rewrite]: #745
-#737 := [monotonicity #747 #732]: #736
-#712 := [trans #737 #733]: #739
-#719 := [monotonicity #712]: #718
-#721 := [trans #719 #725]: #718
-#607 := [quant-inst]: #714
-#722 := [mp #607 #721]: #713
-#641 := [unit-resolution #722 #953]: #746
-#620 := (not #746)
-#623 := (or #620 #723)
-#627 := [th-lemma]: #623
-#629 := [unit-resolution #627 #641]: #723
-#726 := (>= #130 1::int)
-#630 := (or #620 #726)
-#628 := [th-lemma]: #630
-#631 := [unit-resolution #628 #641]: #726
-#611 := [th-lemma #631 #629 #610 #612]: #618
-#587 := [monotonicity #611]: #605
-#596 := [symm #587]: #595
-#581 := (= #103 #127)
-decl uf_5 :: (-> T4 T3)
-decl uf_8 :: T4
-#33 := uf_8
-#34 := (uf_5 uf_8)
-#822 := (uf_4 #34)
-#824 := (+ #130 #822)
-#666 := (uf_3 #824)
-#593 := (= #666 #127)
-#589 := (= #127 #666)
-#624 := (= 1::int #824)
-#619 := (= #824 1::int)
-#789 := (<= #822 0::int)
-#787 := (= #822 0::int)
-#632 := (= #822 #858)
-#35 := (= #34 #31)
-#162 := (= #31 #34)
-#163 := (iff #35 #162)
-#164 := [rewrite]: #163
-#160 := [asserted]: #35
-#167 := [mp #160 #164]: #162
-#662 := [symm #167]: #35
-#633 := [monotonicity #662]: #632
-#634 := [trans #633 #644]: #787
-#635 := (not #787)
-#637 := (or #635 #789)
-#638 := [th-lemma]: #637
-#639 := [unit-resolution #638 #634]: #789
-#781 := (>= #822 0::int)
-#481 := (or #635 #781)
-#640 := [th-lemma]: #481
-#636 := [unit-resolution #640 #634]: #781
-#622 := [th-lemma #631 #629 #636 #639]: #619
-#625 := [symm #622]: #624
-#590 := [monotonicity #625]: #589
-#594 := [symm #590]: #593
-#579 := (= #103 #666)
-decl uf_6 :: (-> int T4 T4)
-#539 := (uf_6 3::int uf_8)
-#836 := (uf_5 #539)
-#810 := (= #836 #666)
-#813 := (= #666 #836)
-#20 := (:var 0 T4)
-#19 := (:var 1 int)
-#21 := (uf_6 #19 #20)
-#872 := (pattern #21)
-#23 := (uf_5 #20)
-#24 := (uf_4 #23)
-#146 := (+ #24 #130)
-#150 := (uf_3 #146)
-#22 := (uf_5 #21)
-#153 := (= #22 #150)
-#873 := (forall (vars (?x3 int) (?x4 T4)) (:pat #872) #153)
-#156 := (forall (vars (?x3 int) (?x4 T4)) #153)
-#876 := (iff #156 #873)
-#874 := (iff #153 #153)
-#875 := [refl]: #874
-#877 := [quant-intro #875]: #876
-#328 := (~ #156 #156)
-#364 := (~ #153 #153)
-#365 := [refl]: #364
-#326 := [nnf-pos #365]: #328
-#25 := (+ #24 #14)
-#26 := (uf_3 #25)
-#27 := (= #22 #26)
-#28 := (forall (vars (?x3 int) (?x4 T4)) #27)
-#157 := (iff #28 #156)
-#154 := (iff #27 #153)
-#151 := (= #26 #150)
-#148 := (= #25 #146)
-#149 := [monotonicity #132]: #148
-#152 := [monotonicity #149]: #151
-#155 := [monotonicity #152]: #154
-#158 := [quant-intro #155]: #157
-#145 := [asserted]: #28
-#161 := [mp #145 #158]: #156
-#366 := [mp~ #161 #326]: #156
-#878 := [mp #366 #877]: #873
-#809 := (not #873)
-#816 := (or #809 #813)
-#817 := (+ #822 #130)
-#818 := (uf_3 #817)
-#823 := (= #836 #818)
-#645 := (or #809 #823)
-#648 := (iff #645 #816)
-#802 := (iff #816 #816)
-#804 := [rewrite]: #802
-#814 := (iff #823 #813)
-#807 := (iff #810 #813)
-#808 := [rewrite]: #807
-#811 := (iff #823 #810)
-#667 := (= #818 #666)
-#819 := (= #817 #824)
-#825 := [rewrite]: #819
-#668 := [monotonicity #825]: #667
-#812 := [monotonicity #668]: #811
-#815 := [trans #812 #808]: #814
-#801 := [monotonicity #815]: #648
-#805 := [trans #801 #804]: #648
-#647 := [quant-inst]: #645
-#803 := [mp #647 #805]: #816
-#658 := [unit-resolution #803 #878]: #813
-#592 := [symm #658]: #810
-#600 := (= #103 #836)
-decl uf_14 :: (-> T4 T3)
-#540 := (uf_14 #539)
-#548 := (= #540 #836)
-#69 := (uf_14 #20)
-#916 := (pattern #69)
-#915 := (pattern #23)
-#230 := (= #23 #69)
-#917 := (forall (vars (?x14 T4)) (:pat #915 #916) #230)
-#234 := (forall (vars (?x14 T4)) #230)
-#920 := (iff #234 #917)
-#918 := (iff #230 #230)
-#919 := [refl]: #918
-#921 := [quant-intro #919]: #920
-#343 := (~ #234 #234)
-#378 := (~ #230 #230)
-#379 := [refl]: #378
-#344 := [nnf-pos #379]: #343
-#70 := (= #69 #23)
-#71 := (forall (vars (?x14 T4)) #70)
-#235 := (iff #71 #234)
-#232 := (iff #70 #230)
-#233 := [rewrite]: #232
-#236 := [quant-intro #233]: #235
-#229 := [asserted]: #71
-#239 := [mp #229 #236]: #234
-#380 := [mp~ #239 #344]: #234
-#922 := [mp #380 #921]: #917
-#541 := (not #917)
-#828 := (or #541 #548)
-#833 := (= #836 #540)
-#829 := (or #541 #833)
-#826 := (iff #829 #828)
-#827 := (iff #828 #828)
-#831 := [rewrite]: #827
-#549 := (iff #833 #548)
-#550 := [rewrite]: #549
-#830 := [monotonicity #550]: #826
-#820 := [trans #830 #831]: #826
-#543 := [quant-inst]: #829
-#821 := [mp #543 #820]: #828
-#657 := [unit-resolution #821 #922]: #548
-#521 := (= #103 #540)
-#75 := (uf_6 #72 uf_8)
-#924 := (pattern #75)
-#73 := (uf_16 #72)
-#923 := (pattern #73)
-#76 := (uf_14 #75)
-#74 := (uf_15 #73)
-#77 := (= #74 #76)
-#925 := (forall (vars (?x15 int)) (:pat #923 #924) #77)
-#78 := (forall (vars (?x15 int)) #77)
-#928 := (iff #78 #925)
-#926 := (iff #77 #77)
-#927 := [refl]: #926
-#929 := [quant-intro #927]: #928
-#345 := (~ #78 #78)
-#381 := (~ #77 #77)
-#382 := [refl]: #381
-#346 := [nnf-pos #382]: #345
-#237 := [asserted]: #78
-#383 := [mp~ #237 #346]: #78
-#930 := [mp #383 #929]: #925
-#515 := (not #925)
-#646 := (or #515 #521)
-#853 := [quant-inst]: #646
-#603 := [unit-resolution #853 #930]: #521
-#577 := [trans #603 #657]: #600
-#580 := [trans #577 #592]: #579
-#582 := [trans #580 #594]: #581
-#584 := [trans #582 #596]: #583
-#578 := [trans #584 #598]: #585
-#571 := [trans #578 #615]: #586
-#572 := [trans #571 #588]: #107
-#108 := (not #107)
-#325 := [asserted]: #108
-[unit-resolution #325 #572]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_nat_arith_01 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,13 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T1)
-:extrafuns (
- (uf_1 Int T1)
- (uf_2 T1 Int)
- (uf_3 T1)
- )
-:assumption (forall (?x1 T1) (= (uf_1 (uf_2 ?x1)) ?x1))
-:assumption (forall (?x2 Int) (implies (<= 0 ?x2) (= (uf_2 (uf_1 ?x2)) ?x2)))
-:assumption (forall (?x3 Int) (implies (< ?x3 0) (= (uf_2 (uf_1 ?x3)) 0)))
-:assumption (= (uf_1 (* 2 (uf_2 uf_3))) (uf_1 1))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_nat_arith_01.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,246 +0,0 @@
-#2 := false
-#9 := 0::int
-decl uf_2 :: (-> T1 int)
-decl uf_1 :: (-> int T1)
-decl uf_3 :: T1
-#22 := uf_3
-#23 := (uf_2 uf_3)
-#21 := 2::int
-#24 := (* 2::int #23)
-#25 := (uf_1 #24)
-#293 := (uf_2 #25)
-#292 := -1::int
-#296 := (* -1::int #293)
-#275 := (+ #24 #296)
-#258 := (<= #275 0::int)
-#611 := (= #275 0::int)
-#204 := (>= #24 0::int)
-#596 := (= #293 0::int)
-#541 := (not #596)
-#300 := (<= #293 0::int)
-#460 := (not #300)
-#26 := 1::int
-#570 := (>= #293 1::int)
-#569 := (= #293 1::int)
-#27 := (uf_1 1::int)
-#318 := (uf_2 #27)
-#311 := (= #318 1::int)
-#10 := (:var 0 int)
-#12 := (uf_1 #10)
-#627 := (pattern #12)
-#70 := (>= #10 0::int)
-#71 := (not #70)
-#13 := (uf_2 #12)
-#52 := (= #10 #13)
-#77 := (or #52 #71)
-#628 := (forall (vars (?x2 int)) (:pat #627) #77)
-#82 := (forall (vars (?x2 int)) #77)
-#631 := (iff #82 #628)
-#629 := (iff #77 #77)
-#630 := [refl]: #629
-#632 := [quant-intro #630]: #631
-#132 := (~ #82 #82)
-#144 := (~ #77 #77)
-#145 := [refl]: #144
-#130 := [nnf-pos #145]: #132
-#14 := (= #13 #10)
-#11 := (<= 0::int #10)
-#15 := (implies #11 #14)
-#16 := (forall (vars (?x2 int)) #15)
-#85 := (iff #16 #82)
-#59 := (not #11)
-#60 := (or #59 #52)
-#65 := (forall (vars (?x2 int)) #60)
-#83 := (iff #65 #82)
-#80 := (iff #60 #77)
-#74 := (or #71 #52)
-#78 := (iff #74 #77)
-#79 := [rewrite]: #78
-#75 := (iff #60 #74)
-#72 := (iff #59 #71)
-#68 := (iff #11 #70)
-#69 := [rewrite]: #68
-#73 := [monotonicity #69]: #72
-#76 := [monotonicity #73]: #75
-#81 := [trans #76 #79]: #80
-#84 := [quant-intro #81]: #83
-#66 := (iff #16 #65)
-#63 := (iff #15 #60)
-#56 := (implies #11 #52)
-#61 := (iff #56 #60)
-#62 := [rewrite]: #61
-#57 := (iff #15 #56)
-#54 := (iff #14 #52)
-#55 := [rewrite]: #54
-#58 := [monotonicity #55]: #57
-#64 := [trans #58 #62]: #63
-#67 := [quant-intro #64]: #66
-#86 := [trans #67 #84]: #85
-#51 := [asserted]: #16
-#87 := [mp #51 #86]: #82
-#146 := [mp~ #87 #130]: #82
-#633 := [mp #146 #632]: #628
-#612 := (not #628)
-#575 := (or #612 #311)
-#316 := (>= 1::int 0::int)
-#317 := (not #316)
-#211 := (= 1::int #318)
-#588 := (or #211 #317)
-#576 := (or #612 #588)
-#572 := (iff #576 #575)
-#578 := (iff #575 #575)
-#573 := [rewrite]: #578
-#585 := (iff #588 #311)
-#583 := (or #311 false)
-#584 := (iff #583 #311)
-#581 := [rewrite]: #584
-#297 := (iff #588 #583)
-#304 := (iff #317 false)
-#1 := true
-#587 := (not true)
-#302 := (iff #587 false)
-#303 := [rewrite]: #302
-#591 := (iff #317 #587)
-#586 := (iff #316 true)
-#590 := [rewrite]: #586
-#301 := [monotonicity #590]: #591
-#582 := [trans #301 #303]: #304
-#589 := (iff #211 #311)
-#312 := [rewrite]: #589
-#580 := [monotonicity #312 #582]: #297
-#574 := [trans #580 #581]: #585
-#577 := [monotonicity #574]: #572
-#579 := [trans #577 #573]: #572
-#571 := [quant-inst]: #576
-#420 := [mp #571 #579]: #575
-#437 := [unit-resolution #420 #633]: #311
-#452 := (= #293 #318)
-#28 := (= #25 #27)
-#129 := [asserted]: #28
-#454 := [monotonicity #129]: #452
-#455 := [trans #454 #437]: #569
-#448 := (not #569)
-#456 := (or #448 #570)
-#457 := [th-lemma]: #456
-#458 := [unit-resolution #457 #455]: #570
-#459 := (not #570)
-#553 := (or #459 #460)
-#550 := [th-lemma]: #553
-#554 := [unit-resolution #550 #458]: #460
-#543 := (or #541 #300)
-#535 := [th-lemma]: #543
-#532 := [unit-resolution #535 #554]: #541
-#598 := (or #204 #596)
-#18 := (= #13 0::int)
-#118 := (or #18 #70)
-#634 := (forall (vars (?x3 int)) (:pat #627) #118)
-#123 := (forall (vars (?x3 int)) #118)
-#637 := (iff #123 #634)
-#635 := (iff #118 #118)
-#636 := [refl]: #635
-#638 := [quant-intro #636]: #637
-#133 := (~ #123 #123)
-#147 := (~ #118 #118)
-#148 := [refl]: #147
-#134 := [nnf-pos #148]: #133
-#17 := (< #10 0::int)
-#19 := (implies #17 #18)
-#20 := (forall (vars (?x3 int)) #19)
-#126 := (iff #20 #123)
-#89 := (= 0::int #13)
-#95 := (not #17)
-#96 := (or #95 #89)
-#101 := (forall (vars (?x3 int)) #96)
-#124 := (iff #101 #123)
-#121 := (iff #96 #118)
-#115 := (or #70 #18)
-#119 := (iff #115 #118)
-#120 := [rewrite]: #119
-#116 := (iff #96 #115)
-#113 := (iff #89 #18)
-#114 := [rewrite]: #113
-#111 := (iff #95 #70)
-#106 := (not #71)
-#109 := (iff #106 #70)
-#110 := [rewrite]: #109
-#107 := (iff #95 #106)
-#104 := (iff #17 #71)
-#105 := [rewrite]: #104
-#108 := [monotonicity #105]: #107
-#112 := [trans #108 #110]: #111
-#117 := [monotonicity #112 #114]: #116
-#122 := [trans #117 #120]: #121
-#125 := [quant-intro #122]: #124
-#102 := (iff #20 #101)
-#99 := (iff #19 #96)
-#92 := (implies #17 #89)
-#97 := (iff #92 #96)
-#98 := [rewrite]: #97
-#93 := (iff #19 #92)
-#90 := (iff #18 #89)
-#91 := [rewrite]: #90
-#94 := [monotonicity #91]: #93
-#100 := [trans #94 #98]: #99
-#103 := [quant-intro #100]: #102
-#127 := [trans #103 #125]: #126
-#88 := [asserted]: #20
-#128 := [mp #88 #127]: #123
-#149 := [mp~ #128 #134]: #123
-#639 := [mp #149 #638]: #634
-#595 := (not #634)
-#601 := (or #595 #204 #596)
-#597 := (or #596 #204)
-#238 := (or #595 #597)
-#606 := (iff #238 #601)
-#604 := (or #595 #598)
-#605 := (iff #604 #601)
-#603 := [rewrite]: #605
-#243 := (iff #238 #604)
-#599 := (iff #597 #598)
-#600 := [rewrite]: #599
-#244 := [monotonicity #600]: #243
-#592 := [trans #244 #603]: #606
-#602 := [quant-inst]: #238
-#593 := [mp #602 #592]: #601
-#534 := [unit-resolution #593 #639]: #598
-#544 := [unit-resolution #534 #532]: #204
-#290 := (not #204)
-#281 := (or #290 #611)
-#618 := (or #612 #290 #611)
-#294 := (= #24 #293)
-#295 := (or #294 #290)
-#608 := (or #612 #295)
-#594 := (iff #608 #618)
-#272 := (or #612 #281)
-#610 := (iff #272 #618)
-#252 := [rewrite]: #610
-#609 := (iff #608 #272)
-#616 := (iff #295 #281)
-#400 := (or #611 #290)
-#614 := (iff #400 #281)
-#615 := [rewrite]: #614
-#607 := (iff #295 #400)
-#613 := (iff #294 #611)
-#269 := [rewrite]: #613
-#280 := [monotonicity #269]: #607
-#617 := [trans #280 #615]: #616
-#268 := [monotonicity #617]: #609
-#256 := [trans #268 #252]: #594
-#267 := [quant-inst]: #608
-#257 := [mp #267 #256]: #618
-#545 := [unit-resolution #257 #633]: #281
-#546 := [unit-resolution #545 #544]: #611
-#542 := (not #611)
-#547 := (or #542 #258)
-#536 := [th-lemma]: #547
-#537 := [unit-resolution #536 #546]: #258
-#259 := (>= #275 0::int)
-#538 := (or #542 #259)
-#539 := [th-lemma]: #538
-#533 := [unit-resolution #539 #546]: #259
-#563 := (<= #293 1::int)
-#540 := (or #448 #563)
-#524 := [th-lemma]: #540
-#525 := [unit-resolution #524 #455]: #563
-[th-lemma #458 #525 #533 #537]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_nat_arith_02 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,14 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T1)
-:extrafuns (
- (uf_1 Int T1)
- (uf_2 T1 Int)
- (uf_3 T1)
- )
-:assumption (forall (?x1 T1) (= (uf_1 (uf_2 ?x1)) ?x1))
-:assumption (forall (?x2 Int) (implies (<= 0 ?x2) (= (uf_2 (uf_1 ?x2)) ?x2)))
-:assumption (forall (?x3 Int) (implies (< ?x3 0) (= (uf_2 (uf_1 ?x3)) 0)))
-:assumption (< (uf_2 uf_3) 3)
-:assumption (not (< (uf_2 (uf_1 (* 2 (uf_2 uf_3)))) 7))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_nat_arith_02.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,199 +0,0 @@
-#2 := false
-#23 := 3::int
-decl uf_2 :: (-> T1 int)
-decl uf_3 :: T1
-#21 := uf_3
-#22 := (uf_2 uf_3)
-#137 := (>= #22 3::int)
-#135 := (not #137)
-#24 := (< #22 3::int)
-#136 := (iff #24 #135)
-#138 := [rewrite]: #136
-#132 := [asserted]: #24
-#139 := [mp #132 #138]: #135
-#9 := 0::int
-decl uf_1 :: (-> int T1)
-#25 := 2::int
-#26 := (* 2::int #22)
-#27 := (uf_1 #26)
-#28 := (uf_2 #27)
-#297 := -1::int
-#633 := (* -1::int #28)
-#635 := (+ #26 #633)
-#278 := (>= #635 0::int)
-#291 := (= #635 0::int)
-#315 := (>= #26 0::int)
-#279 := (= #28 0::int)
-#627 := (not #279)
-#624 := (<= #28 0::int)
-#281 := (not #624)
-#29 := 7::int
-#143 := (>= #28 7::int)
-#30 := (< #28 7::int)
-#31 := (not #30)
-#150 := (iff #31 #143)
-#141 := (not #143)
-#145 := (not #141)
-#148 := (iff #145 #143)
-#149 := [rewrite]: #148
-#146 := (iff #31 #145)
-#142 := (iff #30 #141)
-#144 := [rewrite]: #142
-#147 := [monotonicity #144]: #146
-#151 := [trans #147 #149]: #150
-#133 := [asserted]: #31
-#152 := [mp #133 #151]: #143
-#618 := (or #281 #141)
-#265 := [th-lemma]: #618
-#266 := [unit-resolution #265 #152]: #281
-#625 := (or #627 #624)
-#628 := [th-lemma]: #625
-#614 := [unit-resolution #628 #266]: #627
-#10 := (:var 0 int)
-#12 := (uf_1 #10)
-#649 := (pattern #12)
-#73 := (>= #10 0::int)
-#13 := (uf_2 #12)
-#18 := (= #13 0::int)
-#121 := (or #18 #73)
-#656 := (forall (vars (?x3 int)) (:pat #649) #121)
-#126 := (forall (vars (?x3 int)) #121)
-#659 := (iff #126 #656)
-#657 := (iff #121 #121)
-#658 := [refl]: #657
-#660 := [quant-intro #658]: #659
-#154 := (~ #126 #126)
-#170 := (~ #121 #121)
-#171 := [refl]: #170
-#155 := [nnf-pos #171]: #154
-#17 := (< #10 0::int)
-#19 := (implies #17 #18)
-#20 := (forall (vars (?x3 int)) #19)
-#129 := (iff #20 #126)
-#92 := (= 0::int #13)
-#98 := (not #17)
-#99 := (or #98 #92)
-#104 := (forall (vars (?x3 int)) #99)
-#127 := (iff #104 #126)
-#124 := (iff #99 #121)
-#118 := (or #73 #18)
-#122 := (iff #118 #121)
-#123 := [rewrite]: #122
-#119 := (iff #99 #118)
-#116 := (iff #92 #18)
-#117 := [rewrite]: #116
-#114 := (iff #98 #73)
-#74 := (not #73)
-#109 := (not #74)
-#112 := (iff #109 #73)
-#113 := [rewrite]: #112
-#110 := (iff #98 #109)
-#107 := (iff #17 #74)
-#108 := [rewrite]: #107
-#111 := [monotonicity #108]: #110
-#115 := [trans #111 #113]: #114
-#120 := [monotonicity #115 #117]: #119
-#125 := [trans #120 #123]: #124
-#128 := [quant-intro #125]: #127
-#105 := (iff #20 #104)
-#102 := (iff #19 #99)
-#95 := (implies #17 #92)
-#100 := (iff #95 #99)
-#101 := [rewrite]: #100
-#96 := (iff #19 #95)
-#93 := (iff #18 #92)
-#94 := [rewrite]: #93
-#97 := [monotonicity #94]: #96
-#103 := [trans #97 #101]: #102
-#106 := [quant-intro #103]: #105
-#130 := [trans #106 #128]: #129
-#91 := [asserted]: #20
-#131 := [mp #91 #130]: #126
-#172 := [mp~ #131 #155]: #126
-#661 := [mp #172 #660]: #656
-#619 := (not #656)
-#620 := (or #619 #279 #315)
-#280 := (or #279 #315)
-#621 := (or #619 #280)
-#617 := (iff #621 #620)
-#623 := [rewrite]: #617
-#622 := [quant-inst]: #621
-#260 := [mp #622 #623]: #620
-#615 := [unit-resolution #260 #661 #614]: #315
-#316 := (not #315)
-#302 := (or #291 #316)
-#55 := (= #10 #13)
-#80 := (or #55 #74)
-#650 := (forall (vars (?x2 int)) (:pat #649) #80)
-#85 := (forall (vars (?x2 int)) #80)
-#653 := (iff #85 #650)
-#651 := (iff #80 #80)
-#652 := [refl]: #651
-#654 := [quant-intro #652]: #653
-#153 := (~ #85 #85)
-#167 := (~ #80 #80)
-#168 := [refl]: #167
-#134 := [nnf-pos #168]: #153
-#14 := (= #13 #10)
-#11 := (<= 0::int #10)
-#15 := (implies #11 #14)
-#16 := (forall (vars (?x2 int)) #15)
-#88 := (iff #16 #85)
-#62 := (not #11)
-#63 := (or #62 #55)
-#68 := (forall (vars (?x2 int)) #63)
-#86 := (iff #68 #85)
-#83 := (iff #63 #80)
-#77 := (or #74 #55)
-#81 := (iff #77 #80)
-#82 := [rewrite]: #81
-#78 := (iff #63 #77)
-#75 := (iff #62 #74)
-#71 := (iff #11 #73)
-#72 := [rewrite]: #71
-#76 := [monotonicity #72]: #75
-#79 := [monotonicity #76]: #78
-#84 := [trans #79 #82]: #83
-#87 := [quant-intro #84]: #86
-#69 := (iff #16 #68)
-#66 := (iff #15 #63)
-#59 := (implies #11 #55)
-#64 := (iff #59 #63)
-#65 := [rewrite]: #64
-#60 := (iff #15 #59)
-#57 := (iff #14 #55)
-#58 := [rewrite]: #57
-#61 := [monotonicity #58]: #60
-#67 := [trans #61 #65]: #66
-#70 := [quant-intro #67]: #69
-#89 := [trans #70 #87]: #88
-#54 := [asserted]: #16
-#90 := [mp #54 #89]: #85
-#169 := [mp~ #90 #134]: #85
-#655 := [mp #169 #654]: #650
-#637 := (not #650)
-#638 := (or #637 #291 #316)
-#314 := (= #26 #28)
-#318 := (or #314 #316)
-#639 := (or #637 #318)
-#290 := (iff #639 #638)
-#640 := (or #637 #302)
-#294 := (iff #640 #638)
-#631 := [rewrite]: #294
-#630 := (iff #639 #640)
-#303 := (iff #318 #302)
-#422 := (iff #314 #291)
-#629 := [rewrite]: #422
-#636 := [monotonicity #629]: #303
-#289 := [monotonicity #636]: #630
-#632 := [trans #289 #631]: #290
-#634 := [quant-inst]: #639
-#274 := [mp #634 #632]: #638
-#322 := [unit-resolution #274 #655]: #302
-#337 := [unit-resolution #322 #615]: #291
-#338 := (not #291)
-#339 := (or #338 #278)
-#340 := [th-lemma]: #339
-#232 := [unit-resolution #340 #337]: #278
-[th-lemma #152 #232 #139]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_nat_arith_03 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,13 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T1)
-:extrafuns (
- (uf_1 Int T1)
- (uf_2 T1 Int)
- (uf_3 T1)
- )
-:assumption (forall (?x1 T1) (= (uf_1 (uf_2 ?x1)) ?x1))
-:assumption (forall (?x2 Int) (implies (<= 0 ?x2) (= (uf_2 (uf_1 ?x2)) ?x2)))
-:assumption (forall (?x3 Int) (implies (< ?x3 0) (= (uf_2 (uf_1 ?x3)) 0)))
-:assumption (not (let (?x4 (uf_1 (+ 1 (uf_2 uf_3)))) (< (uf_2 (uf_1 (* 0 (uf_2 ?x4)))) (uf_2 (uf_1 (- (uf_2 ?x4) (uf_2 uf_3)))))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_nat_arith_03.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,367 +0,0 @@
-#2 := false
-#9 := 0::int
-decl uf_2 :: (-> T1 int)
-decl uf_3 :: T1
-#22 := uf_3
-#23 := (uf_2 uf_3)
-#469 := (= #23 0::int)
-decl uf_1 :: (-> int T1)
-#251 := (uf_1 #23)
-#557 := (uf_2 #251)
-#558 := (= #557 0::int)
-#556 := (>= #23 0::int)
-#477 := (not #556)
-#144 := -1::int
-#348 := (>= #23 -1::int)
-#628 := (not #348)
-#21 := 1::int
-#24 := (+ 1::int #23)
-#25 := (uf_1 #24)
-#26 := (uf_2 #25)
-#632 := (* -1::int #26)
-#636 := (+ #23 #632)
-#633 := (= #636 -1::int)
-#471 := (not #633)
-#613 := (<= #636 -1::int)
-#527 := (not #613)
-#145 := (* -1::int #23)
-#146 := (+ #145 #26)
-#149 := (uf_1 #146)
-#152 := (uf_2 #149)
-#504 := (+ #632 #152)
-#505 := (+ #23 #504)
-#573 := (>= #505 0::int)
-#502 := (= #505 0::int)
-#599 := (<= #636 0::int)
-#526 := [hypothesis]: #613
-#491 := (or #527 #599)
-#515 := [th-lemma]: #491
-#516 := [unit-resolution #515 #526]: #599
-#587 := (not #599)
-#578 := (or #502 #587)
-#10 := (:var 0 int)
-#12 := (uf_1 #10)
-#673 := (pattern #12)
-#76 := (>= #10 0::int)
-#77 := (not #76)
-#13 := (uf_2 #12)
-#58 := (= #10 #13)
-#83 := (or #58 #77)
-#674 := (forall (vars (?x2 int)) (:pat #673) #83)
-#88 := (forall (vars (?x2 int)) #83)
-#677 := (iff #88 #674)
-#675 := (iff #83 #83)
-#676 := [refl]: #675
-#678 := [quant-intro #676]: #677
-#179 := (~ #88 #88)
-#191 := (~ #83 #83)
-#192 := [refl]: #191
-#177 := [nnf-pos #192]: #179
-#14 := (= #13 #10)
-#11 := (<= 0::int #10)
-#15 := (implies #11 #14)
-#16 := (forall (vars (?x2 int)) #15)
-#91 := (iff #16 #88)
-#65 := (not #11)
-#66 := (or #65 #58)
-#71 := (forall (vars (?x2 int)) #66)
-#89 := (iff #71 #88)
-#86 := (iff #66 #83)
-#80 := (or #77 #58)
-#84 := (iff #80 #83)
-#85 := [rewrite]: #84
-#81 := (iff #66 #80)
-#78 := (iff #65 #77)
-#74 := (iff #11 #76)
-#75 := [rewrite]: #74
-#79 := [monotonicity #75]: #78
-#82 := [monotonicity #79]: #81
-#87 := [trans #82 #85]: #86
-#90 := [quant-intro #87]: #89
-#72 := (iff #16 #71)
-#69 := (iff #15 #66)
-#62 := (implies #11 #58)
-#67 := (iff #62 #66)
-#68 := [rewrite]: #67
-#63 := (iff #15 #62)
-#60 := (iff #14 #58)
-#61 := [rewrite]: #60
-#64 := [monotonicity #61]: #63
-#70 := [trans #64 #68]: #69
-#73 := [quant-intro #70]: #72
-#92 := [trans #73 #90]: #91
-#57 := [asserted]: #16
-#93 := [mp #57 #92]: #88
-#193 := [mp~ #93 #177]: #88
-#679 := [mp #193 #678]: #674
-#644 := (not #674)
-#591 := (or #644 #502 #587)
-#498 := (>= #146 0::int)
-#500 := (not #498)
-#501 := (= #146 #152)
-#494 := (or #501 #500)
-#592 := (or #644 #494)
-#579 := (iff #592 #591)
-#593 := (or #644 #578)
-#584 := (iff #593 #591)
-#585 := [rewrite]: #584
-#582 := (iff #592 #593)
-#580 := (iff #494 #578)
-#589 := (iff #500 #587)
-#596 := (iff #498 #599)
-#600 := [rewrite]: #596
-#581 := [monotonicity #600]: #589
-#503 := (iff #501 #502)
-#506 := [rewrite]: #503
-#590 := [monotonicity #506 #581]: #580
-#583 := [monotonicity #590]: #582
-#586 := [trans #583 #585]: #579
-#588 := [quant-inst]: #592
-#570 := [mp #588 #586]: #591
-#511 := [unit-resolution #570 #679]: #578
-#517 := [unit-resolution #511 #516]: #502
-#485 := (not #502)
-#492 := (or #485 #573)
-#451 := [th-lemma]: #492
-#482 := [unit-resolution #451 #517]: #573
-#554 := (<= #152 0::int)
-#163 := (* -1::int #152)
-#138 := (uf_1 0::int)
-#141 := (uf_2 #138)
-#164 := (+ #141 #163)
-#162 := (>= #164 0::int)
-#30 := (- #26 #23)
-#31 := (uf_1 #30)
-#32 := (uf_2 #31)
-#27 := (* 0::int #26)
-#28 := (uf_1 #27)
-#29 := (uf_2 #28)
-#33 := (< #29 #32)
-#34 := (not #33)
-#174 := (iff #34 #162)
-#155 := (< #141 #152)
-#158 := (not #155)
-#172 := (iff #158 #162)
-#161 := (not #162)
-#167 := (not #161)
-#170 := (iff #167 #162)
-#171 := [rewrite]: #170
-#168 := (iff #158 #167)
-#165 := (iff #155 #161)
-#166 := [rewrite]: #165
-#169 := [monotonicity #166]: #168
-#173 := [trans #169 #171]: #172
-#159 := (iff #34 #158)
-#156 := (iff #33 #155)
-#153 := (= #32 #152)
-#150 := (= #31 #149)
-#147 := (= #30 #146)
-#148 := [rewrite]: #147
-#151 := [monotonicity #148]: #150
-#154 := [monotonicity #151]: #153
-#142 := (= #29 #141)
-#139 := (= #28 #138)
-#136 := (= #27 0::int)
-#137 := [rewrite]: #136
-#140 := [monotonicity #137]: #139
-#143 := [monotonicity #140]: #142
-#157 := [monotonicity #143 #154]: #156
-#160 := [monotonicity #157]: #159
-#175 := [trans #160 #173]: #174
-#135 := [asserted]: #34
-#176 := [mp #135 #175]: #162
-#651 := (<= #141 0::int)
-#662 := (= #141 0::int)
-#645 := (or #644 #662)
-#316 := (>= 0::int 0::int)
-#446 := (not #316)
-#328 := (= 0::int #141)
-#660 := (or #328 #446)
-#646 := (or #644 #660)
-#647 := (iff #646 #645)
-#648 := (iff #645 #645)
-#650 := [rewrite]: #648
-#642 := (iff #660 #662)
-#640 := (or #662 false)
-#305 := (iff #640 #662)
-#306 := [rewrite]: #305
-#303 := (iff #660 #640)
-#656 := (iff #446 false)
-#1 := true
-#654 := (not true)
-#655 := (iff #654 false)
-#315 := [rewrite]: #655
-#314 := (iff #446 #654)
-#658 := (iff #316 true)
-#664 := [rewrite]: #658
-#319 := [monotonicity #664]: #314
-#299 := [trans #319 #315]: #656
-#661 := (iff #328 #662)
-#663 := [rewrite]: #661
-#304 := [monotonicity #663 #299]: #303
-#643 := [trans #304 #306]: #642
-#285 := [monotonicity #643]: #647
-#290 := [trans #285 #650]: #647
-#641 := [quant-inst]: #646
-#291 := [mp #641 #290]: #645
-#484 := [unit-resolution #291 #679]: #662
-#486 := (not #662)
-#493 := (or #486 #651)
-#495 := [th-lemma]: #493
-#496 := [unit-resolution #495 #484]: #651
-#497 := (not #651)
-#507 := (or #554 #497 #161)
-#487 := [th-lemma]: #507
-#508 := [unit-resolution #487 #496 #176]: #554
-#463 := [th-lemma #508 #526 #482]: false
-#464 := [lemma #463]: #527
-#472 := (or #471 #613)
-#473 := [th-lemma]: #472
-#474 := [unit-resolution #473 #464]: #471
-#631 := (or #628 #633)
-#618 := (or #644 #628 #633)
-#634 := (>= #24 0::int)
-#635 := (not #634)
-#357 := (= #24 #26)
-#358 := (or #357 #635)
-#623 := (or #644 #358)
-#610 := (iff #623 #618)
-#619 := (or #644 #631)
-#467 := (iff #619 #618)
-#468 := [rewrite]: #467
-#625 := (iff #623 #619)
-#622 := (iff #358 #631)
-#626 := (or #633 #628)
-#620 := (iff #626 #631)
-#621 := [rewrite]: #620
-#630 := (iff #358 #626)
-#629 := (iff #635 #628)
-#349 := (iff #634 #348)
-#350 := [rewrite]: #349
-#344 := [monotonicity #350]: #629
-#637 := (iff #357 #633)
-#347 := [rewrite]: #637
-#627 := [monotonicity #347 #344]: #630
-#617 := [trans #627 #621]: #622
-#466 := [monotonicity #617]: #625
-#611 := [trans #466 #468]: #610
-#624 := [quant-inst]: #623
-#612 := [mp #624 #611]: #618
-#475 := [unit-resolution #612 #679]: #631
-#476 := [unit-resolution #475 #474]: #628
-#478 := (or #477 #348)
-#479 := [th-lemma]: #478
-#480 := [unit-resolution #479 #476]: #477
-#560 := (or #556 #558)
-#18 := (= #13 0::int)
-#124 := (or #18 #76)
-#680 := (forall (vars (?x3 int)) (:pat #673) #124)
-#129 := (forall (vars (?x3 int)) #124)
-#683 := (iff #129 #680)
-#681 := (iff #124 #124)
-#682 := [refl]: #681
-#684 := [quant-intro #682]: #683
-#180 := (~ #129 #129)
-#194 := (~ #124 #124)
-#195 := [refl]: #194
-#181 := [nnf-pos #195]: #180
-#17 := (< #10 0::int)
-#19 := (implies #17 #18)
-#20 := (forall (vars (?x3 int)) #19)
-#132 := (iff #20 #129)
-#95 := (= 0::int #13)
-#101 := (not #17)
-#102 := (or #101 #95)
-#107 := (forall (vars (?x3 int)) #102)
-#130 := (iff #107 #129)
-#127 := (iff #102 #124)
-#121 := (or #76 #18)
-#125 := (iff #121 #124)
-#126 := [rewrite]: #125
-#122 := (iff #102 #121)
-#119 := (iff #95 #18)
-#120 := [rewrite]: #119
-#117 := (iff #101 #76)
-#112 := (not #77)
-#115 := (iff #112 #76)
-#116 := [rewrite]: #115
-#113 := (iff #101 #112)
-#110 := (iff #17 #77)
-#111 := [rewrite]: #110
-#114 := [monotonicity #111]: #113
-#118 := [trans #114 #116]: #117
-#123 := [monotonicity #118 #120]: #122
-#128 := [trans #123 #126]: #127
-#131 := [quant-intro #128]: #130
-#108 := (iff #20 #107)
-#105 := (iff #19 #102)
-#98 := (implies #17 #95)
-#103 := (iff #98 #102)
-#104 := [rewrite]: #103
-#99 := (iff #19 #98)
-#96 := (iff #18 #95)
-#97 := [rewrite]: #96
-#100 := [monotonicity #97]: #99
-#106 := [trans #100 #104]: #105
-#109 := [quant-intro #106]: #108
-#133 := [trans #109 #131]: #132
-#94 := [asserted]: #20
-#134 := [mp #94 #133]: #129
-#196 := [mp~ #134 #181]: #129
-#685 := [mp #196 #684]: #680
-#604 := (not #680)
-#562 := (or #604 #556 #558)
-#559 := (or #558 #556)
-#540 := (or #604 #559)
-#542 := (iff #540 #562)
-#543 := (or #604 #560)
-#546 := (iff #543 #562)
-#547 := [rewrite]: #546
-#544 := (iff #540 #543)
-#561 := (iff #559 #560)
-#551 := [rewrite]: #561
-#545 := [monotonicity #551]: #544
-#548 := [trans #545 #547]: #542
-#541 := [quant-inst]: #540
-#534 := [mp #541 #548]: #562
-#465 := [unit-resolution #534 #685]: #560
-#481 := [unit-resolution #465 #480]: #558
-#443 := (= #23 #557)
-#337 := (= uf_3 #251)
-#4 := (:var 0 T1)
-#5 := (uf_2 #4)
-#665 := (pattern #5)
-#6 := (uf_1 #5)
-#51 := (= #4 #6)
-#666 := (forall (vars (?x1 T1)) (:pat #665) #51)
-#54 := (forall (vars (?x1 T1)) #51)
-#667 := (iff #54 #666)
-#669 := (iff #666 #666)
-#670 := [rewrite]: #669
-#668 := [rewrite]: #667
-#671 := [trans #668 #670]: #667
-#188 := (~ #54 #54)
-#186 := (~ #51 #51)
-#187 := [refl]: #186
-#189 := [nnf-pos #187]: #188
-#7 := (= #6 #4)
-#8 := (forall (vars (?x1 T1)) #7)
-#55 := (iff #8 #54)
-#52 := (iff #7 #51)
-#53 := [rewrite]: #52
-#56 := [quant-intro #53]: #55
-#50 := [asserted]: #8
-#59 := [mp #50 #56]: #54
-#190 := [mp~ #59 #189]: #54
-#672 := [mp #190 #671]: #666
-#252 := (not #666)
-#342 := (or #252 #337)
-#339 := [quant-inst]: #342
-#442 := [unit-resolution #339 #672]: #337
-#450 := [monotonicity #442]: #443
-#452 := [trans #450 #481]: #469
-#453 := (not #469)
-#454 := (or #453 #556)
-#456 := [th-lemma]: #454
-[unit-resolution #456 #480 #452]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_nat_arith_04 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,13 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T1 T2)
-:extrafuns (
- (uf_1 Int T1)
- (uf_2 T1 Int)
- (uf_3 T1)
- )
-:assumption (forall (?x1 T1) (= (uf_1 (uf_2 ?x1)) ?x1))
-:assumption (forall (?x2 Int) (implies (<= 0 ?x2) (= (uf_2 (uf_1 ?x2)) ?x2)))
-:assumption (forall (?x3 Int) (implies (< ?x3 0) (= (uf_2 (uf_1 ?x3)) 0)))
-:assumption (not (let (?x4 (uf_1 (+ 1 (uf_2 uf_3)))) (flet ($x5 (if_then_else (< 0 (uf_2 ?x4)) true false)) (or (iff $x5 (= (uf_1 (- (uf_2 ?x4) 1)) uf_3)) $x5))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_nat_arith_04.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,302 +0,0 @@
-#2 := false
-#9 := 0::int
-decl uf_2 :: (-> T1 int)
-decl uf_1 :: (-> int T1)
-decl uf_3 :: T1
-#22 := uf_3
-#23 := (uf_2 uf_3)
-#21 := 1::int
-#24 := (+ 1::int #23)
-#25 := (uf_1 #24)
-#26 := (uf_2 #25)
-#138 := -1::int
-#139 := (+ -1::int #26)
-#142 := (uf_1 #139)
-#289 := (uf_2 #142)
-#383 := (* -1::int #289)
-#542 := (+ #23 #383)
-#544 := (>= #542 0::int)
-#541 := (= #23 #289)
-#148 := (= uf_3 #142)
-#167 := (<= #26 0::int)
-#168 := (not #167)
-#174 := (iff #148 #168)
-#189 := (not #174)
-#220 := (iff #189 #148)
-#210 := (not #148)
-#215 := (not #210)
-#218 := (iff #215 #148)
-#219 := [rewrite]: #218
-#216 := (iff #189 #215)
-#213 := (iff #174 #210)
-#207 := (iff #148 false)
-#211 := (iff #207 #210)
-#212 := [rewrite]: #211
-#208 := (iff #174 #207)
-#205 := (iff #168 false)
-#1 := true
-#200 := (not true)
-#203 := (iff #200 false)
-#204 := [rewrite]: #203
-#201 := (iff #168 #200)
-#198 := (iff #167 true)
-#179 := (or #168 #174)
-#182 := (not #179)
-#27 := (< 0::int #26)
-#28 := (ite #27 true false)
-#29 := (- #26 1::int)
-#30 := (uf_1 #29)
-#31 := (= #30 uf_3)
-#32 := (iff #28 #31)
-#33 := (or #32 #28)
-#34 := (not #33)
-#185 := (iff #34 #182)
-#153 := (iff #27 #148)
-#159 := (or #27 #153)
-#164 := (not #159)
-#183 := (iff #164 #182)
-#180 := (iff #159 #179)
-#177 := (iff #153 #174)
-#171 := (iff #168 #148)
-#175 := (iff #171 #174)
-#176 := [rewrite]: #175
-#172 := (iff #153 #171)
-#169 := (iff #27 #168)
-#170 := [rewrite]: #169
-#173 := [monotonicity #170]: #172
-#178 := [trans #173 #176]: #177
-#181 := [monotonicity #170 #178]: #180
-#184 := [monotonicity #181]: #183
-#165 := (iff #34 #164)
-#162 := (iff #33 #159)
-#156 := (or #153 #27)
-#160 := (iff #156 #159)
-#161 := [rewrite]: #160
-#157 := (iff #33 #156)
-#136 := (iff #28 #27)
-#137 := [rewrite]: #136
-#154 := (iff #32 #153)
-#151 := (iff #31 #148)
-#145 := (= #142 uf_3)
-#149 := (iff #145 #148)
-#150 := [rewrite]: #149
-#146 := (iff #31 #145)
-#143 := (= #30 #142)
-#140 := (= #29 #139)
-#141 := [rewrite]: #140
-#144 := [monotonicity #141]: #143
-#147 := [monotonicity #144]: #146
-#152 := [trans #147 #150]: #151
-#155 := [monotonicity #137 #152]: #154
-#158 := [monotonicity #155 #137]: #157
-#163 := [trans #158 #161]: #162
-#166 := [monotonicity #163]: #165
-#186 := [trans #166 #184]: #185
-#135 := [asserted]: #34
-#187 := [mp #135 #186]: #182
-#188 := [not-or-elim #187]: #167
-#199 := [iff-true #188]: #198
-#202 := [monotonicity #199]: #201
-#206 := [trans #202 #204]: #205
-#209 := [monotonicity #206]: #208
-#214 := [trans #209 #212]: #213
-#217 := [monotonicity #214]: #216
-#221 := [trans #217 #219]: #220
-#190 := [not-or-elim #187]: #189
-#222 := [mp #190 #221]: #148
-#624 := [monotonicity #222]: #541
-#618 := (not #541)
-#625 := (or #618 #544)
-#609 := [th-lemma]: #625
-#610 := [unit-resolution #609 #624]: #544
-#698 := (* -1::int #26)
-#355 := (+ #23 #698)
-#324 := (<= #355 -1::int)
-#485 := (= #355 -1::int)
-#367 := (>= #23 -1::int)
-#533 := (>= #289 0::int)
-#643 := (= #289 0::int)
-#659 := (>= #26 1::int)
-#656 := (not #659)
-#612 := (or #656 #168)
-#613 := [th-lemma]: #612
-#614 := [unit-resolution #613 #188]: #656
-#10 := (:var 0 int)
-#12 := (uf_1 #10)
-#712 := (pattern #12)
-#76 := (>= #10 0::int)
-#13 := (uf_2 #12)
-#18 := (= #13 0::int)
-#124 := (or #18 #76)
-#719 := (forall (vars (?x3 int)) (:pat #712) #124)
-#129 := (forall (vars (?x3 int)) #124)
-#722 := (iff #129 #719)
-#720 := (iff #124 #124)
-#721 := [refl]: #720
-#723 := [quant-intro #721]: #722
-#229 := (~ #129 #129)
-#227 := (~ #124 #124)
-#228 := [refl]: #227
-#230 := [nnf-pos #228]: #229
-#17 := (< #10 0::int)
-#19 := (implies #17 #18)
-#20 := (forall (vars (?x3 int)) #19)
-#132 := (iff #20 #129)
-#95 := (= 0::int #13)
-#101 := (not #17)
-#102 := (or #101 #95)
-#107 := (forall (vars (?x3 int)) #102)
-#130 := (iff #107 #129)
-#127 := (iff #102 #124)
-#121 := (or #76 #18)
-#125 := (iff #121 #124)
-#126 := [rewrite]: #125
-#122 := (iff #102 #121)
-#119 := (iff #95 #18)
-#120 := [rewrite]: #119
-#117 := (iff #101 #76)
-#77 := (not #76)
-#112 := (not #77)
-#115 := (iff #112 #76)
-#116 := [rewrite]: #115
-#113 := (iff #101 #112)
-#110 := (iff #17 #77)
-#111 := [rewrite]: #110
-#114 := [monotonicity #111]: #113
-#118 := [trans #114 #116]: #117
-#123 := [monotonicity #118 #120]: #122
-#128 := [trans #123 #126]: #127
-#131 := [quant-intro #128]: #130
-#108 := (iff #20 #107)
-#105 := (iff #19 #102)
-#98 := (implies #17 #95)
-#103 := (iff #98 #102)
-#104 := [rewrite]: #103
-#99 := (iff #19 #98)
-#96 := (iff #18 #95)
-#97 := [rewrite]: #96
-#100 := [monotonicity #97]: #99
-#106 := [trans #100 #104]: #105
-#109 := [quant-intro #106]: #108
-#133 := [trans #109 #131]: #132
-#94 := [asserted]: #20
-#134 := [mp #94 #133]: #129
-#231 := [mp~ #134 #230]: #129
-#724 := [mp #231 #723]: #719
-#402 := (not #719)
-#528 := (or #402 #643 #659)
-#388 := (>= #139 0::int)
-#644 := (or #643 #388)
-#529 := (or #402 #644)
-#522 := (iff #529 #528)
-#642 := (or #643 #659)
-#636 := (or #402 #642)
-#634 := (iff #636 #528)
-#637 := [rewrite]: #634
-#538 := (iff #529 #636)
-#645 := (iff #644 #642)
-#660 := (iff #388 #659)
-#661 := [rewrite]: #660
-#527 := [monotonicity #661]: #645
-#633 := [monotonicity #527]: #538
-#537 := [trans #633 #637]: #522
-#488 := [quant-inst]: #529
-#539 := [mp #488 #537]: #528
-#615 := [unit-resolution #539 #724 #614]: #643
-#611 := (not #643)
-#616 := (or #611 #533)
-#602 := [th-lemma]: #616
-#603 := [unit-resolution #602 #615]: #533
-#606 := (not #544)
-#605 := (not #533)
-#607 := (or #367 #605 #606)
-#604 := [th-lemma]: #607
-#608 := [unit-resolution #604 #603 #610]: #367
-#701 := (not #367)
-#358 := (or #701 #485)
-#58 := (= #10 #13)
-#83 := (or #58 #77)
-#713 := (forall (vars (?x2 int)) (:pat #712) #83)
-#88 := (forall (vars (?x2 int)) #83)
-#716 := (iff #88 #713)
-#714 := (iff #83 #83)
-#715 := [refl]: #714
-#717 := [quant-intro #715]: #716
-#191 := (~ #88 #88)
-#195 := (~ #83 #83)
-#193 := [refl]: #195
-#225 := [nnf-pos #193]: #191
-#14 := (= #13 #10)
-#11 := (<= 0::int #10)
-#15 := (implies #11 #14)
-#16 := (forall (vars (?x2 int)) #15)
-#91 := (iff #16 #88)
-#65 := (not #11)
-#66 := (or #65 #58)
-#71 := (forall (vars (?x2 int)) #66)
-#89 := (iff #71 #88)
-#86 := (iff #66 #83)
-#80 := (or #77 #58)
-#84 := (iff #80 #83)
-#85 := [rewrite]: #84
-#81 := (iff #66 #80)
-#78 := (iff #65 #77)
-#74 := (iff #11 #76)
-#75 := [rewrite]: #74
-#79 := [monotonicity #75]: #78
-#82 := [monotonicity #79]: #81
-#87 := [trans #82 #85]: #86
-#90 := [quant-intro #87]: #89
-#72 := (iff #16 #71)
-#69 := (iff #15 #66)
-#62 := (implies #11 #58)
-#67 := (iff #62 #66)
-#68 := [rewrite]: #67
-#63 := (iff #15 #62)
-#60 := (iff #14 #58)
-#61 := [rewrite]: #60
-#64 := [monotonicity #61]: #63
-#70 := [trans #64 #68]: #69
-#73 := [quant-intro #70]: #72
-#92 := [trans #73 #90]: #91
-#57 := [asserted]: #16
-#93 := [mp #57 #92]: #88
-#226 := [mp~ #93 #225]: #88
-#718 := [mp #226 #717]: #713
-#679 := (not #713)
-#342 := (or #679 #701 #485)
-#380 := (>= #24 0::int)
-#381 := (not #380)
-#361 := (= #24 #26)
-#696 := (or #361 #381)
-#343 := (or #679 #696)
-#685 := (iff #343 #342)
-#345 := (or #679 #358)
-#683 := (iff #345 #342)
-#684 := [rewrite]: #683
-#681 := (iff #343 #345)
-#695 := (iff #696 #358)
-#703 := (or #485 #701)
-#694 := (iff #703 #358)
-#354 := [rewrite]: #694
-#693 := (iff #696 #703)
-#702 := (iff #381 #701)
-#699 := (iff #380 #367)
-#700 := [rewrite]: #699
-#697 := [monotonicity #700]: #702
-#692 := (iff #361 #485)
-#366 := [rewrite]: #692
-#353 := [monotonicity #366 #697]: #693
-#338 := [trans #353 #354]: #695
-#682 := [monotonicity #338]: #681
-#680 := [trans #682 #684]: #685
-#344 := [quant-inst]: #343
-#686 := [mp #344 #680]: #342
-#588 := [unit-resolution #686 #718]: #358
-#589 := [unit-resolution #588 #608]: #485
-#591 := (not #485)
-#592 := (or #591 #324)
-#593 := [th-lemma]: #592
-#594 := [unit-resolution #593 #589]: #324
-[th-lemma #603 #188 #594 #610]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_nat_arith_05 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,13 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T1)
-:extrafuns (
- (uf_1 Int T1)
- (uf_2 T1 Int)
- (uf_3 T1)
- )
-:assumption (forall (?x1 T1) (= (uf_1 (uf_2 ?x1)) ?x1))
-:assumption (forall (?x2 Int) (implies (<= 0 ?x2) (= (uf_2 (uf_1 ?x2)) ?x2)))
-:assumption (forall (?x3 Int) (implies (< ?x3 0) (= (uf_2 (uf_1 ?x3)) 0)))
-:assumption (not (distinct (uf_1 (+ (uf_2 uf_3) 1)) (uf_1 (+ (uf_2 (uf_1 (* (uf_2 uf_3) 2))) 3)) (uf_1 (- (uf_2 uf_3) (uf_2 uf_3)))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_nat_arith_05.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,525 +0,0 @@
-#2 := false
-decl uf_2 :: (-> T1 int)
-decl uf_1 :: (-> int T1)
-decl uf_3 :: T1
-#21 := uf_3
-#22 := (uf_2 uf_3)
-#23 := 1::int
-#138 := (+ 1::int #22)
-#141 := (uf_1 #138)
-#656 := (uf_2 #141)
-#26 := 2::int
-#144 := (* 2::int #22)
-#147 := (uf_1 #144)
-#150 := (uf_2 #147)
-#30 := 3::int
-#156 := (+ 3::int #150)
-#161 := (uf_1 #156)
-#494 := (uf_2 #161)
-#288 := (= #494 #656)
-#266 := (= #161 #141)
-#249 := (= #141 #161)
-#9 := 0::int
-#166 := (uf_1 0::int)
-#251 := (= #161 #166)
-#337 := (not #251)
-#567 := (= #494 0::int)
-#543 := (uf_2 #166)
-#547 := (= #543 0::int)
-#10 := (:var 0 int)
-#12 := (uf_1 #10)
-#673 := (pattern #12)
-#78 := (>= #10 0::int)
-#79 := (not #78)
-#13 := (uf_2 #12)
-#60 := (= #10 #13)
-#85 := (or #60 #79)
-#674 := (forall (vars (?x2 int)) (:pat #673) #85)
-#90 := (forall (vars (?x2 int)) #85)
-#677 := (iff #90 #674)
-#675 := (iff #85 #85)
-#676 := [refl]: #675
-#678 := [quant-intro #676]: #677
-#178 := (~ #90 #90)
-#190 := (~ #85 #85)
-#191 := [refl]: #190
-#175 := [nnf-pos #191]: #178
-#14 := (= #13 #10)
-#11 := (<= 0::int #10)
-#15 := (implies #11 #14)
-#16 := (forall (vars (?x2 int)) #15)
-#93 := (iff #16 #90)
-#67 := (not #11)
-#68 := (or #67 #60)
-#73 := (forall (vars (?x2 int)) #68)
-#91 := (iff #73 #90)
-#88 := (iff #68 #85)
-#82 := (or #79 #60)
-#86 := (iff #82 #85)
-#87 := [rewrite]: #86
-#83 := (iff #68 #82)
-#80 := (iff #67 #79)
-#76 := (iff #11 #78)
-#77 := [rewrite]: #76
-#81 := [monotonicity #77]: #80
-#84 := [monotonicity #81]: #83
-#89 := [trans #84 #87]: #88
-#92 := [quant-intro #89]: #91
-#74 := (iff #16 #73)
-#71 := (iff #15 #68)
-#64 := (implies #11 #60)
-#69 := (iff #64 #68)
-#70 := [rewrite]: #69
-#65 := (iff #15 #64)
-#62 := (iff #14 #60)
-#63 := [rewrite]: #62
-#66 := [monotonicity #63]: #65
-#72 := [trans #66 #70]: #71
-#75 := [quant-intro #72]: #74
-#94 := [trans #75 #92]: #93
-#59 := [asserted]: #16
-#95 := [mp #59 #94]: #90
-#192 := [mp~ #95 #175]: #90
-#679 := [mp #192 #678]: #674
-#290 := (not #674)
-#519 := (or #290 #547)
-#540 := (>= 0::int 0::int)
-#541 := (not #540)
-#544 := (= 0::int #543)
-#545 := (or #544 #541)
-#520 := (or #290 #545)
-#521 := (iff #520 #519)
-#523 := (iff #519 #519)
-#526 := [rewrite]: #523
-#407 := (iff #545 #547)
-#533 := (or #547 false)
-#513 := (iff #533 #547)
-#514 := [rewrite]: #513
-#539 := (iff #545 #533)
-#537 := (iff #541 false)
-#1 := true
-#530 := (not true)
-#535 := (iff #530 false)
-#536 := [rewrite]: #535
-#531 := (iff #541 #530)
-#548 := (iff #540 true)
-#534 := [rewrite]: #548
-#532 := [monotonicity #534]: #531
-#538 := [trans #532 #536]: #537
-#546 := (iff #544 #547)
-#542 := [rewrite]: #546
-#512 := [monotonicity #542 #538]: #539
-#518 := [trans #512 #514]: #407
-#522 := [monotonicity #518]: #521
-#527 := [trans #522 #526]: #521
-#525 := [quant-inst]: #520
-#528 := [mp #525 #527]: #519
-#316 := [unit-resolution #528 #679]: #547
-#286 := (= #494 #543)
-#287 := [hypothesis]: #251
-#292 := [monotonicity #287]: #286
-#267 := [trans #292 #316]: #567
-#296 := (not #567)
-#551 := (<= #494 0::int)
-#300 := (not #551)
-#501 := (>= #150 0::int)
-#622 := (>= #144 0::int)
-#302 := -1::int
-#303 := (* -1::int #656)
-#304 := (+ #22 #303)
-#635 := (>= #304 -1::int)
-#305 := (= #304 -1::int)
-#644 := (>= #22 -1::int)
-#511 := (>= #22 0::int)
-#487 := (= #22 0::int)
-#660 := (uf_1 #22)
-#517 := (uf_2 #660)
-#485 := (= #517 0::int)
-#389 := (not #511)
-#390 := [hypothesis]: #389
-#492 := (or #485 #511)
-#18 := (= #13 0::int)
-#126 := (or #18 #78)
-#680 := (forall (vars (?x3 int)) (:pat #673) #126)
-#131 := (forall (vars (?x3 int)) #126)
-#683 := (iff #131 #680)
-#681 := (iff #126 #126)
-#682 := [refl]: #681
-#684 := [quant-intro #682]: #683
-#179 := (~ #131 #131)
-#193 := (~ #126 #126)
-#194 := [refl]: #193
-#180 := [nnf-pos #194]: #179
-#17 := (< #10 0::int)
-#19 := (implies #17 #18)
-#20 := (forall (vars (?x3 int)) #19)
-#134 := (iff #20 #131)
-#97 := (= 0::int #13)
-#103 := (not #17)
-#104 := (or #103 #97)
-#109 := (forall (vars (?x3 int)) #104)
-#132 := (iff #109 #131)
-#129 := (iff #104 #126)
-#123 := (or #78 #18)
-#127 := (iff #123 #126)
-#128 := [rewrite]: #127
-#124 := (iff #104 #123)
-#121 := (iff #97 #18)
-#122 := [rewrite]: #121
-#119 := (iff #103 #78)
-#114 := (not #79)
-#117 := (iff #114 #78)
-#118 := [rewrite]: #117
-#115 := (iff #103 #114)
-#112 := (iff #17 #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 #20 #109)
-#107 := (iff #19 #104)
-#100 := (implies #17 #97)
-#105 := (iff #100 #104)
-#106 := [rewrite]: #105
-#101 := (iff #19 #100)
-#98 := (iff #18 #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]: #20
-#136 := [mp #96 #135]: #131
-#195 := [mp~ #136 #180]: #131
-#685 := [mp #195 #684]: #680
-#637 := (not #680)
-#484 := (or #637 #485 #511)
-#486 := (or #637 #492)
-#495 := (iff #486 #484)
-#496 := [rewrite]: #495
-#493 := [quant-inst]: #486
-#497 := [mp #493 #496]: #484
-#391 := [unit-resolution #497 #685]: #492
-#392 := [unit-resolution #391 #390]: #485
-#394 := (= #22 #517)
-#661 := (= uf_3 #660)
-#4 := (:var 0 T1)
-#5 := (uf_2 #4)
-#665 := (pattern #5)
-#6 := (uf_1 #5)
-#53 := (= #4 #6)
-#666 := (forall (vars (?x1 T1)) (:pat #665) #53)
-#56 := (forall (vars (?x1 T1)) #53)
-#667 := (iff #56 #666)
-#669 := (iff #666 #666)
-#670 := [rewrite]: #669
-#668 := [rewrite]: #667
-#671 := [trans #668 #670]: #667
-#187 := (~ #56 #56)
-#185 := (~ #53 #53)
-#186 := [refl]: #185
-#188 := [nnf-pos #186]: #187
-#7 := (= #6 #4)
-#8 := (forall (vars (?x1 T1)) #7)
-#57 := (iff #8 #56)
-#54 := (iff #7 #53)
-#55 := [rewrite]: #54
-#58 := [quant-intro #55]: #57
-#52 := [asserted]: #8
-#61 := [mp #52 #58]: #56
-#189 := [mp~ #61 #188]: #56
-#672 := [mp #189 #671]: #666
-#658 := (not #666)
-#664 := (or #658 #661)
-#654 := [quant-inst]: #664
-#393 := [unit-resolution #654 #672]: #661
-#395 := [monotonicity #393]: #394
-#396 := [trans #395 #392]: #487
-#397 := (not #487)
-#398 := (or #397 #511)
-#399 := [th-lemma]: #398
-#388 := [unit-resolution #399 #390 #396]: false
-#400 := [lemma #388]: #511
-#366 := (or #389 #644)
-#367 := [th-lemma]: #366
-#352 := [unit-resolution #367 #400]: #644
-#641 := (not #644)
-#648 := (or #305 #641)
-#651 := (or #290 #305 #641)
-#313 := (>= #138 0::int)
-#318 := (not #313)
-#298 := (= #138 #656)
-#640 := (or #298 #318)
-#649 := (or #290 #640)
-#363 := (iff #649 #651)
-#638 := (or #290 #648)
-#361 := (iff #638 #651)
-#362 := [rewrite]: #361
-#639 := (iff #649 #638)
-#650 := (iff #640 #648)
-#647 := (iff #318 #641)
-#645 := (iff #313 #644)
-#646 := [rewrite]: #645
-#284 := [monotonicity #646]: #647
-#642 := (iff #298 #305)
-#643 := [rewrite]: #642
-#289 := [monotonicity #643 #284]: #650
-#346 := [monotonicity #289]: #639
-#364 := [trans #346 #362]: #363
-#652 := [quant-inst]: #649
-#257 := [mp #652 #364]: #651
-#424 := [unit-resolution #257 #679]: #648
-#353 := [unit-resolution #424 #352]: #305
-#439 := (not #305)
-#281 := (or #439 #635)
-#440 := [th-lemma]: #281
-#330 := [unit-resolution #440 #353]: #635
-#620 := (<= #656 0::int)
-#441 := (not #620)
-#634 := (<= #304 -1::int)
-#344 := (or #439 #634)
-#354 := [th-lemma]: #344
-#355 := [unit-resolution #354 #353]: #634
-#345 := (not #634)
-#356 := (or #441 #389 #345)
-#322 := [th-lemma]: #356
-#324 := [unit-resolution #322 #355 #400]: #441
-#432 := (not #635)
-#331 := (or #622 #432 #620)
-#319 := [th-lemma]: #331
-#320 := [unit-resolution #319 #324 #330]: #622
-#624 := (* -1::int #150)
-#619 := (+ #144 #624)
-#606 := (<= #619 0::int)
-#625 := (= #619 0::int)
-#617 := (not #622)
-#612 := (or #617 #625)
-#615 := (or #290 #617 #625)
-#618 := (= #144 #150)
-#623 := (or #618 #617)
-#609 := (or #290 #623)
-#604 := (iff #609 #615)
-#445 := (or #290 #612)
-#601 := (iff #445 #615)
-#602 := [rewrite]: #601
-#447 := (iff #609 #445)
-#608 := (iff #623 #612)
-#468 := (or #625 #617)
-#613 := (iff #468 #612)
-#607 := [rewrite]: #613
-#610 := (iff #623 #468)
-#466 := (iff #618 #625)
-#467 := [rewrite]: #466
-#611 := [monotonicity #467]: #610
-#614 := [trans #611 #607]: #608
-#448 := [monotonicity #614]: #447
-#605 := [trans #448 #602]: #604
-#616 := [quant-inst]: #609
-#603 := [mp #616 #605]: #615
-#480 := [unit-resolution #603 #679]: #612
-#299 := [unit-resolution #480 #320]: #625
-#406 := (not #625)
-#408 := (or #406 #606)
-#409 := [th-lemma]: #408
-#301 := [unit-resolution #409 #299]: #606
-#413 := (not #606)
-#306 := (or #501 #413 #617)
-#307 := [th-lemma]: #306
-#308 := [unit-resolution #307 #301 #320]: #501
-#506 := -3::int
-#504 := (* -1::int #494)
-#505 := (+ #150 #504)
-#564 := (<= #505 -3::int)
-#599 := (= #505 -3::int)
-#587 := (>= #150 -3::int)
-#417 := (or #587 #413 #617)
-#410 := [th-lemma]: #417
-#309 := [unit-resolution #410 #301 #320]: #587
-#578 := (not #587)
-#593 := (or #578 #599)
-#579 := (or #290 #578 #599)
-#449 := (>= #156 0::int)
-#597 := (not #449)
-#502 := (= #156 #494)
-#503 := (or #502 #597)
-#586 := (or #290 #503)
-#572 := (iff #586 #579)
-#571 := (or #290 #593)
-#575 := (iff #571 #579)
-#576 := [rewrite]: #575
-#573 := (iff #586 #571)
-#584 := (iff #503 #593)
-#591 := (or #599 #578)
-#582 := (iff #591 #593)
-#583 := [rewrite]: #582
-#592 := (iff #503 #591)
-#580 := (iff #597 #578)
-#589 := (iff #449 #587)
-#581 := [rewrite]: #589
-#590 := [monotonicity #581]: #580
-#596 := (iff #502 #599)
-#600 := [rewrite]: #596
-#588 := [monotonicity #600 #590]: #592
-#585 := [trans #588 #583]: #584
-#574 := [monotonicity #585]: #573
-#577 := [trans #574 #576]: #572
-#570 := [quant-inst]: #586
-#563 := [mp #570 #577]: #579
-#458 := [unit-resolution #563 #679]: #593
-#310 := [unit-resolution #458 #309]: #599
-#460 := (not #599)
-#461 := (or #460 #564)
-#444 := [th-lemma]: #461
-#311 := [unit-resolution #444 #310]: #564
-#434 := (not #564)
-#453 := (not #501)
-#312 := (or #300 #453 #434)
-#293 := [th-lemma]: #312
-#295 := [unit-resolution #293 #311 #308]: #300
-#294 := (or #296 #551)
-#297 := [th-lemma]: #294
-#285 := [unit-resolution #297 #295]: #296
-#271 := [unit-resolution #285 #267]: false
-#272 := [lemma #271]: #337
-#282 := (or #249 #251)
-#250 := (= #141 #166)
-#336 := (not #250)
-#357 := (= #656 0::int)
-#332 := (= #656 #543)
-#329 := [hypothesis]: #250
-#333 := [monotonicity #329]: #332
-#323 := [trans #333 #316]: #357
-#429 := (not #357)
-#430 := (or #429 #620)
-#428 := [th-lemma]: #430
-#325 := [unit-resolution #428 #324]: #429
-#334 := [unit-resolution #325 #323]: false
-#317 := [lemma #334]: #336
-#279 := (or #249 #250 #251)
-#335 := (not #249)
-#328 := (and #335 #336 #337)
-#339 := (not #328)
-#169 := (distinct #141 #161 #166)
-#172 := (not #169)
-#33 := (- #22 #22)
-#34 := (uf_1 #33)
-#27 := (* #22 2::int)
-#28 := (uf_1 #27)
-#29 := (uf_2 #28)
-#31 := (+ #29 3::int)
-#32 := (uf_1 #31)
-#24 := (+ #22 1::int)
-#25 := (uf_1 #24)
-#35 := (distinct #25 #32 #34)
-#36 := (not #35)
-#173 := (iff #36 #172)
-#170 := (iff #35 #169)
-#167 := (= #34 #166)
-#164 := (= #33 0::int)
-#165 := [rewrite]: #164
-#168 := [monotonicity #165]: #167
-#162 := (= #32 #161)
-#159 := (= #31 #156)
-#153 := (+ #150 3::int)
-#157 := (= #153 #156)
-#158 := [rewrite]: #157
-#154 := (= #31 #153)
-#151 := (= #29 #150)
-#148 := (= #28 #147)
-#145 := (= #27 #144)
-#146 := [rewrite]: #145
-#149 := [monotonicity #146]: #148
-#152 := [monotonicity #149]: #151
-#155 := [monotonicity #152]: #154
-#160 := [trans #155 #158]: #159
-#163 := [monotonicity #160]: #162
-#142 := (= #25 #141)
-#139 := (= #24 #138)
-#140 := [rewrite]: #139
-#143 := [monotonicity #140]: #142
-#171 := [monotonicity #143 #163 #168]: #170
-#174 := [monotonicity #171]: #173
-#137 := [asserted]: #36
-#177 := [mp #137 #174]: #172
-#326 := (or #169 #339)
-#327 := [def-axiom]: #326
-#277 := [unit-resolution #327 #177]: #339
-#659 := (or #328 #249 #250 #251)
-#315 := [def-axiom]: #659
-#280 := [unit-resolution #315 #277]: #279
-#278 := [unit-resolution #280 #317]: #282
-#283 := [unit-resolution #278 #272]: #249
-#269 := [symm #283]: #266
-#270 := [monotonicity #269]: #288
-#508 := (+ #494 #303)
-#473 := (<= #508 0::int)
-#433 := (not #473)
-#477 := [hypothesis]: #473
-#421 := (or #622 #433)
-#489 := (= #150 0::int)
-#478 := [hypothesis]: #617
-#490 := (or #489 #622)
-#499 := (or #637 #489 #622)
-#594 := (or #637 #490)
-#598 := (iff #594 #499)
-#483 := [rewrite]: #598
-#595 := [quant-inst]: #594
-#498 := [mp #595 #483]: #499
-#465 := [unit-resolution #498 #685]: #490
-#481 := [unit-resolution #465 #478]: #489
-#442 := (not #489)
-#443 := (or #442 #501)
-#450 := [th-lemma]: #443
-#452 := [unit-resolution #450 #481]: #501
-#454 := (or #453 #587)
-#456 := [th-lemma]: #454
-#457 := [unit-resolution #456 #452]: #587
-#459 := [unit-resolution #458 #457]: #599
-#462 := [unit-resolution #444 #459]: #564
-#435 := (or #432 #622 #433 #453 #434)
-#437 := [th-lemma]: #435
-#438 := [unit-resolution #437 #478 #452 #462 #477]: #432
-#436 := [unit-resolution #440 #438]: #439
-#420 := (or #441 #433 #453 #434)
-#423 := [th-lemma]: #420
-#427 := [unit-resolution #423 #452 #462 #477]: #441
-#431 := [unit-resolution #428 #427]: #429
-#632 := (or #357 #644)
-#347 := (or #637 #357 #644)
-#358 := (or #357 #313)
-#348 := (or #637 #358)
-#630 := (iff #348 #347)
-#350 := (or #637 #632)
-#343 := (iff #350 #347)
-#626 := [rewrite]: #343
-#628 := (iff #348 #350)
-#636 := (iff #358 #632)
-#633 := [monotonicity #646]: #636
-#629 := [monotonicity #633]: #628
-#627 := [trans #629 #626]: #630
-#349 := [quant-inst]: #348
-#631 := [mp #349 #627]: #347
-#419 := [unit-resolution #631 #685]: #632
-#422 := [unit-resolution #419 #431]: #644
-#425 := [unit-resolution #424 #422 #436]: false
-#426 := [lemma #425]: #421
-#479 := [unit-resolution #426 #477]: #622
-#416 := [unit-resolution #480 #479]: #625
-#412 := [unit-resolution #409 #416]: #606
-#418 := [unit-resolution #410 #412 #479]: #587
-#411 := [unit-resolution #458 #418]: #599
-#414 := [unit-resolution #444 #411]: #564
-#415 := (or #644 #617)
-#401 := [th-lemma]: #415
-#403 := [unit-resolution #401 #479]: #644
-#404 := [unit-resolution #424 #403]: #305
-#402 := [unit-resolution #440 #404]: #635
-#405 := [th-lemma #418 #402 #477 #414 #412]: false
-#387 := [lemma #405]: #433
-#273 := (not #288)
-#274 := (or #273 #473)
-#275 := [th-lemma]: #274
-[unit-resolution #275 #387 #270]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_nat_arith_06 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,13 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T1)
-:extrafuns (
- (uf_1 Int T1)
- (uf_2 T1 Int)
- (uf_3 Int)
- )
-:assumption (forall (?x1 T1) (= (uf_1 (uf_2 ?x1)) ?x1))
-:assumption (forall (?x2 Int) (implies (<= 0 ?x2) (= (uf_2 (uf_1 ?x2)) ?x2)))
-:assumption (forall (?x3 Int) (implies (< ?x3 0) (= (uf_2 (uf_1 ?x3)) 0)))
-:assumption (not (= (uf_2 (uf_1 (ite (< uf_3 0) (~ uf_3) uf_3))) (ite (< uf_3 0) (~ uf_3) uf_3)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_nat_arith_06.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,161 +0,0 @@
-#2 := false
-#9 := 0::int
-decl uf_3 :: int
-#21 := uf_3
-#130 := -1::int
-#131 := (* -1::int uf_3)
-#154 := (>= uf_3 0::int)
-#161 := (ite #154 uf_3 #131)
-#648 := (* -1::int #161)
-#651 := (+ #131 #648)
-#657 := (<= #651 0::int)
-#341 := (= #131 #161)
-#155 := (not #154)
-#649 := (+ uf_3 #648)
-#650 := (<= #649 0::int)
-#254 := (= uf_3 #161)
-#646 := [hypothesis]: #154
-#255 := (or #155 #254)
-#342 := [def-axiom]: #255
-#652 := [unit-resolution #342 #646]: #254
-#290 := (not #254)
-#653 := (or #290 #650)
-#655 := [th-lemma]: #653
-#295 := [unit-resolution #655 #652]: #650
-#346 := (>= #161 0::int)
-#274 := (not #346)
-decl uf_2 :: (-> T1 int)
-decl uf_1 :: (-> int T1)
-#166 := (uf_1 #161)
-#169 := (uf_2 #166)
-#172 := (= #161 #169)
-#175 := (not #172)
-#23 := (- uf_3)
-#22 := (< uf_3 0::int)
-#24 := (ite #22 #23 uf_3)
-#25 := (uf_1 #24)
-#26 := (uf_2 #25)
-#27 := (= #26 #24)
-#28 := (not #27)
-#178 := (iff #28 #175)
-#134 := (ite #22 #131 uf_3)
-#137 := (uf_1 #134)
-#140 := (uf_2 #137)
-#146 := (= #134 #140)
-#151 := (not #146)
-#176 := (iff #151 #175)
-#173 := (iff #146 #172)
-#170 := (= #140 #169)
-#167 := (= #137 #166)
-#164 := (= #134 #161)
-#158 := (ite #155 #131 uf_3)
-#162 := (= #158 #161)
-#163 := [rewrite]: #162
-#159 := (= #134 #158)
-#156 := (iff #22 #155)
-#157 := [rewrite]: #156
-#160 := [monotonicity #157]: #159
-#165 := [trans #160 #163]: #164
-#168 := [monotonicity #165]: #167
-#171 := [monotonicity #168]: #170
-#174 := [monotonicity #165 #171]: #173
-#177 := [monotonicity #174]: #176
-#152 := (iff #28 #151)
-#149 := (iff #27 #146)
-#143 := (= #140 #134)
-#147 := (iff #143 #146)
-#148 := [rewrite]: #147
-#144 := (iff #27 #143)
-#135 := (= #24 #134)
-#132 := (= #23 #131)
-#133 := [rewrite]: #132
-#136 := [monotonicity #133]: #135
-#141 := (= #26 #140)
-#138 := (= #25 #137)
-#139 := [monotonicity #136]: #138
-#142 := [monotonicity #139]: #141
-#145 := [monotonicity #142 #136]: #144
-#150 := [trans #145 #148]: #149
-#153 := [monotonicity #150]: #152
-#179 := [trans #153 #177]: #178
-#129 := [asserted]: #28
-#180 := [mp #129 #179]: #175
-#10 := (:var 0 int)
-#12 := (uf_1 #10)
-#678 := (pattern #12)
-#70 := (>= #10 0::int)
-#71 := (not #70)
-#13 := (uf_2 #12)
-#52 := (= #10 #13)
-#77 := (or #52 #71)
-#679 := (forall (vars (?x2 int)) (:pat #678) #77)
-#82 := (forall (vars (?x2 int)) #77)
-#682 := (iff #82 #679)
-#680 := (iff #77 #77)
-#681 := [refl]: #680
-#683 := [quant-intro #681]: #682
-#183 := (~ #82 #82)
-#195 := (~ #77 #77)
-#196 := [refl]: #195
-#181 := [nnf-pos #196]: #183
-#14 := (= #13 #10)
-#11 := (<= 0::int #10)
-#15 := (implies #11 #14)
-#16 := (forall (vars (?x2 int)) #15)
-#85 := (iff #16 #82)
-#59 := (not #11)
-#60 := (or #59 #52)
-#65 := (forall (vars (?x2 int)) #60)
-#83 := (iff #65 #82)
-#80 := (iff #60 #77)
-#74 := (or #71 #52)
-#78 := (iff #74 #77)
-#79 := [rewrite]: #78
-#75 := (iff #60 #74)
-#72 := (iff #59 #71)
-#68 := (iff #11 #70)
-#69 := [rewrite]: #68
-#73 := [monotonicity #69]: #72
-#76 := [monotonicity #73]: #75
-#81 := [trans #76 #79]: #80
-#84 := [quant-intro #81]: #83
-#66 := (iff #16 #65)
-#63 := (iff #15 #60)
-#56 := (implies #11 #52)
-#61 := (iff #56 #60)
-#62 := [rewrite]: #61
-#57 := (iff #15 #56)
-#54 := (iff #14 #52)
-#55 := [rewrite]: #54
-#58 := [monotonicity #55]: #57
-#64 := [trans #58 #62]: #63
-#67 := [quant-intro #64]: #66
-#86 := [trans #67 #84]: #85
-#51 := [asserted]: #16
-#87 := [mp #51 #86]: #82
-#197 := [mp~ #87 #181]: #82
-#684 := [mp #197 #683]: #679
-#321 := (not #679)
-#451 := (or #321 #172 #274)
-#327 := (or #172 #274)
-#658 := (or #321 #327)
-#333 := (iff #658 #451)
-#665 := [rewrite]: #333
-#332 := [quant-inst]: #658
-#666 := [mp #332 #665]: #451
-#296 := [unit-resolution #666 #684 #180]: #274
-#656 := [th-lemma #646 #296 #295]: false
-#654 := [lemma #656]: #155
-#256 := (or #154 #341)
-#343 := [def-axiom]: #256
-#644 := [unit-resolution #343 #654]: #341
-#366 := (not #341)
-#367 := (or #366 #657)
-#368 := [th-lemma]: #367
-#369 := [unit-resolution #368 #644]: #657
-#647 := (<= #161 0::int)
-#262 := (or #647 #346)
-#639 := [th-lemma]: #262
-#640 := [unit-resolution #639 #296]: #647
-[th-lemma #654 #640 #369]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_nat_arith_07 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,19 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T1 T2)
-:extrafuns (
- (uf_1 Int T1)
- (uf_2 T1 Int)
- (uf_5 T1)
- )
-:extrapreds (
- (up_3 T1)
- (up_4 T1 T1)
- )
-:assumption (forall (?x1 T1) (= (uf_1 (uf_2 ?x1)) ?x1))
-:assumption (forall (?x2 Int) (implies (<= 0 ?x2) (= (uf_2 (uf_1 ?x2)) ?x2)))
-:assumption (forall (?x3 Int) (implies (< ?x3 0) (= (uf_2 (uf_1 ?x3)) 0)))
-:assumption (forall (?x4 T1) (iff (up_3 ?x4) (and (< 1 (uf_2 ?x4)) (forall (?x5 T1) (implies (up_4 ?x5 ?x4) (or (= ?x5 (uf_1 1)) (= ?x5 ?x4)))))))
-:assumption (up_3 (uf_1 (+ (uf_2 (uf_1 (* 4 (uf_2 uf_5)))) 1)))
-:assumption (not (<= 1 (uf_2 uf_5)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_nat_arith_07.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,553 +0,0 @@
-#2 := false
-#9 := 0::int
-decl uf_2 :: (-> T1 int)
-decl uf_1 :: (-> int T1)
-decl uf_5 :: T1
-#36 := uf_5
-#37 := (uf_2 uf_5)
-#35 := 4::int
-#38 := (* 4::int #37)
-#39 := (uf_1 #38)
-#40 := (uf_2 #39)
-#549 := (= #40 0::int)
-#963 := (not #549)
-#537 := (<= #40 0::int)
-#958 := (not #537)
-#22 := 1::int
-#186 := (+ 1::int #40)
-#189 := (uf_1 #186)
-#524 := (uf_2 #189)
-#452 := (<= #524 1::int)
-#874 := (not #452)
-decl up_4 :: (-> T1 T1 bool)
-#4 := (:var 0 T1)
-#456 := (up_4 #4 #189)
-#440 := (pattern #456)
-#446 := (not #456)
-#455 := (= #4 #189)
-#26 := (uf_1 1::int)
-#27 := (= #4 #26)
-#434 := (or #27 #455 #446)
-#416 := (forall (vars (?x5 T1)) (:pat #440) #434)
-#417 := (not #416)
-#409 := (or #417 #452)
-#400 := (not #409)
-decl up_3 :: (-> T1 bool)
-#192 := (up_3 #189)
-#429 := (not #192)
-#405 := (or #429 #400)
-#389 := (not #405)
-decl ?x5!0 :: (-> T1 T1)
-#478 := (?x5!0 #189)
-#479 := (= #26 #478)
-#468 := (= #189 #478)
-#445 := (up_4 #478 #189)
-#447 := (not #445)
-#396 := (or #447 #468 #479)
-#391 := (not #396)
-#386 := (or #192 #391 #452)
-#377 := (not #386)
-#843 := (or #377 #389)
-#848 := (not #843)
-#5 := (uf_2 #4)
-#788 := (pattern #5)
-#21 := (up_3 #4)
-#836 := (pattern #21)
-#210 := (?x5!0 #4)
-#274 := (= #4 #210)
-#271 := (= #26 #210)
-#232 := (up_4 #210 #4)
-#233 := (not #232)
-#277 := (or #233 #271 #274)
-#280 := (not #277)
-#163 := (<= #5 1::int)
-#289 := (or #21 #163 #280)
-#304 := (not #289)
-#24 := (:var 1 T1)
-#25 := (up_4 #4 #24)
-#809 := (pattern #25)
-#28 := (= #4 #24)
-#147 := (not #25)
-#167 := (or #147 #27 #28)
-#810 := (forall (vars (?x5 T1)) (:pat #809) #167)
-#815 := (not #810)
-#818 := (or #163 #815)
-#821 := (not #818)
-#253 := (not #21)
-#824 := (or #253 #821)
-#827 := (not #824)
-#830 := (or #827 #304)
-#833 := (not #830)
-#837 := (forall (vars (?x4 T1)) (:pat #836 #788) #833)
-#170 := (forall (vars (?x5 T1)) #167)
-#236 := (not #170)
-#239 := (or #163 #236)
-#240 := (not #239)
-#215 := (or #253 #240)
-#303 := (not #215)
-#305 := (or #303 #304)
-#306 := (not #305)
-#311 := (forall (vars (?x4 T1)) #306)
-#838 := (iff #311 #837)
-#834 := (iff #306 #833)
-#831 := (iff #305 #830)
-#828 := (iff #303 #827)
-#825 := (iff #215 #824)
-#822 := (iff #240 #821)
-#819 := (iff #239 #818)
-#816 := (iff #236 #815)
-#813 := (iff #170 #810)
-#811 := (iff #167 #167)
-#812 := [refl]: #811
-#814 := [quant-intro #812]: #813
-#817 := [monotonicity #814]: #816
-#820 := [monotonicity #817]: #819
-#823 := [monotonicity #820]: #822
-#826 := [monotonicity #823]: #825
-#829 := [monotonicity #826]: #828
-#832 := [monotonicity #829]: #831
-#835 := [monotonicity #832]: #834
-#839 := [quant-intro #835]: #838
-#164 := (not #163)
-#173 := (and #164 #170)
-#259 := (or #253 #173)
-#294 := (and #259 #289)
-#297 := (forall (vars (?x4 T1)) #294)
-#312 := (iff #297 #311)
-#309 := (iff #294 #306)
-#214 := (and #215 #289)
-#307 := (iff #214 #306)
-#308 := [rewrite]: #307
-#301 := (iff #294 #214)
-#216 := (iff #259 #215)
-#268 := (iff #173 #240)
-#300 := [rewrite]: #268
-#213 := [monotonicity #300]: #216
-#302 := [monotonicity #213]: #301
-#310 := [trans #302 #308]: #309
-#313 := [quant-intro #310]: #312
-#230 := (= #210 #4)
-#231 := (= #210 #26)
-#234 := (or #233 #231 #230)
-#235 := (not #234)
-#228 := (not #164)
-#241 := (or #228 #235)
-#258 := (or #21 #241)
-#260 := (and #259 #258)
-#263 := (forall (vars (?x4 T1)) #260)
-#298 := (iff #263 #297)
-#295 := (iff #260 #294)
-#292 := (iff #258 #289)
-#283 := (or #163 #280)
-#286 := (or #21 #283)
-#290 := (iff #286 #289)
-#291 := [rewrite]: #290
-#287 := (iff #258 #286)
-#284 := (iff #241 #283)
-#281 := (iff #235 #280)
-#278 := (iff #234 #277)
-#275 := (iff #230 #274)
-#276 := [rewrite]: #275
-#272 := (iff #231 #271)
-#273 := [rewrite]: #272
-#279 := [monotonicity #273 #276]: #278
-#282 := [monotonicity #279]: #281
-#269 := (iff #228 #163)
-#270 := [rewrite]: #269
-#285 := [monotonicity #270 #282]: #284
-#288 := [monotonicity #285]: #287
-#293 := [trans #288 #291]: #292
-#296 := [monotonicity #293]: #295
-#299 := [quant-intro #296]: #298
-#176 := (iff #21 #173)
-#179 := (forall (vars (?x4 T1)) #176)
-#264 := (~ #179 #263)
-#261 := (~ #176 #260)
-#251 := (~ #173 #173)
-#249 := (~ #170 #170)
-#247 := (~ #167 #167)
-#248 := [refl]: #247
-#250 := [nnf-pos #248]: #249
-#245 := (~ #164 #164)
-#246 := [refl]: #245
-#252 := [monotonicity #246 #250]: #251
-#242 := (not #173)
-#243 := (~ #242 #241)
-#237 := (~ #236 #235)
-#238 := [sk]: #237
-#229 := (~ #228 #228)
-#209 := [refl]: #229
-#244 := [nnf-neg #209 #238]: #243
-#256 := (~ #21 #21)
-#257 := [refl]: #256
-#254 := (~ #253 #253)
-#255 := [refl]: #254
-#262 := [nnf-pos #255 #257 #244 #252]: #261
-#265 := [nnf-pos #262]: #264
-#29 := (or #27 #28)
-#30 := (implies #25 #29)
-#31 := (forall (vars (?x5 T1)) #30)
-#23 := (< 1::int #5)
-#32 := (and #23 #31)
-#33 := (iff #21 #32)
-#34 := (forall (vars (?x4 T1)) #33)
-#182 := (iff #34 #179)
-#148 := (or #147 #29)
-#151 := (forall (vars (?x5 T1)) #148)
-#154 := (and #23 #151)
-#157 := (iff #21 #154)
-#160 := (forall (vars (?x4 T1)) #157)
-#180 := (iff #160 #179)
-#177 := (iff #157 #176)
-#174 := (iff #154 #173)
-#171 := (iff #151 #170)
-#168 := (iff #148 #167)
-#169 := [rewrite]: #168
-#172 := [quant-intro #169]: #171
-#165 := (iff #23 #164)
-#166 := [rewrite]: #165
-#175 := [monotonicity #166 #172]: #174
-#178 := [monotonicity #175]: #177
-#181 := [quant-intro #178]: #180
-#161 := (iff #34 #160)
-#158 := (iff #33 #157)
-#155 := (iff #32 #154)
-#152 := (iff #31 #151)
-#149 := (iff #30 #148)
-#150 := [rewrite]: #149
-#153 := [quant-intro #150]: #152
-#156 := [monotonicity #153]: #155
-#159 := [monotonicity #156]: #158
-#162 := [quant-intro #159]: #161
-#183 := [trans #162 #181]: #182
-#146 := [asserted]: #34
-#184 := [mp #146 #183]: #179
-#266 := [mp~ #184 #265]: #263
-#267 := [mp #266 #299]: #297
-#314 := [mp #267 #313]: #311
-#840 := [mp #314 #839]: #837
-#754 := (not #837)
-#851 := (or #754 #848)
-#448 := (or #447 #479 #468)
-#439 := (not #448)
-#453 := (or #192 #452 #439)
-#454 := (not #453)
-#457 := (or #446 #27 #455)
-#442 := (forall (vars (?x5 T1)) (:pat #440) #457)
-#443 := (not #442)
-#422 := (or #452 #443)
-#424 := (not #422)
-#430 := (or #429 #424)
-#431 := (not #430)
-#432 := (or #431 #454)
-#433 := (not #432)
-#852 := (or #754 #433)
-#854 := (iff #852 #851)
-#856 := (iff #851 #851)
-#857 := [rewrite]: #856
-#849 := (iff #433 #848)
-#846 := (iff #432 #843)
-#379 := (or #389 #377)
-#844 := (iff #379 #843)
-#845 := [rewrite]: #844
-#841 := (iff #432 #379)
-#378 := (iff #454 #377)
-#388 := (iff #453 #386)
-#381 := (or #192 #452 #391)
-#387 := (iff #381 #386)
-#383 := [rewrite]: #387
-#382 := (iff #453 #381)
-#399 := (iff #439 #391)
-#397 := (iff #448 #396)
-#398 := [rewrite]: #397
-#384 := [monotonicity #398]: #399
-#385 := [monotonicity #384]: #382
-#375 := [trans #385 #383]: #388
-#376 := [monotonicity #375]: #378
-#392 := (iff #431 #389)
-#401 := (iff #430 #405)
-#402 := (iff #424 #400)
-#394 := (iff #422 #409)
-#410 := (or #452 #417)
-#415 := (iff #410 #409)
-#390 := [rewrite]: #415
-#411 := (iff #422 #410)
-#420 := (iff #443 #417)
-#418 := (iff #442 #416)
-#423 := (iff #457 #434)
-#435 := [rewrite]: #423
-#419 := [quant-intro #435]: #418
-#408 := [monotonicity #419]: #420
-#414 := [monotonicity #408]: #411
-#395 := [trans #414 #390]: #394
-#403 := [monotonicity #395]: #402
-#406 := [monotonicity #403]: #401
-#393 := [monotonicity #406]: #392
-#842 := [monotonicity #393 #376]: #841
-#847 := [trans #842 #845]: #846
-#850 := [monotonicity #847]: #849
-#855 := [monotonicity #850]: #854
-#858 := [trans #855 #857]: #854
-#853 := [quant-inst]: #852
-#859 := [mp #853 #858]: #851
-#934 := [unit-resolution #859 #840]: #848
-#893 := (or #843 #405)
-#894 := [def-axiom]: #893
-#935 := [unit-resolution #894 #934]: #405
-#938 := (or #389 #400)
-#41 := (+ #40 1::int)
-#42 := (uf_1 #41)
-#43 := (up_3 #42)
-#193 := (iff #43 #192)
-#190 := (= #42 #189)
-#187 := (= #41 #186)
-#188 := [rewrite]: #187
-#191 := [monotonicity #188]: #190
-#194 := [monotonicity #191]: #193
-#185 := [asserted]: #43
-#197 := [mp #185 #194]: #192
-#889 := (or #389 #429 #400)
-#890 := [def-axiom]: #889
-#939 := [unit-resolution #890 #197]: #938
-#940 := [unit-resolution #939 #935]: #400
-#881 := (or #409 #874)
-#882 := [def-axiom]: #881
-#941 := [unit-resolution #882 #940]: #874
-#555 := -1::int
-#525 := (* -1::int #524)
-#528 := (+ #40 #525)
-#494 := (>= #528 -1::int)
-#510 := (= #528 -1::int)
-#514 := (>= #40 -1::int)
-#495 := (= #524 0::int)
-#946 := (not #495)
-#467 := (<= #524 0::int)
-#942 := (not #467)
-#943 := (or #942 #452)
-#944 := [th-lemma]: #943
-#945 := [unit-resolution #944 #941]: #942
-#947 := (or #946 #467)
-#948 := [th-lemma]: #947
-#949 := [unit-resolution #948 #945]: #946
-#498 := (or #495 #514)
-#10 := (:var 0 int)
-#12 := (uf_1 #10)
-#796 := (pattern #12)
-#87 := (>= #10 0::int)
-#13 := (uf_2 #12)
-#18 := (= #13 0::int)
-#135 := (or #18 #87)
-#803 := (forall (vars (?x3 int)) (:pat #796) #135)
-#140 := (forall (vars (?x3 int)) #135)
-#806 := (iff #140 #803)
-#804 := (iff #135 #135)
-#805 := [refl]: #804
-#807 := [quant-intro #805]: #806
-#207 := (~ #140 #140)
-#225 := (~ #135 #135)
-#226 := [refl]: #225
-#208 := [nnf-pos #226]: #207
-#17 := (< #10 0::int)
-#19 := (implies #17 #18)
-#20 := (forall (vars (?x3 int)) #19)
-#143 := (iff #20 #140)
-#106 := (= 0::int #13)
-#112 := (not #17)
-#113 := (or #112 #106)
-#118 := (forall (vars (?x3 int)) #113)
-#141 := (iff #118 #140)
-#138 := (iff #113 #135)
-#132 := (or #87 #18)
-#136 := (iff #132 #135)
-#137 := [rewrite]: #136
-#133 := (iff #113 #132)
-#130 := (iff #106 #18)
-#131 := [rewrite]: #130
-#128 := (iff #112 #87)
-#88 := (not #87)
-#123 := (not #88)
-#126 := (iff #123 #87)
-#127 := [rewrite]: #126
-#124 := (iff #112 #123)
-#121 := (iff #17 #88)
-#122 := [rewrite]: #121
-#125 := [monotonicity #122]: #124
-#129 := [trans #125 #127]: #128
-#134 := [monotonicity #129 #131]: #133
-#139 := [trans #134 #137]: #138
-#142 := [quant-intro #139]: #141
-#119 := (iff #20 #118)
-#116 := (iff #19 #113)
-#109 := (implies #17 #106)
-#114 := (iff #109 #113)
-#115 := [rewrite]: #114
-#110 := (iff #19 #109)
-#107 := (iff #18 #106)
-#108 := [rewrite]: #107
-#111 := [monotonicity #108]: #110
-#117 := [trans #111 #115]: #116
-#120 := [quant-intro #117]: #119
-#144 := [trans #120 #142]: #143
-#105 := [asserted]: #20
-#145 := [mp #105 #144]: #140
-#227 := [mp~ #145 #208]: #140
-#808 := [mp #227 #807]: #803
-#532 := (not #803)
-#488 := (or #532 #495 #514)
-#529 := (>= #186 0::int)
-#496 := (or #495 #529)
-#489 := (or #532 #496)
-#474 := (iff #489 #488)
-#482 := (or #532 #498)
-#483 := (iff #482 #488)
-#493 := [rewrite]: #483
-#491 := (iff #489 #482)
-#497 := (iff #496 #498)
-#515 := (iff #529 #514)
-#516 := [rewrite]: #515
-#499 := [monotonicity #516]: #497
-#492 := [monotonicity #499]: #491
-#475 := [trans #492 #493]: #474
-#490 := [quant-inst]: #489
-#476 := [mp #490 #475]: #488
-#950 := [unit-resolution #476 #808]: #498
-#951 := [unit-resolution #950 #949]: #514
-#517 := (not #514)
-#520 := (or #510 #517)
-#69 := (= #10 #13)
-#94 := (or #69 #88)
-#797 := (forall (vars (?x2 int)) (:pat #796) #94)
-#99 := (forall (vars (?x2 int)) #94)
-#800 := (iff #99 #797)
-#798 := (iff #94 #94)
-#799 := [refl]: #798
-#801 := [quant-intro #799]: #800
-#206 := (~ #99 #99)
-#222 := (~ #94 #94)
-#223 := [refl]: #222
-#196 := [nnf-pos #223]: #206
-#14 := (= #13 #10)
-#11 := (<= 0::int #10)
-#15 := (implies #11 #14)
-#16 := (forall (vars (?x2 int)) #15)
-#102 := (iff #16 #99)
-#76 := (not #11)
-#77 := (or #76 #69)
-#82 := (forall (vars (?x2 int)) #77)
-#100 := (iff #82 #99)
-#97 := (iff #77 #94)
-#91 := (or #88 #69)
-#95 := (iff #91 #94)
-#96 := [rewrite]: #95
-#92 := (iff #77 #91)
-#89 := (iff #76 #88)
-#85 := (iff #11 #87)
-#86 := [rewrite]: #85
-#90 := [monotonicity #86]: #89
-#93 := [monotonicity #90]: #92
-#98 := [trans #93 #96]: #97
-#101 := [quant-intro #98]: #100
-#83 := (iff #16 #82)
-#80 := (iff #15 #77)
-#73 := (implies #11 #69)
-#78 := (iff #73 #77)
-#79 := [rewrite]: #78
-#74 := (iff #15 #73)
-#71 := (iff #14 #69)
-#72 := [rewrite]: #71
-#75 := [monotonicity #72]: #74
-#81 := [trans #75 #79]: #80
-#84 := [quant-intro #81]: #83
-#103 := [trans #84 #101]: #102
-#68 := [asserted]: #16
-#104 := [mp #68 #103]: #99
-#224 := [mp~ #104 #196]: #99
-#802 := [mp #224 #801]: #797
-#559 := (not #797)
-#511 := (or #559 #510 #517)
-#531 := (not #529)
-#526 := (= #186 #524)
-#527 := (or #526 #531)
-#523 := (or #559 #527)
-#507 := (iff #523 #511)
-#502 := (or #559 #520)
-#505 := (iff #502 #511)
-#506 := [rewrite]: #505
-#503 := (iff #523 #502)
-#521 := (iff #527 #520)
-#518 := (iff #531 #517)
-#519 := [monotonicity #516]: #518
-#512 := (iff #526 #510)
-#513 := [rewrite]: #512
-#522 := [monotonicity #513 #519]: #521
-#504 := [monotonicity #522]: #503
-#508 := [trans #504 #506]: #507
-#500 := [quant-inst]: #523
-#501 := [mp #500 #508]: #511
-#952 := [unit-resolution #501 #802]: #520
-#953 := [unit-resolution #952 #951]: #510
-#954 := (not #510)
-#955 := (or #954 #494)
-#956 := [th-lemma]: #955
-#957 := [unit-resolution #956 #953]: #494
-#959 := (not #494)
-#960 := (or #958 #452 #959)
-#961 := [th-lemma]: #960
-#962 := [unit-resolution #961 #957 #941]: #958
-#964 := (or #963 #537)
-#965 := [th-lemma]: #964
-#966 := [unit-resolution #965 #962]: #963
-#583 := (>= #38 0::int)
-#584 := (not #583)
-#556 := (* -1::int #40)
-#557 := (+ #38 #556)
-#558 := (= #557 0::int)
-#971 := (not #558)
-#544 := (>= #557 0::int)
-#967 := (not #544)
-#201 := (>= #37 1::int)
-#202 := (not #201)
-#44 := (<= 1::int #37)
-#45 := (not #44)
-#203 := (iff #45 #202)
-#199 := (iff #44 #201)
-#200 := [rewrite]: #199
-#204 := [monotonicity #200]: #203
-#195 := [asserted]: #45
-#205 := [mp #195 #204]: #202
-#968 := (or #967 #201 #452 #959)
-#969 := [th-lemma]: #968
-#970 := [unit-resolution #969 #205 #957 #941]: #967
-#972 := (or #971 #544)
-#973 := [th-lemma]: #972
-#974 := [unit-resolution #973 #970]: #971
-#562 := (or #558 #584)
-#564 := (or #559 #558 #584)
-#567 := (= #38 #40)
-#585 := (or #567 #584)
-#543 := (or #559 #585)
-#542 := (iff #543 #564)
-#550 := (or #559 #562)
-#551 := (iff #550 #564)
-#554 := [rewrite]: #551
-#552 := (iff #543 #550)
-#404 := (iff #585 #562)
-#560 := (iff #567 #558)
-#561 := [rewrite]: #560
-#563 := [monotonicity #561]: #404
-#553 := [monotonicity #563]: #552
-#545 := [trans #553 #554]: #542
-#546 := [quant-inst]: #543
-#547 := [mp #546 #545]: #564
-#975 := [unit-resolution #547 #802]: #562
-#976 := [unit-resolution #975 #974]: #584
-#539 := (or #549 #583)
-#535 := (or #532 #549 #583)
-#536 := (or #532 #539)
-#533 := (iff #536 #535)
-#541 := [rewrite]: #533
-#540 := [quant-inst]: #536
-#534 := [mp #540 #541]: #535
-#977 := [unit-resolution #534 #808]: #539
-[unit-resolution #977 #976 #966]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_nlarith_01 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,10 +0,0 @@
-(benchmark Isabelle
-:extrafuns (
- (uf_1 Int)
- (uf_2 Int)
- )
-:assumption (< 0 uf_1)
-:assumption (< 0 (* uf_1 uf_2))
-:assumption (not (< 0 uf_2))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_nlarith_01.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_nlarith_02 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,9 +0,0 @@
-(benchmark Isabelle
-:extrafuns (
- (uf_1 Int)
- (uf_2 Int)
- (uf_3 Int)
- )
-:assumption (not (= (* uf_1 (+ (+ uf_2 1) uf_3)) (+ (* uf_1 uf_2) (* uf_1 (+ uf_3 1)))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_nlarith_02.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,75 +0,0 @@
-#2 := false
-#6 := 1::int
-decl uf_3 :: int
-#8 := uf_3
-#12 := (+ uf_3 1::int)
-decl uf_1 :: int
-#4 := uf_1
-#13 := (* uf_1 #12)
-decl uf_2 :: int
-#5 := uf_2
-#11 := (* uf_1 uf_2)
-#14 := (+ #11 #13)
-#7 := (+ uf_2 1::int)
-#9 := (+ #7 uf_3)
-#10 := (* uf_1 #9)
-#15 := (= #10 #14)
-#16 := (not #15)
-#85 := (iff #16 false)
-#1 := true
-#80 := (not true)
-#83 := (iff #80 false)
-#84 := [rewrite]: #83
-#81 := (iff #16 #80)
-#78 := (iff #15 true)
-#48 := (* uf_1 uf_3)
-#49 := (+ #11 #48)
-#50 := (+ uf_1 #49)
-#73 := (= #50 #50)
-#76 := (iff #73 true)
-#77 := [rewrite]: #76
-#74 := (iff #15 #73)
-#71 := (= #14 #50)
-#61 := (+ uf_1 #48)
-#66 := (+ #11 #61)
-#69 := (= #66 #50)
-#70 := [rewrite]: #69
-#67 := (= #14 #66)
-#64 := (= #13 #61)
-#55 := (+ 1::int uf_3)
-#58 := (* uf_1 #55)
-#62 := (= #58 #61)
-#63 := [rewrite]: #62
-#59 := (= #13 #58)
-#56 := (= #12 #55)
-#57 := [rewrite]: #56
-#60 := [monotonicity #57]: #59
-#65 := [trans #60 #63]: #64
-#68 := [monotonicity #65]: #67
-#72 := [trans #68 #70]: #71
-#53 := (= #10 #50)
-#39 := (+ uf_2 uf_3)
-#40 := (+ 1::int #39)
-#45 := (* uf_1 #40)
-#51 := (= #45 #50)
-#52 := [rewrite]: #51
-#46 := (= #10 #45)
-#43 := (= #9 #40)
-#33 := (+ 1::int uf_2)
-#36 := (+ #33 uf_3)
-#41 := (= #36 #40)
-#42 := [rewrite]: #41
-#37 := (= #9 #36)
-#34 := (= #7 #33)
-#35 := [rewrite]: #34
-#38 := [monotonicity #35]: #37
-#44 := [trans #38 #42]: #43
-#47 := [monotonicity #44]: #46
-#54 := [trans #47 #52]: #53
-#75 := [monotonicity #54 #72]: #74
-#79 := [trans #75 #77]: #78
-#82 := [monotonicity #79]: #81
-#86 := [trans #82 #84]: #85
-#32 := [asserted]: #16
-[mp #32 #86]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_nlarith_03 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,8 +0,0 @@
-(benchmark Isabelle
-:extrafuns (
- (uf_1 Real)
- (uf_2 Real)
- )
-:assumption (not (= (- (* uf_1 (+ 1.0 uf_2)) (* uf_1 (- 1.0 uf_2))) (* (* 2.0 uf_1) uf_2)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_nlarith_03.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,62 +0,0 @@
-#2 := false
-decl uf_2 :: real
-#6 := uf_2
-decl uf_1 :: real
-#4 := uf_1
-#12 := 2::real
-#13 := (* 2::real uf_1)
-#14 := (* #13 uf_2)
-#5 := 1::real
-#9 := (- 1::real uf_2)
-#10 := (* uf_1 #9)
-#7 := (+ 1::real uf_2)
-#8 := (* uf_1 #7)
-#11 := (- #8 #10)
-#15 := (= #11 #14)
-#16 := (not #15)
-#73 := (iff #16 false)
-#1 := true
-#68 := (not true)
-#71 := (iff #68 false)
-#72 := [rewrite]: #71
-#69 := (iff #16 #68)
-#66 := (iff #15 true)
-#33 := (* uf_1 uf_2)
-#55 := (* 2::real #33)
-#61 := (= #55 #55)
-#64 := (iff #61 true)
-#65 := [rewrite]: #64
-#62 := (iff #15 #61)
-#59 := (= #14 #55)
-#60 := [rewrite]: #59
-#57 := (= #11 #55)
-#37 := -1::real
-#45 := (* -1::real #33)
-#46 := (+ uf_1 #45)
-#34 := (+ uf_1 #33)
-#51 := (- #34 #46)
-#54 := (= #51 #55)
-#56 := [rewrite]: #54
-#52 := (= #11 #51)
-#49 := (= #10 #46)
-#38 := (* -1::real uf_2)
-#39 := (+ 1::real #38)
-#42 := (* uf_1 #39)
-#47 := (= #42 #46)
-#48 := [rewrite]: #47
-#43 := (= #10 #42)
-#40 := (= #9 #39)
-#41 := [rewrite]: #40
-#44 := [monotonicity #41]: #43
-#50 := [trans #44 #48]: #49
-#35 := (= #8 #34)
-#36 := [rewrite]: #35
-#53 := [monotonicity #36 #50]: #52
-#58 := [trans #53 #56]: #57
-#63 := [monotonicity #58 #60]: #62
-#67 := [trans #63 #65]: #66
-#70 := [monotonicity #67]: #69
-#74 := [trans #70 #72]: #73
-#32 := [asserted]: #16
-[mp #32 #74]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_nlarith_04 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,11 +0,0 @@
-(benchmark Isabelle
-:extrafuns (
- (uf_1 Int)
- (uf_3 Int)
- (uf_5 Int)
- (uf_4 Int)
- (uf_2 Int)
- )
-:assumption (not (= (+ (+ uf_1 (* (+ 1 uf_2) (+ uf_3 uf_4))) (* uf_2 uf_5)) (- (+ uf_1 (+ (+ (* (* 2 (+ 1 uf_2)) (+ uf_3 uf_4)) (* (+ 1 uf_2) uf_5)) (* uf_5 uf_2))) (* (+ 1 uf_2) (+ (+ uf_3 uf_5) uf_4)))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_nlarith_04.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,141 +0,0 @@
-#2 := false
-decl uf_4 :: int
-#9 := uf_4
-decl uf_5 :: int
-#13 := uf_5
-decl uf_3 :: int
-#8 := uf_3
-#24 := (+ uf_3 uf_5)
-#25 := (+ #24 uf_4)
-decl uf_2 :: int
-#6 := uf_2
-#5 := 1::int
-#7 := (+ 1::int uf_2)
-#26 := (* #7 #25)
-#21 := (* uf_5 uf_2)
-#19 := (* #7 uf_5)
-#10 := (+ uf_3 uf_4)
-#16 := 2::int
-#17 := (* 2::int #7)
-#18 := (* #17 #10)
-#20 := (+ #18 #19)
-#22 := (+ #20 #21)
-decl uf_1 :: int
-#4 := uf_1
-#23 := (+ uf_1 #22)
-#27 := (- #23 #26)
-#14 := (* uf_2 uf_5)
-#11 := (* #7 #10)
-#12 := (+ uf_1 #11)
-#15 := (+ #12 #14)
-#28 := (= #15 #27)
-#29 := (not #28)
-#149 := (iff #29 false)
-#1 := true
-#144 := (not true)
-#147 := (iff #144 false)
-#148 := [rewrite]: #147
-#145 := (iff #29 #144)
-#142 := (iff #28 true)
-#47 := (* uf_2 uf_4)
-#46 := (* uf_2 uf_3)
-#48 := (+ #46 #47)
-#59 := (+ #14 #48)
-#60 := (+ uf_4 #59)
-#61 := (+ uf_3 #60)
-#62 := (+ uf_1 #61)
-#136 := (= #62 #62)
-#140 := (iff #136 true)
-#141 := [rewrite]: #140
-#135 := (iff #28 #136)
-#138 := (= #27 #62)
-#123 := (+ uf_5 #59)
-#124 := (+ uf_4 #123)
-#125 := (+ uf_3 #124)
-#77 := (* 2::int #47)
-#75 := (* 2::int #46)
-#78 := (+ #75 #77)
-#104 := (* 2::int #14)
-#105 := (+ #104 #78)
-#106 := (+ uf_5 #105)
-#76 := (* 2::int uf_4)
-#107 := (+ #76 #106)
-#74 := (* 2::int uf_3)
-#108 := (+ #74 #107)
-#113 := (+ uf_1 #108)
-#130 := (- #113 #125)
-#133 := (= #130 #62)
-#139 := [rewrite]: #133
-#131 := (= #27 #130)
-#128 := (= #26 #125)
-#116 := (+ uf_4 uf_5)
-#117 := (+ uf_3 #116)
-#120 := (* #7 #117)
-#126 := (= #120 #125)
-#127 := [rewrite]: #126
-#121 := (= #26 #120)
-#118 := (= #25 #117)
-#119 := [rewrite]: #118
-#122 := [monotonicity #119]: #121
-#129 := [trans #122 #127]: #128
-#114 := (= #23 #113)
-#111 := (= #22 #108)
-#91 := (+ #14 #78)
-#92 := (+ uf_5 #91)
-#93 := (+ #76 #92)
-#94 := (+ #74 #93)
-#101 := (+ #94 #14)
-#109 := (= #101 #108)
-#110 := [rewrite]: #109
-#102 := (= #22 #101)
-#99 := (= #21 #14)
-#100 := [rewrite]: #99
-#97 := (= #20 #94)
-#85 := (+ uf_5 #14)
-#79 := (+ #76 #78)
-#80 := (+ #74 #79)
-#88 := (+ #80 #85)
-#95 := (= #88 #94)
-#96 := [rewrite]: #95
-#89 := (= #20 #88)
-#86 := (= #19 #85)
-#87 := [rewrite]: #86
-#83 := (= #18 #80)
-#67 := (* 2::int uf_2)
-#68 := (+ 2::int #67)
-#71 := (* #68 #10)
-#81 := (= #71 #80)
-#82 := [rewrite]: #81
-#72 := (= #18 #71)
-#69 := (= #17 #68)
-#70 := [rewrite]: #69
-#73 := [monotonicity #70]: #72
-#84 := [trans #73 #82]: #83
-#90 := [monotonicity #84 #87]: #89
-#98 := [trans #90 #96]: #97
-#103 := [monotonicity #98 #100]: #102
-#112 := [trans #103 #110]: #111
-#115 := [monotonicity #112]: #114
-#132 := [monotonicity #115 #129]: #131
-#137 := [trans #132 #139]: #138
-#65 := (= #15 #62)
-#49 := (+ uf_4 #48)
-#50 := (+ uf_3 #49)
-#53 := (+ uf_1 #50)
-#56 := (+ #53 #14)
-#63 := (= #56 #62)
-#64 := [rewrite]: #63
-#57 := (= #15 #56)
-#54 := (= #12 #53)
-#51 := (= #11 #50)
-#52 := [rewrite]: #51
-#55 := [monotonicity #52]: #54
-#58 := [monotonicity #55]: #57
-#66 := [trans #58 #64]: #65
-#134 := [monotonicity #66 #137]: #135
-#143 := [trans #134 #141]: #142
-#146 := [monotonicity #143]: #145
-#150 := [trans #146 #148]: #149
-#45 := [asserted]: #29
-[mp #45 #150]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_pair_01 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,17 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T2 T3 T1)
-:extrafuns (
- (uf_2 T1 T2)
- (uf_3 T1 T3)
- (uf_1 T2 T3 T1)
- (uf_6 T2)
- (uf_4 T2)
- (uf_5 T3)
- )
-:assumption (forall (?x1 T1) (= (uf_1 (uf_2 ?x1) (uf_3 ?x1)) ?x1))
-:assumption (forall (?x2 T2) (?x3 T3) (= (uf_3 (uf_1 ?x2 ?x3)) ?x3))
-:assumption (forall (?x4 T2) (?x5 T3) (= (uf_2 (uf_1 ?x4 ?x5)) ?x4))
-:assumption (= (uf_2 (uf_1 uf_4 uf_5)) uf_6)
-:assumption (not (= uf_4 uf_6))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_pair_01.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,50 +0,0 @@
-#2 := false
-decl uf_6 :: T2
-#23 := uf_6
-decl uf_4 :: T2
-#19 := uf_4
-#25 := (= uf_4 uf_6)
-decl uf_2 :: (-> T1 T2)
-decl uf_1 :: (-> T2 T3 T1)
-decl uf_5 :: T3
-#20 := uf_5
-#21 := (uf_1 uf_4 uf_5)
-#22 := (uf_2 #21)
-#24 := (= #22 uf_6)
-#65 := [asserted]: #24
-#143 := (= uf_4 #22)
-#11 := (:var 0 T3)
-#10 := (:var 1 T2)
-#12 := (uf_1 #10 #11)
-#567 := (pattern #12)
-#16 := (uf_2 #12)
-#58 := (= #10 #16)
-#574 := (forall (vars (?x4 T2) (?x5 T3)) (:pat #567) #58)
-#62 := (forall (vars (?x4 T2) (?x5 T3)) #58)
-#577 := (iff #62 #574)
-#575 := (iff #58 #58)
-#576 := [refl]: #575
-#578 := [quant-intro #576]: #577
-#71 := (~ #62 #62)
-#87 := (~ #58 #58)
-#88 := [refl]: #87
-#72 := [nnf-pos #88]: #71
-#17 := (= #16 #10)
-#18 := (forall (vars (?x4 T2) (?x5 T3)) #17)
-#63 := (iff #18 #62)
-#60 := (iff #17 #58)
-#61 := [rewrite]: #60
-#64 := [quant-intro #61]: #63
-#57 := [asserted]: #18
-#67 := [mp #57 #64]: #62
-#89 := [mp~ #67 #72]: #62
-#579 := [mp #89 #578]: #574
-#214 := (not #574)
-#551 := (or #214 #143)
-#553 := [quant-inst]: #551
-#233 := [unit-resolution #553 #579]: #143
-#235 := [trans #233 #65]: #25
-#26 := (not #25)
-#66 := [asserted]: #26
-[unit-resolution #66 #235]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_pair_02 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,24 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T2 T3 T1 T4)
-:extrafuns (
- (uf_2 T1 T2)
- (uf_5 T4 T3)
- (uf_3 T1 T3)
- (uf_6 T4 T2)
- (uf_1 T2 T3 T1)
- (uf_4 T3 T2 T4)
- (uf_8 T2)
- (uf_9 T3)
- (uf_7 T1)
- (uf_10 T4)
- )
-:assumption (forall (?x1 T1) (= (uf_1 (uf_2 ?x1) (uf_3 ?x1)) ?x1))
-:assumption (forall (?x2 T4) (= (uf_4 (uf_5 ?x2) (uf_6 ?x2)) ?x2))
-:assumption (forall (?x3 T2) (?x4 T3) (= (uf_3 (uf_1 ?x3 ?x4)) ?x4))
-:assumption (forall (?x5 T3) (?x6 T2) (= (uf_6 (uf_4 ?x5 ?x6)) ?x6))
-:assumption (forall (?x7 T2) (?x8 T3) (= (uf_2 (uf_1 ?x7 ?x8)) ?x7))
-:assumption (forall (?x9 T3) (?x10 T2) (= (uf_5 (uf_4 ?x9 ?x10)) ?x9))
-:assumption (and (= uf_7 (uf_1 uf_8 uf_9)) (= uf_10 (uf_4 uf_9 uf_8)))
-:assumption (not (= (uf_2 uf_7) (uf_6 uf_10)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_pair_02.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,105 +0,0 @@
-#2 := false
-decl uf_6 :: (-> T4 T2)
-decl uf_10 :: T4
-#39 := uf_10
-#44 := (uf_6 uf_10)
-decl uf_2 :: (-> T1 T2)
-decl uf_7 :: T1
-#34 := uf_7
-#43 := (uf_2 uf_7)
-#45 := (= #43 #44)
-decl uf_4 :: (-> T3 T2 T4)
-decl uf_8 :: T2
-#35 := uf_8
-decl uf_9 :: T3
-#36 := uf_9
-#40 := (uf_4 uf_9 uf_8)
-#204 := (uf_6 #40)
-#598 := (= #204 #44)
-#595 := (= #44 #204)
-#41 := (= uf_10 #40)
-decl uf_1 :: (-> T2 T3 T1)
-#37 := (uf_1 uf_8 uf_9)
-#38 := (= uf_7 #37)
-#42 := (and #38 #41)
-#109 := [asserted]: #42
-#114 := [and-elim #109]: #41
-#256 := [monotonicity #114]: #595
-#599 := [symm #256]: #598
-#596 := (= #43 #204)
-#269 := (= uf_8 #204)
-#23 := (:var 0 T2)
-#22 := (:var 1 T3)
-#24 := (uf_4 #22 #23)
-#643 := (pattern #24)
-#25 := (uf_6 #24)
-#86 := (= #23 #25)
-#644 := (forall (vars (?x5 T3) (?x6 T2)) (:pat #643) #86)
-#90 := (forall (vars (?x5 T3) (?x6 T2)) #86)
-#647 := (iff #90 #644)
-#645 := (iff #86 #86)
-#646 := [refl]: #645
-#648 := [quant-intro #646]: #647
-#119 := (~ #90 #90)
-#144 := (~ #86 #86)
-#145 := [refl]: #144
-#120 := [nnf-pos #145]: #119
-#26 := (= #25 #23)
-#27 := (forall (vars (?x5 T3) (?x6 T2)) #26)
-#91 := (iff #27 #90)
-#88 := (iff #26 #86)
-#89 := [rewrite]: #88
-#92 := [quant-intro #89]: #91
-#85 := [asserted]: #27
-#95 := [mp #85 #92]: #90
-#146 := [mp~ #95 #120]: #90
-#649 := [mp #146 #648]: #644
-#613 := (not #644)
-#619 := (or #613 #269)
-#609 := [quant-inst]: #619
-#267 := [unit-resolution #609 #649]: #269
-#600 := (= #43 uf_8)
-#289 := (uf_2 #37)
-#259 := (= #289 uf_8)
-#296 := (= uf_8 #289)
-#17 := (:var 0 T3)
-#16 := (:var 1 T2)
-#18 := (uf_1 #16 #17)
-#636 := (pattern #18)
-#28 := (uf_2 #18)
-#94 := (= #16 #28)
-#650 := (forall (vars (?x7 T2) (?x8 T3)) (:pat #636) #94)
-#98 := (forall (vars (?x7 T2) (?x8 T3)) #94)
-#653 := (iff #98 #650)
-#651 := (iff #94 #94)
-#652 := [refl]: #651
-#654 := [quant-intro #652]: #653
-#121 := (~ #98 #98)
-#147 := (~ #94 #94)
-#148 := [refl]: #147
-#122 := [nnf-pos #148]: #121
-#29 := (= #28 #16)
-#30 := (forall (vars (?x7 T2) (?x8 T3)) #29)
-#99 := (iff #30 #98)
-#96 := (iff #29 #94)
-#97 := [rewrite]: #96
-#100 := [quant-intro #97]: #99
-#93 := [asserted]: #30
-#103 := [mp #93 #100]: #98
-#149 := [mp~ #103 #122]: #98
-#655 := [mp #149 #654]: #650
-#615 := (not #650)
-#616 := (or #615 #296)
-#617 := [quant-inst]: #616
-#618 := [unit-resolution #617 #655]: #296
-#597 := [symm #618]: #259
-#611 := (= #43 #289)
-#113 := [and-elim #109]: #38
-#252 := [monotonicity #113]: #611
-#601 := [trans #252 #597]: #600
-#602 := [trans #601 #267]: #596
-#238 := [trans #602 #599]: #45
-#46 := (not #45)
-#110 := [asserted]: #46
-[unit-resolution #110 #238]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_prop_01 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,4 +0,0 @@
-(benchmark Isabelle
-:assumption (not true)
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_prop_01.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,8 +0,0 @@
-#2 := false
-#1 := true
-#4 := (not true)
-#21 := (iff #4 false)
-#22 := [rewrite]: #21
-#20 := [asserted]: #4
-[mp #20 #22]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_prop_02 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,8 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T1)
-:extrapreds (
- (up_1)
- )
-:assumption (not (or up_1 (not up_1)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_prop_02.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,19 +0,0 @@
-#2 := false
-decl up_1 :: bool
-#4 := up_1
-#5 := (not up_1)
-#6 := (or up_1 #5)
-#7 := (not #6)
-#31 := (iff #7 false)
-#1 := true
-#26 := (not true)
-#29 := (iff #26 false)
-#30 := [rewrite]: #29
-#27 := (iff #7 #26)
-#24 := (iff #6 true)
-#25 := [rewrite]: #24
-#28 := [monotonicity #25]: #27
-#32 := [trans #28 #30]: #31
-#23 := [asserted]: #7
-[mp #23 #32]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_prop_03 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,8 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T1)
-:extrapreds (
- (up_1)
- )
-:assumption (not (iff (and up_1 true) up_1))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_prop_03.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,25 +0,0 @@
-#2 := false
-decl up_1 :: bool
-#4 := up_1
-#1 := true
-#5 := (and up_1 true)
-#6 := (iff #5 up_1)
-#7 := (not #6)
-#37 := (iff #7 false)
-#32 := (not true)
-#35 := (iff #32 false)
-#36 := [rewrite]: #35
-#33 := (iff #7 #32)
-#30 := (iff #6 true)
-#25 := (iff up_1 up_1)
-#28 := (iff #25 true)
-#29 := [rewrite]: #28
-#26 := (iff #6 #25)
-#24 := [rewrite]: #6
-#27 := [monotonicity #24]: #26
-#31 := [trans #27 #29]: #30
-#34 := [monotonicity #31]: #33
-#38 := [trans #34 #36]: #37
-#23 := [asserted]: #7
-[mp #23 #38]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_prop_04 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,10 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T1)
-:extrapreds (
- (up_1)
- (up_2)
- )
-:assumption (and (or up_1 up_2) (not up_1))
-:assumption (not up_2)
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_prop_04.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,38 +0,0 @@
-#2 := false
-decl up_2 :: bool
-#5 := up_2
-decl up_1 :: bool
-#4 := up_1
-#6 := (or up_1 up_2)
-#51 := (iff #6 false)
-#46 := (or false false)
-#49 := (iff #46 false)
-#50 := [rewrite]: #49
-#47 := (iff #6 #46)
-#40 := (iff up_2 false)
-#9 := (not up_2)
-#43 := (iff #9 #40)
-#41 := (iff #40 #9)
-#42 := [rewrite]: #41
-#44 := [symm #42]: #43
-#32 := [asserted]: #9
-#45 := [mp #32 #44]: #40
-#35 := (iff up_1 false)
-#7 := (not up_1)
-#37 := (iff #7 #35)
-#33 := (iff #35 #7)
-#36 := [rewrite]: #33
-#38 := [symm #36]: #37
-#26 := (and #7 #6)
-#8 := (and #6 #7)
-#27 := (iff #8 #26)
-#28 := [rewrite]: #27
-#25 := [asserted]: #8
-#31 := [mp #25 #28]: #26
-#29 := [and-elim #31]: #7
-#39 := [mp #29 #38]: #35
-#48 := [monotonicity #39 #45]: #47
-#52 := [trans #48 #50]: #51
-#30 := [and-elim #31]: #6
-[mp #30 #52]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_prop_05 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,12 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T1)
-:extrapreds (
- (up_1)
- (up_2)
- (up_3)
- (up_4)
- )
-:assumption (or (and up_1 up_2) (and up_3 up_4))
-:assumption (not (or (and up_1 up_2) (and up_3 up_4)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_prop_05.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1 +0,0 @@
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_prop_06 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,10 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T1)
-:extrapreds (
- (up_1)
- (up_2)
- (up_3)
- )
-:assumption (not (implies (or (and up_1 up_2) up_3) (or (implies up_1 (or (and up_3 up_2) (and up_1 up_3))) up_1)))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_prop_06.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,71 +0,0 @@
-#2 := false
-decl up_1 :: bool
-#4 := up_1
-decl up_3 :: bool
-#7 := up_3
-#10 := (and up_1 up_3)
-decl up_2 :: bool
-#5 := up_2
-#9 := (and up_3 up_2)
-#11 := (or #9 #10)
-#12 := (implies up_1 #11)
-#13 := (or #12 up_1)
-#6 := (and up_1 up_2)
-#8 := (or #6 up_3)
-#14 := (implies #8 #13)
-#15 := (not #14)
-#81 := (iff #15 false)
-#32 := (and up_2 up_3)
-#38 := (or #10 #32)
-#46 := (not up_1)
-#47 := (or #46 #38)
-#55 := (or up_1 #47)
-#63 := (not #8)
-#64 := (or #63 #55)
-#69 := (not #64)
-#79 := (iff #69 false)
-#1 := true
-#74 := (not true)
-#77 := (iff #74 false)
-#78 := [rewrite]: #77
-#75 := (iff #69 #74)
-#72 := (iff #64 true)
-#73 := [rewrite]: #72
-#76 := [monotonicity #73]: #75
-#80 := [trans #76 #78]: #79
-#70 := (iff #15 #69)
-#67 := (iff #14 #64)
-#60 := (implies #8 #55)
-#65 := (iff #60 #64)
-#66 := [rewrite]: #65
-#61 := (iff #14 #60)
-#58 := (iff #13 #55)
-#52 := (or #47 up_1)
-#56 := (iff #52 #55)
-#57 := [rewrite]: #56
-#53 := (iff #13 #52)
-#50 := (iff #12 #47)
-#43 := (implies up_1 #38)
-#48 := (iff #43 #47)
-#49 := [rewrite]: #48
-#44 := (iff #12 #43)
-#41 := (iff #11 #38)
-#35 := (or #32 #10)
-#39 := (iff #35 #38)
-#40 := [rewrite]: #39
-#36 := (iff #11 #35)
-#33 := (iff #9 #32)
-#34 := [rewrite]: #33
-#37 := [monotonicity #34]: #36
-#42 := [trans #37 #40]: #41
-#45 := [monotonicity #42]: #44
-#51 := [trans #45 #49]: #50
-#54 := [monotonicity #51]: #53
-#59 := [trans #54 #57]: #58
-#62 := [monotonicity #59]: #61
-#68 := [trans #62 #66]: #67
-#71 := [monotonicity #68]: #70
-#82 := [trans #71 #80]: #81
-#31 := [asserted]: #15
-[mp #31 #82]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_prop_07 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,8 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T1)
-:extrapreds (
- (up_1)
- )
-:assumption (not (iff (iff (iff (iff (iff (iff (iff (iff (iff up_1 up_1) up_1) up_1) up_1) up_1) up_1) up_1) up_1) up_1))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_prop_07.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,57 +0,0 @@
-#2 := false
-decl up_1 :: bool
-#4 := up_1
-#5 := (iff up_1 up_1)
-#6 := (iff #5 up_1)
-#7 := (iff #6 up_1)
-#8 := (iff #7 up_1)
-#9 := (iff #8 up_1)
-#10 := (iff #9 up_1)
-#11 := (iff #10 up_1)
-#12 := (iff #11 up_1)
-#13 := (iff #12 up_1)
-#14 := (not #13)
-#69 := (iff #14 false)
-#1 := true
-#64 := (not true)
-#67 := (iff #64 false)
-#68 := [rewrite]: #67
-#65 := (iff #14 #64)
-#62 := (iff #13 true)
-#31 := (iff #5 true)
-#32 := [rewrite]: #31
-#60 := (iff #13 #5)
-#33 := (iff true up_1)
-#36 := (iff #33 up_1)
-#37 := [rewrite]: #36
-#57 := (iff #12 #33)
-#55 := (iff #11 true)
-#53 := (iff #11 #5)
-#50 := (iff #10 #33)
-#48 := (iff #9 true)
-#46 := (iff #9 #5)
-#43 := (iff #8 #33)
-#41 := (iff #7 true)
-#39 := (iff #7 #5)
-#34 := (iff #6 #33)
-#35 := [monotonicity #32]: #34
-#38 := [trans #35 #37]: #7
-#40 := [monotonicity #38]: #39
-#42 := [trans #40 #32]: #41
-#44 := [monotonicity #42]: #43
-#45 := [trans #44 #37]: #9
-#47 := [monotonicity #45]: #46
-#49 := [trans #47 #32]: #48
-#51 := [monotonicity #49]: #50
-#52 := [trans #51 #37]: #11
-#54 := [monotonicity #52]: #53
-#56 := [trans #54 #32]: #55
-#58 := [monotonicity #56]: #57
-#59 := [trans #58 #37]: #13
-#61 := [monotonicity #59]: #60
-#63 := [trans #61 #32]: #62
-#66 := [monotonicity #63]: #65
-#70 := [trans #66 #68]: #69
-#30 := [asserted]: #14
-[mp #30 #70]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_prop_08 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,22 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T1)
-:extrapreds (
- (up_1)
- (up_2)
- (up_3)
- (up_4)
- (up_5)
- (up_6)
- (up_8)
- (up_9)
- (up_7)
- )
-:assumption (or up_1 (or up_2 (or up_3 up_4)))
-:assumption (or up_5 (or up_6 (and up_1 up_4)))
-:assumption (or (not (or up_1 (and up_3 (not up_3)))) up_2)
-:assumption (or (not (and up_2 (or up_7 (not up_7)))) up_3)
-:assumption (or (not (or up_4 false)) up_3)
-:assumption (not (or up_3 (and (not up_8) (or up_8 (and up_9 (not up_9))))))
-:assumption (not false)
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_prop_08.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,194 +0,0 @@
-#2 := false
-decl up_1 :: bool
-#4 := up_1
-#75 := (not up_1)
-#246 := (iff #75 false)
-#1 := true
-#214 := (not true)
-#217 := (iff #214 false)
-#218 := [rewrite]: #217
-#244 := (iff #75 #214)
-#238 := (iff up_1 true)
-#241 := (iff up_1 #238)
-#239 := (iff #238 up_1)
-#240 := [rewrite]: #239
-#242 := [symm #240]: #241
-decl up_4 :: bool
-#7 := up_4
-decl up_2 :: bool
-#5 := up_2
-#161 := (or up_1 up_2 up_4)
-#200 := (iff #161 up_1)
-#195 := (or up_1 false false)
-#198 := (iff #195 up_1)
-#199 := [rewrite]: #198
-#196 := (iff #161 #195)
-#189 := (iff up_4 false)
-#102 := (not up_4)
-#192 := (iff #102 #189)
-#190 := (iff #189 #102)
-#191 := [rewrite]: #190
-#193 := [symm #191]: #192
-decl up_3 :: bool
-#6 := up_3
-#108 := (or up_3 #102)
-#180 := (iff #108 #102)
-#175 := (or false #102)
-#178 := (iff #175 #102)
-#179 := [rewrite]: #178
-#176 := (iff #108 #175)
-#152 := (iff up_3 false)
-#16 := (not up_3)
-#155 := (iff #16 #152)
-#153 := (iff #152 #16)
-#154 := [rewrite]: #153
-#156 := [symm #154]: #155
-decl up_9 :: bool
-#32 := up_9
-#33 := (not up_9)
-#34 := (and up_9 #33)
-decl up_8 :: bool
-#30 := up_8
-#35 := (or up_8 #34)
-#31 := (not up_8)
-#36 := (and #31 #35)
-#37 := (or up_3 #36)
-#38 := (not #37)
-#138 := (iff #38 #16)
-#136 := (iff #37 up_3)
-#131 := (or up_3 false)
-#134 := (iff #131 up_3)
-#135 := [rewrite]: #134
-#132 := (iff #37 #131)
-#129 := (iff #36 false)
-#124 := (and #31 up_8)
-#127 := (iff #124 false)
-#128 := [rewrite]: #127
-#125 := (iff #36 #124)
-#122 := (iff #35 up_8)
-#117 := (or up_8 false)
-#120 := (iff #117 up_8)
-#121 := [rewrite]: #120
-#118 := (iff #35 #117)
-#114 := (iff #34 false)
-#116 := [rewrite]: #114
-#119 := [monotonicity #116]: #118
-#123 := [trans #119 #121]: #122
-#126 := [monotonicity #123]: #125
-#130 := [trans #126 #128]: #129
-#133 := [monotonicity #130]: #132
-#137 := [trans #133 #135]: #136
-#139 := [monotonicity #137]: #138
-#113 := [asserted]: #38
-#142 := [mp #113 #139]: #16
-#157 := [mp #142 #156]: #152
-#177 := [monotonicity #157]: #176
-#181 := [trans #177 #179]: #180
-#27 := (or up_4 false)
-#28 := (not #27)
-#29 := (or #28 up_3)
-#111 := (iff #29 #108)
-#105 := (or #102 up_3)
-#109 := (iff #105 #108)
-#110 := [rewrite]: #109
-#106 := (iff #29 #105)
-#103 := (iff #28 #102)
-#99 := (iff #27 up_4)
-#101 := [rewrite]: #99
-#104 := [monotonicity #101]: #103
-#107 := [monotonicity #104]: #106
-#112 := [trans #107 #110]: #111
-#98 := [asserted]: #29
-#115 := [mp #98 #112]: #108
-#182 := [mp #115 #181]: #102
-#194 := [mp #182 #193]: #189
-#183 := (iff up_2 false)
-#92 := (not up_2)
-#186 := (iff #92 #183)
-#184 := (iff #183 #92)
-#185 := [rewrite]: #184
-#187 := [symm #185]: #186
-#95 := (or #92 up_3)
-#172 := (iff #95 #92)
-#167 := (or #92 false)
-#170 := (iff #167 #92)
-#171 := [rewrite]: #170
-#168 := (iff #95 #167)
-#169 := [monotonicity #157]: #168
-#173 := [trans #169 #171]: #172
-decl up_7 :: bool
-#21 := up_7
-#22 := (not up_7)
-#23 := (or up_7 #22)
-#24 := (and up_2 #23)
-#25 := (not #24)
-#26 := (or #25 up_3)
-#96 := (iff #26 #95)
-#93 := (iff #25 #92)
-#90 := (iff #24 up_2)
-#85 := (and up_2 true)
-#88 := (iff #85 up_2)
-#89 := [rewrite]: #88
-#86 := (iff #24 #85)
-#82 := (iff #23 true)
-#84 := [rewrite]: #82
-#87 := [monotonicity #84]: #86
-#91 := [trans #87 #89]: #90
-#94 := [monotonicity #91]: #93
-#97 := [monotonicity #94]: #96
-#81 := [asserted]: #26
-#100 := [mp #81 #97]: #95
-#174 := [mp #100 #173]: #92
-#188 := [mp #174 #187]: #183
-#197 := [monotonicity #188 #194]: #196
-#201 := [trans #197 #199]: #200
-#58 := (or up_1 up_2 up_3 up_4)
-#164 := (iff #58 #161)
-#158 := (or up_1 up_2 false up_4)
-#162 := (iff #158 #161)
-#163 := [rewrite]: #162
-#159 := (iff #58 #158)
-#160 := [monotonicity #157]: #159
-#165 := [trans #160 #163]: #164
-#8 := (or up_3 up_4)
-#9 := (or up_2 #8)
-#10 := (or up_1 #9)
-#59 := (iff #10 #58)
-#60 := [rewrite]: #59
-#55 := [asserted]: #10
-#61 := [mp #55 #60]: #58
-#166 := [mp #61 #165]: #161
-#202 := [mp #166 #201]: up_1
-#243 := [mp #202 #242]: #238
-#245 := [monotonicity #243]: #244
-#247 := [trans #245 #218]: #246
-#78 := (or #75 up_2)
-#235 := (iff #78 #75)
-#230 := (or #75 false)
-#233 := (iff #230 #75)
-#234 := [rewrite]: #233
-#231 := (iff #78 #230)
-#232 := [monotonicity #188]: #231
-#236 := [trans #232 #234]: #235
-#17 := (and up_3 #16)
-#18 := (or up_1 #17)
-#19 := (not #18)
-#20 := (or #19 up_2)
-#79 := (iff #20 #78)
-#76 := (iff #19 #75)
-#73 := (iff #18 up_1)
-#68 := (or up_1 false)
-#71 := (iff #68 up_1)
-#72 := [rewrite]: #71
-#69 := (iff #18 #68)
-#62 := (iff #17 false)
-#67 := [rewrite]: #62
-#70 := [monotonicity #67]: #69
-#74 := [trans #70 #72]: #73
-#77 := [monotonicity #74]: #76
-#80 := [monotonicity #77]: #79
-#57 := [asserted]: #20
-#83 := [mp #57 #80]: #78
-#237 := [mp #83 #236]: #75
-[mp #237 #247]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_prop_09 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,11 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T1)
-:extrafuns (
- (uf_1 T1 T1 T1)
- (uf_2 T1)
- (uf_3 T1)
- )
-:assumption (forall (?x1 T1) (?x2 T1) (= (uf_1 ?x1 ?x2) (uf_1 ?x2 ?x1)))
-:assumption (not (and (= uf_2 uf_2) (= (uf_1 uf_2 uf_3) (uf_1 uf_3 uf_2))))
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_prop_09.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,52 +0,0 @@
-#2 := false
-decl uf_1 :: (-> T1 T1 T1)
-decl uf_2 :: T1
-#10 := uf_2
-decl uf_3 :: T1
-#12 := uf_3
-#14 := (uf_1 uf_3 uf_2)
-#13 := (uf_1 uf_2 uf_3)
-#15 := (= #13 #14)
-#44 := (not #15)
-#11 := (= uf_2 uf_2)
-#16 := (and #11 #15)
-#17 := (not #16)
-#45 := (iff #17 #44)
-#42 := (iff #16 #15)
-#1 := true
-#37 := (and true #15)
-#40 := (iff #37 #15)
-#41 := [rewrite]: #40
-#38 := (iff #16 #37)
-#35 := (iff #11 true)
-#36 := [rewrite]: #35
-#39 := [monotonicity #36]: #38
-#43 := [trans #39 #41]: #42
-#46 := [monotonicity #43]: #45
-#34 := [asserted]: #17
-#49 := [mp #34 #46]: #44
-#4 := (:var 1 T1)
-#5 := (:var 0 T1)
-#7 := (uf_1 #5 #4)
-#530 := (pattern #7)
-#6 := (uf_1 #4 #5)
-#529 := (pattern #6)
-#8 := (= #6 #7)
-#531 := (forall (vars (?x1 T1) (?x2 T1)) (:pat #529 #530) #8)
-#9 := (forall (vars (?x1 T1) (?x2 T1)) #8)
-#534 := (iff #9 #531)
-#532 := (iff #8 #8)
-#533 := [refl]: #532
-#535 := [quant-intro #533]: #534
-#55 := (~ #9 #9)
-#53 := (~ #8 #8)
-#54 := [refl]: #53
-#56 := [nnf-pos #54]: #55
-#33 := [asserted]: #9
-#57 := [mp~ #33 #56]: #9
-#536 := [mp #57 #535]: #531
-#112 := (not #531)
-#199 := (or #112 #15)
-#113 := [quant-inst]: #199
-[unit-resolution #113 #536 #49]: false
-unsat
--- a/src/HOL/SMT/Examples/cert/z3_prop_10 Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,251 +0,0 @@
-(benchmark Isabelle
-:extrasorts ( T1)
-:extrapreds (
- (up_1)
- (up_5)
- (up_7)
- (up_9)
- (up_11)
- (up_14)
- (up_16)
- (up_18)
- (up_20)
- (up_22)
- (up_25)
- (up_27)
- (up_29)
- (up_31)
- (up_33)
- (up_36)
- (up_38)
- (up_40)
- (up_42)
- (up_44)
- (up_47)
- (up_49)
- (up_51)
- (up_53)
- (up_55)
- (up_57)
- (up_58)
- (up_59)
- (up_60)
- (up_3)
- (up_2)
- (up_6)
- (up_8)
- (up_10)
- (up_12)
- (up_13)
- (up_15)
- (up_17)
- (up_19)
- (up_21)
- (up_23)
- (up_24)
- (up_26)
- (up_28)
- (up_30)
- (up_32)
- (up_34)
- (up_35)
- (up_37)
- (up_39)
- (up_41)
- (up_43)
- (up_45)
- (up_46)
- (up_48)
- (up_50)
- (up_52)
- (up_54)
- (up_56)
- (up_4)
- )
-:assumption (not up_1)
-:assumption (not up_2)
-:assumption (not up_3)
-:assumption (not up_4)
-:assumption (or up_5 (or up_6 up_1))
-:assumption (or up_7 (or up_8 up_5))
-:assumption (or up_9 (or up_10 up_7))
-:assumption (or up_11 (or up_12 up_9))
-:assumption (or up_13 up_11)
-:assumption (or up_14 (or up_15 up_2))
-:assumption (or up_16 (or up_17 (or up_14 up_6)))
-:assumption (or up_18 (or up_19 (or up_16 up_8)))
-:assumption (or up_20 (or up_21 (or up_18 up_10)))
-:assumption (or up_22 (or up_23 (or up_20 up_12)))
-:assumption (or up_24 (or up_22 up_13))
-:assumption (or up_25 (or up_26 up_15))
-:assumption (or up_27 (or up_28 (or up_25 up_17)))
-:assumption (or up_29 (or up_30 (or up_27 up_19)))
-:assumption (or up_31 (or up_32 (or up_29 up_21)))
-:assumption (or up_33 (or up_34 (or up_31 up_23)))
-:assumption (or up_35 (or up_33 up_24))
-:assumption (or up_36 (or up_37 up_26))
-:assumption (or up_38 (or up_39 (or up_36 up_28)))
-:assumption (or up_40 (or up_41 (or up_38 up_30)))
-:assumption (or up_42 (or up_43 (or up_40 up_32)))
-:assumption (or up_44 (or up_45 (or up_42 up_34)))
-:assumption (or up_46 (or up_44 up_35))
-:assumption (or up_47 (or up_48 up_37))
-:assumption (or up_49 (or up_50 (or up_47 up_39)))
-:assumption (or up_51 (or up_52 (or up_49 up_41)))
-:assumption (or up_53 (or up_54 (or up_51 up_43)))
-:assumption (or up_55 (or up_56 (or up_53 up_45)))
-:assumption (or up_4 (or up_55 up_46))
-:assumption (or up_57 up_48)
-:assumption (or up_58 (or up_57 up_50))
-:assumption (or up_59 (or up_58 up_52))
-:assumption (or up_60 (or up_59 up_54))
-:assumption (or up_3 (or up_60 up_56))
-:assumption (or (not up_5) (not up_6))
-:assumption (or (not up_5) (not up_1))
-:assumption (or (not up_6) (not up_1))
-:assumption (or (not up_7) (not up_8))
-:assumption (or (not up_7) (not up_5))
-:assumption (or (not up_8) (not up_5))
-:assumption (or (not up_9) (not up_10))
-:assumption (or (not up_9) (not up_7))
-:assumption (or (not up_10) (not up_7))
-:assumption (or (not up_11) (not up_12))
-:assumption (or (not up_11) (not up_9))
-:assumption (or (not up_12) (not up_9))
-:assumption (or (not up_13) (not up_11))
-:assumption (or (not up_14) (not up_15))
-:assumption (or (not up_14) (not up_2))
-:assumption (or (not up_15) (not up_2))
-:assumption (or (not up_16) (not up_17))
-:assumption (or (not up_16) (not up_14))
-:assumption (or (not up_16) (not up_6))
-:assumption (or (not up_17) (not up_14))
-:assumption (or (not up_17) (not up_6))
-:assumption (or (not up_14) (not up_6))
-:assumption (or (not up_18) (not up_19))
-:assumption (or (not up_18) (not up_16))
-:assumption (or (not up_18) (not up_8))
-:assumption (or (not up_19) (not up_16))
-:assumption (or (not up_19) (not up_8))
-:assumption (or (not up_16) (not up_8))
-:assumption (or (not up_20) (not up_21))
-:assumption (or (not up_20) (not up_18))
-:assumption (or (not up_20) (not up_10))
-:assumption (or (not up_21) (not up_18))
-:assumption (or (not up_21) (not up_10))
-:assumption (or (not up_18) (not up_10))
-:assumption (or (not up_22) (not up_23))
-:assumption (or (not up_22) (not up_20))
-:assumption (or (not up_22) (not up_12))
-:assumption (or (not up_23) (not up_20))
-:assumption (or (not up_23) (not up_12))
-:assumption (or (not up_20) (not up_12))
-:assumption (or (not up_24) (not up_22))
-:assumption (or (not up_24) (not up_13))
-:assumption (or (not up_22) (not up_13))
-:assumption (or (not up_25) (not up_26))
-:assumption (or (not up_25) (not up_15))
-:assumption (or (not up_26) (not up_15))
-:assumption (or (not up_27) (not up_28))
-:assumption (or (not up_27) (not up_25))
-:assumption (or (not up_27) (not up_17))
-:assumption (or (not up_28) (not up_25))
-:assumption (or (not up_28) (not up_17))
-:assumption (or (not up_25) (not up_17))
-:assumption (or (not up_29) (not up_30))
-:assumption (or (not up_29) (not up_27))
-:assumption (or (not up_29) (not up_19))
-:assumption (or (not up_30) (not up_27))
-:assumption (or (not up_30) (not up_19))
-:assumption (or (not up_27) (not up_19))
-:assumption (or (not up_31) (not up_32))
-:assumption (or (not up_31) (not up_29))
-:assumption (or (not up_31) (not up_21))
-:assumption (or (not up_32) (not up_29))
-:assumption (or (not up_32) (not up_21))
-:assumption (or (not up_29) (not up_21))
-:assumption (or (not up_33) (not up_34))
-:assumption (or (not up_33) (not up_31))
-:assumption (or (not up_33) (not up_23))
-:assumption (or (not up_34) (not up_31))
-:assumption (or (not up_34) (not up_23))
-:assumption (or (not up_31) (not up_23))
-:assumption (or (not up_35) (not up_33))
-:assumption (or (not up_35) (not up_24))
-:assumption (or (not up_33) (not up_24))
-:assumption (or (not up_36) (not up_37))
-:assumption (or (not up_36) (not up_26))
-:assumption (or (not up_37) (not up_26))
-:assumption (or (not up_38) (not up_39))
-:assumption (or (not up_38) (not up_36))
-:assumption (or (not up_38) (not up_28))
-:assumption (or (not up_39) (not up_36))
-:assumption (or (not up_39) (not up_28))
-:assumption (or (not up_36) (not up_28))
-:assumption (or (not up_40) (not up_41))
-:assumption (or (not up_40) (not up_38))
-:assumption (or (not up_40) (not up_30))
-:assumption (or (not up_41) (not up_38))
-:assumption (or (not up_41) (not up_30))
-:assumption (or (not up_38) (not up_30))
-:assumption (or (not up_42) (not up_43))
-:assumption (or (not up_42) (not up_40))
-:assumption (or (not up_42) (not up_32))
-:assumption (or (not up_43) (not up_40))
-:assumption (or (not up_43) (not up_32))
-:assumption (or (not up_40) (not up_32))
-:assumption (or (not up_44) (not up_45))
-:assumption (or (not up_44) (not up_42))
-:assumption (or (not up_44) (not up_34))
-:assumption (or (not up_45) (not up_42))
-:assumption (or (not up_45) (not up_34))
-:assumption (or (not up_42) (not up_34))
-:assumption (or (not up_46) (not up_44))
-:assumption (or (not up_46) (not up_35))
-:assumption (or (not up_44) (not up_35))
-:assumption (or (not up_47) (not up_48))
-:assumption (or (not up_47) (not up_37))
-:assumption (or (not up_48) (not up_37))
-:assumption (or (not up_49) (not up_50))
-:assumption (or (not up_49) (not up_47))
-:assumption (or (not up_49) (not up_39))
-:assumption (or (not up_50) (not up_47))
-:assumption (or (not up_50) (not up_39))
-:assumption (or (not up_47) (not up_39))
-:assumption (or (not up_51) (not up_52))
-:assumption (or (not up_51) (not up_49))
-:assumption (or (not up_51) (not up_41))
-:assumption (or (not up_52) (not up_49))
-:assumption (or (not up_52) (not up_41))
-:assumption (or (not up_49) (not up_41))
-:assumption (or (not up_53) (not up_54))
-:assumption (or (not up_53) (not up_51))
-:assumption (or (not up_53) (not up_43))
-:assumption (or (not up_54) (not up_51))
-:assumption (or (not up_54) (not up_43))
-:assumption (or (not up_51) (not up_43))
-:assumption (or (not up_55) (not up_56))
-:assumption (or (not up_55) (not up_53))
-:assumption (or (not up_55) (not up_45))
-:assumption (or (not up_56) (not up_53))
-:assumption (or (not up_56) (not up_45))
-:assumption (or (not up_53) (not up_45))
-:assumption (or (not up_4) (not up_55))
-:assumption (or (not up_4) (not up_46))
-:assumption (or (not up_55) (not up_46))
-:assumption (or (not up_57) (not up_48))
-:assumption (or (not up_58) (not up_57))
-:assumption (or (not up_58) (not up_50))
-:assumption (or (not up_57) (not up_50))
-:assumption (or (not up_59) (not up_58))
-:assumption (or (not up_59) (not up_52))
-:assumption (or (not up_58) (not up_52))
-:assumption (or (not up_60) (not up_59))
-:assumption (or (not up_60) (not up_54))
-:assumption (or (not up_59) (not up_54))
-:assumption (or (not up_3) (not up_60))
-:assumption (or (not up_3) (not up_56))
-:assumption (or (not up_60) (not up_56))
-:assumption (not false)
-:formula true
-)
--- a/src/HOL/SMT/Examples/cert/z3_prop_10.proof Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,1667 +0,0 @@
-#2 := false
-decl up_54 :: bool
-#126 := up_54
-#317 := (not up_54)
-decl up_60 :: bool
-#145 := up_60
-decl up_56 :: bool
-#131 := up_56
-#325 := (not up_56)
-decl up_55 :: bool
-#130 := up_55
-decl up_46 :: bool
-#108 := up_46
-#291 := (not up_46)
-decl up_35 :: bool
-#81 := up_35
-decl up_29 :: bool
-#66 := up_29
-decl up_32 :: bool
-#72 := up_32
-#235 := (not up_32)
-decl up_34 :: bool
-#77 := up_34
-#243 := (not up_34)
-decl up_33 :: bool
-#76 := up_33
-#250 := (not up_35)
-#1611 := [hypothesis]: #250
-decl up_24 :: bool
-#54 := up_24
-#209 := (not up_24)
-decl up_13 :: bool
-#28 := up_13
-decl up_11 :: bool
-#24 := up_11
-#165 := (not up_11)
-decl up_12 :: bool
-#25 := up_12
-#2327 := (or up_12 up_35)
-#345 := (not up_60)
-decl up_59 :: bool
-#142 := up_59
-decl up_19 :: bool
-#40 := up_19
-decl up_8 :: bool
-#17 := up_8
-#156 := (not up_8)
-decl up_7 :: bool
-#16 := up_7
-#166 := (not up_12)
-#1457 := [hypothesis]: #166
-#2183 := (or up_7 up_12 up_35)
-#155 := (not up_7)
-#1612 := [hypothesis]: #155
-decl up_10 :: bool
-#21 := up_10
-#161 := (not up_10)
-decl up_20 :: bool
-#44 := up_20
-decl up_23 :: bool
-#50 := up_23
-#202 := (not up_23)
-#2170 := (or up_34 up_7 up_35 up_12)
-#1605 := [hypothesis]: #243
-#2164 := (or up_29 up_34 up_7 up_35 up_12)
-decl up_42 :: bool
-#98 := up_42
-#275 := (not up_42)
-#226 := (not up_29)
-#907 := [hypothesis]: #226
-#2136 := (or up_29 up_12 up_7 up_35 up_32)
-decl up_22 :: bool
-#49 := up_22
-#895 := [hypothesis]: #235
-#1624 := (or up_29 up_22 up_12 up_32 up_35 up_7)
-decl up_21 :: bool
-#45 := up_21
-decl up_31 :: bool
-#71 := up_31
-#234 := (not up_31)
-decl up_9 :: bool
-#20 := up_9
-#201 := (not up_22)
-#1456 := [hypothesis]: #201
-#847 := (or #161 up_32 up_29 up_22 up_12)
-#193 := (not up_20)
-#1400 := [hypothesis]: up_10
-#964 := (or #161 #193)
-#197 := (or #193 #161)
-#966 := (iff #197 #964)
-#967 := [rewrite]: #966
-#963 := [asserted]: #197
-#970 := [mp #963 #967]: #964
-#1399 := [unit-resolution #970 #1400]: #193
-#500 := (or up_12 up_20 up_22 up_23)
-#51 := (or up_20 up_12)
-#52 := (or up_23 #51)
-#53 := (or up_22 #52)
-#503 := (iff #53 #500)
-#491 := (or up_12 up_20)
-#494 := (or up_23 #491)
-#497 := (or up_22 #494)
-#501 := (iff #497 #500)
-#502 := [rewrite]: #501
-#498 := (iff #53 #497)
-#495 := (iff #52 #494)
-#492 := (iff #51 #491)
-#493 := [rewrite]: #492
-#496 := [monotonicity #493]: #495
-#499 := [monotonicity #496]: #498
-#504 := [trans #499 #502]: #503
-#490 := [asserted]: #53
-#505 := [mp #490 #504]: #500
-#900 := [unit-resolution #505 #1399 #1456 #1457]: up_23
-#194 := (not up_21)
-#974 := (or #161 #194)
-#199 := (or #194 #161)
-#976 := (iff #199 #974)
-#977 := [rewrite]: #976
-#973 := [asserted]: #199
-#980 := [mp #973 #977]: #974
-#902 := [unit-resolution #980 #1400]: #194
-#574 := (or up_21 up_29 up_31 up_32)
-#73 := (or up_29 up_21)
-#74 := (or up_32 #73)
-#75 := (or up_31 #74)
-#577 := (iff #75 #574)
-#565 := (or up_21 up_29)
-#568 := (or up_32 #565)
-#571 := (or up_31 #568)
-#575 := (iff #571 #574)
-#576 := [rewrite]: #575
-#572 := (iff #75 #571)
-#569 := (iff #74 #568)
-#566 := (iff #73 #565)
-#567 := [rewrite]: #566
-#570 := [monotonicity #567]: #569
-#573 := [monotonicity #570]: #572
-#578 := [trans #573 #576]: #577
-#564 := [asserted]: #75
-#579 := [mp #564 #578]: #574
-#851 := [unit-resolution #579 #902 #895 #907]: up_31
-#1135 := (or #202 #234)
-#249 := (or #234 #202)
-#1137 := (iff #249 #1135)
-#1138 := [rewrite]: #1137
-#1134 := [asserted]: #249
-#1141 := [mp #1134 #1138]: #1135
-#858 := [unit-resolution #1141 #851 #900]: false
-#853 := [lemma #858]: #847
-#1613 := [unit-resolution #853 #907 #1456 #895 #1457]: #161
-#405 := (or up_7 up_9 up_10)
-#22 := (or up_10 up_7)
-#23 := (or up_9 #22)
-#408 := (iff #23 #405)
-#399 := (or up_7 up_10)
-#402 := (or up_9 #399)
-#406 := (iff #402 #405)
-#407 := [rewrite]: #406
-#403 := (iff #23 #402)
-#400 := (iff #22 #399)
-#401 := [rewrite]: #400
-#404 := [monotonicity #401]: #403
-#409 := [trans #404 #407]: #408
-#398 := [asserted]: #23
-#410 := [mp #398 #409]: #405
-#1614 := [unit-resolution #410 #1613 #1612]: up_9
-#160 := (not up_9)
-#881 := (or #160 #165)
-#168 := (or #165 #160)
-#882 := (iff #168 #881)
-#883 := [rewrite]: #882
-#879 := [asserted]: #168
-#886 := [mp #879 #883]: #881
-#1615 := [unit-resolution #886 #1614]: #165
-#425 := (or up_11 up_13)
-#29 := (or up_13 up_11)
-#426 := (iff #29 #425)
-#427 := [rewrite]: #426
-#424 := [asserted]: #29
-#430 := [mp #424 #427]: #425
-#1616 := [unit-resolution #430 #1615]: up_13
-#170 := (not up_13)
-#1015 := (or #170 #209)
-#211 := (or #209 #170)
-#1017 := (iff #211 #1015)
-#1018 := [rewrite]: #1017
-#1014 := [asserted]: #211
-#1021 := [mp #1014 #1018]: #1015
-#1617 := [unit-resolution #1021 #1616]: #209
-#603 := (or up_24 up_33 up_35)
-#82 := (or up_33 up_24)
-#83 := (or up_35 #82)
-#606 := (iff #83 #603)
-#597 := (or up_24 up_33)
-#600 := (or up_35 #597)
-#604 := (iff #600 #603)
-#605 := [rewrite]: #604
-#601 := (iff #83 #600)
-#598 := (iff #82 #597)
-#599 := [rewrite]: #598
-#602 := [monotonicity #599]: #601
-#607 := [trans #602 #605]: #606
-#596 := [asserted]: #83
-#608 := [mp #596 #607]: #603
-#1618 := [unit-resolution #608 #1617 #1611]: up_33
-#242 := (not up_33)
-#1116 := (or #234 #242)
-#245 := (or #242 #234)
-#1117 := (iff #245 #1116)
-#1118 := [rewrite]: #1117
-#1114 := [asserted]: #245
-#1121 := [mp #1114 #1118]: #1116
-#1619 := [unit-resolution #1121 #1618]: #234
-#1620 := [unit-resolution #579 #1619 #895 #907]: up_21
-#1120 := (or #202 #242)
-#246 := (or #242 #202)
-#1122 := (iff #246 #1120)
-#1123 := [rewrite]: #1122
-#1119 := [asserted]: #246
-#1126 := [mp #1119 #1123]: #1120
-#1621 := [unit-resolution #1126 #1618]: #202
-#1622 := [unit-resolution #505 #1621 #1456 #1457]: up_20
-#195 := (or #193 #194)
-#957 := [asserted]: #195
-#1623 := [unit-resolution #957 #1622 #1620]: false
-#1625 := [lemma #1623]: #1624
-#2132 := [unit-resolution #1625 #907 #1611 #1457 #895 #1612]: up_22
-#1978 := (or up_32 up_35 up_29 up_21 up_12 up_7)
-#1972 := [unit-resolution #1625 #895 #907 #1457 #1611 #1612]: up_22
-#1010 := (or #201 #209)
-#210 := (or #209 #201)
-#1012 := (iff #210 #1010)
-#1013 := [rewrite]: #1012
-#1009 := [asserted]: #210
-#1016 := [mp #1009 #1013]: #1010
-#1973 := [unit-resolution #1016 #1972]: #209
-#1974 := [hypothesis]: #194
-#1975 := [unit-resolution #579 #895 #907 #1974]: up_31
-#1976 := [unit-resolution #1121 #1975]: #242
-#1977 := [unit-resolution #608 #1976 #1973 #1611]: false
-#1979 := [lemma #1977]: #1978
-#2133 := [unit-resolution #1979 #907 #1611 #1457 #895 #1612]: up_21
-#1682 := (or #194 up_7 up_12 up_23)
-#1673 := [hypothesis]: #202
-#1674 := [hypothesis]: up_21
-#1675 := [unit-resolution #957 #1674]: #193
-#1676 := [unit-resolution #505 #1675 #1457 #1673]: up_22
-#1020 := (or #170 #201)
-#212 := (or #201 #170)
-#1022 := (iff #212 #1020)
-#1023 := [rewrite]: #1022
-#1019 := [asserted]: #212
-#1026 := [mp #1019 #1023]: #1020
-#1677 := [unit-resolution #1026 #1676]: #170
-#1678 := [unit-resolution #980 #1674]: #161
-#1679 := [unit-resolution #410 #1678 #1612]: up_9
-#1680 := [unit-resolution #886 #1679]: #165
-#1681 := [unit-resolution #430 #1680 #1677]: false
-#1683 := [lemma #1681]: #1682
-#2134 := [unit-resolution #1683 #2133 #1457 #1612]: up_23
-#203 := (or #201 #202)
-#983 := [asserted]: #203
-#2135 := [unit-resolution #983 #2134 #2132]: false
-#2137 := [lemma #2135]: #2136
-#2156 := [unit-resolution #2137 #907 #1612 #1611 #1457]: up_32
-#1224 := (or #235 #275)
-#279 := (or #275 #235)
-#1226 := (iff #279 #1224)
-#1227 := [rewrite]: #1226
-#1223 := [asserted]: #279
-#1230 := [mp #1223 #1227]: #1224
-#2157 := [unit-resolution #1230 #2156]: #275
-#2158 := (or up_12 up_29 up_7 up_54)
-decl up_26 :: bool
-#58 := up_26
-#214 := (not up_26)
-decl up_15 :: bool
-#31 := up_15
-decl up_14 :: bool
-#30 := up_14
-#172 := (not up_14)
-decl up_6 :: bool
-#13 := up_6
-decl up_5 :: bool
-#12 := up_5
-#150 := (not up_5)
-decl up_25 :: bool
-#57 := up_25
-#2099 := [hypothesis]: up_5
-#859 := (or #150 #155)
-#158 := (or #155 #150)
-#860 := (iff #158 #859)
-#861 := [rewrite]: #860
-#857 := [asserted]: #158
-#864 := [mp #857 #861]: #859
-#2100 := [unit-resolution #864 #2099]: #155
-#863 := (or #150 #156)
-#159 := (or #156 #150)
-#865 := (iff #159 #863)
-#866 := [rewrite]: #865
-#862 := [asserted]: #159
-#869 := [mp #862 #866]: #863
-#2101 := [unit-resolution #869 #2099]: #156
-#2097 := (or up_12 up_7 up_8)
-#1626 := [hypothesis]: #156
-#2054 := (or up_54 up_7 up_8)
-decl up_16 :: bool
-#34 := up_16
-#1597 := [hypothesis]: #317
-#1888 := (or up_16 up_8 up_7 up_54)
-decl up_45 :: bool
-#104 := up_45
-#284 := (not up_45)
-decl up_52 :: bool
-#121 := up_52
-#309 := (not up_52)
-decl up_51 :: bool
-#120 := up_51
-#177 := (not up_16)
-#1627 := [hypothesis]: #177
-#1733 := (or up_51 up_7 up_54 up_8 up_16)
-decl up_53 :: bool
-#125 := up_53
-#308 := (not up_51)
-#1598 := [hypothesis]: #308
-decl up_43 :: bool
-#99 := up_43
-#276 := (not up_43)
-#1710 := (or up_32 up_16 up_8 up_7 up_51 up_54)
-#1671 := (or up_35 up_16 up_8 up_32 up_7)
-#1655 := (or #166 up_32 up_16 up_8 up_35 up_7)
-#1642 := [hypothesis]: up_12
-#885 := (or #160 #166)
-#169 := (or #166 #160)
-#887 := (iff #169 #885)
-#888 := [rewrite]: #887
-#884 := [asserted]: #169
-#891 := [mp #884 #888]: #885
-#1643 := [unit-resolution #891 #1642]: #160
-#1644 := [unit-resolution #410 #1643 #1612]: up_10
-#1645 := [unit-resolution #980 #1644]: #194
-#167 := (or #165 #166)
-#878 := [asserted]: #167
-#1646 := [unit-resolution #878 #1642]: #165
-#1647 := [unit-resolution #430 #1646]: up_13
-#1648 := [unit-resolution #1021 #1647]: #209
-#1649 := [unit-resolution #608 #1648 #1611]: up_33
-#1650 := [unit-resolution #1121 #1649]: #234
-decl up_18 :: bool
-#39 := up_18
-#185 := (not up_18)
-#979 := (or #161 #185)
-#200 := (or #185 #161)
-#981 := (iff #200 #979)
-#982 := [rewrite]: #981
-#978 := [asserted]: #200
-#985 := [mp #978 #982]: #979
-#1651 := [unit-resolution #985 #1644]: #185
-#468 := (or up_8 up_16 up_18 up_19)
-#41 := (or up_16 up_8)
-#42 := (or up_19 #41)
-#43 := (or up_18 #42)
-#471 := (iff #43 #468)
-#459 := (or up_8 up_16)
-#462 := (or up_19 #459)
-#465 := (or up_18 #462)
-#469 := (iff #465 #468)
-#470 := [rewrite]: #469
-#466 := (iff #43 #465)
-#463 := (iff #42 #462)
-#460 := (iff #41 #459)
-#461 := [rewrite]: #460
-#464 := [monotonicity #461]: #463
-#467 := [monotonicity #464]: #466
-#472 := [trans #467 #470]: #471
-#458 := [asserted]: #43
-#473 := [mp #458 #472]: #468
-#1652 := [unit-resolution #473 #1651 #1627 #1626]: up_19
-#186 := (not up_19)
-#1068 := (or #186 #226)
-#230 := (or #226 #186)
-#1070 := (iff #230 #1068)
-#1071 := [rewrite]: #1070
-#1067 := [asserted]: #230
-#1074 := [mp #1067 #1071]: #1068
-#1653 := [unit-resolution #1074 #1652]: #226
-#1654 := [unit-resolution #579 #1653 #1650 #895 #1645]: false
-#1656 := [lemma #1654]: #1655
-#1657 := [unit-resolution #1656 #1611 #1627 #1626 #895 #1612]: #166
-#1640 := (or up_12 up_35 up_7 up_22 up_16 up_8 up_32)
-#1628 := [unit-resolution #1625 #1457 #1456 #895 #1611 #1612]: up_29
-#1629 := [unit-resolution #1074 #1628]: #186
-#1630 := [unit-resolution #473 #1629 #1627 #1626]: up_18
-#960 := (or #185 #193)
-#196 := (or #193 #185)
-#961 := (iff #196 #960)
-#962 := [rewrite]: #961
-#958 := [asserted]: #196
-#965 := [mp #958 #962]: #960
-#1631 := [unit-resolution #965 #1630]: #193
-#1632 := [unit-resolution #505 #1631 #1456 #1457]: up_23
-#1633 := [unit-resolution #1126 #1632]: #242
-#1634 := [unit-resolution #608 #1633 #1611]: up_24
-#1635 := [unit-resolution #985 #1630]: #161
-#1636 := [unit-resolution #410 #1635 #1612]: up_9
-#1637 := [unit-resolution #886 #1636]: #165
-#1638 := [unit-resolution #430 #1637]: up_13
-#1639 := [unit-resolution #1021 #1638 #1634]: false
-#1641 := [lemma #1639]: #1640
-#1658 := [unit-resolution #1641 #1657 #1612 #1611 #1627 #1626 #895]: up_22
-#1659 := [unit-resolution #1016 #1658]: #209
-#1660 := [unit-resolution #608 #1659 #1611]: up_33
-#1661 := [unit-resolution #1121 #1660]: #234
-#1662 := [unit-resolution #1026 #1658]: #170
-#1663 := [unit-resolution #430 #1662]: up_11
-#1664 := [unit-resolution #886 #1663]: #160
-#1665 := [unit-resolution #410 #1664 #1612]: up_10
-#1666 := [unit-resolution #980 #1665]: #194
-#1667 := [unit-resolution #579 #1666 #895 #1661]: up_29
-#1668 := [unit-resolution #985 #1665]: #185
-#1669 := [unit-resolution #473 #1668 #1627 #1626]: up_19
-#1670 := [unit-resolution #1074 #1669 #1667]: false
-#1672 := [lemma #1670]: #1671
-#1698 := [unit-resolution #1672 #895 #1626 #1627 #1612]: up_35
-#1609 := (or #250 up_34 up_51 up_54)
-#316 := (not up_53)
-#1599 := [hypothesis]: up_35
-#1275 := (or #250 #291)
-#293 := (or #291 #250)
-#1277 := (iff #293 #1275)
-#1278 := [rewrite]: #1277
-#1274 := [asserted]: #293
-#1281 := [mp #1274 #1278]: #1275
-#1600 := [unit-resolution #1281 #1599]: #291
-#777 := (or up_46 up_55)
-decl up_4 :: bool
-#10 := up_4
-#783 := (or up_4 up_46 up_55)
-#1514 := (iff #783 #777)
-#1509 := (or false up_46 up_55)
-#1512 := (iff #1509 #777)
-#1513 := [rewrite]: #1512
-#1510 := (iff #783 #1509)
-#1485 := (iff up_4 false)
-#11 := (not up_4)
-#1488 := (iff #11 #1485)
-#1486 := (iff #1485 #11)
-#1487 := [rewrite]: #1486
-#1489 := [symm #1487]: #1488
-#371 := [asserted]: #11
-#1490 := [mp #371 #1489]: #1485
-#1511 := [monotonicity #1490]: #1510
-#1515 := [trans #1511 #1513]: #1514
-#135 := (or up_55 up_46)
-#136 := (or up_4 #135)
-#786 := (iff #136 #783)
-#780 := (or up_4 #777)
-#784 := (iff #780 #783)
-#785 := [rewrite]: #784
-#781 := (iff #136 #780)
-#778 := (iff #135 #777)
-#779 := [rewrite]: #778
-#782 := [monotonicity #779]: #781
-#787 := [trans #782 #785]: #786
-#776 := [asserted]: #136
-#788 := [mp #776 #787]: #783
-#1516 := [mp #788 #1515]: #777
-#1601 := [unit-resolution #1516 #1600]: up_55
-#324 := (not up_55)
-#1376 := (or #316 #324)
-#327 := (or #324 #316)
-#1377 := (iff #327 #1376)
-#1378 := [rewrite]: #1377
-#1374 := [asserted]: #327
-#1381 := [mp #1374 #1378]: #1376
-#1602 := [unit-resolution #1381 #1601]: #316
-#754 := (or up_43 up_51 up_53 up_54)
-#127 := (or up_51 up_43)
-#128 := (or up_54 #127)
-#129 := (or up_53 #128)
-#757 := (iff #129 #754)
-#745 := (or up_43 up_51)
-#748 := (or up_54 #745)
-#751 := (or up_53 #748)
-#755 := (iff #751 #754)
-#756 := [rewrite]: #755
-#752 := (iff #129 #751)
-#749 := (iff #128 #748)
-#746 := (iff #127 #745)
-#747 := [rewrite]: #746
-#750 := [monotonicity #747]: #749
-#753 := [monotonicity #750]: #752
-#758 := [trans #753 #756]: #757
-#744 := [asserted]: #129
-#759 := [mp #744 #758]: #754
-#1603 := [unit-resolution #759 #1602 #1598 #1597]: up_43
-decl up_44 :: bool
-#103 := up_44
-#283 := (not up_44)
-#1280 := (or #250 #283)
-#294 := (or #283 #250)
-#1282 := (iff #294 #1280)
-#1283 := [rewrite]: #1282
-#1279 := [asserted]: #294
-#1286 := [mp #1279 #1283]: #1280
-#1604 := [unit-resolution #1286 #1599]: #283
-#1380 := (or #284 #324)
-#328 := (or #324 #284)
-#1382 := (iff #328 #1380)
-#1383 := [rewrite]: #1382
-#1379 := [asserted]: #328
-#1386 := [mp #1379 #1383]: #1380
-#1606 := [unit-resolution #1386 #1601]: #284
-#680 := (or up_34 up_42 up_44 up_45)
-#105 := (or up_42 up_34)
-#106 := (or up_45 #105)
-#107 := (or up_44 #106)
-#683 := (iff #107 #680)
-#671 := (or up_34 up_42)
-#674 := (or up_45 #671)
-#677 := (or up_44 #674)
-#681 := (iff #677 #680)
-#682 := [rewrite]: #681
-#678 := (iff #107 #677)
-#675 := (iff #106 #674)
-#672 := (iff #105 #671)
-#673 := [rewrite]: #672
-#676 := [monotonicity #673]: #675
-#679 := [monotonicity #676]: #678
-#684 := [trans #679 #682]: #683
-#670 := [asserted]: #107
-#685 := [mp #670 #684]: #680
-#1607 := [unit-resolution #685 #1606 #1605 #1604]: up_42
-#277 := (or #275 #276)
-#1217 := [asserted]: #277
-#1608 := [unit-resolution #1217 #1607 #1603]: false
-#1610 := [lemma #1608]: #1609
-#1699 := [unit-resolution #1610 #1698 #1598 #1597]: up_34
-#1125 := (or #234 #243)
-#247 := (or #243 #234)
-#1127 := (iff #247 #1125)
-#1128 := [rewrite]: #1127
-#1124 := [asserted]: #247
-#1131 := [mp #1124 #1128]: #1125
-#1700 := [unit-resolution #1131 #1699]: #234
-#1130 := (or #202 #243)
-#248 := (or #243 #202)
-#1132 := (iff #248 #1130)
-#1133 := [rewrite]: #1132
-#1129 := [asserted]: #248
-#1136 := [mp #1129 #1133]: #1130
-#1701 := [unit-resolution #1136 #1699]: #202
-#1696 := (or up_12 up_7 up_23 up_16 up_8 up_32 up_31)
-#1684 := [hypothesis]: #234
-#1685 := [unit-resolution #1683 #1457 #1612 #1673]: #194
-#1686 := [unit-resolution #579 #1685 #895 #1684]: up_29
-#1687 := [unit-resolution #1074 #1686]: #186
-#1688 := [unit-resolution #473 #1687 #1627 #1626]: up_18
-#1689 := [unit-resolution #965 #1688]: #193
-#1690 := [unit-resolution #505 #1689 #1457 #1673]: up_22
-#1691 := [unit-resolution #1026 #1690]: #170
-#1692 := [unit-resolution #985 #1688]: #161
-#1693 := [unit-resolution #410 #1692 #1612]: up_9
-#1694 := [unit-resolution #886 #1693]: #165
-#1695 := [unit-resolution #430 #1694 #1691]: false
-#1697 := [lemma #1695]: #1696
-#1702 := [unit-resolution #1697 #1701 #1612 #1627 #1626 #895 #1700]: up_12
-#1703 := [unit-resolution #891 #1702]: #160
-#1704 := [unit-resolution #410 #1703 #1612]: up_10
-#1705 := [unit-resolution #980 #1704]: #194
-#1706 := [unit-resolution #579 #1705 #895 #1700]: up_29
-#1707 := [unit-resolution #985 #1704]: #185
-#1708 := [unit-resolution #473 #1707 #1627 #1626]: up_19
-#1709 := [unit-resolution #1074 #1708 #1706]: false
-#1711 := [lemma #1709]: #1710
-#1712 := [unit-resolution #1711 #1598 #1626 #1612 #1627 #1597]: up_32
-#1234 := (or #235 #276)
-#281 := (or #276 #235)
-#1236 := (iff #281 #1234)
-#1237 := [rewrite]: #1236
-#1233 := [asserted]: #281
-#1240 := [mp #1233 #1237]: #1234
-#1713 := [unit-resolution #1240 #1712]: #276
-#1714 := [unit-resolution #759 #1713 #1598 #1597]: up_53
-#1395 := (or #284 #316)
-#331 := (or #316 #284)
-#1397 := (iff #331 #1395)
-#1398 := [rewrite]: #1397
-#1394 := [asserted]: #331
-#1401 := [mp #1394 #1398]: #1395
-#1715 := [unit-resolution #1401 #1714]: #284
-#1716 := [unit-resolution #1230 #1712]: #275
-#1717 := [unit-resolution #1381 #1714]: #324
-#1718 := [unit-resolution #1516 #1717]: up_46
-#1270 := (or #283 #291)
-#292 := (or #291 #283)
-#1272 := (iff #292 #1270)
-#1273 := [rewrite]: #1272
-#1269 := [asserted]: #292
-#1276 := [mp #1269 #1273]: #1270
-#1719 := [unit-resolution #1276 #1718]: #283
-#1720 := [unit-resolution #685 #1719 #1716 #1715]: up_34
-#1721 := [unit-resolution #1136 #1720]: #202
-#1722 := [unit-resolution #1281 #1718]: #250
-#244 := (or #242 #243)
-#1113 := [asserted]: #244
-#1723 := [unit-resolution #1113 #1720]: #242
-#1724 := [unit-resolution #608 #1723 #1722]: up_24
-#1725 := [unit-resolution #1016 #1724]: #201
-#1726 := [unit-resolution #1021 #1724]: #170
-#1727 := [unit-resolution #430 #1726]: up_11
-#1728 := [unit-resolution #878 #1727]: #166
-#1729 := [unit-resolution #505 #1728 #1725 #1721]: up_20
-#1730 := [unit-resolution #886 #1727]: #160
-#1731 := [unit-resolution #410 #1730 #1612]: up_10
-#1732 := [unit-resolution #970 #1731 #1729]: false
-#1734 := [lemma #1732]: #1733
-#1858 := [unit-resolution #1734 #1627 #1597 #1626 #1612]: up_51
-#310 := (or #308 #309)
-#1321 := [asserted]: #310
-#1859 := [unit-resolution #1321 #1858]: #309
-decl up_58 :: bool
-#139 := up_58
-#337 := (not up_58)
-decl up_49 :: bool
-#115 := up_49
-#300 := (not up_49)
-#1324 := (or #300 #308)
-#311 := (or #308 #300)
-#1325 := (iff #311 #1324)
-#1326 := [rewrite]: #1325
-#1322 := [asserted]: #311
-#1329 := [mp #1322 #1326]: #1324
-#1860 := [unit-resolution #1329 #1858]: #300
-decl up_39 :: bool
-#89 := up_39
-#260 := (not up_39)
-decl up_38 :: bool
-#88 := up_38
-decl up_40 :: bool
-#93 := up_40
-#267 := (not up_40)
-decl up_41 :: bool
-#94 := up_41
-#268 := (not up_41)
-#1328 := (or #268 #308)
-#312 := (or #308 #268)
-#1330 := (iff #312 #1328)
-#1331 := [rewrite]: #1330
-#1327 := [asserted]: #312
-#1334 := [mp #1327 #1331]: #1328
-#1861 := [unit-resolution #1334 #1858]: #268
-#1771 := (or up_32 up_16 up_8 up_41 up_49 up_52 up_7)
-#1735 := [unit-resolution #1281 #1698]: #291
-#1736 := [unit-resolution #1516 #1735]: up_55
-#1737 := [unit-resolution #1386 #1736]: #284
-#1738 := [unit-resolution #1286 #1698]: #283
-#259 := (not up_38)
-decl up_50 :: bool
-#116 := up_50
-#301 := (not up_50)
-#1739 := [hypothesis]: #309
-#341 := (not up_59)
-#326 := (or #324 #325)
-#1373 := [asserted]: #326
-#1740 := [unit-resolution #1373 #1736]: #325
-#834 := (or up_56 up_60)
-decl up_3 :: bool
-#8 := up_3
-#840 := (or up_3 up_56 up_60)
-#1522 := (iff #840 #834)
-#1517 := (or false up_56 up_60)
-#1520 := (iff #1517 #834)
-#1521 := [rewrite]: #1520
-#1518 := (iff #840 #1517)
-#1479 := (iff up_3 false)
-#9 := (not up_3)
-#1482 := (iff #9 #1479)
-#1480 := (iff #1479 #9)
-#1481 := [rewrite]: #1480
-#1483 := [symm #1481]: #1482
-#370 := [asserted]: #9
-#1484 := [mp #370 #1483]: #1479
-#1519 := [monotonicity #1484]: #1518
-#1523 := [trans #1519 #1521]: #1522
-#148 := (or up_60 up_56)
-#149 := (or up_3 #148)
-#843 := (iff #149 #840)
-#837 := (or up_3 #834)
-#841 := (iff #837 #840)
-#842 := [rewrite]: #841
-#838 := (iff #149 #837)
-#835 := (iff #148 #834)
-#836 := [rewrite]: #835
-#839 := [monotonicity #836]: #838
-#844 := [trans #839 #842]: #843
-#833 := [asserted]: #149
-#845 := [mp #833 #844]: #840
-#1524 := [mp #845 #1523]: #834
-#1741 := [unit-resolution #1524 #1740]: up_60
-#1442 := (or #341 #345)
-#346 := (or #345 #341)
-#1444 := (iff #346 #1442)
-#1445 := [rewrite]: #1444
-#1441 := [asserted]: #346
-#1448 := [mp #1441 #1445]: #1442
-#1742 := [unit-resolution #1448 #1741]: #341
-#814 := (or up_52 up_58 up_59)
-#143 := (or up_58 up_52)
-#144 := (or up_59 #143)
-#817 := (iff #144 #814)
-#808 := (or up_52 up_58)
-#811 := (or up_59 #808)
-#815 := (iff #811 #814)
-#816 := [rewrite]: #815
-#812 := (iff #144 #811)
-#809 := (iff #143 #808)
-#810 := [rewrite]: #809
-#813 := [monotonicity #810]: #812
-#818 := [trans #813 #816]: #817
-#807 := [asserted]: #144
-#819 := [mp #807 #818]: #814
-#1743 := [unit-resolution #819 #1742 #1739]: up_58
-#1417 := (or #301 #337)
-#339 := (or #337 #301)
-#1419 := (iff #339 #1417)
-#1420 := [rewrite]: #1419
-#1416 := [asserted]: #339
-#1423 := [mp #1416 #1420]: #1417
-#1744 := [unit-resolution #1423 #1743]: #301
-#1745 := [hypothesis]: #300
-decl up_47 :: bool
-#111 := up_47
-#295 := (not up_47)
-decl up_48 :: bool
-#112 := up_48
-decl up_57 :: bool
-#137 := up_57
-#335 := (not up_57)
-#1412 := (or #335 #337)
-#338 := (or #337 #335)
-#1414 := (iff #338 #1412)
-#1415 := [rewrite]: #1414
-#1411 := [asserted]: #338
-#1418 := [mp #1411 #1415]: #1412
-#1746 := [unit-resolution #1418 #1743]: #335
-#790 := (or up_48 up_57)
-#138 := (or up_57 up_48)
-#791 := (iff #138 #790)
-#792 := [rewrite]: #791
-#789 := [asserted]: #138
-#795 := [mp #789 #792]: #790
-#1747 := [unit-resolution #795 #1746]: up_48
-#296 := (not up_48)
-#297 := (or #295 #296)
-#1284 := [asserted]: #297
-#1748 := [unit-resolution #1284 #1747]: #295
-#722 := (or up_39 up_47 up_49 up_50)
-#117 := (or up_47 up_39)
-#118 := (or up_50 #117)
-#119 := (or up_49 #118)
-#725 := (iff #119 #722)
-#713 := (or up_39 up_47)
-#716 := (or up_50 #713)
-#719 := (or up_49 #716)
-#723 := (iff #719 #722)
-#724 := [rewrite]: #723
-#720 := (iff #119 #719)
-#717 := (iff #118 #716)
-#714 := (iff #117 #713)
-#715 := [rewrite]: #714
-#718 := [monotonicity #715]: #717
-#721 := [monotonicity #718]: #720
-#726 := [trans #721 #724]: #725
-#712 := [asserted]: #119
-#727 := [mp #712 #726]: #722
-#1749 := [unit-resolution #727 #1748 #1745 #1744]: up_39
-#261 := (or #259 #260)
-#1165 := [asserted]: #261
-#1750 := [unit-resolution #1165 #1749]: #259
-#1751 := [hypothesis]: #268
-decl up_30 :: bool
-#67 := up_30
-#227 := (not up_30)
-decl up_27 :: bool
-#61 := up_27
-#213 := (not up_25)
-decl up_37 :: bool
-#85 := up_37
-#255 := (not up_37)
-#1291 := (or #255 #296)
-#299 := (or #296 #255)
-#1293 := (iff #299 #1291)
-#1294 := [rewrite]: #1293
-#1290 := [asserted]: #299
-#1297 := [mp #1290 #1294]: #1291
-#1752 := [unit-resolution #1297 #1747]: #255
-decl up_36 :: bool
-#84 := up_36
-#254 := (not up_36)
-#1177 := (or #254 #260)
-#264 := (or #260 #254)
-#1179 := (iff #264 #1177)
-#1180 := [rewrite]: #1179
-#1176 := [asserted]: #264
-#1183 := [mp #1176 #1180]: #1177
-#1753 := [unit-resolution #1183 #1749]: #254
-#616 := (or up_26 up_36 up_37)
-#86 := (or up_37 up_26)
-#87 := (or up_36 #86)
-#619 := (iff #87 #616)
-#610 := (or up_26 up_37)
-#613 := (or up_36 #610)
-#617 := (iff #613 #616)
-#618 := [rewrite]: #617
-#614 := (iff #87 #613)
-#611 := (iff #86 #610)
-#612 := [rewrite]: #611
-#615 := [monotonicity #612]: #614
-#620 := [trans #615 #618]: #619
-#609 := [asserted]: #87
-#621 := [mp #609 #620]: #616
-#1754 := [unit-resolution #621 #1753 #1752]: up_26
-#215 := (or #213 #214)
-#1024 := [asserted]: #215
-#1755 := [unit-resolution #1024 #1754]: #213
-decl up_28 :: bool
-#62 := up_28
-#219 := (not up_28)
-#1182 := (or #219 #260)
-#265 := (or #260 #219)
-#1184 := (iff #265 #1182)
-#1185 := [rewrite]: #1184
-#1181 := [asserted]: #265
-#1188 := [mp #1181 #1185]: #1182
-#1756 := [unit-resolution #1188 #1749]: #219
-decl up_17 :: bool
-#35 := up_17
-#178 := (not up_17)
-#173 := (not up_15)
-#1031 := (or #173 #214)
-#217 := (or #214 #173)
-#1033 := (iff #217 #1031)
-#1034 := [rewrite]: #1033
-#1030 := [asserted]: #217
-#1037 := [mp #1030 #1034]: #1031
-#1757 := [unit-resolution #1037 #1754]: #173
-#1503 := (or up_14 up_15)
-decl up_2 :: bool
-#6 := up_2
-#436 := (or up_2 up_14 up_15)
-#1506 := (iff #436 #1503)
-#1500 := (or false up_14 up_15)
-#1504 := (iff #1500 #1503)
-#1505 := [rewrite]: #1504
-#1501 := (iff #436 #1500)
-#1473 := (iff up_2 false)
-#7 := (not up_2)
-#1476 := (iff #7 #1473)
-#1474 := (iff #1473 #7)
-#1475 := [rewrite]: #1474
-#1477 := [symm #1475]: #1476
-#369 := [asserted]: #7
-#1478 := [mp #369 #1477]: #1473
-#1502 := [monotonicity #1478]: #1501
-#1507 := [trans #1502 #1505]: #1506
-#32 := (or up_15 up_2)
-#33 := (or up_14 #32)
-#439 := (iff #33 #436)
-#429 := (or up_2 up_15)
-#433 := (or up_14 #429)
-#437 := (iff #433 #436)
-#438 := [rewrite]: #437
-#434 := (iff #33 #433)
-#431 := (iff #32 #429)
-#432 := [rewrite]: #431
-#435 := [monotonicity #432]: #434
-#440 := [trans #435 #438]: #439
-#428 := [asserted]: #33
-#441 := [mp #428 #440]: #436
-#1508 := [mp #441 #1507]: #1503
-#1758 := [unit-resolution #1508 #1757]: up_14
-#917 := (or #172 #178)
-#182 := (or #178 #172)
-#919 := (iff #182 #917)
-#920 := [rewrite]: #919
-#916 := [asserted]: #182
-#923 := [mp #916 #920]: #917
-#1759 := [unit-resolution #923 #1758]: #178
-#542 := (or up_17 up_25 up_27 up_28)
-#63 := (or up_25 up_17)
-#64 := (or up_28 #63)
-#65 := (or up_27 #64)
-#545 := (iff #65 #542)
-#533 := (or up_17 up_25)
-#536 := (or up_28 #533)
-#539 := (or up_27 #536)
-#543 := (iff #539 #542)
-#544 := [rewrite]: #543
-#540 := (iff #65 #539)
-#537 := (iff #64 #536)
-#534 := (iff #63 #533)
-#535 := [rewrite]: #534
-#538 := [monotonicity #535]: #537
-#541 := [monotonicity #538]: #540
-#546 := [trans #541 #544]: #545
-#532 := [asserted]: #65
-#547 := [mp #532 #546]: #542
-#1760 := [unit-resolution #547 #1759 #1756 #1755]: up_27
-#218 := (not up_27)
-#1073 := (or #218 #227)
-#231 := (or #227 #218)
-#1075 := (iff #231 #1073)
-#1076 := [rewrite]: #1075
-#1072 := [asserted]: #231
-#1079 := [mp #1072 #1076]: #1073
-#1761 := [unit-resolution #1079 #1760]: #227
-#648 := (or up_30 up_38 up_40 up_41)
-#95 := (or up_38 up_30)
-#96 := (or up_41 #95)
-#97 := (or up_40 #96)
-#651 := (iff #97 #648)
-#639 := (or up_30 up_38)
-#642 := (or up_41 #639)
-#645 := (or up_40 #642)
-#649 := (iff #645 #648)
-#650 := [rewrite]: #649
-#646 := (iff #97 #645)
-#643 := (iff #96 #642)
-#640 := (iff #95 #639)
-#641 := [rewrite]: #640
-#644 := [monotonicity #641]: #643
-#647 := [monotonicity #644]: #646
-#652 := [trans #647 #650]: #651
-#638 := [asserted]: #97
-#653 := [mp #638 #652]: #648
-#1762 := [unit-resolution #653 #1761 #1751 #1750]: up_40
-#1220 := (or #267 #275)
-#278 := (or #275 #267)
-#1221 := (iff #278 #1220)
-#1222 := [rewrite]: #1221
-#1218 := [asserted]: #278
-#1225 := [mp #1218 #1222]: #1220
-#1763 := [unit-resolution #1225 #1762]: #275
-#1764 := [unit-resolution #685 #1763 #1738 #1737]: up_34
-#1064 := (or #218 #226)
-#229 := (or #226 #218)
-#1065 := (iff #229 #1064)
-#1066 := [rewrite]: #1065
-#1062 := [asserted]: #229
-#1069 := [mp #1062 #1066]: #1064
-#1765 := [unit-resolution #1069 #1760]: #226
-#1083 := (or #186 #218)
-#233 := (or #218 #186)
-#1085 := (iff #233 #1083)
-#1086 := [rewrite]: #1085
-#1082 := [asserted]: #233
-#1089 := [mp #1082 #1086]: #1083
-#1766 := [unit-resolution #1089 #1760]: #186
-#1767 := [unit-resolution #473 #1766 #1627 #1626]: up_18
-#969 := (or #185 #194)
-#198 := (or #194 #185)
-#971 := (iff #198 #969)
-#972 := [rewrite]: #971
-#968 := [asserted]: #198
-#975 := [mp #968 #972]: #969
-#1768 := [unit-resolution #975 #1767]: #194
-#1769 := [unit-resolution #579 #1768 #895 #1765]: up_31
-#1770 := [unit-resolution #1131 #1769 #1764]: false
-#1772 := [lemma #1770]: #1771
-#1862 := [unit-resolution #1772 #1627 #1626 #1861 #1860 #1859 #1612]: up_32
-#1239 := (or #235 #267)
-#282 := (or #267 #235)
-#1241 := (iff #282 #1239)
-#1242 := [rewrite]: #1241
-#1238 := [asserted]: #282
-#1245 := [mp #1238 #1242]: #1239
-#1863 := [unit-resolution #1245 #1862]: #267
-#1856 := (or up_12 up_52 up_49 up_41 up_16 up_8 up_7)
-#1828 := [unit-resolution #1772 #1627 #1626 #1751 #1745 #1739 #1612]: up_32
-#1829 := [unit-resolution #1245 #1828]: #267
-#1830 := [unit-resolution #1230 #1828]: #275
-#1826 := (or #170 up_41 up_40 up_16 up_8 up_49 up_12 up_52 up_42)
-#1804 := [hypothesis]: up_13
-#1805 := [unit-resolution #1026 #1804]: #201
-#1806 := [unit-resolution #1021 #1804]: #209
-#1798 := [hypothesis]: #275
-#1782 := [hypothesis]: #267
-#1802 := (or #242 up_42 up_52 up_49 up_41 up_40 up_16 up_8 up_12 up_22)
-#1783 := [hypothesis]: up_33
-#1784 := [unit-resolution #1126 #1783]: #202
-#1785 := [unit-resolution #505 #1784 #1457 #1456]: up_20
-#1786 := [unit-resolution #965 #1785]: #185
-#1787 := [unit-resolution #473 #1786 #1627 #1626]: up_19
-#1078 := (or #186 #227)
-#232 := (or #227 #186)
-#1080 := (iff #232 #1078)
-#1081 := [rewrite]: #1080
-#1077 := [asserted]: #232
-#1084 := [mp #1077 #1081]: #1078
-#1788 := [unit-resolution #1084 #1787]: #227
-#1789 := [unit-resolution #653 #1788 #1751 #1782]: up_38
-#1790 := [unit-resolution #1165 #1789]: #260
-#1780 := (or #337 up_49 up_39)
-#1773 := [hypothesis]: up_58
-#1774 := [unit-resolution #1418 #1773]: #335
-#1775 := [unit-resolution #795 #1774]: up_48
-#1776 := [hypothesis]: #260
-#1777 := [unit-resolution #1423 #1773]: #301
-#1778 := [unit-resolution #727 #1777 #1745 #1776]: up_47
-#1779 := [unit-resolution #1284 #1778 #1775]: false
-#1781 := [lemma #1779]: #1780
-#1791 := [unit-resolution #1781 #1790 #1745]: #337
-#1792 := [unit-resolution #819 #1791 #1739]: up_59
-#1793 := [unit-resolution #1448 #1792]: #345
-#1794 := [unit-resolution #1524 #1793]: up_56
-#1795 := [unit-resolution #1373 #1794]: #324
-#1796 := [unit-resolution #1516 #1795]: up_46
-#1797 := [unit-resolution #1113 #1783]: #243
-#1390 := (or #284 #325)
-#330 := (or #325 #284)
-#1392 := (iff #330 #1390)
-#1393 := [rewrite]: #1392
-#1389 := [asserted]: #330
-#1396 := [mp #1389 #1393]: #1390
-#1799 := [unit-resolution #1396 #1794]: #284
-#1800 := [unit-resolution #685 #1799 #1798 #1797]: up_44
-#1801 := [unit-resolution #1276 #1800 #1796]: false
-#1803 := [lemma #1801]: #1802
-#1807 := [unit-resolution #1803 #1805 #1739 #1745 #1751 #1782 #1627 #1626 #1457 #1798]: #242
-#1808 := [unit-resolution #608 #1807 #1806]: up_35
-#1809 := [unit-resolution #1286 #1808]: #283
-#1810 := [unit-resolution #1281 #1808]: #291
-#1811 := [unit-resolution #1516 #1810]: up_55
-#1812 := [unit-resolution #1386 #1811]: #284
-#1813 := [unit-resolution #685 #1812 #1798 #1809]: up_34
-#1814 := [unit-resolution #1136 #1813]: #202
-#1815 := [unit-resolution #505 #1814 #1457 #1805]: up_20
-#1816 := [unit-resolution #965 #1815]: #185
-#1817 := [unit-resolution #473 #1816 #1627 #1626]: up_19
-#1818 := [unit-resolution #1373 #1811]: #325
-#1819 := [unit-resolution #1524 #1818]: up_60
-#1820 := [unit-resolution #1448 #1819]: #341
-#1821 := [unit-resolution #819 #1820 #1739]: up_58
-#1822 := [unit-resolution #1781 #1821 #1745]: up_39
-#1823 := [unit-resolution #1165 #1822]: #259
-#1824 := [unit-resolution #653 #1823 #1751 #1782]: up_30
-#1825 := [unit-resolution #1084 #1824 #1817]: false
-#1827 := [lemma #1825]: #1826
-#1831 := [unit-resolution #1827 #1457 #1829 #1627 #1626 #1745 #1751 #1739 #1830]: #170
-#1832 := [unit-resolution #430 #1831]: up_11
-#1833 := [unit-resolution #886 #1832]: #160
-#1834 := [unit-resolution #410 #1833 #1612]: up_10
-#1835 := [unit-resolution #985 #1834]: #185
-#1836 := [unit-resolution #473 #1835 #1627 #1626]: up_19
-#1837 := [unit-resolution #1084 #1836]: #227
-#1838 := [unit-resolution #653 #1837 #1751 #1829]: up_38
-#1839 := [unit-resolution #1165 #1838]: #260
-#1840 := [unit-resolution #1781 #1839 #1745]: #337
-#1841 := [unit-resolution #819 #1840 #1739]: up_59
-#1842 := [unit-resolution #1448 #1841]: #345
-#1843 := [unit-resolution #1524 #1842]: up_56
-#1844 := [unit-resolution #1373 #1843]: #324
-#1845 := [unit-resolution #1516 #1844]: up_46
-#1846 := [unit-resolution #1281 #1845]: #250
-#1847 := [unit-resolution #1396 #1843]: #284
-#1848 := [unit-resolution #1276 #1845]: #283
-#1849 := [unit-resolution #685 #1848 #1830 #1847]: up_34
-#1850 := [unit-resolution #1113 #1849]: #242
-#1851 := [unit-resolution #608 #1850 #1846]: up_24
-#1852 := [unit-resolution #970 #1834]: #193
-#1853 := [unit-resolution #1136 #1849]: #202
-#1854 := [unit-resolution #505 #1853 #1457 #1852]: up_22
-#1855 := [unit-resolution #1016 #1854 #1851]: false
-#1857 := [lemma #1855]: #1856
-#1864 := [unit-resolution #1857 #1859 #1860 #1861 #1627 #1626 #1612]: up_12
-#1865 := [unit-resolution #891 #1864]: #160
-#1866 := [unit-resolution #410 #1865 #1612]: up_10
-#1867 := [unit-resolution #985 #1866]: #185
-#1868 := [unit-resolution #473 #1867 #1627 #1626]: up_19
-#1869 := [unit-resolution #1084 #1868]: #227
-#1870 := [unit-resolution #653 #1869 #1861 #1863]: up_38
-#1871 := [unit-resolution #1165 #1870]: #260
-#1872 := [unit-resolution #1781 #1871 #1860]: #337
-#1873 := [unit-resolution #819 #1872 #1859]: up_59
-#1874 := [unit-resolution #1448 #1873]: #345
-#1875 := [unit-resolution #1524 #1874]: up_56
-#1876 := [unit-resolution #1396 #1875]: #284
-#1877 := [unit-resolution #1230 #1862]: #275
-#1878 := [unit-resolution #1373 #1875]: #324
-#1879 := [unit-resolution #1516 #1878]: up_46
-#1880 := [unit-resolution #1276 #1879]: #283
-#1881 := [unit-resolution #685 #1880 #1877 #1876]: up_34
-#1882 := [unit-resolution #878 #1864]: #165
-#1883 := [unit-resolution #430 #1882]: up_13
-#1884 := [unit-resolution #1021 #1883]: #209
-#1885 := [unit-resolution #1281 #1879]: #250
-#1886 := [unit-resolution #608 #1885 #1884]: up_33
-#1887 := [unit-resolution #1113 #1886 #1881]: false
-#1889 := [lemma #1887]: #1888
-#2026 := [unit-resolution #1889 #1597 #1612 #1626]: up_16
-#908 := (or #172 #177)
-#180 := (or #177 #172)
-#909 := (iff #180 #908)
-#910 := [rewrite]: #909
-#906 := [asserted]: #180
-#913 := [mp #906 #910]: #908
-#2027 := [unit-resolution #913 #2026]: #172
-#2028 := [unit-resolution #1508 #2027]: up_15
-#2029 := [unit-resolution #1037 #2028]: #214
-#1027 := (or #173 #213)
-#216 := (or #213 #173)
-#1028 := (iff #216 #1027)
-#1029 := [rewrite]: #1028
-#1025 := [asserted]: #216
-#1032 := [mp #1025 #1029]: #1027
-#2030 := [unit-resolution #1032 #2028]: #213
-#179 := (or #177 #178)
-#905 := [asserted]: #179
-#2031 := [unit-resolution #905 #2026]: #178
-#1917 := (or #226 up_54 up_26 up_17 up_25)
-#1890 := [hypothesis]: #214
-#1891 := [hypothesis]: #213
-#1892 := [hypothesis]: #178
-#1893 := [hypothesis]: up_29
-#1894 := [unit-resolution #1069 #1893]: #218
-#1895 := [unit-resolution #547 #1894 #1892 #1891]: up_28
-#1187 := (or #219 #254)
-#266 := (or #254 #219)
-#1189 := (iff #266 #1187)
-#1190 := [rewrite]: #1189
-#1186 := [asserted]: #266
-#1193 := [mp #1186 #1190]: #1187
-#1896 := [unit-resolution #1193 #1895]: #254
-#1897 := [unit-resolution #621 #1896 #1890]: up_37
-#1898 := [unit-resolution #1297 #1897]: #296
-#1899 := [unit-resolution #795 #1898]: up_57
-#1900 := [unit-resolution #1418 #1899]: #337
-#1901 := [unit-resolution #1188 #1895]: #260
-#1287 := (or #255 #295)
-#298 := (or #295 #255)
-#1288 := (iff #298 #1287)
-#1289 := [rewrite]: #1288
-#1285 := [asserted]: #298
-#1292 := [mp #1285 #1289]: #1287
-#1902 := [unit-resolution #1292 #1897]: #295
-#1422 := (or #301 #335)
-#340 := (or #335 #301)
-#1424 := (iff #340 #1422)
-#1425 := [rewrite]: #1424
-#1421 := [asserted]: #340
-#1428 := [mp #1421 #1425]: #1422
-#1903 := [unit-resolution #1428 #1899]: #301
-#1904 := [unit-resolution #727 #1903 #1902 #1901]: up_49
-#1333 := (or #300 #309)
-#313 := (or #309 #300)
-#1335 := (iff #313 #1333)
-#1336 := [rewrite]: #1335
-#1332 := [asserted]: #313
-#1339 := [mp #1332 #1336]: #1333
-#1905 := [unit-resolution #1339 #1904]: #309
-#1906 := [unit-resolution #819 #1905 #1900]: up_59
-#1907 := [unit-resolution #1448 #1906]: #345
-#1908 := [unit-resolution #1524 #1907]: up_56
-#1909 := [unit-resolution #1329 #1904]: #308
-#1172 := (or #219 #259)
-#263 := (or #259 #219)
-#1174 := (iff #263 #1172)
-#1175 := [rewrite]: #1174
-#1171 := [asserted]: #263
-#1178 := [mp #1171 #1175]: #1172
-#1910 := [unit-resolution #1178 #1895]: #259
-#228 := (or #226 #227)
-#1061 := [asserted]: #228
-#1911 := [unit-resolution #1061 #1893]: #227
-#1343 := (or #268 #300)
-#315 := (or #300 #268)
-#1345 := (iff #315 #1343)
-#1346 := [rewrite]: #1345
-#1342 := [asserted]: #315
-#1349 := [mp #1342 #1346]: #1343
-#1912 := [unit-resolution #1349 #1904]: #268
-#1913 := [unit-resolution #653 #1912 #1911 #1910]: up_40
-#1229 := (or #267 #276)
-#280 := (or #276 #267)
-#1231 := (iff #280 #1229)
-#1232 := [rewrite]: #1231
-#1228 := [asserted]: #280
-#1235 := [mp #1228 #1232]: #1229
-#1914 := [unit-resolution #1235 #1913]: #276
-#1915 := [unit-resolution #759 #1914 #1909 #1597]: up_53
-#1385 := (or #316 #325)
-#329 := (or #325 #316)
-#1387 := (iff #329 #1385)
-#1388 := [rewrite]: #1387
-#1384 := [asserted]: #329
-#1391 := [mp #1384 #1388]: #1385
-#1916 := [unit-resolution #1391 #1915 #1908]: false
-#1918 := [lemma #1916]: #1917
-#2032 := [unit-resolution #1918 #1597 #2029 #2031 #2030]: #226
-#2010 := (or up_12 up_29 up_7 up_54 up_26)
-#1993 := (or up_35 up_12 up_54 up_26 up_29 up_7)
-#1955 := (or #170 up_54 up_26 up_29 up_12 up_35 up_7)
-#1940 := [unit-resolution #1625 #1805 #907 #1457 #1611 #1612]: up_32
-#1941 := [unit-resolution #1240 #1940]: #276
-#1942 := [unit-resolution #1230 #1940]: #275
-#1943 := [unit-resolution #608 #1806 #1611]: up_33
-#1944 := [unit-resolution #1113 #1943]: #243
-#1925 := (or #325 up_34 up_42)
-#1919 := [hypothesis]: up_56
-#1920 := [unit-resolution #1373 #1919]: #324
-#1921 := [unit-resolution #1516 #1920]: up_46
-#1922 := [unit-resolution #1396 #1919]: #284
-#1923 := [unit-resolution #685 #1922 #1605 #1798]: up_44
-#1924 := [unit-resolution #1276 #1923 #1921]: false
-#1926 := [lemma #1924]: #1925
-#1945 := [unit-resolution #1926 #1944 #1942]: #325
-#1946 := [unit-resolution #1524 #1945]: up_60
-#1947 := [unit-resolution #1448 #1946]: #341
-#1938 := (or #308 up_26 up_59)
-#1927 := [hypothesis]: up_51
-#1928 := [unit-resolution #1329 #1927]: #300
-#1929 := [hypothesis]: #341
-#1930 := [unit-resolution #1321 #1927]: #309
-#1931 := [unit-resolution #819 #1930 #1929]: up_58
-#1932 := [unit-resolution #1781 #1931 #1928]: up_39
-#1933 := [unit-resolution #1183 #1932]: #254
-#1934 := [unit-resolution #1418 #1931]: #335
-#1935 := [unit-resolution #795 #1934]: up_48
-#1936 := [unit-resolution #1297 #1935]: #255
-#1937 := [unit-resolution #621 #1936 #1933 #1890]: false
-#1939 := [lemma #1937]: #1938
-#1948 := [unit-resolution #1939 #1947 #1890]: #308
-#1949 := [unit-resolution #759 #1948 #1941 #1597]: up_53
-#1950 := [unit-resolution #1381 #1949]: #324
-#1951 := [unit-resolution #1516 #1950]: up_46
-#1952 := [unit-resolution #1401 #1949]: #284
-#1953 := [unit-resolution #685 #1952 #1944 #1942]: up_44
-#1954 := [unit-resolution #1276 #1953 #1951]: false
-#1956 := [lemma #1954]: #1955
-#1980 := [unit-resolution #1956 #1611 #1890 #907 #1457 #1597 #1612]: #170
-#1981 := [unit-resolution #430 #1980]: up_11
-#1982 := [unit-resolution #886 #1981]: #160
-#1983 := [unit-resolution #410 #1982 #1612]: up_10
-#1984 := [unit-resolution #980 #1983]: #194
-#1985 := [unit-resolution #1979 #1611 #907 #1984 #1457 #1612]: up_32
-#1970 := (or #235 up_34 up_54 up_26)
-#1957 := [hypothesis]: up_32
-#1958 := [unit-resolution #1240 #1957]: #276
-#1959 := [unit-resolution #1230 #1957]: #275
-#1960 := [unit-resolution #1926 #1959 #1605]: #325
-#1961 := [unit-resolution #1524 #1960]: up_60
-#1962 := [unit-resolution #1448 #1961]: #341
-#1963 := [unit-resolution #1939 #1962 #1890]: #308
-#1964 := [unit-resolution #759 #1963 #1958 #1597]: up_53
-#1965 := [unit-resolution #1381 #1964]: #324
-#1966 := [unit-resolution #1516 #1965]: up_46
-#1967 := [unit-resolution #1401 #1964]: #284
-#1968 := [unit-resolution #685 #1967 #1605 #1959]: up_44
-#1969 := [unit-resolution #1276 #1968 #1966]: false
-#1971 := [lemma #1969]: #1970
-#1986 := [unit-resolution #1971 #1985 #1597 #1890]: up_34
-#1987 := [unit-resolution #1113 #1986]: #242
-#1988 := [unit-resolution #608 #1987 #1611]: up_24
-#1989 := [unit-resolution #970 #1983]: #193
-#1990 := [unit-resolution #1136 #1986]: #202
-#1991 := [unit-resolution #505 #1990 #1457 #1989]: up_22
-#1992 := [unit-resolution #1016 #1991 #1988]: false
-#1994 := [lemma #1992]: #1993
-#1995 := [unit-resolution #1994 #1457 #1597 #1890 #907 #1612]: up_35
-#1996 := [unit-resolution #1281 #1995]: #291
-#1997 := [unit-resolution #1516 #1996]: up_55
-#1998 := [unit-resolution #1373 #1997]: #325
-#1999 := [unit-resolution #1524 #1998]: up_60
-#2000 := [unit-resolution #1448 #1999]: #341
-#2001 := [unit-resolution #1939 #2000 #1890]: #308
-#2002 := [unit-resolution #1610 #2001 #1995 #1597]: up_34
-#2003 := [unit-resolution #1131 #2002]: #234
-#2004 := [unit-resolution #1381 #1997]: #316
-#2005 := [unit-resolution #759 #2001 #2004 #1597]: up_43
-#2006 := [unit-resolution #1240 #2005]: #235
-#2007 := [unit-resolution #1136 #2002]: #202
-#2008 := [unit-resolution #1683 #2007 #1612 #1457]: #194
-#2009 := [unit-resolution #579 #2008 #2006 #907 #2003]: false
-#2011 := [lemma #2009]: #2010
-#2033 := [unit-resolution #2011 #2032 #1612 #1597 #2029]: up_12
-#2034 := [unit-resolution #891 #2033]: #160
-#2035 := [unit-resolution #410 #2034 #1612]: up_10
-#2036 := [unit-resolution #980 #2035]: #194
-#2037 := [unit-resolution #878 #2033]: #165
-#2038 := [unit-resolution #430 #2037]: up_13
-#2039 := [unit-resolution #1021 #2038]: #209
-#2024 := (or #234 up_26 up_54 up_24)
-#2012 := [hypothesis]: #209
-#2013 := [hypothesis]: up_31
-#2014 := [unit-resolution #1121 #2013]: #242
-#2015 := [unit-resolution #608 #2014 #2012]: up_35
-#2016 := [unit-resolution #1131 #2013]: #243
-#2017 := [unit-resolution #1610 #2016 #2015 #1597]: up_51
-#2018 := [unit-resolution #1939 #2017 #1890]: up_59
-#2019 := [unit-resolution #1448 #2018]: #345
-#2020 := [unit-resolution #1281 #2015]: #291
-#2021 := [unit-resolution #1516 #2020]: up_55
-#2022 := [unit-resolution #1373 #2021]: #325
-#2023 := [unit-resolution #1524 #2022 #2019]: false
-#2025 := [lemma #2023]: #2024
-#2040 := [unit-resolution #2025 #2029 #1597 #2039]: #234
-#2041 := [unit-resolution #579 #2040 #2032 #2036]: up_32
-#2042 := [unit-resolution #1240 #2041]: #276
-#2043 := [unit-resolution #1971 #2041 #1597 #2029]: up_34
-#2044 := [unit-resolution #1113 #2043]: #242
-#2045 := [unit-resolution #608 #2044 #2039]: up_35
-#2046 := [unit-resolution #1281 #2045]: #291
-#2047 := [unit-resolution #1516 #2046]: up_55
-#2048 := [unit-resolution #1381 #2047]: #316
-#2049 := [unit-resolution #759 #2048 #2042 #1597]: up_51
-#2050 := [unit-resolution #1373 #2047]: #325
-#2051 := [unit-resolution #1524 #2050]: up_60
-#2052 := [unit-resolution #1448 #2051]: #341
-#2053 := [unit-resolution #1939 #2052 #2049 #2029]: false
-#2055 := [lemma #2053]: #2054
-#2065 := [unit-resolution #2055 #1612 #1626]: up_54
-#1447 := (or #317 #345)
-#347 := (or #345 #317)
-#1449 := (iff #347 #1447)
-#1450 := [rewrite]: #1449
-#1446 := [asserted]: #347
-#1453 := [mp #1446 #1450]: #1447
-#2066 := [unit-resolution #1453 #2065]: #345
-#2067 := [unit-resolution #1524 #2066]: up_56
-#2083 := (or #275 up_7 up_12 up_8)
-#2063 := [hypothesis]: up_42
-#2064 := [unit-resolution #1230 #2063]: #235
-#2068 := [unit-resolution #1373 #2067]: #324
-#2069 := [unit-resolution #1516 #2068]: up_46
-#2070 := [unit-resolution #1281 #2069]: #250
-#2071 := [unit-resolution #1672 #2064 #1626 #2070 #1612]: up_16
-#2072 := [unit-resolution #913 #2071]: #172
-#2073 := [unit-resolution #1508 #2072]: up_15
-#2074 := [unit-resolution #1032 #2073]: #213
-#2075 := [unit-resolution #905 #2071]: #178
-#1452 := (or #317 #341)
-#348 := (or #341 #317)
-#1454 := (iff #348 #1452)
-#1455 := [rewrite]: #1454
-#1451 := [asserted]: #348
-#1458 := [mp #1451 #1455]: #1452
-#2076 := [unit-resolution #1458 #2065]: #341
-#2077 := [unit-resolution #1225 #2063]: #267
-#2061 := (or #226 up_59 up_40 up_17 up_25)
-#2056 := [unit-resolution #653 #1910 #1782 #1911]: up_41
-#2057 := [unit-resolution #1349 #2056]: #300
-#1338 := (or #268 #309)
-#314 := (or #309 #268)
-#1340 := (iff #314 #1338)
-#1341 := [rewrite]: #1340
-#1337 := [asserted]: #314
-#1344 := [mp #1337 #1341]: #1338
-#2058 := [unit-resolution #1344 #2056]: #309
-#2059 := [unit-resolution #819 #2058 #1929]: up_58
-#2060 := [unit-resolution #1781 #2059 #2057 #1901]: false
-#2062 := [lemma #2060]: #2061
-#2078 := [unit-resolution #2062 #2077 #2076 #2075 #2074]: #226
-#2079 := [unit-resolution #1625 #2078 #2070 #1457 #2064 #1612]: up_22
-#2080 := [unit-resolution #1979 #2078 #2070 #1457 #2064 #1612]: up_21
-#2081 := [unit-resolution #1683 #2080 #1612 #1457]: up_23
-#2082 := [unit-resolution #983 #2081 #2079]: false
-#2084 := [lemma #2082]: #2083
-#2085 := [unit-resolution #2084 #1457 #1612 #1626]: #275
-#2086 := [unit-resolution #1926 #2085 #2067]: up_34
-#2087 := [unit-resolution #1136 #2086]: #202
-#2088 := [unit-resolution #1113 #2086]: #242
-#2089 := [unit-resolution #608 #2088 #2070]: up_24
-#2090 := [unit-resolution #1016 #2089]: #201
-#2091 := [unit-resolution #505 #2090 #1457 #2087]: up_20
-#2092 := [unit-resolution #970 #2091]: #161
-#2093 := [unit-resolution #1021 #2089]: #170
-#2094 := [unit-resolution #430 #2093]: up_11
-#2095 := [unit-resolution #886 #2094]: #160
-#2096 := [unit-resolution #410 #2095 #2092 #1612]: false
-#2098 := [lemma #2096]: #2097
-#2102 := [unit-resolution #2098 #2100 #2101]: up_12
-#2103 := [unit-resolution #891 #2102]: #160
-#2104 := [unit-resolution #410 #2103 #2100]: up_10
-#2105 := [unit-resolution #980 #2104]: #194
-#2106 := [unit-resolution #2055 #2100 #2101]: up_54
-#2107 := [unit-resolution #1453 #2106]: #345
-#2108 := [unit-resolution #1524 #2107]: up_56
-#2109 := [unit-resolution #1373 #2108]: #324
-#2110 := [unit-resolution #1516 #2109]: up_46
-#2111 := [unit-resolution #1281 #2110]: #250
-#2112 := [unit-resolution #878 #2102]: #165
-#2113 := [unit-resolution #430 #2112]: up_13
-#2114 := [unit-resolution #1021 #2113]: #209
-#2115 := [unit-resolution #608 #2114 #2111]: up_33
-#2116 := [unit-resolution #1121 #2115]: #234
-#2117 := [unit-resolution #1276 #2110]: #283
-#2118 := [unit-resolution #1396 #2108]: #284
-#2119 := [unit-resolution #1113 #2115]: #243
-#2120 := [unit-resolution #685 #2119 #2118 #2117]: up_42
-#2121 := [unit-resolution #1230 #2120]: #235
-#2122 := [unit-resolution #579 #2121 #2116 #2105]: up_29
-#2123 := [unit-resolution #1225 #2120]: #267
-#2124 := [unit-resolution #1458 #2106]: #341
-#2125 := [unit-resolution #1672 #2121 #2101 #2111 #2100]: up_16
-#2126 := [unit-resolution #905 #2125]: #178
-#2127 := [unit-resolution #2062 #2126 #2124 #2123 #2122]: up_25
-#2128 := [unit-resolution #913 #2125]: #172
-#2129 := [unit-resolution #1508 #2128]: up_15
-#2130 := [unit-resolution #1032 #2129 #2127]: false
-#2131 := [lemma #2130]: #150
-#1494 := (or up_5 up_6)
-decl up_1 :: bool
-#4 := up_1
-#379 := (or up_1 up_5 up_6)
-#1497 := (iff #379 #1494)
-#1491 := (or false up_5 up_6)
-#1495 := (iff #1491 #1494)
-#1496 := [rewrite]: #1495
-#1492 := (iff #379 #1491)
-#1467 := (iff up_1 false)
-#5 := (not up_1)
-#1470 := (iff #5 #1467)
-#1463 := (iff #1467 #5)
-#1468 := [rewrite]: #1463
-#1471 := [symm #1468]: #1470
-#368 := [asserted]: #5
-#1472 := [mp #368 #1471]: #1467
-#1493 := [monotonicity #1472]: #1492
-#1498 := [trans #1493 #1496]: #1497
-#14 := (or up_6 up_1)
-#15 := (or up_5 #14)
-#382 := (iff #15 #379)
-#373 := (or up_1 up_6)
-#376 := (or up_5 #373)
-#380 := (iff #376 #379)
-#381 := [rewrite]: #380
-#377 := (iff #15 #376)
-#374 := (iff #14 #373)
-#375 := [rewrite]: #374
-#378 := [monotonicity #375]: #377
-#383 := [trans #378 #381]: #382
-#372 := [asserted]: #15
-#384 := [mp #372 #383]: #379
-#1499 := [mp #384 #1498]: #1494
-#2138 := [unit-resolution #1499 #2131]: up_6
-#151 := (not up_6)
-#927 := (or #151 #172)
-#184 := (or #172 #151)
-#929 := (iff #184 #927)
-#930 := [rewrite]: #929
-#926 := [asserted]: #184
-#933 := [mp #926 #930]: #927
-#2139 := [unit-resolution #933 #2138]: #172
-#2140 := [unit-resolution #1508 #2139]: up_15
-#2147 := [unit-resolution #1037 #2140]: #214
-#2159 := [unit-resolution #2011 #2147]: #2158
-#2160 := [unit-resolution #2159 #907 #1612 #1457]: up_54
-#2161 := [unit-resolution #1453 #2160]: #345
-#2162 := [unit-resolution #1524 #2161]: up_56
-#2163 := [unit-resolution #1926 #2162 #2157 #1605]: false
-#2165 := [lemma #2163]: #2164
-#2166 := [unit-resolution #2165 #1605 #1612 #1611 #1457]: up_29
-#2148 := (or #226 up_54)
-#2141 := [unit-resolution #1032 #2140]: #213
-#922 := (or #151 #178)
-#183 := (or #178 #151)
-#924 := (iff #183 #922)
-#925 := [rewrite]: #924
-#921 := [asserted]: #183
-#928 := [mp #921 #925]: #922
-#2142 := [unit-resolution #928 #2138]: #178
-#2149 := [unit-resolution #1918 #2147 #2142 #2141]: #2148
-#2167 := [unit-resolution #2149 #2166]: up_54
-#2154 := (or #226 up_34 up_59)
-#2143 := (or #226 up_59 up_40)
-#2144 := [unit-resolution #2062 #2142 #2141]: #2143
-#2145 := [unit-resolution #2144 #1893 #1929]: up_40
-#2146 := [unit-resolution #1225 #2145]: #275
-#2150 := [unit-resolution #2149 #1893]: up_54
-#2151 := [unit-resolution #1453 #2150]: #345
-#2152 := [unit-resolution #1524 #2151]: up_56
-#2153 := [unit-resolution #1926 #2152 #2146 #1605]: false
-#2155 := [lemma #2153]: #2154
-#2168 := [unit-resolution #2155 #2166 #1605]: up_59
-#2169 := [unit-resolution #1458 #2168 #2167]: false
-#2171 := [lemma #2169]: #2170
-#2172 := [unit-resolution #2171 #1612 #1611 #1457]: up_34
-#2173 := [unit-resolution #1136 #2172]: #202
-#2174 := [unit-resolution #1113 #2172]: #242
-#2175 := [unit-resolution #608 #2174 #1611]: up_24
-#2176 := [unit-resolution #1016 #2175]: #201
-#2177 := [unit-resolution #505 #2176 #1457 #2173]: up_20
-#2178 := [unit-resolution #970 #2177]: #161
-#2179 := [unit-resolution #1021 #2175]: #170
-#2180 := [unit-resolution #430 #2179]: up_11
-#2181 := [unit-resolution #886 #2180]: #160
-#2182 := [unit-resolution #410 #2181 #2178 #1612]: false
-#2184 := [lemma #2182]: #2183
-#2235 := [unit-resolution #2184 #1457 #1611]: up_7
-#157 := (or #155 #156)
-#856 := [asserted]: #157
-#2236 := [unit-resolution #856 #2235]: #156
-#2299 := (or up_34 up_35 up_12)
-#2283 := (or #186 up_34)
-#2185 := [hypothesis]: up_19
-#2191 := [unit-resolution #1084 #2185]: #227
-#2186 := [unit-resolution #1089 #2185]: #218
-#2187 := (or up_27 up_28)
-#2188 := [unit-resolution #547 #2142 #2141]: #2187
-#2189 := [unit-resolution #2188 #2186]: up_28
-#2192 := [unit-resolution #1178 #2189]: #259
-#2265 := [unit-resolution #1193 #2189]: #254
-#2266 := (or up_36 up_37)
-#2267 := [unit-resolution #621 #2147]: #2266
-#2268 := [unit-resolution #2267 #2265]: up_37
-#2269 := [unit-resolution #1292 #2268]: #295
-#2190 := [unit-resolution #1188 #2189]: #260
-#2270 := [unit-resolution #1297 #2268]: #296
-#2271 := [unit-resolution #795 #2270]: up_57
-#2272 := [unit-resolution #1428 #2271]: #301
-#2273 := [unit-resolution #727 #2272 #2190 #2269]: up_49
-#2274 := [unit-resolution #1349 #2273]: #268
-#2275 := [unit-resolution #653 #2274 #2192 #2191]: up_40
-#2276 := [unit-resolution #1225 #2275]: #275
-#2277 := [unit-resolution #1418 #2271]: #337
-#2278 := [unit-resolution #1339 #2273]: #309
-#2279 := [unit-resolution #819 #2278 #2277]: up_59
-#2280 := [unit-resolution #1448 #2279]: #345
-#2281 := [unit-resolution #1524 #2280]: up_56
-#2282 := [unit-resolution #1926 #2281 #2276 #1605]: false
-#2284 := [lemma #2282]: #2283
-#2292 := [unit-resolution #2284 #1605]: #186
-#2223 := (or up_8 up_18 up_19)
-#912 := (or #151 #177)
-#181 := (or #177 #151)
-#914 := (iff #181 #912)
-#915 := [rewrite]: #914
-#911 := [asserted]: #181
-#918 := [mp #911 #915]: #912
-#2222 := [unit-resolution #918 #2138]: #177
-#2224 := [unit-resolution #473 #2222]: #2223
-#2293 := [unit-resolution #2224 #2292 #2236]: up_18
-#2257 := (or #235 up_34)
-#2252 := (or #235 up_34 up_54)
-#2253 := [unit-resolution #1971 #2147]: #2252
-#2254 := [unit-resolution #2253 #1957 #1605]: up_54
-#2255 := [unit-resolution #1453 #2254]: #345
-#2256 := [unit-resolution #1524 #2255 #1960]: false
-#2258 := [lemma #2256]: #2257
-#2294 := [unit-resolution #2258 #1605]: #235
-#2290 := (or up_29 up_35 up_32 up_12 #185)
-#2200 := [hypothesis]: up_18
-#2206 := (or #185 up_29 up_32 up_12 up_22)
-#2201 := [unit-resolution #965 #2200]: #193
-#2202 := [unit-resolution #505 #2201 #1457 #1456]: up_23
-#2203 := [unit-resolution #975 #2200]: #194
-#2204 := [unit-resolution #579 #2203 #907 #895]: up_31
-#2205 := [unit-resolution #1141 #2204 #2202]: false
-#2207 := [lemma #2205]: #2206
-#2285 := [unit-resolution #2207 #907 #895 #1457 #2200]: up_22
-#2286 := [unit-resolution #1016 #2285]: #209
-#2287 := [unit-resolution #579 #907 #895 #2203]: up_31
-#2288 := [unit-resolution #1121 #2287]: #242
-#2289 := [unit-resolution #608 #2288 #2286 #1611]: false
-#2291 := [lemma #2289]: #2290
-#2295 := [unit-resolution #2291 #2294 #1611 #1457 #2293]: up_29
-#2296 := [unit-resolution #2149 #2295]: up_54
-#2297 := [unit-resolution #2155 #2295 #1605]: up_59
-#2298 := [unit-resolution #1458 #2297 #2296]: false
-#2300 := [lemma #2298]: #2299
-#2301 := [unit-resolution #2300 #1457 #1611]: up_34
-#2302 := [unit-resolution #1136 #2301]: #202
-#2303 := [unit-resolution #1113 #2301]: #242
-#2304 := [unit-resolution #608 #2303 #1611]: up_24
-#2305 := [unit-resolution #1016 #2304]: #201
-#2306 := [unit-resolution #505 #2305 #1457 #2302]: up_20
-#2307 := [unit-resolution #965 #2306]: #185
-#2308 := [unit-resolution #2224 #2307 #2236]: up_19
-#2309 := [unit-resolution #957 #2306]: #194
-#2310 := [unit-resolution #1131 #2301]: #234
-#2311 := [unit-resolution #1074 #2308]: #226
-#2312 := [unit-resolution #579 #2311 #2310 #2309]: up_32
-#2313 := [unit-resolution #1245 #2312]: #267
-#2198 := (or #186 up_59 up_40)
-#2193 := [unit-resolution #653 #2192 #1782 #2191]: up_41
-#2194 := [unit-resolution #1349 #2193]: #300
-#2195 := [unit-resolution #1344 #2193]: #309
-#2196 := [unit-resolution #819 #2195 #1929]: up_58
-#2197 := [unit-resolution #1781 #2196 #2194 #2190]: false
-#2199 := [lemma #2197]: #2198
-#2314 := [unit-resolution #2199 #2313 #2308]: up_59
-#2315 := [unit-resolution #1448 #2314]: #345
-#2316 := [unit-resolution #1524 #2315]: up_56
-#2317 := [unit-resolution #1084 #2308]: #227
-#2318 := [unit-resolution #1089 #2308]: #218
-#2319 := [unit-resolution #2188 #2318]: up_28
-#2320 := [unit-resolution #1178 #2319]: #259
-#2321 := [unit-resolution #653 #2313 #2320 #2317]: up_41
-#2322 := [unit-resolution #1334 #2321]: #308
-#2323 := [unit-resolution #1240 #2312]: #276
-#2324 := [unit-resolution #1458 #2314]: #317
-#2325 := [unit-resolution #759 #2324 #2323 #2322]: up_53
-#2326 := [unit-resolution #1391 #2325 #2316]: false
-#2328 := [lemma #2326]: #2327
-#2337 := [unit-resolution #2328 #1611]: up_12
-#2338 := [unit-resolution #878 #2337]: #165
-#2339 := [unit-resolution #430 #2338]: up_13
-#2340 := [unit-resolution #1021 #2339]: #209
-#2341 := [unit-resolution #608 #2340 #1611]: up_33
-#2342 := [unit-resolution #1113 #2341]: #243
-#2343 := [unit-resolution #2258 #2342]: #235
-#2344 := [unit-resolution #1121 #2341]: #234
-#2345 := [unit-resolution #2284 #2342]: #186
-#2346 := [unit-resolution #891 #2337]: #160
-#2335 := (or #194 up_9 up_19)
-#2329 := [hypothesis]: #186
-#2330 := [unit-resolution #975 #1674]: #185
-#2331 := [unit-resolution #2224 #2330 #2329]: up_8
-#2332 := [hypothesis]: #160
-#2333 := [unit-resolution #410 #1678 #2332]: up_7
-#2334 := [unit-resolution #856 #2333 #2331]: false
-#2336 := [lemma #2334]: #2335
-#2347 := [unit-resolution #2336 #2346 #2345]: #194
-#2348 := [unit-resolution #579 #2347 #2344 #2343]: up_29
-#2349 := [unit-resolution #2149 #2348]: up_54
-#2350 := [unit-resolution #2155 #2348 #2342]: up_59
-#2351 := [unit-resolution #1458 #2350 #2349]: false
-#2352 := [lemma #2351]: up_35
-#2353 := [unit-resolution #1281 #2352]: #291
-#2354 := [unit-resolution #1516 #2353]: up_55
-#2355 := [unit-resolution #1373 #2354]: #325
-#2356 := [unit-resolution #1524 #2355]: up_60
-#2357 := [unit-resolution #1453 #2356]: #317
-#2358 := [unit-resolution #2149 #2357]: #226
-#2359 := [unit-resolution #1448 #2356]: #341
-#2217 := (or #308 up_59)
-#2218 := [unit-resolution #1939 #2147]: #2217
-#2360 := [unit-resolution #2218 #2359]: #308
-#2361 := [unit-resolution #1381 #2354]: #316
-#2362 := [unit-resolution #759 #2357 #2361 #2360]: up_43
-#2363 := [unit-resolution #1235 #2362]: #267
-#2364 := [unit-resolution #2199 #2363 #2359]: #186
-#1145 := (or #209 #250)
-#252 := (or #250 #209)
-#1147 := (iff #252 #1145)
-#1148 := [rewrite]: #1147
-#1144 := [asserted]: #252
-#1151 := [mp #1144 #1148]: #1145
-#2365 := [unit-resolution #1151 #2352]: #209
-#2230 := (or #234 up_54 up_24)
-#2231 := [unit-resolution #2025 #2147]: #2230
-#2366 := [unit-resolution #2231 #2357 #2365]: #234
-#2367 := [unit-resolution #1240 #2362]: #235
-#2368 := [unit-resolution #579 #2367 #2366 #2358]: up_21
-#2369 := [unit-resolution #2336 #2368 #2364]: up_9
-#870 := (or #155 #160)
-#163 := (or #160 #155)
-#871 := (iff #163 #870)
-#872 := [rewrite]: #871
-#868 := [asserted]: #163
-#875 := [mp #868 #872]: #870
-#2370 := [unit-resolution #875 #2369]: #155
-#2371 := [unit-resolution #891 #2369]: #166
-[unit-resolution #2159 #2371 #2370 #2358 #2357]: false
-unsat
--- a/src/HOL/SMT/SMT.thy Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/SMT/SMT.thy Tue Feb 02 18:16:48 2010 +0100
@@ -39,20 +39,39 @@
subsection {* Z3-specific options *}
+text {* Pass extra command-line arguments to Z3 to control its behaviour: *}
+
+declare [[ z3_options = "" ]]
+
text {* Enable proof reconstruction for Z3: *}
declare [[ z3_proofs = false ]]
-text {* Pass extra command-line arguments to Z3
-to control its behaviour: *}
+text {* Enable or disable tracing of the theorems used for proving a
+proposition: *}
+
+declare [[ z3_trace_assms = false ]]
+
+
+subsection {* Certificates *}
-declare [[ z3_options = "" ]]
+text {* To avoid invocation of an SMT solver for the same problem
+again and again, cache certificates in a file (the filename must
+be given by an absolute path, an empty string disables the usage
+of certificates): *}
+
+declare [[ smt_certificates = "" ]]
+
+text {* Enables or disables the addition of new certificates to
+the current certificates file (when disabled, only existing
+certificates are used and no SMT solver is invoked): *}
+
+declare [[ smt_record = true ]]
subsection {* Special configuration options *}
-text {*
-Trace the problem file, the result of the SMT solver and
+text {* Trace the problem file, the result of the SMT solver and
further information: *}
declare [[ smt_trace = false ]]
@@ -61,11 +80,4 @@
declare [[ smt_unfold_defs = true ]]
-text {*
-Produce or use certificates (to avoid repeated invocation of an
-SMT solver again and again). The value is an absolute path
-pointing to the problem file. See also the examples. *}
-
-declare [[ smt_keep = "", smt_cert = "" ]]
-
end
--- a/src/HOL/SMT/Tools/smt_solver.ML Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/SMT/Tools/smt_solver.ML Tue Feb 02 18:16:48 2010 +0100
@@ -28,8 +28,8 @@
val with_timeout: Proof.context -> ('a -> 'b) -> 'a -> 'b
val trace: bool Config.T
val trace_msg: Proof.context -> ('a -> string) -> 'a -> unit
- val keep: string Config.T
- val cert: string Config.T
+ val record: bool Config.T
+ val certificates: string Config.T
(*solvers*)
type solver = Proof.context -> thm list -> thm
@@ -88,8 +88,11 @@
fun trace_msg ctxt f x =
if Config.get ctxt trace then tracing (f x) else ()
-val (keep, setup_keep) = Attrib.config_string "smt_keep" ""
-val (cert, setup_cert) = Attrib.config_string "smt_cert" ""
+val (record, setup_record) = Attrib.config_bool "smt_record" true
+val no_certificates = ""
+val (certificates, setup_certificates) =
+ Attrib.config_string "smt_certificates" no_certificates
+
(* interface to external solvers *)
@@ -98,18 +101,14 @@
fun with_files ctxt f =
let
- fun make_names n = (n, n ^ ".proof")
-
- val keep' = Config.get ctxt keep
val paths as (problem_path, proof_path) =
- if keep' <> "" andalso File.exists (Path.dir (Path.explode keep'))
- then pairself Path.explode (make_names keep')
- else pairself (File.tmp_path o Path.explode)
- (make_names ("smt-" ^ serial_string ()))
+ "smt-" ^ serial_string ()
+ |> (fn n => (n, n ^ ".proof"))
+ |> pairself (File.tmp_path o Path.explode)
val y = Exn.capture f (problem_path, proof_path)
- val _ = if keep' = "" then (pairself (try File.rm) paths; ()) else ()
+ val _ = pairself (try File.rm) paths
in Exn.release y end
fun invoke ctxt output f (paths as (problem_path, proof_path)) =
@@ -121,59 +120,46 @@
val _ = trace_msg ctxt (pretty "SMT problem:" o split_lines o File.read)
problem_path
- val _ = with_timeout ctxt f paths
+ val (s, _) = with_timeout ctxt f paths
+ val _ = trace_msg ctxt (pretty "SMT solver:") (split_lines s)
+
fun lines_of path = the_default [] (try (File.fold_lines cons path) [])
val ls = rev (dropwhile (equal "") (lines_of proof_path))
val _ = trace_msg ctxt (pretty "SMT result:") ls
in (x, ls) end
-val expand_name = Path.implode o Path.expand o Path.explode
-
-fun run_perl name args ps =
- system_out (space_implode " " ("perl -w" ::
- File.shell_path (Path.explode (getenv name)) ::
- map File.shell_quote args @ ps))
-
-fun use_certificate ctxt ps =
- let val name = Config.get ctxt cert
+fun choose {env_var, remote_name} =
+ let
+ val local_solver = getenv env_var
+ val remote_solver = the_default "" remote_name
+ val remote_url = getenv "REMOTE_SMT_URL"
in
- if name = "" then false
- else
- (tracing ("Using certificate " ^ quote (expand_name name) ^ " ...");
- run_perl "CERT_SMT_SOLVER" [expand_name name] ps;
- true)
+ if local_solver <> ""
+ then (["local", local_solver],
+ "Invoking local SMT solver " ^ quote local_solver ^ " ...")
+ else if remote_solver <> "" andalso remote_url <> ""
+ then (["remote", remote_solver],
+ "Invoking remote SMT solver " ^ quote remote_solver ^ " at " ^
+ quote remote_url ^ " ...")
+ else error ("Undefined Isabelle environment variable: " ^ quote env_var)
end
-fun run_locally f env_var args ps =
- if getenv env_var = ""
- then f ("Undefined Isabelle environment variable: " ^ quote env_var)
- else
- let val app = Path.expand (Path.explode (getenv env_var))
- in
- if not (File.exists app)
- then f ("No such file: " ^ quote (Path.implode app))
- else
- (tracing ("Invoking local SMT solver " ^ quote (Path.implode app) ^
- " ...");
- system_out (space_implode " " (File.shell_path app ::
- map File.shell_quote args @ ps)); ())
- end
-
-fun run_remote remote_name args ps msg =
- (case remote_name of
- NONE => error msg
- | SOME name =>
- let
- val url = getenv "REMOTE_SMT_URL"
- val _ = tracing ("Invoking remote SMT solver " ^ quote name ^ " at " ^
- quote url ^ " ...")
- in (run_perl "REMOTE_SMT_SOLVER" (url :: name :: args) ps; ()) end)
-
-fun run ctxt {env_var, remote_name} args (problem_path, proof_path) =
- let val ps = [File.shell_path problem_path, ">", File.shell_path proof_path]
+fun run ctxt cmd args (problem_path, proof_path) =
+ let
+ val certs = Config.get ctxt certificates
+ val certs' =
+ if certs = no_certificates then "-"
+ else File.shell_path (Path.explode certs)
+ val (solver, msg) =
+ if certs = no_certificates orelse Config.get ctxt record
+ then choose cmd
+ else (["certificate"], "Using certificate from " ^ quote certs' ^ " ...")
+ val _ = tracing msg
in
- if use_certificate ctxt ps then ()
- else run_locally (run_remote remote_name args ps) env_var args ps
+ system_out (space_implode " " ("perl -w" ::
+ File.shell_path (Path.explode (getenv "RUN_SMT_SOLVER")) :: certs' ::
+ map File.shell_quote (solver @ args) @
+ map File.shell_path [problem_path, proof_path]))
end
in
@@ -291,11 +277,12 @@
"SMT solver configuration" #>
setup_timeout #>
setup_trace #>
- setup_keep #>
- setup_cert #>
+ setup_record #>
+ setup_certificates #>
Method.setup (Binding.name "smt") smt_method
"Applies an SMT solver to the current goal."
+
fun print_setup gen =
let
val t = string_of_int (Config.get_generic gen timeout)
--- a/src/HOL/SMT/etc/settings Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/SMT/etc/settings Tue Feb 02 18:16:48 2010 +0100
@@ -1,11 +1,9 @@
ISABELLE_SMT="$COMPONENT"
-REMOTE_SMT_SOLVER="$ISABELLE_SMT/lib/scripts/remote_smt.pl"
+RUN_SMT_SOLVER="$ISABELLE_SMT/lib/scripts/run_smt_solver.pl"
REMOTE_SMT_URL="http://smt.in.tum.de/smt"
-CERT_SMT_SOLVER="$ISABELLE_SMT/lib/scripts/cert_smt.pl"
-
#
# Paths to local SMT solvers:
#
--- a/src/HOL/SMT/lib/scripts/cert_smt.pl Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,33 +0,0 @@
-#
-# Author: Sascha Boehme, TU Muenchen
-#
-# Fake SMT solver: check that input matches previously computed input and
-# and return previously computed output.
-#
-
-use strict;
-use File::Compare;
-
-
-# arguments
-
-my $cert_path = $ARGV[0];
-my $new_problem = $ARGV[1];
-
-
-# check content of new problem file against old problem file
-
-my $old_problem = $cert_path;
-my $old_proof = $cert_path . ".proof";
-
-if (-e $old_problem and compare($old_problem, $new_problem) == 0) {
- if (-e $old_proof) {
- open FILE, "<$old_proof";
- foreach (<FILE>) {
- print $_;
- }
- close FILE;
- }
- else { print "ERROR: unable to open proof file\n"; }
-}
-else { print "ERROR: bad problem\n"; }
--- a/src/HOL/SMT/lib/scripts/remote_smt.pl Sun Jan 31 15:22:40 2010 +0100
+++ /dev/null Thu Jan 01 00:00:00 1970 +0000
@@ -1,36 +0,0 @@
-#
-# Author: Sascha Boehme, TU Muenchen
-#
-# Invoke remote SMT solvers.
-#
-
-use strict;
-use LWP;
-
-
-# arguments
-
-my $url = $ARGV[0];
-my $solver = $ARGV[1];
-my @options = @ARGV[2 .. ($#ARGV - 1)];
-my $problem_file = $ARGV[-1];
-
-
-# call solver
-
-my $agent = LWP::UserAgent->new;
-$agent->agent("SMT-Request");
-$agent->timeout(180);
-my $response = $agent->post($url, [
- "Solver" => $solver,
- "Options" => join(" ", @options),
- "Problem" => [$problem_file] ],
- "Content_Type" => "form-data");
-if (not $response->is_success) {
- print "HTTP-Error: " . $response->message . "\n";
- exit 1;
-}
-else {
- print $response->content;
- exit 0;
-}
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/SMT/lib/scripts/run_smt_solver.pl Tue Feb 02 18:16:48 2010 +0100
@@ -0,0 +1,110 @@
+#
+# Author: Sascha Boehme, TU Muenchen
+#
+# Cache for prover results, based on discriminating problems using MD5.
+
+use strict;
+use warnings;
+use Digest::MD5;
+use LWP;
+
+
+# arguments
+
+my $certs_file = shift @ARGV;
+my $location = shift @ARGV;
+my $result_file = pop @ARGV;
+my $problem_file = $ARGV[-1];
+
+my $use_certs = not ($certs_file eq "-");
+
+
+# create MD5 problem digest
+
+my $problem_digest = "";
+if ($use_certs) {
+ my $md5 = Digest::MD5->new;
+ open FILE, "<$problem_file" or
+ die "ERROR: Failed to open '$problem_file' ($!)";
+ $md5->addfile(*FILE);
+ close FILE;
+ $problem_digest = $md5->b64digest;
+}
+
+
+# lookup problem digest among existing certificates and
+# return a cached result (if there is a certificate for the problem)
+
+if ($use_certs and -e $certs_file) {
+ my $cached = 0;
+ open CERTS, "<$certs_file" or die "ERROR: Failed to open '$certs_file' ($!)";
+ while (<CERTS>) {
+ if (m/(\S+) (\d+)/) {
+ if ($1 eq $problem_digest) {
+ my $start = tell CERTS;
+ open FILE, ">$result_file" or
+ die "ERROR: Failed to open '$result_file ($!)";
+ while ((tell CERTS) - $start < $2) {
+ my $line = <CERTS>;
+ print FILE $line;
+ }
+ close FILE;
+ $cached = 1;
+ last;
+ }
+ else { seek CERTS, $2, 1; }
+ }
+ else { die "ERROR: Invalid file format in '$certs_file'"; }
+ }
+ close CERTS;
+ if ($cached) { exit 0; }
+}
+
+
+# invoke (remote or local) prover
+
+if ($location eq "remote") {
+ # arguments
+ my $solver = $ARGV[0];
+ my @options = @ARGV[1 .. ($#ARGV - 1)];
+
+ # call solver
+ my $agent = LWP::UserAgent->new;
+ $agent->agent("SMT-Request");
+ $agent->timeout(180);
+ my $response = $agent->post($ENV{"REMOTE_SMT_URL"}, [
+ "Solver" => $solver,
+ "Options" => join(" ", @options),
+ "Problem" => [$problem_file] ],
+ "Content_Type" => "form-data");
+ if (not $response->is_success) { die "HTTP-Error: " . $response->message; }
+ else {
+ open FILE, ">$result_file" or
+ die "ERROR: Failed to create '$result_file' ($!)";
+ print FILE $response->content;
+ close FILE;
+ }
+}
+elsif ($location eq "local") {
+ system "@ARGV >$result_file 2>&1";
+}
+else { die "ERROR: No SMT solver invoked"; }
+
+
+# cache result
+
+if ($use_certs) {
+ open CERTS, ">>$certs_file" or
+ die "ERROR: Failed to open '$certs_file' ($!)";
+ print CERTS
+ ("$problem_digest " . ((-s $result_file) + (length "\n")) . "\n");
+
+ open FILE, "<$result_file" or
+ die "ERROR: Failed to open '$result_file' ($!)";
+ foreach (<FILE>) { print CERTS $_; }
+ close FILE;
+
+ print CERTS "\n";
+ close CERTS;
+}
+
--- a/src/HOL/Tools/Nitpick/HISTORY Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/Tools/Nitpick/HISTORY Tue Feb 02 18:16:48 2010 +0100
@@ -1,7 +1,9 @@
Version 2010
* Added and implemented "binary_ints" and "bits" options
- * Fixed soundness bug in "destroy_constrs" optimization
+ * Added "std" option and implemented support for nonstandard models
+ * Fixed soundness bugs related to "destroy_constrs" optimization and record
+ getters
Version 2009-1
--- a/src/HOL/Tools/Nitpick/kodkod.ML Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/Tools/Nitpick/kodkod.ML Tue Feb 02 18:16:48 2010 +0100
@@ -1,6 +1,6 @@
(* Title: HOL/Tools/Nitpick/kodkod.ML
Author: Jasmin Blanchette, TU Muenchen
- Copyright 2008, 2009
+ Copyright 2008, 2009, 2010
ML interface on top of Kodkod.
*)
--- a/src/HOL/Tools/Nitpick/kodkod_sat.ML Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/Tools/Nitpick/kodkod_sat.ML Tue Feb 02 18:16:48 2010 +0100
@@ -1,6 +1,6 @@
(* Title: HOL/Tools/Nitpick/kodkod_sat.ML
Author: Jasmin Blanchette, TU Muenchen
- Copyright 2009
+ Copyright 2009, 2010
Kodkod SAT solver integration.
*)
--- a/src/HOL/Tools/Nitpick/minipick.ML Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/Tools/Nitpick/minipick.ML Tue Feb 02 18:16:48 2010 +0100
@@ -1,6 +1,6 @@
(* Title: HOL/Tools/Nitpick/minipick.ML
Author: Jasmin Blanchette, TU Muenchen
- Copyright 2009
+ Copyright 2009, 2010
Finite model generation for HOL formulas using Kodkod, minimalistic version.
*)
--- a/src/HOL/Tools/Nitpick/nitpick.ML Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick.ML Tue Feb 02 18:16:48 2010 +0100
@@ -1,6 +1,6 @@
(* Title: HOL/Tools/Nitpick/nitpick.ML
Author: Jasmin Blanchette, TU Muenchen
- Copyright 2008, 2009
+ Copyright 2008, 2009, 2010
Finite model generation for HOL formulas using Kodkod.
*)
@@ -16,6 +16,7 @@
bisim_depths: int list,
boxes: (typ option * bool option) list,
monos: (typ option * bool option) list,
+ stds: (typ option * bool) list,
wfs: (styp option * bool option) list,
sat_solver: string,
blocking: bool,
@@ -57,9 +58,10 @@
val register_codatatype : typ -> string -> styp list -> theory -> theory
val unregister_codatatype : typ -> theory -> theory
val pick_nits_in_term :
- Proof.state -> params -> bool -> term list -> term -> string * Proof.state
+ Proof.state -> params -> bool -> int -> int -> int -> term list -> term
+ -> string * Proof.state
val pick_nits_in_subgoal :
- Proof.state -> params -> bool -> int -> string * Proof.state
+ Proof.state -> params -> bool -> int -> int -> string * Proof.state
end;
structure Nitpick : NITPICK =
@@ -85,6 +87,7 @@
bisim_depths: int list,
boxes: (typ option * bool option) list,
monos: (typ option * bool option) list,
+ stds: (typ option * bool) list,
wfs: (styp option * bool option) list,
sat_solver: string,
blocking: bool,
@@ -183,18 +186,21 @@
(* (unit -> string) -> Pretty.T *)
fun plazy f = Pretty.blk (0, pstrs (f ()))
-(* Time.time -> Proof.state -> params -> bool -> term -> string * Proof.state *)
-fun pick_them_nits_in_term deadline state (params : params) auto orig_assm_ts
- orig_t =
+(* Time.time -> Proof.state -> params -> bool -> int -> int -> int -> term
+ -> string * Proof.state *)
+fun pick_them_nits_in_term deadline state (params : params) auto i n step
+ orig_assm_ts orig_t =
let
val timer = Timer.startRealTimer ()
val thy = Proof.theory_of state
val ctxt = Proof.context_of state
+(* FIXME: reintroduce code before new release
val nitpick_thy = ThyInfo.get_theory "Nitpick"
val _ = Theory.subthy (nitpick_thy, thy) orelse
error "You must import the theory \"Nitpick\" to use Nitpick"
+*)
val {cards_assigns, maxes_assigns, iters_assigns, bitss, bisim_depths,
- boxes, monos, wfs, sat_solver, blocking, falsify, debug, verbose,
+ boxes, monos, stds, wfs, sat_solver, blocking, falsify, debug, verbose,
overlord, user_axioms, assms, merge_type_vars, binary_ints,
destroy_constrs, specialize, skolemize, star_linear_preds, uncurry,
fast_descrs, peephole_optim, tac_timeout, sym_break, sharing_depth,
@@ -231,7 +237,16 @@
if passed_deadline deadline then raise TimeLimit.TimeOut
else raise Interrupt
- val _ = print_m (K "Nitpicking...")
+ val _ =
+ if step = 0 then
+ print_m (fn () => "Nitpicking formula...")
+ else
+ pprint_m (fn () => Pretty.chunks (
+ pretties_for_formulas ctxt ("Nitpicking " ^
+ (if i <> 1 orelse n <> 1 then
+ "subgoal " ^ string_of_int i ^ " of " ^ string_of_int n
+ else
+ "goal")) [orig_t]))
val neg_t = if falsify then Logic.mk_implies (orig_t, @{prop False})
else orig_t
val assms_t = if assms orelse auto then
@@ -260,7 +275,7 @@
val ersatz_table = ersatz_table thy
val (ext_ctxt as {wf_cache, ...}) =
{thy = thy, ctxt = ctxt, max_bisim_depth = max_bisim_depth, boxes = boxes,
- user_axioms = user_axioms, debug = debug, wfs = wfs,
+ stds = stds, wfs = wfs, user_axioms = user_axioms, debug = debug,
binary_ints = binary_ints, destroy_constrs = destroy_constrs,
specialize = specialize, skolemize = skolemize,
star_linear_preds = star_linear_preds, uncurry = uncurry,
@@ -311,7 +326,8 @@
fold add_free_and_const_names (core_u :: def_us @ nondef_us) ([], [])
val (sel_names, nonsel_names) =
List.partition (is_sel o nickname_of) const_names
- val would_be_genuine = got_all_user_axioms andalso none_true wfs
+ val genuine_means_genuine = got_all_user_axioms andalso none_true wfs
+ val standard = forall snd stds
(*
val _ = List.app (priority o string_for_nut ctxt)
(core_u :: def_us @ nondef_us)
@@ -322,27 +338,27 @@
fun is_type_always_monotonic T =
(is_datatype thy T andalso not (is_quot_type thy T) andalso
(not (is_pure_typedef thy T) orelse is_univ_typedef thy T)) orelse
- is_number_type thy T orelse is_bit_type T
+ is_number_type thy T orelse is_bit_type T orelse T = @{typ \<xi>}
fun is_type_monotonic T =
unique_scope orelse
case triple_lookup (type_match thy) monos T of
SOME (SOME b) => b
| _ => is_type_always_monotonic T orelse
- formulas_monotonic ext_ctxt T def_ts nondef_ts core_t
+ formulas_monotonic ext_ctxt T Plus def_ts nondef_ts core_t
fun is_deep_datatype T =
is_datatype thy T andalso
(is_word_type T orelse
exists (curry (op =) T o domain_type o type_of) sel_names)
- val Ts = ground_types_in_terms ext_ctxt (core_t :: def_ts @ nondef_ts)
- |> sort TermOrd.typ_ord
+ val all_Ts = ground_types_in_terms ext_ctxt (core_t :: def_ts @ nondef_ts)
+ |> sort TermOrd.typ_ord
val _ = if verbose andalso binary_ints = SOME true andalso
- exists (member (op =) [nat_T, int_T]) Ts then
+ exists (member (op =) [nat_T, int_T]) all_Ts then
print_v (K "The option \"binary_ints\" will be ignored because \
\of the presence of rationals, reals, \"Suc\", \
\\"gcd\", or \"lcm\" in the problem.")
else
()
- val (mono_Ts, nonmono_Ts) = List.partition is_type_monotonic Ts
+ val (mono_Ts, nonmono_Ts) = List.partition is_type_monotonic all_Ts
val _ =
if verbose andalso not unique_scope then
case filter_out is_type_always_monotonic mono_Ts of
@@ -364,7 +380,32 @@
end)
else
()
- val shallow_dataTs = filter_out is_deep_datatype Ts
+ val deep_dataTs = filter is_deep_datatype all_Ts
+ (* FIXME: Implement proper detection of induction datatypes. *)
+ (* string * (Rule_Cases.T * bool) -> bool *)
+ fun is_inductive_case (_, (Rule_Cases.Case {fixes, assumes, ...}, _)) =
+ false
+(*
+ not (null fixes) andalso exists (String.isSuffix ".hyps" o fst) assumes
+*)
+ (* unit -> typ list *)
+ val induct_dataTs =
+ if exists is_inductive_case (ProofContext.cases_of ctxt) then
+ filter (is_datatype thy) all_Ts
+ else
+ []
+ val _ = if standard andalso not (null induct_dataTs) then
+ pprint_m (fn () => Pretty.blk (0,
+ pstrs "Hint: To check that the induction hypothesis is \
+ \general enough, try the following command: " @
+ [Pretty.mark Markup.sendback (Pretty.blk (0,
+ pstrs ("nitpick [" ^
+ commas (map (prefix "non_std " o maybe_quote
+ o unyxml o string_for_type ctxt)
+ induct_dataTs) ^
+ ", show_consts]")))] @ pstrs "."))
+ else
+ ()
(*
val _ = priority "Monotonic types:"
val _ = List.app (priority o string_for_type ctxt) mono_Ts
@@ -413,7 +454,7 @@
string_of_int j0))
(Typtab.dest ofs)
*)
- val all_exact = forall (is_exact_type datatypes) Ts
+ val all_exact = forall (is_exact_type datatypes) all_Ts
(* nut list -> rep NameTable.table -> nut list * rep NameTable.table *)
val repify_consts = choose_reps_for_consts scope all_exact
val main_j0 = offset_of_type ofs bool_T
@@ -536,7 +577,9 @@
fun tuple_set_for_rel univ_card =
KK.TupleSet o map (kk_tuple debug univ_card) o the oo AList.lookup (op =)
- val word_model = if falsify then "counterexample" else "model"
+ val das_wort_model =
+ (if falsify then "counterexample" else "model")
+ |> not standard ? prefix "nonstandard "
val scopes = Unsynchronized.ref []
val generated_scopes = Unsynchronized.ref []
@@ -560,42 +603,49 @@
show_consts = show_consts}
scope formats frees free_names sel_names nonsel_names rel_table
bounds
- val would_be_genuine = would_be_genuine andalso codatatypes_ok
+ val genuine_means_genuine = genuine_means_genuine andalso codatatypes_ok
in
pprint (Pretty.chunks
[Pretty.blk (0,
(pstrs ("Nitpick found a" ^
(if not genuine then " potential "
- else if would_be_genuine then " "
- else " likely genuine ") ^ word_model) @
+ else if genuine_means_genuine then " "
+ else " likely genuine ") ^ das_wort_model) @
(case pretties_for_scope scope verbose of
[] => []
| pretties => pstrs " for " @ pretties) @
[Pretty.str ":\n"])),
Pretty.indent indent_size reconstructed_model]);
if genuine then
- (if check_genuine then
+ (if check_genuine andalso standard then
(case prove_hol_model scope tac_timeout free_names sel_names
rel_table bounds assms_t of
SOME true => print ("Confirmation by \"auto\": The above " ^
- word_model ^ " is really genuine.")
+ das_wort_model ^ " is really genuine.")
| SOME false =>
- if would_be_genuine then
- error ("A supposedly genuine " ^ word_model ^ " was shown to\
- \be spurious by \"auto\".\nThis should never happen.\n\
- \Please send a bug report to blanchet\
+ if genuine_means_genuine then
+ error ("A supposedly genuine " ^ das_wort_model ^ " was \
+ \shown to be spurious by \"auto\".\nThis should never \
+ \happen.\nPlease send a bug report to blanchet\
\te@in.tum.de.")
else
- print ("Refutation by \"auto\": The above " ^ word_model ^
+ print ("Refutation by \"auto\": The above " ^ das_wort_model ^
" is spurious.")
| NONE => print "No confirmation by \"auto\".")
else
();
+ if not standard andalso not (null induct_dataTs) then
+ print "The existence of a nonstandard model suggests that the \
+ \induction hypothesis is not general enough or perhaps even \
+ \wrong. See the \"Inductive Properties\" section of the \
+ \Nitpick manual for details (\"isabelle doc nitpick\")."
+ else
+ ();
if has_syntactic_sorts orig_t then
print "Hint: Maybe you forgot a type constraint?"
else
();
- if not would_be_genuine then
+ if not genuine_means_genuine then
if no_poly_user_axioms then
let
val options =
@@ -611,12 +661,13 @@
in
print ("Try again with " ^
space_implode " " (serial_commas "and" ss) ^
- " to confirm that the " ^ word_model ^ " is genuine.")
+ " to confirm that the " ^ das_wort_model ^
+ " is genuine.")
end
else
print ("Nitpick is unable to guarantee the authenticity of \
- \the " ^ word_model ^ " in the presence of polymorphic \
- \axioms.")
+ \the " ^ das_wort_model ^ " in the presence of \
+ \polymorphic axioms.")
else
();
NONE)
@@ -630,9 +681,9 @@
in
(case status of
SOME true => print ("Confirmation by \"auto\": The above " ^
- word_model ^ " is genuine.")
+ das_wort_model ^ " is genuine.")
| SOME false => print ("Refutation by \"auto\": The above " ^
- word_model ^ " is spurious.")
+ das_wort_model ^ " is spurious.")
| NONE => print "No confirmation by \"auto\".");
status
end
@@ -820,14 +871,20 @@
(print_m (fn () => excipit "ran out of some resource"); "unknown")
else if max_genuine = original_max_genuine then
if max_potential = original_max_potential then
- (print_m (K ("Nitpick found no " ^ word_model ^ ".")); "none")
+ (print_m (fn () =>
+ "Nitpick found no " ^ das_wort_model ^ "." ^
+ (if not standard andalso not (null induct_dataTs) then
+ " This suggests that the induction hypothesis might be \
+ \general enough to prove this subgoal."
+ else
+ "")); "none")
else
- (print_m (K ("Nitpick could not find " ^
- (if max_genuine = 1 then "a better " ^ word_model ^ "."
- else "any better " ^ word_model ^ "s.")));
- "potential")
+ (print_m (fn () =>
+ "Nitpick could not find a" ^
+ (if max_genuine = 1 then " better " ^ das_wort_model ^ "."
+ else "ny better " ^ das_wort_model ^ "s.")); "potential")
else
- if would_be_genuine then "genuine" else "likely_genuine"
+ if genuine_means_genuine then "genuine" else "likely_genuine"
| run_batches j n (batch :: batches) z =
let val (z as (_, max_genuine, _)) = run_batch j n batch z in
run_batches (j + 1) n (if max_genuine > 0 then batches else []) z
@@ -835,7 +892,7 @@
val (skipped, the_scopes) =
all_scopes ext_ctxt sym_break cards_assigns maxes_assigns iters_assigns
- bitss bisim_depths mono_Ts nonmono_Ts shallow_dataTs
+ bitss bisim_depths mono_Ts nonmono_Ts deep_dataTs
val _ = if skipped > 0 then
print_m (fn () => "Too many scopes. Skipping " ^
string_of_int skipped ^ " scope" ^
@@ -868,9 +925,10 @@
else
error "Nitpick was interrupted."
-(* Proof.state -> params -> bool -> term -> string * Proof.state *)
-fun pick_nits_in_term state (params as {debug, timeout, expect, ...}) auto
- orig_assm_ts orig_t =
+(* Proof.state -> params -> bool -> int -> int -> int -> term
+ -> string * Proof.state *)
+fun pick_nits_in_term state (params as {debug, timeout, expect, ...}) auto i n
+ step orig_assm_ts orig_t =
if getenv "KODKODI" = "" then
(if auto then ()
else warning (Pretty.string_of (plazy install_kodkodi_message));
@@ -880,26 +938,29 @@
val deadline = Option.map (curry Time.+ (Time.now ())) timeout
val outcome as (outcome_code, _) =
time_limit (if debug then NONE else timeout)
- (pick_them_nits_in_term deadline state params auto orig_assm_ts)
- orig_t
+ (pick_them_nits_in_term deadline state params auto i n step
+ orig_assm_ts) orig_t
in
if expect = "" orelse outcome_code = expect then outcome
else error ("Unexpected outcome: " ^ quote outcome_code ^ ".")
end
-(* Proof.state -> params -> thm -> int -> string * Proof.state *)
-fun pick_nits_in_subgoal state params auto i =
+(* Proof.state -> params -> bool -> int -> int -> string * Proof.state *)
+fun pick_nits_in_subgoal state params auto i step =
let
val ctxt = Proof.context_of state
val t = state |> Proof.raw_goal |> #goal |> prop_of
in
- if Logic.count_prems t = 0 then
- (priority "No subgoal!"; ("none", state))
- else
+ case Logic.count_prems t of
+ 0 => (priority "No subgoal!"; ("none", state))
+ | n =>
let
val assms = map term_of (Assumption.all_assms_of ctxt)
val (t, frees) = Logic.goal_params t i
- in pick_nits_in_term state params auto assms (subst_bounds (frees, t)) end
+ in
+ pick_nits_in_term state params auto i n step assms
+ (subst_bounds (frees, t))
+ end
end
end;
--- a/src/HOL/Tools/Nitpick/nitpick_hol.ML Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML Tue Feb 02 18:16:48 2010 +0100
@@ -1,6 +1,6 @@
(* Title: HOL/Tools/Nitpick/nitpick_hol.ML
Author: Jasmin Blanchette, TU Muenchen
- Copyright 2008, 2009
+ Copyright 2008, 2009, 2010
Auxiliary HOL-related functions used by Nitpick.
*)
@@ -18,6 +18,7 @@
ctxt: Proof.context,
max_bisim_depth: int,
boxes: (typ option * bool option) list,
+ stds: (typ option * bool) list,
wfs: (styp option * bool option) list,
user_axioms: bool option,
debug: bool,
@@ -165,6 +166,7 @@
ctxt: Proof.context,
max_bisim_depth: int,
boxes: (typ option * bool option) list,
+ stds: (typ option * bool) list,
wfs: (styp option * bool option) list,
user_axioms: bool option,
debug: bool,
@@ -548,7 +550,7 @@
val is_typedef = is_some oo typedef_info
val is_real_datatype = is_some oo Datatype.get_info
(* theory -> typ -> bool *)
-fun is_quot_type _ (Type ("IntEx.my_int", _)) = true (* ### *)
+fun is_quot_type _ (Type ("IntEx.my_int", _)) = true (* FIXME *)
| is_quot_type _ _ = false
fun is_codatatype thy (T as Type (s, _)) =
not (null (AList.lookup (op =) (#codatatypes (Data.get thy)) s
@@ -615,9 +617,9 @@
| NONE => false)
| is_rep_fun _ _ = false
(* Proof.context -> styp -> bool *)
-fun is_quot_abs_fun _ ("IntEx.abs_my_int", _) = true (* ### *)
+fun is_quot_abs_fun _ ("IntEx.abs_my_int", _) = true (* FIXME *)
| is_quot_abs_fun _ _ = false
-fun is_quot_rep_fun _ ("IntEx.rep_my_int", _) = true (* ### *)
+fun is_quot_rep_fun _ ("IntEx.rep_my_int", _) = true (* FIXME *)
| is_quot_rep_fun _ _ = false
(* theory -> styp -> styp *)
@@ -648,12 +650,12 @@
end
handle TYPE ("dest_Type", _, _) => false
fun is_constr_like thy (s, T) =
- s = @{const_name FunBox} orelse s = @{const_name PairBox} orelse
+ member (op =) [@{const_name FunBox}, @{const_name PairBox},
+ @{const_name Quot}, @{const_name Zero_Rep},
+ @{const_name Suc_Rep}] s orelse
let val (x as (s, T)) = (s, unbit_and_unbox_type T) in
Refute.is_IDT_constructor thy x orelse is_record_constr x orelse
(is_abs_fun thy x andalso is_pure_typedef thy (range_type T)) orelse
- s = @{const_name Quot} orelse
- s = @{const_name Zero_Rep} orelse s = @{const_name Suc_Rep} orelse
x = (@{const_name zero_nat_inst.zero_nat}, nat_T) orelse
is_coconstr thy x
end
@@ -710,7 +712,8 @@
box_type ext_ctxt (in_fun_lhs_for boxy) T1
--> box_type ext_ctxt (in_fun_rhs_for boxy) T2
| Type (z as ("*", Ts)) =>
- if should_box_type ext_ctxt boxy z then
+ if boxy <> InConstr andalso boxy <> InSel
+ andalso should_box_type ext_ctxt boxy z then
Type (@{type_name pair_box}, map (box_type ext_ctxt InSel) Ts)
else
Type ("*", map (box_type ext_ctxt
@@ -778,11 +781,11 @@
fun zero_const T = Const (@{const_name zero_nat_inst.zero_nat}, T)
fun suc_const T = Const (@{const_name Suc}, T --> T)
-(* theory -> typ -> styp list *)
-fun uncached_datatype_constrs thy (T as Type (s, Ts)) =
+(* extended_context -> typ -> styp list *)
+fun uncached_datatype_constrs ({thy, stds, ...} : extended_context)
+ (T as Type (s, Ts)) =
(case AList.lookup (op =) (#codatatypes (Data.get thy)) s of
- SOME (_, xs' as (_ :: _)) =>
- map (apsnd (repair_constr_type thy T)) xs'
+ SOME (_, xs' as (_ :: _)) => map (apsnd (repair_constr_type thy T)) xs'
| _ =>
if is_datatype thy T then
case Datatype.get_info thy s of
@@ -793,6 +796,8 @@
map (fn (s', Us) =>
(s', map (Refute.typ_of_dtyp descr (dtyps ~~ Ts)) Us
---> T)) constrs
+ |> (triple_lookup (type_match thy) stds T |> the |> not) ?
+ cons (@{const_name NonStd}, @{typ \<xi>} --> T)
end
| NONE =>
if is_record_type T then
@@ -815,12 +820,11 @@
[])
| uncached_datatype_constrs _ _ = []
(* extended_context -> typ -> styp list *)
-fun datatype_constrs (ext_ctxt as {thy, constr_cache, ...} : extended_context)
- T =
+fun datatype_constrs (ext_ctxt as {constr_cache, ...}) T =
case AList.lookup (op =) (!constr_cache) T of
SOME xs => xs
| NONE =>
- let val xs = uncached_datatype_constrs thy T in
+ let val xs = uncached_datatype_constrs ext_ctxt T in
(Unsynchronized.change constr_cache (cons (T, xs)); xs)
end
fun boxed_datatype_constrs ext_ctxt =
@@ -838,6 +842,7 @@
AList.lookup (op =) (boxed_datatype_constrs ext_ctxt (domain_type T')) s
|> the |> pair s
end
+
(* extended_context -> styp -> term *)
fun discr_term_for_constr ext_ctxt (x as (s, T)) =
let val dataT = body_type T in
@@ -849,7 +854,6 @@
else
Abs (Name.uu, dataT, @{const True})
end
-
(* extended_context -> styp -> term -> term *)
fun discriminate_value (ext_ctxt as {thy, ...}) (x as (_, T)) t =
case strip_comb t of
@@ -919,7 +923,7 @@
(@{const_name Pair}, T1 --> T2 --> T)
end
else
- datatype_constrs ext_ctxt T |> the_single
+ datatype_constrs ext_ctxt T |> hd
val arg_Ts = binder_types T'
in
list_comb (Const x', map2 (select_nth_constr_arg thy x' t)
@@ -1361,10 +1365,9 @@
val normal_y = normal_t $ y_var
val is_unknown_t = Const (@{const_name is_unknown}, rep_T --> bool_T)
in
- [Logic.mk_equals (sel_a_t, normal_t $ sel_a_t),
+ [Logic.mk_equals (normal_t $ sel_a_t, sel_a_t),
Logic.list_implies
- ([@{const Not} $ (HOLogic.mk_eq (x_var, y_var)),
- @{const Not} $ (is_unknown_t $ normal_x),
+ ([@{const Not} $ (is_unknown_t $ normal_x),
@{const Not} $ (is_unknown_t $ normal_y),
equiv_rel $ x_var $ y_var] |> map HOLogic.mk_Trueprop,
Logic.mk_equals (normal_x, normal_y)),
@@ -1388,17 +1391,27 @@
fun optimized_case_def (ext_ctxt as {thy, ...}) dataT res_T =
let
val xs = datatype_constrs ext_ctxt dataT
- val func_Ts = map ((fn T => binder_types T ---> res_T) o snd) xs
- val (xs', x) = split_last xs
+ val xs' = filter_out (fn (s, _) => s = @{const_name NonStd}) xs
+ val func_Ts = map ((fn T => binder_types T ---> res_T) o snd) xs'
in
- constr_case_body thy (1, x)
- |> fold_rev (add_constr_case ext_ctxt res_T) (length xs downto 2 ~~ xs')
+ (if length xs = length xs' then
+ let
+ val (xs'', x) = split_last xs'
+ in
+ constr_case_body thy (1, x)
+ |> fold_rev (add_constr_case ext_ctxt res_T)
+ (length xs' downto 2 ~~ xs'')
+ end
+ else
+ Const (@{const_name undefined}, dataT --> res_T) $ Bound 0
+ |> fold_rev (add_constr_case ext_ctxt res_T)
+ (length xs' downto 1 ~~ xs'))
|> fold_rev (curry absdummy) (func_Ts @ [dataT])
end
(* extended_context -> string -> typ -> typ -> term -> term *)
fun optimized_record_get (ext_ctxt as {thy, ...}) s rec_T res_T t =
- let val constr_x = the_single (datatype_constrs ext_ctxt rec_T) in
+ let val constr_x = hd (datatype_constrs ext_ctxt rec_T) in
case no_of_record_field thy s rec_T of
~1 => (case rec_T of
Type (_, Ts as _ :: _) =>
@@ -1416,7 +1429,7 @@
(* extended_context -> string -> typ -> term -> term -> term *)
fun optimized_record_update (ext_ctxt as {thy, ...}) s rec_T fun_t rec_t =
let
- val constr_x as (_, constr_T) = the_single (datatype_constrs ext_ctxt rec_T)
+ val constr_x as (_, constr_T) = hd (datatype_constrs ext_ctxt rec_T)
val Ts = binder_types constr_T
val n = length Ts
val special_j = no_of_record_field thy s rec_T
@@ -1606,7 +1619,7 @@
case length ts of
0 => (do_term depth Ts (eta_expand Ts t 1), [])
| _ => (optimized_record_get ext_ctxt s (domain_type T)
- (range_type T) (hd ts), tl ts)
+ (range_type T) (do_term depth Ts (hd ts)), tl ts)
else if is_record_update thy x then
case length ts of
2 => (optimized_record_update ext_ctxt
@@ -2077,7 +2090,7 @@
(t0 as Const (@{const_name "=="}, _)) $ t1 $ t2 =>
do_eq_or_imp Ts true def t0 t1 t2 seen
| (t0 as @{const "==>"}) $ t1 $ t2 =>
- do_eq_or_imp Ts false def t0 t1 t2 seen
+ if def then (t, []) else do_eq_or_imp Ts false def t0 t1 t2 seen
| (t0 as Const (@{const_name "op ="}, _)) $ t1 $ t2 =>
do_eq_or_imp Ts true def t0 t1 t2 seen
| (t0 as @{const "op -->"}) $ t1 $ t2 =>
@@ -2126,27 +2139,33 @@
| aux _ t = t
(* bool -> bool -> term -> term -> term -> term *)
and aux_eq careful pass1 t0 t1 t2 =
- (if careful then
- raise SAME ()
- else if axiom andalso is_Var t2 andalso
- num_occs_of_var (dest_Var t2) = 1 then
- @{const True}
- else case strip_comb t2 of
- (Const (x as (s, T)), args) =>
- let val arg_Ts = binder_types T in
- if length arg_Ts = length args andalso
- (is_constr thy x orelse s = @{const_name Pair} orelse
- x = (@{const_name Suc}, nat_T --> nat_T)) andalso
- (not careful orelse not (is_Var t1) orelse
- String.isPrefix val_var_prefix (fst (fst (dest_Var t1)))) then
- discriminate_value ext_ctxt x t1 ::
- map3 (sel_eq x t1) (index_seq 0 (length args)) arg_Ts args
- |> foldr1 s_conj
- |> body_type (type_of t0) = prop_T ? HOLogic.mk_Trueprop
- else
- raise SAME ()
- end
- | _ => raise SAME ())
+ ((if careful then
+ raise SAME ()
+ else if axiom andalso is_Var t2 andalso
+ num_occs_of_var (dest_Var t2) = 1 then
+ @{const True}
+ else case strip_comb t2 of
+ (* The first case is not as general as it could be. *)
+ (Const (@{const_name PairBox}, _),
+ [Const (@{const_name fst}, _) $ Var z1,
+ Const (@{const_name snd}, _) $ Var z2]) =>
+ if z1 = z2 andalso num_occs_of_var z1 = 2 then @{const True}
+ else raise SAME ()
+ | (Const (x as (s, T)), args) =>
+ let val arg_Ts = binder_types T in
+ if length arg_Ts = length args andalso
+ (is_constr thy x orelse s = @{const_name Pair} orelse
+ x = (@{const_name Suc}, nat_T --> nat_T)) andalso
+ (not careful orelse not (is_Var t1) orelse
+ String.isPrefix val_var_prefix (fst (fst (dest_Var t1)))) then
+ discriminate_value ext_ctxt x t1 ::
+ map3 (sel_eq x t1) (index_seq 0 (length args)) arg_Ts args
+ |> foldr1 s_conj
+ else
+ raise SAME ()
+ end
+ | _ => raise SAME ())
+ |> body_type (type_of t0) = prop_T ? HOLogic.mk_Trueprop)
handle SAME () => if pass1 then aux_eq careful false t0 t2 t1
else t0 $ aux false t2 $ aux false t1
(* styp -> term -> int -> typ -> term -> term *)
@@ -2947,7 +2966,7 @@
| Const (s as @{const_name The}, T) => do_description_operator s T
| Const (s as @{const_name Eps}, T) => do_description_operator s T
| Const (@{const_name quot_normal}, Type ("fun", [_, T2])) =>
- let val T' = box_type ext_ctxt InFunRHS1 T2 in
+ let val T' = box_type ext_ctxt InSel T2 in
Const (@{const_name quot_normal}, T' --> T')
end
| Const (s as @{const_name Tha}, T) => do_description_operator s T
@@ -2960,7 +2979,8 @@
T
else if is_constr_like thy x then
box_type ext_ctxt InConstr T
- else if is_sel s orelse is_rep_fun thy x then
+ else if is_sel s
+ orelse is_rep_fun thy x then
box_type ext_ctxt InSel T
else
box_type ext_ctxt InExpr T)
@@ -3495,6 +3515,7 @@
#> simplify_constrs_and_sels thy
#> distribute_quantifiers
#> push_quantifiers_inward thy
+ #> Envir.eta_contract
#> not core ? Refute.close_form
#> shorten_abs_vars
in
--- a/src/HOL/Tools/Nitpick/nitpick_isar.ML Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_isar.ML Tue Feb 02 18:16:48 2010 +0100
@@ -1,6 +1,6 @@
(* Title: HOL/Tools/Nitpick/nitpick_isar.ML
Author: Jasmin Blanchette, TU Muenchen
- Copyright 2008, 2009
+ Copyright 2008, 2009, 2010
Adds the "nitpick" and "nitpick_params" commands to Isabelle/Isar's outer
syntax.
@@ -28,9 +28,8 @@
val _ =
ProofGeneralPgip.add_preference Preferences.category_tracing
- (Preferences.bool_pref auto
- "auto-nitpick"
- "Whether to run Nitpick automatically.")
+ (Preferences.bool_pref auto "auto-nitpick"
+ "Whether to run Nitpick automatically.")
type raw_param = string * string list
@@ -41,6 +40,7 @@
("bisim_depth", ["7"]),
("box", ["smart"]),
("mono", ["smart"]),
+ ("std", ["true"]),
("wf", ["smart"]),
("sat_solver", ["smart"]),
("batch_size", ["smart"]),
@@ -79,6 +79,7 @@
val negated_params =
[("dont_box", "box"),
("non_mono", "mono"),
+ ("non_std", "std"),
("non_wf", "wf"),
("non_blocking", "blocking"),
("satisfy", "falsify"),
@@ -110,8 +111,8 @@
AList.defined (op =) negated_params s orelse
s = "max" orelse s = "eval" orelse s = "expect" orelse
exists (fn p => String.isPrefix (p ^ " ") s)
- ["card", "max", "iter", "box", "dont_box", "mono", "non_mono", "wf",
- "non_wf", "format"]
+ ["card", "max", "iter", "box", "dont_box", "mono", "non_mono", "std",
+ "non_std", "wf", "non_wf", "format"]
(* string * 'a -> unit *)
fun check_raw_param (s, _) =
@@ -244,6 +245,14 @@
value |> stringify_raw_param_value
|> int_seq_from_string name min_int))
(filter (String.isPrefix (prefix ^ " ") o fst) raw_params)
+ (* (string -> 'a) -> string -> ('a option * bool) list *)
+ fun lookup_bool_assigns read prefix =
+ (NONE, lookup_bool prefix)
+ :: map (fn (name, value) =>
+ (SOME (read (String.extract (name, size prefix + 1, NONE))),
+ value |> stringify_raw_param_value
+ |> bool_option_from_string false name |> the))
+ (filter (String.isPrefix (prefix ^ " ") o fst) raw_params)
(* (string -> 'a) -> string -> ('a option * bool option) list *)
fun lookup_bool_option_assigns read prefix =
(NONE, lookup_bool_option prefix)
@@ -297,6 +306,7 @@
NONE
| (NONE, _) => NONE) cards_assigns
val monos = lookup_bool_option_assigns read_type_polymorphic "mono"
+ val stds = lookup_bool_assigns read_type_polymorphic "std"
val wfs = lookup_bool_option_assigns read_const_polymorphic "wf"
val sat_solver = lookup_string "sat_solver"
val blocking = not auto andalso lookup_bool "blocking"
@@ -339,9 +349,10 @@
in
{cards_assigns = cards_assigns, maxes_assigns = maxes_assigns,
iters_assigns = iters_assigns, bitss = bitss, bisim_depths = bisim_depths,
- boxes = boxes, monos = monos, wfs = wfs, sat_solver = sat_solver,
- blocking = blocking, falsify = falsify, debug = debug, verbose = verbose,
- overlord = overlord, user_axioms = user_axioms, assms = assms,
+ boxes = boxes, monos = monos, stds = stds, wfs = wfs,
+ sat_solver = sat_solver, blocking = blocking, falsify = falsify,
+ debug = debug, verbose = verbose, overlord = overlord,
+ user_axioms = user_axioms, assms = assms,
merge_type_vars = merge_type_vars, binary_ints = binary_ints,
destroy_constrs = destroy_constrs, specialize = specialize,
skolemize = skolemize, star_linear_preds = star_linear_preds,
@@ -416,8 +427,9 @@
| Refute.REFUTE (loc, details) =>
error ("Unhandled Refute error (" ^ quote loc ^ "): " ^ details ^ ".")
-(* raw_param list -> bool -> int -> Proof.state -> bool * Proof.state *)
-fun pick_nits override_params auto i state =
+
+(* raw_param list -> bool -> int -> int -> Proof.state -> bool * Proof.state *)
+fun pick_nits override_params auto i step state =
let
val thy = Proof.theory_of state
val ctxt = Proof.context_of state
@@ -431,7 +443,7 @@
|> (if auto then perhaps o try
else if debug then fn f => fn x => f x
else handle_exceptions ctxt)
- (fn (_, state) => pick_nits_in_subgoal state params auto i
+ (fn (_, state) => pick_nits_in_subgoal state params auto i step
|>> curry (op =) "genuine")
in
if auto orelse blocking then go ()
@@ -441,9 +453,9 @@
(* raw_param list option * int option -> Toplevel.transition
-> Toplevel.transition *)
fun nitpick_trans (opt_params, opt_i) =
- Toplevel.keep (K ()
- o snd o pick_nits (these opt_params) false (the_default 1 opt_i)
- o Toplevel.proof_of)
+ Toplevel.keep (fn st =>
+ (pick_nits (these opt_params) false (the_default 1 opt_i)
+ (Toplevel.proof_position_of st) (Toplevel.proof_of st); ()))
(* raw_param -> string *)
fun string_for_raw_param (name, value) =
@@ -477,7 +489,7 @@
(* Proof.state -> bool * Proof.state *)
fun auto_nitpick state =
- if not (!auto) then (false, state) else pick_nits [] true 1 state
+ if not (!auto) then (false, state) else pick_nits [] true 1 0 state
val setup = Auto_Counterexample.register_tool ("nitpick", auto_nitpick)
--- a/src/HOL/Tools/Nitpick/nitpick_kodkod.ML Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_kodkod.ML Tue Feb 02 18:16:48 2010 +0100
@@ -1,6 +1,6 @@
(* Title: HOL/Tools/Nitpick/nitpick_kodkod.ML
Author: Jasmin Blanchette, TU Muenchen
- Copyright 2008, 2009
+ Copyright 2008, 2009, 2010
Kodkod problem generator part of Kodkod.
*)
@@ -285,8 +285,8 @@
(* Proof.context -> bool -> string -> typ -> rep -> string *)
fun bound_comment ctxt debug nick T R =
short_name nick ^
- (if debug then " :: " ^ plain_string_from_yxml (Syntax.string_of_typ ctxt T)
- else "") ^ " : " ^ string_for_rep R
+ (if debug then " :: " ^ unyxml (Syntax.string_of_typ ctxt T) else "") ^
+ " : " ^ string_for_rep R
(* Proof.context -> bool -> nut -> KK.bound *)
fun bound_for_plain_rel ctxt debug (u as FreeRel (x, T, R, nick)) =
@@ -754,7 +754,7 @@
(* extended_context -> kodkod_constrs -> nut NameTable.table -> dtype_spec list
-> dtype_spec -> nfa_entry option *)
fun nfa_entry_for_datatype _ _ _ _ ({co = true, ...} : dtype_spec) = NONE
- | nfa_entry_for_datatype _ _ _ _ {shallow = true, ...} = NONE
+ | nfa_entry_for_datatype _ _ _ _ {deep = false, ...} = NONE
| nfa_entry_for_datatype ext_ctxt kk rel_table dtypes {typ, constrs, ...} =
SOME (typ, maps (nfa_transitions_for_constr ext_ctxt kk rel_table dtypes
o #const) constrs)
@@ -926,7 +926,7 @@
(* extended_context -> int -> int Typtab.table -> kodkod_constrs
-> nut NameTable.table -> dtype_spec -> KK.formula list *)
-fun other_axioms_for_datatype _ _ _ _ _ {shallow = true, ...} = []
+fun other_axioms_for_datatype _ _ _ _ _ {deep = false, ...} = []
| other_axioms_for_datatype ext_ctxt bits ofs kk rel_table
(dtype as {typ, ...}) =
let val j0 = offset_of_type ofs typ in
--- a/src/HOL/Tools/Nitpick/nitpick_model.ML Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_model.ML Tue Feb 02 18:16:48 2010 +0100
@@ -1,6 +1,6 @@
(* Title: HOL/Tools/Nitpick/nitpick_model.ML
Author: Jasmin Blanchette, TU Muenchen
- Copyright 2009
+ Copyright 2009, 2010
Model reconstruction for Nitpick.
*)
@@ -53,7 +53,8 @@
val base_mixfix = "_\<^bsub>base\<^esub>"
val step_mixfix = "_\<^bsub>step\<^esub>"
val abs_mixfix = "\<guillemotleft>_\<guillemotright>"
-val non_opt_name = nitpick_prefix ^ "non_opt"
+val opt_flag = nitpick_prefix ^ "opt"
+val non_opt_flag = nitpick_prefix ^ "non_opt"
(* string -> int -> string *)
fun atom_suffix s j =
@@ -62,7 +63,10 @@
? prefix "\<^isub>,"
(* string -> typ -> int -> string *)
fun atom_name prefix (T as Type (s, _)) j =
- prefix ^ substring (shortest_name s, 0, 1) ^ atom_suffix s j
+ let val s' = shortest_name s in
+ prefix ^ (if String.isPrefix "\\" s' then s' else substring (s', 0, 1)) ^
+ atom_suffix s j
+ end
| atom_name prefix (T as TFree (s, _)) j =
prefix ^ perhaps (try (unprefix "'")) s ^ atom_suffix s j
| atom_name _ T _ = raise TYPE ("Nitpick_Model.atom_name", [T], [])
@@ -139,23 +143,21 @@
let
(* typ -> typ -> (term * term) list -> term *)
fun aux T1 T2 [] =
- Const (if maybe_opt orelse T2 <> bool_T then @{const_name undefined}
- else non_opt_name, T1 --> T2)
+ Const (if maybe_opt then opt_flag else non_opt_flag, T1 --> T2)
| aux T1 T2 ((t1, t2) :: ps) =
Const (@{const_name fun_upd}, (T1 --> T2) --> T1 --> T2 --> T1 --> T2)
$ aux T1 T2 ps $ t1 $ t2
in aux T1 T2 o rev end
(* term -> bool *)
-fun is_plain_fun (Const (s, _)) =
- (s = @{const_name undefined} orelse s = non_opt_name)
+fun is_plain_fun (Const (s, _)) = (s = opt_flag orelse s = non_opt_flag)
| is_plain_fun (Const (@{const_name fun_upd}, _) $ t0 $ _ $ _) =
is_plain_fun t0
| is_plain_fun _ = false
(* term -> bool * (term list * term list) *)
val dest_plain_fun =
let
- (* term -> term list * term list *)
- fun aux (Const (s, _)) = (s <> non_opt_name, ([], []))
+ (* term -> bool * (term list * term list) *)
+ fun aux (Const (s, _)) = (s <> non_opt_flag, ([], []))
| aux (Const (@{const_name fun_upd}, _) $ t0 $ t1 $ t2) =
let val (s, (ts1, ts2)) = aux t0 in (s, (t1 :: ts1, t2 :: ts2)) end
| aux t = raise TERM ("Nitpick_Model.dest_plain_fun", [t])
@@ -300,7 +302,7 @@
aux' ps
in aux end
(* typ list -> term -> term *)
- fun setify_mapify_funs Ts t =
+ fun polish_funs Ts t =
(case fastype_of1 (Ts, t) of
Type ("fun", [T1, T2]) =>
if is_plain_fun t then
@@ -308,7 +310,7 @@
@{typ bool} =>
let
val (maybe_opt, ts_pair) =
- dest_plain_fun t ||> pairself (map (setify_mapify_funs Ts))
+ dest_plain_fun t ||> pairself (map (polish_funs Ts))
in
make_set maybe_opt T1 T2
(sort_wrt (truth_const_sort_key o snd) (op ~~ ts_pair))
@@ -316,7 +318,7 @@
| Type (@{type_name option}, [T2']) =>
let
val ts_pair = snd (dest_plain_fun t)
- |> pairself (map (setify_mapify_funs Ts))
+ |> pairself (map (polish_funs Ts))
in make_map T1 T2 T2' (rev (op ~~ ts_pair)) end
| _ => raise SAME ()
else
@@ -324,9 +326,18 @@
| _ => raise SAME ())
handle SAME () =>
case t of
- t1 $ t2 => setify_mapify_funs Ts t1 $ setify_mapify_funs Ts t2
- | Abs (s, T, t') => Abs (s, T, setify_mapify_funs (T :: Ts) t')
- | _ => t
+ (t1 as Const (@{const_name fun_upd}, _) $ t11 $ _)
+ $ (t2 as Const (s, _)) =>
+ if s = unknown then polish_funs Ts t11
+ else polish_funs Ts t1 $ polish_funs Ts t2
+ | t1 $ t2 => polish_funs Ts t1 $ polish_funs Ts t2
+ | Abs (s, T, t') => Abs (s, T, polish_funs (T :: Ts) t')
+ | Const (s, Type ("fun", [T1, T2])) =>
+ if s = opt_flag orelse s = non_opt_flag then
+ Abs ("x", T1, Const (unknown, T2))
+ else
+ t
+ | t => t
(* bool -> typ -> typ -> typ -> term list -> term list -> term *)
fun make_fun maybe_opt T1 T2 T' ts1 ts2 =
ts1 ~~ ts2 |> sort (nice_term_ord o pairself fst)
@@ -371,7 +382,7 @@
HOLogic.mk_number nat_T j
else case datatype_spec datatypes T of
NONE => atom for_auto T j
- | SOME {shallow = true, ...} => atom for_auto T j
+ | SOME {deep = false, ...} => atom for_auto T j
| SOME {co, constrs, ...} =>
let
(* styp -> int list *)
@@ -437,13 +448,16 @@
| n2 => Const (@{const_name Algebras.divide},
num_T --> num_T --> num_T)
$ mk_num n1 $ mk_num n2)
- | _ => raise TERM ("Nitpick_Model.reconstruct_term.term_\
- \for_atom (Abs_Frac)", ts)
+ | _ => raise TERM ("Nitpick_Model.reconstruct_term.\
+ \term_for_atom (Abs_Frac)", ts)
end
else if not for_auto andalso
(is_abs_fun thy constr_x orelse
constr_s = @{const_name Quot}) then
Const (abs_name, constr_T) $ the_single ts
+ else if not for_auto andalso
+ constr_s = @{const_name NonStd} then
+ Const (fst (dest_Const (the_single ts)), T)
else
list_comb (Const constr_x, ts)
in
@@ -509,8 +523,7 @@
Refute.string_of_typ T ^ " " ^ string_for_rep R ^ " " ^
string_of_int (length jss))
in
- (not for_auto ? setify_mapify_funs []) o unbit_and_unbox_term
- oooo term_for_rep []
+ (not for_auto ? polish_funs []) o unbit_and_unbox_term oooo term_for_rep []
end
(* scope -> nut list -> nut NameTable.table -> KK.raw_bound list -> nut
@@ -522,8 +535,7 @@
(rep_of name)
end
-(* Proof.context
- -> (string * string * string * string * string) * Proof.context *)
+(* Proof.context -> (string * string * string * string) * Proof.context *)
fun add_wacky_syntax ctxt =
let
(* term -> string *)
@@ -573,13 +585,13 @@
-> nut list -> nut list -> nut NameTable.table -> KK.raw_bound list
-> Pretty.T * bool *)
fun reconstruct_hol_model {show_skolems, show_datatypes, show_consts}
- ({ext_ctxt as {thy, ctxt, max_bisim_depth, boxes, wfs, user_axioms,
- debug, binary_ints, destroy_constrs, specialize,
- skolemize, star_linear_preds, uncurry, fast_descrs,
- tac_timeout, evals, case_names, def_table, nondef_table,
- user_nondefs, simp_table, psimp_table, intro_table,
- ground_thm_table, ersatz_table, skolems, special_funs,
- unrolled_preds, wf_cache, constr_cache},
+ ({ext_ctxt as {thy, ctxt, max_bisim_depth, boxes, stds, wfs,
+ user_axioms, debug, binary_ints, destroy_constrs,
+ specialize, skolemize, star_linear_preds, uncurry,
+ fast_descrs, tac_timeout, evals, case_names, def_table,
+ nondef_table, user_nondefs, simp_table, psimp_table,
+ intro_table, ground_thm_table, ersatz_table, skolems,
+ special_funs, unrolled_preds, wf_cache, constr_cache},
card_assigns, bits, bisim_depth, datatypes, ofs} : scope)
formats all_frees free_names sel_names nonsel_names rel_table bounds =
let
@@ -587,7 +599,7 @@
add_wacky_syntax ctxt
val ext_ctxt =
{thy = thy, ctxt = ctxt, max_bisim_depth = max_bisim_depth, boxes = boxes,
- wfs = wfs, user_axioms = user_axioms, debug = debug,
+ stds = stds, wfs = wfs, user_axioms = user_axioms, debug = debug,
binary_ints = binary_ints, destroy_constrs = destroy_constrs,
specialize = specialize, skolemize = skolemize,
star_linear_preds = star_linear_preds, uncurry = uncurry,
@@ -650,16 +662,15 @@
Pretty.block (Pretty.breaks
[Syntax.pretty_typ ctxt (unbit_and_unbox_type typ), Pretty.str "=",
Pretty.enum "," "{" "}"
- (map (Syntax.pretty_term ctxt) (all_values_of_type card typ)
- @ (if complete then [] else [Pretty.str unrep]))])
+ (map (Syntax.pretty_term ctxt) (all_values_of_type card typ) @
+ (if complete then [] else [Pretty.str unrep]))])
(* typ -> dtype_spec list *)
fun integer_datatype T =
[{typ = T, card = card_of_type card_assigns T, co = false,
- complete = false, concrete = true, shallow = false, constrs = []}]
+ complete = false, concrete = true, deep = true, constrs = []}]
handle TYPE ("Nitpick_HOL.card_of_type", _, _) => []
val (codatatypes, datatypes) =
- datatypes |> filter_out #shallow
- |> List.partition #co
+ datatypes |> filter #deep |> List.partition #co
||> append (integer_datatype nat_T @ integer_datatype int_T)
val block_of_datatypes =
if show_datatypes andalso not (null datatypes) then
--- a/src/HOL/Tools/Nitpick/nitpick_mono.ML Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_mono.ML Tue Feb 02 18:16:48 2010 +0100
@@ -1,16 +1,17 @@
(* Title: HOL/Tools/Nitpick/nitpick_mono.ML
Author: Jasmin Blanchette, TU Muenchen
- Copyright 2009
+ Copyright 2009, 2010
Monotonicity predicate for higher-order logic.
*)
signature NITPICK_MONO =
sig
+ datatype sign = Plus | Minus
type extended_context = Nitpick_HOL.extended_context
val formulas_monotonic :
- extended_context -> typ -> term list -> term list -> term -> bool
+ extended_context -> typ -> sign -> term list -> term list -> term -> bool
end;
structure Nitpick_Mono : NITPICK_MONO =
@@ -21,7 +22,7 @@
type var = int
-datatype sign = Pos | Neg
+datatype sign = Plus | Minus
datatype sign_atom = S of sign | V of var
type literal = var * sign
@@ -54,13 +55,13 @@
if null xs then "" else "\<^bsub>" ^ string_for_vars sep xs ^ "\<^esub>"
(* sign -> string *)
-fun string_for_sign Pos = "+"
- | string_for_sign Neg = "-"
+fun string_for_sign Plus = "+"
+ | string_for_sign Minus = "-"
(* sign -> sign -> sign *)
-fun xor sn1 sn2 = if sn1 = sn2 then Pos else Neg
+fun xor sn1 sn2 = if sn1 = sn2 then Plus else Minus
(* sign -> sign *)
-val negate = xor Neg
+val negate = xor Minus
(* sign_atom -> string *)
fun string_for_sign_atom (S sn) = string_for_sign sn
@@ -152,7 +153,7 @@
(* string * typ list -> ctype list -> ctype *)
fun constr_ctype_for_binders z Cs =
- fold_rev (fn C => curry3 CFun C (S Neg)) Cs (CRec z)
+ fold_rev (fn C => curry3 CFun C (S Minus)) Cs (CRec z)
(* ((string * typ list) * ctype) list -> ctype list -> ctype -> ctype *)
fun repair_ctype _ _ CAlpha = CAlpha
@@ -199,7 +200,7 @@
exists_alpha_sub_ctype_fresh C1 then
V (Unsynchronized.inc max_fresh)
else
- S Neg
+ S Minus
in CFun (C1, a, C2) end
(* typ -> ctype *)
and do_type T =
@@ -252,13 +253,13 @@
fun prodC_factors (CPair (C1, C2)) = maps prodC_factors [C1, C2]
| prodC_factors C = [C]
(* ctype -> ctype list * ctype *)
-fun curried_strip_ctype (CFun (C1, S Neg, C2)) =
+fun curried_strip_ctype (CFun (C1, S Minus, C2)) =
curried_strip_ctype C2 |>> append (prodC_factors C1)
| curried_strip_ctype C = ([], C)
(* string -> ctype -> ctype *)
fun sel_ctype_from_constr_ctype s C =
let val (arg_Cs, dataC) = curried_strip_ctype C in
- CFun (dataC, S Neg,
+ CFun (dataC, S Minus,
case sel_no_from_name s of ~1 => bool_C | n => nth arg_Cs n)
end
@@ -268,8 +269,13 @@
if could_exist_alpha_sub_ctype thy alpha_T T then
case AList.lookup (op =) (!constr_cache) x of
SOME C => C
- | NONE => (fresh_ctype_for_type cdata (body_type T);
- AList.lookup (op =) (!constr_cache) x |> the)
+ | NONE => if T = alpha_T then
+ let val C = fresh_ctype_for_type cdata T in
+ (Unsynchronized.change constr_cache (cons (x, C)); C)
+ end
+ else
+ (fresh_ctype_for_type cdata (body_type T);
+ AList.lookup (op =) (!constr_cache) x |> the)
else
fresh_ctype_for_type cdata T
fun ctype_for_sel (cdata as {ext_ctxt, ...}) (x as (s, _)) =
@@ -332,9 +338,9 @@
| _ => do_sign_atom_comp Eq [] a2 a1 accum)
| do_sign_atom_comp Leq [] a1 a2 (accum as (lits, comps)) =
(case (a1, a2) of
- (_, S Neg) => SOME accum
- | (S Pos, _) => SOME accum
- | (S Neg, S Pos) => NONE
+ (_, S Minus) => SOME accum
+ | (S Plus, _) => SOME accum
+ | (S Minus, S Plus) => NONE
| (V _, V _) => SOME (lits, insert (op =) (a1, a2, Leq, []) comps)
| _ => do_sign_atom_comp Eq [] a1 a2 accum)
| do_sign_atom_comp cmp xs a1 a2 (accum as (lits, comps)) =
@@ -354,8 +360,8 @@
accum |> do_sign_atom_comp Leq xs a1 a2
|> do_ctype_comp Leq xs C21 C11
|> (case a2 of
- S Neg => I
- | S Pos => do_ctype_comp Leq xs C11 C21
+ S Minus => I
+ | S Plus => do_ctype_comp Leq xs C11 C21
| V x => do_ctype_comp Leq (x :: xs) C11 C21)
else
SOME accum)
@@ -386,29 +392,33 @@
(* sign -> sign_expr -> ctype -> (literal list * sign_expr list) option
-> (literal list * sign_expr list) option *)
fun do_notin_ctype_fv _ _ _ NONE = NONE
- | do_notin_ctype_fv Neg _ CAlpha accum = accum
- | do_notin_ctype_fv Pos [] CAlpha _ = NONE
- | do_notin_ctype_fv Pos [(x, sn)] CAlpha (SOME (lits, sexps)) =
+ | do_notin_ctype_fv Minus _ CAlpha accum = accum
+ | do_notin_ctype_fv Plus [] CAlpha _ = NONE
+ | do_notin_ctype_fv Plus [(x, sn)] CAlpha (SOME (lits, sexps)) =
SOME lits |> do_literal (x, sn) |> Option.map (rpair sexps)
- | do_notin_ctype_fv Pos sexp CAlpha (SOME (lits, sexps)) =
+ | do_notin_ctype_fv Plus sexp CAlpha (SOME (lits, sexps)) =
SOME (lits, insert (op =) sexp sexps)
| do_notin_ctype_fv sn sexp (CFun (C1, S sn', C2)) accum =
- accum |> (if sn' = Pos andalso sn = Pos then do_notin_ctype_fv Pos sexp C1
- else I)
- |> (if sn' = Neg orelse sn = Pos then do_notin_ctype_fv Neg sexp C1
- else I)
+ accum |> (if sn' = Plus andalso sn = Plus then
+ do_notin_ctype_fv Plus sexp C1
+ else
+ I)
+ |> (if sn' = Minus orelse sn = Plus then
+ do_notin_ctype_fv Minus sexp C1
+ else
+ I)
|> do_notin_ctype_fv sn sexp C2
- | do_notin_ctype_fv Pos sexp (CFun (C1, V x, C2)) accum =
- accum |> (case do_literal (x, Neg) (SOME sexp) of
+ | do_notin_ctype_fv Plus sexp (CFun (C1, V x, C2)) accum =
+ accum |> (case do_literal (x, Minus) (SOME sexp) of
NONE => I
- | SOME sexp' => do_notin_ctype_fv Pos sexp' C1)
- |> do_notin_ctype_fv Neg sexp C1
- |> do_notin_ctype_fv Pos sexp C2
- | do_notin_ctype_fv Neg sexp (CFun (C1, V x, C2)) accum =
- accum |> (case do_literal (x, Pos) (SOME sexp) of
+ | SOME sexp' => do_notin_ctype_fv Plus sexp' C1)
+ |> do_notin_ctype_fv Minus sexp C1
+ |> do_notin_ctype_fv Plus sexp C2
+ | do_notin_ctype_fv Minus sexp (CFun (C1, V x, C2)) accum =
+ accum |> (case do_literal (x, Plus) (SOME sexp) of
NONE => I
- | SOME sexp' => do_notin_ctype_fv Pos sexp' C1)
- |> do_notin_ctype_fv Neg sexp C2
+ | SOME sexp' => do_notin_ctype_fv Plus sexp' C1)
+ |> do_notin_ctype_fv Minus sexp C2
| do_notin_ctype_fv sn sexp (CPair (C1, C2)) accum =
accum |> fold (do_notin_ctype_fv sn sexp) [C1, C2]
| do_notin_ctype_fv sn sexp (CType (_, Cs)) accum =
@@ -420,14 +430,14 @@
fun add_notin_ctype_fv _ _ UnsolvableCSet = UnsolvableCSet
| add_notin_ctype_fv sn C (CSet (lits, comps, sexps)) =
(print_g ("*** Add " ^ string_for_ctype C ^ " is right-" ^
- (case sn of Neg => "unique" | Pos => "total") ^ ".");
+ (case sn of Minus => "unique" | Plus => "total") ^ ".");
case do_notin_ctype_fv sn [] C (SOME (lits, sexps)) of
NONE => (print_g "**** Unsolvable"; UnsolvableCSet)
| SOME (lits, sexps) => CSet (lits, comps, sexps))
(* ctype -> constraint_set -> constraint_set *)
-val add_ctype_is_right_unique = add_notin_ctype_fv Neg
-val add_ctype_is_right_total = add_notin_ctype_fv Pos
+val add_ctype_is_right_unique = add_notin_ctype_fv Minus
+val add_ctype_is_right_total = add_notin_ctype_fv Plus
(* constraint_set -> constraint_set -> constraint_set *)
fun unite (CSet (lits1, comps1, sexps1)) (CSet (lits2, comps2, sexps2)) =
@@ -437,11 +447,11 @@
| unite _ _ = UnsolvableCSet
(* sign -> bool *)
-fun bool_from_sign Pos = false
- | bool_from_sign Neg = true
+fun bool_from_sign Plus = false
+ | bool_from_sign Minus = true
(* bool -> sign *)
-fun sign_from_bool false = Pos
- | sign_from_bool true = Neg
+fun sign_from_bool false = Plus
+ | sign_from_bool true = Minus
(* literal -> PropLogic.prop_formula *)
fun prop_for_literal (x, sn) =
@@ -460,10 +470,10 @@
PropLogic.SAnd (prop_for_comp (a1, a2, Leq, []),
prop_for_comp (a2, a1, Leq, []))
| prop_for_comp (a1, a2, Leq, []) =
- PropLogic.SOr (prop_for_sign_atom_eq (a1, Pos),
- prop_for_sign_atom_eq (a2, Neg))
+ PropLogic.SOr (prop_for_sign_atom_eq (a1, Plus),
+ prop_for_sign_atom_eq (a2, Minus))
| prop_for_comp (a1, a2, cmp, xs) =
- PropLogic.SOr (prop_for_exists_eq xs Neg, prop_for_comp (a1, a2, cmp, []))
+ PropLogic.SOr (prop_for_exists_eq xs Minus, prop_for_comp (a1, a2, cmp, []))
(* var -> (int -> bool option) -> literal list -> literal list *)
fun literals_from_assignments max_var assigns lits =
@@ -491,7 +501,7 @@
(* literal list -> unit *)
fun print_solution lits =
- let val (pos, neg) = List.partition (curry (op =) Pos o snd) lits in
+ let val (pos, neg) = List.partition (curry (op =) Plus o snd) lits in
print_g ("*** Solution:\n" ^
"+: " ^ commas (map (string_for_var o fst) pos) ^ "\n" ^
"-: " ^ commas (map (string_for_var o fst) neg))
@@ -523,7 +533,7 @@
type ctype_context =
{bounds: ctype list,
frees: (styp * ctype) list,
- consts: (styp * ctype_schema) list}
+ consts: (styp * ctype) list}
type accumulator = ctype_context * constraint_set
@@ -546,20 +556,20 @@
val ctype_for = fresh_ctype_for_type cdata
(* ctype -> ctype *)
fun pos_set_ctype_for_dom C =
- CFun (C, S (if exists_alpha_sub_ctype C then Pos else Neg), bool_C)
+ CFun (C, S (if exists_alpha_sub_ctype C then Plus else Minus), bool_C)
(* typ -> accumulator -> ctype * accumulator *)
fun do_quantifier T (gamma, cset) =
let
val abs_C = ctype_for (domain_type (domain_type T))
val body_C = ctype_for (range_type T)
in
- (CFun (CFun (abs_C, S Neg, body_C), S Neg, body_C),
+ (CFun (CFun (abs_C, S Minus, body_C), S Minus, body_C),
(gamma, cset |> add_ctype_is_right_total abs_C))
end
fun do_equals T (gamma, cset) =
let val C = ctype_for (domain_type T) in
- (CFun (C, S Neg, CFun (C, V (Unsynchronized.inc max_fresh),
- ctype_for (nth_range_type 2 T))),
+ (CFun (C, S Minus, CFun (C, V (Unsynchronized.inc max_fresh),
+ ctype_for (nth_range_type 2 T))),
(gamma, cset |> add_ctype_is_right_unique C))
end
fun do_robust_set_operation T (gamma, cset) =
@@ -569,7 +579,7 @@
val C2 = ctype_for set_T
val C3 = ctype_for set_T
in
- (CFun (C1, S Neg, CFun (C2, S Neg, C3)),
+ (CFun (C1, S Minus, CFun (C2, S Minus, C3)),
(gamma, cset |> add_is_sub_ctype C1 C3 |> add_is_sub_ctype C2 C3))
end
fun do_fragile_set_operation T (gamma, cset) =
@@ -579,7 +589,7 @@
(* typ -> ctype *)
fun custom_ctype_for (T as Type ("fun", [T1, T2])) =
if T = set_T then set_C
- else CFun (custom_ctype_for T1, S Neg, custom_ctype_for T2)
+ else CFun (custom_ctype_for T1, S Minus, custom_ctype_for T2)
| custom_ctype_for T = ctype_for T
in
(custom_ctype_for T, (gamma, cset |> add_ctype_is_right_unique set_C))
@@ -588,13 +598,13 @@
fun do_pair_constr T accum =
case ctype_for (nth_range_type 2 T) of
C as CPair (a_C, b_C) =>
- (CFun (a_C, S Neg, CFun (b_C, S Neg, C)), accum)
+ (CFun (a_C, S Minus, CFun (b_C, S Minus, C)), accum)
| C => raise CTYPE ("Nitpick_Mono.consider_term.do_pair_constr", [C])
(* int -> typ -> accumulator -> ctype * accumulator *)
fun do_nth_pair_sel n T =
case ctype_for (domain_type T) of
C as CPair (a_C, b_C) =>
- pair (CFun (C, S Neg, if n = 0 then a_C else b_C))
+ pair (CFun (C, S Minus, if n = 0 then a_C else b_C))
| C => raise CTYPE ("Nitpick_Mono.consider_term.do_nth_pair_sel", [C])
val unsolvable = (CType ("unsolvable", []), unsolvable_accum)
(* typ -> term -> accumulator -> ctype * accumulator *)
@@ -613,7 +623,7 @@
(case t of
Const (x as (s, T)) =>
(case AList.lookup (op =) consts x of
- SOME (C, cset') => (C, (gamma, cset |> unite cset'))
+ SOME C => (C, accum)
| NONE =>
if not (could_exist_alpha_subtype alpha_T T) then
(ctype_for T, accum)
@@ -627,12 +637,12 @@
| @{const_name Eps} => (print_g "*** Eps"; unsolvable)
| @{const_name If} =>
do_robust_set_operation (range_type T) accum
- |>> curry3 CFun bool_C (S Neg)
+ |>> curry3 CFun bool_C (S Minus)
| @{const_name Pair} => do_pair_constr T accum
| @{const_name fst} => do_nth_pair_sel 0 T accum
| @{const_name snd} => do_nth_pair_sel 1 T accum
| @{const_name Id} =>
- (CFun (ctype_for (domain_type T), S Neg, bool_C), accum)
+ (CFun (ctype_for (domain_type T), S Minus, bool_C), accum)
| @{const_name insert} =>
let
val set_T = domain_type (range_type T)
@@ -641,7 +651,7 @@
val C2 = ctype_for set_T
val C3 = ctype_for set_T
in
- (CFun (C1, S Neg, CFun (C2, S Neg, C3)),
+ (CFun (C1, S Minus, CFun (C2, S Minus, C3)),
(gamma, cset |> add_ctype_is_right_unique C1
|> add_is_sub_ctype C1' C3
|> add_is_sub_ctype C2 C3))
@@ -654,7 +664,7 @@
CFun (ctype_for (domain_type T), V x, bool_C)
val ab_set_C = domain_type T |> ctype_for_set
val ba_set_C = range_type T |> ctype_for_set
- in (CFun (ab_set_C, S Neg, ba_set_C), accum) end
+ in (CFun (ab_set_C, S Minus, ba_set_C), accum) end
| @{const_name trancl} => do_fragile_set_operation T accum
| @{const_name rtrancl} => (print_g "*** rtrancl"; unsolvable)
| @{const_name lower_semilattice_fun_inst.inf_fun} =>
@@ -663,7 +673,7 @@
do_robust_set_operation T accum
| @{const_name finite} =>
let val C1 = ctype_for (domain_type (domain_type T)) in
- (CFun (pos_set_ctype_for_dom C1, S Neg, bool_C), accum)
+ (CFun (pos_set_ctype_for_dom C1, S Minus, bool_C), accum)
end
| @{const_name rel_comp} =>
let
@@ -675,7 +685,7 @@
val ab_set_C = domain_type (range_type T) |> ctype_for_set
val ac_set_C = nth_range_type 2 T |> ctype_for_set
in
- (CFun (bc_set_C, S Neg, CFun (ab_set_C, S Neg, ac_set_C)),
+ (CFun (bc_set_C, S Minus, CFun (ab_set_C, S Minus, ac_set_C)),
accum)
end
| @{const_name image} =>
@@ -683,8 +693,8 @@
val a_C = ctype_for (domain_type (domain_type T))
val b_C = ctype_for (range_type (domain_type T))
in
- (CFun (CFun (a_C, S Neg, b_C), S Neg,
- CFun (pos_set_ctype_for_dom a_C, S Neg,
+ (CFun (CFun (a_C, S Minus, b_C), S Minus,
+ CFun (pos_set_ctype_for_dom a_C, S Minus,
pos_set_ctype_for_dom b_C)), accum)
end
| @{const_name Sigma} =>
@@ -698,11 +708,11 @@
val b_set_C = ctype_for_set (range_type (domain_type
(range_type T)))
val a_set_C = ctype_for_set a_set_T
- val a_to_b_set_C = CFun (a_C, S Neg, b_set_C)
+ val a_to_b_set_C = CFun (a_C, S Minus, b_set_C)
val ab_set_C = ctype_for_set (nth_range_type 2 T)
in
- (CFun (a_set_C, S Neg, CFun (a_to_b_set_C, S Neg, ab_set_C)),
- accum)
+ (CFun (a_set_C, S Minus,
+ CFun (a_to_b_set_C, S Minus, ab_set_C)), accum)
end
| @{const_name minus_fun_inst.minus_fun} =>
let
@@ -710,8 +720,8 @@
val left_set_C = ctype_for set_T
val right_set_C = ctype_for set_T
in
- (CFun (left_set_C, S Neg,
- CFun (right_set_C, S Neg, left_set_C)),
+ (CFun (left_set_C, S Minus,
+ CFun (right_set_C, S Minus, left_set_C)),
(gamma, cset |> add_ctype_is_right_unique right_set_C
|> add_is_sub_ctype right_set_C left_set_C))
end
@@ -721,15 +731,15 @@
let
val a_C = ctype_for (domain_type (domain_type T))
val a_set_C = pos_set_ctype_for_dom a_C
- in (CFun (a_set_C, S Neg, a_C), accum) end
+ in (CFun (a_set_C, S Minus, a_C), accum) end
| @{const_name FunBox} =>
let val dom_C = ctype_for (domain_type T) in
- (CFun (dom_C, S Neg, dom_C), accum)
+ (CFun (dom_C, S Minus, dom_C), accum)
end
| _ => if is_sel s then
if constr_name_for_sel_like s = @{const_name FunBox} then
let val dom_C = ctype_for (domain_type T) in
- (CFun (dom_C, S Neg, dom_C), accum)
+ (CFun (dom_C, S Minus, dom_C), accum)
end
else
(ctype_for_sel cdata x, accum)
@@ -740,7 +750,10 @@
SOME t' => do_term t' accum
| NONE => (print_g ("*** built-in " ^ s); unsolvable)
else
- (ctype_for T, accum))
+ let val C = ctype_for T in
+ (C, ({bounds = bounds, frees = frees,
+ consts = (x, C) :: consts}, cset))
+ end)
| Free (x as (_, T)) =>
(case AList.lookup (op =) frees x of
SOME C => (C, accum)
@@ -756,7 +769,7 @@
let
val C = ctype_for T
val (C', accum) = do_term t' (accum |>> push_bound C)
- in (CFun (C, S Neg, C'), accum |>> pop_bound) end
+ in (CFun (C, S Minus, C'), accum |>> pop_bound) end
| Const (@{const_name All}, _)
$ Abs (_, T', @{const "op -->"} $ (t1 $ Bound 0) $ t2) =>
do_bounded_quantifier T' t1 t2 accum
@@ -802,7 +815,7 @@
fun do_quantifier quant_s abs_T body_t =
let
val abs_C = ctype_for abs_T
- val side_cond = ((sn = Neg) = (quant_s = @{const_name Ex}))
+ val side_cond = ((sn = Minus) = (quant_s = @{const_name Ex}))
val cset = cset |> side_cond ? add_ctype_is_right_total abs_C
in
(gamma |> push_bound abs_C, cset) |> do_co_formula body_t
@@ -815,11 +828,11 @@
(* term -> term -> accumulator *)
fun do_equals t1 t2 =
case sn of
- Pos => do_term t accum |> snd
- | Neg => let
- val (C1, accum) = do_term t1 accum
- val (C2, accum) = do_term t2 accum
- in accum ||> add_ctypes_equal C1 C2 end
+ Plus => do_term t accum |> snd
+ | Minus => let
+ val (C1, accum) = do_term t1 accum
+ val (C2, accum) = do_term t2 accum
+ in accum ||> add_ctypes_equal C1 C2 end
in
case t of
Const (s0 as @{const_name all}, _) $ Abs (_, T1, t1) =>
@@ -876,7 +889,7 @@
(* cdata -> term -> accumulator -> accumulator *)
fun consider_definitional_axiom (cdata as {ext_ctxt as {thy, ...}, ...}) t =
if not (is_constr_pattern_formula thy t) then
- consider_nondefinitional_axiom cdata Pos t
+ consider_nondefinitional_axiom cdata Plus t
else if is_harmless_axiom t then
I
else
@@ -921,11 +934,11 @@
(* theory -> literal list -> ctype_context -> unit *)
fun print_ctype_context ctxt lits ({frees, consts, ...} : ctype_context) =
map (fn (x, C) => string_for_ctype_of_term ctxt lits (Free x) C) frees @
- map (fn (x, (C, _)) => string_for_ctype_of_term ctxt lits (Const x) C) consts
+ map (fn (x, C) => string_for_ctype_of_term ctxt lits (Const x) C) consts
|> cat_lines |> print_g
-(* extended_context -> typ -> term list -> term list -> term -> bool *)
-fun formulas_monotonic (ext_ctxt as {ctxt, ...}) alpha_T def_ts nondef_ts
+(* extended_context -> typ -> sign -> term list -> term list -> term -> bool *)
+fun formulas_monotonic (ext_ctxt as {ctxt, ...}) alpha_T sn def_ts nondef_ts
core_t =
let
val _ = print_g ("****** " ^ string_for_ctype CAlpha ^ " is " ^
@@ -934,8 +947,8 @@
val (gamma, cset) =
(initial_gamma, slack)
|> fold (consider_definitional_axiom cdata) def_ts
- |> fold (consider_nondefinitional_axiom cdata Pos) nondef_ts
- |> consider_general_formula cdata Pos core_t
+ |> fold (consider_nondefinitional_axiom cdata Plus) nondef_ts
+ |> consider_general_formula cdata sn core_t
in
case solve (!max_fresh) cset of
SOME lits => (print_ctype_context ctxt lits gamma; true)
--- a/src/HOL/Tools/Nitpick/nitpick_nut.ML Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_nut.ML Tue Feb 02 18:16:48 2010 +0100
@@ -1,6 +1,6 @@
(* Title: HOL/Tools/Nitpick/nitpick_nut.ML
Author: Jasmin Blanchette, TU Muenchen
- Copyright 2008, 2009
+ Copyright 2008, 2009, 2010
Nitpick underlying terms (nuts).
*)
@@ -766,7 +766,7 @@
(~1 upto num_sels_for_constr_type T - 1)
(* scope -> dtype_spec -> nut list * rep NameTable.table
-> nut list * rep NameTable.table *)
-fun choose_rep_for_sels_of_datatype _ ({shallow = true, ...} : dtype_spec) = I
+fun choose_rep_for_sels_of_datatype _ ({deep = false, ...} : dtype_spec) = I
| choose_rep_for_sels_of_datatype scope {constrs, ...} =
fold_rev (choose_rep_for_sels_for_constr scope o #const) constrs
(* scope -> rep NameTable.table -> nut list * rep NameTable.table *)
--- a/src/HOL/Tools/Nitpick/nitpick_peephole.ML Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_peephole.ML Tue Feb 02 18:16:48 2010 +0100
@@ -1,6 +1,6 @@
(* Title: HOL/Tools/Nitpick/nitpick_peephole.ML
Author: Jasmin Blanchette, TU Muenchen
- Copyright 2008, 2009
+ Copyright 2008, 2009, 2010
Peephole optimizer for Nitpick.
*)
--- a/src/HOL/Tools/Nitpick/nitpick_rep.ML Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_rep.ML Tue Feb 02 18:16:48 2010 +0100
@@ -1,6 +1,6 @@
(* Title: HOL/Tools/Nitpick/nitpick_rep.ML
Author: Jasmin Blanchette, TU Muenchen
- Copyright 2008, 2009
+ Copyright 2008, 2009, 2010
Kodkod representations of Nitpick terms.
*)
--- a/src/HOL/Tools/Nitpick/nitpick_scope.ML Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_scope.ML Tue Feb 02 18:16:48 2010 +0100
@@ -1,6 +1,6 @@
(* Title: HOL/Tools/Nitpick/nitpick_scope.ML
Author: Jasmin Blanchette, TU Muenchen
- Copyright 2008, 2009
+ Copyright 2008, 2009, 2010
Scope enumerator for Nitpick.
*)
@@ -24,7 +24,7 @@
co: bool,
complete: bool,
concrete: bool,
- shallow: bool,
+ deep: bool,
constrs: constr_spec list}
type scope = {
@@ -73,7 +73,7 @@
co: bool,
complete: bool,
concrete: bool,
- shallow: bool,
+ deep: bool,
constrs: constr_spec list}
type scope = {
@@ -134,7 +134,7 @@
fun quintuple_for_scope quote ({ext_ctxt as {thy, ctxt, ...}, card_assigns,
bits, bisim_depth, datatypes, ...} : scope) =
let
- val boring_Ts = [@{typ unsigned_bit}, @{typ signed_bit},
+ val boring_Ts = [@{typ unsigned_bit}, @{typ signed_bit}, @{typ \<xi>},
@{typ bisim_iterator}]
val (iter_assigns, card_assigns) =
card_assigns |> filter_out (member (op =) boring_Ts o fst)
@@ -429,9 +429,10 @@
else if not co andalso num_self_recs > 0 then
if not self_rec andalso num_non_self_recs = 1 andalso
domain_card 2 card_assigns T = 1 then
- {delta = 0, epsilon = 1, exclusive = (s = @{const_name Nil}),
+ {delta = 0, epsilon = 1,
+ exclusive = (s = @{const_name Nil} andalso length constrs = 2),
total = true}
- else if s = @{const_name Cons} then
+ else if s = @{const_name Cons} andalso length constrs = 2 then
{delta = 1, epsilon = card, exclusive = true, total = false}
else
{delta = 0, epsilon = card, exclusive = false, total = false}
@@ -455,10 +456,10 @@
end
(* extended_context -> typ list -> scope_desc -> typ * int -> dtype_spec *)
-fun datatype_spec_from_scope_descriptor (ext_ctxt as {thy, ...}) shallow_dataTs
+fun datatype_spec_from_scope_descriptor (ext_ctxt as {thy, ...}) deep_dataTs
(desc as (card_assigns, _)) (T, card) =
let
- val shallow = member (op =) shallow_dataTs T
+ val deep = member (op =) deep_dataTs T
val co = is_codatatype thy T
val xs = boxed_datatype_constrs ext_ctxt T
val self_recs = map (is_self_recursive_constr_type o snd) xs
@@ -475,7 +476,7 @@
num_non_self_recs) (self_recs ~~ xs) []
in
{typ = T, card = card, co = co, complete = complete, concrete = concrete,
- shallow = shallow, constrs = constrs}
+ deep = deep, constrs = constrs}
end
(* int -> int *)
@@ -487,11 +488,11 @@
(map snd card_assigns @ map snd max_assigns) 0)
(* extended_context -> int -> typ list -> scope_desc -> scope *)
-fun scope_from_descriptor (ext_ctxt as {thy, ...}) sym_break shallow_dataTs
+fun scope_from_descriptor (ext_ctxt as {thy, ...}) sym_break deep_dataTs
(desc as (card_assigns, _)) =
let
val datatypes =
- map (datatype_spec_from_scope_descriptor ext_ctxt shallow_dataTs desc)
+ map (datatype_spec_from_scope_descriptor ext_ctxt deep_dataTs desc)
(filter (is_datatype thy o fst) card_assigns)
val bits = card_of_type card_assigns @{typ signed_bit} - 1
handle TYPE ("Nitpick_HOL.card_of_type", _, _) =>
@@ -524,8 +525,7 @@
-> (styp option * int list) list -> (styp option * int list) list -> int list
-> typ list -> typ list -> typ list -> int * scope list *)
fun all_scopes (ext_ctxt as {thy, ...}) sym_break cards_assigns maxes_assigns
- iters_assigns bitss bisim_depths mono_Ts nonmono_Ts
- shallow_dataTs =
+ iters_assigns bitss bisim_depths mono_Ts nonmono_Ts deep_dataTs =
let
val cards_assigns = repair_cards_assigns_wrt_boxing thy mono_Ts
cards_assigns
@@ -540,7 +540,7 @@
in
(length all - length head,
descs |> length descs <= distinct_threshold ? distinct (op =)
- |> map (scope_from_descriptor ext_ctxt sym_break shallow_dataTs))
+ |> map (scope_from_descriptor ext_ctxt sym_break deep_dataTs))
end
end;
--- a/src/HOL/Tools/Nitpick/nitpick_tests.ML Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_tests.ML Tue Feb 02 18:16:48 2010 +0100
@@ -1,6 +1,6 @@
(* Title: HOL/Tools/Nitpick/nitpick_tests.ML
Author: Jasmin Blanchette, TU Muenchen
- Copyright 2008, 2009
+ Copyright 2008, 2009, 2010
Unit tests for Nitpick.
*)
--- a/src/HOL/Tools/Nitpick/nitpick_util.ML Sun Jan 31 15:22:40 2010 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_util.ML Tue Feb 02 18:16:48 2010 +0100
@@ -1,6 +1,6 @@
(* Title: HOL/Tools/Nitpick/nitpick_util.ML
Author: Jasmin Blanchette, TU Muenchen
- Copyright 2008, 2009
+ Copyright 2008, 2009, 2010
General-purpose functions used by the Nitpick modules.
*)
@@ -58,7 +58,7 @@
val setmp_show_all_types : ('a -> 'b) -> 'a -> 'b
val indent_size : int
val pstrs : string -> Pretty.T list
- val plain_string_from_yxml : string -> string
+ val unyxml : string -> string
val maybe_quote : string -> string
end
@@ -278,13 +278,13 @@
fun plain_string_from_xml_tree t =
Buffer.empty |> XML.add_content t |> Buffer.content
(* string -> string *)
-val plain_string_from_yxml = plain_string_from_xml_tree o YXML.parse
+val unyxml = plain_string_from_xml_tree o YXML.parse
(* string -> bool *)
val is_long_identifier = forall Syntax.is_identifier o space_explode "."
(* string -> string *)
fun maybe_quote y =
- let val s = plain_string_from_yxml y in
+ let val s = unyxml y in
y |> (not (is_long_identifier (perhaps (try (unprefix "'")) s)) orelse
OuterKeyword.is_keyword s) ? quote
end
--- a/src/Pure/Isar/proof_context.ML Sun Jan 31 15:22:40 2010 +0100
+++ b/src/Pure/Isar/proof_context.ML Tue Feb 02 18:16:48 2010 +0100
@@ -33,6 +33,7 @@
val set_syntax_mode: Syntax.mode -> Proof.context -> Proof.context
val restore_syntax_mode: Proof.context -> Proof.context -> Proof.context
val facts_of: Proof.context -> Facts.T
+ val cases_of: Proof.context -> (string * (Rule_Cases.T * bool)) list
val transfer_syntax: theory -> Proof.context -> Proof.context
val transfer: theory -> Proof.context -> Proof.context
val theory: (theory -> theory) -> Proof.context -> Proof.context